| author | wenzelm | 
| Thu, 12 Oct 2017 21:22:02 +0200 | |
| changeset 66852 | d20a668b394e | 
| parent 66827 | c94531b5007d | 
| child 66884 | c2128ab11f61 | 
| permissions | -rw-r--r-- | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1 | (* Title: HOL/Analysis/Convex_Euclidean_Space.thy | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 2 | Author: L C Paulson, University of Cambridge | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 3 | Author: Robert Himmelmann, TU Muenchen | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 4 | Author: Bogdan Grechuk, University of Edinburgh | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 5 | Author: Armin Heller, TU Muenchen | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
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: 
63092diff
changeset | 9 | section \<open>Convex sets, functions and related things\<close> | 
| 33175 | 10 | |
| 11 | theory Convex_Euclidean_Space | |
| 44132 | 12 | imports | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: 
66793diff
changeset | 13 | Connected | 
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
66289diff
changeset | 14 | "HOL-Library.Set_Algebras" | 
| 33175 | 15 | begin | 
| 16 | ||
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 17 | lemma swap_continuous: (*move to Topological_Spaces?*) | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
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: 
64394diff
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: 
64394diff
changeset | 20 | proof - | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
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: 
64394diff
changeset | 22 | by auto | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 23 | then show ?thesis | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 24 | apply (rule ssubst) | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 25 | apply (rule continuous_on_compose) | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 26 | apply (simp add: split_def) | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 27 | apply (rule continuous_intros | simp add: assms)+ | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 28 | done | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 29 | qed | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
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: 
63332diff
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: 
63092diff
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: 
50104diff
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: 
50104diff
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: 
50104diff
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: 
50104diff
changeset | 89 | qed | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 90 | |
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
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: 
50104diff
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: 
62381diff
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: 
62381diff
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: 
62381diff
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: 
61609diff
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: 
61609diff
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: 
61609diff
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: 
61609diff
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: 
63967diff
changeset | 178 | subsection \<open>Convexity\<close> | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 179 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 180 | definition convex :: "'a::real_vector set \<Rightarrow> bool" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
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: 
63967diff
changeset | 182 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 183 | lemma convexI: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
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: 
63967diff
changeset | 185 | shows "convex s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 186 | using assms unfolding convex_def by fast | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 187 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 188 | lemma convexD: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
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: 
63967diff
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: 
63967diff
changeset | 191 | using assms unfolding convex_def by fast | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 192 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
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: 
63967diff
changeset | 194 | (is "_ \<longleftrightarrow> ?alt") | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 195 | proof | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 196 | show "convex s" if alt: ?alt | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 197 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 198 |     {
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 199 | fix x y and u v :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 200 | assume mem: "x \<in> s" "y \<in> s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 201 | assume "0 \<le> u" "0 \<le> v" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 202 | moreover | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 203 | assume "u + v = 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 204 | then have "u = 1 - v" by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
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: 
63967diff
changeset | 206 | using alt [rule_format, OF mem] by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 207 | } | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 208 | then show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 209 | unfolding convex_def by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 210 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 211 | show ?alt if "convex s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 212 | using that by (auto simp: convex_def) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 213 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 214 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 215 | lemma convexD_alt: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
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: 
63967diff
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: 
63967diff
changeset | 218 | using assms unfolding convex_alt by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 219 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 220 | lemma mem_convex_alt: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
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: 
63967diff
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: 
63967diff
changeset | 223 | apply (rule convexD) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 224 | using assms | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
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: 
63967diff
changeset | 226 | done | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 227 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 228 | lemma convex_empty[intro,simp]: "convex {}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 229 | unfolding convex_def by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 230 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 231 | lemma convex_singleton[intro,simp]: "convex {a}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 232 | unfolding convex_def by (auto simp: scaleR_left_distrib[symmetric]) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 233 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 234 | lemma convex_UNIV[intro,simp]: "convex UNIV" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 235 | unfolding convex_def by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 236 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
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: 
63967diff
changeset | 238 | unfolding convex_def by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 239 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
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: 
63967diff
changeset | 241 | unfolding convex_def by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 242 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
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: 
63967diff
changeset | 244 | unfolding convex_def by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 245 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
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: 
63967diff
changeset | 247 | unfolding convex_def by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 248 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
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: 
63967diff
changeset | 250 | unfolding convex_def | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 251 | by (auto simp: inner_add intro!: convex_bound_le) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 252 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
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: 
63967diff
changeset | 254 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
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: 
63967diff
changeset | 256 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 257 | show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 258 | unfolding * using convex_halfspace_le[of "-a" "-b"] by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 259 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 260 | |
| 65583 
8d53b3bebab4
Further new material. The simprule status of some exp and ln identities was reverted.
 paulson <lp15@cam.ac.uk> parents: 
65057diff
changeset | 261 | lemma convex_halfspace_abs_le: "convex {x. \<bar>inner a x\<bar> \<le> b}"
 | 
| 
8d53b3bebab4
Further new material. The simprule status of some exp and ln identities was reverted.
 paulson <lp15@cam.ac.uk> parents: 
65057diff
changeset | 262 | proof - | 
| 
8d53b3bebab4
Further new material. The simprule status of some exp and ln identities was reverted.
 paulson <lp15@cam.ac.uk> parents: 
65057diff
changeset | 263 |   have *: "{x. \<bar>inner a x\<bar> \<le> b} = {x. inner a x \<le> b} \<inter> {x. -b \<le> inner a x}"
 | 
| 
8d53b3bebab4
Further new material. The simprule status of some exp and ln identities was reverted.
 paulson <lp15@cam.ac.uk> parents: 
65057diff
changeset | 264 | by auto | 
| 
8d53b3bebab4
Further new material. The simprule status of some exp and ln identities was reverted.
 paulson <lp15@cam.ac.uk> parents: 
65057diff
changeset | 265 | show ?thesis | 
| 
8d53b3bebab4
Further new material. The simprule status of some exp and ln identities was reverted.
 paulson <lp15@cam.ac.uk> parents: 
65057diff
changeset | 266 | unfolding * by (simp add: convex_Int convex_halfspace_ge convex_halfspace_le) | 
| 
8d53b3bebab4
Further new material. The simprule status of some exp and ln identities was reverted.
 paulson <lp15@cam.ac.uk> parents: 
65057diff
changeset | 267 | qed | 
| 
8d53b3bebab4
Further new material. The simprule status of some exp and ln identities was reverted.
 paulson <lp15@cam.ac.uk> parents: 
65057diff
changeset | 268 | |
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 269 | lemma convex_hyperplane: "convex {x. inner a x = b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 270 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 271 |   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: 
63967diff
changeset | 272 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 273 | show ?thesis using convex_halfspace_le convex_halfspace_ge | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 274 | by (auto intro!: convex_Int simp: *) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 275 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 276 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 277 | lemma convex_halfspace_lt: "convex {x. inner a x < b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 278 | unfolding convex_def | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 279 | by (auto simp: convex_bound_lt inner_add) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 280 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 281 | lemma convex_halfspace_gt: "convex {x. inner a x > b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 282 | using convex_halfspace_lt[of "-a" "-b"] by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 283 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 284 | lemma convex_real_interval [iff]: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 285 | fixes a b :: "real" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 286 |   shows "convex {a..}" and "convex {..b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 287 |     and "convex {a<..}" and "convex {..<b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 288 |     and "convex {a..b}" and "convex {a<..b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 289 |     and "convex {a..<b}" and "convex {a<..<b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 290 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 291 |   have "{a..} = {x. a \<le> inner 1 x}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 292 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 293 |   then show 1: "convex {a..}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 294 | by (simp only: convex_halfspace_ge) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 295 |   have "{..b} = {x. inner 1 x \<le> b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 296 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 297 |   then show 2: "convex {..b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 298 | by (simp only: convex_halfspace_le) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 299 |   have "{a<..} = {x. a < inner 1 x}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 300 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 301 |   then show 3: "convex {a<..}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 302 | by (simp only: convex_halfspace_gt) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 303 |   have "{..<b} = {x. inner 1 x < b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 304 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 305 |   then show 4: "convex {..<b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 306 | by (simp only: convex_halfspace_lt) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 307 |   have "{a..b} = {a..} \<inter> {..b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 308 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 309 |   then show "convex {a..b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 310 | by (simp only: convex_Int 1 2) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 311 |   have "{a<..b} = {a<..} \<inter> {..b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 312 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 313 |   then show "convex {a<..b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 314 | by (simp only: convex_Int 3 2) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 315 |   have "{a..<b} = {a..} \<inter> {..<b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 316 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 317 |   then show "convex {a..<b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 318 | by (simp only: convex_Int 1 4) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 319 |   have "{a<..<b} = {a<..} \<inter> {..<b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 320 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 321 |   then show "convex {a<..<b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 322 | by (simp only: convex_Int 3 4) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 323 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 324 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 325 | lemma convex_Reals: "convex \<real>" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 326 | by (simp add: convex_def scaleR_conv_of_real) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 327 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 328 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 329 | subsection \<open>Explicit expressions for convexity in terms of arbitrary sums\<close> | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 330 | |
| 64267 | 331 | lemma convex_sum: | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 332 | fixes C :: "'a::real_vector set" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 333 | assumes "finite s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 334 | and "convex C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 335 | and "(\<Sum> i \<in> s. a i) = 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 336 | assumes "\<And>i. i \<in> s \<Longrightarrow> a i \<ge> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 337 | and "\<And>i. i \<in> s \<Longrightarrow> y i \<in> C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 338 | 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: 
63967diff
changeset | 339 | using assms(1,3,4,5) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 340 | proof (induct arbitrary: a set: finite) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 341 | case empty | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 342 | then show ?case by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 343 | next | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 344 | case (insert i s) note IH = this(3) | 
| 64267 | 345 | have "a i + sum a s = 1" | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 346 | and "0 \<le> a i" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 347 | and "\<forall>j\<in>s. 0 \<le> a j" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 348 | and "y i \<in> C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 349 | and "\<forall>j\<in>s. y j \<in> C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 350 | using insert.hyps(1,2) insert.prems by simp_all | 
| 64267 | 351 | then have "0 \<le> sum a s" | 
| 352 | by (simp add: sum_nonneg) | |
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 353 | have "a i *\<^sub>R y i + (\<Sum>j\<in>s. a j *\<^sub>R y j) \<in> C" | 
| 64267 | 354 | proof (cases "sum a s = 0") | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 355 | case True | 
| 64267 | 356 | 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: 
63967diff
changeset | 357 | by simp | 
| 64267 | 358 | 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: 
63967diff
changeset | 359 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 360 | 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: 
63967diff
changeset | 361 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 362 | next | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 363 | case False | 
| 64267 | 364 | 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: 
63967diff
changeset | 365 | by simp | 
| 64267 | 366 | 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: 
63967diff
changeset | 367 | using \<open>\<forall>j\<in>s. 0 \<le> a j\<close> and \<open>\<forall>j\<in>s. y j \<in> C\<close> | 
| 64267 | 368 | by (simp add: IH sum_divide_distrib [symmetric]) | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 369 | from \<open>convex C\<close> and \<open>y i \<in> C\<close> and this and \<open>0 \<le> a i\<close> | 
| 64267 | 370 | and \<open>0 \<le> sum a s\<close> and \<open>a i + sum a s = 1\<close> | 
| 371 | 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: 
63967diff
changeset | 372 | by (rule convexD) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 373 | then show ?thesis | 
| 64267 | 374 | by (simp add: scaleR_sum_right False) | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 375 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 376 | 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: 
63967diff
changeset | 377 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 378 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 379 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 380 | lemma convex: | 
| 64267 | 381 |   "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)
 | 
| 382 |       \<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: 
63967diff
changeset | 383 | proof safe | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 384 | fix k :: nat | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 385 | fix u :: "nat \<Rightarrow> real" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 386 | fix x | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 387 | assume "convex s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 388 | "\<forall>i. 1 \<le> i \<and> i \<le> k \<longrightarrow> 0 \<le> u i \<and> x i \<in> s" | 
| 64267 | 389 |     "sum u {1..k} = 1"
 | 
| 390 |   with convex_sum[of "{1 .. k}" s] show "(\<Sum>j\<in>{1 .. k}. u j *\<^sub>R x j) \<in> s"
 | |
| 391 | by auto | |
| 392 | next | |
| 393 |   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: 
63967diff
changeset | 394 | \<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: 
63967diff
changeset | 395 |   {
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 396 | fix \<mu> :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 397 | fix x y :: 'a | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 398 | assume xy: "x \<in> s" "y \<in> s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 399 | assume mu: "\<mu> \<ge> 0" "\<mu> \<le> 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 400 | 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: 
63967diff
changeset | 401 | let ?x = "\<lambda>i. if (i :: nat) = 1 then x else y" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 402 |     have "{1 :: nat .. 2} \<inter> - {x. x = 1} = {2}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 403 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 404 |     then have card: "card ({1 :: nat .. 2} \<inter> - {x. x = 1}) = 1"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 405 | by simp | 
| 64267 | 406 |     then have "sum ?u {1 .. 2} = 1"
 | 
| 407 |       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: 
63967diff
changeset | 408 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 409 |     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: 
63967diff
changeset | 410 | using mu xy by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 411 |     have grarr: "(\<Sum>j \<in> {Suc (Suc 0)..2}. ?u j *\<^sub>R ?x j) = (1 - \<mu>) *\<^sub>R y"
 | 
| 64267 | 412 | using sum_head_Suc[of "Suc (Suc 0)" 2 "\<lambda> j. (1 - \<mu>) *\<^sub>R y"] by auto | 
| 413 | 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: 
63967diff
changeset | 414 |     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: 
63967diff
changeset | 415 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 416 | 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: 
63967diff
changeset | 417 | using s by (auto simp: add.commute) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 418 | } | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 419 | then show "convex s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 420 | unfolding convex_alt by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 421 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 422 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 423 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 424 | lemma convex_explicit: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 425 | fixes s :: "'a::real_vector set" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 426 | shows "convex s \<longleftrightarrow> | 
| 64267 | 427 | (\<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: 
63967diff
changeset | 428 | proof safe | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 429 | fix t | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 430 | fix u :: "'a \<Rightarrow> real" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 431 | assume "convex s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 432 | and "finite t" | 
| 64267 | 433 | 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: 
63967diff
changeset | 434 | then show "(\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s" | 
| 64267 | 435 | 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: 
63967diff
changeset | 436 | next | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 437 | assume *: "\<forall>t. \<forall> u. finite t \<and> t \<subseteq> s \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and> | 
| 64267 | 438 | 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: 
63967diff
changeset | 439 | show "convex s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 440 | unfolding convex_alt | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 441 | proof safe | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 442 | fix x y | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 443 | fix \<mu> :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 444 | 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: 
63967diff
changeset | 445 | show "(1 - \<mu>) *\<^sub>R x + \<mu> *\<^sub>R y \<in> s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 446 | proof (cases "x = y") | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 447 | case False | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 448 | then show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 449 |         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: 
63967diff
changeset | 450 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 451 | next | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 452 | case True | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 453 | then show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 454 |         using *[rule_format, of "{x, y}" "\<lambda> z. 1"] **
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 455 | by (auto simp: field_simps real_vector.scale_left_diff_distrib) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 456 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 457 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 458 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 459 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 460 | lemma convex_finite: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 461 | assumes "finite s" | 
| 64267 | 462 | 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: 
63967diff
changeset | 463 | unfolding convex_explicit | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 464 | apply safe | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 465 | 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: 
63967diff
changeset | 466 | subgoal for t u | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 467 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 468 | 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: 
63967diff
changeset | 469 | by simp | 
| 64267 | 470 | 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" | 
| 471 | 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: 
63967diff
changeset | 472 | assume "t \<subseteq> s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 473 | then have "s \<inter> t = t" by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 474 | 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 | 475 | 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: 
63967diff
changeset | 476 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 477 | done | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 478 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 479 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 480 | subsection \<open>Functions that are convex on a set\<close> | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 481 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 482 | 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: 
63967diff
changeset | 483 | where "convex_on s f \<longleftrightarrow> | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 484 | (\<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: 
63967diff
changeset | 485 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 486 | lemma convex_onI [intro?]: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 487 | 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: 
63967diff
changeset | 488 | 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: 
63967diff
changeset | 489 | shows "convex_on A f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 490 | unfolding convex_on_def | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 491 | proof clarify | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 492 | fix x y | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 493 | fix u v :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 494 | 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: 
63967diff
changeset | 495 | from A(5) have [simp]: "v = 1 - u" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 496 | by (simp add: algebra_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 497 | 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: 
63967diff
changeset | 498 | using assms[of u y x] | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 499 | by (cases "u = 0 \<or> u = 1") (auto simp: algebra_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 500 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 501 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 502 | lemma convex_on_linorderI [intro?]: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 503 |   fixes A :: "('a::{linorder,real_vector}) set"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 504 | 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: 
63967diff
changeset | 505 | 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: 
63967diff
changeset | 506 | shows "convex_on A f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 507 | proof | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 508 | fix x y | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 509 | fix t :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 510 | assume A: "x \<in> A" "y \<in> A" "t > 0" "t < 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 511 | with assms [of t x y] assms [of "1 - t" y x] | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 512 | 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: 
63967diff
changeset | 513 | by (cases x y rule: linorder_cases) (auto simp: algebra_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 514 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 515 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 516 | lemma convex_onD: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 517 | assumes "convex_on A f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 518 | 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: 
63967diff
changeset | 519 | 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: 
63967diff
changeset | 520 | using assms by (auto simp: convex_on_def) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 521 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 522 | lemma convex_onD_Icc: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 523 |   assumes "convex_on {x..y} f" "x \<le> (y :: _ :: {real_vector,preorder})"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 524 | shows "\<And>t. t \<ge> 0 \<Longrightarrow> t \<le> 1 \<Longrightarrow> | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 525 | 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: 
63967diff
changeset | 526 | using assms(2) by (intro convex_onD [OF assms(1)]) simp_all | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 527 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 528 | 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: 
63967diff
changeset | 529 | unfolding convex_on_def by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 530 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 531 | lemma convex_on_add [intro]: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 532 | assumes "convex_on s f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 533 | and "convex_on s g" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 534 | shows "convex_on s (\<lambda>x. f x + g x)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 535 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 536 |   {
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 537 | fix x y | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 538 | assume "x \<in> s" "y \<in> s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 539 | moreover | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 540 | fix u v :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 541 | assume "0 \<le> u" "0 \<le> v" "u + v = 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 542 | ultimately | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 543 | 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: 
63967diff
changeset | 544 | using assms unfolding convex_on_def by (auto simp: add_mono) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 545 | 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: 
63967diff
changeset | 546 | by (simp add: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 547 | } | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 548 | then show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 549 | unfolding convex_on_def by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 550 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 551 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 552 | lemma convex_on_cmul [intro]: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 553 | fixes c :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 554 | assumes "0 \<le> c" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 555 | and "convex_on s f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 556 | shows "convex_on s (\<lambda>x. c * f x)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 557 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 558 | have *: "u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 559 | for u c fx v fy :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 560 | by (simp add: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 561 | show ?thesis using assms(2) and mult_left_mono [OF _ assms(1)] | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 562 | unfolding convex_on_def and * by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 563 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 564 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 565 | lemma convex_lower: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 566 | assumes "convex_on s f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 567 | and "x \<in> s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 568 | and "y \<in> s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 569 | and "0 \<le> u" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 570 | and "0 \<le> v" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 571 | and "u + v = 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 572 | 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: 
63967diff
changeset | 573 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 574 | let ?m = "max (f x) (f y)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 575 | 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: 
63967diff
changeset | 576 | using assms(4,5) by (auto simp: mult_left_mono add_mono) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 577 | also have "\<dots> = max (f x) (f y)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 578 | using assms(6) by (simp add: distrib_right [symmetric]) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 579 | finally show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 580 | using assms unfolding convex_on_def by fastforce | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 581 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 582 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 583 | lemma convex_on_dist [intro]: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 584 | fixes s :: "'a::real_normed_vector set" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 585 | shows "convex_on s (\<lambda>x. dist a x)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 586 | proof (auto simp: convex_on_def dist_norm) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 587 | fix x y | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 588 | assume "x \<in> s" "y \<in> s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 589 | fix u v :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 590 | assume "0 \<le> u" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 591 | assume "0 \<le> v" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 592 | assume "u + v = 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 593 | have "a = u *\<^sub>R a + v *\<^sub>R a" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 594 | 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: 
63967diff
changeset | 595 | 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: 
63967diff
changeset | 596 | by (auto simp: algebra_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 597 | 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: 
63967diff
changeset | 598 | 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: 
63967diff
changeset | 599 | 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: 
63967diff
changeset | 600 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 601 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 602 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 603 | subsection \<open>Arithmetic operations on sets preserve convexity\<close> | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 604 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 605 | lemma convex_linear_image: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 606 | assumes "linear f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 607 | and "convex s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 608 | shows "convex (f ` s)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 609 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 610 | interpret f: linear f by fact | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 611 | from \<open>convex s\<close> show "convex (f ` s)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 612 | by (simp add: convex_def f.scaleR [symmetric] f.add [symmetric]) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 613 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 614 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 615 | lemma convex_linear_vimage: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 616 | assumes "linear f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 617 | and "convex s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 618 | shows "convex (f -` s)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 619 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 620 | interpret f: linear f by fact | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 621 | from \<open>convex s\<close> show "convex (f -` s)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 622 | by (simp add: convex_def f.add f.scaleR) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 623 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 624 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 625 | lemma convex_scaling: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 626 | assumes "convex s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 627 | shows "convex ((\<lambda>x. c *\<^sub>R x) ` s)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 628 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 629 | have "linear (\<lambda>x. c *\<^sub>R x)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 630 | by (simp add: linearI scaleR_add_right) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 631 | then show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 632 | using \<open>convex s\<close> by (rule convex_linear_image) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 633 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 634 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 635 | lemma convex_scaled: | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 636 | assumes "convex S" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 637 | shows "convex ((\<lambda>x. x *\<^sub>R c) ` S)" | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 638 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 639 | have "linear (\<lambda>x. x *\<^sub>R c)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 640 | by (simp add: linearI scaleR_add_left) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 641 | then show ?thesis | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 642 | using \<open>convex S\<close> by (rule convex_linear_image) | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 643 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 644 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 645 | lemma convex_negations: | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 646 | assumes "convex S" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 647 | shows "convex ((\<lambda>x. - x) ` S)" | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 648 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 649 | have "linear (\<lambda>x. - x)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 650 | by (simp add: linearI) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 651 | then show ?thesis | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 652 | using \<open>convex S\<close> by (rule convex_linear_image) | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 653 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 654 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 655 | lemma convex_sums: | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 656 | assumes "convex S" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 657 | and "convex T" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 658 |   shows "convex (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})"
 | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 659 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 660 | have "linear (\<lambda>(x, y). x + y)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 661 | by (auto intro: linearI simp: scaleR_add_right) | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 662 | with assms have "convex ((\<lambda>(x, y). x + y) ` (S \<times> T))" | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 663 | by (intro convex_linear_image convex_Times) | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 664 |   also have "((\<lambda>(x, y). x + y) ` (S \<times> T)) = (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})"
 | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 665 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 666 | finally show ?thesis . | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 667 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 668 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 669 | lemma convex_differences: | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 670 | assumes "convex S" "convex T" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 671 |   shows "convex (\<Union>x\<in> S. \<Union>y \<in> T. {x - y})"
 | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 672 | proof - | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 673 |   have "{x - y| x y. x \<in> S \<and> y \<in> T} = {x + y |x y. x \<in> S \<and> y \<in> uminus ` T}"
 | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 674 | 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: 
63967diff
changeset | 675 | then show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 676 | 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: 
63967diff
changeset | 677 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 678 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 679 | lemma convex_translation: | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 680 | assumes "convex S" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 681 | shows "convex ((\<lambda>x. a + x) ` S)" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 682 | proof - | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 683 |   have "(\<Union> x\<in> {a}. \<Union>y \<in> S. {x + y}) = (\<lambda>x. a + x) ` S"
 | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 684 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 685 | then show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 686 | using convex_sums[OF convex_singleton[of a] assms] by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 687 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 688 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 689 | lemma convex_affinity: | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 690 | assumes "convex S" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 691 | shows "convex ((\<lambda>x. a + c *\<^sub>R x) ` S)" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 692 | proof - | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 693 | have "(\<lambda>x. a + c *\<^sub>R x) ` S = op + a ` op *\<^sub>R c ` S" | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 694 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 695 | then show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 696 | 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: 
63967diff
changeset | 697 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 698 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 699 | lemma pos_is_convex: "convex {0 :: real <..}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 700 | unfolding convex_alt | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 701 | proof safe | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 702 | fix y x \<mu> :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 703 | assume *: "y > 0" "x > 0" "\<mu> \<ge> 0" "\<mu> \<le> 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 704 |   {
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 705 | assume "\<mu> = 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 706 | then have "\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y = y" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 707 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 708 | then have "\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y > 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 709 | using * by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 710 | } | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 711 | moreover | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 712 |   {
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 713 | assume "\<mu> = 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
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: 
63967diff
changeset | 715 | using * by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 716 | } | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 717 | moreover | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 718 |   {
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 719 | assume "\<mu> \<noteq> 1" "\<mu> \<noteq> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 720 | then have "\<mu> > 0" "(1 - \<mu>) > 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 721 | using * by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 722 | then have "\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y > 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 723 | using * by (auto simp: add_pos_pos) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 724 | } | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 725 | ultimately show "(1 - \<mu>) *\<^sub>R y + \<mu> *\<^sub>R x > 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 726 | by fastforce | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 727 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 728 | |
| 64267 | 729 | lemma convex_on_sum: | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 730 | fixes a :: "'a \<Rightarrow> real" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 731 | and y :: "'a \<Rightarrow> 'b::real_vector" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 732 | and f :: "'b \<Rightarrow> real" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 733 |   assumes "finite s" "s \<noteq> {}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 734 | and "convex_on C f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 735 | and "convex C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 736 | and "(\<Sum> i \<in> s. a i) = 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 737 | and "\<And>i. i \<in> s \<Longrightarrow> a i \<ge> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 738 | and "\<And>i. i \<in> s \<Longrightarrow> y i \<in> C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 739 | 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: 
63967diff
changeset | 740 | using assms | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 741 | proof (induct s arbitrary: a rule: finite_ne_induct) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 742 | case (singleton i) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 743 | then have ai: "a i = 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 744 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 745 | then show ?case | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 746 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 747 | next | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 748 | case (insert i s) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 749 | then have "convex_on C f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 750 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 751 | from this[unfolded convex_on_def, rule_format] | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 752 | 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: 
63967diff
changeset | 753 | 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: 
63967diff
changeset | 754 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 755 | show ?case | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 756 | proof (cases "a i = 1") | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 757 | case True | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 758 | then have "(\<Sum> j \<in> s. a j) = 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 759 | using insert by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 760 | then have "\<And>j. j \<in> s \<Longrightarrow> a j = 0" | 
| 64267 | 761 | using insert by (fastforce simp: sum_nonneg_eq_0_iff) | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 762 | then show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 763 | using insert by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 764 | next | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 765 | case False | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 766 | from insert have yai: "y i \<in> C" "a i \<ge> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 767 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 768 | have fis: "finite (insert i s)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 769 | using insert by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 770 | then have ai1: "a i \<le> 1" | 
| 64267 | 771 | 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: 
63967diff
changeset | 772 | then have "a i < 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 773 | using False by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 774 | then have i0: "1 - a i > 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 775 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 776 | let ?a = "\<lambda>j. a j / (1 - a i)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 777 | 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: 
63967diff
changeset | 778 | using i0 insert that by fastforce | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 779 | have "(\<Sum> j \<in> insert i s. a j) = 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 780 | using insert by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 781 | then have "(\<Sum> j \<in> s. a j) = 1 - a i" | 
| 64267 | 782 | using sum.insert insert by fastforce | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 783 | then have "(\<Sum> j \<in> s. a j) / (1 - a i) = 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 784 | using i0 by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 785 | then have a1: "(\<Sum> j \<in> s. ?a j) = 1" | 
| 64267 | 786 | unfolding sum_divide_distrib by simp | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 787 | have "convex C" using insert by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 788 | then have asum: "(\<Sum> j \<in> s. ?a j *\<^sub>R y j) \<in> C" | 
| 64267 | 789 | 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: 
63967diff
changeset | 790 | 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: 
63967diff
changeset | 791 | using a_nonneg a1 insert by blast | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 792 | 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 | 793 | 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: 
63967diff
changeset | 794 | by (auto simp only: add.commute) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 795 | 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: 
63967diff
changeset | 796 | using i0 by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 797 | 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 | 798 | 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: 
63967diff
changeset | 799 | by (auto simp: algebra_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 800 | 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: 
63967diff
changeset | 801 | by (auto simp: divide_inverse) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 802 | 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: 
63967diff
changeset | 803 | 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: 
63967diff
changeset | 804 | by (auto simp: add.commute) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 805 | 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: 
63967diff
changeset | 806 | using add_right_mono [OF mult_left_mono [of _ _ "1 - a i", | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 807 | 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: 
63967diff
changeset | 808 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 809 | also have "\<dots> = (\<Sum> j \<in> s. (1 - a i) * ?a j * f (y j)) + a i * f (y i)" | 
| 64267 | 810 | 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: 
63967diff
changeset | 811 | using i0 by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 812 | 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: 
63967diff
changeset | 813 | using i0 by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 814 | 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: 
63967diff
changeset | 815 | using insert by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 816 | finally show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 817 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 818 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 819 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 820 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 821 | lemma convex_on_alt: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 822 | fixes C :: "'a::real_vector set" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 823 | assumes "convex C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 824 | shows "convex_on C f \<longleftrightarrow> | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 825 | (\<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: 
63967diff
changeset | 826 | 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: 
63967diff
changeset | 827 | proof safe | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 828 | fix x y | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 829 | fix \<mu> :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 830 | 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: 
63967diff
changeset | 831 | from this[unfolded convex_on_def, rule_format] | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 832 | 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: 
63967diff
changeset | 833 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 834 | from this [of "\<mu>" "1 - \<mu>", simplified] * | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 835 | 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: 
63967diff
changeset | 836 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 837 | next | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 838 | 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: 
63967diff
changeset | 839 | 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: 
63967diff
changeset | 840 |   {
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 841 | fix x y | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 842 | fix u v :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 843 | 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: 
63967diff
changeset | 844 | then have[simp]: "1 - u = v" by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 845 | from *[rule_format, of x y u] | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 846 | 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: 
63967diff
changeset | 847 | using ** by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 848 | } | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 849 | then show "convex_on C f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 850 | unfolding convex_on_def by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 851 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 852 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 853 | lemma convex_on_diff: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 854 | fixes f :: "real \<Rightarrow> real" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 855 | assumes f: "convex_on I f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 856 | and I: "x \<in> I" "y \<in> I" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 857 | and t: "x < t" "t < y" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 858 | 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: 
63967diff
changeset | 859 | 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: 
63967diff
changeset | 860 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 861 | define a where "a \<equiv> (t - y) / (x - y)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 862 | with t have "0 \<le> a" "0 \<le> 1 - a" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 863 | by (auto simp: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 864 | 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: 
63967diff
changeset | 865 | by (auto simp: convex_on_def) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 866 | have "a * x + (1 - a) * y = a * (x - y) + y" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 867 | by (simp add: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 868 | also have "\<dots> = t" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 869 | 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: 
63967diff
changeset | 870 | finally have "f t \<le> a * f x + (1 - a) * f y" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 871 | using cvx by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 872 | also have "\<dots> = a * (f x - f y) + f y" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 873 | by (simp add: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 874 | finally have "f t - f y \<le> a * (f x - f y)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 875 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 876 | 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: 
63967diff
changeset | 877 | 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: 
63967diff
changeset | 878 | 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: 
63967diff
changeset | 879 | by (simp add: le_divide_eq divide_le_eq field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 880 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 881 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 882 | lemma pos_convex_function: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 883 | fixes f :: "real \<Rightarrow> real" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 884 | assumes "convex C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 885 | 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: 
63967diff
changeset | 886 | shows "convex_on C f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 887 | unfolding convex_on_alt[OF assms(1)] | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 888 | using assms | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 889 | proof safe | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 890 | fix x y \<mu> :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 891 | let ?x = "\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 892 | 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: 
63967diff
changeset | 893 | then have "1 - \<mu> \<ge> 0" by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 894 | then have xpos: "?x \<in> C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 895 | using * unfolding convex_alt by fastforce | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 896 | 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: 
63967diff
changeset | 897 | \<mu> * f' ?x * (x - ?x) + (1 - \<mu>) * f' ?x * (y - ?x)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 898 | 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: 
63967diff
changeset | 899 | 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: 
63967diff
changeset | 900 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 901 | 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: 
63967diff
changeset | 902 | by (auto simp: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 903 | 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: 
63967diff
changeset | 904 | using convex_on_alt by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 905 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 906 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 907 | lemma atMostAtLeast_subset_convex: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 908 | fixes C :: "real set" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 909 | assumes "convex C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 910 | and "x \<in> C" "y \<in> C" "x < y" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 911 |   shows "{x .. y} \<subseteq> C"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 912 | proof safe | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 913 |   fix z assume z: "z \<in> {x .. y}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 914 | have less: "z \<in> C" if *: "x < z" "z < y" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 915 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 916 | let ?\<mu> = "(y - z) / (y - x)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 917 | have "0 \<le> ?\<mu>" "?\<mu> \<le> 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 918 | using assms * by (auto simp: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 919 | then have comb: "?\<mu> * x + (1 - ?\<mu>) * y \<in> C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 920 | 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: 
63967diff
changeset | 921 | by (simp add: algebra_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 922 | 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: 
63967diff
changeset | 923 | by (auto simp: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 924 | 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: 
63967diff
changeset | 925 | using assms by (simp only: add_divide_distrib) (auto simp: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 926 | also have "\<dots> = z" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 927 | using assms by (auto simp: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 928 | finally show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 929 | using comb by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 930 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 931 | show "z \<in> C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 932 | using z less assms by (auto simp: le_less) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 933 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 934 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 935 | lemma f''_imp_f': | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 936 | fixes f :: "real \<Rightarrow> real" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 937 | assumes "convex C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 938 | 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: 
63967diff
changeset | 939 | 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: 
63967diff
changeset | 940 | and pos: "\<And>x. x \<in> C \<Longrightarrow> f'' x \<ge> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 941 | and x: "x \<in> C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 942 | and y: "y \<in> C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 943 | shows "f' x * (y - x) \<le> f y - f x" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 944 | using assms | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 945 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 946 | 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: 
63967diff
changeset | 947 | 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: 
63967diff
changeset | 948 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 949 | from * have ge: "y - x > 0" "y - x \<ge> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 950 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 951 | from * have le: "x - y < 0" "x - y \<le> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 952 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 953 | 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: 
63967diff
changeset | 954 | 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: 
63967diff
changeset | 955 | 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: 
63967diff
changeset | 956 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 957 | then have "z1 \<in> C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 958 | 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: 
63967diff
changeset | 959 | by fastforce | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 960 | from z1 have z1': "f x - f y = (x - y) * f' z1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 961 | by (simp add: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 962 | 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: 
63967diff
changeset | 963 | 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: 
63967diff
changeset | 964 | 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: 
63967diff
changeset | 965 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 966 | 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: 
63967diff
changeset | 967 | 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: 
63967diff
changeset | 968 | 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: 
63967diff
changeset | 969 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 970 | 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: 
63967diff
changeset | 971 | using * z1' by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 972 | also have "\<dots> = (y - z1) * f'' z3" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 973 | using z3 by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 974 | 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: 
63967diff
changeset | 975 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 976 | have A': "y - z1 \<ge> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 977 | using z1 by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 978 | have "z3 \<in> C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 979 | 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: 
63967diff
changeset | 980 | by fastforce | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 981 | then have B': "f'' z3 \<ge> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 982 | using assms by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 983 | from A' B' have "(y - z1) * f'' z3 \<ge> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 984 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 985 | 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: 
63967diff
changeset | 986 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 987 | from mult_right_mono_neg[OF this le(2)] | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 988 | 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: 
63967diff
changeset | 989 | by (simp add: algebra_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 990 | then have "f' y * (x - y) - (f x - f y) \<le> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 991 | using le by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 992 | then have res: "f' y * (x - y) \<le> f x - f y" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 993 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 994 | 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: 
63967diff
changeset | 995 | using * z1 by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 996 | also have "\<dots> = (z1 - x) * f'' z2" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 997 | using z2 by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 998 | 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: 
63967diff
changeset | 999 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1000 | have A: "z1 - x \<ge> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1001 | using z1 by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1002 | have "z2 \<in> C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1003 | 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: 
63967diff
changeset | 1004 | by fastforce | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1005 | then have B: "f'' z2 \<ge> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1006 | using assms by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1007 | from A B have "(z1 - x) * f'' z2 \<ge> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1008 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1009 | 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: 
63967diff
changeset | 1010 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1011 | from mult_right_mono[OF this ge(2)] | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1012 | 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: 
63967diff
changeset | 1013 | by (simp add: algebra_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1014 | then have "f y - f x - f' x * (y - x) \<ge> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1015 | using ge by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1016 | 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: 
63967diff
changeset | 1017 | using res by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1018 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1019 | show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1020 | proof (cases "x = y") | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1021 | case True | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1022 | with x y show ?thesis by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1023 | next | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1024 | case False | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1025 | with less_imp x y show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1026 | by (auto simp: neq_iff) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1027 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1028 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1029 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1030 | lemma f''_ge0_imp_convex: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1031 | fixes f :: "real \<Rightarrow> real" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1032 | assumes conv: "convex C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1033 | 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: 
63967diff
changeset | 1034 | 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: 
63967diff
changeset | 1035 | and pos: "\<And>x. x \<in> C \<Longrightarrow> f'' x \<ge> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1036 | shows "convex_on C f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1037 | 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: 
63967diff
changeset | 1038 | by fastforce | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1039 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1040 | lemma minus_log_convex: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1041 | fixes b :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1042 | assumes "b > 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1043 |   shows "convex_on {0 <..} (\<lambda> x. - log b x)"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1044 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1045 | 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: 
63967diff
changeset | 1046 | using DERIV_log by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1047 | 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: 
63967diff
changeset | 1048 | by (auto simp: DERIV_minus) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1049 | 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: 
63967diff
changeset | 1050 | using less_imp_neq[THEN not_sym, THEN DERIV_inverse] by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1051 | from this[THEN DERIV_cmult, of _ "- 1 / ln b"] | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1052 | have "\<And>z::real. z > 0 \<Longrightarrow> | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1053 | 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: 
63967diff
changeset | 1054 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1055 | then have f''0: "\<And>z::real. z > 0 \<Longrightarrow> | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1056 | 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: 
63967diff
changeset | 1057 | unfolding inverse_eq_divide by (auto simp: mult.assoc) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1058 | 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: 
63967diff
changeset | 1059 | using \<open>b > 1\<close> by (auto intro!: less_imp_le) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1060 | 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: 
63967diff
changeset | 1061 | show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1062 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1063 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1064 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1065 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1066 | subsection \<open>Convexity of real functions\<close> | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1067 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1068 | lemma convex_on_realI: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1069 | assumes "connected A" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1070 | 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: 
63967diff
changeset | 1071 | 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: 
63967diff
changeset | 1072 | shows "convex_on A f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1073 | proof (rule convex_on_linorderI) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1074 | fix t x y :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1075 | assume t: "t > 0" "t < 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1076 | assume xy: "x \<in> A" "y \<in> A" "x < y" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1077 | define z where "z = (1 - t) * x + t * y" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1078 |   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: 
63967diff
changeset | 1079 | using connected_contains_Icc by blast | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1080 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1081 | from xy t have xz: "z > x" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1082 | by (simp add: z_def algebra_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1083 | have "y - z = (1 - t) * (y - x)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1084 | by (simp add: z_def algebra_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1085 | also from xy t have "\<dots> > 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1086 | by (intro mult_pos_pos) simp_all | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1087 | finally have yz: "z < y" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1088 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1089 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1090 | 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: 
63967diff
changeset | 1091 | by (intro MVT2) (auto intro!: assms(2)) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1092 | 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: 
63967diff
changeset | 1093 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1094 | 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: 
63967diff
changeset | 1095 | by (intro MVT2) (auto intro!: assms(2)) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1096 | 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: 
63967diff
changeset | 1097 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1098 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1099 | from \<eta>(3) have "(f y - f z) / (y - z) = f' \<eta>" .. | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1100 | also from \<xi> \<eta> ivl have "\<xi> \<in> A" "\<eta> \<in> A" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1101 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1102 | with \<xi> \<eta> have "f' \<eta> \<ge> f' \<xi>" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1103 | by (intro assms(3)) auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1104 | 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: 
63967diff
changeset | 1105 | 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: 
63967diff
changeset | 1106 | using xz yz by (simp add: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1107 | also have "z - x = t * (y - x)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1108 | by (simp add: z_def algebra_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1109 | also have "y - z = (1 - t) * (y - x)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1110 | by (simp add: z_def algebra_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1111 | 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: 
63967diff
changeset | 1112 | using xy by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1113 | 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: 
63967diff
changeset | 1114 | by (simp add: z_def algebra_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1115 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1116 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1117 | lemma convex_on_inverse: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1118 |   assumes "A \<subseteq> {0<..}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1119 | shows "convex_on A (inverse :: real \<Rightarrow> real)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1120 | 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: 
63967diff
changeset | 1121 | fix u v :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1122 |   assume "u \<in> {0<..}" "v \<in> {0<..}" "u \<le> v"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1123 | with assms show "-inverse (u^2) \<le> -inverse (v^2)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1124 | 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: 
63967diff
changeset | 1125 | 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: 
63967diff
changeset | 1126 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1127 | lemma convex_onD_Icc': | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1128 |   assumes "convex_on {x..y} f" "c \<in> {x..y}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1129 | defines "d \<equiv> y - x" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1130 | 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: 
63967diff
changeset | 1131 | proof (cases x y rule: linorder_cases) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1132 | case less | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1133 | then have d: "d > 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1134 | by (simp add: d_def) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1135 | 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: 
63967diff
changeset | 1136 | by (simp_all add: d_def divide_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1137 | have "f c = f (x + (c - x) * 1)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1138 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1139 | also from less have "1 = ((y - x) / d)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1140 | by (simp add: d_def) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1141 | 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: 
63967diff
changeset | 1142 | by (simp add: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1143 | 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: 
63967diff
changeset | 1144 | using assms less by (intro convex_onD_Icc) simp_all | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1145 | 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: 
63967diff
changeset | 1146 | by (simp add: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1147 | finally show ?thesis . | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1148 | qed (insert assms(2), simp_all) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1149 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1150 | lemma convex_onD_Icc'': | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1151 |   assumes "convex_on {x..y} f" "c \<in> {x..y}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1152 | defines "d \<equiv> y - x" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1153 | 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: 
63967diff
changeset | 1154 | proof (cases x y rule: linorder_cases) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1155 | case less | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1156 | then have d: "d > 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1157 | by (simp add: d_def) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1158 | 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: 
63967diff
changeset | 1159 | by (simp_all add: d_def divide_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1160 | have "f c = f (y - (y - c) * 1)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1161 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1162 | also from less have "1 = ((y - x) / d)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1163 | by (simp add: d_def) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1164 | 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: 
63967diff
changeset | 1165 | by (simp add: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1166 | 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: 
63967diff
changeset | 1167 | 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: 
63967diff
changeset | 1168 | 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: 
63967diff
changeset | 1169 | by (simp add: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1170 | finally show ?thesis . | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1171 | qed (insert assms(2), simp_all) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1172 | |
| 64267 | 1173 | lemma convex_supp_sum: | 
| 1174 | assumes "convex S" and 1: "supp_sum u I = 1" | |
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1175 | and "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> u i \<and> (u i = 0 \<or> f i \<in> S)" | 
| 64267 | 1176 | 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: 
63967diff
changeset | 1177 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1178 |   have fin: "finite {i \<in> I. u i \<noteq> 0}"
 | 
| 64267 | 1179 | using 1 sum.infinite by (force simp: supp_sum_def support_on_def) | 
| 1180 |   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}"
 | |
| 1181 | 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: 
63967diff
changeset | 1182 | show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1183 | apply (simp add: eq) | 
| 64267 | 1184 | apply (rule convex_sum [OF fin \<open>convex S\<close>]) | 
| 1185 | 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: 
63967diff
changeset | 1186 | done | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1187 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1188 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1189 | 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: 
63967diff
changeset | 1190 | by (metis convex_translation translation_galois) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1191 | |
| 61694 
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 1192 | 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: 
61609diff
changeset | 1193 | 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: 
61609diff
changeset | 1194 | 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: 
61609diff
changeset | 1195 | 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: 
61609diff
changeset | 1196 | |
| 40377 | 1197 | lemma basis_to_basis_subspace_isomorphism: | 
| 1198 |   assumes s: "subspace (S:: ('n::euclidean_space) set)"
 | |
| 49529 | 1199 |     and t: "subspace (T :: ('m::euclidean_space) set)"
 | 
| 1200 | and d: "dim S = dim T" | |
| 53333 | 1201 | and B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "card B = dim S" | 
| 1202 | and C: "C \<subseteq> T" "independent C" "T \<subseteq> span C" "card C = dim T" | |
| 1203 | shows "\<exists>f. linear f \<and> f ` B = C \<and> f ` S = T \<and> inj_on f S" | |
| 49529 | 1204 | proof - | 
| 53333 | 1205 | from B independent_bound have fB: "finite B" | 
| 1206 | by blast | |
| 1207 | from C independent_bound have fC: "finite C" | |
| 1208 | by blast | |
| 40377 | 1209 | from B(4) C(4) card_le_inj[of B C] d obtain f where | 
| 60420 | 1210 | f: "f ` B \<subseteq> C" "inj_on f B" using \<open>finite B\<close> \<open>finite C\<close> by auto | 
| 40377 | 1211 | from linear_independent_extend[OF B(2)] obtain g where | 
| 53333 | 1212 | g: "linear g" "\<forall>x \<in> B. g x = f x" by blast | 
| 40377 | 1213 | from inj_on_iff_eq_card[OF fB, of f] f(2) | 
| 1214 | have "card (f ` B) = card B" by simp | |
| 1215 | with B(4) C(4) have ceq: "card (f ` B) = card C" using d | |
| 1216 | by simp | |
| 1217 | have "g ` B = f ` B" using g(2) | |
| 1218 | by (auto simp add: image_iff) | |
| 1219 | also have "\<dots> = C" using card_subset_eq[OF fC f(1) ceq] . | |
| 1220 | finally have gBC: "g ` B = C" . | |
| 1221 | have gi: "inj_on g B" using f(2) g(2) | |
| 1222 | by (auto simp add: inj_on_def) | |
| 1223 | note g0 = linear_indep_image_lemma[OF g(1) fB, unfolded gBC, OF C(2) gi] | |
| 53333 | 1224 |   {
 | 
| 1225 | fix x y | |
| 49529 | 1226 | assume x: "x \<in> S" and y: "y \<in> S" and gxy: "g x = g y" | 
| 53333 | 1227 | from B(3) x y have x': "x \<in> span B" and y': "y \<in> span B" | 
| 1228 | by blast+ | |
| 1229 | 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: 
63332diff
changeset | 1230 | by (simp add: linear_diff[OF g(1)]) | 
| 53333 | 1231 | have th1: "x - y \<in> span B" using x' y' | 
| 63938 | 1232 | by (metis span_diff) | 
| 53333 | 1233 | have "x = y" using g0[OF th1 th0] by simp | 
| 1234 | } | |
| 1235 | then have giS: "inj_on g S" unfolding inj_on_def by blast | |
| 40377 | 1236 | from span_subspace[OF B(1,3) s] | 
| 53333 | 1237 | have "g ` S = span (g ` B)" | 
| 1238 | by (simp add: span_linear_image[OF g(1)]) | |
| 1239 | also have "\<dots> = span C" | |
| 1240 | unfolding gBC .. | |
| 1241 | also have "\<dots> = T" | |
| 1242 | using span_subspace[OF C(1,3) t] . | |
| 40377 | 1243 | finally have gS: "g ` S = T" . | 
| 53333 | 1244 | from g(1) gS giS gBC show ?thesis | 
| 1245 | by blast | |
| 40377 | 1246 | qed | 
| 1247 | ||
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1248 | lemma closure_bounded_linear_image_subset: | 
| 44524 | 1249 | assumes f: "bounded_linear f" | 
| 53333 | 1250 | shows "f ` closure S \<subseteq> closure (f ` S)" | 
| 44524 | 1251 | using linear_continuous_on [OF f] closed_closure closure_subset | 
| 1252 | by (rule image_closure_subset) | |
| 1253 | ||
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1254 | lemma closure_linear_image_subset: | 
| 53339 | 1255 | fixes f :: "'m::euclidean_space \<Rightarrow> 'n::real_normed_vector" | 
| 49529 | 1256 | assumes "linear f" | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1257 | shows "f ` (closure S) \<subseteq> closure (f ` S)" | 
| 44524 | 1258 | using assms unfolding linear_conv_bounded_linear | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1259 | by (rule closure_bounded_linear_image_subset) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1260 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1261 | lemma closed_injective_linear_image: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1262 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1263 | assumes S: "closed S" and f: "linear f" "inj f" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1264 | shows "closed (f ` S)" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1265 | proof - | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1266 | obtain g where g: "linear g" "g \<circ> f = id" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1267 | using linear_injective_left_inverse [OF f] by blast | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1268 | then have confg: "continuous_on (range f) g" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1269 | using linear_continuous_on linear_conv_bounded_linear by blast | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1270 | have [simp]: "g ` f ` S = S" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1271 | using g by (simp add: image_comp) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1272 | have cgf: "closed (g ` f ` S)" | 
| 61808 | 1273 | 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: 
61426diff
changeset | 1274 |   have [simp]: "{x \<in> range f. g x \<in> S} = f ` S"
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1275 | using g by (simp add: o_def id_def image_def) metis | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1276 | show ?thesis | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1277 | apply (rule closedin_closed_trans [of "range f"]) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1278 | apply (rule continuous_closedin_preimage [OF confg cgf, simplified]) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1279 | apply (rule closed_injective_image_subspace) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1280 | using f | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1281 | apply (auto simp: linear_linear linear_injective_0) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1282 | done | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1283 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1284 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1285 | lemma closed_injective_linear_image_eq: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1286 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1287 | assumes f: "linear f" "inj f" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1288 | shows "(closed(image f s) \<longleftrightarrow> closed s)" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1289 | by (metis closed_injective_linear_image closure_eq closure_linear_image_subset closure_subset_eq f(1) f(2) inj_image_subset_iff) | 
| 40377 | 1290 | |
| 1291 | lemma closure_injective_linear_image: | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1292 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1293 | 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: 
61426diff
changeset | 1294 | apply (rule subset_antisym) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1295 | apply (simp add: closure_linear_image_subset) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1296 | 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: 
61426diff
changeset | 1297 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1298 | lemma closure_bounded_linear_image: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1299 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1300 | 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: 
61426diff
changeset | 1301 | apply (rule subset_antisym, simp add: closure_linear_image_subset) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1302 | apply (rule closure_minimal, simp add: closure_subset image_mono) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1303 | by (meson bounded_closure closed_closure compact_continuous_image compact_eq_bounded_closed linear_continuous_on linear_conv_bounded_linear) | 
| 40377 | 1304 | |
| 44524 | 1305 | lemma closure_scaleR: | 
| 53339 | 1306 | fixes S :: "'a::real_normed_vector set" | 
| 44524 | 1307 | shows "(op *\<^sub>R c) ` (closure S) = closure ((op *\<^sub>R c) ` S)" | 
| 1308 | proof | |
| 1309 | show "(op *\<^sub>R c) ` (closure S) \<subseteq> closure ((op *\<^sub>R c) ` S)" | |
| 53333 | 1310 | using bounded_linear_scaleR_right | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1311 | by (rule closure_bounded_linear_image_subset) | 
| 44524 | 1312 | show "closure ((op *\<^sub>R c) ` S) \<subseteq> (op *\<^sub>R c) ` (closure S)" | 
| 49529 | 1313 | by (intro closure_minimal image_mono closure_subset closed_scaling closed_closure) | 
| 1314 | qed | |
| 1315 | ||
| 1316 | lemma fst_linear: "linear fst" | |
| 53600 
8fda7ad57466
make 'linear' into a sublocale of 'bounded_linear';
 huffman parents: 
53406diff
changeset | 1317 | unfolding linear_iff by (simp add: algebra_simps) | 
| 49529 | 1318 | |
| 1319 | lemma snd_linear: "linear snd" | |
| 53600 
8fda7ad57466
make 'linear' into a sublocale of 'bounded_linear';
 huffman parents: 
53406diff
changeset | 1320 | unfolding linear_iff by (simp add: algebra_simps) | 
| 49529 | 1321 | |
| 54465 | 1322 | lemma fst_snd_linear: "linear (\<lambda>(x,y). x + y)" | 
| 53600 
8fda7ad57466
make 'linear' into a sublocale of 'bounded_linear';
 huffman parents: 
53406diff
changeset | 1323 | unfolding linear_iff by (simp add: algebra_simps) | 
| 40377 | 1324 | |
| 49529 | 1325 | lemma vector_choose_size: | 
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1326 | assumes "0 \<le> c" | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1327 |   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: 
62131diff
changeset | 1328 | proof - | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1329 | obtain a::'a where "a \<noteq> 0" | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1330 | 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: 
62131diff
changeset | 1331 | then show ?thesis | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1332 | 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: 
62131diff
changeset | 1333 | qed | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1334 | |
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1335 | lemma vector_choose_dist: | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1336 | assumes "0 \<le> c" | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1337 |   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: 
62131diff
changeset | 1338 | 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: 
62131diff
changeset | 1339 | |
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1340 | lemma sphere_eq_empty [simp]: | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1341 |   fixes a :: "'a::{real_normed_vector, perfect_space}"
 | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1342 |   shows "sphere a r = {} \<longleftrightarrow> r < 0"
 | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1343 | by (auto simp: sphere_def dist_norm) (metis dist_norm le_less_linear vector_choose_dist) | 
| 49529 | 1344 | |
| 64267 | 1345 | lemma sum_delta_notmem: | 
| 49529 | 1346 | assumes "x \<notin> s" | 
| 64267 | 1347 | shows "sum (\<lambda>y. if (y = x) then P x else Q y) s = sum Q s" | 
| 1348 | and "sum (\<lambda>y. if (x = y) then P x else Q y) s = sum Q s" | |
| 1349 | and "sum (\<lambda>y. if (y = x) then P y else Q y) s = sum Q s" | |
| 1350 | and "sum (\<lambda>y. if (x = y) then P y else Q y) s = sum Q s" | |
| 1351 | apply (rule_tac [!] sum.cong) | |
| 53333 | 1352 | using assms | 
| 1353 | apply auto | |
| 49529 | 1354 | done | 
| 33175 | 1355 | |
| 64267 | 1356 | lemma sum_delta'': | 
| 49529 | 1357 | fixes s::"'a::real_vector set" | 
| 1358 | assumes "finite s" | |
| 33175 | 1359 | 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 | 1360 | proof - | 
| 1361 | 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)" | |
| 1362 | by auto | |
| 1363 | show ?thesis | |
| 64267 | 1364 | unfolding * using sum.delta[OF assms, of y "\<lambda>x. f x *\<^sub>R x"] by auto | 
| 33175 | 1365 | qed | 
| 1366 | ||
| 53333 | 1367 | 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 | 1368 | by (fact if_distrib) | 
| 33175 | 1369 | |
| 1370 | lemma dist_triangle_eq: | |
| 44361 
75ec83d45303
remove unnecessary euclidean_space class constraints
 huffman parents: 
44349diff
changeset | 1371 | fixes x y z :: "'a::real_inner" | 
| 53333 | 1372 | shows "dist x z = dist x y + dist y z \<longleftrightarrow> | 
| 1373 | norm (x - y) *\<^sub>R (y - z) = norm (y - z) *\<^sub>R (x - y)" | |
| 49529 | 1374 | proof - | 
| 1375 | 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: 
36844diff
changeset | 1376 | show ?thesis unfolding dist_norm norm_triangle_eq[of "x - y" "y - z", unfolded *] | 
| 49529 | 1377 | by (auto simp add:norm_minus_commute) | 
| 1378 | qed | |
| 33175 | 1379 | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 1380 | |
| 60420 | 1381 | subsection \<open>Affine set and affine hull\<close> | 
| 33175 | 1382 | |
| 49529 | 1383 | definition affine :: "'a::real_vector set \<Rightarrow> bool" | 
| 1384 | 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 | 1385 | |
| 1386 | 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 | 1387 | unfolding affine_def by (metis eq_diff_eq') | 
| 33175 | 1388 | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1389 | lemma affine_empty [iff]: "affine {}"
 | 
| 33175 | 1390 | unfolding affine_def by auto | 
| 1391 | ||
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1392 | lemma affine_sing [iff]: "affine {x}"
 | 
| 33175 | 1393 | unfolding affine_alt by (auto simp add: scaleR_left_distrib [symmetric]) | 
| 1394 | ||
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1395 | lemma affine_UNIV [iff]: "affine UNIV" | 
| 33175 | 1396 | unfolding affine_def by auto | 
| 1397 | ||
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 1398 | lemma affine_Inter [intro]: "(\<And>s. s\<in>f \<Longrightarrow> affine s) \<Longrightarrow> affine (\<Inter>f)" | 
| 49531 | 1399 | unfolding affine_def by auto | 
| 33175 | 1400 | |
| 60303 | 1401 | lemma affine_Int[intro]: "affine s \<Longrightarrow> affine t \<Longrightarrow> affine (s \<inter> t)" | 
| 33175 | 1402 | unfolding affine_def by auto | 
| 1403 | ||
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 1404 | 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: 
63092diff
changeset | 1405 | apply (clarsimp simp add: affine_def) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 1406 | 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: 
63092diff
changeset | 1407 | apply (auto simp: algebra_simps) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 1408 | done | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 1409 | |
| 60303 | 1410 | lemma affine_affine_hull [simp]: "affine(affine hull s)" | 
| 49529 | 1411 | unfolding hull_def | 
| 1412 |   using affine_Inter[of "{t. affine t \<and> s \<subseteq> t}"] by auto
 | |
| 33175 | 1413 | |
| 1414 | lemma affine_hull_eq[simp]: "(affine hull s = s) \<longleftrightarrow> affine s" | |
| 49529 | 1415 | by (metis affine_affine_hull hull_same) | 
| 1416 | ||
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1417 | lemma affine_hyperplane: "affine {x. a \<bullet> x = b}"
 | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1418 | 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: 
62843diff
changeset | 1419 | |
| 33175 | 1420 | |
| 60420 | 1421 | subsubsection \<open>Some explicit formulations (from Lars Schewe)\<close> | 
| 33175 | 1422 | |
| 49529 | 1423 | lemma affine: | 
| 1424 | fixes V::"'a::real_vector set" | |
| 1425 | shows "affine V \<longleftrightarrow> | |
| 64267 | 1426 |     (\<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 | 1427 | unfolding affine_def | 
| 1428 | apply rule | |
| 1429 | apply(rule, rule, rule) | |
| 49531 | 1430 | apply(erule conjE)+ | 
| 49529 | 1431 | defer | 
| 1432 | apply (rule, rule, rule, rule, rule) | |
| 1433 | proof - | |
| 1434 | fix x y u v | |
| 1435 | assume as: "x \<in> V" "y \<in> V" "u + v = (1::real)" | |
| 64267 | 1436 |     "\<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 | 1437 | then show "u *\<^sub>R x + v *\<^sub>R y \<in> V" | 
| 1438 | apply (cases "x = y") | |
| 1439 |     using as(4)[THEN spec[where x="{x,y}"], THEN spec[where x="\<lambda>w. if w = x then u else v"]]
 | |
| 1440 | and as(1-3) | |
| 53333 | 1441 | apply (auto simp add: scaleR_left_distrib[symmetric]) | 
| 1442 | done | |
| 33175 | 1443 | next | 
| 49529 | 1444 | fix s u | 
| 1445 | 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 | 1446 |     "finite s" "s \<noteq> {}" "s \<subseteq> V" "sum u s = (1::real)"
 | 
| 63040 | 1447 | define n where "n = card s" | 
| 33175 | 1448 | have "card s = 0 \<or> card s = 1 \<or> card s = 2 \<or> card s > 2" by auto | 
| 49529 | 1449 | then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V" | 
| 1450 | proof (auto simp only: disjE) | |
| 1451 | assume "card s = 2" | |
| 53333 | 1452 | then have "card s = Suc (Suc 0)" | 
| 1453 | by auto | |
| 1454 |     then obtain a b where "s = {a, b}"
 | |
| 1455 | unfolding card_Suc_eq by auto | |
| 49529 | 1456 | then show ?thesis | 
| 1457 | using as(1)[THEN bspec[where x=a], THEN bspec[where x=b]] using as(4,5) | |
| 64267 | 1458 | by (auto simp add: sum_clauses(2)) | 
| 49529 | 1459 | next | 
| 1460 | assume "card s > 2" | |
| 1461 | then show ?thesis using as and n_def | |
| 1462 | proof (induct n arbitrary: u s) | |
| 1463 | case 0 | |
| 1464 | then show ?case by auto | |
| 1465 | next | |
| 1466 | case (Suc n) | |
| 1467 | fix s :: "'a set" and u :: "'a \<Rightarrow> real" | |
| 1468 | assume IA: | |
| 1469 | "\<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 | 1470 |           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 | 1471 | and as: | 
| 1472 | "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 | 1473 |            "finite s" "s \<noteq> {}" "s \<subseteq> V" "sum u s = 1"
 | 
| 49529 | 1474 | have "\<exists>x\<in>s. u x \<noteq> 1" | 
| 1475 | proof (rule ccontr) | |
| 1476 | assume "\<not> ?thesis" | |
| 64267 | 1477 | then have "sum u s = real_of_nat (card s)" | 
| 1478 | unfolding card_eq_sum by auto | |
| 49529 | 1479 | then show False | 
| 60420 | 1480 | using as(7) and \<open>card s > 2\<close> | 
| 49529 | 1481 | 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: 
45051diff
changeset | 1482 | qed | 
| 53339 | 1483 | then obtain x where x:"x \<in> s" "u x \<noteq> 1" by auto | 
| 33175 | 1484 | |
| 49529 | 1485 |       have c: "card (s - {x}) = card s - 1"
 | 
| 53333 | 1486 | apply (rule card_Diff_singleton) | 
| 60420 | 1487 | using \<open>x\<in>s\<close> as(4) | 
| 53333 | 1488 | apply auto | 
| 1489 | done | |
| 49529 | 1490 |       have *: "s = insert x (s - {x})" "finite (s - {x})"
 | 
| 60420 | 1491 | using \<open>x\<in>s\<close> and as(4) by auto | 
| 64267 | 1492 |       have **: "sum u (s - {x}) = 1 - u x"
 | 
| 1493 | using sum_clauses(2)[OF *(2), of u x, unfolded *(1)[symmetric] as(7)] by auto | |
| 1494 |       have ***: "inverse (1 - u x) * sum u (s - {x}) = 1"
 | |
| 60420 | 1495 | unfolding ** using \<open>u x \<noteq> 1\<close> by auto | 
| 49529 | 1496 |       have "(\<Sum>xa\<in>s - {x}. (inverse (1 - u x) * u xa) *\<^sub>R xa) \<in> V"
 | 
| 1497 |       proof (cases "card (s - {x}) > 2")
 | |
| 1498 | case True | |
| 1499 |         then have "s - {x} \<noteq> {}" "card (s - {x}) = n"
 | |
| 1500 | unfolding c and as(1)[symmetric] | |
| 49531 | 1501 | proof (rule_tac ccontr) | 
| 49529 | 1502 |           assume "\<not> s - {x} \<noteq> {}"
 | 
| 49531 | 1503 |           then have "card (s - {x}) = 0" unfolding card_0_eq[OF *(2)] by simp
 | 
| 49529 | 1504 | then show False using True by auto | 
| 1505 | qed auto | |
| 1506 | then show ?thesis | |
| 1507 |           apply (rule_tac IA[of "s - {x}" "\<lambda>y. (inverse (1 - u x) * u y)"])
 | |
| 64267 | 1508 | unfolding sum_distrib_left[symmetric] | 
| 53333 | 1509 | using as and *** and True | 
| 49529 | 1510 | apply auto | 
| 1511 | done | |
| 1512 | next | |
| 1513 | case False | |
| 53333 | 1514 |         then have "card (s - {x}) = Suc (Suc 0)"
 | 
| 1515 | using as(2) and c by auto | |
| 1516 |         then obtain a b where "(s - {x}) = {a, b}" "a\<noteq>b"
 | |
| 1517 | unfolding card_Suc_eq by auto | |
| 1518 | then show ?thesis | |
| 1519 | using as(3)[THEN bspec[where x=a], THEN bspec[where x=b]] | |
| 60420 | 1520 | using *** *(2) and \<open>s \<subseteq> V\<close> | 
| 64267 | 1521 | unfolding sum_distrib_left | 
| 1522 | by (auto simp add: sum_clauses(2)) | |
| 49529 | 1523 | qed | 
| 1524 | then have "u x + (1 - u x) = 1 \<Longrightarrow> | |
| 1525 |           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"
 | |
| 1526 | apply - | |
| 1527 | apply (rule as(3)[rule_format]) | |
| 64267 | 1528 | unfolding Real_Vector_Spaces.scaleR_right.sum | 
| 53333 | 1529 | using x(1) as(6) | 
| 1530 | apply auto | |
| 49529 | 1531 | done | 
| 1532 | then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V" | |
| 64267 | 1533 | unfolding scaleR_scaleR[symmetric] and scaleR_right.sum [symmetric] | 
| 49529 | 1534 | apply (subst *) | 
| 64267 | 1535 | unfolding sum_clauses(2)[OF *(2)] | 
| 60420 | 1536 | using \<open>u x \<noteq> 1\<close> | 
| 53333 | 1537 | apply auto | 
| 49529 | 1538 | done | 
| 1539 | qed | |
| 1540 | next | |
| 1541 | assume "card s = 1" | |
| 53333 | 1542 |     then obtain a where "s={a}"
 | 
| 1543 | by (auto simp add: card_Suc_eq) | |
| 1544 | then show ?thesis | |
| 1545 | using as(4,5) by simp | |
| 60420 | 1546 |   qed (insert \<open>s\<noteq>{}\<close> \<open>finite s\<close>, auto)
 | 
| 33175 | 1547 | qed | 
| 1548 | ||
| 1549 | lemma affine_hull_explicit: | |
| 53333 | 1550 | "affine hull p = | 
| 64267 | 1551 |     {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 | 1552 | apply (rule hull_unique) | 
| 1553 | apply (subst subset_eq) | |
| 1554 | prefer 3 | |
| 1555 | apply rule | |
| 1556 | unfolding mem_Collect_eq | |
| 1557 | apply (erule exE)+ | |
| 1558 | apply (erule conjE)+ | |
| 1559 | prefer 2 | |
| 1560 | apply rule | |
| 1561 | proof - | |
| 1562 | fix x | |
| 1563 | assume "x\<in>p" | |
| 64267 | 1564 |   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 | 1565 |     apply (rule_tac x="{x}" in exI)
 | 
| 1566 | apply (rule_tac x="\<lambda>x. 1" in exI) | |
| 49529 | 1567 | apply auto | 
| 1568 | done | |
| 33175 | 1569 | next | 
| 49529 | 1570 | fix t x s u | 
| 53333 | 1571 |   assume as: "p \<subseteq> t" "affine t" "finite s" "s \<noteq> {}"
 | 
| 64267 | 1572 | "s \<subseteq> p" "sum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x" | 
| 49529 | 1573 | then show "x \<in> t" | 
| 53333 | 1574 | using as(2)[unfolded affine, THEN spec[where x=s], THEN spec[where x=u]] | 
| 1575 | by auto | |
| 33175 | 1576 | next | 
| 64267 | 1577 |   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 | 1578 | unfolding affine_def | 
| 1579 | apply (rule, rule, rule, rule, rule) | |
| 1580 | unfolding mem_Collect_eq | |
| 1581 | proof - | |
| 1582 | fix u v :: real | |
| 1583 | assume uv: "u + v = 1" | |
| 1584 | fix x | |
| 64267 | 1585 |     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 | 1586 | then obtain sx ux where | 
| 64267 | 1587 |       x: "finite sx" "sx \<noteq> {}" "sx \<subseteq> p" "sum ux sx = 1" "(\<Sum>v\<in>sx. ux v *\<^sub>R v) = x"
 | 
| 53333 | 1588 | by auto | 
| 1589 | fix y | |
| 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) = y"
 | 
| 49529 | 1591 | then obtain sy uy where | 
| 64267 | 1592 |       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 | 1593 | have xy: "finite (sx \<union> sy)" | 
| 1594 | using x(1) y(1) by auto | |
| 1595 | have **: "(sx \<union> sy) \<inter> sx = sx" "(sx \<union> sy) \<inter> sy = sy" | |
| 1596 | by auto | |
| 49529 | 1597 |     show "\<exists>s ua. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and>
 | 
| 64267 | 1598 | sum ua s = 1 \<and> (\<Sum>v\<in>s. ua v *\<^sub>R v) = u *\<^sub>R x + v *\<^sub>R y" | 
| 49529 | 1599 | apply (rule_tac x="sx \<union> sy" in exI) | 
| 1600 | 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 | 1601 | unfolding scaleR_left_distrib sum.distrib if_smult scaleR_zero_left | 
| 1602 | ** sum.inter_restrict[OF xy, symmetric] | |
| 1603 | unfolding scaleR_scaleR[symmetric] Real_Vector_Spaces.scaleR_right.sum [symmetric] | |
| 1604 | and sum_distrib_left[symmetric] | |
| 49529 | 1605 | unfolding x y | 
| 53333 | 1606 | using x(1-3) y(1-3) uv | 
| 1607 | apply simp | |
| 49529 | 1608 | done | 
| 1609 | qed | |
| 1610 | qed | |
| 33175 | 1611 | |
| 1612 | lemma affine_hull_finite: | |
| 1613 | assumes "finite s" | |
| 64267 | 1614 |   shows "affine hull s = {y. \<exists>u. sum u s = 1 \<and> sum (\<lambda>v. u v *\<^sub>R v) s = y}"
 | 
| 53333 | 1615 | unfolding affine_hull_explicit and set_eq_iff and mem_Collect_eq | 
| 1616 | apply (rule, rule) | |
| 1617 | apply (erule exE)+ | |
| 1618 | apply (erule conjE)+ | |
| 49529 | 1619 | defer | 
| 1620 | apply (erule exE) | |
| 1621 | apply (erule conjE) | |
| 1622 | proof - | |
| 1623 | fix x u | |
| 64267 | 1624 | assume "sum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x" | 
| 49529 | 1625 | then show "\<exists>sa u. finite sa \<and> | 
| 64267 | 1626 |       \<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 | 1627 | apply (rule_tac x=s in exI, rule_tac x=u in exI) | 
| 53333 | 1628 | using assms | 
| 1629 | apply auto | |
| 49529 | 1630 | done | 
| 33175 | 1631 | next | 
| 49529 | 1632 | fix x t u | 
| 1633 | assume "t \<subseteq> s" | |
| 53333 | 1634 | then have *: "s \<inter> t = t" | 
| 1635 | by auto | |
| 64267 | 1636 |   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"
 | 
| 1637 | then show "\<exists>u. sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x" | |
| 49529 | 1638 | apply (rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI) | 
| 64267 | 1639 | unfolding if_smult scaleR_zero_left and sum.inter_restrict[OF assms, symmetric] and * | 
| 49529 | 1640 | apply auto | 
| 1641 | done | |
| 1642 | qed | |
| 1643 | ||
| 33175 | 1644 | |
| 60420 | 1645 | subsubsection \<open>Stepping theorems and hence small special cases\<close> | 
| 33175 | 1646 | |
| 1647 | lemma affine_hull_empty[simp]: "affine hull {} = {}"
 | |
| 49529 | 1648 | by (rule hull_unique) auto | 
| 33175 | 1649 | |
| 64267 | 1650 | (*could delete: it simply rewrites sum expressions, but it's used twice*) | 
| 33175 | 1651 | lemma affine_hull_finite_step: | 
| 1652 | fixes y :: "'a::real_vector" | |
| 49529 | 1653 | shows | 
| 64267 | 1654 |     "(\<exists>u. sum u {} = w \<and> sum (\<lambda>x. u x *\<^sub>R x) {} = y) \<longleftrightarrow> w = 0 \<and> y = 0" (is ?th1)
 | 
| 53347 | 1655 | and | 
| 49529 | 1656 | "finite s \<Longrightarrow> | 
| 64267 | 1657 | (\<exists>u. sum u (insert a s) = w \<and> sum (\<lambda>x. u x *\<^sub>R x) (insert a s) = y) \<longleftrightarrow> | 
| 1658 | (\<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 | 1659 | proof - | 
| 33175 | 1660 | show ?th1 by simp | 
| 53347 | 1661 | assume fin: "finite s" | 
| 1662 | show "?lhs = ?rhs" | |
| 1663 | proof | |
| 53302 | 1664 | assume ?lhs | 
| 64267 | 1665 | 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 | 1666 | by auto | 
| 53347 | 1667 | show ?rhs | 
| 49529 | 1668 | proof (cases "a \<in> s") | 
| 1669 | case True | |
| 1670 | then have *: "insert a s = s" by auto | |
| 53302 | 1671 | show ?thesis | 
| 1672 | using u[unfolded *] | |
| 1673 | apply(rule_tac x=0 in exI) | |
| 1674 | apply auto | |
| 1675 | done | |
| 33175 | 1676 | next | 
| 49529 | 1677 | case False | 
| 1678 | then show ?thesis | |
| 1679 | apply (rule_tac x="u a" in exI) | |
| 53347 | 1680 | using u and fin | 
| 53302 | 1681 | apply auto | 
| 49529 | 1682 | done | 
| 53302 | 1683 | qed | 
| 53347 | 1684 | next | 
| 53302 | 1685 | assume ?rhs | 
| 64267 | 1686 | 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 | 1687 | by auto | 
| 1688 | 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)" | |
| 1689 | by auto | |
| 53347 | 1690 | show ?lhs | 
| 49529 | 1691 | proof (cases "a \<in> s") | 
| 1692 | case True | |
| 1693 | then show ?thesis | |
| 1694 | apply (rule_tac x="\<lambda>x. (if x=a then v else 0) + u x" in exI) | |
| 64267 | 1695 | unfolding sum_clauses(2)[OF fin] | 
| 53333 | 1696 | apply simp | 
| 64267 | 1697 | unfolding scaleR_left_distrib and sum.distrib | 
| 33175 | 1698 | unfolding vu and * and scaleR_zero_left | 
| 64267 | 1699 | apply (auto simp add: sum.delta[OF fin]) | 
| 49529 | 1700 | done | 
| 33175 | 1701 | next | 
| 49531 | 1702 | case False | 
| 49529 | 1703 | then have **: | 
| 1704 | "\<And>x. x \<in> s \<Longrightarrow> u x = (if x = a then v else u x)" | |
| 1705 | "\<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 | 1706 | from False show ?thesis | 
| 49529 | 1707 | apply (rule_tac x="\<lambda>x. if x=a then v else u x" in exI) | 
| 64267 | 1708 | unfolding sum_clauses(2)[OF fin] and * using vu | 
| 1709 | 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)] | |
| 1710 | using sum.cong [of s _ u "\<lambda>x. if x = a then v else u x", OF _ **(1)] | |
| 49529 | 1711 | apply auto | 
| 1712 | done | |
| 1713 | qed | |
| 53347 | 1714 | qed | 
| 33175 | 1715 | qed | 
| 1716 | ||
| 1717 | lemma affine_hull_2: | |
| 1718 | fixes a b :: "'a::real_vector" | |
| 53302 | 1719 |   shows "affine hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b| u v. (u + v = 1)}"
 | 
| 1720 | (is "?lhs = ?rhs") | |
| 49529 | 1721 | proof - | 
| 1722 | have *: | |
| 49531 | 1723 | "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)" | 
| 49529 | 1724 | "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto | 
| 64267 | 1725 |   have "?lhs = {y. \<exists>u. sum u {a, b} = 1 \<and> (\<Sum>v\<in>{a, b}. u v *\<^sub>R v) = y}"
 | 
| 33175 | 1726 |     using affine_hull_finite[of "{a,b}"] by auto
 | 
| 1727 |   also have "\<dots> = {y. \<exists>v u. u b = 1 - v \<and> u b *\<^sub>R b = y - v *\<^sub>R a}"
 | |
| 49529 | 1728 |     by (simp add: affine_hull_finite_step(2)[of "{b}" a])
 | 
| 33175 | 1729 | also have "\<dots> = ?rhs" unfolding * by auto | 
| 1730 | finally show ?thesis by auto | |
| 1731 | qed | |
| 1732 | ||
| 1733 | lemma affine_hull_3: | |
| 1734 | fixes a b c :: "'a::real_vector" | |
| 53302 | 1735 |   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 | 1736 | proof - | 
| 1737 | have *: | |
| 49531 | 1738 | "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)" | 
| 49529 | 1739 | "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto | 
| 1740 | show ?thesis | |
| 1741 | apply (simp add: affine_hull_finite affine_hull_finite_step) | |
| 1742 | unfolding * | |
| 1743 | apply auto | |
| 53302 | 1744 | apply (rule_tac x=v in exI) | 
| 1745 | apply (rule_tac x=va in exI) | |
| 1746 | apply auto | |
| 1747 | apply (rule_tac x=u in exI) | |
| 1748 | apply force | |
| 49529 | 1749 | done | 
| 33175 | 1750 | qed | 
| 1751 | ||
| 40377 | 1752 | lemma mem_affine: | 
| 53333 | 1753 | assumes "affine S" "x \<in> S" "y \<in> S" "u + v = 1" | 
| 53347 | 1754 | shows "u *\<^sub>R x + v *\<^sub>R y \<in> S" | 
| 40377 | 1755 | using assms affine_def[of S] by auto | 
| 1756 | ||
| 1757 | lemma mem_affine_3: | |
| 53333 | 1758 | assumes "affine S" "x \<in> S" "y \<in> S" "z \<in> S" "u + v + w = 1" | 
| 53347 | 1759 | shows "u *\<^sub>R x + v *\<^sub>R y + w *\<^sub>R z \<in> S" | 
| 49529 | 1760 | proof - | 
| 53347 | 1761 |   have "u *\<^sub>R x + v *\<^sub>R y + w *\<^sub>R z \<in> affine hull {x, y, z}"
 | 
| 49529 | 1762 | using affine_hull_3[of x y z] assms by auto | 
| 1763 | moreover | |
| 53347 | 1764 |   have "affine hull {x, y, z} \<subseteq> affine hull S"
 | 
| 49529 | 1765 |     using hull_mono[of "{x, y, z}" "S"] assms by auto
 | 
| 1766 | moreover | |
| 53347 | 1767 | have "affine hull S = S" | 
| 1768 | using assms affine_hull_eq[of S] by auto | |
| 49531 | 1769 | ultimately show ?thesis by auto | 
| 40377 | 1770 | qed | 
| 1771 | ||
| 1772 | lemma mem_affine_3_minus: | |
| 53333 | 1773 | assumes "affine S" "x \<in> S" "y \<in> S" "z \<in> S" | 
| 1774 | shows "x + v *\<^sub>R (y-z) \<in> S" | |
| 1775 | using mem_affine_3[of S x y z 1 v "-v"] assms | |
| 1776 | by (simp add: algebra_simps) | |
| 40377 | 1777 | |
| 60307 
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
 paulson <lp15@cam.ac.uk> parents: 
60303diff
changeset | 1778 | corollary mem_affine_3_minus2: | 
| 
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
 paulson <lp15@cam.ac.uk> parents: 
60303diff
changeset | 1779 | "\<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: 
60303diff
changeset | 1780 | 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: 
60303diff
changeset | 1781 | |
| 40377 | 1782 | |
| 60420 | 1783 | subsubsection \<open>Some relations between affine hull and subspaces\<close> | 
| 33175 | 1784 | |
| 1785 | lemma affine_hull_insert_subset_span: | |
| 49529 | 1786 |   "affine hull (insert a s) \<subseteq> {a + v| v . v \<in> span {x - a | x . x \<in> s}}"
 | 
| 1787 | unfolding subset_eq Ball_def | |
| 1788 | unfolding affine_hull_explicit span_explicit mem_Collect_eq | |
| 50804 | 1789 | apply (rule, rule) | 
| 1790 | apply (erule exE)+ | |
| 1791 | apply (erule conjE)+ | |
| 49529 | 1792 | proof - | 
| 1793 | fix x t u | |
| 64267 | 1794 |   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 | 1795 |   have "(\<lambda>x. x - a) ` (t - {a}) \<subseteq> {x - a |x. x \<in> s}"
 | 
| 1796 | using as(3) by auto | |
| 49529 | 1797 |   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)"
 | 
| 1798 | apply (rule_tac x="x - a" in exI) | |
| 33175 | 1799 | apply (rule conjI, simp) | 
| 49529 | 1800 |     apply (rule_tac x="(\<lambda>x. x - a) ` (t - {a})" in exI)
 | 
| 1801 | apply (rule_tac x="\<lambda>x. u (x + a)" in exI) | |
| 33175 | 1802 | apply (rule conjI) using as(1) apply simp | 
| 1803 | apply (erule conjI) | |
| 1804 | using as(1) | |
| 64267 | 1805 | apply (simp add: sum.reindex[unfolded inj_on_def] scaleR_right_diff_distrib | 
| 1806 | sum_subtractf scaleR_left.sum[symmetric] sum_diff1 scaleR_left_diff_distrib) | |
| 49529 | 1807 | unfolding as | 
| 1808 | apply simp | |
| 1809 | done | |
| 1810 | qed | |
| 33175 | 1811 | |
| 1812 | lemma affine_hull_insert_span: | |
| 1813 | assumes "a \<notin> s" | |
| 49529 | 1814 |   shows "affine hull (insert a s) = {a + v | v . v \<in> span {x - a | x.  x \<in> s}}"
 | 
| 1815 | apply (rule, rule affine_hull_insert_subset_span) | |
| 1816 | unfolding subset_eq Ball_def | |
| 1817 | unfolding affine_hull_explicit and mem_Collect_eq | |
| 1818 | proof (rule, rule, erule exE, erule conjE) | |
| 49531 | 1819 | fix y v | 
| 49529 | 1820 |   assume "y = a + v" "v \<in> span {x - a |x. x \<in> s}"
 | 
| 53339 | 1821 |   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 | 1822 | unfolding span_explicit by auto | 
| 63040 | 1823 | define f where "f = (\<lambda>x. x + a) ` t" | 
| 53333 | 1824 | have f: "finite f" "f \<subseteq> s" "(\<Sum>v\<in>f. u (v - a) *\<^sub>R (v - a)) = y - a" | 
| 64267 | 1825 | unfolding f_def using obt by (auto simp add: sum.reindex[unfolded inj_on_def]) | 
| 53333 | 1826 |   have *: "f \<inter> {a} = {}" "f \<inter> - {a} = f"
 | 
| 1827 | using f(2) assms by auto | |
| 64267 | 1828 |   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 | 1829 | apply (rule_tac x = "insert a f" in exI) | 
| 64267 | 1830 | 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 | 1831 | using assms and f | 
| 64267 | 1832 | unfolding sum_clauses(2)[OF f(1)] and if_smult | 
| 1833 | unfolding sum.If_cases[OF f(1), of "\<lambda>x. x = a"] | |
| 1834 | apply (auto simp add: sum_subtractf scaleR_left.sum algebra_simps *) | |
| 49529 | 1835 | done | 
| 1836 | qed | |
| 33175 | 1837 | |
| 1838 | lemma affine_hull_span: | |
| 1839 | assumes "a \<in> s" | |
| 1840 |   shows "affine hull s = {a + v | v. v \<in> span {x - a | x. x \<in> s - {a}}}"
 | |
| 1841 |   using affine_hull_insert_span[of a "s - {a}", unfolded insert_Diff[OF assms]] by auto
 | |
| 1842 | ||
| 49529 | 1843 | |
| 60420 | 1844 | subsubsection \<open>Parallel affine sets\<close> | 
| 40377 | 1845 | |
| 53347 | 1846 | definition affine_parallel :: "'a::real_vector set \<Rightarrow> 'a::real_vector set \<Rightarrow> bool" | 
| 53339 | 1847 | where "affine_parallel S T \<longleftrightarrow> (\<exists>a. T = (\<lambda>x. a + x) ` S)" | 
| 40377 | 1848 | |
| 1849 | lemma affine_parallel_expl_aux: | |
| 49529 | 1850 | fixes S T :: "'a::real_vector set" | 
| 53339 | 1851 | assumes "\<forall>x. x \<in> S \<longleftrightarrow> a + x \<in> T" | 
| 1852 | shows "T = (\<lambda>x. a + x) ` S" | |
| 49529 | 1853 | proof - | 
| 53302 | 1854 |   {
 | 
| 1855 | fix x | |
| 53339 | 1856 | assume "x \<in> T" | 
| 1857 | then have "( - a) + x \<in> S" | |
| 1858 | using assms by auto | |
| 1859 | then have "x \<in> ((\<lambda>x. a + x) ` S)" | |
| 53333 | 1860 | using imageI[of "-a+x" S "(\<lambda>x. a+x)"] by auto | 
| 53302 | 1861 | } | 
| 53339 | 1862 | moreover have "T \<ge> (\<lambda>x. a + x) ` S" | 
| 53333 | 1863 | using assms by auto | 
| 49529 | 1864 | ultimately show ?thesis by auto | 
| 1865 | qed | |
| 1866 | ||
| 53339 | 1867 | lemma affine_parallel_expl: "affine_parallel S T \<longleftrightarrow> (\<exists>a. \<forall>x. x \<in> S \<longleftrightarrow> a + x \<in> T)" | 
| 49529 | 1868 | unfolding affine_parallel_def | 
| 1869 | using affine_parallel_expl_aux[of S _ T] by auto | |
| 1870 | ||
| 1871 | lemma affine_parallel_reflex: "affine_parallel S S" | |
| 53302 | 1872 | unfolding affine_parallel_def | 
| 1873 | apply (rule exI[of _ "0"]) | |
| 1874 | apply auto | |
| 1875 | done | |
| 40377 | 1876 | |
| 1877 | lemma affine_parallel_commut: | |
| 49529 | 1878 | assumes "affine_parallel A B" | 
| 1879 | shows "affine_parallel B A" | |
| 1880 | proof - | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53676diff
changeset | 1881 | from assms obtain a where B: "B = (\<lambda>x. a + x) ` A" | 
| 49529 | 1882 | unfolding affine_parallel_def by auto | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53676diff
changeset | 1883 | 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: 
53676diff
changeset | 1884 | from B show ?thesis | 
| 53333 | 1885 | using translation_galois [of B a A] | 
| 1886 | unfolding affine_parallel_def by auto | |
| 40377 | 1887 | qed | 
| 1888 | ||
| 1889 | lemma affine_parallel_assoc: | |
| 53339 | 1890 | assumes "affine_parallel A B" | 
| 1891 | and "affine_parallel B C" | |
| 49531 | 1892 | shows "affine_parallel A C" | 
| 49529 | 1893 | proof - | 
| 53333 | 1894 | from assms obtain ab where "B = (\<lambda>x. ab + x) ` A" | 
| 49531 | 1895 | unfolding affine_parallel_def by auto | 
| 1896 | moreover | |
| 53333 | 1897 | from assms obtain bc where "C = (\<lambda>x. bc + x) ` B" | 
| 49529 | 1898 | unfolding affine_parallel_def by auto | 
| 1899 | ultimately show ?thesis | |
| 1900 | using translation_assoc[of bc ab A] unfolding affine_parallel_def by auto | |
| 40377 | 1901 | qed | 
| 1902 | ||
| 1903 | lemma affine_translation_aux: | |
| 1904 | fixes a :: "'a::real_vector" | |
| 53333 | 1905 | assumes "affine ((\<lambda>x. a + x) ` S)" | 
| 1906 | shows "affine S" | |
| 53302 | 1907 | proof - | 
| 1908 |   {
 | |
| 1909 | fix x y u v | |
| 53333 | 1910 | assume xy: "x \<in> S" "y \<in> S" "(u :: real) + v = 1" | 
| 1911 | then have "(a + x) \<in> ((\<lambda>x. a + x) ` S)" "(a + y) \<in> ((\<lambda>x. a + x) ` S)" | |
| 1912 | by auto | |
| 53339 | 1913 | then have h1: "u *\<^sub>R (a + x) + v *\<^sub>R (a + y) \<in> (\<lambda>x. a + x) ` S" | 
| 49529 | 1914 | using xy assms unfolding affine_def by auto | 
| 53339 | 1915 | 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 | 1916 | by (simp add: algebra_simps) | 
| 53339 | 1917 | also have "\<dots> = a + (u *\<^sub>R x + v *\<^sub>R y)" | 
| 60420 | 1918 | using \<open>u + v = 1\<close> by auto | 
| 53339 | 1919 | ultimately have "a + (u *\<^sub>R x + v *\<^sub>R y) \<in> (\<lambda>x. a + x) ` S" | 
| 53333 | 1920 | using h1 by auto | 
| 49529 | 1921 | then have "u *\<^sub>R x + v *\<^sub>R y : S" by auto | 
| 1922 | } | |
| 1923 | then show ?thesis unfolding affine_def by auto | |
| 40377 | 1924 | qed | 
| 1925 | ||
| 1926 | lemma affine_translation: | |
| 1927 | fixes a :: "'a::real_vector" | |
| 53339 | 1928 | shows "affine S \<longleftrightarrow> affine ((\<lambda>x. a + x) ` S)" | 
| 49529 | 1929 | proof - | 
| 53339 | 1930 | have "affine S \<Longrightarrow> affine ((\<lambda>x. a + x) ` S)" | 
| 1931 | using affine_translation_aux[of "-a" "((\<lambda>x. a + x) ` S)"] | |
| 49529 | 1932 | using translation_assoc[of "-a" a S] by auto | 
| 1933 | then show ?thesis using affine_translation_aux by auto | |
| 40377 | 1934 | qed | 
| 1935 | ||
| 1936 | lemma parallel_is_affine: | |
| 49529 | 1937 | fixes S T :: "'a::real_vector set" | 
| 1938 | assumes "affine S" "affine_parallel S T" | |
| 1939 | shows "affine T" | |
| 1940 | proof - | |
| 53339 | 1941 | from assms obtain a where "T = (\<lambda>x. a + x) ` S" | 
| 49531 | 1942 | unfolding affine_parallel_def by auto | 
| 53339 | 1943 | then show ?thesis | 
| 1944 | using affine_translation assms by auto | |
| 40377 | 1945 | qed | 
| 1946 | ||
| 44361 
75ec83d45303
remove unnecessary euclidean_space class constraints
 huffman parents: 
44349diff
changeset | 1947 | lemma subspace_imp_affine: "subspace s \<Longrightarrow> affine s" | 
| 40377 | 1948 | unfolding subspace_def affine_def by auto | 
| 1949 | ||
| 49529 | 1950 | |
| 60420 | 1951 | subsubsection \<open>Subspace parallel to an affine set\<close> | 
| 40377 | 1952 | |
| 53339 | 1953 | lemma subspace_affine: "subspace S \<longleftrightarrow> affine S \<and> 0 \<in> S" | 
| 49529 | 1954 | proof - | 
| 53333 | 1955 | have h0: "subspace S \<Longrightarrow> affine S \<and> 0 \<in> S" | 
| 49529 | 1956 | using subspace_imp_affine[of S] subspace_0 by auto | 
| 53302 | 1957 |   {
 | 
| 53333 | 1958 | assume assm: "affine S \<and> 0 \<in> S" | 
| 53302 | 1959 |     {
 | 
| 1960 | fix c :: real | |
| 54465 | 1961 | fix x | 
| 1962 | assume x: "x \<in> S" | |
| 49529 | 1963 | have "c *\<^sub>R x = (1-c) *\<^sub>R 0 + c *\<^sub>R x" by auto | 
| 1964 | moreover | |
| 53339 | 1965 | have "(1 - c) *\<^sub>R 0 + c *\<^sub>R x \<in> S" | 
| 54465 | 1966 | using affine_alt[of S] assm x by auto | 
| 53333 | 1967 | ultimately have "c *\<^sub>R x \<in> S" by auto | 
| 49529 | 1968 | } | 
| 53333 | 1969 | then have h1: "\<forall>c. \<forall>x \<in> S. c *\<^sub>R x \<in> S" by auto | 
| 49529 | 1970 | |
| 53302 | 1971 |     {
 | 
| 1972 | fix x y | |
| 54465 | 1973 | assume xy: "x \<in> S" "y \<in> S" | 
| 63040 | 1974 | define u where "u = (1 :: real)/2" | 
| 53302 | 1975 | have "(1/2) *\<^sub>R (x+y) = (1/2) *\<^sub>R (x+y)" | 
| 1976 | by auto | |
| 49529 | 1977 | moreover | 
| 53302 | 1978 | have "(1/2) *\<^sub>R (x+y)=(1/2) *\<^sub>R x + (1-(1/2)) *\<^sub>R y" | 
| 1979 | by (simp add: algebra_simps) | |
| 49529 | 1980 | moreover | 
| 54465 | 1981 | have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> S" | 
| 1982 | using affine_alt[of S] assm xy by auto | |
| 49529 | 1983 | ultimately | 
| 53333 | 1984 | have "(1/2) *\<^sub>R (x+y) \<in> S" | 
| 53302 | 1985 | using u_def by auto | 
| 49529 | 1986 | moreover | 
| 54465 | 1987 | have "x + y = 2 *\<^sub>R ((1/2) *\<^sub>R (x+y))" | 
| 53302 | 1988 | by auto | 
| 49529 | 1989 | ultimately | 
| 54465 | 1990 | have "x + y \<in> S" | 
| 53302 | 1991 | using h1[rule_format, of "(1/2) *\<^sub>R (x+y)" "2"] by auto | 
| 49529 | 1992 | } | 
| 53302 | 1993 | then have "\<forall>x \<in> S. \<forall>y \<in> S. x + y \<in> S" | 
| 1994 | by auto | |
| 1995 | then have "subspace S" | |
| 1996 | using h1 assm unfolding subspace_def by auto | |
| 49529 | 1997 | } | 
| 1998 | then show ?thesis using h0 by metis | |
| 40377 | 1999 | qed | 
| 2000 | ||
| 2001 | lemma affine_diffs_subspace: | |
| 53333 | 2002 | assumes "affine S" "a \<in> S" | 
| 53302 | 2003 | shows "subspace ((\<lambda>x. (-a)+x) ` S)" | 
| 49529 | 2004 | proof - | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53676diff
changeset | 2005 | have [simp]: "(\<lambda>x. x - a) = plus (- a)" by (simp add: fun_eq_iff) | 
| 53302 | 2006 | have "affine ((\<lambda>x. (-a)+x) ` S)" | 
| 49531 | 2007 | using affine_translation assms by auto | 
| 53302 | 2008 | moreover have "0 : ((\<lambda>x. (-a)+x) ` S)" | 
| 53333 | 2009 | using assms exI[of "(\<lambda>x. x\<in>S \<and> -a+x = 0)" a] by auto | 
| 49531 | 2010 | ultimately show ?thesis using subspace_affine by auto | 
| 40377 | 2011 | qed | 
| 2012 | ||
| 2013 | lemma parallel_subspace_explicit: | |
| 54465 | 2014 | assumes "affine S" | 
| 2015 | and "a \<in> S" | |
| 2016 |   assumes "L \<equiv> {y. \<exists>x \<in> S. (-a) + x = y}"
 | |
| 2017 | shows "subspace L \<and> affine_parallel S L" | |
| 49529 | 2018 | proof - | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53676diff
changeset | 2019 | from assms have "L = plus (- a) ` S" by auto | 
| 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53676diff
changeset | 2020 | then have par: "affine_parallel S L" | 
| 54465 | 2021 | unfolding affine_parallel_def .. | 
| 49531 | 2022 | then have "affine L" using assms parallel_is_affine by auto | 
| 53302 | 2023 | moreover have "0 \<in> L" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53676diff
changeset | 2024 | using assms by auto | 
| 53302 | 2025 | ultimately show ?thesis | 
| 2026 | using subspace_affine par by auto | |
| 40377 | 2027 | qed | 
| 2028 | ||
| 2029 | lemma parallel_subspace_aux: | |
| 53302 | 2030 | assumes "subspace A" | 
| 2031 | and "subspace B" | |
| 2032 | and "affine_parallel A B" | |
| 2033 | shows "A \<supseteq> B" | |
| 49529 | 2034 | proof - | 
| 54465 | 2035 | from assms obtain a where a: "\<forall>x. x \<in> A \<longleftrightarrow> a + x \<in> B" | 
| 49529 | 2036 | using affine_parallel_expl[of A B] by auto | 
| 53302 | 2037 | then have "-a \<in> A" | 
| 2038 | using assms subspace_0[of B] by auto | |
| 2039 | then have "a \<in> A" | |
| 2040 | using assms subspace_neg[of A "-a"] by auto | |
| 2041 | then show ?thesis | |
| 54465 | 2042 | using assms a unfolding subspace_def by auto | 
| 40377 | 2043 | qed | 
| 2044 | ||
| 2045 | lemma parallel_subspace: | |
| 53302 | 2046 | assumes "subspace A" | 
| 2047 | and "subspace B" | |
| 2048 | and "affine_parallel A B" | |
| 49529 | 2049 | shows "A = B" | 
| 2050 | proof | |
| 53302 | 2051 | show "A \<supseteq> B" | 
| 49529 | 2052 | using assms parallel_subspace_aux by auto | 
| 53302 | 2053 | show "A \<subseteq> B" | 
| 49529 | 2054 | using assms parallel_subspace_aux[of B A] affine_parallel_commut by auto | 
| 40377 | 2055 | qed | 
| 2056 | ||
| 2057 | lemma affine_parallel_subspace: | |
| 53302 | 2058 |   assumes "affine S" "S \<noteq> {}"
 | 
| 53339 | 2059 | shows "\<exists>!L. subspace L \<and> affine_parallel S L" | 
| 49529 | 2060 | proof - | 
| 53339 | 2061 | have ex: "\<exists>L. subspace L \<and> affine_parallel S L" | 
| 49531 | 2062 | using assms parallel_subspace_explicit by auto | 
| 53302 | 2063 |   {
 | 
| 2064 | fix L1 L2 | |
| 53339 | 2065 | assume ass: "subspace L1 \<and> affine_parallel S L1" "subspace L2 \<and> affine_parallel S L2" | 
| 49529 | 2066 | then have "affine_parallel L1 L2" | 
| 2067 | using affine_parallel_commut[of S L1] affine_parallel_assoc[of L1 S L2] by auto | |
| 2068 | then have "L1 = L2" | |
| 2069 | using ass parallel_subspace by auto | |
| 2070 | } | |
| 2071 | then show ?thesis using ex by auto | |
| 2072 | qed | |
| 2073 | ||
| 40377 | 2074 | |
| 60420 | 2075 | subsection \<open>Cones\<close> | 
| 33175 | 2076 | |
| 49529 | 2077 | definition cone :: "'a::real_vector set \<Rightarrow> bool" | 
| 53339 | 2078 | where "cone s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>c\<ge>0. c *\<^sub>R x \<in> s)" | 
| 33175 | 2079 | |
| 2080 | lemma cone_empty[intro, simp]: "cone {}"
 | |
| 2081 | unfolding cone_def by auto | |
| 2082 | ||
| 2083 | lemma cone_univ[intro, simp]: "cone UNIV" | |
| 2084 | unfolding cone_def by auto | |
| 2085 | ||
| 53339 | 2086 | lemma cone_Inter[intro]: "\<forall>s\<in>f. cone s \<Longrightarrow> cone (\<Inter>f)" | 
| 33175 | 2087 | unfolding cone_def by auto | 
| 2088 | ||
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 2089 | 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: 
63332diff
changeset | 2090 | by (simp add: cone_def subspace_mul) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 2091 | |
| 49529 | 2092 | |
| 60420 | 2093 | subsubsection \<open>Conic hull\<close> | 
| 33175 | 2094 | |
| 2095 | lemma cone_cone_hull: "cone (cone hull s)" | |
| 44170 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 huffman parents: 
44142diff
changeset | 2096 | unfolding hull_def by auto | 
| 33175 | 2097 | |
| 53302 | 2098 | lemma cone_hull_eq: "cone hull s = s \<longleftrightarrow> cone s" | 
| 49529 | 2099 | apply (rule hull_eq) | 
| 53302 | 2100 | using cone_Inter | 
| 2101 | unfolding subset_eq | |
| 2102 | apply auto | |
| 49529 | 2103 | done | 
| 33175 | 2104 | |
| 40377 | 2105 | lemma mem_cone: | 
| 53302 | 2106 | assumes "cone S" "x \<in> S" "c \<ge> 0" | 
| 40377 | 2107 | shows "c *\<^sub>R x : S" | 
| 2108 | using assms cone_def[of S] by auto | |
| 2109 | ||
| 2110 | lemma cone_contains_0: | |
| 49529 | 2111 | assumes "cone S" | 
| 53302 | 2112 |   shows "S \<noteq> {} \<longleftrightarrow> 0 \<in> S"
 | 
| 49529 | 2113 | proof - | 
| 53302 | 2114 |   {
 | 
| 2115 |     assume "S \<noteq> {}"
 | |
| 2116 | then obtain a where "a \<in> S" by auto | |
| 2117 | then have "0 \<in> S" | |
| 2118 | using assms mem_cone[of S a 0] by auto | |
| 49529 | 2119 | } | 
| 2120 | then show ?thesis by auto | |
| 40377 | 2121 | qed | 
| 2122 | ||
| 44361 
75ec83d45303
remove unnecessary euclidean_space class constraints
 huffman parents: 
44349diff
changeset | 2123 | lemma cone_0: "cone {0}"
 | 
| 49529 | 2124 | unfolding cone_def by auto | 
| 40377 | 2125 | |
| 61952 | 2126 | lemma cone_Union[intro]: "(\<forall>s\<in>f. cone s) \<longrightarrow> cone (\<Union>f)" | 
| 40377 | 2127 | unfolding cone_def by blast | 
| 2128 | ||
| 2129 | lemma cone_iff: | |
| 53347 | 2130 |   assumes "S \<noteq> {}"
 | 
| 2131 | shows "cone S \<longleftrightarrow> 0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (op *\<^sub>R c) ` S = S)" | |
| 49529 | 2132 | proof - | 
| 53302 | 2133 |   {
 | 
| 2134 | assume "cone S" | |
| 2135 |     {
 | |
| 53347 | 2136 | fix c :: real | 
| 2137 | assume "c > 0" | |
| 53302 | 2138 |       {
 | 
| 2139 | fix x | |
| 53347 | 2140 | assume "x \<in> S" | 
| 2141 | then have "x \<in> (op *\<^sub>R c) ` S" | |
| 49529 | 2142 | unfolding image_def | 
| 60420 | 2143 | using \<open>cone S\<close> \<open>c>0\<close> mem_cone[of S x "1/c"] | 
| 54465 | 2144 | exI[of "(\<lambda>t. t \<in> S \<and> x = c *\<^sub>R t)" "(1 / c) *\<^sub>R x"] | 
| 53347 | 2145 | by auto | 
| 49529 | 2146 | } | 
| 2147 | moreover | |
| 53302 | 2148 |       {
 | 
| 2149 | fix x | |
| 53347 | 2150 | assume "x \<in> (op *\<^sub>R c) ` S" | 
| 2151 | then have "x \<in> S" | |
| 60420 | 2152 | using \<open>cone S\<close> \<open>c > 0\<close> | 
| 2153 | unfolding cone_def image_def \<open>c > 0\<close> by auto | |
| 49529 | 2154 | } | 
| 53302 | 2155 | ultimately have "(op *\<^sub>R c) ` S = S" by auto | 
| 40377 | 2156 | } | 
| 53339 | 2157 | then have "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (op *\<^sub>R c) ` S = S)" | 
| 60420 | 2158 | using \<open>cone S\<close> cone_contains_0[of S] assms by auto | 
| 49529 | 2159 | } | 
| 2160 | moreover | |
| 53302 | 2161 |   {
 | 
| 53339 | 2162 | assume a: "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (op *\<^sub>R c) ` S = S)" | 
| 53302 | 2163 |     {
 | 
| 2164 | fix x | |
| 2165 | assume "x \<in> S" | |
| 53347 | 2166 | fix c1 :: real | 
| 2167 | assume "c1 \<ge> 0" | |
| 2168 | then have "c1 = 0 \<or> c1 > 0" by auto | |
| 60420 | 2169 | then have "c1 *\<^sub>R x \<in> S" using a \<open>x \<in> S\<close> by auto | 
| 49529 | 2170 | } | 
| 2171 | then have "cone S" unfolding cone_def by auto | |
| 40377 | 2172 | } | 
| 49529 | 2173 | ultimately show ?thesis by blast | 
| 2174 | qed | |
| 2175 | ||
| 2176 | lemma cone_hull_empty: "cone hull {} = {}"
 | |
| 2177 | by (metis cone_empty cone_hull_eq) | |
| 2178 | ||
| 53302 | 2179 | lemma cone_hull_empty_iff: "S = {} \<longleftrightarrow> cone hull S = {}"
 | 
| 49529 | 2180 | by (metis bot_least cone_hull_empty hull_subset xtrans(5)) | 
| 2181 | ||
| 53302 | 2182 | lemma cone_hull_contains_0: "S \<noteq> {} \<longleftrightarrow> 0 \<in> cone hull S"
 | 
| 49529 | 2183 | using cone_cone_hull[of S] cone_contains_0[of "cone hull S"] cone_hull_empty_iff[of S] | 
| 2184 | by auto | |
| 40377 | 2185 | |
| 2186 | lemma mem_cone_hull: | |
| 53347 | 2187 | assumes "x \<in> S" "c \<ge> 0" | 
| 53302 | 2188 | shows "c *\<^sub>R x \<in> cone hull S" | 
| 49529 | 2189 | by (metis assms cone_cone_hull hull_inc mem_cone) | 
| 2190 | ||
| 53339 | 2191 | lemma cone_hull_expl: "cone hull S = {c *\<^sub>R x | c x. c \<ge> 0 \<and> x \<in> S}"
 | 
| 2192 | (is "?lhs = ?rhs") | |
| 49529 | 2193 | proof - | 
| 53302 | 2194 |   {
 | 
| 2195 | fix x | |
| 2196 | assume "x \<in> ?rhs" | |
| 54465 | 2197 | then obtain cx :: real and xx where x: "x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S" | 
| 49529 | 2198 | by auto | 
| 53347 | 2199 | fix c :: real | 
| 2200 | assume c: "c \<ge> 0" | |
| 53339 | 2201 | then have "c *\<^sub>R x = (c * cx) *\<^sub>R xx" | 
| 54465 | 2202 | using x by (simp add: algebra_simps) | 
| 49529 | 2203 | moreover | 
| 56536 | 2204 | have "c * cx \<ge> 0" using c x by auto | 
| 49529 | 2205 | ultimately | 
| 54465 | 2206 | have "c *\<^sub>R x \<in> ?rhs" using x by auto | 
| 53302 | 2207 | } | 
| 53347 | 2208 | then have "cone ?rhs" | 
| 2209 | unfolding cone_def by auto | |
| 2210 | then have "?rhs \<in> Collect cone" | |
| 2211 | unfolding mem_Collect_eq by auto | |
| 53302 | 2212 |   {
 | 
| 2213 | fix x | |
| 2214 | assume "x \<in> S" | |
| 2215 | then have "1 *\<^sub>R x \<in> ?rhs" | |
| 49531 | 2216 | apply auto | 
| 53347 | 2217 | apply (rule_tac x = 1 in exI) | 
| 49529 | 2218 | apply auto | 
| 2219 | done | |
| 53302 | 2220 | then have "x \<in> ?rhs" by auto | 
| 53347 | 2221 | } | 
| 2222 | then have "S \<subseteq> ?rhs" by auto | |
| 53302 | 2223 | then have "?lhs \<subseteq> ?rhs" | 
| 60420 | 2224 | using \<open>?rhs \<in> Collect cone\<close> hull_minimal[of S "?rhs" "cone"] by auto | 
| 49529 | 2225 | moreover | 
| 53302 | 2226 |   {
 | 
| 2227 | fix x | |
| 2228 | assume "x \<in> ?rhs" | |
| 54465 | 2229 | then obtain cx :: real and xx where x: "x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S" | 
| 53339 | 2230 | by auto | 
| 2231 | then have "xx \<in> cone hull S" | |
| 2232 | using hull_subset[of S] by auto | |
| 53302 | 2233 | then have "x \<in> ?lhs" | 
| 54465 | 2234 | using x cone_cone_hull[of S] cone_def[of "cone hull S"] by auto | 
| 49529 | 2235 | } | 
| 2236 | ultimately show ?thesis by auto | |
| 40377 | 2237 | qed | 
| 2238 | ||
| 2239 | lemma cone_closure: | |
| 53347 | 2240 | fixes S :: "'a::real_normed_vector set" | 
| 49529 | 2241 | assumes "cone S" | 
| 2242 | shows "cone (closure S)" | |
| 2243 | proof (cases "S = {}")
 | |
| 2244 | case True | |
| 2245 | then show ?thesis by auto | |
| 2246 | next | |
| 2247 | case False | |
| 53339 | 2248 | then have "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` S = S)" | 
| 49529 | 2249 | using cone_iff[of S] assms by auto | 
| 53339 | 2250 | then have "0 \<in> closure S \<and> (\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` closure S = closure S)" | 
| 49529 | 2251 | using closure_subset by (auto simp add: closure_scaleR) | 
| 53339 | 2252 | then show ?thesis | 
| 60974 
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
 paulson <lp15@cam.ac.uk> parents: 
60809diff
changeset | 2253 | using False cone_iff[of "closure S"] by auto | 
| 49529 | 2254 | qed | 
| 2255 | ||
| 40377 | 2256 | |
| 60420 | 2257 | subsection \<open>Affine dependence and consequential theorems (from Lars Schewe)\<close> | 
| 33175 | 2258 | |
| 49529 | 2259 | definition affine_dependent :: "'a::real_vector set \<Rightarrow> bool" | 
| 53339 | 2260 |   where "affine_dependent s \<longleftrightarrow> (\<exists>x\<in>s. x \<in> affine hull (s - {x}))"
 | 
| 33175 | 2261 | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2262 | lemma affine_dependent_subset: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2263 | "\<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: 
62950diff
changeset | 2264 | 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: 
62950diff
changeset | 2265 | 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: 
62950diff
changeset | 2266 | done | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2267 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2268 | lemma affine_independent_subset: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2269 | 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: 
62950diff
changeset | 2270 | by (metis affine_dependent_subset) | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2271 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2272 | lemma affine_independent_Diff: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2273 | "~ 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: 
62950diff
changeset | 2274 | by (meson Diff_subset affine_dependent_subset) | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2275 | |
| 33175 | 2276 | lemma affine_dependent_explicit: | 
| 2277 | "affine_dependent p \<longleftrightarrow> | |
| 64267 | 2278 | (\<exists>s u. finite s \<and> s \<subseteq> p \<and> sum u s = 0 \<and> | 
| 2279 | (\<exists>v\<in>s. u v \<noteq> 0) \<and> sum (\<lambda>v. u v *\<^sub>R v) s = 0)" | |
| 49529 | 2280 | unfolding affine_dependent_def affine_hull_explicit mem_Collect_eq | 
| 2281 | apply rule | |
| 2282 | apply (erule bexE, erule exE, erule exE) | |
| 2283 | apply (erule conjE)+ | |
| 2284 | defer | |
| 2285 | apply (erule exE, erule exE) | |
| 2286 | apply (erule conjE)+ | |
| 2287 | apply (erule bexE) | |
| 2288 | proof - | |
| 2289 | fix x s u | |
| 64267 | 2290 |   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 | 2291 | have "x \<notin> s" using as(1,4) by auto | 
| 64267 | 2292 | 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 | 2293 | 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 | 2294 | unfolding if_smult and sum_clauses(2)[OF as(2)] and sum_delta_notmem[OF \<open>x\<notin>s\<close>] and as | 
| 53339 | 2295 | using as | 
| 2296 | apply auto | |
| 49529 | 2297 | done | 
| 33175 | 2298 | next | 
| 49529 | 2299 | fix s u v | 
| 64267 | 2300 | 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 | 2301 |   have "s \<noteq> {v}"
 | 
| 2302 | using as(3,6) by auto | |
| 64267 | 2303 |   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 | 2304 | apply (rule_tac x=v in bexI) | 
| 2305 |     apply (rule_tac x="s - {v}" in exI)
 | |
| 2306 | apply (rule_tac x="\<lambda>x. - (1 / u v) * u x" in exI) | |
| 64267 | 2307 | unfolding scaleR_scaleR[symmetric] and scaleR_right.sum [symmetric] | 
| 2308 | unfolding sum_distrib_left[symmetric] and sum_diff1[OF as(1)] | |
| 53302 | 2309 | using as | 
| 2310 | apply auto | |
| 49529 | 2311 | done | 
| 33175 | 2312 | qed | 
| 2313 | ||
| 2314 | lemma affine_dependent_explicit_finite: | |
| 49529 | 2315 | fixes s :: "'a::real_vector set" | 
| 2316 | assumes "finite s" | |
| 53302 | 2317 | shows "affine_dependent s \<longleftrightarrow> | 
| 64267 | 2318 | (\<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 | 2319 | (is "?lhs = ?rhs") | 
| 2320 | proof | |
| 53347 | 2321 | 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 | 2322 | by auto | 
| 33175 | 2323 | assume ?lhs | 
| 49529 | 2324 | then obtain t u v where | 
| 64267 | 2325 | "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 | 2326 | unfolding affine_dependent_explicit by auto | 
| 49529 | 2327 | then show ?rhs | 
| 2328 | apply (rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI) | |
| 64267 | 2329 | apply auto unfolding * and sum.inter_restrict[OF assms, symmetric] | 
| 60420 | 2330 | unfolding Int_absorb1[OF \<open>t\<subseteq>s\<close>] | 
| 49529 | 2331 | apply auto | 
| 2332 | done | |
| 33175 | 2333 | next | 
| 2334 | assume ?rhs | |
| 64267 | 2335 | 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 | 2336 | by auto | 
| 49529 | 2337 | then show ?lhs unfolding affine_dependent_explicit | 
| 2338 | using assms by auto | |
| 2339 | qed | |
| 2340 | ||
| 33175 | 2341 | |
| 60420 | 2342 | subsection \<open>Connectedness of convex sets\<close> | 
| 44465 
fa56622bb7bc
move connected_real_lemma to the one place it is used
 huffman parents: 
44457diff
changeset | 2343 | |
| 51480 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2344 | lemma connectedD: | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2345 |   "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: 
61222diff
changeset | 2346 | by (rule Topological_Spaces.topological_space_class.connectedD) | 
| 33175 | 2347 | |
| 2348 | lemma convex_connected: | |
| 2349 | fixes s :: "'a::real_normed_vector set" | |
| 53302 | 2350 | assumes "convex s" | 
| 2351 | shows "connected s" | |
| 51480 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2352 | proof (rule connectedI) | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2353 | fix A B | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2354 |   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: 
51475diff
changeset | 2355 | moreover | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2356 |   assume "A \<inter> s \<noteq> {}" "B \<inter> s \<noteq> {}"
 | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2357 | then obtain a b where a: "a \<in> A" "a \<in> s" and b: "b \<in> B" "b \<in> s" by auto | 
| 63040 | 2358 | 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: 
51475diff
changeset | 2359 |   then have "continuous_on {0 .. 1} f"
 | 
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56369diff
changeset | 2360 | by (auto intro!: continuous_intros) | 
| 51480 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2361 |   then have "connected (f ` {0 .. 1})"
 | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2362 | by (auto intro!: connected_continuous_image) | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2363 | note connectedD[OF this, of A B] | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2364 |   moreover have "a \<in> A \<inter> f ` {0 .. 1}"
 | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2365 | 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: 
51475diff
changeset | 2366 |   moreover have "b \<in> B \<inter> f ` {0 .. 1}"
 | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2367 | 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: 
51475diff
changeset | 2368 |   moreover have "f ` {0 .. 1} \<subseteq> s"
 | 
| 60420 | 2369 | 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: 
51475diff
changeset | 2370 | ultimately show False by auto | 
| 33175 | 2371 | qed | 
| 2372 | ||
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61222diff
changeset | 2373 | 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: 
61222diff
changeset | 2374 | by(simp add: convex_connected) | 
| 33175 | 2375 | |
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62097diff
changeset | 2376 | proposition clopen: | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62097diff
changeset | 2377 | fixes s :: "'a :: real_normed_vector set" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62097diff
changeset | 2378 |   shows "closed s \<and> open s \<longleftrightarrow> s = {} \<or> s = UNIV"
 | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62097diff
changeset | 2379 | apply (rule iffI) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62097diff
changeset | 2380 | apply (rule connected_UNIV [unfolded connected_clopen, rule_format]) | 
| 64122 | 2381 | apply (force simp add: closed_closedin, force) | 
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62097diff
changeset | 2382 | done | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62097diff
changeset | 2383 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62097diff
changeset | 2384 | corollary compact_open: | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62097diff
changeset | 2385 | fixes s :: "'a :: euclidean_space set" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62097diff
changeset | 2386 |   shows "compact s \<and> open s \<longleftrightarrow> s = {}"
 | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62097diff
changeset | 2387 | by (auto simp: compact_eq_bounded_closed clopen) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62097diff
changeset | 2388 | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2389 | corollary finite_imp_not_open: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2390 |     fixes S :: "'a::{real_normed_vector, perfect_space} set"
 | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2391 |     shows "\<lbrakk>finite S; open S\<rbrakk> \<Longrightarrow> S={}"
 | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2392 | 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: 
62843diff
changeset | 2393 | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2394 | corollary empty_interior_finite: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2395 |     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: 
62950diff
changeset | 2396 |     shows "finite S \<Longrightarrow> interior S = {}"
 | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2397 | 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: 
62950diff
changeset | 2398 | |
| 60420 | 2399 | text \<open>Balls, being convex, are connected.\<close> | 
| 33175 | 2400 | |
| 56188 | 2401 | lemma convex_prod: | 
| 53347 | 2402 |   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: 
50104diff
changeset | 2403 |   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: 
50104diff
changeset | 2404 | 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: 
50104diff
changeset | 2405 | 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: 
50104diff
changeset | 2406 | |
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 2407 | lemma convex_positive_orthant: "convex {x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i)}"
 | 
| 56188 | 2408 | by (rule convex_prod) (simp add: atLeast_def[symmetric] convex_real_interval) | 
| 33175 | 2409 | |
| 2410 | lemma convex_local_global_minimum: | |
| 2411 | fixes s :: "'a::real_normed_vector set" | |
| 53347 | 2412 | assumes "e > 0" | 
| 2413 | and "convex_on s f" | |
| 2414 | and "ball x e \<subseteq> s" | |
| 2415 | and "\<forall>y\<in>ball x e. f x \<le> f y" | |
| 33175 | 2416 | shows "\<forall>y\<in>s. f x \<le> f y" | 
| 53302 | 2417 | proof (rule ccontr) | 
| 2418 | have "x \<in> s" using assms(1,3) by auto | |
| 2419 | assume "\<not> ?thesis" | |
| 2420 | 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: 
61952diff
changeset | 2421 | then have xy: "0 < dist x y" by auto | 
| 53347 | 2422 | then obtain u where "0 < u" "u \<le> 1" and u: "u < e / dist x y" | 
| 60420 | 2423 | using real_lbound_gt_zero[of 1 "e / dist x y"] xy \<open>e>0\<close> by auto | 
| 53302 | 2424 | then have "f ((1-u) *\<^sub>R x + u *\<^sub>R y) \<le> (1-u) * f x + u * f y" | 
| 60420 | 2425 | using \<open>x\<in>s\<close> \<open>y\<in>s\<close> | 
| 53302 | 2426 | using assms(2)[unfolded convex_on_def, | 
| 2427 | THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x="1-u"]] | |
| 50804 | 2428 | by auto | 
| 33175 | 2429 | moreover | 
| 50804 | 2430 | have *: "x - ((1 - u) *\<^sub>R x + u *\<^sub>R y) = u *\<^sub>R (x - y)" | 
| 2431 | by (simp add: algebra_simps) | |
| 2432 | have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> ball x e" | |
| 53302 | 2433 | unfolding mem_ball dist_norm | 
| 60420 | 2434 | unfolding * and norm_scaleR and abs_of_pos[OF \<open>0<u\<close>] | 
| 50804 | 2435 | unfolding dist_norm[symmetric] | 
| 53302 | 2436 | using u | 
| 2437 | unfolding pos_less_divide_eq[OF xy] | |
| 2438 | by auto | |
| 2439 | then have "f x \<le> f ((1 - u) *\<^sub>R x + u *\<^sub>R y)" | |
| 2440 | using assms(4) by auto | |
| 50804 | 2441 | ultimately show False | 
| 60420 | 2442 | using mult_strict_left_mono[OF y \<open>u>0\<close>] | 
| 53302 | 2443 | unfolding left_diff_distrib | 
| 2444 | by auto | |
| 33175 | 2445 | qed | 
| 2446 | ||
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 2447 | lemma convex_ball [iff]: | 
| 33175 | 2448 | fixes x :: "'a::real_normed_vector" | 
| 49531 | 2449 | shows "convex (ball x e)" | 
| 50804 | 2450 | proof (auto simp add: convex_def) | 
| 2451 | fix y z | |
| 2452 | assume yz: "dist x y < e" "dist x z < e" | |
| 2453 | fix u v :: real | |
| 2454 | assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1" | |
| 2455 | have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z" | |
| 2456 | using uv yz | |
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2457 | using convex_on_dist [of "ball x e" x, unfolded convex_on_def, | 
| 53302 | 2458 | THEN bspec[where x=y], THEN bspec[where x=z]] | 
| 50804 | 2459 | by auto | 
| 2460 | then show "dist x (u *\<^sub>R y + v *\<^sub>R z) < e" | |
| 2461 | using convex_bound_lt[OF yz uv] by auto | |
| 33175 | 2462 | qed | 
| 2463 | ||
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 2464 | lemma convex_cball [iff]: | 
| 33175 | 2465 | fixes x :: "'a::real_normed_vector" | 
| 53347 | 2466 | shows "convex (cball x e)" | 
| 2467 | proof - | |
| 2468 |   {
 | |
| 2469 | fix y z | |
| 2470 | assume yz: "dist x y \<le> e" "dist x z \<le> e" | |
| 2471 | fix u v :: real | |
| 2472 | assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1" | |
| 2473 | have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z" | |
| 2474 | using uv yz | |
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2475 | using convex_on_dist [of "cball x e" x, unfolded convex_on_def, | 
| 53347 | 2476 | THEN bspec[where x=y], THEN bspec[where x=z]] | 
| 2477 | by auto | |
| 2478 | then have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> e" | |
| 2479 | using convex_bound_le[OF yz uv] by auto | |
| 2480 | } | |
| 2481 | then show ?thesis by (auto simp add: convex_def Ball_def) | |
| 33175 | 2482 | qed | 
| 2483 | ||
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2484 | lemma connected_ball [iff]: | 
| 33175 | 2485 | fixes x :: "'a::real_normed_vector" | 
| 2486 | shows "connected (ball x e)" | |
| 2487 | using convex_connected convex_ball by auto | |
| 2488 | ||
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2489 | lemma connected_cball [iff]: | 
| 33175 | 2490 | fixes x :: "'a::real_normed_vector" | 
| 53302 | 2491 | shows "connected (cball x e)" | 
| 33175 | 2492 | using convex_connected convex_cball by auto | 
| 2493 | ||
| 50804 | 2494 | |
| 60420 | 2495 | subsection \<open>Convex hull\<close> | 
| 33175 | 2496 | |
| 60762 | 2497 | lemma convex_convex_hull [iff]: "convex (convex hull s)" | 
| 53302 | 2498 | unfolding hull_def | 
| 2499 |   using convex_Inter[of "{t. convex t \<and> s \<subseteq> t}"]
 | |
| 44170 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 huffman parents: 
44142diff
changeset | 2500 | by auto | 
| 33175 | 2501 | |
| 63016 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 2502 | lemma convex_hull_subset: | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 2503 | "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: 
63007diff
changeset | 2504 | by (simp add: convex_convex_hull subset_hull) | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 2505 | |
| 34064 
eee04bbbae7e
avoid dependency on implicit dest rule predicate1D in proofs
 haftmann parents: 
33758diff
changeset | 2506 | lemma convex_hull_eq: "convex hull s = s \<longleftrightarrow> convex s" | 
| 50804 | 2507 | by (metis convex_convex_hull hull_same) | 
| 33175 | 2508 | |
| 2509 | lemma bounded_convex_hull: | |
| 2510 | fixes s :: "'a::real_normed_vector set" | |
| 53347 | 2511 | assumes "bounded s" | 
| 2512 | shows "bounded (convex hull s)" | |
| 50804 | 2513 | proof - | 
| 2514 | from assms obtain B where B: "\<forall>x\<in>s. norm x \<le> B" | |
| 2515 | unfolding bounded_iff by auto | |
| 2516 | show ?thesis | |
| 2517 | apply (rule bounded_subset[OF bounded_cball, of _ 0 B]) | |
| 44170 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 huffman parents: 
44142diff
changeset | 2518 | unfolding subset_hull[of convex, OF convex_cball] | 
| 53302 | 2519 | unfolding subset_eq mem_cball dist_norm using B | 
| 2520 | apply auto | |
| 50804 | 2521 | done | 
| 2522 | qed | |
| 33175 | 2523 | |
| 2524 | lemma finite_imp_bounded_convex_hull: | |
| 2525 | fixes s :: "'a::real_normed_vector set" | |
| 53302 | 2526 | shows "finite s \<Longrightarrow> bounded (convex hull s)" | 
| 2527 | using bounded_convex_hull finite_imp_bounded | |
| 2528 | by auto | |
| 33175 | 2529 | |
| 50804 | 2530 | |
| 60420 | 2531 | subsubsection \<open>Convex hull is "preserved" by a linear function\<close> | 
| 40377 | 2532 | |
| 2533 | lemma convex_hull_linear_image: | |
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2534 | assumes f: "linear f" | 
| 40377 | 2535 | shows "f ` (convex hull s) = convex hull (f ` s)" | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2536 | proof | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2537 | show "convex hull (f ` s) \<subseteq> f ` (convex hull s)" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2538 | 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: 
53600diff
changeset | 2539 | show "f ` (convex hull s) \<subseteq> convex hull (f ` s)" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2540 | proof (unfold image_subset_iff_subset_vimage, rule hull_minimal) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2541 | show "s \<subseteq> f -` (convex hull (f ` s))" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2542 | by (fast intro: hull_inc) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2543 | show "convex (f -` (convex hull (f ` s)))" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2544 | by (intro convex_linear_vimage [OF f] convex_convex_hull) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2545 | qed | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2546 | qed | 
| 40377 | 2547 | |
| 2548 | lemma in_convex_hull_linear_image: | |
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2549 | assumes "linear f" | 
| 53347 | 2550 | and "x \<in> convex hull s" | 
| 53339 | 2551 | shows "f x \<in> convex hull (f ` s)" | 
| 50804 | 2552 | using convex_hull_linear_image[OF assms(1)] assms(2) by auto | 
| 2553 | ||
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2554 | lemma convex_hull_Times: | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2555 | "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: 
53600diff
changeset | 2556 | proof | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2557 | 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: 
53600diff
changeset | 2558 | 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: 
53600diff
changeset | 2559 | 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: 
53600diff
changeset | 2560 | proof (intro hull_induct) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2561 | fix x y assume "x \<in> s" and "y \<in> t" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2562 | then show "(x, y) \<in> convex hull (s \<times> t)" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2563 | by (simp add: hull_inc) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2564 | next | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2565 | 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: 
53600diff
changeset | 2566 | have "convex ?S" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2567 | by (intro convex_linear_vimage convex_translation convex_convex_hull, | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2568 | simp add: linear_iff) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2569 |     also have "?S = {y. (x, y) \<in> convex hull (s \<times> t)}"
 | 
| 57865 | 2570 | by (auto simp add: image_def Bex_def) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2571 |     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: 
53600diff
changeset | 2572 | next | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2573 |     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: 
53600diff
changeset | 2574 | proof (unfold Collect_ball_eq, rule convex_INT [rule_format]) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2575 | 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: 
53600diff
changeset | 2576 | have "convex ?S" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2577 | by (intro convex_linear_vimage convex_translation convex_convex_hull, | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2578 | simp add: linear_iff) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2579 |       also have "?S = {x. (x, y) \<in> convex hull (s \<times> t)}"
 | 
| 57865 | 2580 | by (auto simp add: image_def Bex_def) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2581 |       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: 
53600diff
changeset | 2582 | qed | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2583 | qed | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2584 | 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: 
53600diff
changeset | 2585 | unfolding subset_eq split_paired_Ball_Sigma . | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2586 | qed | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2587 | |
| 40377 | 2588 | |
| 60420 | 2589 | subsubsection \<open>Stepping theorems for convex hulls of finite sets\<close> | 
| 33175 | 2590 | |
| 2591 | lemma convex_hull_empty[simp]: "convex hull {} = {}"
 | |
| 50804 | 2592 | by (rule hull_unique) auto | 
| 33175 | 2593 | |
| 2594 | lemma convex_hull_singleton[simp]: "convex hull {a} = {a}"
 | |
| 50804 | 2595 | by (rule hull_unique) auto | 
| 33175 | 2596 | |
| 2597 | lemma convex_hull_insert: | |
| 2598 | fixes s :: "'a::real_vector set" | |
| 2599 |   assumes "s \<noteq> {}"
 | |
| 50804 | 2600 | shows "convex hull (insert a s) = | 
| 2601 |     {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 | 2602 | (is "_ = ?hull") | 
| 50804 | 2603 | apply (rule, rule hull_minimal, rule) | 
| 2604 | unfolding insert_iff | |
| 2605 | prefer 3 | |
| 2606 | apply rule | |
| 2607 | proof - | |
| 2608 | fix x | |
| 2609 | assume x: "x = a \<or> x \<in> s" | |
| 2610 | then show "x \<in> ?hull" | |
| 2611 | apply rule | |
| 2612 | unfolding mem_Collect_eq | |
| 2613 | apply (rule_tac x=1 in exI) | |
| 2614 | defer | |
| 2615 | apply (rule_tac x=0 in exI) | |
| 2616 | using assms hull_subset[of s convex] | |
| 2617 | apply auto | |
| 2618 | done | |
| 33175 | 2619 | next | 
| 50804 | 2620 | fix x | 
| 2621 | assume "x \<in> ?hull" | |
| 2622 | 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" | |
| 2623 | by auto | |
| 53339 | 2624 | have "a \<in> convex hull insert a s" "b \<in> convex hull insert a s" | 
| 50804 | 2625 |     using hull_mono[of s "insert a s" convex] hull_mono[of "{a}" "insert a s" convex] and obt(4)
 | 
| 2626 | by auto | |
| 2627 | then show "x \<in> convex hull insert a s" | |
| 53676 | 2628 | unfolding obt(5) using obt(1-3) | 
| 2629 | by (rule convexD [OF convex_convex_hull]) | |
| 33175 | 2630 | next | 
| 50804 | 2631 | show "convex ?hull" | 
| 53676 | 2632 | proof (rule convexI) | 
| 50804 | 2633 | fix x y u v | 
| 2634 | assume as: "(0::real) \<le> u" "0 \<le> v" "u + v = 1" "x\<in>?hull" "y\<in>?hull" | |
| 53339 | 2635 | from as(4) obtain u1 v1 b1 where | 
| 2636 | obt1: "u1\<ge>0" "v1\<ge>0" "u1 + v1 = 1" "b1 \<in> convex hull s" "x = u1 *\<^sub>R a + v1 *\<^sub>R b1" | |
| 2637 | by auto | |
| 2638 | from as(5) obtain u2 v2 b2 where | |
| 2639 | obt2: "u2\<ge>0" "v2\<ge>0" "u2 + v2 = 1" "b2 \<in> convex hull s" "y = u2 *\<^sub>R a + v2 *\<^sub>R b2" | |
| 2640 | by auto | |
| 50804 | 2641 | have *: "\<And>(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" | 
| 2642 | by (auto simp add: algebra_simps) | |
| 2643 | have **: "\<exists>b \<in> convex hull s. u *\<^sub>R x + v *\<^sub>R y = | |
| 2644 | (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (b - (u * u1) *\<^sub>R b - (v * u2) *\<^sub>R b)" | |
| 2645 | proof (cases "u * v1 + v * v2 = 0") | |
| 2646 | case True | |
| 2647 | have *: "\<And>(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" | |
| 2648 | by (auto simp add: algebra_simps) | |
| 2649 | from True have ***: "u * v1 = 0" "v * v2 = 0" | |
| 60420 | 2650 | 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 | 2651 | by arith+ | 
| 50804 | 2652 | then have "u * u1 + v * u2 = 1" | 
| 2653 | using as(3) obt1(3) obt2(3) by auto | |
| 2654 | then show ?thesis | |
| 2655 | unfolding obt1(5) obt2(5) * | |
| 2656 | using assms hull_subset[of s convex] | |
| 2657 | by (auto simp add: *** scaleR_right_distrib) | |
| 33175 | 2658 | next | 
| 50804 | 2659 | case False | 
| 2660 | have "1 - (u * u1 + v * u2) = (u + v) - (u * u1 + v * u2)" | |
| 2661 | using as(3) obt1(3) obt2(3) by (auto simp add: field_simps) | |
| 2662 | also have "\<dots> = u * (v1 + u1 - u1) + v * (v2 + u2 - u2)" | |
| 2663 | using as(3) obt1(3) obt2(3) by (auto simp add: field_simps) | |
| 2664 | also have "\<dots> = u * v1 + v * v2" | |
| 2665 | by simp | |
| 2666 | finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" by auto | |
| 2667 | have "0 \<le> u * v1 + v * v2" "0 \<le> u * v1" "0 \<le> u * v1 + v * v2" "0 \<le> v * v2" | |
| 56536 | 2668 | using as(1,2) obt1(1,2) obt2(1,2) by auto | 
| 50804 | 2669 | then show ?thesis | 
| 2670 | unfolding obt1(5) obt2(5) | |
| 2671 | unfolding * and ** | |
| 2672 | using False | |
| 53339 | 2673 | apply (rule_tac | 
| 2674 | x = "((u * v1) / (u * v1 + v * v2)) *\<^sub>R b1 + ((v * v2) / (u * v1 + v * v2)) *\<^sub>R b2" in bexI) | |
| 50804 | 2675 | defer | 
| 53676 | 2676 | apply (rule convexD [OF convex_convex_hull]) | 
| 50804 | 2677 | using obt1(4) obt2(4) | 
| 49530 | 2678 | unfolding add_divide_distrib[symmetric] and zero_le_divide_iff | 
| 50804 | 2679 | apply (auto simp add: scaleR_left_distrib scaleR_right_distrib) | 
| 2680 | done | |
| 2681 | qed | |
| 2682 | have u1: "u1 \<le> 1" | |
| 2683 | unfolding obt1(3)[symmetric] and not_le using obt1(2) by auto | |
| 2684 | have u2: "u2 \<le> 1" | |
| 2685 | unfolding obt2(3)[symmetric] and not_le using obt2(2) by auto | |
| 53339 | 2686 | have "u1 * u + u2 * v \<le> max u1 u2 * u + max u1 u2 * v" | 
| 50804 | 2687 | apply (rule add_mono) | 
| 2688 | apply (rule_tac [!] mult_right_mono) | |
| 2689 | using as(1,2) obt1(1,2) obt2(1,2) | |
| 2690 | apply auto | |
| 2691 | done | |
| 2692 | also have "\<dots> \<le> 1" | |
| 2693 | unfolding distrib_left[symmetric] and as(3) using u1 u2 by auto | |
| 2694 | finally show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull" | |
| 2695 | unfolding mem_Collect_eq | |
| 2696 | apply (rule_tac x="u * u1 + v * u2" in exI) | |
| 2697 | apply (rule conjI) | |
| 2698 | defer | |
| 2699 | apply (rule_tac x="1 - u * u1 - v * u2" in exI) | |
| 2700 | unfolding Bex_def | |
| 2701 | using as(1,2) obt1(1,2) obt2(1,2) ** | |
| 56536 | 2702 | apply (auto simp add: algebra_simps) | 
| 50804 | 2703 | done | 
| 33175 | 2704 | qed | 
| 2705 | qed | |
| 2706 | ||
| 66287 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65719diff
changeset | 2707 | lemma convex_hull_insert_alt: | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65719diff
changeset | 2708 | "convex hull (insert a S) = | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65719diff
changeset | 2709 |       (if S = {} then {a}
 | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65719diff
changeset | 2710 |       else {(1 - u) *\<^sub>R a + u *\<^sub>R x |x u. 0 \<le> u \<and> u \<le> 1 \<and> x \<in> convex hull S})"
 | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65719diff
changeset | 2711 | apply (auto simp: convex_hull_insert) | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65719diff
changeset | 2712 | using diff_eq_eq apply fastforce | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65719diff
changeset | 2713 | by (metis add.group_left_neutral add_le_imp_le_diff diff_add_cancel) | 
| 33175 | 2714 | |
| 60420 | 2715 | subsubsection \<open>Explicit expression for convex hull\<close> | 
| 33175 | 2716 | |
| 2717 | lemma convex_hull_indexed: | |
| 2718 | fixes s :: "'a::real_vector set" | |
| 50804 | 2719 | shows "convex hull s = | 
| 53347 | 2720 |     {y. \<exists>k u x.
 | 
| 2721 |       (\<forall>i\<in>{1::nat .. k}. 0 \<le> u i \<and> x i \<in> s) \<and>
 | |
| 64267 | 2722 |       (sum u {1..k} = 1) \<and> (sum (\<lambda>i. u i *\<^sub>R x i) {1..k} = y)}"
 | 
| 53339 | 2723 | (is "?xyz = ?hull") | 
| 50804 | 2724 | apply (rule hull_unique) | 
| 2725 | apply rule | |
| 2726 | defer | |
| 53676 | 2727 | apply (rule convexI) | 
| 50804 | 2728 | proof - | 
| 2729 | fix x | |
| 2730 | assume "x\<in>s" | |
| 2731 | then show "x \<in> ?hull" | |
| 2732 | unfolding mem_Collect_eq | |
| 2733 | apply (rule_tac x=1 in exI, rule_tac x="\<lambda>x. 1" in exI) | |
| 2734 | apply auto | |
| 2735 | done | |
| 33175 | 2736 | next | 
| 50804 | 2737 | fix t | 
| 2738 | assume as: "s \<subseteq> t" "convex t" | |
| 2739 | show "?hull \<subseteq> t" | |
| 2740 | apply rule | |
| 2741 | unfolding mem_Collect_eq | |
| 53302 | 2742 | apply (elim exE conjE) | 
| 50804 | 2743 | proof - | 
| 2744 | fix x k u y | |
| 2745 | assume assm: | |
| 2746 |       "\<forall>i\<in>{1::nat..k}. 0 \<le> u i \<and> y i \<in> s"
 | |
| 64267 | 2747 |       "sum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R y i) = x"
 | 
| 50804 | 2748 | show "x\<in>t" | 
| 2749 | unfolding assm(3) [symmetric] | |
| 2750 | apply (rule as(2)[unfolded convex, rule_format]) | |
| 2751 | using assm(1,2) as(1) apply auto | |
| 2752 | done | |
| 2753 | qed | |
| 33175 | 2754 | next | 
| 50804 | 2755 | fix x y u v | 
| 53347 | 2756 | assume uv: "0 \<le> u" "0 \<le> v" "u + v = (1::real)" | 
| 2757 | assume xy: "x \<in> ?hull" "y \<in> ?hull" | |
| 50804 | 2758 | from xy obtain k1 u1 x1 where | 
| 64267 | 2759 |     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 | 2760 | by auto | 
| 2761 | from xy obtain k2 u2 x2 where | |
| 64267 | 2762 |     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 | 2763 | by auto | 
| 2764 | have *: "\<And>P (x1::'a) x2 s1 s2 i. | |
| 2765 | (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 | 2766 |     "{1..k1 + k2} \<inter> {1..k1} = {1..k1}" "{1..k1 + k2} \<inter> - {1..k1} = (\<lambda>i. i + k1) ` {1..k2}"
 | 
| 50804 | 2767 | prefer 3 | 
| 2768 | apply (rule, rule) | |
| 2769 | unfolding image_iff | |
| 2770 | apply (rule_tac x = "x - k1" in bexI) | |
| 2771 | apply (auto simp add: not_le) | |
| 2772 | done | |
| 2773 |   have inj: "inj_on (\<lambda>i. i + k1) {1..k2}"
 | |
| 2774 | unfolding inj_on_def by auto | |
| 2775 | show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull" | |
| 2776 | apply rule | |
| 2777 | apply (rule_tac x="k1 + k2" in exI) | |
| 2778 |     apply (rule_tac x="\<lambda>i. if i \<in> {1..k1} then u * u1 i else v * u2 (i - k1)" in exI)
 | |
| 2779 |     apply (rule_tac x="\<lambda>i. if i \<in> {1..k1} then x1 i else x2 (i - k1)" in exI)
 | |
| 2780 | apply (rule, rule) | |
| 2781 | defer | |
| 2782 | apply rule | |
| 64267 | 2783 | unfolding * and sum.If_cases[OF finite_atLeastAtMost[of 1 "k1 + k2"]] and | 
| 2784 | sum.reindex[OF inj] and o_def Collect_mem_eq | |
| 2785 | unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] sum_distrib_left[symmetric] | |
| 50804 | 2786 | proof - | 
| 2787 | fix i | |
| 2788 |     assume i: "i \<in> {1..k1+k2}"
 | |
| 2789 |     show "0 \<le> (if i \<in> {1..k1} then u * u1 i else v * u2 (i - k1)) \<and>
 | |
| 2790 |       (if i \<in> {1..k1} then x1 i else x2 (i - k1)) \<in> s"
 | |
| 2791 |     proof (cases "i\<in>{1..k1}")
 | |
| 2792 | case True | |
| 2793 | then show ?thesis | |
| 56536 | 2794 | using uv(1) x(1)[THEN bspec[where x=i]] by auto | 
| 50804 | 2795 | next | 
| 2796 | case False | |
| 63040 | 2797 | define j where "j = i - k1" | 
| 53347 | 2798 |       from i False have "j \<in> {1..k2}"
 | 
| 2799 | unfolding j_def by auto | |
| 50804 | 2800 | then show ?thesis | 
| 56536 | 2801 | using False uv(2) y(1)[THEN bspec[where x=j]] | 
| 2802 | by (auto simp: j_def[symmetric]) | |
| 50804 | 2803 | qed | 
| 2804 | qed (auto simp add: not_le x(2,3) y(2,3) uv(3)) | |
| 33175 | 2805 | qed | 
| 2806 | ||
| 2807 | lemma convex_hull_finite: | |
| 2808 | fixes s :: "'a::real_vector set" | |
| 2809 | assumes "finite s" | |
| 2810 |   shows "convex hull s = {y. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and>
 | |
| 64267 | 2811 | sum u s = 1 \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y}" | 
| 53339 | 2812 | (is "?HULL = ?set") | 
| 50804 | 2813 | proof (rule hull_unique, auto simp add: convex_def[of ?set]) | 
| 2814 | fix x | |
| 2815 | assume "x \<in> s" | |
| 64267 | 2816 | 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 | 2817 | apply (rule_tac x="\<lambda>y. if x=y then 1 else 0" in exI) | 
| 2818 | apply auto | |
| 64267 | 2819 | unfolding sum.delta'[OF assms] and sum_delta''[OF assms] | 
| 50804 | 2820 | apply auto | 
| 2821 | done | |
| 33175 | 2822 | next | 
| 50804 | 2823 | fix u v :: real | 
| 2824 | assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1" | |
| 64267 | 2825 | fix ux assume ux: "\<forall>x\<in>s. 0 \<le> ux x" "sum ux s = (1::real)" | 
| 2826 | fix uy assume uy: "\<forall>x\<in>s. 0 \<le> uy x" "sum uy s = (1::real)" | |
| 53339 | 2827 |   {
 | 
| 2828 | fix x | |
| 50804 | 2829 | assume "x\<in>s" | 
| 2830 | then have "0 \<le> u * ux x + v * uy x" | |
| 2831 | using ux(1)[THEN bspec[where x=x]] uy(1)[THEN bspec[where x=x]] and uv(1,2) | |
| 56536 | 2832 | by auto | 
| 50804 | 2833 | } | 
| 2834 | moreover | |
| 2835 | have "(\<Sum>x\<in>s. u * ux x + v * uy x) = 1" | |
| 64267 | 2836 | unfolding sum.distrib and sum_distrib_left[symmetric] and ux(2) uy(2) | 
| 53302 | 2837 | using uv(3) by auto | 
| 50804 | 2838 | moreover | 
| 2839 | 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 | 2840 | unfolding scaleR_left_distrib and sum.distrib and scaleR_scaleR[symmetric] | 
| 2841 | and scaleR_right.sum [symmetric] | |
| 50804 | 2842 | by auto | 
| 2843 | ultimately | |
| 64267 | 2844 | show "\<exists>uc. (\<forall>x\<in>s. 0 \<le> uc x) \<and> sum uc s = 1 \<and> | 
| 50804 | 2845 | (\<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)" | 
| 2846 | apply (rule_tac x="\<lambda>x. u * ux x + v * uy x" in exI) | |
| 2847 | apply auto | |
| 2848 | done | |
| 33175 | 2849 | next | 
| 50804 | 2850 | fix t | 
| 2851 | assume t: "s \<subseteq> t" "convex t" | |
| 2852 | fix u | |
| 64267 | 2853 | assume u: "\<forall>x\<in>s. 0 \<le> u x" "sum u s = (1::real)" | 
| 50804 | 2854 | then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> t" | 
| 2855 | using t(2)[unfolded convex_explicit, THEN spec[where x=s], THEN spec[where x=u]] | |
| 33175 | 2856 | using assms and t(1) by auto | 
| 2857 | qed | |
| 2858 | ||
| 50804 | 2859 | |
| 60420 | 2860 | subsubsection \<open>Another formulation from Lars Schewe\<close> | 
| 33175 | 2861 | |
| 2862 | lemma convex_hull_explicit: | |
| 2863 | fixes p :: "'a::real_vector set" | |
| 53347 | 2864 | shows "convex hull p = | 
| 64267 | 2865 |     {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 | 2866 | (is "?lhs = ?rhs") | 
| 50804 | 2867 | proof - | 
| 53302 | 2868 |   {
 | 
| 2869 | fix x | |
| 2870 | assume "x\<in>?lhs" | |
| 50804 | 2871 | then obtain k u y where | 
| 64267 | 2872 |         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 | 2873 | unfolding convex_hull_indexed by auto | 
| 2874 | ||
| 50804 | 2875 |     have fin: "finite {1..k}" by auto
 | 
| 2876 |     have fin': "\<And>v. finite {i \<in> {1..k}. y i = v}" by auto
 | |
| 53302 | 2877 |     {
 | 
| 2878 | fix j | |
| 50804 | 2879 |       assume "j\<in>{1..k}"
 | 
| 64267 | 2880 |       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 | 2881 | using obt(1)[THEN bspec[where x=j]] and obt(2) | 
| 2882 | apply simp | |
| 64267 | 2883 | apply (rule sum_nonneg) | 
| 50804 | 2884 | using obt(1) | 
| 2885 | apply auto | |
| 2886 | done | |
| 2887 | } | |
| 33175 | 2888 | moreover | 
| 64267 | 2889 |     have "(\<Sum>v\<in>y ` {1..k}. sum u {i \<in> {1..k}. y i = v}) = 1"
 | 
| 2890 | unfolding sum_image_gen[OF fin, symmetric] using obt(2) by auto | |
| 2891 |     moreover have "(\<Sum>v\<in>y ` {1..k}. sum u {i \<in> {1..k}. y i = v} *\<^sub>R v) = x"
 | |
| 2892 | using sum_image_gen[OF fin, of "\<lambda>i. u i *\<^sub>R y i" y, symmetric] | |
| 2893 | unfolding scaleR_left.sum using obt(3) by auto | |
| 50804 | 2894 | ultimately | 
| 64267 | 2895 | 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 | 2896 |       apply (rule_tac x="y ` {1..k}" in exI)
 | 
| 64267 | 2897 |       apply (rule_tac x="\<lambda>v. sum u {i\<in>{1..k}. y i = v}" in exI)
 | 
| 50804 | 2898 | apply auto | 
| 2899 | done | |
| 2900 | then have "x\<in>?rhs" by auto | |
| 2901 | } | |
| 33175 | 2902 | moreover | 
| 53302 | 2903 |   {
 | 
| 2904 | fix y | |
| 2905 | assume "y\<in>?rhs" | |
| 50804 | 2906 | then obtain s u where | 
| 64267 | 2907 | 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 | 2908 | by auto | 
| 50804 | 2909 | |
| 2910 |     obtain f where f: "inj_on f {1..card s}" "f ` {1..card s} = s"
 | |
| 2911 | using ex_bij_betw_nat_finite_1[OF obt(1)] unfolding bij_betw_def by auto | |
| 2912 | ||
| 53302 | 2913 |     {
 | 
| 2914 | fix i :: nat | |
| 50804 | 2915 |       assume "i\<in>{1..card s}"
 | 
| 2916 | then have "f i \<in> s" | |
| 2917 | apply (subst f(2)[symmetric]) | |
| 2918 | apply auto | |
| 2919 | done | |
| 2920 | then have "0 \<le> u (f i)" "f i \<in> p" using obt(2,3) by auto | |
| 2921 | } | |
| 53347 | 2922 |     moreover have *: "finite {1..card s}" by auto
 | 
| 53302 | 2923 |     {
 | 
| 2924 | fix y | |
| 50804 | 2925 | assume "y\<in>s" | 
| 53302 | 2926 |       then obtain i where "i\<in>{1..card s}" "f i = y"
 | 
| 2927 |         using f using image_iff[of y f "{1..card s}"]
 | |
| 50804 | 2928 | by auto | 
| 2929 |       then have "{x. Suc 0 \<le> x \<and> x \<le> card s \<and> f x = y} = {i}"
 | |
| 2930 | apply auto | |
| 2931 | using f(1)[unfolded inj_on_def] | |
| 2932 | apply(erule_tac x=x in ballE) | |
| 2933 | apply auto | |
| 2934 | done | |
| 2935 |       then have "card {x. Suc 0 \<le> x \<and> x \<le> card s \<and> f x = y} = 1" by auto
 | |
| 2936 |       then have "(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x)) = u y"
 | |
| 2937 |           "(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y"
 | |
| 64267 | 2938 | by (auto simp add: sum_constant_scaleR) | 
| 50804 | 2939 | } | 
| 2940 | 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 | 2941 | unfolding sum_image_gen[OF *(1), of "\<lambda>x. u (f x) *\<^sub>R f x" f] | 
| 2942 | and sum_image_gen[OF *(1), of "\<lambda>x. u (f x)" f] | |
| 53339 | 2943 | unfolding f | 
| 64267 | 2944 |       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"]
 | 
| 2945 |       using sum.cong [of s s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x))" u]
 | |
| 53302 | 2946 | unfolding obt(4,5) | 
| 2947 | by auto | |
| 50804 | 2948 | ultimately | 
| 64267 | 2949 |     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 | 2950 | (\<Sum>i::nat = 1..k. u i *\<^sub>R x i) = y" | 
| 2951 | apply (rule_tac x="card s" in exI) | |
| 2952 | apply (rule_tac x="u \<circ> f" in exI) | |
| 2953 | apply (rule_tac x=f in exI) | |
| 2954 | apply fastforce | |
| 2955 | done | |
| 53302 | 2956 | then have "y \<in> ?lhs" | 
| 2957 | unfolding convex_hull_indexed by auto | |
| 50804 | 2958 | } | 
| 53302 | 2959 | ultimately show ?thesis | 
| 2960 | unfolding set_eq_iff by blast | |
| 33175 | 2961 | qed | 
| 2962 | ||
| 50804 | 2963 | |
| 60420 | 2964 | subsubsection \<open>A stepping theorem for that expansion\<close> | 
| 33175 | 2965 | |
| 2966 | lemma convex_hull_finite_step: | |
| 50804 | 2967 | fixes s :: "'a::real_vector set" | 
| 2968 | assumes "finite s" | |
| 53302 | 2969 | shows | 
| 64267 | 2970 | "(\<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) | 
| 2971 | \<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 | 2972 | (is "?lhs = ?rhs") | 
| 50804 | 2973 | proof (rule, case_tac[!] "a\<in>s") | 
| 53302 | 2974 | assume "a \<in> s" | 
| 53339 | 2975 | then have *: "insert a s = s" by auto | 
| 50804 | 2976 | assume ?lhs | 
| 2977 | then show ?rhs | |
| 2978 | unfolding * | |
| 2979 | apply (rule_tac x=0 in exI) | |
| 2980 | apply auto | |
| 2981 | done | |
| 33175 | 2982 | next | 
| 50804 | 2983 | assume ?lhs | 
| 53302 | 2984 | then obtain u where | 
| 64267 | 2985 | 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 | 2986 | by auto | 
| 2987 | assume "a \<notin> s" | |
| 2988 | then show ?rhs | |
| 2989 | apply (rule_tac x="u a" in exI) | |
| 2990 | using u(1)[THEN bspec[where x=a]] | |
| 2991 | apply simp | |
| 2992 | apply (rule_tac x=u in exI) | |
| 64267 | 2993 | using u[unfolded sum_clauses(2)[OF assms]] and \<open>a\<notin>s\<close> | 
| 50804 | 2994 | apply auto | 
| 2995 | done | |
| 33175 | 2996 | next | 
| 50804 | 2997 | assume "a \<in> s" | 
| 2998 | then have *: "insert a s = s" by auto | |
| 2999 | have fin: "finite (insert a s)" using assms by auto | |
| 3000 | assume ?rhs | |
| 64267 | 3001 | 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 | 3002 | by auto | 
| 3003 | show ?lhs | |
| 3004 | apply (rule_tac x = "\<lambda>x. (if a = x then v else 0) + u x" in exI) | |
| 64267 | 3005 | unfolding scaleR_left_distrib and sum.distrib and sum_delta''[OF fin] and sum.delta'[OF fin] | 
| 3006 | unfolding sum_clauses(2)[OF assms] | |
| 60420 | 3007 | using uv and uv(2)[THEN bspec[where x=a]] and \<open>a\<in>s\<close> | 
| 50804 | 3008 | apply auto | 
| 3009 | done | |
| 33175 | 3010 | next | 
| 50804 | 3011 | assume ?rhs | 
| 53339 | 3012 | then obtain v u where | 
| 64267 | 3013 | 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 | 3014 | by auto | 
| 3015 | moreover | |
| 3016 | assume "a \<notin> s" | |
| 3017 | moreover | |
| 64267 | 3018 | have "(\<Sum>x\<in>s. if a = x then v else u x) = sum u s" | 
| 53302 | 3019 | 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 | 3020 | apply (rule_tac sum.cong) apply rule | 
| 50804 | 3021 | defer | 
| 64267 | 3022 | apply (rule_tac sum.cong) apply rule | 
| 60420 | 3023 | using \<open>a \<notin> s\<close> | 
| 50804 | 3024 | apply auto | 
| 3025 | done | |
| 3026 | ultimately show ?lhs | |
| 3027 | apply (rule_tac x="\<lambda>x. if a = x then v else u x" in exI) | |
| 64267 | 3028 | unfolding sum_clauses(2)[OF assms] | 
| 50804 | 3029 | apply auto | 
| 3030 | done | |
| 3031 | qed | |
| 3032 | ||
| 33175 | 3033 | |
| 60420 | 3034 | subsubsection \<open>Hence some special cases\<close> | 
| 33175 | 3035 | |
| 3036 | lemma convex_hull_2: | |
| 3037 |   "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 | 3038 | proof - | 
| 3039 |   have *: "\<And>u. (\<forall>x\<in>{a, b}. 0 \<le> u x) \<longleftrightarrow> 0 \<le> u a \<and> 0 \<le> u b"
 | |
| 3040 | by auto | |
| 3041 |   have **: "finite {b}" by auto
 | |
| 3042 | show ?thesis | |
| 3043 | apply (simp add: convex_hull_finite) | |
| 3044 | unfolding convex_hull_finite_step[OF **, of a 1, unfolded * conj_assoc] | |
| 3045 | apply auto | |
| 3046 | apply (rule_tac x=v in exI) | |
| 3047 | apply (rule_tac x="1 - v" in exI) | |
| 3048 | apply simp | |
| 3049 | apply (rule_tac x=u in exI) | |
| 3050 | apply simp | |
| 3051 | apply (rule_tac x="\<lambda>x. v" in exI) | |
| 3052 | apply simp | |
| 3053 | done | |
| 3054 | qed | |
| 33175 | 3055 | |
| 3056 | 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: 
44142diff
changeset | 3057 | unfolding convex_hull_2 | 
| 53302 | 3058 | proof (rule Collect_cong) | 
| 3059 | have *: "\<And>x y ::real. x + y = 1 \<longleftrightarrow> x = 1 - y" | |
| 3060 | by auto | |
| 3061 | fix x | |
| 3062 | 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> | |
| 3063 | (\<exists>u. x = a + u *\<^sub>R (b - a) \<and> 0 \<le> u \<and> u \<le> 1)" | |
| 3064 | unfolding * | |
| 3065 | apply auto | |
| 3066 | apply (rule_tac[!] x=u in exI) | |
| 3067 | apply (auto simp add: algebra_simps) | |
| 3068 | done | |
| 3069 | qed | |
| 33175 | 3070 | |
| 3071 | lemma convex_hull_3: | |
| 3072 |   "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 | 3073 | proof - | 
| 3074 |   have fin: "finite {a,b,c}" "finite {b,c}" "finite {c}"
 | |
| 3075 | by auto | |
| 3076 | 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: 
44349diff
changeset | 3077 | by (auto simp add: field_simps) | 
| 53302 | 3078 | show ?thesis | 
| 3079 | unfolding convex_hull_finite[OF fin(1)] and convex_hull_finite_step[OF fin(2)] and * | |
| 3080 | unfolding convex_hull_finite_step[OF fin(3)] | |
| 3081 | apply (rule Collect_cong) | |
| 3082 | apply simp | |
| 3083 | apply auto | |
| 3084 | apply (rule_tac x=va in exI) | |
| 3085 | apply (rule_tac x="u c" in exI) | |
| 3086 | apply simp | |
| 3087 | apply (rule_tac x="1 - v - w" in exI) | |
| 3088 | apply simp | |
| 3089 | apply (rule_tac x=v in exI) | |
| 3090 | apply simp | |
| 3091 | apply (rule_tac x="\<lambda>x. w" in exI) | |
| 3092 | apply simp | |
| 3093 | done | |
| 3094 | qed | |
| 33175 | 3095 | |
| 3096 | lemma convex_hull_3_alt: | |
| 3097 |   "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 | 3098 | proof - | 
| 3099 | have *: "\<And>x y z ::real. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z" | |
| 3100 | by auto | |
| 3101 | show ?thesis | |
| 3102 | unfolding convex_hull_3 | |
| 3103 | apply (auto simp add: *) | |
| 3104 | apply (rule_tac x=v in exI) | |
| 3105 | apply (rule_tac x=w in exI) | |
| 3106 | apply (simp add: algebra_simps) | |
| 3107 | apply (rule_tac x=u in exI) | |
| 3108 | apply (rule_tac x=v in exI) | |
| 3109 | apply (simp add: algebra_simps) | |
| 3110 | done | |
| 3111 | qed | |
| 3112 | ||
| 33175 | 3113 | |
| 60420 | 3114 | subsection \<open>Relations among closure notions and corresponding hulls\<close> | 
| 33175 | 3115 | |
| 3116 | lemma affine_imp_convex: "affine s \<Longrightarrow> convex s" | |
| 3117 | unfolding affine_def convex_def by auto | |
| 3118 | ||
| 64394 | 3119 | lemma convex_affine_hull [simp]: "convex (affine hull S)" | 
| 3120 | by (simp add: affine_imp_convex) | |
| 3121 | ||
| 44361 
75ec83d45303
remove unnecessary euclidean_space class constraints
 huffman parents: 
44349diff
changeset | 3122 | lemma subspace_imp_convex: "subspace s \<Longrightarrow> convex s" | 
| 33175 | 3123 | using subspace_imp_affine affine_imp_convex by auto | 
| 3124 | ||
| 44361 
75ec83d45303
remove unnecessary euclidean_space class constraints
 huffman parents: 
44349diff
changeset | 3125 | lemma affine_hull_subset_span: "(affine hull s) \<subseteq> (span s)" | 
| 53302 | 3126 | by (metis hull_minimal span_inc subspace_imp_affine subspace_span) | 
| 33175 | 3127 | |
| 44361 
75ec83d45303
remove unnecessary euclidean_space class constraints
 huffman parents: 
44349diff
changeset | 3128 | lemma convex_hull_subset_span: "(convex hull s) \<subseteq> (span s)" | 
| 53302 | 3129 | by (metis hull_minimal span_inc subspace_imp_convex subspace_span) | 
| 33175 | 3130 | |
| 3131 | lemma convex_hull_subset_affine_hull: "(convex hull s) \<subseteq> (affine hull s)" | |
| 53302 | 3132 | by (metis affine_affine_hull affine_imp_convex hull_minimal hull_subset) | 
| 3133 | ||
| 3134 | lemma affine_dependent_imp_dependent: "affine_dependent s \<Longrightarrow> dependent s" | |
| 49531 | 3135 | unfolding affine_dependent_def dependent_def | 
| 33175 | 3136 | using affine_hull_subset_span by auto | 
| 3137 | ||
| 3138 | lemma dependent_imp_affine_dependent: | |
| 53302 | 3139 |   assumes "dependent {x - a| x . x \<in> s}"
 | 
| 3140 | and "a \<notin> s" | |
| 33175 | 3141 | shows "affine_dependent (insert a s)" | 
| 53302 | 3142 | proof - | 
| 49531 | 3143 | from assms(1)[unfolded dependent_explicit] obtain S u v | 
| 53347 | 3144 |     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"
 | 
| 3145 | by auto | |
| 63040 | 3146 | define t where "t = (\<lambda>x. x + a) ` S" | 
| 33175 | 3147 | |
| 53347 | 3148 | have inj: "inj_on (\<lambda>x. x + a) S" | 
| 53302 | 3149 | unfolding inj_on_def by auto | 
| 3150 | have "0 \<notin> S" | |
| 3151 | using obt(2) assms(2) unfolding subset_eq by auto | |
| 53347 | 3152 | have fin: "finite t" and "t \<subseteq> s" | 
| 53302 | 3153 | unfolding t_def using obt(1,2) by auto | 
| 3154 | then have "finite (insert a t)" and "insert a t \<subseteq> insert a s" | |
| 3155 | by auto | |
| 3156 | 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 | 3157 | apply (rule sum.cong) | 
| 60420 | 3158 | using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close> | 
| 53302 | 3159 | apply auto | 
| 3160 | done | |
| 33175 | 3161 | have "(\<Sum>x\<in>insert a t. if x = a then - (\<Sum>x\<in>t. u (x - a)) else u (x - a)) = 0" | 
| 64267 | 3162 | unfolding sum_clauses(2)[OF fin] | 
| 60420 | 3163 | using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close> | 
| 53302 | 3164 | apply auto | 
| 3165 | unfolding * | |
| 3166 | apply auto | |
| 3167 | done | |
| 33175 | 3168 | 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 | 3169 | apply (rule_tac x="v + a" in bexI) | 
| 60420 | 3170 | using obt(3,4) and \<open>0\<notin>S\<close> | 
| 53302 | 3171 | unfolding t_def | 
| 3172 | apply auto | |
| 3173 | done | |
| 3174 | 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 | 3175 | apply (rule sum.cong) | 
| 60420 | 3176 | using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close> | 
| 53302 | 3177 | apply auto | 
| 3178 | done | |
| 49531 | 3179 | have "(\<Sum>x\<in>t. u (x - a)) *\<^sub>R a = (\<Sum>v\<in>t. u (v - a) *\<^sub>R v)" | 
| 64267 | 3180 | unfolding scaleR_left.sum | 
| 3181 | unfolding t_def and sum.reindex[OF inj] and o_def | |
| 53302 | 3182 | using obt(5) | 
| 64267 | 3183 | by (auto simp add: sum.distrib scaleR_right_distrib) | 
| 53302 | 3184 | 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 | 3185 | unfolding sum_clauses(2)[OF fin] | 
| 60420 | 3186 | using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close> | 
| 53302 | 3187 | by (auto simp add: *) | 
| 3188 | ultimately show ?thesis | |
| 3189 | unfolding affine_dependent_explicit | |
| 3190 | apply (rule_tac x="insert a t" in exI) | |
| 3191 | apply auto | |
| 3192 | done | |
| 33175 | 3193 | qed | 
| 3194 | ||
| 3195 | lemma convex_cone: | |
| 53302 | 3196 | "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)" | 
| 3197 | (is "?lhs = ?rhs") | |
| 3198 | proof - | |
| 3199 |   {
 | |
| 3200 | fix x y | |
| 3201 | assume "x\<in>s" "y\<in>s" and ?lhs | |
| 3202 | then have "2 *\<^sub>R x \<in>s" "2 *\<^sub>R y \<in> s" | |
| 3203 | unfolding cone_def by auto | |
| 3204 | then have "x + y \<in> s" | |
| 60420 | 3205 | using \<open>?lhs\<close>[unfolded convex_def, THEN conjunct1] | 
| 53302 | 3206 | apply (erule_tac x="2*\<^sub>R x" in ballE) | 
| 3207 | apply (erule_tac x="2*\<^sub>R y" in ballE) | |
| 3208 | apply (erule_tac x="1/2" in allE) | |
| 3209 | apply simp | |
| 3210 | apply (erule_tac x="1/2" in allE) | |
| 3211 | apply auto | |
| 3212 | done | |
| 3213 | } | |
| 3214 | then show ?thesis | |
| 3215 | unfolding convex_def cone_def by blast | |
| 3216 | qed | |
| 3217 | ||
| 3218 | lemma affine_dependent_biggerset: | |
| 53347 | 3219 | 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: 
36844diff
changeset | 3220 |   assumes "finite s" "card s \<ge> DIM('a) + 2"
 | 
| 33175 | 3221 | shows "affine_dependent s" | 
| 53302 | 3222 | proof - | 
| 3223 |   have "s \<noteq> {}" using assms by auto
 | |
| 3224 | then obtain a where "a\<in>s" by auto | |
| 3225 |   have *: "{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})"
 | |
| 3226 | by auto | |
| 3227 |   have "card {x - a |x. x \<in> s - {a}} = card (s - {a})"
 | |
| 3228 | unfolding * | |
| 3229 | apply (rule card_image) | |
| 3230 | unfolding inj_on_def | |
| 3231 | apply auto | |
| 3232 | done | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 3233 |   also have "\<dots> > DIM('a)" using assms(2)
 | 
| 60420 | 3234 | unfolding card_Diff_singleton[OF assms(1) \<open>a\<in>s\<close>] by auto | 
| 53302 | 3235 | finally show ?thesis | 
| 60420 | 3236 | apply (subst insert_Diff[OF \<open>a\<in>s\<close>, symmetric]) | 
| 53302 | 3237 | apply (rule dependent_imp_affine_dependent) | 
| 3238 | apply (rule dependent_biggerset) | |
| 3239 | apply auto | |
| 3240 | done | |
| 3241 | qed | |
| 33175 | 3242 | |
| 3243 | lemma affine_dependent_biggerset_general: | |
| 53347 | 3244 | assumes "finite (s :: 'a::euclidean_space set)" | 
| 3245 | and "card s \<ge> dim s + 2" | |
| 33175 | 3246 | shows "affine_dependent s" | 
| 53302 | 3247 | proof - | 
| 33175 | 3248 |   from assms(2) have "s \<noteq> {}" by auto
 | 
| 3249 | then obtain a where "a\<in>s" by auto | |
| 53302 | 3250 |   have *: "{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})"
 | 
| 3251 | by auto | |
| 3252 |   have **: "card {x - a |x. x \<in> s - {a}} = card (s - {a})"
 | |
| 3253 | unfolding * | |
| 3254 | apply (rule card_image) | |
| 3255 | unfolding inj_on_def | |
| 3256 | apply auto | |
| 3257 | done | |
| 33175 | 3258 |   have "dim {x - a |x. x \<in> s - {a}} \<le> dim s"
 | 
| 53302 | 3259 | apply (rule subset_le_dim) | 
| 3260 | unfolding subset_eq | |
| 60420 | 3261 | using \<open>a\<in>s\<close> | 
| 63938 | 3262 | apply (auto simp add:span_superset span_diff) | 
| 53302 | 3263 | done | 
| 33175 | 3264 | also have "\<dots> < dim s + 1" by auto | 
| 53302 | 3265 |   also have "\<dots> \<le> card (s - {a})"
 | 
| 3266 | using assms | |
| 60420 | 3267 | using card_Diff_singleton[OF assms(1) \<open>a\<in>s\<close>] | 
| 53302 | 3268 | by auto | 
| 3269 | finally show ?thesis | |
| 60420 | 3270 | apply (subst insert_Diff[OF \<open>a\<in>s\<close>, symmetric]) | 
| 53302 | 3271 | apply (rule dependent_imp_affine_dependent) | 
| 3272 | apply (rule dependent_biggerset_general) | |
| 3273 | unfolding ** | |
| 3274 | apply auto | |
| 3275 | done | |
| 3276 | qed | |
| 3277 | ||
| 33175 | 3278 | |
| 60420 | 3279 | subsection \<open>Some Properties of Affine Dependent Sets\<close> | 
| 40377 | 3280 | |
| 66287 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65719diff
changeset | 3281 | lemma affine_independent_0 [simp]: "\<not> affine_dependent {}"
 | 
| 40377 | 3282 | by (simp add: affine_dependent_def) | 
| 3283 | ||
| 66287 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65719diff
changeset | 3284 | lemma affine_independent_1 [simp]: "\<not> affine_dependent {a}"
 | 
| 53302 | 3285 | by (simp add: affine_dependent_def) | 
| 3286 | ||
| 66287 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65719diff
changeset | 3287 | lemma affine_independent_2 [simp]: "\<not> affine_dependent {a,b}"
 | 
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 3288 | 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: 
62950diff
changeset | 3289 | |
| 53302 | 3290 | lemma affine_hull_translation: "affine hull ((\<lambda>x. a + x) ` S) = (\<lambda>x. a + x) ` (affine hull S)" | 
| 3291 | proof - | |
| 3292 | have "affine ((\<lambda>x. a + x) ` (affine hull S))" | |
| 60303 | 3293 | using affine_translation affine_affine_hull by blast | 
| 53347 | 3294 | moreover have "(\<lambda>x. a + x) ` S \<subseteq> (\<lambda>x. a + x) ` (affine hull S)" | 
| 53302 | 3295 | using hull_subset[of S] by auto | 
| 53347 | 3296 | ultimately have h1: "affine hull ((\<lambda>x. a + x) ` S) \<subseteq> (\<lambda>x. a + x) ` (affine hull S)" | 
| 53302 | 3297 | by (metis hull_minimal) | 
| 3298 | have "affine((\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) ` S)))" | |
| 60303 | 3299 | using affine_translation affine_affine_hull by blast | 
| 53347 | 3300 | moreover have "(\<lambda>x. -a + x) ` (\<lambda>x. a + x) ` S \<subseteq> (\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) ` S))" | 
| 53302 | 3301 | using hull_subset[of "(\<lambda>x. a + x) ` S"] by auto | 
| 53347 | 3302 | moreover have "S = (\<lambda>x. -a + x) ` (\<lambda>x. a + x) ` S" | 
| 53302 | 3303 | using translation_assoc[of "-a" a] by auto | 
| 3304 | ultimately have "(\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) ` S)) >= (affine hull S)" | |
| 3305 | by (metis hull_minimal) | |
| 3306 | then have "affine hull ((\<lambda>x. a + x) ` S) >= (\<lambda>x. a + x) ` (affine hull S)" | |
| 3307 | by auto | |
| 54465 | 3308 | then show ?thesis using h1 by auto | 
| 40377 | 3309 | qed | 
| 3310 | ||
| 3311 | lemma affine_dependent_translation: | |
| 3312 | assumes "affine_dependent S" | |
| 53339 | 3313 | shows "affine_dependent ((\<lambda>x. a + x) ` S)" | 
| 53302 | 3314 | proof - | 
| 54465 | 3315 |   obtain x where x: "x \<in> S \<and> x \<in> affine hull (S - {x})"
 | 
| 53302 | 3316 | using assms affine_dependent_def by auto | 
| 3317 |   have "op + a ` (S - {x}) = op + a ` S - {a + x}"
 | |
| 3318 | by auto | |
| 53347 | 3319 |   then have "a + x \<in> affine hull ((\<lambda>x. a + x) ` S - {a + x})"
 | 
| 54465 | 3320 |     using affine_hull_translation[of a "S - {x}"] x by auto
 | 
| 53347 | 3321 | moreover have "a + x \<in> (\<lambda>x. a + x) ` S" | 
| 54465 | 3322 | using x by auto | 
| 53302 | 3323 | ultimately show ?thesis | 
| 3324 | unfolding affine_dependent_def by auto | |
| 40377 | 3325 | qed | 
| 3326 | ||
| 3327 | lemma affine_dependent_translation_eq: | |
| 54465 | 3328 | "affine_dependent S \<longleftrightarrow> affine_dependent ((\<lambda>x. a + x) ` S)" | 
| 53302 | 3329 | proof - | 
| 3330 |   {
 | |
| 53339 | 3331 | assume "affine_dependent ((\<lambda>x. a + x) ` S)" | 
| 53302 | 3332 | then have "affine_dependent S" | 
| 53339 | 3333 | using affine_dependent_translation[of "((\<lambda>x. a + x) ` S)" "-a"] translation_assoc[of "-a" a] | 
| 53302 | 3334 | by auto | 
| 3335 | } | |
| 3336 | then show ?thesis | |
| 3337 | using affine_dependent_translation by auto | |
| 40377 | 3338 | qed | 
| 3339 | ||
| 3340 | lemma affine_hull_0_dependent: | |
| 53339 | 3341 | assumes "0 \<in> affine hull S" | 
| 40377 | 3342 | shows "dependent S" | 
| 53302 | 3343 | proof - | 
| 64267 | 3344 |   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 | 3345 | using assms affine_hull_explicit[of S] by auto | 
| 53339 | 3346 | then have "\<exists>v\<in>s. u v \<noteq> 0" | 
| 64267 | 3347 | using sum_not_0[of "u" "s"] by auto | 
| 53339 | 3348 | 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 | 3349 | using s_u by auto | 
| 53302 | 3350 | then show ?thesis | 
| 3351 | unfolding dependent_explicit[of S] by auto | |
| 40377 | 3352 | qed | 
| 3353 | ||
| 3354 | lemma affine_dependent_imp_dependent2: | |
| 3355 | assumes "affine_dependent (insert 0 S)" | |
| 3356 | shows "dependent S" | |
| 53302 | 3357 | proof - | 
| 54465 | 3358 |   obtain x where x: "x \<in> insert 0 S \<and> x \<in> affine hull (insert 0 S - {x})"
 | 
| 53302 | 3359 | using affine_dependent_def[of "(insert 0 S)"] assms by blast | 
| 3360 |   then have "x \<in> span (insert 0 S - {x})"
 | |
| 3361 | using affine_hull_subset_span by auto | |
| 3362 |   moreover have "span (insert 0 S - {x}) = span (S - {x})"
 | |
| 3363 |     using insert_Diff_if[of "0" S "{x}"] span_insert_0[of "S-{x}"] by auto
 | |
| 3364 |   ultimately have "x \<in> span (S - {x})" by auto
 | |
| 3365 | then have "x \<noteq> 0 \<Longrightarrow> dependent S" | |
| 54465 | 3366 | using x dependent_def by auto | 
| 53302 | 3367 | moreover | 
| 3368 |   {
 | |
| 3369 | assume "x = 0" | |
| 3370 | then have "0 \<in> affine hull S" | |
| 54465 | 3371 |       using x hull_mono[of "S - {0}" S] by auto
 | 
| 53302 | 3372 | then have "dependent S" | 
| 3373 | using affine_hull_0_dependent by auto | |
| 3374 | } | |
| 3375 | ultimately show ?thesis by auto | |
| 40377 | 3376 | qed | 
| 3377 | ||
| 3378 | lemma affine_dependent_iff_dependent: | |
| 53302 | 3379 | assumes "a \<notin> S" | 
| 3380 | shows "affine_dependent (insert a S) \<longleftrightarrow> dependent ((\<lambda>x. -a + x) ` S)" | |
| 3381 | proof - | |
| 3382 |   have "(op + (- a) ` S) = {x - a| x . x : S}" by auto
 | |
| 3383 | then show ?thesis | |
| 3384 | using affine_dependent_translation_eq[of "(insert a S)" "-a"] | |
| 49531 | 3385 | affine_dependent_imp_dependent2 assms | 
| 53302 | 3386 | dependent_imp_affine_dependent[of a S] | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53676diff
changeset | 3387 | by (auto simp del: uminus_add_conv_diff) | 
| 40377 | 3388 | qed | 
| 3389 | ||
| 3390 | lemma affine_dependent_iff_dependent2: | |
| 53339 | 3391 | assumes "a \<in> S" | 
| 3392 |   shows "affine_dependent S \<longleftrightarrow> dependent ((\<lambda>x. -a + x) ` (S-{a}))"
 | |
| 53302 | 3393 | proof - | 
| 53339 | 3394 |   have "insert a (S - {a}) = S"
 | 
| 53302 | 3395 | using assms by auto | 
| 3396 | then show ?thesis | |
| 3397 |     using assms affine_dependent_iff_dependent[of a "S-{a}"] by auto
 | |
| 40377 | 3398 | qed | 
| 3399 | ||
| 3400 | lemma affine_hull_insert_span_gen: | |
| 53339 | 3401 | "affine hull (insert a s) = (\<lambda>x. a + x) ` span ((\<lambda>x. - a + x) ` s)" | 
| 53302 | 3402 | proof - | 
| 53339 | 3403 |   have h1: "{x - a |x. x \<in> s} = ((\<lambda>x. -a+x) ` s)"
 | 
| 53302 | 3404 | by auto | 
| 3405 |   {
 | |
| 3406 | assume "a \<notin> s" | |
| 3407 | then have ?thesis | |
| 3408 | using affine_hull_insert_span[of a s] h1 by auto | |
| 3409 | } | |
| 3410 | moreover | |
| 3411 |   {
 | |
| 3412 | assume a1: "a \<in> s" | |
| 53339 | 3413 | have "\<exists>x. x \<in> s \<and> -a+x=0" | 
| 53302 | 3414 | apply (rule exI[of _ a]) | 
| 3415 | using a1 | |
| 3416 | apply auto | |
| 3417 | done | |
| 53339 | 3418 |     then have "insert 0 ((\<lambda>x. -a+x) ` (s - {a})) = (\<lambda>x. -a+x) ` s"
 | 
| 53302 | 3419 | by auto | 
| 53339 | 3420 |     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: 
53676diff
changeset | 3421 |       using span_insert_0[of "op + (- a) ` (s - {a})"] by (auto simp del: uminus_add_conv_diff)
 | 
| 53339 | 3422 |     moreover have "{x - a |x. x \<in> (s - {a})} = ((\<lambda>x. -a+x) ` (s - {a}))"
 | 
| 53302 | 3423 | by auto | 
| 53339 | 3424 |     moreover have "insert a (s - {a}) = insert a s"
 | 
| 63092 | 3425 | by auto | 
| 53302 | 3426 | ultimately have ?thesis | 
| 63092 | 3427 |       using affine_hull_insert_span[of "a" "s-{a}"] by auto
 | 
| 53302 | 3428 | } | 
| 3429 | ultimately show ?thesis by auto | |
| 40377 | 3430 | qed | 
| 3431 | ||
| 3432 | lemma affine_hull_span2: | |
| 53302 | 3433 | assumes "a \<in> s" | 
| 3434 |   shows "affine hull s = (\<lambda>x. a+x) ` span ((\<lambda>x. -a+x) ` (s-{a}))"
 | |
| 3435 |   using affine_hull_insert_span_gen[of a "s - {a}", unfolded insert_Diff[OF assms]]
 | |
| 3436 | by auto | |
| 40377 | 3437 | |
| 3438 | lemma affine_hull_span_gen: | |
| 53339 | 3439 | assumes "a \<in> affine hull s" | 
| 3440 | shows "affine hull s = (\<lambda>x. a+x) ` span ((\<lambda>x. -a+x) ` s)" | |
| 53302 | 3441 | proof - | 
| 3442 | have "affine hull (insert a s) = affine hull s" | |
| 3443 | using hull_redundant[of a affine s] assms by auto | |
| 3444 | then show ?thesis | |
| 3445 | using affine_hull_insert_span_gen[of a "s"] by auto | |
| 40377 | 3446 | qed | 
| 3447 | ||
| 3448 | lemma affine_hull_span_0: | |
| 53339 | 3449 | assumes "0 \<in> affine hull S" | 
| 40377 | 3450 | shows "affine hull S = span S" | 
| 53302 | 3451 | using affine_hull_span_gen[of "0" S] assms by auto | 
| 40377 | 3452 | |
| 63016 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3453 | lemma extend_to_affine_basis_nonempty: | 
| 53339 | 3454 | fixes S V :: "'n::euclidean_space set" | 
| 3455 |   assumes "\<not> affine_dependent S" "S \<subseteq> V" "S \<noteq> {}"
 | |
| 3456 | shows "\<exists>T. \<not> affine_dependent T \<and> S \<subseteq> T \<and> T \<subseteq> V \<and> affine hull T = affine hull V" | |
| 53302 | 3457 | proof - | 
| 54465 | 3458 | obtain a where a: "a \<in> S" | 
| 53302 | 3459 | using assms by auto | 
| 53339 | 3460 |   then have h0: "independent  ((\<lambda>x. -a + x) ` (S-{a}))"
 | 
| 53302 | 3461 | using affine_dependent_iff_dependent2 assms by auto | 
| 54465 | 3462 | then obtain B where B: | 
| 53339 | 3463 |     "(\<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"
 | 
| 3464 |      using maximal_independent_subset_extend[of "(\<lambda>x. -a+x) ` (S-{a})" "(\<lambda>x. -a + x) ` V"] assms
 | |
| 53302 | 3465 | by blast | 
| 63040 | 3466 | define T where "T = (\<lambda>x. a+x) ` insert 0 B" | 
| 53339 | 3467 | then have "T = insert a ((\<lambda>x. a+x) ` B)" | 
| 3468 | by auto | |
| 3469 | then have "affine hull T = (\<lambda>x. a+x) ` span B" | |
| 3470 | using affine_hull_insert_span_gen[of a "((\<lambda>x. a+x) ` B)"] translation_assoc[of "-a" a B] | |
| 53302 | 3471 | by auto | 
| 53347 | 3472 | then have "V \<subseteq> affine hull T" | 
| 54465 | 3473 | using B assms translation_inverse_subset[of a V "span B"] | 
| 53302 | 3474 | by auto | 
| 53339 | 3475 | moreover have "T \<subseteq> V" | 
| 54465 | 3476 | using T_def B a assms by auto | 
| 53302 | 3477 | ultimately have "affine hull T = affine hull V" | 
| 44457 
d366fa5551ef
declare euclidean_simps [simp] at the point they are proved;
 huffman parents: 
44365diff
changeset | 3478 | by (metis Int_absorb1 Int_absorb2 hull_hull hull_mono) | 
| 53347 | 3479 | moreover have "S \<subseteq> T" | 
| 54465 | 3480 |     using T_def B translation_inverse_subset[of a "S-{a}" B]
 | 
| 53302 | 3481 | by auto | 
| 3482 | moreover have "\<not> affine_dependent T" | |
| 53339 | 3483 | using T_def affine_dependent_translation_eq[of "insert 0 B"] | 
| 54465 | 3484 | affine_dependent_imp_dependent2 B | 
| 53302 | 3485 | by auto | 
| 60420 | 3486 | ultimately show ?thesis using \<open>T \<subseteq> V\<close> by auto | 
| 40377 | 3487 | qed | 
| 3488 | ||
| 49531 | 3489 | lemma affine_basis_exists: | 
| 53339 | 3490 | fixes V :: "'n::euclidean_space set" | 
| 3491 | shows "\<exists>B. B \<subseteq> V \<and> \<not> affine_dependent B \<and> affine hull V = affine hull B" | |
| 53302 | 3492 | proof (cases "V = {}")
 | 
| 3493 | case True | |
| 3494 | then show ?thesis | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 3495 | using affine_independent_0 by auto | 
| 53302 | 3496 | next | 
| 3497 | case False | |
| 3498 | then obtain x where "x \<in> V" by auto | |
| 3499 | then show ?thesis | |
| 63016 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3500 |     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: 
63007diff
changeset | 3501 | by auto | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3502 | qed | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3503 | |
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3504 | proposition extend_to_affine_basis: | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3505 | fixes S V :: "'n::euclidean_space set" | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3506 | assumes "\<not> affine_dependent S" "S \<subseteq> V" | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3507 | 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: 
63007diff
changeset | 3508 | proof (cases "S = {}")
 | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3509 | case True then show ?thesis | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3510 | using affine_basis_exists by (metis empty_subsetI that) | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3511 | next | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3512 | case False | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3513 | then show ?thesis by (metis assms extend_to_affine_basis_nonempty that) | 
| 53302 | 3514 | qed | 
| 3515 | ||
| 40377 | 3516 | |
| 60420 | 3517 | subsection \<open>Affine Dimension of a Set\<close> | 
| 40377 | 3518 | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3519 | definition aff_dim :: "('a::euclidean_space) set \<Rightarrow> int"
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3520 | where "aff_dim V = | 
| 53339 | 3521 | (SOME d :: int. | 
| 3522 | \<exists>B. affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> of_nat (card B) = d + 1)" | |
| 40377 | 3523 | |
| 3524 | lemma aff_dim_basis_exists: | |
| 49531 | 3525 |   fixes V :: "('n::euclidean_space) set"
 | 
| 53339 | 3526 | shows "\<exists>B. affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> of_nat (card B) = aff_dim V + 1" | 
| 53302 | 3527 | proof - | 
| 53347 | 3528 | obtain B where "\<not> affine_dependent B \<and> affine hull B = affine hull V" | 
| 53302 | 3529 | using affine_basis_exists[of V] by auto | 
| 3530 | then show ?thesis | |
| 53339 | 3531 | unfolding aff_dim_def | 
| 53347 | 3532 | 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 | 3533 | apply auto | 
| 53339 | 3534 | apply (rule exI[of _ "int (card B) - (1 :: int)"]) | 
| 53302 | 3535 | apply (rule exI[of _ "B"]) | 
| 3536 | apply auto | |
| 3537 | done | |
| 3538 | qed | |
| 3539 | ||
| 3540 | lemma affine_hull_nonempty: "S \<noteq> {} \<longleftrightarrow> affine hull S \<noteq> {}"
 | |
| 3541 | proof - | |
| 3542 |   have "S = {} \<Longrightarrow> affine hull S = {}"
 | |
| 3543 | using affine_hull_empty by auto | |
| 3544 |   moreover have "affine hull S = {} \<Longrightarrow> S = {}"
 | |
| 3545 | unfolding hull_def by auto | |
| 3546 | ultimately show ?thesis by blast | |
| 40377 | 3547 | qed | 
| 3548 | ||
| 3549 | lemma aff_dim_parallel_subspace_aux: | |
| 53347 | 3550 | fixes B :: "'n::euclidean_space set" | 
| 53302 | 3551 | assumes "\<not> affine_dependent B" "a \<in> B" | 
| 53339 | 3552 |   shows "finite B \<and> ((card B) - 1 = dim (span ((\<lambda>x. -a+x) ` (B-{a}))))"
 | 
| 53302 | 3553 | proof - | 
| 53339 | 3554 |   have "independent ((\<lambda>x. -a + x) ` (B-{a}))"
 | 
| 53302 | 3555 | using affine_dependent_iff_dependent2 assms by auto | 
| 53339 | 3556 |   then have fin: "dim (span ((\<lambda>x. -a+x) ` (B-{a}))) = card ((\<lambda>x. -a + x) ` (B-{a}))"
 | 
| 3557 |     "finite ((\<lambda>x. -a + x) ` (B - {a}))"
 | |
| 53347 | 3558 |     using indep_card_eq_dim_span[of "(\<lambda>x. -a+x) ` (B-{a})"] by auto
 | 
| 53302 | 3559 | show ?thesis | 
| 53339 | 3560 |   proof (cases "(\<lambda>x. -a + x) ` (B - {a}) = {}")
 | 
| 53302 | 3561 | case True | 
| 53339 | 3562 |     have "B = insert a ((\<lambda>x. a + x) ` (\<lambda>x. -a + x) ` (B - {a}))"
 | 
| 53302 | 3563 |       using translation_assoc[of "a" "-a" "(B - {a})"] assms by auto
 | 
| 53339 | 3564 |     then have "B = {a}" using True by auto
 | 
| 53302 | 3565 | then show ?thesis using assms fin by auto | 
| 3566 | next | |
| 3567 | case False | |
| 53339 | 3568 |     then have "card ((\<lambda>x. -a + x) ` (B - {a})) > 0"
 | 
| 53302 | 3569 | using fin by auto | 
| 53339 | 3570 |     moreover have h1: "card ((\<lambda>x. -a + x) ` (B-{a})) = card (B-{a})"
 | 
| 53302 | 3571 | apply (rule card_image) | 
| 3572 | using translate_inj_on | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53676diff
changeset | 3573 | apply (auto simp del: uminus_add_conv_diff) | 
| 53302 | 3574 | done | 
| 53339 | 3575 |     ultimately have "card (B-{a}) > 0" by auto
 | 
| 3576 |     then have *: "finite (B - {a})"
 | |
| 53302 | 3577 |       using card_gt_0_iff[of "(B - {a})"] by auto
 | 
| 53339 | 3578 |     then have "card (B - {a}) = card B - 1"
 | 
| 53302 | 3579 | using card_Diff_singleton assms by auto | 
| 3580 | with * show ?thesis using fin h1 by auto | |
| 3581 | qed | |
| 40377 | 3582 | qed | 
| 3583 | ||
| 3584 | lemma aff_dim_parallel_subspace: | |
| 53339 | 3585 | fixes V L :: "'n::euclidean_space set" | 
| 53302 | 3586 |   assumes "V \<noteq> {}"
 | 
| 53339 | 3587 | and "subspace L" | 
| 3588 | and "affine_parallel (affine hull V) L" | |
| 53302 | 3589 | shows "aff_dim V = int (dim L)" | 
| 3590 | proof - | |
| 53339 | 3591 | obtain B where | 
| 54465 | 3592 | B: "affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> int (card B) = aff_dim V + 1" | 
| 53302 | 3593 | using aff_dim_basis_exists by auto | 
| 3594 |   then have "B \<noteq> {}"
 | |
| 54465 | 3595 | using assms B affine_hull_nonempty[of V] affine_hull_nonempty[of B] | 
| 53302 | 3596 | by auto | 
| 54465 | 3597 | then obtain a where a: "a \<in> B" by auto | 
| 63040 | 3598 |   define Lb where "Lb = span ((\<lambda>x. -a+x) ` (B-{a}))"
 | 
| 40377 | 3599 | moreover have "affine_parallel (affine hull B) Lb" | 
| 54465 | 3600 | using Lb_def B assms affine_hull_span2[of a B] a | 
| 53339 | 3601 | affine_parallel_commut[of "Lb" "(affine hull B)"] | 
| 3602 | unfolding affine_parallel_def | |
| 3603 | by auto | |
| 53302 | 3604 | moreover have "subspace Lb" | 
| 3605 | using Lb_def subspace_span by auto | |
| 3606 |   moreover have "affine hull B \<noteq> {}"
 | |
| 54465 | 3607 | using assms B affine_hull_nonempty[of V] by auto | 
| 53302 | 3608 | ultimately have "L = Lb" | 
| 54465 | 3609 | using assms affine_parallel_subspace[of "affine hull B"] affine_affine_hull[of B] B | 
| 53302 | 3610 | by auto | 
| 53339 | 3611 | then have "dim L = dim Lb" | 
| 3612 | by auto | |
| 3613 | moreover have "card B - 1 = dim Lb" and "finite B" | |
| 54465 | 3614 | using Lb_def aff_dim_parallel_subspace_aux a B by auto | 
| 53302 | 3615 | ultimately show ?thesis | 
| 60420 | 3616 |     using B \<open>B \<noteq> {}\<close> card_gt_0_iff[of B] by auto
 | 
| 40377 | 3617 | qed | 
| 3618 | ||
| 3619 | lemma aff_independent_finite: | |
| 53339 | 3620 | fixes B :: "'n::euclidean_space set" | 
| 3621 | assumes "\<not> affine_dependent B" | |
| 53302 | 3622 | shows "finite B" | 
| 3623 | proof - | |
| 3624 |   {
 | |
| 3625 |     assume "B \<noteq> {}"
 | |
| 3626 | then obtain a where "a \<in> B" by auto | |
| 3627 | then have ?thesis | |
| 3628 | using aff_dim_parallel_subspace_aux assms by auto | |
| 3629 | } | |
| 3630 | then show ?thesis by auto | |
| 40377 | 3631 | qed | 
| 3632 | ||
| 3633 | lemma independent_finite: | |
| 53339 | 3634 | fixes B :: "'n::euclidean_space set" | 
| 53302 | 3635 | assumes "independent B" | 
| 3636 | shows "finite B" | |
| 3637 | using affine_dependent_imp_dependent[of B] aff_independent_finite[of B] assms | |
| 3638 | by auto | |
| 40377 | 3639 | |
| 3640 | lemma subspace_dim_equal: | |
| 53339 | 3641 |   assumes "subspace (S :: ('n::euclidean_space) set)"
 | 
| 3642 | and "subspace T" | |
| 3643 | and "S \<subseteq> T" | |
| 3644 | and "dim S \<ge> dim T" | |
| 53302 | 3645 | shows "S = T" | 
| 3646 | proof - | |
| 53347 | 3647 | obtain B where B: "B \<le> S" "independent B \<and> S \<subseteq> span B" "card B = dim S" | 
| 53339 | 3648 | using basis_exists[of S] by auto | 
| 3649 | then have "span B \<subseteq> S" | |
| 3650 | using span_mono[of B S] span_eq[of S] assms by metis | |
| 3651 | then have "span B = S" | |
| 53347 | 3652 | using B by auto | 
| 53339 | 3653 | have "dim S = dim T" | 
| 3654 | using assms dim_subset[of S T] by auto | |
| 3655 | then have "T \<subseteq> span B" | |
| 53347 | 3656 | using card_eq_dim[of B T] B independent_finite assms by auto | 
| 53339 | 3657 | then show ?thesis | 
| 60420 | 3658 | using assms \<open>span B = S\<close> by auto | 
| 40377 | 3659 | qed | 
| 3660 | ||
| 63016 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3661 | corollary dim_eq_span: | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3662 | fixes S :: "'a::euclidean_space set" | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3663 | 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: 
63007diff
changeset | 3664 | by (simp add: span_mono subspace_dim_equal subspace_span) | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3665 | |
| 63075 
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
 paulson <lp15@cam.ac.uk> parents: 
63072diff
changeset | 3666 | lemma dim_eq_full: | 
| 
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
 paulson <lp15@cam.ac.uk> parents: 
63072diff
changeset | 3667 | fixes S :: "'a :: euclidean_space set" | 
| 
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
 paulson <lp15@cam.ac.uk> parents: 
63072diff
changeset | 3668 |     shows "dim S = DIM('a) \<longleftrightarrow> span S = UNIV"
 | 
| 
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
 paulson <lp15@cam.ac.uk> parents: 
63072diff
changeset | 3669 | apply (rule iffI) | 
| 
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
 paulson <lp15@cam.ac.uk> parents: 
63072diff
changeset | 3670 | 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: 
63072diff
changeset | 3671 | by (metis dim_UNIV dim_span) | 
| 
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
 paulson <lp15@cam.ac.uk> parents: 
63072diff
changeset | 3672 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 3673 | 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: 
50104diff
changeset | 3674 | assumes d: "d \<subseteq> Basis" | 
| 53347 | 3675 |   shows "span d = {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
 | 
| 3676 | (is "_ = ?B") | |
| 53339 | 3677 | proof - | 
| 3678 | have "d \<subseteq> ?B" | |
| 3679 | using d by (auto simp: inner_Basis) | |
| 3680 | moreover have s: "subspace ?B" | |
| 3681 | using subspace_substandard[of "\<lambda>i. i \<notin> d"] . | |
| 3682 | ultimately have "span d \<subseteq> ?B" | |
| 3683 | using span_mono[of d "?B"] span_eq[of "?B"] by blast | |
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53348diff
changeset | 3684 | moreover have *: "card d \<le> dim (span d)" | 
| 53339 | 3685 | using independent_card_le_dim[of d "span d"] independent_substdbasis[OF assms] span_inc[of d] | 
| 3686 | by auto | |
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53348diff
changeset | 3687 | moreover from * have "dim ?B \<le> dim (span d)" | 
| 53339 | 3688 | using dim_substandard[OF assms] by auto | 
| 3689 | ultimately show ?thesis | |
| 3690 | using s subspace_dim_equal[of "span d" "?B"] subspace_span[of d] by auto | |
| 40377 | 3691 | qed | 
| 3692 | ||
| 3693 | lemma basis_to_substdbasis_subspace_isomorphism: | |
| 53339 | 3694 | fixes B :: "'a::euclidean_space set" | 
| 3695 | assumes "independent B" | |
| 3696 | shows "\<exists>f d::'a set. card d = card B \<and> linear f \<and> f ` B = d \<and> | |
| 3697 |     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"
 | |
| 3698 | proof - | |
| 3699 | have B: "card B = dim B" | |
| 3700 | using dim_unique[of B B "card B"] assms span_inc[of B] by auto | |
| 3701 | have "dim B \<le> card (Basis :: 'a set)" | |
| 3702 | using dim_subset_UNIV[of B] by simp | |
| 3703 | from ex_card[OF this] obtain d :: "'a set" where d: "d \<subseteq> Basis" and t: "card d = dim B" | |
| 3704 | by auto | |
| 53347 | 3705 |   let ?t = "{x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
 | 
| 53339 | 3706 | 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: 
50104diff
changeset | 3707 | apply (rule basis_to_basis_subspace_isomorphism[of "span B" ?t B "d"]) | 
| 53339 | 3708 | apply (rule subspace_span) | 
| 3709 | apply (rule subspace_substandard) | |
| 3710 | defer | |
| 3711 | apply (rule span_inc) | |
| 3712 | apply (rule assms) | |
| 3713 | defer | |
| 3714 | unfolding dim_span[of B] | |
| 3715 | apply(rule B) | |
| 54465 | 3716 | unfolding span_substd_basis[OF d, symmetric] | 
| 53339 | 3717 | apply (rule span_inc) | 
| 3718 | apply (rule independent_substdbasis[OF d]) | |
| 3719 | apply rule | |
| 3720 | apply assumption | |
| 3721 | unfolding t[symmetric] span_substd_basis[OF d] dim_substandard[OF d] | |
| 3722 | apply auto | |
| 3723 | done | |
| 60420 | 3724 | with t \<open>card B = dim B\<close> d show ?thesis by auto | 
| 40377 | 3725 | qed | 
| 3726 | ||
| 3727 | lemma aff_dim_empty: | |
| 53339 | 3728 | fixes S :: "'n::euclidean_space set" | 
| 3729 |   shows "S = {} \<longleftrightarrow> aff_dim S = -1"
 | |
| 3730 | proof - | |
| 3731 | obtain B where *: "affine hull B = affine hull S" | |
| 3732 | and "\<not> affine_dependent B" | |
| 3733 | and "int (card B) = aff_dim S + 1" | |
| 3734 | using aff_dim_basis_exists by auto | |
| 3735 | moreover | |
| 3736 |   from * have "S = {} \<longleftrightarrow> B = {}"
 | |
| 3737 | using affine_hull_nonempty[of B] affine_hull_nonempty[of S] by auto | |
| 3738 | ultimately show ?thesis | |
| 3739 | using aff_independent_finite[of B] card_gt_0_iff[of B] by auto | |
| 3740 | qed | |
| 3741 | ||
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3742 | 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: 
61426diff
changeset | 3743 | by (simp add: aff_dim_empty [symmetric]) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3744 | |
| 63075 
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
 paulson <lp15@cam.ac.uk> parents: 
63072diff
changeset | 3745 | lemma aff_dim_affine_hull [simp]: "aff_dim (affine hull S) = aff_dim S" | 
| 53339 | 3746 | unfolding aff_dim_def using hull_hull[of _ S] by auto | 
| 40377 | 3747 | |
| 3748 | lemma aff_dim_affine_hull2: | |
| 53339 | 3749 | assumes "affine hull S = affine hull T" | 
| 3750 | shows "aff_dim S = aff_dim T" | |
| 3751 | unfolding aff_dim_def using assms by auto | |
| 40377 | 3752 | |
| 49531 | 3753 | lemma aff_dim_unique: | 
| 53339 | 3754 | fixes B V :: "'n::euclidean_space set" | 
| 3755 | assumes "affine hull B = affine hull V \<and> \<not> affine_dependent B" | |
| 3756 | shows "of_nat (card B) = aff_dim V + 1" | |
| 3757 | proof (cases "B = {}")
 | |
| 3758 | case True | |
| 3759 |   then have "V = {}"
 | |
| 3760 | using affine_hull_nonempty[of V] affine_hull_nonempty[of B] assms | |
| 3761 | by auto | |
| 3762 | then have "aff_dim V = (-1::int)" | |
| 3763 | using aff_dim_empty by auto | |
| 3764 | then show ?thesis | |
| 60420 | 3765 |     using \<open>B = {}\<close> by auto
 | 
| 53339 | 3766 | next | 
| 3767 | case False | |
| 54465 | 3768 | then obtain a where a: "a \<in> B" by auto | 
| 63040 | 3769 |   define Lb where "Lb = span ((\<lambda>x. -a+x) ` (B-{a}))"
 | 
| 40377 | 3770 | have "affine_parallel (affine hull B) Lb" | 
| 54465 | 3771 | using Lb_def affine_hull_span2[of a B] a | 
| 53339 | 3772 | affine_parallel_commut[of "Lb" "(affine hull B)"] | 
| 3773 | unfolding affine_parallel_def by auto | |
| 3774 | moreover have "subspace Lb" | |
| 3775 | using Lb_def subspace_span by auto | |
| 3776 | ultimately have "aff_dim B = int(dim Lb)" | |
| 60420 | 3777 |     using aff_dim_parallel_subspace[of B Lb] \<open>B \<noteq> {}\<close> by auto
 | 
| 53339 | 3778 | moreover have "(card B) - 1 = dim Lb" "finite B" | 
| 54465 | 3779 | using Lb_def aff_dim_parallel_subspace_aux a assms by auto | 
| 53339 | 3780 | ultimately have "of_nat (card B) = aff_dim B + 1" | 
| 60420 | 3781 |     using \<open>B \<noteq> {}\<close> card_gt_0_iff[of B] by auto
 | 
| 53339 | 3782 | then show ?thesis | 
| 3783 | using aff_dim_affine_hull2 assms by auto | |
| 40377 | 3784 | qed | 
| 3785 | ||
| 49531 | 3786 | lemma aff_dim_affine_independent: | 
| 53339 | 3787 | fixes B :: "'n::euclidean_space set" | 
| 3788 | assumes "\<not> affine_dependent B" | |
| 3789 | shows "of_nat (card B) = aff_dim B + 1" | |
| 40377 | 3790 | using aff_dim_unique[of B B] assms by auto | 
| 3791 | ||
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3792 | lemma affine_independent_iff_card: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3793 | fixes s :: "'a::euclidean_space set" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3794 | 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: 
61426diff
changeset | 3795 | apply (rule iffI) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3796 | apply (simp add: aff_dim_affine_independent aff_independent_finite) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3797 | 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: 
61426diff
changeset | 3798 | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 3799 | lemma aff_dim_sing [simp]: | 
| 53339 | 3800 | fixes a :: "'n::euclidean_space" | 
| 3801 |   shows "aff_dim {a} = 0"
 | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 3802 |   using aff_dim_affine_independent[of "{a}"] affine_independent_1 by auto
 | 
| 40377 | 3803 | |
| 63881 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3804 | 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: 
63627diff
changeset | 3805 | proof (clarsimp) | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3806 | assume "a \<noteq> b" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3807 |   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: 
63627diff
changeset | 3808 | 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: 
63627diff
changeset | 3809 | also have "... = 1" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3810 | using \<open>a \<noteq> b\<close> by simp | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3811 |   finally show "aff_dim {a, b} = 1" .
 | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3812 | qed | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3813 | |
| 40377 | 3814 | lemma aff_dim_inner_basis_exists: | 
| 49531 | 3815 |   fixes V :: "('n::euclidean_space) set"
 | 
| 53339 | 3816 | shows "\<exists>B. B \<subseteq> V \<and> affine hull B = affine hull V \<and> | 
| 3817 | \<not> affine_dependent B \<and> of_nat (card B) = aff_dim V + 1" | |
| 3818 | proof - | |
| 53347 | 3819 | obtain B where B: "\<not> affine_dependent B" "B \<subseteq> V" "affine hull B = affine hull V" | 
| 53339 | 3820 | using affine_basis_exists[of V] by auto | 
| 3821 | then have "of_nat(card B) = aff_dim V+1" using aff_dim_unique by auto | |
| 53347 | 3822 | with B show ?thesis by auto | 
| 40377 | 3823 | qed | 
| 3824 | ||
| 3825 | lemma aff_dim_le_card: | |
| 53347 | 3826 | fixes V :: "'n::euclidean_space set" | 
| 53339 | 3827 | assumes "finite V" | 
| 53347 | 3828 | shows "aff_dim V \<le> of_nat (card V) - 1" | 
| 53339 | 3829 | proof - | 
| 53347 | 3830 | obtain B where B: "B \<subseteq> V" "of_nat (card B) = aff_dim V + 1" | 
| 53339 | 3831 | using aff_dim_inner_basis_exists[of V] by auto | 
| 3832 | then have "card B \<le> card V" | |
| 3833 | using assms card_mono by auto | |
| 53347 | 3834 | with B show ?thesis by auto | 
| 40377 | 3835 | qed | 
| 3836 | ||
| 3837 | lemma aff_dim_parallel_eq: | |
| 53339 | 3838 | fixes S T :: "'n::euclidean_space set" | 
| 3839 | assumes "affine_parallel (affine hull S) (affine hull T)" | |
| 3840 | shows "aff_dim S = aff_dim T" | |
| 3841 | proof - | |
| 3842 |   {
 | |
| 3843 |     assume "T \<noteq> {}" "S \<noteq> {}"
 | |
| 53347 | 3844 | then obtain L where L: "subspace L \<and> affine_parallel (affine hull T) L" | 
| 3845 | using affine_parallel_subspace[of "affine hull T"] | |
| 3846 | affine_affine_hull[of T] affine_hull_nonempty | |
| 53339 | 3847 | by auto | 
| 3848 | then have "aff_dim T = int (dim L)" | |
| 60420 | 3849 |       using aff_dim_parallel_subspace \<open>T \<noteq> {}\<close> by auto
 | 
| 53339 | 3850 | moreover have *: "subspace L \<and> affine_parallel (affine hull S) L" | 
| 53347 | 3851 | using L affine_parallel_assoc[of "affine hull S" "affine hull T" L] assms by auto | 
| 53339 | 3852 | moreover from * have "aff_dim S = int (dim L)" | 
| 60420 | 3853 |       using aff_dim_parallel_subspace \<open>S \<noteq> {}\<close> by auto
 | 
| 53339 | 3854 | ultimately have ?thesis by auto | 
| 3855 | } | |
| 3856 | moreover | |
| 3857 |   {
 | |
| 3858 |     assume "S = {}"
 | |
| 3859 |     then have "S = {}" and "T = {}"
 | |
| 3860 | using assms affine_hull_nonempty | |
| 3861 | unfolding affine_parallel_def | |
| 3862 | by auto | |
| 3863 | then have ?thesis using aff_dim_empty by auto | |
| 3864 | } | |
| 3865 | moreover | |
| 3866 |   {
 | |
| 3867 |     assume "T = {}"
 | |
| 3868 |     then have "S = {}" and "T = {}"
 | |
| 3869 | using assms affine_hull_nonempty | |
| 3870 | unfolding affine_parallel_def | |
| 3871 | by auto | |
| 3872 | then have ?thesis | |
| 3873 | using aff_dim_empty by auto | |
| 3874 | } | |
| 3875 | ultimately show ?thesis by blast | |
| 40377 | 3876 | qed | 
| 3877 | ||
| 3878 | lemma aff_dim_translation_eq: | |
| 53339 | 3879 | fixes a :: "'n::euclidean_space" | 
| 3880 | shows "aff_dim ((\<lambda>x. a + x) ` S) = aff_dim S" | |
| 3881 | proof - | |
| 53347 | 3882 | have "affine_parallel (affine hull S) (affine hull ((\<lambda>x. a + x) ` S))" | 
| 53339 | 3883 | unfolding affine_parallel_def | 
| 3884 | apply (rule exI[of _ "a"]) | |
| 3885 | using affine_hull_translation[of a S] | |
| 3886 | apply auto | |
| 3887 | done | |
| 3888 | then show ?thesis | |
| 3889 | using aff_dim_parallel_eq[of S "(\<lambda>x. a + x) ` S"] by auto | |
| 40377 | 3890 | qed | 
| 3891 | ||
| 3892 | lemma aff_dim_affine: | |
| 53339 | 3893 | fixes S L :: "'n::euclidean_space set" | 
| 3894 |   assumes "S \<noteq> {}"
 | |
| 3895 | and "affine S" | |
| 3896 | and "subspace L" | |
| 3897 | and "affine_parallel S L" | |
| 3898 | shows "aff_dim S = int (dim L)" | |
| 3899 | proof - | |
| 3900 | have *: "affine hull S = S" | |
| 3901 | using assms affine_hull_eq[of S] by auto | |
| 3902 | then have "affine_parallel (affine hull S) L" | |
| 3903 | using assms by (simp add: *) | |
| 3904 | then show ?thesis | |
| 3905 | using assms aff_dim_parallel_subspace[of S L] by blast | |
| 40377 | 3906 | qed | 
| 3907 | ||
| 3908 | lemma dim_affine_hull: | |
| 53339 | 3909 | fixes S :: "'n::euclidean_space set" | 
| 3910 | shows "dim (affine hull S) = dim S" | |
| 3911 | proof - | |
| 3912 | have "dim (affine hull S) \<ge> dim S" | |
| 3913 | using dim_subset by auto | |
| 3914 | moreover have "dim (span S) \<ge> dim (affine hull S)" | |
| 60303 | 3915 | using dim_subset affine_hull_subset_span by blast | 
| 53339 | 3916 | moreover have "dim (span S) = dim S" | 
| 3917 | using dim_span by auto | |
| 3918 | ultimately show ?thesis by auto | |
| 40377 | 3919 | qed | 
| 3920 | ||
| 3921 | lemma aff_dim_subspace: | |
| 53339 | 3922 | fixes S :: "'n::euclidean_space set" | 
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 3923 | assumes "subspace S" | 
| 53339 | 3924 | shows "aff_dim S = int (dim S)" | 
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 3925 | proof (cases "S={}")
 | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 3926 | case True with assms show ?thesis | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 3927 | by (simp add: subspace_affine) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 3928 | next | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 3929 | case False | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 3930 | 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: 
63092diff
changeset | 3931 | show ?thesis by auto | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 3932 | qed | 
| 40377 | 3933 | |
| 3934 | lemma aff_dim_zero: | |
| 53339 | 3935 | fixes S :: "'n::euclidean_space set" | 
| 3936 | assumes "0 \<in> affine hull S" | |
| 3937 | shows "aff_dim S = int (dim S)" | |
| 3938 | proof - | |
| 3939 | have "subspace (affine hull S)" | |
| 3940 | using subspace_affine[of "affine hull S"] affine_affine_hull assms | |
| 3941 | by auto | |
| 3942 | then have "aff_dim (affine hull S) = int (dim (affine hull S))" | |
| 3943 | using assms aff_dim_subspace[of "affine hull S"] by auto | |
| 3944 | then show ?thesis | |
| 3945 | using aff_dim_affine_hull[of S] dim_affine_hull[of S] | |
| 3946 | by auto | |
| 40377 | 3947 | qed | 
| 3948 | ||
| 63016 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3949 | lemma aff_dim_eq_dim: | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3950 | fixes S :: "'n::euclidean_space set" | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3951 | assumes "a \<in> affine hull S" | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3952 | 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: 
63007diff
changeset | 3953 | proof - | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3954 | have "0 \<in> affine hull ((\<lambda>x. -a+x) ` S)" | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3955 | unfolding Convex_Euclidean_Space.affine_hull_translation | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3956 | 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: 
63007diff
changeset | 3957 | with aff_dim_zero show ?thesis | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3958 | by (metis aff_dim_translation_eq) | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3959 | qed | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3960 | |
| 63072 | 3961 | lemma aff_dim_UNIV [simp]: "aff_dim (UNIV :: 'n::euclidean_space set) = int(DIM('n))"
 | 
| 53347 | 3962 | using aff_dim_subspace[of "(UNIV :: 'n::euclidean_space set)"] | 
| 53339 | 3963 | dim_UNIV[where 'a="'n::euclidean_space"] | 
| 3964 | by auto | |
| 40377 | 3965 | |
| 3966 | lemma aff_dim_geq: | |
| 53339 | 3967 | fixes V :: "'n::euclidean_space set" | 
| 3968 | shows "aff_dim V \<ge> -1" | |
| 3969 | proof - | |
| 53347 | 3970 | obtain B where "affine hull B = affine hull V" | 
| 3971 | and "\<not> affine_dependent B" | |
| 3972 | and "int (card B) = aff_dim V + 1" | |
| 53339 | 3973 | using aff_dim_basis_exists by auto | 
| 3974 | then show ?thesis by auto | |
| 40377 | 3975 | qed | 
| 3976 | ||
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 3977 | lemma aff_dim_negative_iff [simp]: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 3978 | fixes S :: "'n::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 3979 |   shows "aff_dim S < 0 \<longleftrightarrow>S = {}"
 | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 3980 | 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: 
63332diff
changeset | 3981 | |
| 66641 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 3982 | lemma aff_lowdim_subset_hyperplane: | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 3983 | fixes S :: "'a::euclidean_space set" | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 3984 |   assumes "aff_dim S < DIM('a)"
 | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 3985 |   obtains a b where "a \<noteq> 0" "S \<subseteq> {x. a \<bullet> x = b}"
 | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 3986 | proof (cases "S={}")
 | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 3987 | case True | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 3988 | moreover | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 3989 | have "(SOME b. b \<in> Basis) \<noteq> 0" | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 3990 | by (metis norm_some_Basis norm_zero zero_neq_one) | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 3991 | ultimately show ?thesis | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 3992 | using that by blast | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 3993 | next | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 3994 | case False | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 3995 | then obtain c S' where "c \<notin> S'" "S = insert c S'" | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 3996 | by (meson equals0I mk_disjoint_insert) | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 3997 |   have "dim (op + (-c) ` S) < DIM('a)"
 | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 3998 | by (metis \<open>S = insert c S'\<close> aff_dim_eq_dim assms hull_inc insertI1 of_nat_less_imp_less) | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 3999 |   then obtain a where "a \<noteq> 0" "span (op + (-c) ` S) \<subseteq> {x. a \<bullet> x = 0}"
 | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4000 | using lowdim_subset_hyperplane by blast | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4001 | moreover | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4002 |   have "a \<bullet> w = a \<bullet> c" if "span (op + (- c) ` S) \<subseteq> {x. a \<bullet> x = 0}" "w \<in> S" for w
 | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4003 | proof - | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4004 | have "w-c \<in> span (op + (- c) ` S)" | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4005 | by (simp add: span_superset \<open>w \<in> S\<close>) | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4006 |     with that have "w-c \<in> {x. a \<bullet> x = 0}"
 | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4007 | by blast | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4008 | then show ?thesis | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4009 | by (auto simp: algebra_simps) | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4010 | qed | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4011 |   ultimately have "S \<subseteq> {x. a \<bullet> x = a \<bullet> c}"
 | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4012 | by blast | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4013 | then show ?thesis | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4014 | by (rule that[OF \<open>a \<noteq> 0\<close>]) | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4015 | qed | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4016 | |
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4017 | lemma affine_independent_card_dim_diffs: | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4018 | fixes S :: "'a :: euclidean_space set" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4019 | assumes "~ affine_dependent S" "a \<in> S" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4020 |     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: 
63092diff
changeset | 4021 | proof - | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4022 |   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: 
63092diff
changeset | 4023 |   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: 
63092diff
changeset | 4024 | proof (cases "x = a") | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4025 | case True then show ?thesis by simp | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4026 | next | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4027 | case False then show ?thesis | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4028 | using assms by (blast intro: span_superset that) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4029 | qed | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4030 | have "\<not> affine_dependent (insert a S)" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4031 | by (simp add: assms insert_absorb) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4032 |   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: 
63092diff
changeset | 4033 | using dependent_imp_affine_dependent by fastforce | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4034 |   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: 
63092diff
changeset | 4035 | by blast | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4036 |   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: 
63092diff
changeset | 4037 | by simp | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4038 |   also have "... = card (S - {a})"
 | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4039 | 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: 
63092diff
changeset | 4040 | also have "... = card S - 1" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4041 | by (simp add: aff_independent_finite assms) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4042 |   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: 
63092diff
changeset | 4043 | have "finite S" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4044 | by (meson assms aff_independent_finite) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4045 | 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: 
63092diff
changeset | 4046 |   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: 
63092diff
changeset | 4047 | 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: 
63092diff
changeset | 4048 | ultimately show ?thesis | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4049 | by auto | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4050 | qed | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4051 | |
| 49531 | 4052 | lemma independent_card_le_aff_dim: | 
| 53347 | 4053 | fixes B :: "'n::euclidean_space set" | 
| 4054 | assumes "B \<subseteq> V" | |
| 53339 | 4055 | assumes "\<not> affine_dependent B" | 
| 4056 | shows "int (card B) \<le> aff_dim V + 1" | |
| 63016 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 4057 | proof - | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 4058 | 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: 
63007diff
changeset | 4059 | by (metis assms extend_to_affine_basis[of B V]) | 
| 53339 | 4060 | then have "of_nat (card T) = aff_dim V + 1" | 
| 4061 | using aff_dim_unique by auto | |
| 4062 | then show ?thesis | |
| 53347 | 4063 | using T card_mono[of T B] aff_independent_finite[of T] by auto | 
| 40377 | 4064 | qed | 
| 4065 | ||
| 4066 | lemma aff_dim_subset: | |
| 53347 | 4067 | fixes S T :: "'n::euclidean_space set" | 
| 4068 | assumes "S \<subseteq> T" | |
| 4069 | shows "aff_dim S \<le> aff_dim T" | |
| 53339 | 4070 | proof - | 
| 53347 | 4071 | obtain B where B: "\<not> affine_dependent B" "B \<subseteq> S" "affine hull B = affine hull S" | 
| 4072 | "of_nat (card B) = aff_dim S + 1" | |
| 53339 | 4073 | using aff_dim_inner_basis_exists[of S] by auto | 
| 4074 | then have "int (card B) \<le> aff_dim T + 1" | |
| 4075 | using assms independent_card_le_aff_dim[of B T] by auto | |
| 53347 | 4076 | with B show ?thesis by auto | 
| 40377 | 4077 | qed | 
| 4078 | ||
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4079 | lemma aff_dim_le_DIM: | 
| 53339 | 4080 | fixes S :: "'n::euclidean_space set" | 
| 4081 |   shows "aff_dim S \<le> int (DIM('n))"
 | |
| 49531 | 4082 | proof - | 
| 53339 | 4083 |   have "aff_dim (UNIV :: 'n::euclidean_space set) = int(DIM('n))"
 | 
| 63072 | 4084 | using aff_dim_UNIV by auto | 
| 53339 | 4085 |   then show "aff_dim (S:: 'n::euclidean_space set) \<le> int(DIM('n))"
 | 
| 63092 | 4086 |     using aff_dim_subset[of S "(UNIV :: ('n::euclidean_space) set)"] subset_UNIV by auto
 | 
| 40377 | 4087 | qed | 
| 4088 | ||
| 4089 | lemma affine_dim_equal: | |
| 53347 | 4090 | fixes S :: "'n::euclidean_space set" | 
| 4091 |   assumes "affine S" "affine T" "S \<noteq> {}" "S \<subseteq> T" "aff_dim S = aff_dim T"
 | |
| 4092 | shows "S = T" | |
| 4093 | proof - | |
| 4094 | obtain a where "a \<in> S" using assms by auto | |
| 4095 | then have "a \<in> T" using assms by auto | |
| 63040 | 4096 |   define LS where "LS = {y. \<exists>x \<in> S. (-a) + x = y}"
 | 
| 53347 | 4097 | then have ls: "subspace LS" "affine_parallel S LS" | 
| 60420 | 4098 | using assms parallel_subspace_explicit[of S a LS] \<open>a \<in> S\<close> by auto | 
| 53347 | 4099 | then have h1: "int(dim LS) = aff_dim S" | 
| 4100 | using assms aff_dim_affine[of S LS] by auto | |
| 4101 |   have "T \<noteq> {}" using assms by auto
 | |
| 63040 | 4102 |   define LT where "LT = {y. \<exists>x \<in> T. (-a) + x = y}"
 | 
| 53347 | 4103 | then have lt: "subspace LT \<and> affine_parallel T LT" | 
| 60420 | 4104 | using assms parallel_subspace_explicit[of T a LT] \<open>a \<in> T\<close> by auto | 
| 53347 | 4105 | then have "int(dim LT) = aff_dim T" | 
| 60420 | 4106 |     using assms aff_dim_affine[of T LT] \<open>T \<noteq> {}\<close> by auto
 | 
| 53347 | 4107 | then have "dim LS = dim LT" | 
| 4108 | using h1 assms by auto | |
| 4109 | moreover have "LS \<le> LT" | |
| 4110 | using LS_def LT_def assms by auto | |
| 4111 | ultimately have "LS = LT" | |
| 4112 | using subspace_dim_equal[of LS LT] ls lt by auto | |
| 4113 |   moreover have "S = {x. \<exists>y \<in> LS. a+y=x}"
 | |
| 4114 | using LS_def by auto | |
| 4115 |   moreover have "T = {x. \<exists>y \<in> LT. a+y=x}"
 | |
| 4116 | using LT_def by auto | |
| 4117 | ultimately show ?thesis by auto | |
| 40377 | 4118 | qed | 
| 4119 | ||
| 63881 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4120 | lemma aff_dim_eq_0: | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4121 | fixes S :: "'a::euclidean_space set" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4122 |   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: 
63627diff
changeset | 4123 | proof (cases "S = {}")
 | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4124 | case True | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4125 | then show ?thesis | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4126 | by auto | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4127 | next | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4128 | case False | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4129 | 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: 
63627diff
changeset | 4130 | show ?thesis | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4131 | proof safe | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4132 | assume 0: "aff_dim S = 0" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4133 |     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: 
63627diff
changeset | 4134 | 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: 
63627diff
changeset | 4135 |     then show "\<exists>a. S = {a}"
 | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4136 | using \<open>a \<in> S\<close> by blast | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4137 | qed auto | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4138 | qed | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4139 | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4140 | lemma affine_hull_UNIV: | 
| 53347 | 4141 | fixes S :: "'n::euclidean_space set" | 
| 4142 |   assumes "aff_dim S = int(DIM('n))"
 | |
| 4143 |   shows "affine hull S = (UNIV :: ('n::euclidean_space) set)"
 | |
| 4144 | proof - | |
| 4145 |   have "S \<noteq> {}"
 | |
| 4146 | using assms aff_dim_empty[of S] by auto | |
| 4147 | have h0: "S \<subseteq> affine hull S" | |
| 4148 | using hull_subset[of S _] by auto | |
| 4149 |   have h1: "aff_dim (UNIV :: ('n::euclidean_space) set) = aff_dim S"
 | |
| 63072 | 4150 | using aff_dim_UNIV assms by auto | 
| 53347 | 4151 |   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: 
62950diff
changeset | 4152 | using aff_dim_le_DIM[of "affine hull S"] assms h0 by auto | 
| 53347 | 4153 | have h3: "aff_dim S \<le> aff_dim (affine hull S)" | 
| 4154 | using h0 aff_dim_subset[of S "affine hull S"] assms by auto | |
| 4155 |   then have h4: "aff_dim (affine hull S) = aff_dim (UNIV :: ('n::euclidean_space) set)"
 | |
| 4156 | using h0 h1 h2 by auto | |
| 4157 | then show ?thesis | |
| 4158 |     using affine_dim_equal[of "affine hull S" "(UNIV :: ('n::euclidean_space) set)"]
 | |
| 60420 | 4159 |       affine_affine_hull[of S] affine_UNIV assms h4 h0 \<open>S \<noteq> {}\<close>
 | 
| 53347 | 4160 | by auto | 
| 40377 | 4161 | qed | 
| 4162 | ||
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4163 | lemma disjoint_affine_hull: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4164 | fixes s :: "'n::euclidean_space set" | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4165 |   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: 
62950diff
changeset | 4166 |     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: 
62950diff
changeset | 4167 | proof - | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4168 | 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: 
62950diff
changeset | 4169 | 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: 
62950diff
changeset | 4170 |   { fix y
 | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4171 | 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: 
62950diff
changeset | 4172 | then obtain a b | 
| 64267 | 4173 | where a1 [simp]: "sum a t = 1" and [simp]: "sum (\<lambda>v. a v *\<^sub>R v) t = y" | 
| 4174 | 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: 
62950diff
changeset | 4175 | by (auto simp: affine_hull_finite \<open>finite t\<close> \<open>finite u\<close>) | 
| 63040 | 4176 | 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: 
62950diff
changeset | 4177 | have [simp]: "s \<inter> t = t" "s \<inter> - t \<inter> u = u" using assms by auto | 
| 64267 | 4178 | have "sum c s = 0" | 
| 4179 | 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: 
62950diff
changeset | 4180 | moreover have "~ (\<forall>v\<in>s. c v = 0)" | 
| 64267 | 4181 | 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: 
62950diff
changeset | 4182 | moreover have "(\<Sum>v\<in>s. c v *\<^sub>R v) = 0" | 
| 64267 | 4183 | by (simp add: c_def if_smult sum_negf | 
| 4184 | 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: 
62950diff
changeset | 4185 | ultimately have False | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4186 | 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: 
62950diff
changeset | 4187 | } | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4188 | then show ?thesis by blast | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4189 | qed | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4190 | |
| 40377 | 4191 | lemma aff_dim_convex_hull: | 
| 53347 | 4192 | fixes S :: "'n::euclidean_space set" | 
| 4193 | shows "aff_dim (convex hull S) = aff_dim S" | |
| 49531 | 4194 | using aff_dim_affine_hull[of S] convex_hull_subset_affine_hull[of S] | 
| 53347 | 4195 | hull_subset[of S "convex"] aff_dim_subset[of S "convex hull S"] | 
| 4196 | aff_dim_subset[of "convex hull S" "affine hull S"] | |
| 4197 | by auto | |
| 40377 | 4198 | |
| 4199 | lemma aff_dim_cball: | |
| 53347 | 4200 | fixes a :: "'n::euclidean_space" | 
| 4201 | assumes "e > 0" | |
| 4202 |   shows "aff_dim (cball a e) = int (DIM('n))"
 | |
| 4203 | proof - | |
| 4204 | have "(\<lambda>x. a + x) ` (cball 0 e) \<subseteq> cball a e" | |
| 4205 | unfolding cball_def dist_norm by auto | |
| 4206 | then have "aff_dim (cball (0 :: 'n::euclidean_space) e) \<le> aff_dim (cball a e)" | |
| 4207 | using aff_dim_translation_eq[of a "cball 0 e"] | |
| 4208 | aff_dim_subset[of "op + a ` cball 0 e" "cball a e"] | |
| 4209 | by auto | |
| 4210 |   moreover have "aff_dim (cball (0 :: 'n::euclidean_space) e) = int (DIM('n))"
 | |
| 4211 | using hull_inc[of "(0 :: 'n::euclidean_space)" "cball 0 e"] | |
| 4212 | centre_in_cball[of "(0 :: 'n::euclidean_space)"] assms | |
| 4213 | by (simp add: dim_cball[of e] aff_dim_zero[of "cball 0 e"]) | |
| 4214 | ultimately show ?thesis | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4215 | using aff_dim_le_DIM[of "cball a e"] by auto | 
| 40377 | 4216 | qed | 
| 4217 | ||
| 4218 | lemma aff_dim_open: | |
| 53347 | 4219 | fixes S :: "'n::euclidean_space set" | 
| 4220 | assumes "open S" | |
| 4221 |     and "S \<noteq> {}"
 | |
| 4222 |   shows "aff_dim S = int (DIM('n))"
 | |
| 4223 | proof - | |
| 4224 | obtain x where "x \<in> S" | |
| 4225 | using assms by auto | |
| 4226 | then obtain e where e: "e > 0" "cball x e \<subseteq> S" | |
| 4227 | using open_contains_cball[of S] assms by auto | |
| 4228 | then have "aff_dim (cball x e) \<le> aff_dim S" | |
| 4229 | using aff_dim_subset by auto | |
| 4230 | with e show ?thesis | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4231 | using aff_dim_cball[of e x] aff_dim_le_DIM[of S] by auto | 
| 40377 | 4232 | qed | 
| 4233 | ||
| 4234 | lemma low_dim_interior: | |
| 53347 | 4235 | fixes S :: "'n::euclidean_space set" | 
| 4236 |   assumes "\<not> aff_dim S = int (DIM('n))"
 | |
| 4237 |   shows "interior S = {}"
 | |
| 4238 | proof - | |
| 4239 | have "aff_dim(interior S) \<le> aff_dim S" | |
| 4240 | using interior_subset aff_dim_subset[of "interior S" S] by auto | |
| 4241 | then show ?thesis | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4242 | using aff_dim_open[of "interior S"] aff_dim_le_DIM[of S] assms by auto | 
| 40377 | 4243 | qed | 
| 4244 | ||
| 60307 
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
 paulson <lp15@cam.ac.uk> parents: 
60303diff
changeset | 4245 | corollary empty_interior_lowdim: | 
| 
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
 paulson <lp15@cam.ac.uk> parents: 
60303diff
changeset | 4246 | fixes S :: "'n::euclidean_space set" | 
| 
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
 paulson <lp15@cam.ac.uk> parents: 
60303diff
changeset | 4247 |   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: 
62950diff
changeset | 4248 | 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: 
60303diff
changeset | 4249 | |
| 63016 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 4250 | corollary aff_dim_nonempty_interior: | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 4251 | fixes S :: "'a::euclidean_space set" | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 4252 |   shows "interior S \<noteq> {} \<Longrightarrow> aff_dim S = DIM('a)"
 | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 4253 | by (metis low_dim_interior) | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 4254 | |
| 63881 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4255 | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4256 | subsection \<open>Caratheodory's theorem.\<close> | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4257 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4258 | lemma convex_hull_caratheodory_aff_dim: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4259 |   fixes p :: "('a::euclidean_space) set"
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4260 | shows "convex hull p = | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4261 |     {y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> aff_dim p + 1 \<and>
 | 
| 64267 | 4262 | (\<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: 
61426diff
changeset | 4263 | unfolding convex_hull_explicit set_eq_iff mem_Collect_eq | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4264 | proof (intro allI iffI) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4265 | fix y | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4266 | 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 | 4267 | sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y" | 
| 4268 | 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: 
61426diff
changeset | 4269 | then obtain N where "?P N" by auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4270 | 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: 
61426diff
changeset | 4271 | apply (rule_tac ex_least_nat_le) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4272 | apply auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4273 | done | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4274 | 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: 
61426diff
changeset | 4275 | by blast | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4276 | then obtain s u where obt: "finite s" "card s = n" "s\<subseteq>p" "\<forall>x\<in>s. 0 \<le> u x" | 
| 64267 | 4277 | "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: 
61426diff
changeset | 4278 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4279 | have "card s \<le> aff_dim p + 1" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4280 | proof (rule ccontr, simp only: not_le) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4281 | assume "aff_dim p + 1 < card s" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4282 | then have "affine_dependent s" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4283 | 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: 
61426diff
changeset | 4284 | by blast | 
| 64267 | 4285 | 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: 
61426diff
changeset | 4286 | using affine_dependent_explicit_finite[OF obt(1)] by auto | 
| 63040 | 4287 |     define i where "i = (\<lambda>v. (u v) / (- w v)) ` {v\<in>s. w v < 0}"
 | 
| 4288 | define t where "t = Min i" | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4289 | have "\<exists>x\<in>s. w x < 0" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4290 | proof (rule ccontr, simp add: not_less) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4291 | assume as:"\<forall>x\<in>s. 0 \<le> w x" | 
| 64267 | 4292 |       then have "sum w (s - {v}) \<ge> 0"
 | 
| 4293 | apply (rule_tac sum_nonneg) | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4294 | apply auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4295 | done | 
| 64267 | 4296 | then have "sum w s > 0" | 
| 4297 | unfolding sum.remove[OF obt(1) \<open>v\<in>s\<close>] | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4298 | 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: 
61426diff
changeset | 4299 | then show False using wv(1) by auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4300 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4301 |     then have "i \<noteq> {}" unfolding i_def by auto
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4302 | then have "t \<ge> 0" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4303 | using Min_ge_iff[of i 0 ] and obt(1) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4304 | unfolding t_def i_def | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4305 | using obt(4)[unfolded le_less] | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4306 | by (auto simp: divide_le_0_iff) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4307 | 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: 
61426diff
changeset | 4308 | proof | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4309 | fix v | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4310 | assume "v \<in> s" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4311 | then have v: "0 \<le> u v" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4312 | using obt(4)[THEN bspec[where x=v]] by auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4313 | show "0 \<le> u v + t * w v" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4314 | proof (cases "w v < 0") | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4315 | case False | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4316 | thus ?thesis using v \<open>t\<ge>0\<close> by auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4317 | next | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4318 | case True | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4319 | then have "t \<le> u v / (- w v)" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4320 | using \<open>v\<in>s\<close> unfolding t_def i_def | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4321 | apply (rule_tac Min_le) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4322 | using obt(1) apply auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4323 | done | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4324 | then show ?thesis | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4325 | unfolding real_0_le_add_iff | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4326 | 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: 
61426diff
changeset | 4327 | by auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4328 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4329 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4330 | 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: 
61426diff
changeset | 4331 |       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: 
61426diff
changeset | 4332 | then have a: "a \<in> s" "u a + t * w a = 0" by auto | 
| 64267 | 4333 |     have *: "\<And>f. sum f (s - {a}) = sum f s - ((f a)::'b::ab_group_add)"
 | 
| 4334 | 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: 
61426diff
changeset | 4335 | have "(\<Sum>v\<in>s. u v + t * w v) = 1" | 
| 64267 | 4336 | 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: 
61426diff
changeset | 4337 | 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 | 4338 | 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: 
61426diff
changeset | 4339 | 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: 
61426diff
changeset | 4340 | ultimately have "?P (n - 1)" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4341 |       apply (rule_tac x="(s - {a})" in exI)
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4342 | 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: 
61426diff
changeset | 4343 | using obt(1-3) and t and a | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4344 | apply (auto simp add: * scaleR_left_distrib) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4345 | done | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4346 | then show False | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4347 | using smallest[THEN spec[where x="n - 1"]] by auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4348 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4349 | then show "\<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> aff_dim p + 1 \<and> | 
| 64267 | 4350 | (\<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: 
61426diff
changeset | 4351 | using obt by auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4352 | qed auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4353 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4354 | lemma caratheodory_aff_dim: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4355 |   fixes p :: "('a::euclidean_space) set"
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4356 |   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: 
61426diff
changeset | 4357 | (is "?lhs = ?rhs") | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4358 | proof | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4359 | show "?lhs \<subseteq> ?rhs" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4360 | apply (subst convex_hull_caratheodory_aff_dim) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4361 | apply clarify | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4362 | apply (rule_tac x="s" in exI) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4363 | 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: 
61426diff
changeset | 4364 | done | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4365 | next | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4366 | show "?rhs \<subseteq> ?lhs" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4367 | using hull_mono by blast | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4368 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4369 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4370 | lemma convex_hull_caratheodory: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4371 |   fixes p :: "('a::euclidean_space) set"
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4372 | shows "convex hull p = | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4373 |             {y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> DIM('a) + 1 \<and>
 | 
| 64267 | 4374 | (\<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: 
61426diff
changeset | 4375 | (is "?lhs = ?rhs") | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4376 | proof (intro set_eqI iffI) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4377 | fix x | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4378 | assume "x \<in> ?lhs" then show "x \<in> ?rhs" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4379 | apply (simp only: convex_hull_caratheodory_aff_dim Set.mem_Collect_eq) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4380 | apply (erule ex_forward)+ | 
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4381 | using aff_dim_le_DIM [of p] | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4382 | apply simp | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4383 | done | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4384 | next | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4385 | fix x | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4386 | assume "x \<in> ?rhs" then show "x \<in> ?lhs" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4387 | by (auto simp add: convex_hull_explicit) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4388 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4389 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4390 | theorem caratheodory: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4391 | "convex hull p = | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4392 |     {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: 
61426diff
changeset | 4393 |       card s \<le> DIM('a) + 1 \<and> x \<in> convex hull s}"
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4394 | proof safe | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4395 | fix x | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4396 | assume "x \<in> convex hull p" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4397 |   then obtain s u where "finite s" "s \<subseteq> p" "card s \<le> DIM('a) + 1"
 | 
| 64267 | 4398 | "\<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: 
61426diff
changeset | 4399 | unfolding convex_hull_caratheodory by auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4400 |   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: 
61426diff
changeset | 4401 | apply (rule_tac x=s in exI) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4402 | using hull_subset[of s convex] | 
| 63170 | 4403 | using convex_convex_hull[simplified convex_explicit, of s, | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4404 | THEN spec[where x=s], THEN spec[where x=u]] | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4405 | apply auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4406 | done | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4407 | next | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4408 | fix x s | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4409 |   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: 
61426diff
changeset | 4410 | then show "x \<in> convex hull p" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4411 | using hull_mono[OF \<open>s\<subseteq>p\<close>] by auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4412 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4413 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4414 | |
| 60420 | 4415 | subsection \<open>Relative interior of a set\<close> | 
| 40377 | 4416 | |
| 53347 | 4417 | definition "rel_interior S = | 
| 4418 |   {x. \<exists>T. openin (subtopology euclidean (affine hull S)) T \<and> x \<in> T \<and> T \<subseteq> S}"
 | |
| 4419 | ||
| 64287 | 4420 | lemma rel_interior_mono: | 
| 4421 | "\<lbrakk>S \<subseteq> T; affine hull S = affine hull T\<rbrakk> | |
| 4422 | \<Longrightarrow> (rel_interior S) \<subseteq> (rel_interior T)" | |
| 4423 | by (auto simp: rel_interior_def) | |
| 4424 | ||
| 4425 | lemma rel_interior_maximal: | |
| 4426 | "\<lbrakk>T \<subseteq> S; openin(subtopology euclidean (affine hull S)) T\<rbrakk> \<Longrightarrow> T \<subseteq> (rel_interior S)" | |
| 4427 | by (auto simp: rel_interior_def) | |
| 4428 | ||
| 53347 | 4429 | lemma rel_interior: | 
| 4430 |   "rel_interior S = {x \<in> S. \<exists>T. open T \<and> x \<in> T \<and> T \<inter> affine hull S \<subseteq> S}"
 | |
| 4431 | unfolding rel_interior_def[of S] openin_open[of "affine hull S"] | |
| 4432 | apply auto | |
| 4433 | proof - | |
| 4434 | fix x T | |
| 4435 | assume *: "x \<in> S" "open T" "x \<in> T" "T \<inter> affine hull S \<subseteq> S" | |
| 4436 | then have **: "x \<in> T \<inter> affine hull S" | |
| 4437 | using hull_inc by auto | |
| 54465 | 4438 | show "\<exists>Tb. (\<exists>Ta. open Ta \<and> Tb = affine hull S \<inter> Ta) \<and> x \<in> Tb \<and> Tb \<subseteq> S" | 
| 4439 | apply (rule_tac x = "T \<inter> (affine hull S)" in exI) | |
| 53347 | 4440 | using * ** | 
| 4441 | apply auto | |
| 4442 | done | |
| 4443 | qed | |
| 4444 | ||
| 4445 | 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)" | |
| 4446 | by (auto simp add: rel_interior) | |
| 4447 | ||
| 4448 | lemma mem_rel_interior_ball: | |
| 4449 | "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 | 4450 | apply (simp add: rel_interior, safe) | 
| 4451 | apply (force simp add: open_contains_ball) | |
| 53347 | 4452 | apply (rule_tac x = "ball x e" in exI) | 
| 44457 
d366fa5551ef
declare euclidean_simps [simp] at the point they are proved;
 huffman parents: 
44365diff
changeset | 4453 | apply simp | 
| 40377 | 4454 | done | 
| 4455 | ||
| 49531 | 4456 | lemma rel_interior_ball: | 
| 53347 | 4457 |   "rel_interior S = {x \<in> S. \<exists>e. e > 0 \<and> ball x e \<inter> affine hull S \<subseteq> S}"
 | 
| 4458 | using mem_rel_interior_ball [of _ S] by auto | |
| 4459 | ||
| 4460 | lemma mem_rel_interior_cball: | |
| 4461 | "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 | 4462 | apply (simp add: rel_interior, safe) | 
| 40377 | 4463 | apply (force simp add: open_contains_cball) | 
| 53347 | 4464 | apply (rule_tac x = "ball x e" in exI) | 
| 44457 
d366fa5551ef
declare euclidean_simps [simp] at the point they are proved;
 huffman parents: 
44365diff
changeset | 4465 | apply (simp add: subset_trans [OF ball_subset_cball]) | 
| 40377 | 4466 | apply auto | 
| 4467 | done | |
| 4468 | ||
| 53347 | 4469 | lemma rel_interior_cball: | 
| 4470 |   "rel_interior S = {x \<in> S. \<exists>e. e > 0 \<and> cball x e \<inter> affine hull S \<subseteq> S}"
 | |
| 4471 | using mem_rel_interior_cball [of _ S] by auto | |
| 40377 | 4472 | |
| 60303 | 4473 | lemma rel_interior_empty [simp]: "rel_interior {} = {}"
 | 
| 49531 | 4474 | by (auto simp add: rel_interior_def) | 
| 40377 | 4475 | |
| 60303 | 4476 | lemma affine_hull_sing [simp]: "affine hull {a :: 'n::euclidean_space} = {a}"
 | 
| 53347 | 4477 | by (metis affine_hull_eq affine_sing) | 
| 40377 | 4478 | |
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4479 | lemma rel_interior_sing [simp]: | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4480 |     fixes a :: "'n::euclidean_space"  shows "rel_interior {a} = {a}"
 | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4481 | apply (auto simp: rel_interior_ball) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4482 | apply (rule_tac x=1 in exI) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4483 | apply force | 
| 53347 | 4484 | done | 
| 40377 | 4485 | |
| 4486 | lemma subset_rel_interior: | |
| 53347 | 4487 | fixes S T :: "'n::euclidean_space set" | 
| 4488 | assumes "S \<subseteq> T" | |
| 4489 | and "affine hull S = affine hull T" | |
| 4490 | shows "rel_interior S \<subseteq> rel_interior T" | |
| 49531 | 4491 | using assms by (auto simp add: rel_interior_def) | 
| 4492 | ||
| 53347 | 4493 | lemma rel_interior_subset: "rel_interior S \<subseteq> S" | 
| 4494 | by (auto simp add: rel_interior_def) | |
| 4495 | ||
| 4496 | lemma rel_interior_subset_closure: "rel_interior S \<subseteq> closure S" | |
| 4497 | using rel_interior_subset by (auto simp add: closure_def) | |
| 4498 | ||
| 4499 | lemma interior_subset_rel_interior: "interior S \<subseteq> rel_interior S" | |
| 4500 | by (auto simp add: rel_interior interior_def) | |
| 40377 | 4501 | |
| 4502 | lemma interior_rel_interior: | |
| 53347 | 4503 | fixes S :: "'n::euclidean_space set" | 
| 4504 |   assumes "aff_dim S = int(DIM('n))"
 | |
| 4505 | shows "rel_interior S = interior S" | |
| 40377 | 4506 | proof - | 
| 53347 | 4507 | have "affine hull S = UNIV" | 
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4508 | using assms affine_hull_UNIV[of S] by auto | 
| 53347 | 4509 | then show ?thesis | 
| 4510 | unfolding rel_interior interior_def by auto | |
| 40377 | 4511 | qed | 
| 4512 | ||
| 60303 | 4513 | lemma rel_interior_interior: | 
| 4514 | fixes S :: "'n::euclidean_space set" | |
| 4515 | assumes "affine hull S = UNIV" | |
| 4516 | shows "rel_interior S = interior S" | |
| 4517 | using assms unfolding rel_interior interior_def by auto | |
| 4518 | ||
| 40377 | 4519 | lemma rel_interior_open: | 
| 53347 | 4520 | fixes S :: "'n::euclidean_space set" | 
| 4521 | assumes "open S" | |
| 4522 | shows "rel_interior S = S" | |
| 4523 | by (metis assms interior_eq interior_subset_rel_interior rel_interior_subset set_eq_subset) | |
| 40377 | 4524 | |
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 4525 | 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: 
60762diff
changeset | 4526 | by (simp add: interior_open) | 
| 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 4527 | |
| 40377 | 4528 | lemma interior_rel_interior_gen: | 
| 53347 | 4529 | fixes S :: "'n::euclidean_space set" | 
| 4530 |   shows "interior S = (if aff_dim S = int(DIM('n)) then rel_interior S else {})"
 | |
| 4531 | by (metis interior_rel_interior low_dim_interior) | |
| 40377 | 4532 | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4533 | lemma rel_interior_nonempty_interior: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4534 | fixes S :: "'n::euclidean_space set" | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4535 |   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: 
62950diff
changeset | 4536 | by (metis interior_rel_interior_gen) | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4537 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4538 | lemma affine_hull_nonempty_interior: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4539 | fixes S :: "'n::euclidean_space set" | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4540 |   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: 
62950diff
changeset | 4541 | 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: 
62950diff
changeset | 4542 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4543 | lemma rel_interior_affine_hull [simp]: | 
| 53347 | 4544 | fixes S :: "'n::euclidean_space set" | 
| 4545 | shows "rel_interior (affine hull S) = affine hull S" | |
| 4546 | proof - | |
| 4547 | have *: "rel_interior (affine hull S) \<subseteq> affine hull S" | |
| 4548 | using rel_interior_subset by auto | |
| 4549 |   {
 | |
| 4550 | fix x | |
| 4551 | assume x: "x \<in> affine hull S" | |
| 63040 | 4552 | define e :: real where "e = 1" | 
| 53347 | 4553 | then have "e > 0" "ball x e \<inter> affine hull (affine hull S) \<subseteq> affine hull S" | 
| 4554 | using hull_hull[of _ S] by auto | |
| 4555 | then have "x \<in> rel_interior (affine hull S)" | |
| 4556 | using x rel_interior_ball[of "affine hull S"] by auto | |
| 4557 | } | |
| 4558 | then show ?thesis using * by auto | |
| 40377 | 4559 | qed | 
| 4560 | ||
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4561 | lemma rel_interior_UNIV [simp]: "rel_interior (UNIV :: ('n::euclidean_space) set) = UNIV"
 | 
| 53347 | 4562 | by (metis open_UNIV rel_interior_open) | 
| 40377 | 4563 | |
| 4564 | lemma rel_interior_convex_shrink: | |
| 53347 | 4565 | fixes S :: "'a::euclidean_space set" | 
| 4566 | assumes "convex S" | |
| 4567 | and "c \<in> rel_interior S" | |
| 4568 | and "x \<in> S" | |
| 4569 | and "0 < e" | |
| 4570 | and "e \<le> 1" | |
| 4571 | shows "x - e *\<^sub>R (x - c) \<in> rel_interior S" | |
| 4572 | proof - | |
| 54465 | 4573 | obtain d where "d > 0" and d: "ball c d \<inter> affine hull S \<subseteq> S" | 
| 53347 | 4574 | using assms(2) unfolding mem_rel_interior_ball by auto | 
| 4575 |   {
 | |
| 4576 | fix y | |
| 4577 | assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d" "y \<in> affine hull S" | |
| 4578 | have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" | |
| 60420 | 4579 | using \<open>e > 0\<close> by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib) | 
| 53347 | 4580 | have "x \<in> affine hull S" | 
| 4581 | using assms hull_subset[of S] by auto | |
| 49531 | 4582 | moreover have "1 / e + - ((1 - e) / e) = 1" | 
| 60420 | 4583 | using \<open>e > 0\<close> left_diff_distrib[of "1" "(1-e)" "1/e"] by auto | 
| 53347 | 4584 | ultimately have **: "(1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x \<in> affine hull S" | 
| 4585 | using as affine_affine_hull[of S] mem_affine[of "affine hull S" y x "(1 / e)" "-((1 - e) / e)"] | |
| 4586 | by (simp add: algebra_simps) | |
| 61945 | 4587 | 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 | 4588 | unfolding dist_norm norm_scaleR[symmetric] | 
| 4589 | apply (rule arg_cong[where f=norm]) | |
| 60420 | 4590 | using \<open>e > 0\<close> | 
| 53347 | 4591 | apply (auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps) | 
| 4592 | done | |
| 61945 | 4593 | also have "\<dots> = \<bar>1/e\<bar> * norm (x - e *\<^sub>R (x - c) - y)" | 
| 53347 | 4594 | by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) | 
| 4595 | also have "\<dots> < d" | |
| 60420 | 4596 | using as[unfolded dist_norm] and \<open>e > 0\<close> | 
| 4597 | by (auto simp add:pos_divide_less_eq[OF \<open>e > 0\<close>] mult.commute) | |
| 53347 | 4598 | finally have "y \<in> S" | 
| 4599 | apply (subst *) | |
| 4600 | apply (rule assms(1)[unfolded convex_alt,rule_format]) | |
| 4601 | apply (rule d[unfolded subset_eq,rule_format]) | |
| 4602 | unfolding mem_ball | |
| 4603 | using assms(3-5) ** | |
| 4604 | apply auto | |
| 4605 | done | |
| 4606 | } | |
| 4607 | then have "ball (x - e *\<^sub>R (x - c)) (e*d) \<inter> affine hull S \<subseteq> S" | |
| 4608 | by auto | |
| 4609 | moreover have "e * d > 0" | |
| 60420 | 4610 | using \<open>e > 0\<close> \<open>d > 0\<close> by simp | 
| 53347 | 4611 | moreover have c: "c \<in> S" | 
| 4612 | using assms rel_interior_subset by auto | |
| 4613 | 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: 
61222diff
changeset | 4614 | using convexD_alt[of S x c e] | 
| 53347 | 4615 | apply (simp add: algebra_simps) | 
| 4616 | using assms | |
| 4617 | apply auto | |
| 4618 | done | |
| 4619 | ultimately show ?thesis | |
| 60420 | 4620 | using mem_rel_interior_ball[of "x - e *\<^sub>R (x - c)" S] \<open>e > 0\<close> by auto | 
| 40377 | 4621 | qed | 
| 4622 | ||
| 4623 | lemma interior_real_semiline: | |
| 53347 | 4624 | fixes a :: real | 
| 4625 |   shows "interior {a..} = {a<..}"
 | |
| 4626 | proof - | |
| 4627 |   {
 | |
| 4628 | fix y | |
| 4629 | assume "a < y" | |
| 4630 |     then have "y \<in> interior {a..}"
 | |
| 4631 | apply (simp add: mem_interior) | |
| 4632 | apply (rule_tac x="(y-a)" in exI) | |
| 4633 | apply (auto simp add: dist_norm) | |
| 4634 | done | |
| 4635 | } | |
| 4636 | moreover | |
| 4637 |   {
 | |
| 4638 | fix y | |
| 4639 |     assume "y \<in> interior {a..}"
 | |
| 4640 |     then obtain e where e: "e > 0" "cball y e \<subseteq> {a..}"
 | |
| 4641 |       using mem_interior_cball[of y "{a..}"] by auto
 | |
| 4642 | moreover from e have "y - e \<in> cball y e" | |
| 4643 | 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: 
60303diff
changeset | 4644 | ultimately have "a \<le> y - e" by blast | 
| 53347 | 4645 | then have "a < y" using e by auto | 
| 4646 | } | |
| 4647 | ultimately show ?thesis by auto | |
| 40377 | 4648 | qed | 
| 4649 | ||
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4650 | lemma continuous_ge_on_Ioo: | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4651 |   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: 
61848diff
changeset | 4652 | shows "g (x::real) \<ge> (a::real)" | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4653 | proof- | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4654 |   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: 
61848diff
changeset | 4655 |   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: 
61848diff
changeset | 4656 |   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: 
61848diff
changeset | 4657 |   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: 
61848diff
changeset | 4658 | by (auto simp: continuous_on_closed_vimage) | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4659 |   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: 
61952diff
changeset | 4660 |   finally show ?thesis using \<open>x \<in> {c..d}\<close> by auto
 | 
| 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61952diff
changeset | 4661 | qed | 
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4662 | |
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4663 | lemma interior_real_semiline': | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4664 | fixes a :: real | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4665 |   shows "interior {..a} = {..<a}"
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4666 | proof - | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4667 |   {
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4668 | fix y | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4669 | assume "a > y" | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4670 |     then have "y \<in> interior {..a}"
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4671 | apply (simp add: mem_interior) | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4672 | apply (rule_tac x="(a-y)" in exI) | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4673 | apply (auto simp add: dist_norm) | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4674 | done | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4675 | } | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4676 | moreover | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4677 |   {
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4678 | fix y | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4679 |     assume "y \<in> interior {..a}"
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4680 |     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: 
61848diff
changeset | 4681 |       using mem_interior_cball[of y "{..a}"] by auto
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4682 | moreover from e have "y + e \<in> cball y e" | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4683 | by (auto simp add: cball_def dist_norm) | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4684 | ultimately have "a \<ge> y + e" by auto | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4685 | then have "a > y" using e by auto | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4686 | } | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4687 | ultimately show ?thesis by auto | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4688 | qed | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4689 | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 4690 | 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: 
61848diff
changeset | 4691 | proof- | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4692 |   have "{a..b} = {a..} \<inter> {..b}" by auto
 | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61952diff
changeset | 4693 |   also have "interior ... = {a<..} \<inter> {..<b}"
 | 
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4694 | by (simp add: interior_real_semiline interior_real_semiline') | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4695 |   also have "... = {a<..<b}" by auto
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4696 | finally show ?thesis . | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4697 | qed | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4698 | |
| 66793 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 4699 | lemma interior_atLeastLessThan [simp]: | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 4700 |   fixes a::real shows "interior {a..<b} = {a<..<b}"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 4701 | by (metis atLeastLessThan_def greaterThanLessThan_def interior_atLeastAtMost_real interior_Int interior_interior interior_real_semiline) | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 4702 | |
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 4703 | lemma interior_lessThanAtMost [simp]: | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 4704 |   fixes a::real shows "interior {a<..b} = {a<..<b}"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 4705 | by (metis atLeastAtMost_def greaterThanAtMost_def interior_atLeastAtMost_real interior_Int | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 4706 | interior_interior interior_real_semiline) | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 4707 | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 4708 | lemma interior_greaterThanLessThan_real [simp]: "interior {a<..<b} = {a<..<b :: real}"
 | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 4709 | by (metis interior_atLeastAtMost_real interior_interior) | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 4710 | |
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 4711 | lemma frontier_real_Iic [simp]: | 
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4712 | fixes a :: real | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4713 |   shows "frontier {..a} = {a}"
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4714 | unfolding frontier_def by (auto simp add: interior_real_semiline') | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4715 | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 4716 | lemma rel_interior_real_box [simp]: | 
| 53347 | 4717 | fixes a b :: real | 
| 4718 | assumes "a < b" | |
| 56188 | 4719 |   shows "rel_interior {a .. b} = {a <..< b}"
 | 
| 53347 | 4720 | proof - | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54465diff
changeset | 4721 |   have "box a b \<noteq> {}"
 | 
| 53347 | 4722 | using assms | 
| 4723 | unfolding set_eq_iff | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 4724 | by (auto intro!: exI[of _ "(a + b) / 2"] simp: box_def) | 
| 40377 | 4725 | then show ?thesis | 
| 56188 | 4726 | using interior_rel_interior_gen[of "cbox a b", symmetric] | 
| 62390 | 4727 | by (simp split: if_split_asm del: box_real add: box_real[symmetric] interior_cbox) | 
| 40377 | 4728 | qed | 
| 4729 | ||
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 4730 | lemma rel_interior_real_semiline [simp]: | 
| 53347 | 4731 | fixes a :: real | 
| 4732 |   shows "rel_interior {a..} = {a<..}"
 | |
| 4733 | proof - | |
| 4734 |   have *: "{a<..} \<noteq> {}"
 | |
| 4735 | unfolding set_eq_iff by (auto intro!: exI[of _ "a + 1"]) | |
| 4736 |   then show ?thesis using interior_real_semiline interior_rel_interior_gen[of "{a..}"]
 | |
| 62390 | 4737 | by (auto split: if_split_asm) | 
| 40377 | 4738 | qed | 
| 4739 | ||
| 60420 | 4740 | subsubsection \<open>Relative open sets\<close> | 
| 40377 | 4741 | |
| 53347 | 4742 | definition "rel_open S \<longleftrightarrow> rel_interior S = S" | 
| 4743 | ||
| 4744 | lemma rel_open: "rel_open S \<longleftrightarrow> openin (subtopology euclidean (affine hull S)) S" | |
| 4745 | unfolding rel_open_def rel_interior_def | |
| 4746 | apply auto | |
| 4747 | using openin_subopen[of "subtopology euclidean (affine hull S)" S] | |
| 4748 | apply auto | |
| 4749 | done | |
| 4750 | ||
| 63072 | 4751 | lemma openin_rel_interior: "openin (subtopology euclidean (affine hull S)) (rel_interior S)" | 
| 40377 | 4752 | apply (simp add: rel_interior_def) | 
| 53347 | 4753 | apply (subst openin_subopen) | 
| 4754 | apply blast | |
| 4755 | done | |
| 40377 | 4756 | |
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4757 | lemma openin_set_rel_interior: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4758 | "openin (subtopology euclidean S) (rel_interior S)" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4759 | 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: 
63332diff
changeset | 4760 | |
| 49531 | 4761 | lemma affine_rel_open: | 
| 53347 | 4762 | fixes S :: "'n::euclidean_space set" | 
| 4763 | assumes "affine S" | |
| 4764 | shows "rel_open S" | |
| 4765 | unfolding rel_open_def | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4766 | using assms rel_interior_affine_hull[of S] affine_hull_eq[of S] | 
| 53347 | 4767 | by metis | 
| 40377 | 4768 | |
| 49531 | 4769 | lemma affine_closed: | 
| 53347 | 4770 | fixes S :: "'n::euclidean_space set" | 
| 4771 | assumes "affine S" | |
| 4772 | shows "closed S" | |
| 4773 | proof - | |
| 4774 |   {
 | |
| 4775 |     assume "S \<noteq> {}"
 | |
| 4776 | then obtain L where L: "subspace L" "affine_parallel S L" | |
| 4777 | using assms affine_parallel_subspace[of S] by auto | |
| 4778 | then obtain a where a: "S = (op + a ` L)" | |
| 4779 | using affine_parallel_def[of L S] affine_parallel_commut by auto | |
| 4780 | from L have "closed L" using closed_subspace by auto | |
| 4781 | then have "closed S" | |
| 4782 | using closed_translation a by auto | |
| 4783 | } | |
| 4784 | then show ?thesis by auto | |
| 40377 | 4785 | qed | 
| 4786 | ||
| 4787 | lemma closure_affine_hull: | |
| 53347 | 4788 | fixes S :: "'n::euclidean_space set" | 
| 4789 | shows "closure S \<subseteq> affine hull S" | |
| 44524 | 4790 | by (intro closure_minimal hull_subset affine_closed affine_affine_hull) | 
| 40377 | 4791 | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 4792 | lemma closure_same_affine_hull [simp]: | 
| 53347 | 4793 | fixes S :: "'n::euclidean_space set" | 
| 40377 | 4794 | shows "affine hull (closure S) = affine hull S" | 
| 53347 | 4795 | proof - | 
| 4796 | have "affine hull (closure S) \<subseteq> affine hull S" | |
| 4797 | using hull_mono[of "closure S" "affine hull S" "affine"] | |
| 4798 | closure_affine_hull[of S] hull_hull[of "affine" S] | |
| 4799 | by auto | |
| 4800 | moreover have "affine hull (closure S) \<supseteq> affine hull S" | |
| 4801 | using hull_mono[of "S" "closure S" "affine"] closure_subset by auto | |
| 4802 | ultimately show ?thesis by auto | |
| 49531 | 4803 | qed | 
| 4804 | ||
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4805 | lemma closure_aff_dim [simp]: | 
| 53347 | 4806 | fixes S :: "'n::euclidean_space set" | 
| 40377 | 4807 | shows "aff_dim (closure S) = aff_dim S" | 
| 53347 | 4808 | proof - | 
| 4809 | have "aff_dim S \<le> aff_dim (closure S)" | |
| 4810 | using aff_dim_subset closure_subset by auto | |
| 4811 | 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: 
63072diff
changeset | 4812 | using aff_dim_subset closure_affine_hull by blast | 
| 53347 | 4813 | moreover have "aff_dim (affine hull S) = aff_dim S" | 
| 4814 | using aff_dim_affine_hull by auto | |
| 4815 | ultimately show ?thesis by auto | |
| 40377 | 4816 | qed | 
| 4817 | ||
| 4818 | lemma rel_interior_closure_convex_shrink: | |
| 53347 | 4819 | fixes S :: "_::euclidean_space set" | 
| 4820 | assumes "convex S" | |
| 4821 | and "c \<in> rel_interior S" | |
| 4822 | and "x \<in> closure S" | |
| 4823 | and "e > 0" | |
| 4824 | and "e \<le> 1" | |
| 4825 | shows "x - e *\<^sub>R (x - c) \<in> rel_interior S" | |
| 4826 | proof - | |
| 4827 | obtain d where "d > 0" and d: "ball c d \<inter> affine hull S \<subseteq> S" | |
| 4828 | using assms(2) unfolding mem_rel_interior_ball by auto | |
| 4829 | have "\<exists>y \<in> S. norm (y - x) * (1 - e) < e * d" | |
| 4830 | proof (cases "x \<in> S") | |
| 4831 | case True | |
| 60420 | 4832 | then show ?thesis using \<open>e > 0\<close> \<open>d > 0\<close> | 
| 53347 | 4833 | apply (rule_tac bexI[where x=x]) | 
| 56544 | 4834 | apply (auto) | 
| 53347 | 4835 | done | 
| 4836 | next | |
| 4837 | case False | |
| 4838 | then have x: "x islimpt S" | |
| 4839 | using assms(3)[unfolded closure_def] by auto | |
| 4840 | show ?thesis | |
| 4841 | proof (cases "e = 1") | |
| 4842 | case True | |
| 4843 | obtain y where "y \<in> S" "y \<noteq> x" "dist y x < 1" | |
| 40377 | 4844 | using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto | 
| 53347 | 4845 | then show ?thesis | 
| 4846 | apply (rule_tac x=y in bexI) | |
| 4847 | unfolding True | |
| 60420 | 4848 | using \<open>d > 0\<close> | 
| 53347 | 4849 | apply auto | 
| 4850 | done | |
| 4851 | next | |
| 4852 | case False | |
| 4853 | then have "0 < e * d / (1 - e)" and *: "1 - e > 0" | |
| 60420 | 4854 | using \<open>e \<le> 1\<close> \<open>e > 0\<close> \<open>d > 0\<close> by (auto) | 
| 53347 | 4855 | then obtain y where "y \<in> S" "y \<noteq> x" "dist y x < e * d / (1 - e)" | 
| 40377 | 4856 | using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto | 
| 53347 | 4857 | then show ?thesis | 
| 4858 | apply (rule_tac x=y in bexI) | |
| 4859 | unfolding dist_norm | |
| 4860 | using pos_less_divide_eq[OF *] | |
| 4861 | apply auto | |
| 4862 | done | |
| 4863 | qed | |
| 4864 | qed | |
| 4865 | then obtain y where "y \<in> S" and y: "norm (y - x) * (1 - e) < e * d" | |
| 4866 | by auto | |
| 63040 | 4867 | define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)" | 
| 53347 | 4868 | have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" | 
| 60420 | 4869 | unfolding z_def using \<open>e > 0\<close> | 
| 53347 | 4870 | by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) | 
| 4871 | have zball: "z \<in> ball c d" | |
| 4872 | using mem_ball z_def dist_norm[of c] | |
| 4873 | using y and assms(4,5) | |
| 4874 | by (auto simp add:field_simps norm_minus_commute) | |
| 4875 | have "x \<in> affine hull S" | |
| 4876 | using closure_affine_hull assms by auto | |
| 4877 | moreover have "y \<in> affine hull S" | |
| 60420 | 4878 | using \<open>y \<in> S\<close> hull_subset[of S] by auto | 
| 53347 | 4879 | moreover have "c \<in> affine hull S" | 
| 4880 | using assms rel_interior_subset hull_subset[of S] by auto | |
| 4881 | ultimately have "z \<in> affine hull S" | |
| 49531 | 4882 | using z_def affine_affine_hull[of S] | 
| 53347 | 4883 | mem_affine_3_minus [of "affine hull S" c x y "(1 - e) / e"] | 
| 4884 | assms | |
| 4885 | by (auto simp add: field_simps) | |
| 4886 | then have "z \<in> S" using d zball by auto | |
| 4887 | obtain d1 where "d1 > 0" and d1: "ball z d1 \<le> ball c d" | |
| 40377 | 4888 | using zball open_ball[of c d] openE[of "ball c d" z] by auto | 
| 53347 | 4889 | then have "ball z d1 \<inter> affine hull S \<subseteq> ball c d \<inter> affine hull S" | 
| 4890 | by auto | |
| 4891 | then have "ball z d1 \<inter> affine hull S \<subseteq> S" | |
| 4892 | using d by auto | |
| 4893 | then have "z \<in> rel_interior S" | |
| 60420 | 4894 | using mem_rel_interior_ball using \<open>d1 > 0\<close> \<open>z \<in> S\<close> by auto | 
| 53347 | 4895 | then have "y - e *\<^sub>R (y - z) \<in> rel_interior S" | 
| 60420 | 4896 | using rel_interior_convex_shrink[of S z y e] assms \<open>y \<in> S\<close> by auto | 
| 53347 | 4897 | then show ?thesis using * by auto | 
| 4898 | qed | |
| 4899 | ||
| 62620 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 4900 | lemma rel_interior_eq: | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 4901 | "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: 
62618diff
changeset | 4902 | 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: 
62618diff
changeset | 4903 | |
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 4904 | lemma rel_interior_openin: | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 4905 | "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: 
62618diff
changeset | 4906 | by (simp add: rel_interior_eq) | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 4907 | |
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4908 | lemma rel_interior_affine: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4909 | fixes S :: "'n::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4910 | 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: 
63332diff
changeset | 4911 | 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: 
63332diff
changeset | 4912 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4913 | lemma rel_interior_eq_closure: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4914 | fixes S :: "'n::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4915 | 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: 
63332diff
changeset | 4916 | proof (cases "S = {}")
 | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4917 | case True | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4918 | then show ?thesis | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4919 | by auto | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4920 | next | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4921 | case False show ?thesis | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4922 | proof | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4923 | assume eq: "rel_interior S = closure S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4924 |     have "S = {} \<or> S = affine hull S"
 | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4925 | 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: 
63332diff
changeset | 4926 | 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: 
63332diff
changeset | 4927 | apply (rule conjI) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4928 | 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: 
63332diff
changeset | 4929 | 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: 
63332diff
changeset | 4930 | done | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4931 | with False have "affine hull S = S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4932 | by auto | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4933 | then show "affine S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4934 | by (metis affine_hull_eq) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4935 | next | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4936 | assume "affine S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4937 | then show "rel_interior S = closure S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4938 | 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: 
63332diff
changeset | 4939 | qed | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4940 | qed | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4941 | |
| 40377 | 4942 | |
| 60420 | 4943 | subsubsection\<open>Relative interior preserves under linear transformations\<close> | 
| 40377 | 4944 | |
| 4945 | lemma rel_interior_translation_aux: | |
| 53347 | 4946 | fixes a :: "'n::euclidean_space" | 
| 4947 | shows "((\<lambda>x. a + x) ` rel_interior S) \<subseteq> rel_interior ((\<lambda>x. a + x) ` S)" | |
| 4948 | proof - | |
| 4949 |   {
 | |
| 4950 | fix x | |
| 4951 | assume x: "x \<in> rel_interior S" | |
| 4952 | then obtain T where "open T" "x \<in> T \<inter> S" "T \<inter> affine hull S \<subseteq> S" | |
| 4953 | using mem_rel_interior[of x S] by auto | |
| 4954 | then have "open ((\<lambda>x. a + x) ` T)" | |
| 4955 | and "a + x \<in> ((\<lambda>x. a + x) ` T) \<inter> ((\<lambda>x. a + x) ` S)" | |
| 4956 | and "((\<lambda>x. a + x) ` T) \<inter> affine hull ((\<lambda>x. a + x) ` S) \<subseteq> (\<lambda>x. a + x) ` S" | |
| 4957 | using affine_hull_translation[of a S] open_translation[of T a] x by auto | |
| 4958 | then have "a + x \<in> rel_interior ((\<lambda>x. a + x) ` S)" | |
| 4959 | using mem_rel_interior[of "a+x" "((\<lambda>x. a + x) ` S)"] by auto | |
| 4960 | } | |
| 4961 | then show ?thesis by auto | |
| 60809 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60800diff
changeset | 4962 | qed | 
| 40377 | 4963 | |
| 4964 | lemma rel_interior_translation: | |
| 53347 | 4965 | fixes a :: "'n::euclidean_space" | 
| 4966 | shows "rel_interior ((\<lambda>x. a + x) ` S) = (\<lambda>x. a + x) ` rel_interior S" | |
| 4967 | proof - | |
| 4968 | have "(\<lambda>x. (-a) + x) ` rel_interior ((\<lambda>x. a + x) ` S) \<subseteq> rel_interior S" | |
| 4969 | using rel_interior_translation_aux[of "-a" "(\<lambda>x. a + x) ` S"] | |
| 4970 | translation_assoc[of "-a" "a"] | |
| 4971 | by auto | |
| 4972 | then have "((\<lambda>x. a + x) ` rel_interior S) \<supseteq> rel_interior ((\<lambda>x. a + x) ` S)" | |
| 4973 | using translation_inverse_subset[of a "rel_interior (op + a ` S)" "rel_interior S"] | |
| 4974 | by auto | |
| 4975 | then show ?thesis | |
| 4976 | using rel_interior_translation_aux[of a S] by auto | |
| 40377 | 4977 | qed | 
| 4978 | ||
| 4979 | ||
| 4980 | lemma affine_hull_linear_image: | |
| 53347 | 4981 | assumes "bounded_linear f" | 
| 4982 | shows "f ` (affine hull s) = affine hull f ` s" | |
| 4983 | apply rule | |
| 4984 | unfolding subset_eq ball_simps | |
| 4985 | apply (rule_tac[!] hull_induct, rule hull_inc) | |
| 4986 | prefer 3 | |
| 4987 | apply (erule imageE) | |
| 4988 | apply (rule_tac x=xa in image_eqI) | |
| 4989 | apply assumption | |
| 4990 | apply (rule hull_subset[unfolded subset_eq, rule_format]) | |
| 4991 | apply assumption | |
| 4992 | proof - | |
| 40377 | 4993 | interpret f: bounded_linear f by fact | 
| 53347 | 4994 |   show "affine {x. f x \<in> affine hull f ` s}"
 | 
| 4995 | unfolding affine_def | |
| 4996 | by (auto simp add: f.scaleR f.add affine_affine_hull[unfolded affine_def, rule_format]) | |
| 4997 |   show "affine {x. x \<in> f ` (affine hull s)}"
 | |
| 4998 | using affine_affine_hull[unfolded affine_def, of s] | |
| 40377 | 4999 | unfolding affine_def by (auto simp add: f.scaleR [symmetric] f.add [symmetric]) | 
| 5000 | qed auto | |
| 5001 | ||
| 5002 | ||
| 5003 | lemma rel_interior_injective_on_span_linear_image: | |
| 53347 | 5004 | fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space" | 
| 5005 | and S :: "'m::euclidean_space set" | |
| 5006 | assumes "bounded_linear f" | |
| 5007 | and "inj_on f (span S)" | |
| 5008 | shows "rel_interior (f ` S) = f ` (rel_interior S)" | |
| 5009 | proof - | |
| 5010 |   {
 | |
| 5011 | fix z | |
| 5012 | assume z: "z \<in> rel_interior (f ` S)" | |
| 5013 | then have "z \<in> f ` S" | |
| 5014 | using rel_interior_subset[of "f ` S"] by auto | |
| 5015 | then obtain x where x: "x \<in> S" "f x = z" by auto | |
| 5016 | obtain e2 where e2: "e2 > 0" "cball z e2 \<inter> affine hull (f ` S) \<subseteq> (f ` S)" | |
| 5017 | using z rel_interior_cball[of "f ` S"] by auto | |
| 5018 | obtain K where K: "K > 0" "\<And>x. norm (f x) \<le> norm x * K" | |
| 5019 | using assms Real_Vector_Spaces.bounded_linear.pos_bounded[of f] by auto | |
| 63040 | 5020 | define e1 where "e1 = 1 / K" | 
| 53347 | 5021 | then have e1: "e1 > 0" "\<And>x. e1 * norm (f x) \<le> norm x" | 
| 5022 | using K pos_le_divide_eq[of e1] by auto | |
| 63040 | 5023 | define e where "e = e1 * e2" | 
| 56544 | 5024 | then have "e > 0" using e1 e2 by auto | 
| 53347 | 5025 |     {
 | 
| 5026 | fix y | |
| 5027 | assume y: "y \<in> cball x e \<inter> affine hull S" | |
| 5028 | then have h1: "f y \<in> affine hull (f ` S)" | |
| 5029 | using affine_hull_linear_image[of f S] assms by auto | |
| 5030 | from y have "norm (x-y) \<le> e1 * e2" | |
| 5031 | using cball_def[of x e] dist_norm[of x y] e_def by auto | |
| 5032 | 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: 
63332diff
changeset | 5033 | using assms linear_diff[of f x y] linear_conv_bounded_linear[of f] by auto | 
| 53347 | 5034 | moreover have "e1 * norm (f (x-y)) \<le> norm (x - y)" | 
| 5035 | using e1 by auto | |
| 5036 | ultimately have "e1 * norm ((f x)-(f y)) \<le> e1 * e2" | |
| 5037 | by auto | |
| 5038 | then have "f y \<in> cball z e2" | |
| 5039 | using cball_def[of "f x" e2] dist_norm[of "f x" "f y"] e1 x by auto | |
| 5040 | then have "f y \<in> f ` S" | |
| 5041 | using y e2 h1 by auto | |
| 5042 | then have "y \<in> S" | |
| 5043 | 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: 
61518diff
changeset | 5044 | 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: 
61518diff
changeset | 5045 | by (metis Int_iff span_inc subsetCE) | 
| 53347 | 5046 | } | 
| 5047 | then have "z \<in> f ` (rel_interior S)" | |
| 60420 | 5048 | using mem_rel_interior_cball[of x S] \<open>e > 0\<close> x by auto | 
| 49531 | 5049 | } | 
| 53347 | 5050 | moreover | 
| 5051 |   {
 | |
| 5052 | fix x | |
| 5053 | assume x: "x \<in> rel_interior S" | |
| 54465 | 5054 | then obtain e2 where e2: "e2 > 0" "cball x e2 \<inter> affine hull S \<subseteq> S" | 
| 53347 | 5055 | using rel_interior_cball[of S] by auto | 
| 5056 | have "x \<in> S" using x rel_interior_subset by auto | |
| 5057 | then have *: "f x \<in> f ` S" by auto | |
| 5058 | have "\<forall>x\<in>span S. f x = 0 \<longrightarrow> x = 0" | |
| 5059 | using assms subspace_span linear_conv_bounded_linear[of f] | |
| 5060 | linear_injective_on_subspace_0[of f "span S"] | |
| 5061 | by auto | |
| 5062 | then obtain e1 where e1: "e1 > 0" "\<forall>x \<in> span S. e1 * norm x \<le> norm (f x)" | |
| 5063 | using assms injective_imp_isometric[of "span S" f] | |
| 5064 | subspace_span[of S] closed_subspace[of "span S"] | |
| 5065 | by auto | |
| 63040 | 5066 | define e where "e = e1 * e2" | 
| 56544 | 5067 | hence "e > 0" using e1 e2 by auto | 
| 53347 | 5068 |     {
 | 
| 5069 | fix y | |
| 5070 | assume y: "y \<in> cball (f x) e \<inter> affine hull (f ` S)" | |
| 5071 | then have "y \<in> f ` (affine hull S)" | |
| 5072 | using affine_hull_linear_image[of f S] assms by auto | |
| 5073 | then obtain xy where xy: "xy \<in> affine hull S" "f xy = y" by auto | |
| 5074 | with y have "norm (f x - f xy) \<le> e1 * e2" | |
| 5075 | using cball_def[of "f x" e] dist_norm[of "f x" y] e_def by auto | |
| 5076 | 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: 
63332diff
changeset | 5077 | using assms linear_diff[of f x xy] linear_conv_bounded_linear[of f] by auto | 
| 53347 | 5078 | moreover have *: "x - xy \<in> span S" | 
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 5079 | using subspace_diff[of "span S" x xy] subspace_span \<open>x \<in> S\<close> xy | 
| 53347 | 5080 | affine_hull_subset_span[of S] span_inc | 
| 5081 | by auto | |
| 5082 | moreover from * have "e1 * norm (x - xy) \<le> norm (f (x - xy))" | |
| 5083 | using e1 by auto | |
| 5084 | ultimately have "e1 * norm (x - xy) \<le> e1 * e2" | |
| 5085 | by auto | |
| 5086 | then have "xy \<in> cball x e2" | |
| 5087 | using cball_def[of x e2] dist_norm[of x xy] e1 by auto | |
| 5088 | then have "y \<in> f ` S" | |
| 5089 | using xy e2 by auto | |
| 5090 | } | |
| 5091 | then have "f x \<in> rel_interior (f ` S)" | |
| 60420 | 5092 | using mem_rel_interior_cball[of "(f x)" "(f ` S)"] * \<open>e > 0\<close> by auto | 
| 49531 | 5093 | } | 
| 53347 | 5094 | ultimately show ?thesis by auto | 
| 40377 | 5095 | qed | 
| 5096 | ||
| 5097 | lemma rel_interior_injective_linear_image: | |
| 53347 | 5098 | fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space" | 
| 5099 | assumes "bounded_linear f" | |
| 5100 | and "inj f" | |
| 5101 | shows "rel_interior (f ` S) = f ` (rel_interior S)" | |
| 5102 | using assms rel_interior_injective_on_span_linear_image[of f S] | |
| 5103 | subset_inj_on[of f "UNIV" "span S"] | |
| 5104 | by auto | |
| 5105 | ||
| 40377 | 5106 | |
| 60420 | 5107 | subsection\<open>Some Properties of subset of standard basis\<close> | 
| 40377 | 5108 | |
| 53347 | 5109 | lemma affine_hull_substd_basis: | 
| 5110 | assumes "d \<subseteq> Basis" | |
| 5111 |   shows "affine hull (insert 0 d) = {x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
 | |
| 5112 | (is "affine hull (insert 0 ?A) = ?B") | |
| 5113 | proof - | |
| 61076 | 5114 | have *: "\<And>A. op + (0::'a) ` A = A" "\<And>A. op + (- (0::'a)) ` A = A" | 
| 53347 | 5115 | by auto | 
| 5116 | show ?thesis | |
| 5117 | unfolding affine_hull_insert_span_gen span_substd_basis[OF assms,symmetric] * .. | |
| 40377 | 5118 | qed | 
| 5119 | ||
| 60303 | 5120 | lemma affine_hull_convex_hull [simp]: "affine hull (convex hull S) = affine hull S" | 
| 53347 | 5121 | by (metis Int_absorb1 Int_absorb2 convex_hull_subset_affine_hull hull_hull hull_mono hull_subset) | 
| 5122 | ||
| 40377 | 5123 | |
| 60420 | 5124 | subsection \<open>Openness and compactness are preserved by convex hull operation.\<close> | 
| 33175 | 5125 | |
| 34964 | 5126 | lemma open_convex_hull[intro]: | 
| 33175 | 5127 | fixes s :: "'a::real_normed_vector set" | 
| 5128 | assumes "open s" | |
| 53347 | 5129 | shows "open (convex hull s)" | 
| 5130 | unfolding open_contains_cball convex_hull_explicit | |
| 5131 | unfolding mem_Collect_eq ball_simps(8) | |
| 5132 | proof (rule, rule) | |
| 5133 | fix a | |
| 64267 | 5134 | 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" | 
| 5135 | 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 | 5136 | by auto | 
| 5137 | ||
| 5138 | from assms[unfolded open_contains_cball] obtain b | |
| 5139 | where b: "\<forall>x\<in>s. 0 < b x \<and> cball x (b x) \<subseteq> s" | |
| 5140 | using bchoice[of s "\<lambda>x e. e > 0 \<and> cball x e \<subseteq> s"] by auto | |
| 5141 |   have "b ` t \<noteq> {}"
 | |
| 56889 
48a745e1bde7
avoid the Complex constructor, use the more natural Re/Im view; moved csqrt to Complex.
 hoelzl parents: 
56571diff
changeset | 5142 | using obt by auto | 
| 63040 | 5143 | define i where "i = b ` t" | 
| 53347 | 5144 | |
| 5145 | show "\<exists>e > 0. | |
| 64267 | 5146 |     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 | 5147 | apply (rule_tac x = "Min i" in exI) | 
| 5148 | unfolding subset_eq | |
| 5149 | apply rule | |
| 5150 | defer | |
| 5151 | apply rule | |
| 5152 | unfolding mem_Collect_eq | |
| 5153 | proof - | |
| 5154 | show "0 < Min i" | |
| 60420 | 5155 |       unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] \<open>b ` t\<noteq>{}\<close>]
 | 
| 53347 | 5156 | using b | 
| 5157 | apply simp | |
| 5158 | apply rule | |
| 5159 | apply (erule_tac x=x in ballE) | |
| 60420 | 5160 | using \<open>t\<subseteq>s\<close> | 
| 53347 | 5161 | apply auto | 
| 5162 | done | |
| 5163 | next | |
| 5164 | fix y | |
| 5165 | assume "y \<in> cball a (Min i)" | |
| 5166 | then have y: "norm (a - y) \<le> Min i" | |
| 5167 | unfolding dist_norm[symmetric] by auto | |
| 5168 |     {
 | |
| 5169 | fix x | |
| 5170 | assume "x \<in> t" | |
| 5171 | then have "Min i \<le> b x" | |
| 5172 | unfolding i_def | |
| 5173 | apply (rule_tac Min_le) | |
| 5174 | using obt(1) | |
| 5175 | apply auto | |
| 5176 | done | |
| 5177 | then have "x + (y - a) \<in> cball x (b x)" | |
| 5178 | using y unfolding mem_cball dist_norm by auto | |
| 60420 | 5179 | moreover from \<open>x\<in>t\<close> have "x \<in> s" | 
| 53347 | 5180 | using obt(2) by auto | 
| 5181 | ultimately have "x + (y - a) \<in> s" | |
| 54465 | 5182 | using y and b[THEN bspec[where x=x]] unfolding subset_eq by fast | 
| 53347 | 5183 | } | 
| 33175 | 5184 | moreover | 
| 53347 | 5185 | have *: "inj_on (\<lambda>v. v + (y - a)) t" | 
| 5186 | unfolding inj_on_def by auto | |
| 33175 | 5187 | have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a))) = 1" | 
| 64267 | 5188 | unfolding sum.reindex[OF *] o_def using obt(4) by auto | 
| 33175 | 5189 | moreover have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a)) *\<^sub>R v) = y" | 
| 64267 | 5190 | unfolding sum.reindex[OF *] o_def using obt(4,5) | 
| 5191 | by (simp add: sum.distrib sum_subtractf scaleR_left.sum[symmetric] scaleR_right_distrib) | |
| 53347 | 5192 | ultimately | 
| 64267 | 5193 | 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 | 5194 | apply (rule_tac x="(\<lambda>v. v + (y - a)) ` t" in exI) | 
| 5195 | apply (rule_tac x="\<lambda>v. u (v - (y - a))" in exI) | |
| 5196 | using obt(1, 3) | |
| 5197 | apply auto | |
| 5198 | done | |
| 33175 | 5199 | qed | 
| 5200 | qed | |
| 5201 | ||
| 5202 | lemma compact_convex_combinations: | |
| 5203 | fixes s t :: "'a::real_normed_vector set" | |
| 5204 | assumes "compact s" "compact t" | |
| 5205 |   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 | 5206 | proof - | 
| 33175 | 5207 |   let ?X = "{0..1} \<times> s \<times> t"
 | 
| 5208 | let ?h = "(\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" | |
| 53347 | 5209 |   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"
 | 
| 5210 | apply (rule set_eqI) | |
| 5211 | unfolding image_iff mem_Collect_eq | |
| 5212 | apply rule | |
| 5213 | apply auto | |
| 5214 | apply (rule_tac x=u in rev_bexI) | |
| 5215 | apply simp | |
| 5216 | apply (erule rev_bexI) | |
| 5217 | apply (erule rev_bexI) | |
| 5218 | apply simp | |
| 5219 | apply auto | |
| 5220 | done | |
| 56188 | 5221 | have "continuous_on ?X (\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" | 
| 33175 | 5222 | unfolding continuous_on by (rule ballI) (intro tendsto_intros) | 
| 53347 | 5223 | then show ?thesis | 
| 5224 | unfolding * | |
| 33175 | 5225 | apply (rule compact_continuous_image) | 
| 56188 | 5226 | apply (intro compact_Times compact_Icc assms) | 
| 33175 | 5227 | done | 
| 5228 | qed | |
| 5229 | ||
| 44525 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5230 | lemma finite_imp_compact_convex_hull: | 
| 53347 | 5231 | fixes s :: "'a::real_normed_vector set" | 
| 5232 | assumes "finite s" | |
| 5233 | shows "compact (convex hull s)" | |
| 44525 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5234 | proof (cases "s = {}")
 | 
| 53347 | 5235 | case True | 
| 5236 | then show ?thesis by simp | |
| 44525 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5237 | next | 
| 53347 | 5238 | case False | 
| 5239 | with assms show ?thesis | |
| 44525 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5240 | proof (induct rule: finite_ne_induct) | 
| 53347 | 5241 | case (singleton x) | 
| 5242 | show ?case by simp | |
| 44525 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5243 | next | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5244 | case (insert x A) | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5245 | 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: 
44524diff
changeset | 5246 |     let ?T = "{0..1::real} \<times> (convex hull A)"
 | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5247 | have "continuous_on ?T ?f" | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5248 | unfolding split_def continuous_on by (intro ballI tendsto_intros) | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5249 | moreover have "compact ?T" | 
| 56188 | 5250 | by (intro compact_Times compact_Icc insert) | 
| 44525 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5251 | ultimately have "compact (?f ` ?T)" | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5252 | by (rule compact_continuous_image) | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5253 | also have "?f ` ?T = convex hull (insert x A)" | 
| 60420 | 5254 |       unfolding convex_hull_insert [OF \<open>A \<noteq> {}\<close>]
 | 
| 44525 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5255 | apply safe | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5256 | apply (rule_tac x=a in exI, simp) | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5257 | apply (rule_tac x="1 - a" in exI, simp) | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5258 | apply fast | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5259 | apply (rule_tac x="(u, b)" in image_eqI, simp_all) | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5260 | done | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5261 | finally show "compact (convex hull (insert x A))" . | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5262 | qed | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5263 | qed | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5264 | |
| 53347 | 5265 | lemma compact_convex_hull: | 
| 5266 | fixes s :: "'a::euclidean_space set" | |
| 5267 | assumes "compact s" | |
| 5268 | shows "compact (convex hull s)" | |
| 5269 | proof (cases "s = {}")
 | |
| 5270 | case True | |
| 5271 | then show ?thesis using compact_empty by simp | |
| 33175 | 5272 | next | 
| 53347 | 5273 | case False | 
| 5274 | then obtain w where "w \<in> s" by auto | |
| 5275 | show ?thesis | |
| 5276 | unfolding caratheodory[of s] | |
| 5277 |   proof (induct ("DIM('a) + 1"))
 | |
| 5278 | case 0 | |
| 5279 |     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: 
36341diff
changeset | 5280 | using compact_empty by auto | 
| 53347 | 5281 | from 0 show ?case unfolding * by simp | 
| 33175 | 5282 | next | 
| 5283 | case (Suc n) | |
| 53347 | 5284 | show ?case | 
| 5285 | proof (cases "n = 0") | |
| 5286 | case True | |
| 5287 |       have "{x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t} = s"
 | |
| 5288 | unfolding set_eq_iff and mem_Collect_eq | |
| 5289 | proof (rule, rule) | |
| 5290 | fix x | |
| 5291 | assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t" | |
| 5292 | then obtain t where t: "finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t" | |
| 5293 | by auto | |
| 5294 | show "x \<in> s" | |
| 5295 | proof (cases "card t = 0") | |
| 5296 | case True | |
| 5297 | then show ?thesis | |
| 5298 | using t(4) unfolding card_0_eq[OF t(1)] by simp | |
| 33175 | 5299 | next | 
| 53347 | 5300 | case False | 
| 60420 | 5301 | then have "card t = Suc 0" using t(3) \<open>n=0\<close> by auto | 
| 33175 | 5302 |           then obtain a where "t = {a}" unfolding card_Suc_eq by auto
 | 
| 53347 | 5303 | then show ?thesis using t(2,4) by simp | 
| 33175 | 5304 | qed | 
| 5305 | next | |
| 5306 | fix x assume "x\<in>s" | |
| 53347 | 5307 | then show "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t" | 
| 5308 |           apply (rule_tac x="{x}" in exI)
 | |
| 5309 | unfolding convex_hull_singleton | |
| 5310 | apply auto | |
| 5311 | done | |
| 5312 | qed | |
| 5313 | then show ?thesis using assms by simp | |
| 33175 | 5314 | next | 
| 53347 | 5315 | case False | 
| 5316 |       have "{x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t} =
 | |
| 5317 |         {(1 - u) *\<^sub>R x + u *\<^sub>R y | x y u.
 | |
| 5318 |           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}}"
 | |
| 5319 | unfolding set_eq_iff and mem_Collect_eq | |
| 5320 | proof (rule, rule) | |
| 5321 | fix x | |
| 5322 | assume "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and> | |
| 33175 | 5323 | 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 | 5324 | then obtain u v c t where obt: "x = (1 - c) *\<^sub>R u + c *\<^sub>R v" | 
| 5325 | "0 \<le> c \<and> c \<le> 1" "u \<in> s" "finite t" "t \<subseteq> s" "card t \<le> n" "v \<in> convex hull t" | |
| 5326 | by auto | |
| 33175 | 5327 | 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: 
61222diff
changeset | 5328 | apply (rule convexD_alt) | 
| 53347 | 5329 | using obt(2) and convex_convex_hull and hull_subset[of "insert u t" convex] | 
| 5330 | using obt(7) and hull_mono[of t "insert u t"] | |
| 5331 | apply auto | |
| 5332 | done | |
| 33175 | 5333 | ultimately show "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t" | 
| 53347 | 5334 | apply (rule_tac x="insert u t" in exI) | 
| 5335 | apply (auto simp add: card_insert_if) | |
| 5336 | done | |
| 33175 | 5337 | next | 
| 53347 | 5338 | fix x | 
| 5339 | assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t" | |
| 5340 | then obtain t where t: "finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t" | |
| 5341 | by auto | |
| 5342 | show "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and> | |
| 33175 | 5343 | 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 | 5344 | proof (cases "card t = Suc n") | 
| 5345 | case False | |
| 5346 | then have "card t \<le> n" using t(3) by auto | |
| 5347 | then show ?thesis | |
| 5348 | apply (rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI) | |
| 60420 | 5349 | using \<open>w\<in>s\<close> and t | 
| 53347 | 5350 | apply (auto intro!: exI[where x=t]) | 
| 5351 | done | |
| 33175 | 5352 | next | 
| 53347 | 5353 | case True | 
| 5354 | then obtain a u where au: "t = insert a u" "a\<notin>u" | |
| 5355 | apply (drule_tac card_eq_SucD) | |
| 5356 | apply auto | |
| 5357 | done | |
| 5358 | show ?thesis | |
| 5359 |           proof (cases "u = {}")
 | |
| 5360 | case True | |
| 5361 | then have "x = a" using t(4)[unfolded au] by auto | |
| 60420 | 5362 | show ?thesis unfolding \<open>x = a\<close> | 
| 53347 | 5363 | apply (rule_tac x=a in exI) | 
| 5364 | apply (rule_tac x=a in exI) | |
| 5365 | apply (rule_tac x=1 in exI) | |
| 60420 | 5366 | using t and \<open>n \<noteq> 0\<close> | 
| 53347 | 5367 | unfolding au | 
| 5368 |               apply (auto intro!: exI[where x="{a}"])
 | |
| 5369 | done | |
| 33175 | 5370 | next | 
| 53347 | 5371 | case False | 
| 5372 | obtain ux vx b where obt: "ux\<ge>0" "vx\<ge>0" "ux + vx = 1" | |
| 5373 | "b \<in> convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b" | |
| 5374 | using t(4)[unfolded au convex_hull_insert[OF False]] | |
| 5375 | by auto | |
| 5376 | have *: "1 - vx = ux" using obt(3) by auto | |
| 5377 | show ?thesis | |
| 5378 | apply (rule_tac x=a in exI) | |
| 5379 | apply (rule_tac x=b in exI) | |
| 5380 | apply (rule_tac x=vx in exI) | |
| 5381 | using obt and t(1-3) | |
| 5382 | unfolding au and * using card_insert_disjoint[OF _ au(2)] | |
| 5383 | apply (auto intro!: exI[where x=u]) | |
| 5384 | done | |
| 33175 | 5385 | qed | 
| 5386 | qed | |
| 5387 | qed | |
| 53347 | 5388 | then show ?thesis | 
| 5389 | using compact_convex_combinations[OF assms Suc] by simp | |
| 33175 | 5390 | qed | 
| 36362 
06475a1547cb
fix lots of looping simp calls and other warnings
 huffman parents: 
36341diff
changeset | 5391 | qed | 
| 33175 | 5392 | qed | 
| 5393 | ||
| 53347 | 5394 | |
| 60420 | 5395 | subsection \<open>Extremal points of a simplex are some vertices.\<close> | 
| 33175 | 5396 | |
| 5397 | lemma dist_increases_online: | |
| 5398 | fixes a b d :: "'a::real_inner" | |
| 5399 | assumes "d \<noteq> 0" | |
| 5400 | shows "dist a (b + d) > dist a b \<or> dist a (b - d) > dist a b" | |
| 53347 | 5401 | proof (cases "inner a d - inner b d > 0") | 
| 5402 | case True | |
| 5403 | then have "0 < inner d d + (inner a d * 2 - inner b d * 2)" | |
| 5404 | apply (rule_tac add_pos_pos) | |
| 5405 | using assms | |
| 5406 | apply auto | |
| 5407 | done | |
| 5408 | then show ?thesis | |
| 5409 | apply (rule_tac disjI2) | |
| 5410 | unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff | |
| 5411 | apply (simp add: algebra_simps inner_commute) | |
| 5412 | done | |
| 33175 | 5413 | next | 
| 53347 | 5414 | case False | 
| 5415 | then have "0 < inner d d + (inner b d * 2 - inner a d * 2)" | |
| 5416 | apply (rule_tac add_pos_nonneg) | |
| 5417 | using assms | |
| 5418 | apply auto | |
| 5419 | done | |
| 5420 | then show ?thesis | |
| 5421 | apply (rule_tac disjI1) | |
| 5422 | unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff | |
| 5423 | apply (simp add: algebra_simps inner_commute) | |
| 5424 | done | |
| 33175 | 5425 | qed | 
| 5426 | ||
| 5427 | lemma norm_increases_online: | |
| 5428 | fixes d :: "'a::real_inner" | |
| 53347 | 5429 | shows "d \<noteq> 0 \<Longrightarrow> norm (a + d) > norm a \<or> norm(a - d) > norm a" | 
| 33175 | 5430 | using dist_increases_online[of d a 0] unfolding dist_norm by auto | 
| 5431 | ||
| 5432 | lemma simplex_furthest_lt: | |
| 53347 | 5433 | fixes s :: "'a::real_inner set" | 
| 5434 | assumes "finite s" | |
| 5435 | shows "\<forall>x \<in> convex hull s. x \<notin> s \<longrightarrow> (\<exists>y \<in> convex hull s. norm (x - a) < norm(y - a))" | |
| 5436 | using assms | |
| 5437 | proof induct | |
| 5438 | fix x s | |
| 5439 | 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))" | |
| 5440 | show "\<forall>xa\<in>convex hull insert x s. xa \<notin> insert x s \<longrightarrow> | |
| 5441 | (\<exists>y\<in>convex hull insert x s. norm (xa - a) < norm (y - a))" | |
| 5442 |   proof (rule, rule, cases "s = {}")
 | |
| 5443 | case False | |
| 5444 | fix y | |
| 5445 | assume y: "y \<in> convex hull insert x s" "y \<notin> insert x s" | |
| 5446 | 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 | 5447 | using y(1)[unfolded convex_hull_insert[OF False]] by auto | 
| 5448 | show "\<exists>z\<in>convex hull insert x s. norm (y - a) < norm (z - a)" | |
| 53347 | 5449 | proof (cases "y \<in> convex hull s") | 
| 5450 | case True | |
| 5451 | then obtain z where "z \<in> convex hull s" "norm (y - a) < norm (z - a)" | |
| 33175 | 5452 | using as(3)[THEN bspec[where x=y]] and y(2) by auto | 
| 53347 | 5453 | then show ?thesis | 
| 5454 | apply (rule_tac x=z in bexI) | |
| 5455 | unfolding convex_hull_insert[OF False] | |
| 5456 | apply auto | |
| 5457 | done | |
| 33175 | 5458 | next | 
| 53347 | 5459 | case False | 
| 5460 | show ?thesis | |
| 5461 | using obt(3) | |
| 5462 | proof (cases "u = 0", case_tac[!] "v = 0") | |
| 5463 | assume "u = 0" "v \<noteq> 0" | |
| 5464 | then have "y = b" using obt by auto | |
| 5465 | then show ?thesis using False and obt(4) by auto | |
| 33175 | 5466 | next | 
| 53347 | 5467 | assume "u \<noteq> 0" "v = 0" | 
| 5468 | then have "y = x" using obt by auto | |
| 5469 | then show ?thesis using y(2) by auto | |
| 5470 | next | |
| 5471 | assume "u \<noteq> 0" "v \<noteq> 0" | |
| 5472 | then obtain w where w: "w>0" "w<u" "w<v" | |
| 5473 | using real_lbound_gt_zero[of u v] and obt(1,2) by auto | |
| 5474 | have "x \<noteq> b" | |
| 5475 | proof | |
| 5476 | assume "x = b" | |
| 5477 | then have "y = b" unfolding obt(5) | |
| 5478 | using obt(3) by (auto simp add: scaleR_left_distrib[symmetric]) | |
| 5479 | then show False using obt(4) and False by simp | |
| 5480 | qed | |
| 5481 | then have *: "w *\<^sub>R (x - b) \<noteq> 0" using w(1) by auto | |
| 5482 | show ?thesis | |
| 5483 | using dist_increases_online[OF *, of a y] | |
| 5484 | proof (elim disjE) | |
| 33175 | 5485 | assume "dist a y < dist a (y + w *\<^sub>R (x - b))" | 
| 53347 | 5486 | then have "norm (y - a) < norm ((u + w) *\<^sub>R x + (v - w) *\<^sub>R b - a)" | 
| 5487 | unfolding dist_commute[of a] | |
| 5488 | unfolding dist_norm obt(5) | |
| 5489 | by (simp add: algebra_simps) | |
| 33175 | 5490 | moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \<in> convex hull insert x s" | 
| 60420 | 5491 |             unfolding convex_hull_insert[OF \<open>s\<noteq>{}\<close>] and mem_Collect_eq
 | 
| 53347 | 5492 | apply (rule_tac x="u + w" in exI) | 
| 5493 | apply rule | |
| 5494 | defer | |
| 5495 | apply (rule_tac x="v - w" in exI) | |
| 60420 | 5496 | using \<open>u \<ge> 0\<close> and w and obt(3,4) | 
| 53347 | 5497 | apply auto | 
| 5498 | done | |
| 33175 | 5499 | ultimately show ?thesis by auto | 
| 5500 | next | |
| 5501 | assume "dist a y < dist a (y - w *\<^sub>R (x - b))" | |
| 53347 | 5502 | then have "norm (y - a) < norm ((u - w) *\<^sub>R x + (v + w) *\<^sub>R b - a)" | 
| 5503 | unfolding dist_commute[of a] | |
| 5504 | unfolding dist_norm obt(5) | |
| 5505 | by (simp add: algebra_simps) | |
| 33175 | 5506 | moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \<in> convex hull insert x s" | 
| 60420 | 5507 |             unfolding convex_hull_insert[OF \<open>s\<noteq>{}\<close>] and mem_Collect_eq
 | 
| 53347 | 5508 | apply (rule_tac x="u - w" in exI) | 
| 5509 | apply rule | |
| 5510 | defer | |
| 5511 | apply (rule_tac x="v + w" in exI) | |
| 60420 | 5512 | using \<open>u \<ge> 0\<close> and w and obt(3,4) | 
| 53347 | 5513 | apply auto | 
| 5514 | done | |
| 33175 | 5515 | ultimately show ?thesis by auto | 
| 5516 | qed | |
| 5517 | qed auto | |
| 5518 | qed | |
| 5519 | qed auto | |
| 5520 | qed (auto simp add: assms) | |
| 5521 | ||
| 5522 | lemma simplex_furthest_le: | |
| 53347 | 5523 | fixes s :: "'a::real_inner set" | 
| 5524 | assumes "finite s" | |
| 5525 |     and "s \<noteq> {}"
 | |
| 5526 | shows "\<exists>y\<in>s. \<forall>x\<in> convex hull s. norm (x - a) \<le> norm (y - a)" | |
| 5527 | proof - | |
| 5528 |   have "convex hull s \<noteq> {}"
 | |
| 5529 | using hull_subset[of s convex] and assms(2) by auto | |
| 5530 | then obtain x where x: "x \<in> convex hull s" "\<forall>y\<in>convex hull s. norm (y - a) \<le> norm (x - a)" | |
| 33175 | 5531 | using distance_attains_sup[OF finite_imp_compact_convex_hull[OF assms(1)], of a] | 
| 53347 | 5532 | unfolding dist_commute[of a] | 
| 5533 | unfolding dist_norm | |
| 5534 | by auto | |
| 5535 | show ?thesis | |
| 5536 | proof (cases "x \<in> s") | |
| 5537 | case False | |
| 5538 | then obtain y where "y \<in> convex hull s" "norm (x - a) < norm (y - a)" | |
| 5539 | using simplex_furthest_lt[OF assms(1), THEN bspec[where x=x]] and x(1) | |
| 5540 | by auto | |
| 5541 | then show ?thesis | |
| 5542 | using x(2)[THEN bspec[where x=y]] by auto | |
| 5543 | next | |
| 5544 | case True | |
| 5545 | with x show ?thesis by auto | |
| 5546 | qed | |
| 33175 | 5547 | qed | 
| 5548 | ||
| 5549 | lemma simplex_furthest_le_exists: | |
| 44525 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5550 |   fixes s :: "('a::real_inner) set"
 | 
| 53347 | 5551 | shows "finite s \<Longrightarrow> \<forall>x\<in>(convex hull s). \<exists>y\<in>s. norm (x - a) \<le> norm (y - a)" | 
| 5552 |   using simplex_furthest_le[of s] by (cases "s = {}") auto
 | |
| 33175 | 5553 | |
| 5554 | lemma simplex_extremal_le: | |
| 53347 | 5555 | fixes s :: "'a::real_inner set" | 
| 5556 | assumes "finite s" | |
| 5557 |     and "s \<noteq> {}"
 | |
| 5558 | 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)" | |
| 5559 | proof - | |
| 5560 |   have "convex hull s \<noteq> {}"
 | |
| 5561 | using hull_subset[of s convex] and assms(2) by auto | |
| 5562 | then obtain u v where obt: "u \<in> convex hull s" "v \<in> convex hull s" | |
| 33175 | 5563 | "\<forall>x\<in>convex hull s. \<forall>y\<in>convex hull s. norm (x - y) \<le> norm (u - v)" | 
| 53347 | 5564 | using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]] | 
| 5565 | by (auto simp: dist_norm) | |
| 5566 | then show ?thesis | |
| 5567 | proof (cases "u\<notin>s \<or> v\<notin>s", elim disjE) | |
| 5568 | assume "u \<notin> s" | |
| 5569 | then obtain y where "y \<in> convex hull s" "norm (u - v) < norm (y - v)" | |
| 5570 | using simplex_furthest_lt[OF assms(1), THEN bspec[where x=u]] and obt(1) | |
| 5571 | by auto | |
| 5572 | then show ?thesis | |
| 5573 | using obt(3)[THEN bspec[where x=y], THEN bspec[where x=v]] and obt(2) | |
| 5574 | by auto | |
| 33175 | 5575 | next | 
| 53347 | 5576 | assume "v \<notin> s" | 
| 5577 | then obtain y where "y \<in> convex hull s" "norm (v - u) < norm (y - u)" | |
| 5578 | using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2) | |
| 5579 | by auto | |
| 5580 | then show ?thesis | |
| 5581 | using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1) | |
| 33175 | 5582 | by (auto simp add: norm_minus_commute) | 
| 5583 | qed auto | |
| 49531 | 5584 | qed | 
| 33175 | 5585 | |
| 5586 | lemma simplex_extremal_le_exists: | |
| 53347 | 5587 | fixes s :: "'a::real_inner set" | 
| 5588 | shows "finite s \<Longrightarrow> x \<in> convex hull s \<Longrightarrow> y \<in> convex hull s \<Longrightarrow> | |
| 5589 | \<exists>u\<in>s. \<exists>v\<in>s. norm (x - y) \<le> norm (u - v)" | |
| 5590 | using convex_hull_empty simplex_extremal_le[of s] | |
| 5591 |   by(cases "s = {}") auto
 | |
| 5592 | ||
| 33175 | 5593 | |
| 60420 | 5594 | subsection \<open>Closest point of a convex set is unique, with a continuous projection.\<close> | 
| 33175 | 5595 | |
| 53347 | 5596 | definition closest_point :: "'a::{real_inner,heine_borel} set \<Rightarrow> 'a \<Rightarrow> 'a"
 | 
| 5597 | where "closest_point s a = (SOME x. x \<in> s \<and> (\<forall>y\<in>s. dist a x \<le> dist a y))" | |
| 33175 | 5598 | |
| 5599 | lemma closest_point_exists: | |
| 53347 | 5600 | assumes "closed s" | 
| 5601 |     and "s \<noteq> {}"
 | |
| 5602 | shows "closest_point s a \<in> s" | |
| 5603 | and "\<forall>y\<in>s. dist a (closest_point s a) \<le> dist a y" | |
| 5604 | unfolding closest_point_def | |
| 5605 | apply(rule_tac[!] someI2_ex) | |
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 5606 | apply (auto intro: distance_attains_inf[OF assms(1,2), of a]) | 
| 53347 | 5607 | done | 
| 5608 | ||
| 5609 | lemma closest_point_in_set: "closed s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> closest_point s a \<in> s"
 | |
| 5610 | by (meson closest_point_exists) | |
| 5611 | ||
| 5612 | lemma closest_point_le: "closed s \<Longrightarrow> x \<in> s \<Longrightarrow> dist a (closest_point s a) \<le> dist a x" | |
| 33175 | 5613 | using closest_point_exists[of s] by auto | 
| 5614 | ||
| 5615 | lemma closest_point_self: | |
| 53347 | 5616 | assumes "x \<in> s" | 
| 5617 | shows "closest_point s x = x" | |
| 5618 | unfolding closest_point_def | |
| 5619 | apply (rule some1_equality, rule ex1I[of _ x]) | |
| 5620 | using assms | |
| 5621 | apply auto | |
| 5622 | done | |
| 5623 | ||
| 5624 | lemma closest_point_refl: "closed s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> closest_point s x = x \<longleftrightarrow> x \<in> s"
 | |
| 5625 | using closest_point_in_set[of s x] closest_point_self[of x s] | |
| 5626 | by auto | |
| 33175 | 5627 | |
| 36337 | 5628 | lemma closer_points_lemma: | 
| 33175 | 5629 | assumes "inner y z > 0" | 
| 5630 | shows "\<exists>u>0. \<forall>v>0. v \<le> u \<longrightarrow> norm(v *\<^sub>R z - y) < norm y" | |
| 53347 | 5631 | proof - | 
| 5632 | have z: "inner z z > 0" | |
| 5633 | unfolding inner_gt_zero_iff using assms by auto | |
| 5634 | then show ?thesis | |
| 5635 | using assms | |
| 5636 | apply (rule_tac x = "inner y z / inner z z" in exI) | |
| 5637 | apply rule | |
| 5638 | defer | |
| 5639 | proof rule+ | |
| 5640 | fix v | |
| 5641 | assume "0 < v" and "v \<le> inner y z / inner z z" | |
| 5642 | then show "norm (v *\<^sub>R z - y) < norm y" | |
| 5643 | unfolding norm_lt using z and assms | |
| 60420 | 5644 | by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ \<open>0<v\<close>]) | 
| 56541 | 5645 | qed auto | 
| 53347 | 5646 | qed | 
| 33175 | 5647 | |
| 5648 | lemma closer_point_lemma: | |
| 5649 | assumes "inner (y - x) (z - x) > 0" | |
| 5650 | shows "\<exists>u>0. u \<le> 1 \<and> dist (x + u *\<^sub>R (z - x)) y < dist x y" | |
| 53347 | 5651 | proof - | 
| 5652 | obtain u where "u > 0" | |
| 5653 | and u: "\<forall>v>0. v \<le> u \<longrightarrow> norm (v *\<^sub>R (z - x) - (y - x)) < norm (y - x)" | |
| 33175 | 5654 | using closer_points_lemma[OF assms] by auto | 
| 53347 | 5655 | show ?thesis | 
| 5656 | apply (rule_tac x="min u 1" in exI) | |
| 60420 | 5657 | using u[THEN spec[where x="min u 1"]] and \<open>u > 0\<close> | 
| 53347 | 5658 | unfolding dist_norm by (auto simp add: norm_minus_commute field_simps) | 
| 5659 | qed | |
| 33175 | 5660 | |
| 5661 | lemma any_closest_point_dot: | |
| 5662 | assumes "convex s" "closed s" "x \<in> s" "y \<in> s" "\<forall>z\<in>s. dist a x \<le> dist a z" | |
| 5663 | shows "inner (a - x) (y - x) \<le> 0" | |
| 53347 | 5664 | proof (rule ccontr) | 
| 5665 | assume "\<not> ?thesis" | |
| 5666 | then obtain u where u: "u>0" "u\<le>1" "dist (x + u *\<^sub>R (y - x)) a < dist x a" | |
| 5667 | using closer_point_lemma[of a x y] by auto | |
| 5668 | let ?z = "(1 - u) *\<^sub>R x + u *\<^sub>R y" | |
| 5669 | 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: 
61222diff
changeset | 5670 | using convexD_alt[OF assms(1,3,4), of u] using u by auto | 
| 53347 | 5671 | then show False | 
| 5672 | using assms(5)[THEN bspec[where x="?z"]] and u(3) | |
| 5673 | by (auto simp add: dist_commute algebra_simps) | |
| 5674 | qed | |
| 33175 | 5675 | |
| 5676 | lemma any_closest_point_unique: | |
| 36337 | 5677 | fixes x :: "'a::real_inner" | 
| 33175 | 5678 | assumes "convex s" "closed s" "x \<in> s" "y \<in> s" | 
| 53347 | 5679 | "\<forall>z\<in>s. dist a x \<le> dist a z" "\<forall>z\<in>s. dist a y \<le> dist a z" | 
| 5680 | shows "x = y" | |
| 5681 | using any_closest_point_dot[OF assms(1-4,5)] and any_closest_point_dot[OF assms(1-2,4,3,6)] | |
| 33175 | 5682 | unfolding norm_pths(1) and norm_le_square | 
| 5683 | by (auto simp add: algebra_simps) | |
| 5684 | ||
| 5685 | lemma closest_point_unique: | |
| 5686 | assumes "convex s" "closed s" "x \<in> s" "\<forall>z\<in>s. dist a x \<le> dist a z" | |
| 5687 | shows "x = closest_point s a" | |
| 49531 | 5688 | using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point s a"] | 
| 33175 | 5689 | using closest_point_exists[OF assms(2)] and assms(3) by auto | 
| 5690 | ||
| 5691 | lemma closest_point_dot: | |
| 5692 | assumes "convex s" "closed s" "x \<in> s" | |
| 5693 | shows "inner (a - closest_point s a) (x - closest_point s a) \<le> 0" | |
| 53347 | 5694 | apply (rule any_closest_point_dot[OF assms(1,2) _ assms(3)]) | 
| 5695 | using closest_point_exists[OF assms(2)] and assms(3) | |
| 5696 | apply auto | |
| 5697 | done | |
| 33175 | 5698 | |
| 5699 | lemma closest_point_lt: | |
| 5700 | assumes "convex s" "closed s" "x \<in> s" "x \<noteq> closest_point s a" | |
| 5701 | shows "dist a (closest_point s a) < dist a x" | |
| 53347 | 5702 | apply (rule ccontr) | 
| 5703 | apply (rule_tac notE[OF assms(4)]) | |
| 5704 | apply (rule closest_point_unique[OF assms(1-3), of a]) | |
| 5705 | using closest_point_le[OF assms(2), of _ a] | |
| 5706 | apply fastforce | |
| 5707 | done | |
| 33175 | 5708 | |
| 5709 | lemma closest_point_lipschitz: | |
| 53347 | 5710 | assumes "convex s" | 
| 5711 |     and "closed s" "s \<noteq> {}"
 | |
| 33175 | 5712 | shows "dist (closest_point s x) (closest_point s y) \<le> dist x y" | 
| 53347 | 5713 | proof - | 
| 33175 | 5714 | have "inner (x - closest_point s x) (closest_point s y - closest_point s x) \<le> 0" | 
| 53347 | 5715 | and "inner (y - closest_point s y) (closest_point s x - closest_point s y) \<le> 0" | 
| 5716 | apply (rule_tac[!] any_closest_point_dot[OF assms(1-2)]) | |
| 5717 | using closest_point_exists[OF assms(2-3)] | |
| 5718 | apply auto | |
| 5719 | done | |
| 5720 | then show ?thesis unfolding dist_norm and norm_le | |
| 33175 | 5721 | using inner_ge_zero[of "(x - closest_point s x) - (y - closest_point s y)"] | 
| 53347 | 5722 | by (simp add: inner_add inner_diff inner_commute) | 
| 5723 | qed | |
| 33175 | 5724 | |
| 5725 | lemma continuous_at_closest_point: | |
| 53347 | 5726 | assumes "convex s" | 
| 5727 | and "closed s" | |
| 5728 |     and "s \<noteq> {}"
 | |
| 33175 | 5729 | shows "continuous (at x) (closest_point s)" | 
| 49531 | 5730 | unfolding continuous_at_eps_delta | 
| 33175 | 5731 | using le_less_trans[OF closest_point_lipschitz[OF assms]] by auto | 
| 5732 | ||
| 5733 | lemma continuous_on_closest_point: | |
| 53347 | 5734 | assumes "convex s" | 
| 5735 | and "closed s" | |
| 5736 |     and "s \<noteq> {}"
 | |
| 33175 | 5737 | shows "continuous_on t (closest_point s)" | 
| 53347 | 5738 | by (metis continuous_at_imp_continuous_on continuous_at_closest_point[OF assms]) | 
| 5739 | ||
| 63881 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5740 | proposition closest_point_in_rel_interior: | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5741 |   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: 
63627diff
changeset | 5742 | 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: 
63627diff
changeset | 5743 | proof (cases "x \<in> S") | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5744 | case True | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5745 | then show ?thesis | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5746 | by (simp add: closest_point_self) | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5747 | next | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5748 | case False | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5749 | 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: 
63627diff
changeset | 5750 | proof - | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5751 | 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: 
63627diff
changeset | 5752 | 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: 
63627diff
changeset | 5753 | using asm mem_rel_interior_cball by blast | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5754 | 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: 
63627diff
changeset | 5755 | using \<open>x \<notin> S\<close> by auto | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5756 | 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: 
63627diff
changeset | 5757 | (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: 
63627diff
changeset | 5758 | 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: 
63627diff
changeset | 5759 | by (simp add: y_def algebra_simps) | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5760 | 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: 
63627diff
changeset | 5761 | by simp | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5762 | also have "... < norm(x - closest_point S x)" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5763 | using clo_notx \<open>e > 0\<close> | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5764 | 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: 
63627diff
changeset | 5765 | 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: 
63627diff
changeset | 5766 | have "y \<in> affine hull S" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5767 | unfolding y_def | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5768 | 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: 
63627diff
changeset | 5769 | 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: 
63627diff
changeset | 5770 | 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: 
63627diff
changeset | 5771 | ultimately have "y \<in> S" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5772 | using subsetD [OF e] by simp | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5773 | 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: 
63627diff
changeset | 5774 | 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: 
63627diff
changeset | 5775 | with no_less show False | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5776 | by (simp add: dist_norm) | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5777 | qed | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5778 | moreover have "x \<notin> rel_interior S" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5779 | using rel_interior_subset False by blast | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5780 | ultimately show ?thesis by blast | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5781 | qed | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5782 | |
| 33175 | 5783 | |
| 60420 | 5784 | subsubsection \<open>Various point-to-set separating/supporting hyperplane theorems.\<close> | 
| 33175 | 5785 | |
| 5786 | lemma supporting_hyperplane_closed_point: | |
| 36337 | 5787 |   fixes z :: "'a::{real_inner,heine_borel}"
 | 
| 53347 | 5788 | assumes "convex s" | 
| 5789 | and "closed s" | |
| 5790 |     and "s \<noteq> {}"
 | |
| 5791 | and "z \<notin> s" | |
| 5792 | 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)" | |
| 5793 | proof - | |
| 5794 | 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: 
63072diff
changeset | 5795 | by (metis distance_attains_inf[OF assms(2-3)]) | 
| 53347 | 5796 | show ?thesis | 
| 5797 | apply (rule_tac x="y - z" in exI) | |
| 5798 | apply (rule_tac x="inner (y - z) y" in exI) | |
| 5799 | apply (rule_tac x=y in bexI) | |
| 5800 | apply rule | |
| 5801 | defer | |
| 5802 | apply rule | |
| 5803 | defer | |
| 5804 | apply rule | |
| 5805 | apply (rule ccontr) | |
| 60420 | 5806 | using \<open>y \<in> s\<close> | 
| 53347 | 5807 | proof - | 
| 5808 | 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: 
61738diff
changeset | 5809 | apply (subst diff_gt_0_iff_gt [symmetric]) | 
| 53347 | 5810 | unfolding inner_diff_right[symmetric] and inner_gt_zero_iff | 
| 60420 | 5811 | using \<open>y\<in>s\<close> \<open>z\<notin>s\<close> | 
| 53347 | 5812 | apply auto | 
| 5813 | done | |
| 33175 | 5814 | next | 
| 53347 | 5815 | fix x | 
| 5816 | assume "x \<in> s" | |
| 5817 | 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 | 5818 | using assms(1)[unfolded convex_alt] and y and \<open>x\<in>s\<close> and \<open>y\<in>s\<close> by auto | 
| 53347 | 5819 | assume "\<not> inner (y - z) y \<le> inner (y - z) x" | 
| 5820 | then obtain v where "v > 0" "v \<le> 1" "dist (y + v *\<^sub>R (x - y)) z < dist y z" | |
| 5821 | using closer_point_lemma[of z y x] by (auto simp add: inner_diff) | |
| 5822 | then show False | |
| 5823 | using *[THEN spec[where x=v]] by (auto simp add: dist_commute algebra_simps) | |
| 33175 | 5824 | qed auto | 
| 5825 | qed | |
| 5826 | ||
| 5827 | lemma separating_hyperplane_closed_point: | |
| 36337 | 5828 |   fixes z :: "'a::{real_inner,heine_borel}"
 | 
| 53347 | 5829 | assumes "convex s" | 
| 5830 | and "closed s" | |
| 5831 | and "z \<notin> s" | |
| 33175 | 5832 | shows "\<exists>a b. inner a z < b \<and> (\<forall>x\<in>s. inner a x > b)" | 
| 53347 | 5833 | proof (cases "s = {}")
 | 
| 5834 | case True | |
| 5835 | then show ?thesis | |
| 5836 | apply (rule_tac x="-z" in exI) | |
| 5837 | apply (rule_tac x=1 in exI) | |
| 5838 | using less_le_trans[OF _ inner_ge_zero[of z]] | |
| 5839 | apply auto | |
| 5840 | done | |
| 33175 | 5841 | next | 
| 53347 | 5842 | case False | 
| 5843 | 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: 
62131diff
changeset | 5844 | by (metis distance_attains_inf[OF assms(2) False]) | 
| 53347 | 5845 | show ?thesis | 
| 5846 | apply (rule_tac x="y - z" in exI) | |
| 5847 | apply (rule_tac x="inner (y - z) z + (norm (y - z))\<^sup>2 / 2" in exI) | |
| 5848 | apply rule | |
| 5849 | defer | |
| 5850 | apply rule | |
| 5851 | proof - | |
| 5852 | fix x | |
| 5853 | assume "x \<in> s" | |
| 5854 | have "\<not> 0 < inner (z - y) (x - y)" | |
| 5855 | apply (rule notI) | |
| 5856 | apply (drule closer_point_lemma) | |
| 5857 | proof - | |
| 33175 | 5858 | assume "\<exists>u>0. u \<le> 1 \<and> dist (y + u *\<^sub>R (x - y)) z < dist y z" | 
| 53347 | 5859 | then obtain u where "u > 0" "u \<le> 1" "dist (y + u *\<^sub>R (x - y)) z < dist y z" | 
| 5860 | by auto | |
| 5861 | then show False using y[THEN bspec[where x="y + u *\<^sub>R (x - y)"]] | |
| 33175 | 5862 | using assms(1)[unfolded convex_alt, THEN bspec[where x=y]] | 
| 60420 | 5863 | using \<open>x\<in>s\<close> \<open>y\<in>s\<close> by (auto simp add: dist_commute algebra_simps) | 
| 53347 | 5864 | qed | 
| 5865 | moreover have "0 < (norm (y - z))\<^sup>2" | |
| 60420 | 5866 | using \<open>y\<in>s\<close> \<open>z\<notin>s\<close> by auto | 
| 53347 | 5867 | then have "0 < inner (y - z) (y - z)" | 
| 5868 | 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: 
51524diff
changeset | 5869 | ultimately show "inner (y - z) z + (norm (y - z))\<^sup>2 / 2 < inner (y - z) x" | 
| 53347 | 5870 | unfolding power2_norm_eq_inner and not_less | 
| 5871 | by (auto simp add: field_simps inner_commute inner_diff) | |
| 60420 | 5872 | qed (insert \<open>y\<in>s\<close> \<open>z\<notin>s\<close>, auto) | 
| 33175 | 5873 | qed | 
| 5874 | ||
| 5875 | lemma separating_hyperplane_closed_0: | |
| 53347 | 5876 |   assumes "convex (s::('a::euclidean_space) set)"
 | 
| 5877 | and "closed s" | |
| 5878 | and "0 \<notin> s" | |
| 33175 | 5879 | shows "\<exists>a b. a \<noteq> 0 \<and> 0 < b \<and> (\<forall>x\<in>s. inner a x > b)" | 
| 53347 | 5880 | 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: 
50104diff
changeset | 5881 | case True | 
| 53347 | 5882 | have "norm ((SOME i. i\<in>Basis)::'a) = 1" "(SOME i. i\<in>Basis) \<noteq> (0::'a)" | 
| 5883 | defer | |
| 5884 | apply (subst norm_le_zero_iff[symmetric]) | |
| 5885 | apply (auto simp: SOME_Basis) | |
| 5886 | done | |
| 5887 | then show ?thesis | |
| 5888 | apply (rule_tac x="SOME i. i\<in>Basis" in exI) | |
| 5889 | apply (rule_tac x=1 in exI) | |
| 5890 | using True using DIM_positive[where 'a='a] | |
| 5891 | apply auto | |
| 5892 | done | |
| 5893 | next | |
| 5894 | case False | |
| 5895 | then show ?thesis | |
| 5896 | using False using separating_hyperplane_closed_point[OF assms] | |
| 5897 | apply (elim exE) | |
| 5898 | unfolding inner_zero_right | |
| 5899 | apply (rule_tac x=a in exI) | |
| 5900 | apply (rule_tac x=b in exI) | |
| 5901 | apply auto | |
| 5902 | done | |
| 5903 | qed | |
| 5904 | ||
| 33175 | 5905 | |
| 60420 | 5906 | subsubsection \<open>Now set-to-set for closed/compact sets\<close> | 
| 33175 | 5907 | |
| 5908 | lemma separating_hyperplane_closed_compact: | |
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5909 | fixes S :: "'a::euclidean_space set" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5910 | assumes "convex S" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5911 | and "closed S" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5912 | and "convex T" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5913 | and "compact T" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5914 |     and "T \<noteq> {}"
 | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5915 |     and "S \<inter> T = {}"
 | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5916 | shows "\<exists>a b. (\<forall>x\<in>S. inner a x < b) \<and> (\<forall>x\<in>T. inner a x > b)" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5917 | proof (cases "S = {}")
 | 
| 33175 | 5918 | case True | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5919 | obtain b where b: "b > 0" "\<forall>x\<in>T. norm x \<le> b" | 
| 53347 | 5920 | using compact_imp_bounded[OF assms(4)] unfolding bounded_pos by auto | 
| 5921 | obtain z :: 'a where z: "norm z = b + 1" | |
| 5922 | using vector_choose_size[of "b + 1"] and b(1) by auto | |
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5923 | then have "z \<notin> T" using b(2)[THEN bspec[where x=z]] by auto | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5924 | then obtain a b where ab: "inner a z < b" "\<forall>x\<in>T. b < inner a x" | 
| 53347 | 5925 | using separating_hyperplane_closed_point[OF assms(3) compact_imp_closed[OF assms(4)], of z] | 
| 5926 | by auto | |
| 5927 | then show ?thesis | |
| 5928 | using True by auto | |
| 33175 | 5929 | next | 
| 53347 | 5930 | case False | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5931 | then obtain y where "y \<in> S" by auto | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5932 |   obtain a b where "0 < b" "\<forall>x \<in> (\<Union>x\<in> S. \<Union>y \<in> T. {x - y}). b < inner a x"
 | 
| 33175 | 5933 | using separating_hyperplane_closed_point[OF convex_differences[OF assms(1,3)], of 0] | 
| 53347 | 5934 | using closed_compact_differences[OF assms(2,4)] | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5935 | using assms(6) by auto | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5936 | then have ab: "\<forall>x\<in>S. \<forall>y\<in>T. b + inner a y < inner a x" | 
| 53347 | 5937 | apply - | 
| 5938 | apply rule | |
| 5939 | apply rule | |
| 5940 | apply (erule_tac x="x - y" in ballE) | |
| 5941 | apply (auto simp add: inner_diff) | |
| 5942 | done | |
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5943 | define k where "k = (SUP x:T. a \<bullet> x)" | 
| 53347 | 5944 | show ?thesis | 
| 5945 | apply (rule_tac x="-a" in exI) | |
| 5946 | 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: 
54258diff
changeset | 5947 | apply (intro conjI ballI) | 
| 53347 | 5948 | unfolding inner_minus_left and neg_less_iff_less | 
| 5949 | proof - | |
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5950 | fix x assume "x \<in> T" | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54258diff
changeset | 5951 | then have "inner a x - b / 2 < k" | 
| 53347 | 5952 | unfolding k_def | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54258diff
changeset | 5953 | proof (subst less_cSUP_iff) | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5954 |       show "T \<noteq> {}" by fact
 | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5955 | show "bdd_above (op \<bullet> a ` T)" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5956 | 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: 
54258diff
changeset | 5957 | by (intro bdd_aboveI2[where M="inner a y - b"]) (auto simp: field_simps intro: less_imp_le) | 
| 60420 | 5958 | 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: 
54258diff
changeset | 5959 | then show "inner a x < k + b / 2" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54258diff
changeset | 5960 | by auto | 
| 33175 | 5961 | next | 
| 53347 | 5962 | fix x | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5963 | assume "x \<in> S" | 
| 53347 | 5964 | then have "k \<le> inner a x - b" | 
| 5965 | unfolding k_def | |
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54258diff
changeset | 5966 | apply (rule_tac cSUP_least) | 
| 53347 | 5967 | using assms(5) | 
| 5968 | using ab[THEN bspec[where x=x]] | |
| 5969 | apply auto | |
| 5970 | done | |
| 5971 | then show "k + b / 2 < inner a x" | |
| 60420 | 5972 | using \<open>0 < b\<close> by auto | 
| 33175 | 5973 | qed | 
| 5974 | qed | |
| 5975 | ||
| 5976 | lemma separating_hyperplane_compact_closed: | |
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5977 | fixes S :: "'a::euclidean_space set" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5978 | assumes "convex S" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5979 | and "compact S" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5980 |     and "S \<noteq> {}"
 | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5981 | and "convex T" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5982 | and "closed T" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5983 |     and "S \<inter> T = {}"
 | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5984 | shows "\<exists>a b. (\<forall>x\<in>S. inner a x < b) \<and> (\<forall>x\<in>T. inner a x > b)" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5985 | proof - | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 5986 | obtain a b where "(\<forall>x\<in>T. inner a x < b) \<and> (\<forall>x\<in>S. b < inner a x)" | 
| 53347 | 5987 | using separating_hyperplane_closed_compact[OF assms(4-5,1-2,3)] and assms(6) | 
| 5988 | by auto | |
| 5989 | then show ?thesis | |
| 5990 | apply (rule_tac x="-a" in exI) | |
| 5991 | apply (rule_tac x="-b" in exI) | |
| 5992 | apply auto | |
| 5993 | done | |
| 5994 | qed | |
| 5995 | ||
| 33175 | 5996 | |
| 60420 | 5997 | subsubsection \<open>General case without assuming closure and getting non-strict separation\<close> | 
| 33175 | 5998 | |
| 5999 | 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: 
36844diff
changeset | 6000 | assumes "convex s" "(0::'a::euclidean_space) \<notin> s" | 
| 33175 | 6001 | shows "\<exists>a. a \<noteq> 0 \<and> (\<forall>x\<in>s. 0 \<le> inner a x)" | 
| 53347 | 6002 | proof - | 
| 6003 |   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: 
62131diff
changeset | 6004 |   have *: "frontier (cball 0 1) \<inter> \<Inter>f \<noteq> {}" if as: "f \<subseteq> ?k ` s" "finite f" for f
 | 
| 53347 | 6005 | proof - | 
| 6006 | obtain c where c: "f = ?k ` c" "c \<subseteq> s" "finite c" | |
| 6007 | using finite_subset_image[OF as(2,1)] by auto | |
| 6008 | then obtain a b where ab: "a \<noteq> 0" "0 < b" "\<forall>x\<in>convex hull c. b < inner a x" | |
| 33175 | 6009 | using separating_hyperplane_closed_0[OF convex_convex_hull, of c] | 
| 6010 | using finite_imp_compact_convex_hull[OF c(3), THEN compact_imp_closed] and assms(2) | |
| 53347 | 6011 | 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: 
61531diff
changeset | 6012 | by force | 
| 53347 | 6013 | then have "\<exists>x. norm x = 1 \<and> (\<forall>y\<in>c. 0 \<le> inner y x)" | 
| 6014 | apply (rule_tac x = "inverse(norm a) *\<^sub>R a" in exI) | |
| 6015 | using hull_subset[of c convex] | |
| 6016 | unfolding subset_eq and inner_scaleR | |
| 56536 | 6017 | by (auto simp add: inner_commute del: ballE elim!: ballE) | 
| 53347 | 6018 |     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: 
62131diff
changeset | 6019 | 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: 
62131diff
changeset | 6020 | qed | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 6021 |   have "frontier (cball 0 1) \<inter> (\<Inter>(?k ` s)) \<noteq> {}"
 | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 6022 | apply (rule compact_imp_fip) | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 6023 | apply (rule compact_frontier[OF compact_cball]) | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 6024 | using * closed_halfspace_ge | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 6025 | by auto | 
| 53347 | 6026 | 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: 
62131diff
changeset | 6027 | unfolding frontier_cball dist_norm sphere_def by auto | 
| 53347 | 6028 | then show ?thesis | 
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 6029 | by (metis inner_commute mem_Collect_eq norm_eq_zero zero_neq_one) | 
| 53347 | 6030 | qed | 
| 33175 | 6031 | |
| 6032 | lemma separating_hyperplane_sets: | |
| 53347 | 6033 | fixes s t :: "'a::euclidean_space set" | 
| 6034 | assumes "convex s" | |
| 6035 | and "convex t" | |
| 6036 |     and "s \<noteq> {}"
 | |
| 6037 |     and "t \<noteq> {}"
 | |
| 6038 |     and "s \<inter> t = {}"
 | |
| 33175 | 6039 | 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 | 6040 | proof - | 
| 6041 | from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]] | |
| 6042 |   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: 
61531diff
changeset | 6043 | using assms(3-5) by fastforce | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54258diff
changeset | 6044 | then have *: "\<And>x y. x \<in> t \<Longrightarrow> y \<in> s \<Longrightarrow> inner a y \<le> inner a x" | 
| 33270 | 6045 | by (force simp add: inner_diff) | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54258diff
changeset | 6046 | then have bdd: "bdd_above ((op \<bullet> a)`s)" | 
| 60420 | 6047 |     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: 
54258diff
changeset | 6048 | show ?thesis | 
| 60420 | 6049 | using \<open>a\<noteq>0\<close> | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54258diff
changeset | 6050 | by (intro exI[of _ a] exI[of _ "SUP x:s. a \<bullet> x"]) | 
| 60420 | 6051 |        (auto intro!: cSUP_upper bdd cSUP_least \<open>a \<noteq> 0\<close> \<open>s \<noteq> {}\<close> *)
 | 
| 6052 | qed | |
| 6053 | ||
| 6054 | ||
| 6055 | subsection \<open>More convexity generalities\<close> | |
| 33175 | 6056 | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 6057 | lemma convex_closure [intro,simp]: | 
| 33175 | 6058 | fixes s :: "'a::real_normed_vector set" | 
| 53347 | 6059 | assumes "convex s" | 
| 6060 | shows "convex (closure s)" | |
| 53676 | 6061 | apply (rule convexI) | 
| 6062 | apply (unfold closure_sequential, elim exE) | |
| 6063 | apply (rule_tac x="\<lambda>n. u *\<^sub>R xa n + v *\<^sub>R xb n" in exI) | |
| 53347 | 6064 | apply (rule,rule) | 
| 53676 | 6065 | apply (rule convexD [OF assms]) | 
| 53347 | 6066 | apply (auto del: tendsto_const intro!: tendsto_intros) | 
| 6067 | done | |
| 33175 | 6068 | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 6069 | lemma convex_interior [intro,simp]: | 
| 33175 | 6070 | fixes s :: "'a::real_normed_vector set" | 
| 53347 | 6071 | assumes "convex s" | 
| 6072 | shows "convex (interior s)" | |
| 6073 | unfolding convex_alt Ball_def mem_interior | |
| 6074 | apply (rule,rule,rule,rule,rule,rule) | |
| 6075 | apply (elim exE conjE) | |
| 6076 | proof - | |
| 6077 | fix x y u | |
| 6078 | assume u: "0 \<le> u" "u \<le> (1::real)" | |
| 6079 | fix e d | |
| 6080 | assume ed: "ball x e \<subseteq> s" "ball y d \<subseteq> s" "0<d" "0<e" | |
| 6081 | show "\<exists>e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \<subseteq> s" | |
| 6082 | apply (rule_tac x="min d e" in exI) | |
| 6083 | apply rule | |
| 6084 | unfolding subset_eq | |
| 6085 | defer | |
| 6086 | apply rule | |
| 6087 | proof - | |
| 6088 | fix z | |
| 6089 | assume "z \<in> ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) (min d e)" | |
| 6090 | then have "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \<in> s" | |
| 6091 | apply (rule_tac assms[unfolded convex_alt, rule_format]) | |
| 6092 | using ed(1,2) and u | |
| 6093 | unfolding subset_eq mem_ball Ball_def dist_norm | |
| 6094 | apply (auto simp add: algebra_simps) | |
| 6095 | done | |
| 6096 | then show "z \<in> s" | |
| 6097 | using u by (auto simp add: algebra_simps) | |
| 6098 | qed(insert u ed(3-4), auto) | |
| 6099 | qed | |
| 33175 | 6100 | |
| 34964 | 6101 | lemma convex_hull_eq_empty[simp]: "convex hull s = {} \<longleftrightarrow> s = {}"
 | 
| 33175 | 6102 | using hull_subset[of s convex] convex_hull_empty by auto | 
| 6103 | ||
| 53347 | 6104 | |
| 60420 | 6105 | subsection \<open>Moving and scaling convex hulls.\<close> | 
| 33175 | 6106 | |
| 53676 | 6107 | lemma convex_hull_set_plus: | 
| 6108 | "convex hull (s + t) = convex hull s + convex hull t" | |
| 6109 | unfolding set_plus_image | |
| 6110 | apply (subst convex_hull_linear_image [symmetric]) | |
| 6111 | apply (simp add: linear_iff scaleR_right_distrib) | |
| 6112 | apply (simp add: convex_hull_Times) | |
| 6113 | done | |
| 6114 | ||
| 6115 | lemma translation_eq_singleton_plus: "(\<lambda>x. a + x) ` t = {a} + t"
 | |
| 6116 | unfolding set_plus_def by auto | |
| 33175 | 6117 | |
| 6118 | lemma convex_hull_translation: | |
| 6119 | "convex hull ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (convex hull s)" | |
| 53676 | 6120 | unfolding translation_eq_singleton_plus | 
| 6121 | by (simp only: convex_hull_set_plus convex_hull_singleton) | |
| 33175 | 6122 | |
| 6123 | lemma convex_hull_scaling: | |
| 6124 | "convex hull ((\<lambda>x. c *\<^sub>R x) ` s) = (\<lambda>x. c *\<^sub>R x) ` (convex hull s)" | |
| 53676 | 6125 | using linear_scaleR by (rule convex_hull_linear_image [symmetric]) | 
| 33175 | 6126 | |
| 6127 | lemma convex_hull_affinity: | |
| 6128 | "convex hull ((\<lambda>x. a + c *\<^sub>R x) ` s) = (\<lambda>x. a + c *\<^sub>R x) ` (convex hull s)" | |
| 53347 | 6129 | by(simp only: image_image[symmetric] convex_hull_scaling convex_hull_translation) | 
| 6130 | ||
| 33175 | 6131 | |
| 60420 | 6132 | subsection \<open>Convexity of cone hulls\<close> | 
| 40377 | 6133 | |
| 6134 | lemma convex_cone_hull: | |
| 53347 | 6135 | assumes "convex S" | 
| 6136 | shows "convex (cone hull S)" | |
| 53676 | 6137 | proof (rule convexI) | 
| 6138 | fix x y | |
| 6139 | assume xy: "x \<in> cone hull S" "y \<in> cone hull S" | |
| 6140 |   then have "S \<noteq> {}"
 | |
| 6141 | using cone_hull_empty_iff[of S] by auto | |
| 6142 | fix u v :: real | |
| 6143 | assume uv: "u \<ge> 0" "v \<ge> 0" "u + v = 1" | |
| 6144 | then have *: "u *\<^sub>R x \<in> cone hull S" "v *\<^sub>R y \<in> cone hull S" | |
| 6145 | using cone_cone_hull[of S] xy cone_def[of "cone hull S"] by auto | |
| 6146 | from * obtain cx :: real and xx where x: "u *\<^sub>R x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S" | |
| 6147 | using cone_hull_expl[of S] by auto | |
| 6148 | from * obtain cy :: real and yy where y: "v *\<^sub>R y = cy *\<^sub>R yy" "cy \<ge> 0" "yy \<in> S" | |
| 6149 | using cone_hull_expl[of S] by auto | |
| 53347 | 6150 |   {
 | 
| 53676 | 6151 | assume "cx + cy \<le> 0" | 
| 6152 | then have "u *\<^sub>R x = 0" and "v *\<^sub>R y = 0" | |
| 6153 | using x y by auto | |
| 6154 | then have "u *\<^sub>R x + v *\<^sub>R y = 0" | |
| 6155 | by auto | |
| 6156 | then have "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S" | |
| 60420 | 6157 |       using cone_hull_contains_0[of S] \<open>S \<noteq> {}\<close> by auto
 | 
| 40377 | 6158 | } | 
| 53676 | 6159 | moreover | 
| 6160 |   {
 | |
| 6161 | assume "cx + cy > 0" | |
| 6162 | then have "(cx / (cx + cy)) *\<^sub>R xx + (cy / (cx + cy)) *\<^sub>R yy \<in> S" | |
| 6163 | using assms mem_convex_alt[of S xx yy cx cy] x y by auto | |
| 6164 | then have "cx *\<^sub>R xx + cy *\<^sub>R yy \<in> cone hull S" | |
| 60420 | 6165 | 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 | 6166 | by (auto simp add: scaleR_right_distrib) | 
| 6167 | then have "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S" | |
| 6168 | using x y by auto | |
| 6169 | } | |
| 6170 | moreover have "cx + cy \<le> 0 \<or> cx + cy > 0" by auto | |
| 6171 | ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S" by blast | |
| 40377 | 6172 | qed | 
| 6173 | ||
| 6174 | lemma cone_convex_hull: | |
| 53347 | 6175 | assumes "cone S" | 
| 6176 | shows "cone (convex hull S)" | |
| 6177 | proof (cases "S = {}")
 | |
| 6178 | case True | |
| 6179 | then show ?thesis by auto | |
| 6180 | next | |
| 6181 | case False | |
| 54465 | 6182 | then have *: "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` S = S)" | 
| 6183 | using cone_iff[of S] assms by auto | |
| 53347 | 6184 |   {
 | 
| 6185 | fix c :: real | |
| 6186 | assume "c > 0" | |
| 6187 | then have "op *\<^sub>R c ` (convex hull S) = convex hull (op *\<^sub>R c ` S)" | |
| 6188 | using convex_hull_scaling[of _ S] by auto | |
| 6189 | also have "\<dots> = convex hull S" | |
| 60420 | 6190 | using * \<open>c > 0\<close> by auto | 
| 53347 | 6191 | finally have "op *\<^sub>R c ` (convex hull S) = convex hull S" | 
| 6192 | by auto | |
| 40377 | 6193 | } | 
| 53347 | 6194 | then have "0 \<in> convex hull S" "\<And>c. c > 0 \<Longrightarrow> (op *\<^sub>R c ` (convex hull S)) = (convex hull S)" | 
| 6195 | using * hull_subset[of S convex] by auto | |
| 6196 | then show ?thesis | |
| 60420 | 6197 |     using \<open>S \<noteq> {}\<close> cone_iff[of "convex hull S"] by auto
 | 
| 6198 | qed | |
| 6199 | ||
| 6200 | subsection \<open>Convex set as intersection of halfspaces\<close> | |
| 33175 | 6201 | |
| 6202 | 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: 
36844diff
changeset | 6203 |   fixes s :: "('a::euclidean_space) set"
 | 
| 33175 | 6204 | assumes "closed s" "convex s" | 
| 60585 | 6205 |   shows "s = \<Inter>{h. s \<subseteq> h \<and> (\<exists>a b. h = {x. inner a x \<le> b})}"
 | 
| 53347 | 6206 | apply (rule set_eqI) | 
| 6207 | apply rule | |
| 6208 | unfolding Inter_iff Ball_def mem_Collect_eq | |
| 6209 | apply (rule,rule,erule conjE) | |
| 6210 | proof - | |
| 54465 | 6211 | fix x | 
| 53347 | 6212 |   assume "\<forall>xa. s \<subseteq> xa \<and> (\<exists>a b. xa = {x. inner a x \<le> b}) \<longrightarrow> x \<in> xa"
 | 
| 6213 |   then have "\<forall>a b. s \<subseteq> {x. inner a x \<le> b} \<longrightarrow> x \<in> {x. inner a x \<le> b}"
 | |
| 6214 | by blast | |
| 6215 | then show "x \<in> s" | |
| 6216 | apply (rule_tac ccontr) | |
| 6217 | apply (drule separating_hyperplane_closed_point[OF assms(2,1)]) | |
| 6218 | apply (erule exE)+ | |
| 6219 | apply (erule_tac x="-a" in allE) | |
| 6220 | apply (erule_tac x="-b" in allE) | |
| 6221 | apply auto | |
| 6222 | done | |
| 33175 | 6223 | qed auto | 
| 6224 | ||
| 53347 | 6225 | |
| 60420 | 6226 | subsection \<open>Radon's theorem (from Lars Schewe)\<close> | 
| 33175 | 6227 | |
| 6228 | lemma radon_ex_lemma: | |
| 6229 | assumes "finite c" "affine_dependent c" | |
| 64267 | 6230 | 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 | 6231 | proof - | 
| 55697 | 6232 | from assms(2)[unfolded affine_dependent_explicit] | 
| 6233 | obtain s u where | |
| 64267 | 6234 | "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 | 6235 | by blast | 
| 53347 | 6236 | then show ?thesis | 
| 6237 | apply (rule_tac x="\<lambda>v. if v\<in>s then u v else 0" in exI) | |
| 64267 | 6238 | unfolding if_smult scaleR_zero_left and sum.inter_restrict[OF assms(1), symmetric] | 
| 53347 | 6239 | apply (auto simp add: Int_absorb1) | 
| 6240 | done | |
| 6241 | qed | |
| 33175 | 6242 | |
| 6243 | lemma radon_s_lemma: | |
| 53347 | 6244 | assumes "finite s" | 
| 64267 | 6245 | and "sum f s = (0::real)" | 
| 6246 |   shows "sum f {x\<in>s. 0 < f x} = - sum f {x\<in>s. f x < 0}"
 | |
| 53347 | 6247 | proof - | 
| 6248 | have *: "\<And>x. (if f x < 0 then f x else 0) + (if 0 < f x then f x else 0) = f x" | |
| 6249 | by auto | |
| 6250 | show ?thesis | |
| 64267 | 6251 | unfolding add_eq_0_iff[symmetric] and sum.inter_filter[OF assms(1)] | 
| 6252 | and sum.distrib[symmetric] and * | |
| 53347 | 6253 | 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: 
61531diff
changeset | 6254 | by assumption | 
| 53347 | 6255 | qed | 
| 33175 | 6256 | |
| 6257 | lemma radon_v_lemma: | |
| 53347 | 6258 | assumes "finite s" | 
| 64267 | 6259 | and "sum f s = 0" | 
| 53347 | 6260 | and "\<forall>x. g x = (0::real) \<longrightarrow> f x = (0::'a::euclidean_space)" | 
| 64267 | 6261 |   shows "(sum f {x\<in>s. 0 < g x}) = - sum f {x\<in>s. g x < 0}"
 | 
| 53347 | 6262 | proof - | 
| 6263 | have *: "\<And>x. (if 0 < g x then f x else 0) + (if g x < 0 then f x else 0) = f x" | |
| 6264 | using assms(3) by auto | |
| 6265 | show ?thesis | |
| 64267 | 6266 | unfolding eq_neg_iff_add_eq_0 and sum.inter_filter[OF assms(1)] | 
| 6267 | and sum.distrib[symmetric] and * | |
| 53347 | 6268 | using assms(2) | 
| 6269 | apply assumption | |
| 6270 | done | |
| 6271 | qed | |
| 33175 | 6272 | |
| 6273 | lemma radon_partition: | |
| 6274 | assumes "finite c" "affine_dependent c" | |
| 53347 | 6275 |   shows "\<exists>m p. m \<inter> p = {} \<and> m \<union> p = c \<and> (convex hull m) \<inter> (convex hull p) \<noteq> {}"
 | 
| 6276 | proof - | |
| 64267 | 6277 | 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 | 6278 | using radon_ex_lemma[OF assms] by auto | 
| 6279 |   have fin: "finite {x \<in> c. 0 < u x}" "finite {x \<in> c. 0 > u x}"
 | |
| 6280 | using assms(1) by auto | |
| 64267 | 6281 |   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}"
 | 
| 6282 |   have "sum u {x \<in> c. 0 < u x} \<noteq> 0"
 | |
| 53347 | 6283 | proof (cases "u v \<ge> 0") | 
| 6284 | case False | |
| 6285 | then have "u v < 0" by auto | |
| 6286 | then show ?thesis | |
| 6287 |     proof (cases "\<exists>w\<in>{x \<in> c. 0 < u x}. u w > 0")
 | |
| 6288 | case True | |
| 6289 | then show ?thesis | |
| 64267 | 6290 | using sum_nonneg_eq_0_iff[of _ u, OF fin(1)] by auto | 
| 33175 | 6291 | next | 
| 53347 | 6292 | case False | 
| 64267 | 6293 | then have "sum u c \<le> sum (\<lambda>x. if x=v then u v else 0) c" | 
| 6294 | apply (rule_tac sum_mono) | |
| 53347 | 6295 | apply auto | 
| 6296 | done | |
| 6297 | then show ?thesis | |
| 64267 | 6298 | unfolding sum.delta[OF assms(1)] using uv(2) and \<open>u v < 0\<close> and uv(1) by auto | 
| 53347 | 6299 | qed | 
| 64267 | 6300 | qed (insert sum_nonneg_eq_0_iff[of _ u, OF fin(1)] uv(2-3), auto) | 
| 6301 | ||
| 6302 |   then have *: "sum u {x\<in>c. u x > 0} > 0"
 | |
| 53347 | 6303 | unfolding less_le | 
| 6304 | apply (rule_tac conjI) | |
| 64267 | 6305 | apply (rule_tac sum_nonneg) | 
| 6306 | apply auto | |
| 6307 | done | |
| 6308 |   moreover have "sum u ({x \<in> c. 0 < u x} \<union> {x \<in> c. u x < 0}) = sum u c"
 | |
| 33175 | 6309 |     "(\<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 | 6310 | using assms(1) | 
| 64267 | 6311 | apply (rule_tac[!] sum.mono_neutral_left) | 
| 6312 | apply auto | |
| 6313 | done | |
| 6314 |   then have "sum u {x \<in> c. 0 < u x} = - sum u {x \<in> c. 0 > u x}"
 | |
| 53347 | 6315 |     "(\<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)"
 | 
| 6316 | unfolding eq_neg_iff_add_eq_0 | |
| 6317 | using uv(1,4) | |
| 64267 | 6318 | by (auto simp add: sum.union_inter_neutral[OF fin, symmetric]) | 
| 6319 |   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 | 6320 | apply rule | 
| 6321 | apply (rule mult_nonneg_nonneg) | |
| 6322 | using * | |
| 6323 | apply auto | |
| 6324 | done | |
| 6325 |   ultimately have "z \<in> convex hull {v \<in> c. u v \<le> 0}"
 | |
| 6326 | unfolding convex_hull_explicit mem_Collect_eq | |
| 6327 |     apply (rule_tac x="{v \<in> c. u v < 0}" in exI)
 | |
| 64267 | 6328 |     apply (rule_tac x="\<lambda>y. inverse (sum u {x\<in>c. u x > 0}) * - u y" in exI)
 | 
| 6329 | using assms(1) unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] and z_def | |
| 6330 | apply (auto simp add: sum_negf sum_distrib_left[symmetric]) | |
| 6331 | done | |
| 6332 |   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 | 6333 | apply rule | 
| 6334 | apply (rule mult_nonneg_nonneg) | |
| 6335 | using * | |
| 6336 | apply auto | |
| 6337 | done | |
| 6338 |   then have "z \<in> convex hull {v \<in> c. u v > 0}"
 | |
| 6339 | unfolding convex_hull_explicit mem_Collect_eq | |
| 6340 |     apply (rule_tac x="{v \<in> c. 0 < u v}" in exI)
 | |
| 64267 | 6341 |     apply (rule_tac x="\<lambda>y. inverse (sum u {x\<in>c. u x > 0}) * u y" in exI)
 | 
| 53347 | 6342 | using assms(1) | 
| 64267 | 6343 | unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] and z_def | 
| 53347 | 6344 | using * | 
| 64267 | 6345 | apply (auto simp add: sum_negf sum_distrib_left[symmetric]) | 
| 53347 | 6346 | done | 
| 6347 | ultimately show ?thesis | |
| 6348 |     apply (rule_tac x="{v\<in>c. u v \<le> 0}" in exI)
 | |
| 6349 |     apply (rule_tac x="{v\<in>c. u v > 0}" in exI)
 | |
| 6350 | apply auto | |
| 6351 | done | |
| 6352 | qed | |
| 6353 | ||
| 6354 | lemma radon: | |
| 6355 | assumes "affine_dependent c" | |
| 6356 |   obtains m p where "m \<subseteq> c" "p \<subseteq> c" "m \<inter> p = {}" "(convex hull m) \<inter> (convex hull p) \<noteq> {}"
 | |
| 6357 | proof - | |
| 55697 | 6358 | from assms[unfolded affine_dependent_explicit] | 
| 6359 | obtain s u where | |
| 64267 | 6360 | "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 | 6361 | by blast | 
| 53347 | 6362 | then have *: "finite s" "affine_dependent s" and s: "s \<subseteq> c" | 
| 6363 | unfolding affine_dependent_explicit by auto | |
| 55697 | 6364 | from radon_partition[OF *] | 
| 6365 |   obtain m p where "m \<inter> p = {}" "m \<union> p = s" "convex hull m \<inter> convex hull p \<noteq> {}"
 | |
| 6366 | by blast | |
| 53347 | 6367 | then show ?thesis | 
| 6368 | apply (rule_tac that[of p m]) | |
| 6369 | using s | |
| 6370 | apply auto | |
| 6371 | done | |
| 6372 | qed | |
| 6373 | ||
| 33175 | 6374 | |
| 60420 | 6375 | subsection \<open>Helly's theorem\<close> | 
| 33175 | 6376 | |
| 53347 | 6377 | lemma helly_induct: | 
| 6378 | fixes f :: "'a::euclidean_space set set" | |
| 6379 | assumes "card f = n" | |
| 6380 |     and "n \<ge> DIM('a) + 1"
 | |
| 60585 | 6381 |     and "\<forall>s\<in>f. convex s" "\<forall>t\<subseteq>f. card t = DIM('a) + 1 \<longrightarrow> \<Inter>t \<noteq> {}"
 | 
| 53347 | 6382 |   shows "\<Inter>f \<noteq> {}"
 | 
| 6383 | using assms | |
| 6384 | proof (induct n arbitrary: f) | |
| 6385 | case 0 | |
| 6386 | then show ?case by auto | |
| 6387 | next | |
| 6388 | case (Suc n) | |
| 6389 | have "finite f" | |
| 60420 | 6390 | using \<open>card f = Suc n\<close> by (auto intro: card_ge_0_finite) | 
| 53347 | 6391 |   show "\<Inter>f \<noteq> {}"
 | 
| 6392 |     apply (cases "n = DIM('a)")
 | |
| 6393 | apply (rule Suc(5)[rule_format]) | |
| 60420 | 6394 | unfolding \<open>card f = Suc n\<close> | 
| 53347 | 6395 | proof - | 
| 6396 |     assume ng: "n \<noteq> DIM('a)"
 | |
| 6397 |     then have "\<exists>X. \<forall>s\<in>f. X s \<in> \<Inter>(f - {s})"
 | |
| 6398 | apply (rule_tac bchoice) | |
| 6399 | unfolding ex_in_conv | |
| 6400 | apply (rule, rule Suc(1)[rule_format]) | |
| 60420 | 6401 | unfolding card_Diff_singleton_if[OF \<open>finite f\<close>] \<open>card f = Suc n\<close> | 
| 53347 | 6402 | defer | 
| 6403 | defer | |
| 6404 | apply (rule Suc(4)[rule_format]) | |
| 6405 | defer | |
| 6406 | apply (rule Suc(5)[rule_format]) | |
| 60420 | 6407 | using Suc(3) \<open>finite f\<close> | 
| 53347 | 6408 | apply auto | 
| 6409 | done | |
| 6410 |     then obtain X where X: "\<forall>s\<in>f. X s \<in> \<Inter>(f - {s})" by auto
 | |
| 6411 | show ?thesis | |
| 6412 | proof (cases "inj_on X f") | |
| 6413 | case False | |
| 6414 | then obtain s t where st: "s\<noteq>t" "s\<in>f" "t\<in>f" "X s = X t" | |
| 6415 | unfolding inj_on_def by auto | |
| 6416 |       then have *: "\<Inter>f = \<Inter>(f - {s}) \<inter> \<Inter>(f - {t})" by auto
 | |
| 6417 | show ?thesis | |
| 6418 | unfolding * | |
| 6419 | unfolding ex_in_conv[symmetric] | |
| 6420 | apply (rule_tac x="X s" in exI) | |
| 6421 | apply rule | |
| 6422 | apply (rule X[rule_format]) | |
| 6423 | using X st | |
| 6424 | apply auto | |
| 6425 | done | |
| 6426 | next | |
| 6427 | case True | |
| 6428 |       then obtain m p where mp: "m \<inter> p = {}" "m \<union> p = X ` f" "convex hull m \<inter> convex hull p \<noteq> {}"
 | |
| 6429 | using radon_partition[of "X ` f"] and affine_dependent_biggerset[of "X ` f"] | |
| 60420 | 6430 | unfolding card_image[OF True] and \<open>card f = Suc n\<close> | 
| 6431 | using Suc(3) \<open>finite f\<close> and ng | |
| 53347 | 6432 | by auto | 
| 6433 | have "m \<subseteq> X ` f" "p \<subseteq> X ` f" | |
| 6434 | using mp(2) by auto | |
| 6435 | then obtain g h where gh:"m = X ` g" "p = X ` h" "g \<subseteq> f" "h \<subseteq> f" | |
| 6436 | unfolding subset_image_iff by auto | |
| 6437 | then have "f \<union> (g \<union> h) = f" by auto | |
| 6438 | then have f: "f = g \<union> h" | |
| 6439 | using inj_on_Un_image_eq_iff[of X f "g \<union> h"] and True | |
| 6440 | unfolding mp(2)[unfolded image_Un[symmetric] gh] | |
| 6441 | by auto | |
| 6442 |       have *: "g \<inter> h = {}"
 | |
| 6443 | using mp(1) | |
| 6444 | unfolding gh | |
| 6445 | using inj_on_image_Int[OF True gh(3,4)] | |
| 6446 | by auto | |
| 6447 | have "convex hull (X ` h) \<subseteq> \<Inter>g" "convex hull (X ` g) \<subseteq> \<Inter>h" | |
| 6448 | apply (rule_tac [!] hull_minimal) | |
| 6449 | using Suc gh(3-4) | |
| 6450 | unfolding subset_eq | |
| 6451 | apply (rule_tac [2] convex_Inter, rule_tac [4] convex_Inter) | |
| 6452 | apply rule | |
| 6453 | prefer 3 | |
| 6454 | apply rule | |
| 6455 | proof - | |
| 6456 | fix x | |
| 6457 | assume "x \<in> X ` g" | |
| 55697 | 6458 | then obtain y where "y \<in> g" "x = X y" | 
| 6459 | unfolding image_iff .. | |
| 53347 | 6460 | then show "x \<in> \<Inter>h" | 
| 6461 | using X[THEN bspec[where x=y]] using * f by auto | |
| 6462 | next | |
| 6463 | fix x | |
| 6464 | assume "x \<in> X ` h" | |
| 55697 | 6465 | then obtain y where "y \<in> h" "x = X y" | 
| 6466 | unfolding image_iff .. | |
| 53347 | 6467 | then show "x \<in> \<Inter>g" | 
| 6468 | using X[THEN bspec[where x=y]] using * f by auto | |
| 6469 | qed auto | |
| 6470 | then show ?thesis | |
| 6471 | unfolding f using mp(3)[unfolded gh] by blast | |
| 6472 | qed | |
| 6473 | qed auto | |
| 6474 | qed | |
| 6475 | ||
| 6476 | lemma helly: | |
| 6477 | 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: 
36844diff
changeset | 6478 |   assumes "card f \<ge> DIM('a) + 1" "\<forall>s\<in>f. convex s"
 | 
| 60585 | 6479 |     and "\<forall>t\<subseteq>f. card t = DIM('a) + 1 \<longrightarrow> \<Inter>t \<noteq> {}"
 | 
| 53347 | 6480 |   shows "\<Inter>f \<noteq> {}"
 | 
| 6481 | apply (rule helly_induct) | |
| 6482 | using assms | |
| 6483 | apply auto | |
| 6484 | done | |
| 6485 | ||
| 33175 | 6486 | |
| 60420 | 6487 | subsection \<open>Epigraphs of convex functions\<close> | 
| 33175 | 6488 | |
| 53348 | 6489 | definition "epigraph s (f :: _ \<Rightarrow> real) = {xy. fst xy \<in> s \<and> f (fst xy) \<le> snd xy}"
 | 
| 6490 | ||
| 6491 | lemma mem_epigraph: "(x, y) \<in> epigraph s f \<longleftrightarrow> x \<in> s \<and> f x \<le> y" | |
| 6492 | unfolding epigraph_def by auto | |
| 6493 | ||
| 6494 | lemma convex_epigraph: "convex (epigraph s f) \<longleftrightarrow> convex_on s f \<and> convex s" | |
| 36338 | 6495 | unfolding convex_def convex_on_def | 
| 6496 | unfolding Ball_def split_paired_All epigraph_def | |
| 6497 | unfolding mem_Collect_eq fst_conv snd_conv fst_add snd_add fst_scaleR snd_scaleR Ball_def[symmetric] | |
| 53348 | 6498 | apply safe | 
| 6499 | defer | |
| 6500 | apply (erule_tac x=x in allE) | |
| 6501 | apply (erule_tac x="f x" in allE) | |
| 6502 | apply safe | |
| 6503 | apply (erule_tac x=xa in allE) | |
| 6504 | apply (erule_tac x="f xa" in allE) | |
| 6505 | prefer 3 | |
| 6506 | apply (rule_tac y="u * f a + v * f aa" in order_trans) | |
| 6507 | defer | |
| 6508 | apply (auto intro!:mult_left_mono add_mono) | |
| 6509 | done | |
| 6510 | ||
| 6511 | lemma convex_epigraphI: "convex_on s f \<Longrightarrow> convex s \<Longrightarrow> convex (epigraph s f)" | |
| 6512 | unfolding convex_epigraph by auto | |
| 6513 | ||
| 6514 | lemma convex_epigraph_convex: "convex s \<Longrightarrow> convex_on s f \<longleftrightarrow> convex(epigraph s f)" | |
| 6515 | by (simp add: convex_epigraph) | |
| 6516 | ||
| 33175 | 6517 | |
| 60420 | 6518 | subsubsection \<open>Use this to derive general bound property of convex function\<close> | 
| 33175 | 6519 | |
| 6520 | lemma convex_on: | |
| 6521 | assumes "convex s" | |
| 53348 | 6522 | shows "convex_on s f \<longleftrightarrow> | 
| 64267 | 6523 |     (\<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>
 | 
| 6524 |       f (sum (\<lambda>i. u i *\<^sub>R x i) {1..k} ) \<le> sum (\<lambda>i. u i * f(x i)) {1..k})"
 | |
| 33175 | 6525 | unfolding convex_epigraph_convex[OF assms] convex epigraph_def Ball_def mem_Collect_eq | 
| 64267 | 6526 | unfolding fst_sum snd_sum fst_scaleR snd_scaleR | 
| 36338 | 6527 | apply safe | 
| 6528 | apply (drule_tac x=k in spec) | |
| 6529 | apply (drule_tac x=u in spec) | |
| 6530 | apply (drule_tac x="\<lambda>i. (x i, f (x i))" in spec) | |
| 6531 | apply simp | |
| 53348 | 6532 | using assms[unfolded convex] | 
| 6533 | apply simp | |
| 6534 | apply (rule_tac y="\<Sum>i = 1..k. u i * f (fst (x i))" in order_trans) | |
| 6535 | defer | |
| 64267 | 6536 | apply (rule sum_mono) | 
| 53348 | 6537 | apply (erule_tac x=i in allE) | 
| 6538 | unfolding real_scaleR_def | |
| 6539 | apply (rule mult_left_mono) | |
| 6540 | using assms[unfolded convex] | |
| 6541 | apply auto | |
| 6542 | done | |
| 33175 | 6543 | |
| 36338 | 6544 | |
| 60420 | 6545 | subsection \<open>Convexity of general and special intervals\<close> | 
| 33175 | 6546 | |
| 6547 | lemma is_interval_convex: | |
| 53348 | 6548 | fixes s :: "'a::euclidean_space set" | 
| 6549 | assumes "is_interval s" | |
| 6550 | shows "convex s" | |
| 37732 
6432bf0d7191
generalize type of is_interval to class euclidean_space
 huffman parents: 
37673diff
changeset | 6551 | proof (rule convexI) | 
| 53348 | 6552 | fix x y and u v :: real | 
| 6553 | assume as: "x \<in> s" "y \<in> s" "0 \<le> u" "0 \<le> v" "u + v = 1" | |
| 6554 | then have *: "u = 1 - v" "1 - v \<ge> 0" and **: "v = 1 - u" "1 - u \<ge> 0" | |
| 6555 | by auto | |
| 6556 |   {
 | |
| 6557 | fix a b | |
| 6558 | assume "\<not> b \<le> u * a + v * b" | |
| 6559 | then have "u * a < (1 - v) * b" | |
| 6560 | unfolding not_le using as(4) by (auto simp add: field_simps) | |
| 6561 | then have "a < b" | |
| 6562 | unfolding * using as(4) *(2) | |
| 6563 | apply (rule_tac mult_left_less_imp_less[of "1 - v"]) | |
| 6564 | apply (auto simp add: field_simps) | |
| 6565 | done | |
| 6566 | then have "a \<le> u * a + v * b" | |
| 6567 | unfolding * using as(4) | |
| 6568 | by (auto simp add: field_simps intro!:mult_right_mono) | |
| 6569 | } | |
| 6570 | moreover | |
| 6571 |   {
 | |
| 6572 | fix a b | |
| 6573 | assume "\<not> u * a + v * b \<le> a" | |
| 6574 | then have "v * b > (1 - u) * a" | |
| 6575 | unfolding not_le using as(4) by (auto simp add: field_simps) | |
| 6576 | then have "a < b" | |
| 6577 | unfolding * using as(4) | |
| 6578 | apply (rule_tac mult_left_less_imp_less) | |
| 6579 | apply (auto simp add: field_simps) | |
| 6580 | done | |
| 6581 | then have "u * a + v * b \<le> b" | |
| 6582 | unfolding ** | |
| 6583 | using **(2) as(3) | |
| 6584 | by (auto simp add: field_simps intro!:mult_right_mono) | |
| 6585 | } | |
| 6586 | ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> s" | |
| 6587 | apply - | |
| 6588 | apply (rule assms[unfolded is_interval_def, rule_format, OF as(1,2)]) | |
| 6589 | using as(3-) DIM_positive[where 'a='a] | |
| 6590 | apply (auto simp: inner_simps) | |
| 6591 | done | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 6592 | qed | 
| 33175 | 6593 | |
| 6594 | lemma is_interval_connected: | |
| 53348 | 6595 | fixes s :: "'a::euclidean_space set" | 
| 33175 | 6596 | shows "is_interval s \<Longrightarrow> connected s" | 
| 6597 | using is_interval_convex convex_connected by auto | |
| 6598 | ||
| 62618 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 6599 | lemma convex_box [simp]: "convex (cbox a b)" "convex (box a (b::'a::euclidean_space))" | 
| 56188 | 6600 | apply (rule_tac[!] is_interval_convex)+ | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 6601 | using is_interval_box is_interval_cbox | 
| 53348 | 6602 | apply auto | 
| 6603 | done | |
| 33175 | 6604 | |
| 63928 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6605 | 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: 
63918diff
changeset | 6606 | lemma connected_imp_perfect: | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6607 | fixes a :: "'a::metric_space" | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6608 |   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: 
63918diff
changeset | 6609 | shows "a islimpt S" | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6610 | proof - | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6611 | 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: 
63918diff
changeset | 6612 | proof - | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6613 | 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: 
63918diff
changeset | 6614 | 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: 
63918diff
changeset | 6615 |     have "openin (subtopology euclidean S) {a}"
 | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6616 | 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: 
63918diff
changeset | 6617 |     moreover have "closedin (subtopology euclidean S) {a}"
 | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6618 | by (simp add: assms) | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6619 | ultimately show "False" | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6620 | using \<open>connected S\<close> connected_clopen S by blast | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6621 | qed | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6622 | then show ?thesis | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6623 | unfolding islimpt_def by blast | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6624 | qed | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6625 | |
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6626 | lemma connected_imp_perfect_aff_dim: | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6627 | "\<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: 
63918diff
changeset | 6628 | using aff_dim_sing connected_imp_perfect by blast | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6629 | |
| 61808 | 6630 | subsection \<open>On \<open>real\<close>, \<open>is_interval\<close>, \<open>convex\<close> and \<open>connected\<close> are all equivalent.\<close> | 
| 33175 | 6631 | |
| 6632 | lemma is_interval_1: | |
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6633 | "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: 
53600diff
changeset | 6634 | unfolding is_interval_def by auto | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6635 | |
| 54465 | 6636 | lemma is_interval_connected_1: | 
| 6637 | fixes s :: "real set" | |
| 6638 | shows "is_interval s \<longleftrightarrow> connected s" | |
| 6639 | apply rule | |
| 6640 | apply (rule is_interval_connected, assumption) | |
| 6641 | unfolding is_interval_1 | |
| 6642 | apply rule | |
| 6643 | apply rule | |
| 6644 | apply rule | |
| 6645 | apply rule | |
| 6646 | apply (erule conjE) | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 6647 | apply (rule ccontr) | 
| 54465 | 6648 | proof - | 
| 6649 | fix a b x | |
| 6650 | assume as: "connected s" "a \<in> s" "b \<in> s" "a \<le> x" "x \<le> b" "x \<notin> s" | |
| 6651 | then have *: "a < x" "x < b" | |
| 6652 | unfolding not_le [symmetric] by auto | |
| 6653 |   let ?halfl = "{..<x} "
 | |
| 6654 |   let ?halfr = "{x<..}"
 | |
| 6655 |   {
 | |
| 6656 | fix y | |
| 6657 | assume "y \<in> s" | |
| 60420 | 6658 | with \<open>x \<notin> s\<close> have "x \<noteq> y" by auto | 
| 54465 | 6659 | then have "y \<in> ?halfr \<union> ?halfl" by auto | 
| 6660 | } | |
| 6661 | moreover have "a \<in> ?halfl" "b \<in> ?halfr" using * by auto | |
| 6662 |   then have "?halfl \<inter> s \<noteq> {}" "?halfr \<inter> s \<noteq> {}"
 | |
| 6663 | using as(2-3) by auto | |
| 6664 | ultimately show False | |
| 6665 | apply (rule_tac notE[OF as(1)[unfolded connected_def]]) | |
| 6666 | apply (rule_tac x = ?halfl in exI) | |
| 6667 | apply (rule_tac x = ?halfr in exI) | |
| 6668 | apply rule | |
| 6669 | apply (rule open_lessThan) | |
| 6670 | apply rule | |
| 6671 | apply (rule open_greaterThan) | |
| 6672 | apply auto | |
| 6673 | done | |
| 6674 | qed | |
| 33175 | 6675 | |
| 6676 | lemma is_interval_convex_1: | |
| 54465 | 6677 | fixes s :: "real set" | 
| 6678 | shows "is_interval s \<longleftrightarrow> convex s" | |
| 6679 | by (metis is_interval_convex convex_connected is_interval_connected_1) | |
| 33175 | 6680 | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 6681 | lemma connected_compact_interval_1: | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 6682 |      "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: 
64394diff
changeset | 6683 | by (auto simp: is_interval_connected_1 [symmetric] is_interval_compact) | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 6684 | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6685 | lemma connected_convex_1: | 
| 54465 | 6686 | fixes s :: "real set" | 
| 6687 | shows "connected s \<longleftrightarrow> convex s" | |
| 6688 | by (metis is_interval_convex convex_connected is_interval_connected_1) | |
| 53348 | 6689 | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6690 | lemma connected_convex_1_gen: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6691 | fixes s :: "'a :: euclidean_space set" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6692 |   assumes "DIM('a) = 1"
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6693 | shows "connected s \<longleftrightarrow> convex s" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6694 | proof - | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6695 | obtain f:: "'a \<Rightarrow> real" where linf: "linear f" and "inj f" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6696 | using subspace_isomorphism [where 'a = 'a and 'b = real] | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6697 | by (metis DIM_real dim_UNIV subspace_UNIV assms) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6698 | then have "f -` (f ` s) = s" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6699 | by (simp add: inj_vimage_image_eq) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6700 | then show ?thesis | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6701 | 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: 
61426diff
changeset | 6702 | qed | 
| 53348 | 6703 | |
| 60420 | 6704 | subsection \<open>Another intermediate value theorem formulation\<close> | 
| 33175 | 6705 | |
| 53348 | 6706 | 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: 
61531diff
changeset | 6707 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" | 
| 53348 | 6708 | assumes "a \<le> b" | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6709 |     and "continuous_on {a..b} f"
 | 
| 53348 | 6710 | 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: 
61426diff
changeset | 6711 |   shows "\<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
 | 
| 56188 | 6712 | proof - | 
| 6713 | have "f a \<in> f ` cbox a b" "f b \<in> f ` cbox a b" | |
| 53348 | 6714 | apply (rule_tac[!] imageI) | 
| 6715 | using assms(1) | |
| 6716 | apply auto | |
| 6717 | done | |
| 6718 | then show ?thesis | |
| 56188 | 6719 | using connected_ivt_component[of "f ` cbox a b" "f a" "f b" k y] | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: 
66793diff
changeset | 6720 | by (simp add: connected_continuous_image assms) | 
| 53348 | 6721 | qed | 
| 6722 | ||
| 6723 | lemma ivt_increasing_component_1: | |
| 6724 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6725 |   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: 
61426diff
changeset | 6726 |     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 | 6727 | by (rule ivt_increasing_component_on_1) (auto simp add: continuous_at_imp_continuous_on) | 
| 6728 | ||
| 6729 | lemma ivt_decreasing_component_on_1: | |
| 6730 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" | |
| 6731 | assumes "a \<le> b" | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6732 |     and "continuous_on {a..b} f"
 | 
| 53348 | 6733 | and "(f b)\<bullet>k \<le> y" | 
| 6734 | and "y \<le> (f a)\<bullet>k" | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6735 |   shows "\<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
 | 
| 53348 | 6736 | apply (subst neg_equal_iff_equal[symmetric]) | 
| 44531 
1d477a2b1572
replace some continuous_on lemmas with more general versions
 huffman parents: 
44525diff
changeset | 6737 | using ivt_increasing_component_on_1[of a b "\<lambda>x. - f x" k "- y"] | 
| 53348 | 6738 | using assms using continuous_on_minus | 
| 6739 | apply auto | |
| 6740 | done | |
| 6741 | ||
| 6742 | lemma ivt_decreasing_component_1: | |
| 6743 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6744 |   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: 
61426diff
changeset | 6745 |     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 | 6746 | by (rule ivt_decreasing_component_on_1) (auto simp: continuous_at_imp_continuous_on) | 
| 6747 | ||
| 33175 | 6748 | |
| 60420 | 6749 | subsection \<open>A bound within a convex hull, and so an interval\<close> | 
| 33175 | 6750 | |
| 6751 | lemma convex_on_convex_hull_bound: | |
| 53348 | 6752 | assumes "convex_on (convex hull s) f" | 
| 6753 | and "\<forall>x\<in>s. f x \<le> b" | |
| 6754 | shows "\<forall>x\<in> convex hull s. f x \<le> b" | |
| 6755 | proof | |
| 6756 | fix x | |
| 6757 | assume "x \<in> convex hull s" | |
| 6758 | then obtain k u v where | |
| 64267 | 6759 |     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 | 6760 | unfolding convex_hull_indexed mem_Collect_eq by auto | 
| 53348 | 6761 | have "(\<Sum>i = 1..k. u i * f (v i)) \<le> b" | 
| 64267 | 6762 |     using sum_mono[of "{1..k}" "\<lambda>i. u i * f (v i)" "\<lambda>i. u i * b"]
 | 
| 6763 | unfolding sum_distrib_right[symmetric] obt(2) mult_1 | |
| 53348 | 6764 | apply (drule_tac meta_mp) | 
| 6765 | apply (rule mult_left_mono) | |
| 6766 | using assms(2) obt(1) | |
| 6767 | apply auto | |
| 6768 | done | |
| 6769 | then show "f x \<le> b" | |
| 6770 | using assms(1)[unfolded convex_on[OF convex_convex_hull], rule_format, of k u v] | |
| 6771 | unfolding obt(2-3) | |
| 6772 | using obt(1) and hull_subset[unfolded subset_eq, rule_format, of _ s] | |
| 6773 | by auto | |
| 6774 | qed | |
| 6775 | ||
| 64267 | 6776 | lemma inner_sum_Basis[simp]: "i \<in> Basis \<Longrightarrow> (\<Sum>Basis) \<bullet> i = 1" | 
| 6777 | 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: 
53600diff
changeset | 6778 | |
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6779 | lemma convex_set_plus: | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 6780 | assumes "convex S" and "convex T" shows "convex (S + T)" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 6781 | proof - | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 6782 |   have "convex (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})"
 | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6783 | using assms by (rule convex_sums) | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 6784 |   moreover have "(\<Union>x\<in> S. \<Union>y \<in> T. {x + y}) = S + T"
 | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6785 | unfolding set_plus_def by auto | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 6786 | finally show "convex (S + T)" . | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6787 | qed | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6788 | |
| 64267 | 6789 | lemma convex_set_sum: | 
| 55929 
91f245c23bc5
remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
 huffman parents: 
55928diff
changeset | 6790 | 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: 
55928diff
changeset | 6791 | shows "convex (\<Sum>i\<in>A. B i)" | 
| 
91f245c23bc5
remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
 huffman parents: 
55928diff
changeset | 6792 | proof (cases "finite A") | 
| 
91f245c23bc5
remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
 huffman parents: 
55928diff
changeset | 6793 | case True then show ?thesis using assms | 
| 
91f245c23bc5
remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
 huffman parents: 
55928diff
changeset | 6794 | by induct (auto simp: convex_set_plus) | 
| 
91f245c23bc5
remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
 huffman parents: 
55928diff
changeset | 6795 | qed auto | 
| 
91f245c23bc5
remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
 huffman parents: 
55928diff
changeset | 6796 | |
| 64267 | 6797 | lemma finite_set_sum: | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6798 | 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: 
53600diff
changeset | 6799 | 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: 
53600diff
changeset | 6800 | |
| 64267 | 6801 | lemma set_sum_eq: | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6802 |   "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: 
53600diff
changeset | 6803 | apply (induct set: finite) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6804 | apply simp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6805 | apply simp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6806 | apply (safe elim!: set_plus_elim) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6807 | apply (rule_tac x="fun_upd f x a" in exI) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6808 | apply simp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6809 | apply (rule_tac f="\<lambda>x. a + x" in arg_cong) | 
| 64267 | 6810 | apply (rule sum.cong [OF refl]) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6811 | apply clarsimp | 
| 57865 | 6812 | apply fast | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6813 | done | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6814 | |
| 64267 | 6815 | lemma box_eq_set_sum_Basis: | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6816 |   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 | 6817 | apply (subst set_sum_eq [OF finite_Basis]) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6818 | apply safe | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6819 | apply (fast intro: euclidean_representation [symmetric]) | 
| 64267 | 6820 | apply (subst inner_sum_left) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6821 | 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: 
53600diff
changeset | 6822 | apply (drule (1) bspec) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6823 | apply clarsimp | 
| 64267 | 6824 | apply (frule sum.remove [OF finite_Basis]) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6825 | apply (erule trans) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6826 | apply simp | 
| 64267 | 6827 | apply (rule sum.neutral) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6828 | apply clarsimp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6829 | apply (frule_tac x=i in bspec, assumption) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6830 | apply (drule_tac x=x in bspec, assumption) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6831 | apply clarsimp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6832 | 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: 
53600diff
changeset | 6833 | apply (rule ccontr) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6834 | apply simp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6835 | done | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6836 | |
| 64267 | 6837 | lemma convex_hull_set_sum: | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6838 | "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: 
53600diff
changeset | 6839 | proof (cases "finite A") | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6840 | assume "finite A" then show ?thesis | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6841 | by (induct set: finite, simp, simp add: convex_hull_set_plus) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6842 | qed simp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6843 | |
| 56188 | 6844 | lemma convex_hull_eq_real_cbox: | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6845 | fixes x y :: real assumes "x \<le> y" | 
| 56188 | 6846 |   shows "convex hull {x, y} = cbox x y"
 | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6847 | proof (rule hull_unique) | 
| 60420 | 6848 |   show "{x, y} \<subseteq> cbox x y" using \<open>x \<le> y\<close> by auto
 | 
| 56188 | 6849 | show "convex (cbox x y)" | 
| 6850 | by (rule convex_box) | |
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6851 | next | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6852 |   fix s assume "{x, y} \<subseteq> s" and "convex s"
 | 
| 56188 | 6853 | then show "cbox x y \<subseteq> s" | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6854 | 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: 
53600diff
changeset | 6855 | by - (clarify, simp (no_asm_use), fast) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6856 | qed | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 6857 | |
| 33175 | 6858 | 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: 
57418diff
changeset | 6859 |   "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: 
36844diff
changeset | 6860 | (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: 
50104diff
changeset | 6861 | proof - | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 6862 | have One[simp]: "\<And>i. i \<in> Basis \<Longrightarrow> One \<bullet> i = 1" | 
| 64267 | 6863 | by (simp add: inner_sum_left sum.If_cases inner_Basis) | 
| 56188 | 6864 |   have "?int = {x. \<forall>i\<in>Basis. x \<bullet> i \<in> cbox 0 1}"
 | 
| 6865 | by (auto simp: cbox_def) | |
| 6866 | also have "\<dots> = (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` cbox 0 1)" | |
| 64267 | 6867 | by (simp only: box_eq_set_sum_Basis) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6868 |   also have "\<dots> = (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` (convex hull {0, 1}))"
 | 
| 56188 | 6869 | 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: 
53600diff
changeset | 6870 |   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: 
53600diff
changeset | 6871 | by (simp only: convex_hull_linear_image linear_scaleR_left) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6872 |   also have "\<dots> = convex hull (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` {0, 1})"
 | 
| 64267 | 6873 | by (simp only: convex_hull_set_sum) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6874 |   also have "\<dots> = convex hull {x. \<forall>i\<in>Basis. x\<bullet>i \<in> {0, 1}}"
 | 
| 64267 | 6875 | by (simp only: box_eq_set_sum_Basis) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6876 |   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: 
53600diff
changeset | 6877 | by simp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6878 | 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: 
50104diff
changeset | 6879 | qed | 
| 33175 | 6880 | |
| 60420 | 6881 | text \<open>And this is a finite set of vertices.\<close> | 
| 33175 | 6882 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 6883 | lemma unit_cube_convex_hull: | 
| 56188 | 6884 | obtains s :: "'a::euclidean_space set" | 
| 6885 | where "finite s" and "cbox 0 (\<Sum>Basis) = convex hull s" | |
| 53348 | 6886 |   apply (rule that[of "{x::'a. \<forall>i\<in>Basis. x\<bullet>i=0 \<or> x\<bullet>i=1}"])
 | 
| 6887 | 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"]) | |
| 6888 | prefer 3 | |
| 6889 | apply (rule unit_interval_convex_hull) | |
| 6890 | apply rule | |
| 6891 | unfolding mem_Collect_eq | |
| 6892 | proof - | |
| 6893 | fix x :: 'a | |
| 6894 | 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: 
50104diff
changeset | 6895 | show "x \<in> (\<lambda>s. \<Sum>i\<in>Basis. (if i\<in>s then 1 else 0) *\<^sub>R i) ` Pow Basis" | 
| 53348 | 6896 |     apply (rule image_eqI[where x="{i. i\<in>Basis \<and> x\<bullet>i = 1}"])
 | 
| 6897 | using as | |
| 6898 | apply (subst euclidean_eq_iff) | |
| 57865 | 6899 | apply auto | 
| 53348 | 6900 | done | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 6901 | qed auto | 
| 33175 | 6902 | |
| 60420 | 6903 | text \<open>Hence any cube (could do any nonempty interval).\<close> | 
| 33175 | 6904 | |
| 6905 | lemma cube_convex_hull: | |
| 53348 | 6906 | assumes "d > 0" | 
| 56188 | 6907 | obtains s :: "'a::euclidean_space set" where | 
| 6908 | "finite s" and "cbox (x - (\<Sum>i\<in>Basis. d*\<^sub>Ri)) (x + (\<Sum>i\<in>Basis. d*\<^sub>Ri)) = convex hull s" | |
| 53348 | 6909 | proof - | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 6910 | let ?d = "(\<Sum>i\<in>Basis. d*\<^sub>Ri)::'a" | 
| 56188 | 6911 | have *: "cbox (x - ?d) (x + ?d) = (\<lambda>y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 (\<Sum>Basis)" | 
| 53348 | 6912 | apply (rule set_eqI, rule) | 
| 6913 | unfolding image_iff | |
| 6914 | defer | |
| 6915 | apply (erule bexE) | |
| 6916 | proof - | |
| 6917 | fix y | |
| 56188 | 6918 | assume as: "y\<in>cbox (x - ?d) (x + ?d)" | 
| 6919 | 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: 
57865diff
changeset | 6920 | using assms by (simp add: mem_box field_simps inner_simps) | 
| 60420 | 6921 | 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: 
57865diff
changeset | 6922 | by (intro bexI[of _ "inverse (2 * d) *\<^sub>R (y - (x - ?d))"]) auto | 
| 33175 | 6923 | next | 
| 53348 | 6924 | fix y z | 
| 56188 | 6925 | 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: 
50104diff
changeset | 6926 | have "\<And>i. i\<in>Basis \<Longrightarrow> 0 \<le> d * (z \<bullet> i) \<and> d * (z \<bullet> i) \<le> d" | 
| 56188 | 6927 | using assms as(1)[unfolded mem_box] | 
| 53348 | 6928 | apply (erule_tac x=i in ballE) | 
| 6929 | apply rule | |
| 56536 | 6930 | prefer 2 | 
| 53348 | 6931 | apply (rule mult_right_le_one_le) | 
| 6932 | using assms | |
| 6933 | apply auto | |
| 6934 | done | |
| 56188 | 6935 | then show "y \<in> cbox (x - ?d) (x + ?d)" | 
| 6936 | unfolding as(2) mem_box | |
| 53348 | 6937 | apply - | 
| 6938 | apply rule | |
| 56188 | 6939 | using as(1)[unfolded mem_box] | 
| 53348 | 6940 | apply (erule_tac x=i in ballE) | 
| 6941 | using assms | |
| 6942 | apply (auto simp: inner_simps) | |
| 6943 | done | |
| 6944 | qed | |
| 56188 | 6945 | obtain s where "finite s" "cbox 0 (\<Sum>Basis::'a) = convex hull s" | 
| 53348 | 6946 | using unit_cube_convex_hull by auto | 
| 6947 | then show ?thesis | |
| 6948 | apply (rule_tac that[of "(\<lambda>y. x - ?d + (2 * d) *\<^sub>R y)` s"]) | |
| 6949 | unfolding * and convex_hull_affinity | |
| 6950 | apply auto | |
| 6951 | done | |
| 6952 | qed | |
| 6953 | ||
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6954 | 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: 
62950diff
changeset | 6955 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6956 | lemma image_stretch_interval: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6957 | "(\<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: 
62950diff
changeset | 6958 |   (if (cbox a b) = {} then {} else
 | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6959 | 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: 
62950diff
changeset | 6960 | (\<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: 
62950diff
changeset | 6961 | proof cases | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6962 |   assume *: "cbox a b \<noteq> {}"
 | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6963 | show ?thesis | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6964 | 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: 
62950diff
changeset | 6965 | 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: 
62950diff
changeset | 6966 | apply (subst choice_Basis_iff[symmetric]) | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6967 | proof (intro allI ball_cong refl) | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6968 | 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: 
62950diff
changeset | 6969 | 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: 
62950diff
changeset | 6970 | unfolding box_ne_empty by auto | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6971 | 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: 
62950diff
changeset | 6972 | 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: 
62950diff
changeset | 6973 | proof (cases "m i = 0") | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6974 | case True | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6975 | 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: 
62950diff
changeset | 6976 | next | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6977 | case False | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6978 | 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: 
62950diff
changeset | 6979 | by (auto simp add: field_simps) | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6980 | from False have | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6981 | "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: 
62950diff
changeset | 6982 | "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: 
62950diff
changeset | 6983 | 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: 
62950diff
changeset | 6984 | 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: 
62950diff
changeset | 6985 | 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: 
62950diff
changeset | 6986 | qed | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6987 | qed | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6988 | qed simp | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6989 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6990 | lemma interval_image_stretch_interval: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6991 | "\<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: 
62950diff
changeset | 6992 | unfolding image_stretch_interval by auto | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6993 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6994 | 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: 
62950diff
changeset | 6995 | 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: 
62950diff
changeset | 6996 | 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: 
62950diff
changeset | 6997 | 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: 
62950diff
changeset | 6998 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6999 | lemma cbox_image_unit_interval: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7000 | fixes a :: "'a::euclidean_space" | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7001 |   assumes "cbox a b \<noteq> {}"
 | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7002 | shows "cbox a b = | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7003 | 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: 
62950diff
changeset | 7004 | using assms | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7005 | apply (simp add: box_ne_empty image_stretch_interval cbox_translation [symmetric]) | 
| 64267 | 7006 | 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: 
62950diff
changeset | 7007 | done | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7008 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7009 | lemma closed_interval_as_convex_hull: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7010 | fixes a :: "'a::euclidean_space" | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7011 | 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: 
62950diff
changeset | 7012 | proof (cases "cbox a b = {}")
 | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7013 | 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: 
62950diff
changeset | 7014 | by blast | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7015 | next | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7016 | case False | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7017 | 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: 
62950diff
changeset | 7018 | by (blast intro: unit_cube_convex_hull) | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7019 | have lin: "linear (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k)" | 
| 64267 | 7020 | 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: 
62950diff
changeset | 7021 | 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: 
62950diff
changeset | 7022 | 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: 
62950diff
changeset | 7023 | then show ?thesis | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7024 | apply (rule that) | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7025 | 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: 
62950diff
changeset | 7026 | 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: 
62950diff
changeset | 7027 | done | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7028 | qed | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7029 | |
| 33175 | 7030 | |
| 60420 | 7031 | subsection \<open>Bounded convex function on open set is continuous\<close> | 
| 33175 | 7032 | |
| 7033 | lemma convex_on_bounded_continuous: | |
| 36338 | 7034 |   fixes s :: "('a::real_normed_vector) set"
 | 
| 53348 | 7035 | assumes "open s" | 
| 7036 | and "convex_on s f" | |
| 61945 | 7037 | and "\<forall>x\<in>s. \<bar>f x\<bar> \<le> b" | 
| 33175 | 7038 | shows "continuous_on s f" | 
| 53348 | 7039 | apply (rule continuous_at_imp_continuous_on) | 
| 7040 | unfolding continuous_at_real_range | |
| 7041 | proof (rule,rule,rule) | |
| 7042 | fix x and e :: real | |
| 7043 | assume "x \<in> s" "e > 0" | |
| 63040 | 7044 | define B where "B = \<bar>b\<bar> + 1" | 
| 61945 | 7045 | have B: "0 < B" "\<And>x. x\<in>s \<Longrightarrow> \<bar>f x\<bar> \<le> B" | 
| 53348 | 7046 | unfolding B_def | 
| 7047 | defer | |
| 7048 | apply (drule assms(3)[rule_format]) | |
| 7049 | apply auto | |
| 7050 | done | |
| 7051 | obtain k where "k > 0" and k: "cball x k \<subseteq> s" | |
| 7052 | using assms(1)[unfolded open_contains_cball, THEN bspec[where x=x]] | |
| 60420 | 7053 | using \<open>x\<in>s\<close> by auto | 
| 33175 | 7054 | show "\<exists>d>0. \<forall>x'. norm (x' - x) < d \<longrightarrow> \<bar>f x' - f x\<bar> < e" | 
| 53348 | 7055 | apply (rule_tac x="min (k / 2) (e / (2 * B) * k)" in exI) | 
| 7056 | apply rule | |
| 7057 | defer | |
| 7058 | proof (rule, rule) | |
| 7059 | fix y | |
| 7060 | assume as: "norm (y - x) < min (k / 2) (e / (2 * B) * k)" | |
| 7061 | show "\<bar>f y - f x\<bar> < e" | |
| 7062 | proof (cases "y = x") | |
| 7063 | case False | |
| 63040 | 7064 | define t where "t = k / norm (y - x)" | 
| 53348 | 7065 | have "2 < t" "0<t" | 
| 60420 | 7066 | unfolding t_def using as False and \<open>k>0\<close> | 
| 53348 | 7067 | by (auto simp add:field_simps) | 
| 7068 | have "y \<in> s" | |
| 7069 | apply (rule k[unfolded subset_eq,rule_format]) | |
| 7070 | unfolding mem_cball dist_norm | |
| 7071 | apply (rule order_trans[of _ "2 * norm (x - y)"]) | |
| 7072 | using as | |
| 7073 | by (auto simp add: field_simps norm_minus_commute) | |
| 7074 |       {
 | |
| 63040 | 7075 | define w where "w = x + t *\<^sub>R (y - x)" | 
| 53348 | 7076 | have "w \<in> s" | 
| 7077 | unfolding w_def | |
| 7078 | apply (rule k[unfolded subset_eq,rule_format]) | |
| 7079 | unfolding mem_cball dist_norm | |
| 7080 | unfolding t_def | |
| 60420 | 7081 | using \<open>k>0\<close> | 
| 53348 | 7082 | apply auto | 
| 7083 | done | |
| 7084 | have "(1 / t) *\<^sub>R x + - x + ((t - 1) / t) *\<^sub>R x = (1 / t - 1 + (t - 1) / t) *\<^sub>R x" | |
| 7085 | by (auto simp add: algebra_simps) | |
| 7086 | also have "\<dots> = 0" | |
| 60420 | 7087 | using \<open>t > 0\<close> by (auto simp add:field_simps) | 
| 53348 | 7088 | finally have w: "(1 / t) *\<^sub>R w + ((t - 1) / t) *\<^sub>R x = y" | 
| 60420 | 7089 | unfolding w_def using False and \<open>t > 0\<close> | 
| 53348 | 7090 | by (auto simp add: algebra_simps) | 
| 7091 | have "2 * B < e * t" | |
| 60420 | 7092 | unfolding t_def using \<open>0 < e\<close> \<open>0 < k\<close> \<open>B > 0\<close> and as and False | 
| 53348 | 7093 | by (auto simp add:field_simps) | 
| 7094 | then have "(f w - f x) / t < e" | |
| 60420 | 7095 | using B(2)[OF \<open>w\<in>s\<close>] and B(2)[OF \<open>x\<in>s\<close>] | 
| 7096 | using \<open>t > 0\<close> by (auto simp add:field_simps) | |
| 53348 | 7097 | then have th1: "f y - f x < e" | 
| 7098 | apply - | |
| 7099 | apply (rule le_less_trans) | |
| 7100 | defer | |
| 7101 | apply assumption | |
| 33175 | 7102 | using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w] | 
| 60420 | 7103 | using \<open>0 < t\<close> \<open>2 < t\<close> and \<open>x \<in> s\<close> \<open>w \<in> s\<close> | 
| 53348 | 7104 | by (auto simp add:field_simps) | 
| 7105 | } | |
| 49531 | 7106 | moreover | 
| 53348 | 7107 |       {
 | 
| 63040 | 7108 | define w where "w = x - t *\<^sub>R (y - x)" | 
| 53348 | 7109 | have "w \<in> s" | 
| 7110 | unfolding w_def | |
| 7111 | apply (rule k[unfolded subset_eq,rule_format]) | |
| 7112 | unfolding mem_cball dist_norm | |
| 7113 | unfolding t_def | |
| 60420 | 7114 | using \<open>k > 0\<close> | 
| 53348 | 7115 | apply auto | 
| 7116 | done | |
| 7117 | have "(1 / (1 + t)) *\<^sub>R x + (t / (1 + t)) *\<^sub>R x = (1 / (1 + t) + t / (1 + t)) *\<^sub>R x" | |
| 7118 | by (auto simp add: algebra_simps) | |
| 7119 | also have "\<dots> = x" | |
| 60420 | 7120 | using \<open>t > 0\<close> by (auto simp add:field_simps) | 
| 53348 | 7121 | finally have w: "(1 / (1+t)) *\<^sub>R w + (t / (1 + t)) *\<^sub>R y = x" | 
| 60420 | 7122 | unfolding w_def using False and \<open>t > 0\<close> | 
| 53348 | 7123 | by (auto simp add: algebra_simps) | 
| 7124 | have "2 * B < e * t" | |
| 7125 | unfolding t_def | |
| 60420 | 7126 | using \<open>0 < e\<close> \<open>0 < k\<close> \<open>B > 0\<close> and as and False | 
| 53348 | 7127 | by (auto simp add:field_simps) | 
| 7128 | then have *: "(f w - f y) / t < e" | |
| 60420 | 7129 | using B(2)[OF \<open>w\<in>s\<close>] and B(2)[OF \<open>y\<in>s\<close>] | 
| 7130 | using \<open>t > 0\<close> | |
| 53348 | 7131 | by (auto simp add:field_simps) | 
| 49531 | 7132 | have "f x \<le> 1 / (1 + t) * f w + (t / (1 + t)) * f y" | 
| 33175 | 7133 | using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w] | 
| 60420 | 7134 | using \<open>0 < t\<close> \<open>2 < t\<close> and \<open>y \<in> s\<close> \<open>w \<in> s\<close> | 
| 53348 | 7135 | by (auto simp add:field_simps) | 
| 7136 | also have "\<dots> = (f w + t * f y) / (1 + t)" | |
| 60420 | 7137 | using \<open>t > 0\<close> by (auto simp add: divide_simps) | 
| 53348 | 7138 | also have "\<dots> < e + f y" | 
| 60420 | 7139 | using \<open>t > 0\<close> * \<open>e > 0\<close> by (auto simp add: field_simps) | 
| 53348 | 7140 | finally have "f x - f y < e" by auto | 
| 7141 | } | |
| 49531 | 7142 | ultimately show ?thesis by auto | 
| 60420 | 7143 | qed (insert \<open>0<e\<close>, auto) | 
| 7144 | qed (insert \<open>0<e\<close> \<open>0<k\<close> \<open>0<B\<close>, auto simp: field_simps) | |
| 7145 | qed | |
| 7146 | ||
| 7147 | ||
| 7148 | subsection \<open>Upper bound on a ball implies upper and lower bounds\<close> | |
| 33175 | 7149 | |
| 7150 | lemma convex_bounds_lemma: | |
| 36338 | 7151 | fixes x :: "'a::real_normed_vector" | 
| 53348 | 7152 | assumes "convex_on (cball x e) f" | 
| 7153 | and "\<forall>y \<in> cball x e. f y \<le> b" | |
| 61945 | 7154 | shows "\<forall>y \<in> cball x e. \<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" | 
| 53348 | 7155 | apply rule | 
| 7156 | proof (cases "0 \<le> e") | |
| 7157 | case True | |
| 7158 | fix y | |
| 7159 | assume y: "y \<in> cball x e" | |
| 63040 | 7160 | define z where "z = 2 *\<^sub>R x - y" | 
| 53348 | 7161 | have *: "x - (2 *\<^sub>R x - y) = y - x" | 
| 7162 | by (simp add: scaleR_2) | |
| 7163 | have z: "z \<in> cball x e" | |
| 7164 | using y unfolding z_def mem_cball dist_norm * by (auto simp add: norm_minus_commute) | |
| 7165 | have "(1 / 2) *\<^sub>R y + (1 / 2) *\<^sub>R z = x" | |
| 7166 | unfolding z_def by (auto simp add: algebra_simps) | |
| 7167 | then show "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" | |
| 7168 | using assms(1)[unfolded convex_on_def,rule_format, OF y z, of "1/2" "1/2"] | |
| 7169 | using assms(2)[rule_format,OF y] assms(2)[rule_format,OF z] | |
| 7170 | by (auto simp add:field_simps) | |
| 7171 | next | |
| 7172 | case False | |
| 7173 | fix y | |
| 7174 | assume "y \<in> cball x e" | |
| 7175 | then have "dist x y < 0" | |
| 7176 | using False unfolding mem_cball not_le by (auto simp del: dist_not_less_zero) | |
| 7177 | then show "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" | |
| 7178 | using zero_le_dist[of x y] by auto | |
| 7179 | qed | |
| 7180 | ||
| 33175 | 7181 | |
| 60420 | 7182 | subsubsection \<open>Hence a convex function on an open set is continuous\<close> | 
| 33175 | 7183 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 7184 | 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: 
50104diff
changeset | 7185 | by auto | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 7186 | |
| 33175 | 7187 | lemma convex_on_continuous: | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7188 |   assumes "open (s::('a::euclidean_space) set)" "convex_on s f"
 | 
| 33175 | 7189 | shows "continuous_on s f" | 
| 53348 | 7190 | unfolding continuous_on_eq_continuous_at[OF assms(1)] | 
| 7191 | proof | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 7192 | note dimge1 = DIM_positive[where 'a='a] | 
| 53348 | 7193 | fix x | 
| 7194 | assume "x \<in> s" | |
| 7195 | then obtain e where e: "cball x e \<subseteq> s" "e > 0" | |
| 7196 | using assms(1) unfolding open_contains_cball by auto | |
| 63040 | 7197 |   define d where "d = e / real DIM('a)"
 | 
| 53348 | 7198 | have "0 < d" | 
| 60420 | 7199 | 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: 
50104diff
changeset | 7200 | 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: 
53600diff
changeset | 7201 | obtain c | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7202 | where c: "finite c" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7203 | and c1: "convex hull c \<subseteq> cball x e" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7204 | and c2: "cball x d \<subseteq> convex hull c" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7205 | proof | 
| 63040 | 7206 |     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: 
53600diff
changeset | 7207 | show "finite c" | 
| 64267 | 7208 | unfolding c_def by (simp add: finite_set_sum) | 
| 56188 | 7209 |     have 1: "convex hull c = {a. \<forall>i\<in>Basis. a \<bullet> i \<in> cbox (x \<bullet> i - d) (x \<bullet> i + d)}"
 | 
| 64267 | 7210 | unfolding box_eq_set_sum_Basis | 
| 7211 | unfolding c_def convex_hull_set_sum | |
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7212 | apply (subst convex_hull_linear_image [symmetric]) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7213 | apply (simp add: linear_iff scaleR_add_left) | 
| 64267 | 7214 | apply (rule sum.cong [OF refl]) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7215 | apply (rule image_cong [OF _ refl]) | 
| 56188 | 7216 | apply (rule convex_hull_eq_real_cbox) | 
| 60420 | 7217 | apply (cut_tac \<open>0 < d\<close>, simp) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7218 | done | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7219 |     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: 
53600diff
changeset | 7220 | by (simp add: dist_norm abs_le_iff algebra_simps) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7221 | show "cball x d \<subseteq> convex hull c" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7222 | unfolding 2 | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7223 | apply clarsimp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7224 | apply (simp only: dist_norm) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7225 | apply (subst inner_diff_left [symmetric]) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7226 | apply simp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7227 | apply (erule (1) order_trans [OF Basis_le_norm]) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7228 | done | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7229 | 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: 
61531diff
changeset | 7230 | by (simp add: d_def DIM_positive) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7231 | show "convex hull c \<subseteq> cball x e" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7232 | unfolding 2 | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7233 | apply clarsimp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7234 | apply (subst euclidean_dist_l2) | 
| 64267 | 7235 | apply (rule order_trans [OF setL2_le_sum]) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7236 | apply (rule zero_le_dist) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7237 | unfolding e' | 
| 64267 | 7238 | apply (rule sum_mono) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7239 | apply simp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7240 | done | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7241 | qed | 
| 63040 | 7242 | define k where "k = Max (f ` c)" | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7243 | 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: 
50104diff
changeset | 7244 | 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: 
50104diff
changeset | 7245 | apply(rule subset_trans[OF _ e(1)]) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7246 | apply(rule c1) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7247 | done | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7248 | then have k: "\<forall>y\<in>convex hull c. f y \<le> k" | 
| 53348 | 7249 | apply (rule_tac convex_on_convex_hull_bound) | 
| 7250 | apply assumption | |
| 7251 | unfolding k_def | |
| 7252 | apply (rule, rule Max_ge) | |
| 7253 | using c(1) | |
| 7254 | apply auto | |
| 7255 | done | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 7256 | 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: 
50104diff
changeset | 7257 | unfolding d_def | 
| 53348 | 7258 | apply (rule mult_imp_div_pos_le) | 
| 60420 | 7259 | 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: 
50104diff
changeset | 7260 | 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: 
50104diff
changeset | 7261 | 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: 
50104diff
changeset | 7262 | done | 
| 53348 | 7263 | then have dsube: "cball x d \<subseteq> cball x e" | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7264 | by (rule subset_cball) | 
| 53348 | 7265 | have conv: "convex_on (cball x d) f" | 
| 7266 | apply (rule convex_on_subset) | |
| 7267 | apply (rule convex_on_subset[OF assms(2)]) | |
| 7268 | apply (rule e(1)) | |
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7269 | apply (rule dsube) | 
| 53348 | 7270 | done | 
| 61945 | 7271 | 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: 
53600diff
changeset | 7272 | apply (rule convex_bounds_lemma) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7273 | apply (rule ballI) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7274 | apply (rule k [rule_format]) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7275 | apply (erule rev_subsetD) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7276 | apply (rule c2) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7277 | done | 
| 53348 | 7278 | then have "continuous_on (ball x d) f" | 
| 7279 | apply (rule_tac convex_on_bounded_continuous) | |
| 7280 | apply (rule open_ball, rule convex_on_subset[OF conv]) | |
| 7281 | apply (rule ball_subset_cball) | |
| 33270 | 7282 | apply force | 
| 7283 | done | |
| 53348 | 7284 | then show "continuous (at x) f" | 
| 7285 | unfolding continuous_on_eq_continuous_at[OF open_ball] | |
| 60420 | 7286 | using \<open>d > 0\<close> by auto | 
| 7287 | qed | |
| 7288 | ||
| 33175 | 7289 | end |