| author | wenzelm | 
| Sun, 07 Jan 2018 16:55:45 +0100 | |
| changeset 67362 | 221612c942de | 
| parent 67155 | 9e5b05d54f9d | 
| child 67399 | eab6ce8368fa | 
| 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}"
 | 
| 67135 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 282 | using convex_halfspace_lt[of "-a" "-b"] by auto | 
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 283 | |
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 284 | lemma convex_halfspace_Re_ge: "convex {x. Re x \<ge> b}"
 | 
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 285 | using convex_halfspace_ge[of b "1::complex"] by simp | 
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 286 | |
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 287 | lemma convex_halfspace_Re_le: "convex {x. Re x \<le> b}"
 | 
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 288 | using convex_halfspace_le[of "1::complex" b] by simp | 
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 289 | |
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 290 | lemma convex_halfspace_Im_ge: "convex {x. Im x \<ge> b}"
 | 
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 291 | using convex_halfspace_ge[of b \<i>] by simp | 
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 292 | |
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 293 | lemma convex_halfspace_Im_le: "convex {x. Im x \<le> b}"
 | 
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 294 | using convex_halfspace_le[of \<i> b] by simp | 
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 295 | |
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 296 | lemma convex_halfspace_Re_gt: "convex {x. Re x > b}"
 | 
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 297 | using convex_halfspace_gt[of b "1::complex"] by simp | 
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 298 | |
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 299 | lemma convex_halfspace_Re_lt: "convex {x. Re x < b}"
 | 
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 300 | using convex_halfspace_lt[of "1::complex" b] by simp | 
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 301 | |
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 302 | lemma convex_halfspace_Im_gt: "convex {x. Im x > b}"
 | 
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 303 | using convex_halfspace_gt[of b \<i>] by simp | 
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 304 | |
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 305 | lemma convex_halfspace_Im_lt: "convex {x. Im x < b}"
 | 
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66939diff
changeset | 306 | using convex_halfspace_lt[of \<i> b] by simp | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 307 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 308 | lemma convex_real_interval [iff]: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 309 | fixes a b :: "real" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 310 |   shows "convex {a..}" and "convex {..b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 311 |     and "convex {a<..}" and "convex {..<b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 312 |     and "convex {a..b}" and "convex {a<..b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 313 |     and "convex {a..<b}" and "convex {a<..<b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 314 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 315 |   have "{a..} = {x. a \<le> inner 1 x}"
 | 
| 
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 1: "convex {a..}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 318 | by (simp only: convex_halfspace_ge) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 319 |   have "{..b} = {x. inner 1 x \<le> 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 2: "convex {..b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 322 | by (simp only: convex_halfspace_le) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 323 |   have "{a<..} = {x. a < inner 1 x}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 324 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 325 |   then show 3: "convex {a<..}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 326 | by (simp only: convex_halfspace_gt) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 327 |   have "{..<b} = {x. inner 1 x < b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 328 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 329 |   then show 4: "convex {..<b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 330 | by (simp only: convex_halfspace_lt) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 331 |   have "{a..b} = {a..} \<inter> {..b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 332 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 333 |   then show "convex {a..b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 334 | by (simp only: convex_Int 1 2) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 335 |   have "{a<..b} = {a<..} \<inter> {..b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 336 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 337 |   then show "convex {a<..b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 338 | by (simp only: convex_Int 3 2) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 339 |   have "{a..<b} = {a..} \<inter> {..<b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 340 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 341 |   then show "convex {a..<b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 342 | by (simp only: convex_Int 1 4) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 343 |   have "{a<..<b} = {a<..} \<inter> {..<b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 344 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 345 |   then show "convex {a<..<b}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 346 | by (simp only: convex_Int 3 4) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 347 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 348 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 349 | lemma convex_Reals: "convex \<real>" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 350 | by (simp add: convex_def scaleR_conv_of_real) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 351 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 352 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 353 | 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 | 354 | |
| 64267 | 355 | lemma convex_sum: | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 356 | fixes C :: "'a::real_vector set" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 357 | assumes "finite s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 358 | and "convex C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 359 | and "(\<Sum> i \<in> s. a i) = 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 360 | 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 | 361 | 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 | 362 | 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 | 363 | using assms(1,3,4,5) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 364 | proof (induct arbitrary: a set: finite) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 365 | case empty | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 366 | then show ?case by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 367 | next | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 368 | case (insert i s) note IH = this(3) | 
| 64267 | 369 | have "a i + sum a s = 1" | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 370 | and "0 \<le> a i" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 371 | and "\<forall>j\<in>s. 0 \<le> a j" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 372 | and "y i \<in> C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 373 | and "\<forall>j\<in>s. y j \<in> C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 374 | using insert.hyps(1,2) insert.prems by simp_all | 
| 64267 | 375 | then have "0 \<le> sum a s" | 
| 376 | by (simp add: sum_nonneg) | |
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 377 | have "a i *\<^sub>R y i + (\<Sum>j\<in>s. a j *\<^sub>R y j) \<in> C" | 
| 64267 | 378 | proof (cases "sum a s = 0") | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 379 | case True | 
| 64267 | 380 | 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 | 381 | by simp | 
| 64267 | 382 | 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 | 383 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 384 | 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 | 385 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 386 | next | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 387 | case False | 
| 64267 | 388 | 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 | 389 | by simp | 
| 64267 | 390 | 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 | 391 | using \<open>\<forall>j\<in>s. 0 \<le> a j\<close> and \<open>\<forall>j\<in>s. y j \<in> C\<close> | 
| 64267 | 392 | by (simp add: IH sum_divide_distrib [symmetric]) | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 393 | from \<open>convex C\<close> and \<open>y i \<in> C\<close> and this and \<open>0 \<le> a i\<close> | 
| 64267 | 394 | and \<open>0 \<le> sum a s\<close> and \<open>a i + sum a s = 1\<close> | 
| 395 | 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 | 396 | by (rule convexD) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 397 | then show ?thesis | 
| 64267 | 398 | by (simp add: scaleR_sum_right False) | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 399 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 400 | 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 | 401 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 402 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 403 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 404 | lemma convex: | 
| 64267 | 405 |   "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)
 | 
| 406 |       \<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 | 407 | proof safe | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 408 | fix k :: nat | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 409 | fix u :: "nat \<Rightarrow> real" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 410 | fix x | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 411 | assume "convex s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 412 | "\<forall>i. 1 \<le> i \<and> i \<le> k \<longrightarrow> 0 \<le> u i \<and> x i \<in> s" | 
| 64267 | 413 |     "sum u {1..k} = 1"
 | 
| 414 |   with convex_sum[of "{1 .. k}" s] show "(\<Sum>j\<in>{1 .. k}. u j *\<^sub>R x j) \<in> s"
 | |
| 415 | by auto | |
| 416 | next | |
| 417 |   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 | 418 | \<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 | 419 |   {
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 420 | fix \<mu> :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 421 | fix x y :: 'a | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 422 | assume xy: "x \<in> s" "y \<in> s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 423 | assume mu: "\<mu> \<ge> 0" "\<mu> \<le> 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 424 | 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 | 425 | 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 | 426 |     have "{1 :: nat .. 2} \<inter> - {x. x = 1} = {2}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 427 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 428 |     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 | 429 | by simp | 
| 64267 | 430 |     then have "sum ?u {1 .. 2} = 1"
 | 
| 431 |       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 | 432 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 433 |     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 | 434 | using mu xy by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 435 |     have grarr: "(\<Sum>j \<in> {Suc (Suc 0)..2}. ?u j *\<^sub>R ?x j) = (1 - \<mu>) *\<^sub>R y"
 | 
| 64267 | 436 | using sum_head_Suc[of "Suc (Suc 0)" 2 "\<lambda> j. (1 - \<mu>) *\<^sub>R y"] by auto | 
| 437 | 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 | 438 |     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 | 439 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 440 | 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 | 441 | using s by (auto simp: add.commute) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 442 | } | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 443 | then show "convex s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 444 | unfolding convex_alt by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 445 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 446 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 447 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 448 | lemma convex_explicit: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 449 | fixes s :: "'a::real_vector set" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 450 | shows "convex s \<longleftrightarrow> | 
| 64267 | 451 | (\<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 | 452 | proof safe | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 453 | fix t | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 454 | fix u :: "'a \<Rightarrow> real" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 455 | assume "convex s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 456 | and "finite t" | 
| 64267 | 457 | 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 | 458 | then show "(\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s" | 
| 64267 | 459 | 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 | 460 | next | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 461 | assume *: "\<forall>t. \<forall> u. finite t \<and> t \<subseteq> s \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and> | 
| 64267 | 462 | 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 | 463 | show "convex s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 464 | unfolding convex_alt | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 465 | proof safe | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 466 | fix x y | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 467 | fix \<mu> :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 468 | 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 | 469 | 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 | 470 | proof (cases "x = y") | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 471 | case False | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 472 | then show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 473 |         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 | 474 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 475 | next | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 476 | case True | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 477 | then show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 478 |         using *[rule_format, of "{x, y}" "\<lambda> z. 1"] **
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 479 | 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 | 480 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 481 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 482 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 483 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 484 | lemma convex_finite: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 485 | assumes "finite s" | 
| 64267 | 486 | 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 | 487 | unfolding convex_explicit | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 488 | apply safe | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 489 | 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 | 490 | subgoal for t u | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 491 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 492 | 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 | 493 | by simp | 
| 64267 | 494 | 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" | 
| 495 | 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 | 496 | assume "t \<subseteq> s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 497 | then have "s \<inter> t = t" by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 498 | 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 | 499 | 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 | 500 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 501 | done | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 502 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 503 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 504 | subsection \<open>Functions that are convex on a set\<close> | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 505 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 506 | 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 | 507 | where "convex_on s f \<longleftrightarrow> | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 508 | (\<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 | 509 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 510 | lemma convex_onI [intro?]: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 511 | 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 | 512 | 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 | shows "convex_on A f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 514 | unfolding convex_on_def | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 515 | proof clarify | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 516 | fix x y | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 517 | fix u v :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 518 | 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 | 519 | from A(5) have [simp]: "v = 1 - u" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 520 | by (simp add: algebra_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 521 | 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 | 522 | using assms[of u y x] | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 523 | 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 | 524 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 525 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 526 | lemma convex_on_linorderI [intro?]: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 527 |   fixes A :: "('a::{linorder,real_vector}) set"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 528 | 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 | 529 | 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 | 530 | shows "convex_on A f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 531 | proof | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 532 | fix x y | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 533 | fix t :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 534 | 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 | 535 | 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 | 536 | 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 | 537 | 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 | 538 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 539 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 540 | lemma convex_onD: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 541 | assumes "convex_on A f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 542 | 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 | 543 | 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 | 544 | using assms by (auto simp: convex_on_def) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 545 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 546 | lemma convex_onD_Icc: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 547 |   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 | 548 | 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 | 549 | 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 | 550 | 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 | 551 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 552 | 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 | 553 | unfolding convex_on_def by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 554 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 555 | lemma convex_on_add [intro]: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 556 | assumes "convex_on s f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 557 | and "convex_on s g" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 558 | shows "convex_on s (\<lambda>x. f x + g x)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 559 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 560 |   {
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 561 | fix x y | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 562 | assume "x \<in> s" "y \<in> s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 563 | moreover | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 564 | fix u v :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 565 | assume "0 \<le> u" "0 \<le> v" "u + v = 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 566 | ultimately | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 567 | 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 | 568 | 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 | 569 | 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 | 570 | by (simp add: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 571 | } | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 572 | then show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 573 | unfolding convex_on_def by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 574 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 575 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 576 | lemma convex_on_cmul [intro]: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 577 | fixes c :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 578 | assumes "0 \<le> c" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 579 | and "convex_on s f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 580 | shows "convex_on s (\<lambda>x. c * f x)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 581 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 582 | 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 | 583 | for u c fx v fy :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 584 | by (simp add: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 585 | 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 | 586 | unfolding convex_on_def and * by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 587 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 588 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 589 | lemma convex_lower: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 590 | assumes "convex_on s f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 591 | and "x \<in> s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 592 | and "y \<in> s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 593 | and "0 \<le> u" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 594 | and "0 \<le> v" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 595 | and "u + v = 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 596 | 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 | 597 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 598 | let ?m = "max (f x) (f y)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 599 | 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 | 600 | 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 | 601 | also have "\<dots> = max (f x) (f y)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 602 | using assms(6) by (simp add: distrib_right [symmetric]) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 603 | finally show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 604 | using assms unfolding convex_on_def by fastforce | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 605 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 606 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 607 | lemma convex_on_dist [intro]: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 608 | fixes s :: "'a::real_normed_vector set" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 609 | shows "convex_on s (\<lambda>x. dist a x)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 610 | proof (auto simp: convex_on_def dist_norm) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 611 | fix x y | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 612 | assume "x \<in> s" "y \<in> s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 613 | fix u v :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 614 | assume "0 \<le> u" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 615 | assume "0 \<le> v" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 616 | assume "u + v = 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 617 | have "a = u *\<^sub>R a + v *\<^sub>R a" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 618 | 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 | 619 | 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 | 620 | by (auto simp: algebra_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 621 | 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 | 622 | 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 | 623 | 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 | 624 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 625 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 626 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 627 | subsection \<open>Arithmetic operations on sets preserve convexity\<close> | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 628 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 629 | lemma convex_linear_image: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 630 | assumes "linear f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 631 | and "convex s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 632 | shows "convex (f ` s)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 633 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 634 | interpret f: linear f by fact | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 635 | from \<open>convex s\<close> show "convex (f ` s)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 636 | 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 | 637 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 638 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 639 | lemma convex_linear_vimage: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 640 | assumes "linear f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 641 | and "convex s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 642 | shows "convex (f -` s)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 643 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 644 | interpret f: linear f by fact | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 645 | from \<open>convex s\<close> show "convex (f -` s)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 646 | by (simp add: convex_def f.add f.scaleR) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 647 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 648 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 649 | lemma convex_scaling: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 650 | assumes "convex s" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 651 | shows "convex ((\<lambda>x. c *\<^sub>R x) ` s)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 652 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 653 | have "linear (\<lambda>x. c *\<^sub>R x)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 654 | by (simp add: linearI scaleR_add_right) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 655 | then show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 656 | using \<open>convex s\<close> by (rule convex_linear_image) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 657 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 658 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 659 | 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 | 660 | 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 | 661 | shows "convex ((\<lambda>x. x *\<^sub>R c) ` S)" | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 662 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 663 | have "linear (\<lambda>x. x *\<^sub>R c)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 664 | by (simp add: linearI scaleR_add_left) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 665 | 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 | 666 | 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 | 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_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 | 670 | 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 | 671 | shows "convex ((\<lambda>x. - x) ` S)" | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 672 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 673 | have "linear (\<lambda>x. - x)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 674 | by (simp add: linearI) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 675 | 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 | 676 | 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 | 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_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 | 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 | 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 | 682 |   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 | 683 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 684 | have "linear (\<lambda>(x, y). x + y)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 685 | 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 | 686 | 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 | 687 | 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 | 688 |   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 | 689 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 690 | finally show ?thesis . | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 691 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 692 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 693 | 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 | 694 | 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 | 695 |   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 | 696 | 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 | 697 |   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 | 698 | 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 | 699 | then show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 700 | 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 | 701 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 702 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 703 | 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 | 704 | 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 | 705 | 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 | 706 | 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 | 707 |   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 | 708 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 709 | then show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 710 | 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 | 711 | qed | 
| 
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 | 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 | 714 | 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 | 715 | 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 | 716 | 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 | 717 | 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 | 718 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 719 | then show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 720 | 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 | 721 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 722 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 723 | lemma pos_is_convex: "convex {0 :: real <..}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 724 | unfolding convex_alt | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 725 | proof safe | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 726 | fix y x \<mu> :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 727 | 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 | 728 |   {
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 729 | assume "\<mu> = 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 730 | 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 | 731 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 732 | 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 | 733 | using * by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 734 | } | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 735 | moreover | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 736 |   {
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 737 | assume "\<mu> = 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 738 | 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 | 739 | using * by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 740 | } | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 741 | moreover | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 742 |   {
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 743 | assume "\<mu> \<noteq> 1" "\<mu> \<noteq> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 744 | then have "\<mu> > 0" "(1 - \<mu>) > 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 745 | using * by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 746 | 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 | 747 | using * by (auto simp: add_pos_pos) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 748 | } | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 749 | 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 | 750 | by fastforce | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 751 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 752 | |
| 64267 | 753 | lemma convex_on_sum: | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 754 | fixes a :: "'a \<Rightarrow> real" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 755 | and y :: "'a \<Rightarrow> 'b::real_vector" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 756 | and f :: "'b \<Rightarrow> real" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 757 |   assumes "finite s" "s \<noteq> {}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 758 | and "convex_on C f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 759 | and "convex C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 760 | and "(\<Sum> i \<in> s. a i) = 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 761 | 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 | 762 | 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 | 763 | 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 | 764 | using assms | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 765 | proof (induct s arbitrary: a rule: finite_ne_induct) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 766 | case (singleton i) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 767 | then have ai: "a i = 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 768 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 769 | then show ?case | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 770 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 771 | next | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 772 | case (insert i s) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 773 | then have "convex_on C f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 774 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 775 | from this[unfolded convex_on_def, rule_format] | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 776 | 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 | 777 | 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 | 778 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 779 | show ?case | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 780 | proof (cases "a i = 1") | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 781 | case True | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 782 | then have "(\<Sum> j \<in> s. a j) = 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 783 | using insert by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 784 | then have "\<And>j. j \<in> s \<Longrightarrow> a j = 0" | 
| 64267 | 785 | 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 | 786 | then show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 787 | using insert by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 788 | next | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 789 | case False | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 790 | 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 | 791 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 792 | have fis: "finite (insert i s)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 793 | using insert by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 794 | then have ai1: "a i \<le> 1" | 
| 64267 | 795 | 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 | 796 | then have "a i < 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 797 | using False by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 798 | then have i0: "1 - a i > 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 799 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 800 | let ?a = "\<lambda>j. a j / (1 - a i)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 801 | 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 | 802 | using i0 insert that by fastforce | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 803 | have "(\<Sum> j \<in> insert i s. a j) = 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 804 | using insert by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 805 | then have "(\<Sum> j \<in> s. a j) = 1 - a i" | 
| 64267 | 806 | using sum.insert insert by fastforce | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 807 | 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 | 808 | using i0 by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 809 | then have a1: "(\<Sum> j \<in> s. ?a j) = 1" | 
| 64267 | 810 | unfolding sum_divide_distrib by simp | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 811 | have "convex C" using insert by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 812 | then have asum: "(\<Sum> j \<in> s. ?a j *\<^sub>R y j) \<in> C" | 
| 64267 | 813 | 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 | 814 | 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 | 815 | using a_nonneg a1 insert by blast | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 816 | 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 | 817 | 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 | 818 | by (auto simp only: add.commute) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 819 | 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 | 820 | using i0 by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 821 | 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 | 822 | 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 | 823 | by (auto simp: algebra_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 824 | 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 | 825 | by (auto simp: divide_inverse) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 826 | 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 | 827 | 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 | 828 | by (auto simp: add.commute) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 829 | 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 | 830 | 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 | 831 | 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 | 832 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 833 | also have "\<dots> = (\<Sum> j \<in> s. (1 - a i) * ?a j * f (y j)) + a i * f (y i)" | 
| 64267 | 834 | 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 | 835 | using i0 by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 836 | 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 | 837 | using i0 by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 838 | 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 | 839 | using insert by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 840 | finally show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 841 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 842 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 843 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 844 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 845 | lemma convex_on_alt: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 846 | fixes C :: "'a::real_vector set" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 847 | assumes "convex C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 848 | shows "convex_on C f \<longleftrightarrow> | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 849 | (\<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 | 850 | 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 | 851 | proof safe | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 852 | fix x y | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 853 | fix \<mu> :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 854 | 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 | 855 | from this[unfolded convex_on_def, rule_format] | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 856 | 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 | 857 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 858 | from this [of "\<mu>" "1 - \<mu>", simplified] * | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 859 | 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 | 860 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 861 | next | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 862 | 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 | 863 | 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 | 864 |   {
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 865 | fix x y | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 866 | fix u v :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 867 | 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 | 868 | then have[simp]: "1 - u = v" by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 869 | from *[rule_format, of x y u] | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 870 | 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 | 871 | using ** by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 872 | } | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 873 | then show "convex_on C f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 874 | unfolding convex_on_def by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 875 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 876 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 877 | lemma convex_on_diff: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 878 | fixes f :: "real \<Rightarrow> real" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 879 | assumes f: "convex_on I f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 880 | and I: "x \<in> I" "y \<in> I" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 881 | and t: "x < t" "t < y" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 882 | 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 | 883 | 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 | 884 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 885 | define a where "a \<equiv> (t - y) / (x - y)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 886 | with t have "0 \<le> a" "0 \<le> 1 - a" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 887 | by (auto simp: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 888 | 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 | 889 | by (auto simp: convex_on_def) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 890 | have "a * x + (1 - a) * y = a * (x - y) + y" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 891 | by (simp add: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 892 | also have "\<dots> = t" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 893 | 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 | 894 | 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 | 895 | using cvx by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 896 | also have "\<dots> = a * (f x - f y) + f y" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 897 | by (simp add: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 898 | 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 | 899 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 900 | 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 | 901 | 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 | 902 | 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 | 903 | 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 | 904 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 905 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 906 | lemma pos_convex_function: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 907 | fixes f :: "real \<Rightarrow> real" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 908 | assumes "convex C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 909 | 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 | 910 | shows "convex_on C f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 911 | unfolding convex_on_alt[OF assms(1)] | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 912 | using assms | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 913 | proof safe | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 914 | fix x y \<mu> :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 915 | 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 | 916 | 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 | 917 | then have "1 - \<mu> \<ge> 0" by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 918 | then have xpos: "?x \<in> C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 919 | using * unfolding convex_alt by fastforce | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 920 | 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 | 921 | \<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 | 922 | 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 | 923 | 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 | 924 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 925 | 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 | 926 | by (auto simp: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 927 | 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 | 928 | using convex_on_alt by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 929 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 930 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 931 | lemma atMostAtLeast_subset_convex: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 932 | fixes C :: "real set" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 933 | assumes "convex C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 934 | and "x \<in> C" "y \<in> C" "x < y" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 935 |   shows "{x .. y} \<subseteq> C"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 936 | proof safe | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 937 |   fix z assume z: "z \<in> {x .. y}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 938 | have less: "z \<in> C" if *: "x < z" "z < y" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 939 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 940 | let ?\<mu> = "(y - z) / (y - x)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 941 | have "0 \<le> ?\<mu>" "?\<mu> \<le> 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 942 | using assms * by (auto simp: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 943 | then have comb: "?\<mu> * x + (1 - ?\<mu>) * y \<in> C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 944 | 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 | 945 | by (simp add: algebra_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 946 | 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 | 947 | by (auto simp: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 948 | 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 | 949 | 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 | 950 | also have "\<dots> = z" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 951 | using assms by (auto simp: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 952 | finally show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 953 | using comb by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 954 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 955 | show "z \<in> C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 956 | using z less assms by (auto simp: le_less) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 957 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 958 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 959 | lemma f''_imp_f': | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 960 | fixes f :: "real \<Rightarrow> real" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 961 | assumes "convex C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 962 | 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 | 963 | 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 | 964 | 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 | 965 | and x: "x \<in> C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 966 | and y: "y \<in> C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 967 | shows "f' x * (y - x) \<le> f y - f x" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 968 | using assms | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 969 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 970 | 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 | 971 | 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 | 972 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 973 | from * have ge: "y - x > 0" "y - x \<ge> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 974 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 975 | from * have le: "x - y < 0" "x - y \<le> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 976 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 977 | 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 | 978 | 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 | 979 | 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 | 980 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 981 | then have "z1 \<in> C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 982 | 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 | 983 | by fastforce | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 984 | 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 | 985 | by (simp add: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 986 | 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 | 987 | 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 | 988 | 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 | 989 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 990 | 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 | 991 | 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 | 992 | 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 | 993 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 994 | 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 | 995 | using * z1' by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 996 | also have "\<dots> = (y - z1) * f'' z3" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 997 | using z3 by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 998 | 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 | 999 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1000 | have A': "y - z1 \<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 "z3 \<in> C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1003 | 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 | 1004 | by fastforce | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1005 | then have B': "f'' z3 \<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 "(y - z1) * f'' z3 \<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 | 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 | 1010 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1011 | from mult_right_mono_neg[OF this le(2)] | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1012 | 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 | 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 * (x - y) - (f x - f y) \<le> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1015 | using le by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1016 | 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 | 1017 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1018 | 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 | 1019 | using * z1 by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1020 | also have "\<dots> = (z1 - x) * f'' z2" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1021 | using z2 by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1022 | 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 | 1023 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1024 | have A: "z1 - x \<ge> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1025 | using z1 by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1026 | have "z2 \<in> C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1027 | 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 | 1028 | by fastforce | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1029 | then have B: "f'' z2 \<ge> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1030 | using assms by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1031 | from A B have "(z1 - x) * f'' z2 \<ge> 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1032 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1033 | 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 | 1034 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1035 | from mult_right_mono[OF this ge(2)] | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1036 | 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 | 1037 | by (simp add: algebra_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1038 | 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 | 1039 | using ge by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1040 | 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 | 1041 | using res by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1042 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1043 | show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1044 | proof (cases "x = y") | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1045 | case True | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1046 | with x y show ?thesis by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1047 | next | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1048 | case False | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1049 | with less_imp x y show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1050 | by (auto simp: neq_iff) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1051 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1052 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1053 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1054 | lemma f''_ge0_imp_convex: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1055 | fixes f :: "real \<Rightarrow> real" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1056 | assumes conv: "convex C" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1057 | 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 | 1058 | 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 | 1059 | 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 | 1060 | shows "convex_on C f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1061 | 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 | 1062 | by fastforce | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1063 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1064 | lemma minus_log_convex: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1065 | fixes b :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1066 | assumes "b > 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1067 |   shows "convex_on {0 <..} (\<lambda> x. - log b x)"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1068 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1069 | 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 | 1070 | using DERIV_log by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1071 | 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 | 1072 | by (auto simp: DERIV_minus) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1073 | 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 | 1074 | 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 | 1075 | from this[THEN DERIV_cmult, of _ "- 1 / ln b"] | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1076 | have "\<And>z::real. z > 0 \<Longrightarrow> | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1077 | 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 | 1078 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1079 | then have f''0: "\<And>z::real. z > 0 \<Longrightarrow> | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1080 | 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 | 1081 | unfolding inverse_eq_divide by (auto simp: mult.assoc) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1082 | 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 | 1083 | 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 | 1084 | 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 | 1085 | show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1086 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1087 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1088 | |
| 
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 | subsection \<open>Convexity of real functions\<close> | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1091 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1092 | lemma convex_on_realI: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1093 | assumes "connected A" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1094 | 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 | 1095 | 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 | 1096 | shows "convex_on A f" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1097 | proof (rule convex_on_linorderI) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1098 | fix t x y :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1099 | assume t: "t > 0" "t < 1" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1100 | assume xy: "x \<in> A" "y \<in> A" "x < y" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1101 | define z where "z = (1 - t) * x + t * y" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1102 |   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 | 1103 | using connected_contains_Icc by blast | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1104 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1105 | from xy t have xz: "z > x" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1106 | by (simp add: z_def algebra_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1107 | have "y - z = (1 - 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 from xy t have "\<dots> > 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1110 | by (intro mult_pos_pos) simp_all | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1111 | finally have yz: "z < y" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1112 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1113 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1114 | 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 | 1115 | by (intro MVT2) (auto intro!: assms(2)) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1116 | 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 | 1117 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1118 | 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 | 1119 | by (intro MVT2) (auto intro!: assms(2)) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1120 | 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 | 1121 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1122 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1123 | 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 | 1124 | 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 | 1125 | by auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1126 | with \<xi> \<eta> have "f' \<eta> \<ge> f' \<xi>" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1127 | by (intro assms(3)) auto | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1128 | 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 | 1129 | 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 | 1130 | using xz yz by (simp add: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1131 | also have "z - x = t * (y - x)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1132 | by (simp add: z_def algebra_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1133 | also have "y - z = (1 - t) * (y - x)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1134 | by (simp add: z_def algebra_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1135 | 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 | 1136 | using xy by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1137 | 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 | 1138 | by (simp add: z_def algebra_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1139 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1140 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1141 | lemma convex_on_inverse: | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1142 |   assumes "A \<subseteq> {0<..}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1143 | shows "convex_on A (inverse :: real \<Rightarrow> real)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1144 | 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 | 1145 | fix u v :: real | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1146 |   assume "u \<in> {0<..}" "v \<in> {0<..}" "u \<le> v"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1147 | with assms show "-inverse (u^2) \<le> -inverse (v^2)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1148 | 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 | 1149 | 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 | 1150 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1151 | lemma convex_onD_Icc': | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1152 |   assumes "convex_on {x..y} f" "c \<in> {x..y}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1153 | defines "d \<equiv> y - x" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1154 | 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 | 1155 | proof (cases x y rule: linorder_cases) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1156 | case less | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1157 | then have d: "d > 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1158 | by (simp add: d_def) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1159 | 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 | 1160 | by (simp_all add: d_def divide_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1161 | have "f c = f (x + (c - x) * 1)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1162 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1163 | also from less have "1 = ((y - x) / d)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1164 | by (simp add: d_def) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1165 | 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 | 1166 | by (simp add: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1167 | 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 | 1168 | using assms less by (intro convex_onD_Icc) simp_all | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1169 | 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 | 1170 | by (simp add: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1171 | finally show ?thesis . | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1172 | qed (insert assms(2), simp_all) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1173 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1174 | lemma convex_onD_Icc'': | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1175 |   assumes "convex_on {x..y} f" "c \<in> {x..y}"
 | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1176 | defines "d \<equiv> y - x" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1177 | 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 | 1178 | proof (cases x y rule: linorder_cases) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1179 | case less | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1180 | then have d: "d > 0" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1181 | by (simp add: d_def) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1182 | 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 | 1183 | by (simp_all add: d_def divide_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1184 | have "f c = f (y - (y - c) * 1)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1185 | by simp | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1186 | also from less have "1 = ((y - x) / d)" | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1187 | by (simp add: d_def) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1188 | 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 | 1189 | by (simp add: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1190 | 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 | 1191 | 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 | 1192 | 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 | 1193 | by (simp add: field_simps) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1194 | finally show ?thesis . | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1195 | qed (insert assms(2), simp_all) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1196 | |
| 64267 | 1197 | lemma convex_supp_sum: | 
| 1198 | 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 | 1199 | and "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> u i \<and> (u i = 0 \<or> f i \<in> S)" | 
| 64267 | 1200 | 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 | 1201 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1202 |   have fin: "finite {i \<in> I. u i \<noteq> 0}"
 | 
| 64267 | 1203 | using 1 sum.infinite by (force simp: supp_sum_def support_on_def) | 
| 1204 |   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}"
 | |
| 1205 | 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 | 1206 | show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1207 | apply (simp add: eq) | 
| 64267 | 1208 | apply (rule convex_sum [OF fin \<open>convex S\<close>]) | 
| 1209 | 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 | 1210 | done | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1211 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1212 | |
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1213 | 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 | 1214 | by (metis convex_translation translation_galois) | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1215 | |
| 61694 
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 1216 | 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 | 1217 | 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 | 1218 | 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 | 1219 | 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 | 1220 | |
| 40377 | 1221 | lemma basis_to_basis_subspace_isomorphism: | 
| 1222 |   assumes s: "subspace (S:: ('n::euclidean_space) set)"
 | |
| 49529 | 1223 |     and t: "subspace (T :: ('m::euclidean_space) set)"
 | 
| 1224 | and d: "dim S = dim T" | |
| 53333 | 1225 | and B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "card B = dim S" | 
| 1226 | and C: "C \<subseteq> T" "independent C" "T \<subseteq> span C" "card C = dim T" | |
| 1227 | shows "\<exists>f. linear f \<and> f ` B = C \<and> f ` S = T \<and> inj_on f S" | |
| 49529 | 1228 | proof - | 
| 53333 | 1229 | from B independent_bound have fB: "finite B" | 
| 1230 | by blast | |
| 1231 | from C independent_bound have fC: "finite C" | |
| 1232 | by blast | |
| 40377 | 1233 | from B(4) C(4) card_le_inj[of B C] d obtain f where | 
| 60420 | 1234 | f: "f ` B \<subseteq> C" "inj_on f B" using \<open>finite B\<close> \<open>finite C\<close> by auto | 
| 40377 | 1235 | from linear_independent_extend[OF B(2)] obtain g where | 
| 53333 | 1236 | g: "linear g" "\<forall>x \<in> B. g x = f x" by blast | 
| 40377 | 1237 | from inj_on_iff_eq_card[OF fB, of f] f(2) | 
| 1238 | have "card (f ` B) = card B" by simp | |
| 1239 | with B(4) C(4) have ceq: "card (f ` B) = card C" using d | |
| 1240 | by simp | |
| 1241 | have "g ` B = f ` B" using g(2) | |
| 1242 | by (auto simp add: image_iff) | |
| 1243 | also have "\<dots> = C" using card_subset_eq[OF fC f(1) ceq] . | |
| 1244 | finally have gBC: "g ` B = C" . | |
| 1245 | have gi: "inj_on g B" using f(2) g(2) | |
| 1246 | by (auto simp add: inj_on_def) | |
| 1247 | note g0 = linear_indep_image_lemma[OF g(1) fB, unfolded gBC, OF C(2) gi] | |
| 53333 | 1248 |   {
 | 
| 1249 | fix x y | |
| 49529 | 1250 | assume x: "x \<in> S" and y: "y \<in> S" and gxy: "g x = g y" | 
| 53333 | 1251 | from B(3) x y have x': "x \<in> span B" and y': "y \<in> span B" | 
| 1252 | by blast+ | |
| 1253 | 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 | 1254 | by (simp add: linear_diff[OF g(1)]) | 
| 53333 | 1255 | have th1: "x - y \<in> span B" using x' y' | 
| 63938 | 1256 | by (metis span_diff) | 
| 53333 | 1257 | have "x = y" using g0[OF th1 th0] by simp | 
| 1258 | } | |
| 1259 | then have giS: "inj_on g S" unfolding inj_on_def by blast | |
| 40377 | 1260 | from span_subspace[OF B(1,3) s] | 
| 53333 | 1261 | have "g ` S = span (g ` B)" | 
| 1262 | by (simp add: span_linear_image[OF g(1)]) | |
| 1263 | also have "\<dots> = span C" | |
| 1264 | unfolding gBC .. | |
| 1265 | also have "\<dots> = T" | |
| 1266 | using span_subspace[OF C(1,3) t] . | |
| 40377 | 1267 | finally have gS: "g ` S = T" . | 
| 53333 | 1268 | from g(1) gS giS gBC show ?thesis | 
| 1269 | by blast | |
| 40377 | 1270 | qed | 
| 1271 | ||
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1272 | lemma closure_bounded_linear_image_subset: | 
| 44524 | 1273 | assumes f: "bounded_linear f" | 
| 53333 | 1274 | shows "f ` closure S \<subseteq> closure (f ` S)" | 
| 44524 | 1275 | using linear_continuous_on [OF f] closed_closure closure_subset | 
| 1276 | by (rule image_closure_subset) | |
| 1277 | ||
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1278 | lemma closure_linear_image_subset: | 
| 53339 | 1279 | fixes f :: "'m::euclidean_space \<Rightarrow> 'n::real_normed_vector" | 
| 49529 | 1280 | assumes "linear f" | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1281 | shows "f ` (closure S) \<subseteq> closure (f ` S)" | 
| 44524 | 1282 | using assms unfolding linear_conv_bounded_linear | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1283 | by (rule closure_bounded_linear_image_subset) | 
| 
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: | 
| 
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 S: "closed S" and f: "linear f" "inj f" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1288 | shows "closed (f ` S)" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1289 | proof - | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1290 | obtain g where g: "linear g" "g \<circ> f = id" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1291 | using linear_injective_left_inverse [OF f] by blast | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1292 | then have confg: "continuous_on (range f) g" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1293 | using linear_continuous_on linear_conv_bounded_linear by blast | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1294 | have [simp]: "g ` f ` S = S" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1295 | using g by (simp add: image_comp) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1296 | have cgf: "closed (g ` f ` S)" | 
| 61808 | 1297 | by (simp add: \<open>g \<circ> f = id\<close> S image_comp) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1298 | have [simp]: "(range f \<inter> g -` S) = f ` S" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1299 | using g unfolding o_def id_def image_def by auto metis+ | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1300 | show ?thesis | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1301 | proof (rule closedin_closed_trans [of "range f"]) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1302 | show "closedin (subtopology euclidean (range f)) (f ` S)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1303 | using continuous_closedin_preimage [OF confg cgf] by simp | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1304 | show "closed (range f)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1305 | apply (rule closed_injective_image_subspace) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1306 | using f apply (auto simp: linear_linear linear_injective_0) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1307 | done | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1308 | qed | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1309 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1310 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1311 | lemma closed_injective_linear_image_eq: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1312 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1313 | assumes f: "linear f" "inj f" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1314 | shows "(closed(image f s) \<longleftrightarrow> closed s)" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1315 | by (metis closed_injective_linear_image closure_eq closure_linear_image_subset closure_subset_eq f(1) f(2) inj_image_subset_iff) | 
| 40377 | 1316 | |
| 1317 | lemma closure_injective_linear_image: | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1318 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1319 | 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 | 1320 | apply (rule subset_antisym) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1321 | apply (simp add: closure_linear_image_subset) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1322 | 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 | 1323 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1324 | lemma closure_bounded_linear_image: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1325 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1326 | 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 | 1327 | apply (rule subset_antisym, simp add: closure_linear_image_subset) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1328 | apply (rule closure_minimal, simp add: closure_subset image_mono) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1329 | by (meson bounded_closure closed_closure compact_continuous_image compact_eq_bounded_closed linear_continuous_on linear_conv_bounded_linear) | 
| 40377 | 1330 | |
| 44524 | 1331 | lemma closure_scaleR: | 
| 53339 | 1332 | fixes S :: "'a::real_normed_vector set" | 
| 44524 | 1333 | shows "(op *\<^sub>R c) ` (closure S) = closure ((op *\<^sub>R c) ` S)" | 
| 1334 | proof | |
| 1335 | show "(op *\<^sub>R c) ` (closure S) \<subseteq> closure ((op *\<^sub>R c) ` S)" | |
| 53333 | 1336 | using bounded_linear_scaleR_right | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1337 | by (rule closure_bounded_linear_image_subset) | 
| 44524 | 1338 | show "closure ((op *\<^sub>R c) ` S) \<subseteq> (op *\<^sub>R c) ` (closure S)" | 
| 49529 | 1339 | by (intro closure_minimal image_mono closure_subset closed_scaling closed_closure) | 
| 1340 | qed | |
| 1341 | ||
| 1342 | lemma fst_linear: "linear fst" | |
| 53600 
8fda7ad57466
make 'linear' into a sublocale of 'bounded_linear';
 huffman parents: 
53406diff
changeset | 1343 | unfolding linear_iff by (simp add: algebra_simps) | 
| 49529 | 1344 | |
| 1345 | lemma snd_linear: "linear snd" | |
| 53600 
8fda7ad57466
make 'linear' into a sublocale of 'bounded_linear';
 huffman parents: 
53406diff
changeset | 1346 | unfolding linear_iff by (simp add: algebra_simps) | 
| 49529 | 1347 | |
| 54465 | 1348 | lemma fst_snd_linear: "linear (\<lambda>(x,y). x + y)" | 
| 53600 
8fda7ad57466
make 'linear' into a sublocale of 'bounded_linear';
 huffman parents: 
53406diff
changeset | 1349 | unfolding linear_iff by (simp add: algebra_simps) | 
| 40377 | 1350 | |
| 49529 | 1351 | lemma vector_choose_size: | 
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1352 | assumes "0 \<le> c" | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1353 |   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 | 1354 | proof - | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1355 | obtain a::'a where "a \<noteq> 0" | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1356 | 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 | 1357 | then show ?thesis | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1358 | 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 | 1359 | qed | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1360 | |
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1361 | lemma vector_choose_dist: | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1362 | assumes "0 \<le> c" | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1363 |   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 | 1364 | 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 | 1365 | |
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1366 | lemma sphere_eq_empty [simp]: | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1367 |   fixes a :: "'a::{real_normed_vector, perfect_space}"
 | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1368 |   shows "sphere a r = {} \<longleftrightarrow> r < 0"
 | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1369 | by (auto simp: sphere_def dist_norm) (metis dist_norm le_less_linear vector_choose_dist) | 
| 49529 | 1370 | |
| 64267 | 1371 | lemma sum_delta_notmem: | 
| 49529 | 1372 | assumes "x \<notin> s" | 
| 64267 | 1373 | shows "sum (\<lambda>y. if (y = x) then P x else Q y) s = sum Q s" | 
| 1374 | and "sum (\<lambda>y. if (x = y) then P x else Q y) s = sum Q s" | |
| 1375 | and "sum (\<lambda>y. if (y = x) then P y else Q y) s = sum Q s" | |
| 1376 | and "sum (\<lambda>y. if (x = y) then P y else Q y) s = sum Q s" | |
| 1377 | apply (rule_tac [!] sum.cong) | |
| 53333 | 1378 | using assms | 
| 1379 | apply auto | |
| 49529 | 1380 | done | 
| 33175 | 1381 | |
| 64267 | 1382 | lemma sum_delta'': | 
| 49529 | 1383 | fixes s::"'a::real_vector set" | 
| 1384 | assumes "finite s" | |
| 33175 | 1385 | 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 | 1386 | proof - | 
| 1387 | 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)" | |
| 1388 | by auto | |
| 1389 | show ?thesis | |
| 64267 | 1390 | unfolding * using sum.delta[OF assms, of y "\<lambda>x. f x *\<^sub>R x"] by auto | 
| 33175 | 1391 | qed | 
| 1392 | ||
| 53333 | 1393 | 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 | 1394 | by (fact if_distrib) | 
| 33175 | 1395 | |
| 1396 | lemma dist_triangle_eq: | |
| 44361 
75ec83d45303
remove unnecessary euclidean_space class constraints
 huffman parents: 
44349diff
changeset | 1397 | fixes x y z :: "'a::real_inner" | 
| 53333 | 1398 | shows "dist x z = dist x y + dist y z \<longleftrightarrow> | 
| 1399 | norm (x - y) *\<^sub>R (y - z) = norm (y - z) *\<^sub>R (x - y)" | |
| 49529 | 1400 | proof - | 
| 1401 | 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 | 1402 | show ?thesis unfolding dist_norm norm_triangle_eq[of "x - y" "y - z", unfolded *] | 
| 49529 | 1403 | by (auto simp add:norm_minus_commute) | 
| 1404 | qed | |
| 33175 | 1405 | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 1406 | |
| 60420 | 1407 | subsection \<open>Affine set and affine hull\<close> | 
| 33175 | 1408 | |
| 49529 | 1409 | definition affine :: "'a::real_vector set \<Rightarrow> bool" | 
| 1410 | 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 | 1411 | |
| 1412 | 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 | 1413 | unfolding affine_def by (metis eq_diff_eq') | 
| 33175 | 1414 | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1415 | lemma affine_empty [iff]: "affine {}"
 | 
| 33175 | 1416 | unfolding affine_def by auto | 
| 1417 | ||
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1418 | lemma affine_sing [iff]: "affine {x}"
 | 
| 33175 | 1419 | unfolding affine_alt by (auto simp add: scaleR_left_distrib [symmetric]) | 
| 1420 | ||
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1421 | lemma affine_UNIV [iff]: "affine UNIV" | 
| 33175 | 1422 | unfolding affine_def by auto | 
| 1423 | ||
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 1424 | lemma affine_Inter [intro]: "(\<And>s. s\<in>f \<Longrightarrow> affine s) \<Longrightarrow> affine (\<Inter>f)" | 
| 49531 | 1425 | unfolding affine_def by auto | 
| 33175 | 1426 | |
| 60303 | 1427 | lemma affine_Int[intro]: "affine s \<Longrightarrow> affine t \<Longrightarrow> affine (s \<inter> t)" | 
| 33175 | 1428 | unfolding affine_def by auto | 
| 1429 | ||
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 1430 | 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 | 1431 | apply (clarsimp simp add: affine_def) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 1432 | 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 | 1433 | apply (auto simp: algebra_simps) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 1434 | done | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 1435 | |
| 60303 | 1436 | lemma affine_affine_hull [simp]: "affine(affine hull s)" | 
| 49529 | 1437 | unfolding hull_def | 
| 1438 |   using affine_Inter[of "{t. affine t \<and> s \<subseteq> t}"] by auto
 | |
| 33175 | 1439 | |
| 1440 | lemma affine_hull_eq[simp]: "(affine hull s = s) \<longleftrightarrow> affine s" | |
| 49529 | 1441 | by (metis affine_affine_hull hull_same) | 
| 1442 | ||
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1443 | 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 | 1444 | 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 | 1445 | |
| 33175 | 1446 | |
| 60420 | 1447 | subsubsection \<open>Some explicit formulations (from Lars Schewe)\<close> | 
| 33175 | 1448 | |
| 49529 | 1449 | lemma affine: | 
| 1450 | fixes V::"'a::real_vector set" | |
| 1451 | shows "affine V \<longleftrightarrow> | |
| 64267 | 1452 |     (\<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 | 1453 | unfolding affine_def | 
| 1454 | apply rule | |
| 1455 | apply(rule, rule, rule) | |
| 49531 | 1456 | apply(erule conjE)+ | 
| 49529 | 1457 | defer | 
| 1458 | apply (rule, rule, rule, rule, rule) | |
| 1459 | proof - | |
| 1460 | fix x y u v | |
| 1461 | assume as: "x \<in> V" "y \<in> V" "u + v = (1::real)" | |
| 64267 | 1462 |     "\<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 | 1463 | then show "u *\<^sub>R x + v *\<^sub>R y \<in> V" | 
| 1464 | apply (cases "x = y") | |
| 1465 |     using as(4)[THEN spec[where x="{x,y}"], THEN spec[where x="\<lambda>w. if w = x then u else v"]]
 | |
| 1466 | and as(1-3) | |
| 53333 | 1467 | apply (auto simp add: scaleR_left_distrib[symmetric]) | 
| 1468 | done | |
| 33175 | 1469 | next | 
| 49529 | 1470 | fix s u | 
| 1471 | 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 | 1472 |     "finite s" "s \<noteq> {}" "s \<subseteq> V" "sum u s = (1::real)"
 | 
| 63040 | 1473 | define n where "n = card s" | 
| 33175 | 1474 | have "card s = 0 \<or> card s = 1 \<or> card s = 2 \<or> card s > 2" by auto | 
| 49529 | 1475 | then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V" | 
| 1476 | proof (auto simp only: disjE) | |
| 1477 | assume "card s = 2" | |
| 53333 | 1478 | then have "card s = Suc (Suc 0)" | 
| 1479 | by auto | |
| 1480 |     then obtain a b where "s = {a, b}"
 | |
| 1481 | unfolding card_Suc_eq by auto | |
| 49529 | 1482 | then show ?thesis | 
| 1483 | using as(1)[THEN bspec[where x=a], THEN bspec[where x=b]] using as(4,5) | |
| 64267 | 1484 | by (auto simp add: sum_clauses(2)) | 
| 49529 | 1485 | next | 
| 1486 | assume "card s > 2" | |
| 1487 | then show ?thesis using as and n_def | |
| 1488 | proof (induct n arbitrary: u s) | |
| 1489 | case 0 | |
| 1490 | then show ?case by auto | |
| 1491 | next | |
| 1492 | case (Suc n) | |
| 1493 | fix s :: "'a set" and u :: "'a \<Rightarrow> real" | |
| 1494 | assume IA: | |
| 1495 | "\<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 | 1496 |           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 | 1497 | and as: | 
| 1498 | "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 | 1499 |            "finite s" "s \<noteq> {}" "s \<subseteq> V" "sum u s = 1"
 | 
| 49529 | 1500 | have "\<exists>x\<in>s. u x \<noteq> 1" | 
| 1501 | proof (rule ccontr) | |
| 1502 | assume "\<not> ?thesis" | |
| 64267 | 1503 | then have "sum u s = real_of_nat (card s)" | 
| 1504 | unfolding card_eq_sum by auto | |
| 49529 | 1505 | then show False | 
| 60420 | 1506 | using as(7) and \<open>card s > 2\<close> | 
| 49529 | 1507 | 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 | 1508 | qed | 
| 53339 | 1509 | then obtain x where x:"x \<in> s" "u x \<noteq> 1" by auto | 
| 33175 | 1510 | |
| 49529 | 1511 |       have c: "card (s - {x}) = card s - 1"
 | 
| 53333 | 1512 | apply (rule card_Diff_singleton) | 
| 60420 | 1513 | using \<open>x\<in>s\<close> as(4) | 
| 53333 | 1514 | apply auto | 
| 1515 | done | |
| 49529 | 1516 |       have *: "s = insert x (s - {x})" "finite (s - {x})"
 | 
| 60420 | 1517 | using \<open>x\<in>s\<close> and as(4) by auto | 
| 64267 | 1518 |       have **: "sum u (s - {x}) = 1 - u x"
 | 
| 1519 | using sum_clauses(2)[OF *(2), of u x, unfolded *(1)[symmetric] as(7)] by auto | |
| 1520 |       have ***: "inverse (1 - u x) * sum u (s - {x}) = 1"
 | |
| 60420 | 1521 | unfolding ** using \<open>u x \<noteq> 1\<close> by auto | 
| 49529 | 1522 |       have "(\<Sum>xa\<in>s - {x}. (inverse (1 - u x) * u xa) *\<^sub>R xa) \<in> V"
 | 
| 1523 |       proof (cases "card (s - {x}) > 2")
 | |
| 1524 | case True | |
| 1525 |         then have "s - {x} \<noteq> {}" "card (s - {x}) = n"
 | |
| 1526 | unfolding c and as(1)[symmetric] | |
| 49531 | 1527 | proof (rule_tac ccontr) | 
| 49529 | 1528 |           assume "\<not> s - {x} \<noteq> {}"
 | 
| 49531 | 1529 |           then have "card (s - {x}) = 0" unfolding card_0_eq[OF *(2)] by simp
 | 
| 49529 | 1530 | then show False using True by auto | 
| 1531 | qed auto | |
| 1532 | then show ?thesis | |
| 1533 |           apply (rule_tac IA[of "s - {x}" "\<lambda>y. (inverse (1 - u x) * u y)"])
 | |
| 64267 | 1534 | unfolding sum_distrib_left[symmetric] | 
| 53333 | 1535 | using as and *** and True | 
| 49529 | 1536 | apply auto | 
| 1537 | done | |
| 1538 | next | |
| 1539 | case False | |
| 53333 | 1540 |         then have "card (s - {x}) = Suc (Suc 0)"
 | 
| 1541 | using as(2) and c by auto | |
| 1542 |         then obtain a b where "(s - {x}) = {a, b}" "a\<noteq>b"
 | |
| 1543 | unfolding card_Suc_eq by auto | |
| 1544 | then show ?thesis | |
| 1545 | using as(3)[THEN bspec[where x=a], THEN bspec[where x=b]] | |
| 60420 | 1546 | using *** *(2) and \<open>s \<subseteq> V\<close> | 
| 64267 | 1547 | unfolding sum_distrib_left | 
| 1548 | by (auto simp add: sum_clauses(2)) | |
| 49529 | 1549 | qed | 
| 1550 | then have "u x + (1 - u x) = 1 \<Longrightarrow> | |
| 1551 |           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"
 | |
| 1552 | apply - | |
| 1553 | apply (rule as(3)[rule_format]) | |
| 64267 | 1554 | unfolding Real_Vector_Spaces.scaleR_right.sum | 
| 53333 | 1555 | using x(1) as(6) | 
| 1556 | apply auto | |
| 49529 | 1557 | done | 
| 1558 | then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V" | |
| 64267 | 1559 | unfolding scaleR_scaleR[symmetric] and scaleR_right.sum [symmetric] | 
| 49529 | 1560 | apply (subst *) | 
| 64267 | 1561 | unfolding sum_clauses(2)[OF *(2)] | 
| 60420 | 1562 | using \<open>u x \<noteq> 1\<close> | 
| 53333 | 1563 | apply auto | 
| 49529 | 1564 | done | 
| 1565 | qed | |
| 1566 | next | |
| 1567 | assume "card s = 1" | |
| 53333 | 1568 |     then obtain a where "s={a}"
 | 
| 1569 | by (auto simp add: card_Suc_eq) | |
| 1570 | then show ?thesis | |
| 1571 | using as(4,5) by simp | |
| 60420 | 1572 |   qed (insert \<open>s\<noteq>{}\<close> \<open>finite s\<close>, auto)
 | 
| 33175 | 1573 | qed | 
| 1574 | ||
| 1575 | lemma affine_hull_explicit: | |
| 53333 | 1576 | "affine hull p = | 
| 64267 | 1577 |     {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 | 1578 | apply (rule hull_unique) | 
| 1579 | apply (subst subset_eq) | |
| 1580 | prefer 3 | |
| 1581 | apply rule | |
| 1582 | unfolding mem_Collect_eq | |
| 1583 | apply (erule exE)+ | |
| 1584 | apply (erule conjE)+ | |
| 1585 | prefer 2 | |
| 1586 | apply rule | |
| 1587 | proof - | |
| 1588 | fix x | |
| 1589 | assume "x\<in>p" | |
| 64267 | 1590 |   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 | 1591 |     apply (rule_tac x="{x}" in exI)
 | 
| 1592 | apply (rule_tac x="\<lambda>x. 1" in exI) | |
| 49529 | 1593 | apply auto | 
| 1594 | done | |
| 33175 | 1595 | next | 
| 49529 | 1596 | fix t x s u | 
| 53333 | 1597 |   assume as: "p \<subseteq> t" "affine t" "finite s" "s \<noteq> {}"
 | 
| 64267 | 1598 | "s \<subseteq> p" "sum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x" | 
| 49529 | 1599 | then show "x \<in> t" | 
| 53333 | 1600 | using as(2)[unfolded affine, THEN spec[where x=s], THEN spec[where x=u]] | 
| 1601 | by auto | |
| 33175 | 1602 | next | 
| 64267 | 1603 |   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 | 1604 | unfolding affine_def | 
| 1605 | apply (rule, rule, rule, rule, rule) | |
| 1606 | unfolding mem_Collect_eq | |
| 1607 | proof - | |
| 1608 | fix u v :: real | |
| 1609 | assume uv: "u + v = 1" | |
| 1610 | fix x | |
| 64267 | 1611 |     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 | 1612 | then obtain sx ux where | 
| 64267 | 1613 |       x: "finite sx" "sx \<noteq> {}" "sx \<subseteq> p" "sum ux sx = 1" "(\<Sum>v\<in>sx. ux v *\<^sub>R v) = x"
 | 
| 53333 | 1614 | by auto | 
| 1615 | fix y | |
| 64267 | 1616 |     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 | 1617 | then obtain sy uy where | 
| 64267 | 1618 |       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 | 1619 | have xy: "finite (sx \<union> sy)" | 
| 1620 | using x(1) y(1) by auto | |
| 1621 | have **: "(sx \<union> sy) \<inter> sx = sx" "(sx \<union> sy) \<inter> sy = sy" | |
| 1622 | by auto | |
| 49529 | 1623 |     show "\<exists>s ua. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and>
 | 
| 64267 | 1624 | sum ua s = 1 \<and> (\<Sum>v\<in>s. ua v *\<^sub>R v) = u *\<^sub>R x + v *\<^sub>R y" | 
| 49529 | 1625 | apply (rule_tac x="sx \<union> sy" in exI) | 
| 1626 | 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 | 1627 | unfolding scaleR_left_distrib sum.distrib if_smult scaleR_zero_left | 
| 1628 | ** sum.inter_restrict[OF xy, symmetric] | |
| 1629 | unfolding scaleR_scaleR[symmetric] Real_Vector_Spaces.scaleR_right.sum [symmetric] | |
| 1630 | and sum_distrib_left[symmetric] | |
| 49529 | 1631 | unfolding x y | 
| 53333 | 1632 | using x(1-3) y(1-3) uv | 
| 1633 | apply simp | |
| 49529 | 1634 | done | 
| 1635 | qed | |
| 1636 | qed | |
| 33175 | 1637 | |
| 1638 | lemma affine_hull_finite: | |
| 1639 | assumes "finite s" | |
| 64267 | 1640 |   shows "affine hull s = {y. \<exists>u. sum u s = 1 \<and> sum (\<lambda>v. u v *\<^sub>R v) s = y}"
 | 
| 53333 | 1641 | unfolding affine_hull_explicit and set_eq_iff and mem_Collect_eq | 
| 1642 | apply (rule, rule) | |
| 1643 | apply (erule exE)+ | |
| 1644 | apply (erule conjE)+ | |
| 49529 | 1645 | defer | 
| 1646 | apply (erule exE) | |
| 1647 | apply (erule conjE) | |
| 1648 | proof - | |
| 1649 | fix x u | |
| 64267 | 1650 | assume "sum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x" | 
| 49529 | 1651 | then show "\<exists>sa u. finite sa \<and> | 
| 64267 | 1652 |       \<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 | 1653 | apply (rule_tac x=s in exI, rule_tac x=u in exI) | 
| 53333 | 1654 | using assms | 
| 1655 | apply auto | |
| 49529 | 1656 | done | 
| 33175 | 1657 | next | 
| 49529 | 1658 | fix x t u | 
| 1659 | assume "t \<subseteq> s" | |
| 53333 | 1660 | then have *: "s \<inter> t = t" | 
| 1661 | by auto | |
| 64267 | 1662 |   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"
 | 
| 1663 | then show "\<exists>u. sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x" | |
| 49529 | 1664 | apply (rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI) | 
| 64267 | 1665 | unfolding if_smult scaleR_zero_left and sum.inter_restrict[OF assms, symmetric] and * | 
| 49529 | 1666 | apply auto | 
| 1667 | done | |
| 1668 | qed | |
| 1669 | ||
| 33175 | 1670 | |
| 60420 | 1671 | subsubsection \<open>Stepping theorems and hence small special cases\<close> | 
| 33175 | 1672 | |
| 1673 | lemma affine_hull_empty[simp]: "affine hull {} = {}"
 | |
| 49529 | 1674 | by (rule hull_unique) auto | 
| 33175 | 1675 | |
| 64267 | 1676 | (*could delete: it simply rewrites sum expressions, but it's used twice*) | 
| 33175 | 1677 | lemma affine_hull_finite_step: | 
| 1678 | fixes y :: "'a::real_vector" | |
| 49529 | 1679 | shows | 
| 64267 | 1680 |     "(\<exists>u. sum u {} = w \<and> sum (\<lambda>x. u x *\<^sub>R x) {} = y) \<longleftrightarrow> w = 0 \<and> y = 0" (is ?th1)
 | 
| 53347 | 1681 | and | 
| 49529 | 1682 | "finite s \<Longrightarrow> | 
| 64267 | 1683 | (\<exists>u. sum u (insert a s) = w \<and> sum (\<lambda>x. u x *\<^sub>R x) (insert a s) = y) \<longleftrightarrow> | 
| 1684 | (\<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 | 1685 | proof - | 
| 33175 | 1686 | show ?th1 by simp | 
| 53347 | 1687 | assume fin: "finite s" | 
| 1688 | show "?lhs = ?rhs" | |
| 1689 | proof | |
| 53302 | 1690 | assume ?lhs | 
| 64267 | 1691 | 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 | 1692 | by auto | 
| 53347 | 1693 | show ?rhs | 
| 49529 | 1694 | proof (cases "a \<in> s") | 
| 1695 | case True | |
| 1696 | then have *: "insert a s = s" by auto | |
| 53302 | 1697 | show ?thesis | 
| 1698 | using u[unfolded *] | |
| 1699 | apply(rule_tac x=0 in exI) | |
| 1700 | apply auto | |
| 1701 | done | |
| 33175 | 1702 | next | 
| 49529 | 1703 | case False | 
| 1704 | then show ?thesis | |
| 1705 | apply (rule_tac x="u a" in exI) | |
| 53347 | 1706 | using u and fin | 
| 53302 | 1707 | apply auto | 
| 49529 | 1708 | done | 
| 53302 | 1709 | qed | 
| 53347 | 1710 | next | 
| 53302 | 1711 | assume ?rhs | 
| 64267 | 1712 | 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 | 1713 | by auto | 
| 1714 | 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)" | |
| 1715 | by auto | |
| 53347 | 1716 | show ?lhs | 
| 49529 | 1717 | proof (cases "a \<in> s") | 
| 1718 | case True | |
| 1719 | then show ?thesis | |
| 1720 | apply (rule_tac x="\<lambda>x. (if x=a then v else 0) + u x" in exI) | |
| 64267 | 1721 | unfolding sum_clauses(2)[OF fin] | 
| 53333 | 1722 | apply simp | 
| 64267 | 1723 | unfolding scaleR_left_distrib and sum.distrib | 
| 33175 | 1724 | unfolding vu and * and scaleR_zero_left | 
| 64267 | 1725 | apply (auto simp add: sum.delta[OF fin]) | 
| 49529 | 1726 | done | 
| 33175 | 1727 | next | 
| 49531 | 1728 | case False | 
| 49529 | 1729 | then have **: | 
| 1730 | "\<And>x. x \<in> s \<Longrightarrow> u x = (if x = a then v else u x)" | |
| 1731 | "\<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 | 1732 | from False show ?thesis | 
| 49529 | 1733 | apply (rule_tac x="\<lambda>x. if x=a then v else u x" in exI) | 
| 64267 | 1734 | unfolding sum_clauses(2)[OF fin] and * using vu | 
| 1735 | 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)] | |
| 1736 | using sum.cong [of s _ u "\<lambda>x. if x = a then v else u x", OF _ **(1)] | |
| 49529 | 1737 | apply auto | 
| 1738 | done | |
| 1739 | qed | |
| 53347 | 1740 | qed | 
| 33175 | 1741 | qed | 
| 1742 | ||
| 1743 | lemma affine_hull_2: | |
| 1744 | fixes a b :: "'a::real_vector" | |
| 53302 | 1745 |   shows "affine hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b| u v. (u + v = 1)}"
 | 
| 1746 | (is "?lhs = ?rhs") | |
| 49529 | 1747 | proof - | 
| 1748 | have *: | |
| 49531 | 1749 | "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)" | 
| 49529 | 1750 | "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto | 
| 64267 | 1751 |   have "?lhs = {y. \<exists>u. sum u {a, b} = 1 \<and> (\<Sum>v\<in>{a, b}. u v *\<^sub>R v) = y}"
 | 
| 33175 | 1752 |     using affine_hull_finite[of "{a,b}"] by auto
 | 
| 1753 |   also have "\<dots> = {y. \<exists>v u. u b = 1 - v \<and> u b *\<^sub>R b = y - v *\<^sub>R a}"
 | |
| 49529 | 1754 |     by (simp add: affine_hull_finite_step(2)[of "{b}" a])
 | 
| 33175 | 1755 | also have "\<dots> = ?rhs" unfolding * by auto | 
| 1756 | finally show ?thesis by auto | |
| 1757 | qed | |
| 1758 | ||
| 1759 | lemma affine_hull_3: | |
| 1760 | fixes a b c :: "'a::real_vector" | |
| 53302 | 1761 |   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 | 1762 | proof - | 
| 1763 | have *: | |
| 49531 | 1764 | "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)" | 
| 49529 | 1765 | "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto | 
| 1766 | show ?thesis | |
| 1767 | apply (simp add: affine_hull_finite affine_hull_finite_step) | |
| 1768 | unfolding * | |
| 1769 | apply auto | |
| 53302 | 1770 | apply (rule_tac x=v in exI) | 
| 1771 | apply (rule_tac x=va in exI) | |
| 1772 | apply auto | |
| 1773 | apply (rule_tac x=u in exI) | |
| 1774 | apply force | |
| 49529 | 1775 | done | 
| 33175 | 1776 | qed | 
| 1777 | ||
| 40377 | 1778 | lemma mem_affine: | 
| 53333 | 1779 | assumes "affine S" "x \<in> S" "y \<in> S" "u + v = 1" | 
| 53347 | 1780 | shows "u *\<^sub>R x + v *\<^sub>R y \<in> S" | 
| 40377 | 1781 | using assms affine_def[of S] by auto | 
| 1782 | ||
| 1783 | lemma mem_affine_3: | |
| 53333 | 1784 | assumes "affine S" "x \<in> S" "y \<in> S" "z \<in> S" "u + v + w = 1" | 
| 53347 | 1785 | shows "u *\<^sub>R x + v *\<^sub>R y + w *\<^sub>R z \<in> S" | 
| 49529 | 1786 | proof - | 
| 53347 | 1787 |   have "u *\<^sub>R x + v *\<^sub>R y + w *\<^sub>R z \<in> affine hull {x, y, z}"
 | 
| 49529 | 1788 | using affine_hull_3[of x y z] assms by auto | 
| 1789 | moreover | |
| 53347 | 1790 |   have "affine hull {x, y, z} \<subseteq> affine hull S"
 | 
| 49529 | 1791 |     using hull_mono[of "{x, y, z}" "S"] assms by auto
 | 
| 1792 | moreover | |
| 53347 | 1793 | have "affine hull S = S" | 
| 1794 | using assms affine_hull_eq[of S] by auto | |
| 49531 | 1795 | ultimately show ?thesis by auto | 
| 40377 | 1796 | qed | 
| 1797 | ||
| 1798 | lemma mem_affine_3_minus: | |
| 53333 | 1799 | assumes "affine S" "x \<in> S" "y \<in> S" "z \<in> S" | 
| 1800 | shows "x + v *\<^sub>R (y-z) \<in> S" | |
| 1801 | using mem_affine_3[of S x y z 1 v "-v"] assms | |
| 1802 | by (simp add: algebra_simps) | |
| 40377 | 1803 | |
| 60307 
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
 paulson <lp15@cam.ac.uk> parents: 
60303diff
changeset | 1804 | 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 | 1805 | "\<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 | 1806 | 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 | 1807 | |
| 40377 | 1808 | |
| 60420 | 1809 | subsubsection \<open>Some relations between affine hull and subspaces\<close> | 
| 33175 | 1810 | |
| 1811 | lemma affine_hull_insert_subset_span: | |
| 49529 | 1812 |   "affine hull (insert a s) \<subseteq> {a + v| v . v \<in> span {x - a | x . x \<in> s}}"
 | 
| 1813 | unfolding subset_eq Ball_def | |
| 1814 | unfolding affine_hull_explicit span_explicit mem_Collect_eq | |
| 50804 | 1815 | apply (rule, rule) | 
| 1816 | apply (erule exE)+ | |
| 1817 | apply (erule conjE)+ | |
| 49529 | 1818 | proof - | 
| 1819 | fix x t u | |
| 64267 | 1820 |   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 | 1821 |   have "(\<lambda>x. x - a) ` (t - {a}) \<subseteq> {x - a |x. x \<in> s}"
 | 
| 1822 | using as(3) by auto | |
| 49529 | 1823 |   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)"
 | 
| 1824 | apply (rule_tac x="x - a" in exI) | |
| 33175 | 1825 | apply (rule conjI, simp) | 
| 49529 | 1826 |     apply (rule_tac x="(\<lambda>x. x - a) ` (t - {a})" in exI)
 | 
| 1827 | apply (rule_tac x="\<lambda>x. u (x + a)" in exI) | |
| 33175 | 1828 | apply (rule conjI) using as(1) apply simp | 
| 1829 | apply (erule conjI) | |
| 1830 | using as(1) | |
| 64267 | 1831 | apply (simp add: sum.reindex[unfolded inj_on_def] scaleR_right_diff_distrib | 
| 1832 | sum_subtractf scaleR_left.sum[symmetric] sum_diff1 scaleR_left_diff_distrib) | |
| 49529 | 1833 | unfolding as | 
| 1834 | apply simp | |
| 1835 | done | |
| 1836 | qed | |
| 33175 | 1837 | |
| 1838 | lemma affine_hull_insert_span: | |
| 1839 | assumes "a \<notin> s" | |
| 49529 | 1840 |   shows "affine hull (insert a s) = {a + v | v . v \<in> span {x - a | x.  x \<in> s}}"
 | 
| 1841 | apply (rule, rule affine_hull_insert_subset_span) | |
| 1842 | unfolding subset_eq Ball_def | |
| 1843 | unfolding affine_hull_explicit and mem_Collect_eq | |
| 1844 | proof (rule, rule, erule exE, erule conjE) | |
| 49531 | 1845 | fix y v | 
| 49529 | 1846 |   assume "y = a + v" "v \<in> span {x - a |x. x \<in> s}"
 | 
| 53339 | 1847 |   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 | 1848 | unfolding span_explicit by auto | 
| 63040 | 1849 | define f where "f = (\<lambda>x. x + a) ` t" | 
| 53333 | 1850 | have f: "finite f" "f \<subseteq> s" "(\<Sum>v\<in>f. u (v - a) *\<^sub>R (v - a)) = y - a" | 
| 64267 | 1851 | unfolding f_def using obt by (auto simp add: sum.reindex[unfolded inj_on_def]) | 
| 53333 | 1852 |   have *: "f \<inter> {a} = {}" "f \<inter> - {a} = f"
 | 
| 1853 | using f(2) assms by auto | |
| 64267 | 1854 |   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 | 1855 | apply (rule_tac x = "insert a f" in exI) | 
| 64267 | 1856 | 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 | 1857 | using assms and f | 
| 64267 | 1858 | unfolding sum_clauses(2)[OF f(1)] and if_smult | 
| 1859 | unfolding sum.If_cases[OF f(1), of "\<lambda>x. x = a"] | |
| 1860 | apply (auto simp add: sum_subtractf scaleR_left.sum algebra_simps *) | |
| 49529 | 1861 | done | 
| 1862 | qed | |
| 33175 | 1863 | |
| 1864 | lemma affine_hull_span: | |
| 1865 | assumes "a \<in> s" | |
| 1866 |   shows "affine hull s = {a + v | v. v \<in> span {x - a | x. x \<in> s - {a}}}"
 | |
| 1867 |   using affine_hull_insert_span[of a "s - {a}", unfolded insert_Diff[OF assms]] by auto
 | |
| 1868 | ||
| 49529 | 1869 | |
| 60420 | 1870 | subsubsection \<open>Parallel affine sets\<close> | 
| 40377 | 1871 | |
| 53347 | 1872 | definition affine_parallel :: "'a::real_vector set \<Rightarrow> 'a::real_vector set \<Rightarrow> bool" | 
| 53339 | 1873 | where "affine_parallel S T \<longleftrightarrow> (\<exists>a. T = (\<lambda>x. a + x) ` S)" | 
| 40377 | 1874 | |
| 1875 | lemma affine_parallel_expl_aux: | |
| 49529 | 1876 | fixes S T :: "'a::real_vector set" | 
| 53339 | 1877 | assumes "\<forall>x. x \<in> S \<longleftrightarrow> a + x \<in> T" | 
| 1878 | shows "T = (\<lambda>x. a + x) ` S" | |
| 49529 | 1879 | proof - | 
| 53302 | 1880 |   {
 | 
| 1881 | fix x | |
| 53339 | 1882 | assume "x \<in> T" | 
| 1883 | then have "( - a) + x \<in> S" | |
| 1884 | using assms by auto | |
| 1885 | then have "x \<in> ((\<lambda>x. a + x) ` S)" | |
| 53333 | 1886 | using imageI[of "-a+x" S "(\<lambda>x. a+x)"] by auto | 
| 53302 | 1887 | } | 
| 53339 | 1888 | moreover have "T \<ge> (\<lambda>x. a + x) ` S" | 
| 53333 | 1889 | using assms by auto | 
| 49529 | 1890 | ultimately show ?thesis by auto | 
| 1891 | qed | |
| 1892 | ||
| 53339 | 1893 | lemma affine_parallel_expl: "affine_parallel S T \<longleftrightarrow> (\<exists>a. \<forall>x. x \<in> S \<longleftrightarrow> a + x \<in> T)" | 
| 49529 | 1894 | unfolding affine_parallel_def | 
| 1895 | using affine_parallel_expl_aux[of S _ T] by auto | |
| 1896 | ||
| 1897 | lemma affine_parallel_reflex: "affine_parallel S S" | |
| 53302 | 1898 | unfolding affine_parallel_def | 
| 1899 | apply (rule exI[of _ "0"]) | |
| 1900 | apply auto | |
| 1901 | done | |
| 40377 | 1902 | |
| 1903 | lemma affine_parallel_commut: | |
| 49529 | 1904 | assumes "affine_parallel A B" | 
| 1905 | shows "affine_parallel B A" | |
| 1906 | proof - | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53676diff
changeset | 1907 | from assms obtain a where B: "B = (\<lambda>x. a + x) ` A" | 
| 49529 | 1908 | unfolding affine_parallel_def by auto | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53676diff
changeset | 1909 | 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 | 1910 | from B show ?thesis | 
| 53333 | 1911 | using translation_galois [of B a A] | 
| 1912 | unfolding affine_parallel_def by auto | |
| 40377 | 1913 | qed | 
| 1914 | ||
| 1915 | lemma affine_parallel_assoc: | |
| 53339 | 1916 | assumes "affine_parallel A B" | 
| 1917 | and "affine_parallel B C" | |
| 49531 | 1918 | shows "affine_parallel A C" | 
| 49529 | 1919 | proof - | 
| 53333 | 1920 | from assms obtain ab where "B = (\<lambda>x. ab + x) ` A" | 
| 49531 | 1921 | unfolding affine_parallel_def by auto | 
| 1922 | moreover | |
| 53333 | 1923 | from assms obtain bc where "C = (\<lambda>x. bc + x) ` B" | 
| 49529 | 1924 | unfolding affine_parallel_def by auto | 
| 1925 | ultimately show ?thesis | |
| 1926 | using translation_assoc[of bc ab A] unfolding affine_parallel_def by auto | |
| 40377 | 1927 | qed | 
| 1928 | ||
| 1929 | lemma affine_translation_aux: | |
| 1930 | fixes a :: "'a::real_vector" | |
| 53333 | 1931 | assumes "affine ((\<lambda>x. a + x) ` S)" | 
| 1932 | shows "affine S" | |
| 53302 | 1933 | proof - | 
| 1934 |   {
 | |
| 1935 | fix x y u v | |
| 53333 | 1936 | assume xy: "x \<in> S" "y \<in> S" "(u :: real) + v = 1" | 
| 1937 | then have "(a + x) \<in> ((\<lambda>x. a + x) ` S)" "(a + y) \<in> ((\<lambda>x. a + x) ` S)" | |
| 1938 | by auto | |
| 53339 | 1939 | then have h1: "u *\<^sub>R (a + x) + v *\<^sub>R (a + y) \<in> (\<lambda>x. a + x) ` S" | 
| 49529 | 1940 | using xy assms unfolding affine_def by auto | 
| 53339 | 1941 | 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 | 1942 | by (simp add: algebra_simps) | 
| 53339 | 1943 | also have "\<dots> = a + (u *\<^sub>R x + v *\<^sub>R y)" | 
| 60420 | 1944 | using \<open>u + v = 1\<close> by auto | 
| 53339 | 1945 | ultimately have "a + (u *\<^sub>R x + v *\<^sub>R y) \<in> (\<lambda>x. a + x) ` S" | 
| 53333 | 1946 | using h1 by auto | 
| 49529 | 1947 | then have "u *\<^sub>R x + v *\<^sub>R y : S" by auto | 
| 1948 | } | |
| 1949 | then show ?thesis unfolding affine_def by auto | |
| 40377 | 1950 | qed | 
| 1951 | ||
| 1952 | lemma affine_translation: | |
| 1953 | fixes a :: "'a::real_vector" | |
| 53339 | 1954 | shows "affine S \<longleftrightarrow> affine ((\<lambda>x. a + x) ` S)" | 
| 49529 | 1955 | proof - | 
| 53339 | 1956 | have "affine S \<Longrightarrow> affine ((\<lambda>x. a + x) ` S)" | 
| 1957 | using affine_translation_aux[of "-a" "((\<lambda>x. a + x) ` S)"] | |
| 49529 | 1958 | using translation_assoc[of "-a" a S] by auto | 
| 1959 | then show ?thesis using affine_translation_aux by auto | |
| 40377 | 1960 | qed | 
| 1961 | ||
| 1962 | lemma parallel_is_affine: | |
| 49529 | 1963 | fixes S T :: "'a::real_vector set" | 
| 1964 | assumes "affine S" "affine_parallel S T" | |
| 1965 | shows "affine T" | |
| 1966 | proof - | |
| 53339 | 1967 | from assms obtain a where "T = (\<lambda>x. a + x) ` S" | 
| 49531 | 1968 | unfolding affine_parallel_def by auto | 
| 53339 | 1969 | then show ?thesis | 
| 1970 | using affine_translation assms by auto | |
| 40377 | 1971 | qed | 
| 1972 | ||
| 44361 
75ec83d45303
remove unnecessary euclidean_space class constraints
 huffman parents: 
44349diff
changeset | 1973 | lemma subspace_imp_affine: "subspace s \<Longrightarrow> affine s" | 
| 40377 | 1974 | unfolding subspace_def affine_def by auto | 
| 1975 | ||
| 49529 | 1976 | |
| 60420 | 1977 | subsubsection \<open>Subspace parallel to an affine set\<close> | 
| 40377 | 1978 | |
| 53339 | 1979 | lemma subspace_affine: "subspace S \<longleftrightarrow> affine S \<and> 0 \<in> S" | 
| 49529 | 1980 | proof - | 
| 53333 | 1981 | have h0: "subspace S \<Longrightarrow> affine S \<and> 0 \<in> S" | 
| 49529 | 1982 | using subspace_imp_affine[of S] subspace_0 by auto | 
| 53302 | 1983 |   {
 | 
| 53333 | 1984 | assume assm: "affine S \<and> 0 \<in> S" | 
| 53302 | 1985 |     {
 | 
| 1986 | fix c :: real | |
| 54465 | 1987 | fix x | 
| 1988 | assume x: "x \<in> S" | |
| 49529 | 1989 | have "c *\<^sub>R x = (1-c) *\<^sub>R 0 + c *\<^sub>R x" by auto | 
| 1990 | moreover | |
| 53339 | 1991 | have "(1 - c) *\<^sub>R 0 + c *\<^sub>R x \<in> S" | 
| 54465 | 1992 | using affine_alt[of S] assm x by auto | 
| 53333 | 1993 | ultimately have "c *\<^sub>R x \<in> S" by auto | 
| 49529 | 1994 | } | 
| 53333 | 1995 | then have h1: "\<forall>c. \<forall>x \<in> S. c *\<^sub>R x \<in> S" by auto | 
| 49529 | 1996 | |
| 53302 | 1997 |     {
 | 
| 1998 | fix x y | |
| 54465 | 1999 | assume xy: "x \<in> S" "y \<in> S" | 
| 63040 | 2000 | define u where "u = (1 :: real)/2" | 
| 53302 | 2001 | have "(1/2) *\<^sub>R (x+y) = (1/2) *\<^sub>R (x+y)" | 
| 2002 | by auto | |
| 49529 | 2003 | moreover | 
| 53302 | 2004 | have "(1/2) *\<^sub>R (x+y)=(1/2) *\<^sub>R x + (1-(1/2)) *\<^sub>R y" | 
| 2005 | by (simp add: algebra_simps) | |
| 49529 | 2006 | moreover | 
| 54465 | 2007 | have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> S" | 
| 2008 | using affine_alt[of S] assm xy by auto | |
| 49529 | 2009 | ultimately | 
| 53333 | 2010 | have "(1/2) *\<^sub>R (x+y) \<in> S" | 
| 53302 | 2011 | using u_def by auto | 
| 49529 | 2012 | moreover | 
| 54465 | 2013 | have "x + y = 2 *\<^sub>R ((1/2) *\<^sub>R (x+y))" | 
| 53302 | 2014 | by auto | 
| 49529 | 2015 | ultimately | 
| 54465 | 2016 | have "x + y \<in> S" | 
| 53302 | 2017 | using h1[rule_format, of "(1/2) *\<^sub>R (x+y)" "2"] by auto | 
| 49529 | 2018 | } | 
| 53302 | 2019 | then have "\<forall>x \<in> S. \<forall>y \<in> S. x + y \<in> S" | 
| 2020 | by auto | |
| 2021 | then have "subspace S" | |
| 2022 | using h1 assm unfolding subspace_def by auto | |
| 49529 | 2023 | } | 
| 2024 | then show ?thesis using h0 by metis | |
| 40377 | 2025 | qed | 
| 2026 | ||
| 2027 | lemma affine_diffs_subspace: | |
| 53333 | 2028 | assumes "affine S" "a \<in> S" | 
| 53302 | 2029 | shows "subspace ((\<lambda>x. (-a)+x) ` S)" | 
| 49529 | 2030 | proof - | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53676diff
changeset | 2031 | have [simp]: "(\<lambda>x. x - a) = plus (- a)" by (simp add: fun_eq_iff) | 
| 53302 | 2032 | have "affine ((\<lambda>x. (-a)+x) ` S)" | 
| 49531 | 2033 | using affine_translation assms by auto | 
| 53302 | 2034 | moreover have "0 : ((\<lambda>x. (-a)+x) ` S)" | 
| 53333 | 2035 | using assms exI[of "(\<lambda>x. x\<in>S \<and> -a+x = 0)" a] by auto | 
| 49531 | 2036 | ultimately show ?thesis using subspace_affine by auto | 
| 40377 | 2037 | qed | 
| 2038 | ||
| 2039 | lemma parallel_subspace_explicit: | |
| 54465 | 2040 | assumes "affine S" | 
| 2041 | and "a \<in> S" | |
| 2042 |   assumes "L \<equiv> {y. \<exists>x \<in> S. (-a) + x = y}"
 | |
| 2043 | shows "subspace L \<and> affine_parallel S L" | |
| 49529 | 2044 | proof - | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53676diff
changeset | 2045 | from assms have "L = plus (- a) ` S" by auto | 
| 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53676diff
changeset | 2046 | then have par: "affine_parallel S L" | 
| 54465 | 2047 | unfolding affine_parallel_def .. | 
| 49531 | 2048 | then have "affine L" using assms parallel_is_affine by auto | 
| 53302 | 2049 | moreover have "0 \<in> L" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53676diff
changeset | 2050 | using assms by auto | 
| 53302 | 2051 | ultimately show ?thesis | 
| 2052 | using subspace_affine par by auto | |
| 40377 | 2053 | qed | 
| 2054 | ||
| 2055 | lemma parallel_subspace_aux: | |
| 53302 | 2056 | assumes "subspace A" | 
| 2057 | and "subspace B" | |
| 2058 | and "affine_parallel A B" | |
| 2059 | shows "A \<supseteq> B" | |
| 49529 | 2060 | proof - | 
| 54465 | 2061 | from assms obtain a where a: "\<forall>x. x \<in> A \<longleftrightarrow> a + x \<in> B" | 
| 49529 | 2062 | using affine_parallel_expl[of A B] by auto | 
| 53302 | 2063 | then have "-a \<in> A" | 
| 2064 | using assms subspace_0[of B] by auto | |
| 2065 | then have "a \<in> A" | |
| 2066 | using assms subspace_neg[of A "-a"] by auto | |
| 2067 | then show ?thesis | |
| 54465 | 2068 | using assms a unfolding subspace_def by auto | 
| 40377 | 2069 | qed | 
| 2070 | ||
| 2071 | lemma parallel_subspace: | |
| 53302 | 2072 | assumes "subspace A" | 
| 2073 | and "subspace B" | |
| 2074 | and "affine_parallel A B" | |
| 49529 | 2075 | shows "A = B" | 
| 2076 | proof | |
| 53302 | 2077 | show "A \<supseteq> B" | 
| 49529 | 2078 | using assms parallel_subspace_aux by auto | 
| 53302 | 2079 | show "A \<subseteq> B" | 
| 49529 | 2080 | using assms parallel_subspace_aux[of B A] affine_parallel_commut by auto | 
| 40377 | 2081 | qed | 
| 2082 | ||
| 2083 | lemma affine_parallel_subspace: | |
| 53302 | 2084 |   assumes "affine S" "S \<noteq> {}"
 | 
| 53339 | 2085 | shows "\<exists>!L. subspace L \<and> affine_parallel S L" | 
| 49529 | 2086 | proof - | 
| 53339 | 2087 | have ex: "\<exists>L. subspace L \<and> affine_parallel S L" | 
| 49531 | 2088 | using assms parallel_subspace_explicit by auto | 
| 53302 | 2089 |   {
 | 
| 2090 | fix L1 L2 | |
| 53339 | 2091 | assume ass: "subspace L1 \<and> affine_parallel S L1" "subspace L2 \<and> affine_parallel S L2" | 
| 49529 | 2092 | then have "affine_parallel L1 L2" | 
| 2093 | using affine_parallel_commut[of S L1] affine_parallel_assoc[of L1 S L2] by auto | |
| 2094 | then have "L1 = L2" | |
| 2095 | using ass parallel_subspace by auto | |
| 2096 | } | |
| 2097 | then show ?thesis using ex by auto | |
| 2098 | qed | |
| 2099 | ||
| 40377 | 2100 | |
| 60420 | 2101 | subsection \<open>Cones\<close> | 
| 33175 | 2102 | |
| 49529 | 2103 | definition cone :: "'a::real_vector set \<Rightarrow> bool" | 
| 53339 | 2104 | where "cone s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>c\<ge>0. c *\<^sub>R x \<in> s)" | 
| 33175 | 2105 | |
| 2106 | lemma cone_empty[intro, simp]: "cone {}"
 | |
| 2107 | unfolding cone_def by auto | |
| 2108 | ||
| 2109 | lemma cone_univ[intro, simp]: "cone UNIV" | |
| 2110 | unfolding cone_def by auto | |
| 2111 | ||
| 53339 | 2112 | lemma cone_Inter[intro]: "\<forall>s\<in>f. cone s \<Longrightarrow> cone (\<Inter>f)" | 
| 33175 | 2113 | unfolding cone_def by auto | 
| 2114 | ||
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 2115 | 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 | 2116 | 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 | 2117 | |
| 49529 | 2118 | |
| 60420 | 2119 | subsubsection \<open>Conic hull\<close> | 
| 33175 | 2120 | |
| 2121 | lemma cone_cone_hull: "cone (cone hull s)" | |
| 44170 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 huffman parents: 
44142diff
changeset | 2122 | unfolding hull_def by auto | 
| 33175 | 2123 | |
| 53302 | 2124 | lemma cone_hull_eq: "cone hull s = s \<longleftrightarrow> cone s" | 
| 49529 | 2125 | apply (rule hull_eq) | 
| 53302 | 2126 | using cone_Inter | 
| 2127 | unfolding subset_eq | |
| 2128 | apply auto | |
| 49529 | 2129 | done | 
| 33175 | 2130 | |
| 40377 | 2131 | lemma mem_cone: | 
| 53302 | 2132 | assumes "cone S" "x \<in> S" "c \<ge> 0" | 
| 40377 | 2133 | shows "c *\<^sub>R x : S" | 
| 2134 | using assms cone_def[of S] by auto | |
| 2135 | ||
| 2136 | lemma cone_contains_0: | |
| 49529 | 2137 | assumes "cone S" | 
| 53302 | 2138 |   shows "S \<noteq> {} \<longleftrightarrow> 0 \<in> S"
 | 
| 49529 | 2139 | proof - | 
| 53302 | 2140 |   {
 | 
| 2141 |     assume "S \<noteq> {}"
 | |
| 2142 | then obtain a where "a \<in> S" by auto | |
| 2143 | then have "0 \<in> S" | |
| 2144 | using assms mem_cone[of S a 0] by auto | |
| 49529 | 2145 | } | 
| 2146 | then show ?thesis by auto | |
| 40377 | 2147 | qed | 
| 2148 | ||
| 44361 
75ec83d45303
remove unnecessary euclidean_space class constraints
 huffman parents: 
44349diff
changeset | 2149 | lemma cone_0: "cone {0}"
 | 
| 49529 | 2150 | unfolding cone_def by auto | 
| 40377 | 2151 | |
| 61952 | 2152 | lemma cone_Union[intro]: "(\<forall>s\<in>f. cone s) \<longrightarrow> cone (\<Union>f)" | 
| 40377 | 2153 | unfolding cone_def by blast | 
| 2154 | ||
| 2155 | lemma cone_iff: | |
| 53347 | 2156 |   assumes "S \<noteq> {}"
 | 
| 2157 | shows "cone S \<longleftrightarrow> 0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (op *\<^sub>R c) ` S = S)" | |
| 49529 | 2158 | proof - | 
| 53302 | 2159 |   {
 | 
| 2160 | assume "cone S" | |
| 2161 |     {
 | |
| 53347 | 2162 | fix c :: real | 
| 2163 | assume "c > 0" | |
| 53302 | 2164 |       {
 | 
| 2165 | fix x | |
| 53347 | 2166 | assume "x \<in> S" | 
| 2167 | then have "x \<in> (op *\<^sub>R c) ` S" | |
| 49529 | 2168 | unfolding image_def | 
| 60420 | 2169 | using \<open>cone S\<close> \<open>c>0\<close> mem_cone[of S x "1/c"] | 
| 54465 | 2170 | exI[of "(\<lambda>t. t \<in> S \<and> x = c *\<^sub>R t)" "(1 / c) *\<^sub>R x"] | 
| 53347 | 2171 | by auto | 
| 49529 | 2172 | } | 
| 2173 | moreover | |
| 53302 | 2174 |       {
 | 
| 2175 | fix x | |
| 53347 | 2176 | assume "x \<in> (op *\<^sub>R c) ` S" | 
| 2177 | then have "x \<in> S" | |
| 60420 | 2178 | using \<open>cone S\<close> \<open>c > 0\<close> | 
| 2179 | unfolding cone_def image_def \<open>c > 0\<close> by auto | |
| 49529 | 2180 | } | 
| 53302 | 2181 | ultimately have "(op *\<^sub>R c) ` S = S" by auto | 
| 40377 | 2182 | } | 
| 53339 | 2183 | then have "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (op *\<^sub>R c) ` S = S)" | 
| 60420 | 2184 | using \<open>cone S\<close> cone_contains_0[of S] assms by auto | 
| 49529 | 2185 | } | 
| 2186 | moreover | |
| 53302 | 2187 |   {
 | 
| 53339 | 2188 | assume a: "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (op *\<^sub>R c) ` S = S)" | 
| 53302 | 2189 |     {
 | 
| 2190 | fix x | |
| 2191 | assume "x \<in> S" | |
| 53347 | 2192 | fix c1 :: real | 
| 2193 | assume "c1 \<ge> 0" | |
| 2194 | then have "c1 = 0 \<or> c1 > 0" by auto | |
| 60420 | 2195 | then have "c1 *\<^sub>R x \<in> S" using a \<open>x \<in> S\<close> by auto | 
| 49529 | 2196 | } | 
| 2197 | then have "cone S" unfolding cone_def by auto | |
| 40377 | 2198 | } | 
| 49529 | 2199 | ultimately show ?thesis by blast | 
| 2200 | qed | |
| 2201 | ||
| 2202 | lemma cone_hull_empty: "cone hull {} = {}"
 | |
| 2203 | by (metis cone_empty cone_hull_eq) | |
| 2204 | ||
| 53302 | 2205 | lemma cone_hull_empty_iff: "S = {} \<longleftrightarrow> cone hull S = {}"
 | 
| 49529 | 2206 | by (metis bot_least cone_hull_empty hull_subset xtrans(5)) | 
| 2207 | ||
| 53302 | 2208 | lemma cone_hull_contains_0: "S \<noteq> {} \<longleftrightarrow> 0 \<in> cone hull S"
 | 
| 49529 | 2209 | using cone_cone_hull[of S] cone_contains_0[of "cone hull S"] cone_hull_empty_iff[of S] | 
| 2210 | by auto | |
| 40377 | 2211 | |
| 2212 | lemma mem_cone_hull: | |
| 53347 | 2213 | assumes "x \<in> S" "c \<ge> 0" | 
| 53302 | 2214 | shows "c *\<^sub>R x \<in> cone hull S" | 
| 49529 | 2215 | by (metis assms cone_cone_hull hull_inc mem_cone) | 
| 2216 | ||
| 53339 | 2217 | lemma cone_hull_expl: "cone hull S = {c *\<^sub>R x | c x. c \<ge> 0 \<and> x \<in> S}"
 | 
| 2218 | (is "?lhs = ?rhs") | |
| 49529 | 2219 | proof - | 
| 53302 | 2220 |   {
 | 
| 2221 | fix x | |
| 2222 | assume "x \<in> ?rhs" | |
| 54465 | 2223 | then obtain cx :: real and xx where x: "x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S" | 
| 49529 | 2224 | by auto | 
| 53347 | 2225 | fix c :: real | 
| 2226 | assume c: "c \<ge> 0" | |
| 53339 | 2227 | then have "c *\<^sub>R x = (c * cx) *\<^sub>R xx" | 
| 54465 | 2228 | using x by (simp add: algebra_simps) | 
| 49529 | 2229 | moreover | 
| 56536 | 2230 | have "c * cx \<ge> 0" using c x by auto | 
| 49529 | 2231 | ultimately | 
| 54465 | 2232 | have "c *\<^sub>R x \<in> ?rhs" using x by auto | 
| 53302 | 2233 | } | 
| 53347 | 2234 | then have "cone ?rhs" | 
| 2235 | unfolding cone_def by auto | |
| 2236 | then have "?rhs \<in> Collect cone" | |
| 2237 | unfolding mem_Collect_eq by auto | |
| 53302 | 2238 |   {
 | 
| 2239 | fix x | |
| 2240 | assume "x \<in> S" | |
| 2241 | then have "1 *\<^sub>R x \<in> ?rhs" | |
| 49531 | 2242 | apply auto | 
| 53347 | 2243 | apply (rule_tac x = 1 in exI) | 
| 49529 | 2244 | apply auto | 
| 2245 | done | |
| 53302 | 2246 | then have "x \<in> ?rhs" by auto | 
| 53347 | 2247 | } | 
| 2248 | then have "S \<subseteq> ?rhs" by auto | |
| 53302 | 2249 | then have "?lhs \<subseteq> ?rhs" | 
| 60420 | 2250 | using \<open>?rhs \<in> Collect cone\<close> hull_minimal[of S "?rhs" "cone"] by auto | 
| 49529 | 2251 | moreover | 
| 53302 | 2252 |   {
 | 
| 2253 | fix x | |
| 2254 | assume "x \<in> ?rhs" | |
| 54465 | 2255 | then obtain cx :: real and xx where x: "x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S" | 
| 53339 | 2256 | by auto | 
| 2257 | then have "xx \<in> cone hull S" | |
| 2258 | using hull_subset[of S] by auto | |
| 53302 | 2259 | then have "x \<in> ?lhs" | 
| 54465 | 2260 | using x cone_cone_hull[of S] cone_def[of "cone hull S"] by auto | 
| 49529 | 2261 | } | 
| 2262 | ultimately show ?thesis by auto | |
| 40377 | 2263 | qed | 
| 2264 | ||
| 2265 | lemma cone_closure: | |
| 53347 | 2266 | fixes S :: "'a::real_normed_vector set" | 
| 49529 | 2267 | assumes "cone S" | 
| 2268 | shows "cone (closure S)" | |
| 2269 | proof (cases "S = {}")
 | |
| 2270 | case True | |
| 2271 | then show ?thesis by auto | |
| 2272 | next | |
| 2273 | case False | |
| 53339 | 2274 | then have "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` S = S)" | 
| 49529 | 2275 | using cone_iff[of S] assms by auto | 
| 53339 | 2276 | then have "0 \<in> closure S \<and> (\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` closure S = closure S)" | 
| 49529 | 2277 | using closure_subset by (auto simp add: closure_scaleR) | 
| 53339 | 2278 | then show ?thesis | 
| 60974 
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
 paulson <lp15@cam.ac.uk> parents: 
60809diff
changeset | 2279 | using False cone_iff[of "closure S"] by auto | 
| 49529 | 2280 | qed | 
| 2281 | ||
| 40377 | 2282 | |
| 60420 | 2283 | subsection \<open>Affine dependence and consequential theorems (from Lars Schewe)\<close> | 
| 33175 | 2284 | |
| 49529 | 2285 | definition affine_dependent :: "'a::real_vector set \<Rightarrow> bool" | 
| 53339 | 2286 |   where "affine_dependent s \<longleftrightarrow> (\<exists>x\<in>s. x \<in> affine hull (s - {x}))"
 | 
| 33175 | 2287 | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2288 | lemma affine_dependent_subset: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2289 | "\<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 | 2290 | 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 | 2291 | 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 | 2292 | done | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2293 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2294 | lemma affine_independent_subset: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2295 | 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 | 2296 | by (metis affine_dependent_subset) | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2297 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2298 | lemma affine_independent_Diff: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2299 | "~ 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 | 2300 | 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 | 2301 | |
| 33175 | 2302 | lemma affine_dependent_explicit: | 
| 2303 | "affine_dependent p \<longleftrightarrow> | |
| 64267 | 2304 | (\<exists>s u. finite s \<and> s \<subseteq> p \<and> sum u s = 0 \<and> | 
| 2305 | (\<exists>v\<in>s. u v \<noteq> 0) \<and> sum (\<lambda>v. u v *\<^sub>R v) s = 0)" | |
| 49529 | 2306 | unfolding affine_dependent_def affine_hull_explicit mem_Collect_eq | 
| 2307 | apply rule | |
| 2308 | apply (erule bexE, erule exE, erule exE) | |
| 2309 | apply (erule conjE)+ | |
| 2310 | defer | |
| 2311 | apply (erule exE, erule exE) | |
| 2312 | apply (erule conjE)+ | |
| 2313 | apply (erule bexE) | |
| 2314 | proof - | |
| 2315 | fix x s u | |
| 64267 | 2316 |   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 | 2317 | have "x \<notin> s" using as(1,4) by auto | 
| 64267 | 2318 | 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 | 2319 | 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 | 2320 | unfolding if_smult and sum_clauses(2)[OF as(2)] and sum_delta_notmem[OF \<open>x\<notin>s\<close>] and as | 
| 53339 | 2321 | using as | 
| 2322 | apply auto | |
| 49529 | 2323 | done | 
| 33175 | 2324 | next | 
| 49529 | 2325 | fix s u v | 
| 64267 | 2326 | 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 | 2327 |   have "s \<noteq> {v}"
 | 
| 2328 | using as(3,6) by auto | |
| 64267 | 2329 |   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 | 2330 | apply (rule_tac x=v in bexI) | 
| 2331 |     apply (rule_tac x="s - {v}" in exI)
 | |
| 2332 | apply (rule_tac x="\<lambda>x. - (1 / u v) * u x" in exI) | |
| 64267 | 2333 | unfolding scaleR_scaleR[symmetric] and scaleR_right.sum [symmetric] | 
| 2334 | unfolding sum_distrib_left[symmetric] and sum_diff1[OF as(1)] | |
| 53302 | 2335 | using as | 
| 2336 | apply auto | |
| 49529 | 2337 | done | 
| 33175 | 2338 | qed | 
| 2339 | ||
| 2340 | lemma affine_dependent_explicit_finite: | |
| 49529 | 2341 | fixes s :: "'a::real_vector set" | 
| 2342 | assumes "finite s" | |
| 53302 | 2343 | shows "affine_dependent s \<longleftrightarrow> | 
| 64267 | 2344 | (\<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 | 2345 | (is "?lhs = ?rhs") | 
| 2346 | proof | |
| 53347 | 2347 | 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 | 2348 | by auto | 
| 33175 | 2349 | assume ?lhs | 
| 49529 | 2350 | then obtain t u v where | 
| 64267 | 2351 | "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 | 2352 | unfolding affine_dependent_explicit by auto | 
| 49529 | 2353 | then show ?rhs | 
| 2354 | apply (rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI) | |
| 64267 | 2355 | apply auto unfolding * and sum.inter_restrict[OF assms, symmetric] | 
| 60420 | 2356 | unfolding Int_absorb1[OF \<open>t\<subseteq>s\<close>] | 
| 49529 | 2357 | apply auto | 
| 2358 | done | |
| 33175 | 2359 | next | 
| 2360 | assume ?rhs | |
| 64267 | 2361 | 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 | 2362 | by auto | 
| 49529 | 2363 | then show ?lhs unfolding affine_dependent_explicit | 
| 2364 | using assms by auto | |
| 2365 | qed | |
| 2366 | ||
| 33175 | 2367 | |
| 60420 | 2368 | subsection \<open>Connectedness of convex sets\<close> | 
| 44465 
fa56622bb7bc
move connected_real_lemma to the one place it is used
 huffman parents: 
44457diff
changeset | 2369 | |
| 51480 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2370 | lemma connectedD: | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2371 |   "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 | 2372 | by (rule Topological_Spaces.topological_space_class.connectedD) | 
| 33175 | 2373 | |
| 2374 | lemma convex_connected: | |
| 2375 | fixes s :: "'a::real_normed_vector set" | |
| 53302 | 2376 | assumes "convex s" | 
| 2377 | shows "connected s" | |
| 51480 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2378 | proof (rule connectedI) | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2379 | fix A B | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2380 |   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 | 2381 | moreover | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2382 |   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 | 2383 | then obtain a b where a: "a \<in> A" "a \<in> s" and b: "b \<in> B" "b \<in> s" by auto | 
| 63040 | 2384 | 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 | 2385 |   then have "continuous_on {0 .. 1} f"
 | 
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56369diff
changeset | 2386 | by (auto intro!: continuous_intros) | 
| 51480 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2387 |   then have "connected (f ` {0 .. 1})"
 | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2388 | by (auto intro!: connected_continuous_image) | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2389 | note connectedD[OF this, of A B] | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51475diff
changeset | 2390 |   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 | 2391 | 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 | 2392 |   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 | 2393 | 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 | 2394 |   moreover have "f ` {0 .. 1} \<subseteq> s"
 | 
| 60420 | 2395 | 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 | 2396 | ultimately show False by auto | 
| 33175 | 2397 | qed | 
| 2398 | ||
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61222diff
changeset | 2399 | corollary connected_UNIV[intro]: "connected (UNIV :: 'a::real_normed_vector set)" | 
| 66939 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 2400 | by (simp add: convex_connected) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 2401 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 2402 | corollary component_complement_connected: | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 2403 | fixes S :: "'a::real_normed_vector set" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 2404 | assumes "connected S" "C \<in> components (-S)" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 2405 | shows "connected(-C)" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 2406 | using component_diff_connected [of S UNIV] assms | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 2407 | by (auto simp: Compl_eq_Diff_UNIV) | 
| 33175 | 2408 | |
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62097diff
changeset | 2409 | proposition clopen: | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 2410 | fixes S :: "'a :: real_normed_vector set" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 2411 |   shows "closed S \<and> open S \<longleftrightarrow> S = {} \<or> S = UNIV"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 2412 | by (force intro!: connected_UNIV [unfolded connected_clopen, rule_format]) | 
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62097diff
changeset | 2413 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62097diff
changeset | 2414 | corollary compact_open: | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 2415 | fixes S :: "'a :: euclidean_space set" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 2416 |   shows "compact S \<and> open S \<longleftrightarrow> S = {}"
 | 
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62097diff
changeset | 2417 | by (auto simp: compact_eq_bounded_closed clopen) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62097diff
changeset | 2418 | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2419 | corollary finite_imp_not_open: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2420 |     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 | 2421 |     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 | 2422 | 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 | 2423 | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2424 | corollary empty_interior_finite: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2425 |     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 | 2426 |     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 | 2427 | 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 | 2428 | |
| 60420 | 2429 | text \<open>Balls, being convex, are connected.\<close> | 
| 33175 | 2430 | |
| 56188 | 2431 | lemma convex_prod: | 
| 53347 | 2432 |   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 | 2433 |   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 | 2434 | 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 | 2435 | 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 | 2436 | |
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 2437 | lemma convex_positive_orthant: "convex {x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i)}"
 | 
| 56188 | 2438 | by (rule convex_prod) (simp add: atLeast_def[symmetric] convex_real_interval) | 
| 33175 | 2439 | |
| 2440 | lemma convex_local_global_minimum: | |
| 2441 | fixes s :: "'a::real_normed_vector set" | |
| 53347 | 2442 | assumes "e > 0" | 
| 2443 | and "convex_on s f" | |
| 2444 | and "ball x e \<subseteq> s" | |
| 2445 | and "\<forall>y\<in>ball x e. f x \<le> f y" | |
| 33175 | 2446 | shows "\<forall>y\<in>s. f x \<le> f y" | 
| 53302 | 2447 | proof (rule ccontr) | 
| 2448 | have "x \<in> s" using assms(1,3) by auto | |
| 2449 | assume "\<not> ?thesis" | |
| 2450 | 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 | 2451 | then have xy: "0 < dist x y" by auto | 
| 53347 | 2452 | then obtain u where "0 < u" "u \<le> 1" and u: "u < e / dist x y" | 
| 60420 | 2453 | using real_lbound_gt_zero[of 1 "e / dist x y"] xy \<open>e>0\<close> by auto | 
| 53302 | 2454 | then have "f ((1-u) *\<^sub>R x + u *\<^sub>R y) \<le> (1-u) * f x + u * f y" | 
| 60420 | 2455 | using \<open>x\<in>s\<close> \<open>y\<in>s\<close> | 
| 53302 | 2456 | using assms(2)[unfolded convex_on_def, | 
| 2457 | THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x="1-u"]] | |
| 50804 | 2458 | by auto | 
| 33175 | 2459 | moreover | 
| 50804 | 2460 | have *: "x - ((1 - u) *\<^sub>R x + u *\<^sub>R y) = u *\<^sub>R (x - y)" | 
| 2461 | by (simp add: algebra_simps) | |
| 2462 | have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> ball x e" | |
| 53302 | 2463 | unfolding mem_ball dist_norm | 
| 60420 | 2464 | unfolding * and norm_scaleR and abs_of_pos[OF \<open>0<u\<close>] | 
| 50804 | 2465 | unfolding dist_norm[symmetric] | 
| 53302 | 2466 | using u | 
| 2467 | unfolding pos_less_divide_eq[OF xy] | |
| 2468 | by auto | |
| 2469 | then have "f x \<le> f ((1 - u) *\<^sub>R x + u *\<^sub>R y)" | |
| 2470 | using assms(4) by auto | |
| 50804 | 2471 | ultimately show False | 
| 60420 | 2472 | using mult_strict_left_mono[OF y \<open>u>0\<close>] | 
| 53302 | 2473 | unfolding left_diff_distrib | 
| 2474 | by auto | |
| 33175 | 2475 | qed | 
| 2476 | ||
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 2477 | lemma convex_ball [iff]: | 
| 33175 | 2478 | fixes x :: "'a::real_normed_vector" | 
| 49531 | 2479 | shows "convex (ball x e)" | 
| 50804 | 2480 | proof (auto simp add: convex_def) | 
| 2481 | fix y z | |
| 2482 | assume yz: "dist x y < e" "dist x z < e" | |
| 2483 | fix u v :: real | |
| 2484 | assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1" | |
| 2485 | have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z" | |
| 2486 | using uv yz | |
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2487 | using convex_on_dist [of "ball x e" x, unfolded convex_on_def, | 
| 53302 | 2488 | THEN bspec[where x=y], THEN bspec[where x=z]] | 
| 50804 | 2489 | by auto | 
| 2490 | then show "dist x (u *\<^sub>R y + v *\<^sub>R z) < e" | |
| 2491 | using convex_bound_lt[OF yz uv] by auto | |
| 33175 | 2492 | qed | 
| 2493 | ||
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 2494 | lemma convex_cball [iff]: | 
| 33175 | 2495 | fixes x :: "'a::real_normed_vector" | 
| 53347 | 2496 | shows "convex (cball x e)" | 
| 2497 | proof - | |
| 2498 |   {
 | |
| 2499 | fix y z | |
| 2500 | assume yz: "dist x y \<le> e" "dist x z \<le> e" | |
| 2501 | fix u v :: real | |
| 2502 | assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1" | |
| 2503 | have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z" | |
| 2504 | using uv yz | |
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2505 | using convex_on_dist [of "cball x e" x, unfolded convex_on_def, | 
| 53347 | 2506 | THEN bspec[where x=y], THEN bspec[where x=z]] | 
| 2507 | by auto | |
| 2508 | then have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> e" | |
| 2509 | using convex_bound_le[OF yz uv] by auto | |
| 2510 | } | |
| 2511 | then show ?thesis by (auto simp add: convex_def Ball_def) | |
| 33175 | 2512 | qed | 
| 2513 | ||
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2514 | lemma connected_ball [iff]: | 
| 33175 | 2515 | fixes x :: "'a::real_normed_vector" | 
| 2516 | shows "connected (ball x e)" | |
| 2517 | using convex_connected convex_ball by auto | |
| 2518 | ||
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2519 | lemma connected_cball [iff]: | 
| 33175 | 2520 | fixes x :: "'a::real_normed_vector" | 
| 53302 | 2521 | shows "connected (cball x e)" | 
| 33175 | 2522 | using convex_connected convex_cball by auto | 
| 2523 | ||
| 50804 | 2524 | |
| 60420 | 2525 | subsection \<open>Convex hull\<close> | 
| 33175 | 2526 | |
| 60762 | 2527 | lemma convex_convex_hull [iff]: "convex (convex hull s)" | 
| 53302 | 2528 | unfolding hull_def | 
| 2529 |   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 | 2530 | by auto | 
| 33175 | 2531 | |
| 63016 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 2532 | lemma convex_hull_subset: | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 2533 | "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 | 2534 | by (simp add: convex_convex_hull subset_hull) | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 2535 | |
| 34064 
eee04bbbae7e
avoid dependency on implicit dest rule predicate1D in proofs
 haftmann parents: 
33758diff
changeset | 2536 | lemma convex_hull_eq: "convex hull s = s \<longleftrightarrow> convex s" | 
| 50804 | 2537 | by (metis convex_convex_hull hull_same) | 
| 33175 | 2538 | |
| 2539 | lemma bounded_convex_hull: | |
| 2540 | fixes s :: "'a::real_normed_vector set" | |
| 53347 | 2541 | assumes "bounded s" | 
| 2542 | shows "bounded (convex hull s)" | |
| 50804 | 2543 | proof - | 
| 2544 | from assms obtain B where B: "\<forall>x\<in>s. norm x \<le> B" | |
| 2545 | unfolding bounded_iff by auto | |
| 2546 | show ?thesis | |
| 2547 | apply (rule bounded_subset[OF bounded_cball, of _ 0 B]) | |
| 44170 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 huffman parents: 
44142diff
changeset | 2548 | unfolding subset_hull[of convex, OF convex_cball] | 
| 53302 | 2549 | unfolding subset_eq mem_cball dist_norm using B | 
| 2550 | apply auto | |
| 50804 | 2551 | done | 
| 2552 | qed | |
| 33175 | 2553 | |
| 2554 | lemma finite_imp_bounded_convex_hull: | |
| 2555 | fixes s :: "'a::real_normed_vector set" | |
| 53302 | 2556 | shows "finite s \<Longrightarrow> bounded (convex hull s)" | 
| 2557 | using bounded_convex_hull finite_imp_bounded | |
| 2558 | by auto | |
| 33175 | 2559 | |
| 50804 | 2560 | |
| 60420 | 2561 | subsubsection \<open>Convex hull is "preserved" by a linear function\<close> | 
| 40377 | 2562 | |
| 2563 | lemma convex_hull_linear_image: | |
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2564 | assumes f: "linear f" | 
| 40377 | 2565 | 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 | 2566 | proof | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2567 | 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 | 2568 | 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 | 2569 | 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 | 2570 | 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 | 2571 | show "s \<subseteq> f -` (convex hull (f ` s))" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2572 | by (fast intro: hull_inc) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2573 | show "convex (f -` (convex hull (f ` s)))" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2574 | 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 | 2575 | qed | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2576 | qed | 
| 40377 | 2577 | |
| 2578 | lemma in_convex_hull_linear_image: | |
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2579 | assumes "linear f" | 
| 53347 | 2580 | and "x \<in> convex hull s" | 
| 53339 | 2581 | shows "f x \<in> convex hull (f ` s)" | 
| 50804 | 2582 | using convex_hull_linear_image[OF assms(1)] assms(2) by auto | 
| 2583 | ||
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2584 | lemma convex_hull_Times: | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2585 | "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 | 2586 | proof | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2587 | 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 | 2588 | 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 | 2589 | 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 | 2590 | proof (intro hull_induct) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2591 | 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 | 2592 | 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 | 2593 | by (simp add: hull_inc) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2594 | next | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2595 | 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 | 2596 | have "convex ?S" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2597 | 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 | 2598 | simp add: linear_iff) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2599 |     also have "?S = {y. (x, y) \<in> convex hull (s \<times> t)}"
 | 
| 57865 | 2600 | by (auto simp add: image_def Bex_def) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2601 |     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 | 2602 | next | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2603 |     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 | 2604 | 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 | 2605 | 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 | 2606 | have "convex ?S" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2607 | 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 | 2608 | simp add: linear_iff) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2609 |       also have "?S = {x. (x, y) \<in> convex hull (s \<times> t)}"
 | 
| 57865 | 2610 | by (auto simp add: image_def Bex_def) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2611 |       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 | 2612 | qed | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2613 | qed | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2614 | 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 | 2615 | unfolding subset_eq split_paired_Ball_Sigma . | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2616 | qed | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2617 | |
| 40377 | 2618 | |
| 60420 | 2619 | subsubsection \<open>Stepping theorems for convex hulls of finite sets\<close> | 
| 33175 | 2620 | |
| 2621 | lemma convex_hull_empty[simp]: "convex hull {} = {}"
 | |
| 50804 | 2622 | by (rule hull_unique) auto | 
| 33175 | 2623 | |
| 2624 | lemma convex_hull_singleton[simp]: "convex hull {a} = {a}"
 | |
| 50804 | 2625 | by (rule hull_unique) auto | 
| 33175 | 2626 | |
| 2627 | lemma convex_hull_insert: | |
| 2628 | fixes s :: "'a::real_vector set" | |
| 2629 |   assumes "s \<noteq> {}"
 | |
| 50804 | 2630 | shows "convex hull (insert a s) = | 
| 2631 |     {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 | 2632 | (is "_ = ?hull") | 
| 50804 | 2633 | apply (rule, rule hull_minimal, rule) | 
| 2634 | unfolding insert_iff | |
| 2635 | prefer 3 | |
| 2636 | apply rule | |
| 2637 | proof - | |
| 2638 | fix x | |
| 2639 | assume x: "x = a \<or> x \<in> s" | |
| 2640 | then show "x \<in> ?hull" | |
| 2641 | apply rule | |
| 2642 | unfolding mem_Collect_eq | |
| 2643 | apply (rule_tac x=1 in exI) | |
| 2644 | defer | |
| 2645 | apply (rule_tac x=0 in exI) | |
| 2646 | using assms hull_subset[of s convex] | |
| 2647 | apply auto | |
| 2648 | done | |
| 33175 | 2649 | next | 
| 50804 | 2650 | fix x | 
| 2651 | assume "x \<in> ?hull" | |
| 2652 | 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" | |
| 2653 | by auto | |
| 53339 | 2654 | have "a \<in> convex hull insert a s" "b \<in> convex hull insert a s" | 
| 50804 | 2655 |     using hull_mono[of s "insert a s" convex] hull_mono[of "{a}" "insert a s" convex] and obt(4)
 | 
| 2656 | by auto | |
| 2657 | then show "x \<in> convex hull insert a s" | |
| 53676 | 2658 | unfolding obt(5) using obt(1-3) | 
| 2659 | by (rule convexD [OF convex_convex_hull]) | |
| 33175 | 2660 | next | 
| 50804 | 2661 | show "convex ?hull" | 
| 53676 | 2662 | proof (rule convexI) | 
| 50804 | 2663 | fix x y u v | 
| 2664 | assume as: "(0::real) \<le> u" "0 \<le> v" "u + v = 1" "x\<in>?hull" "y\<in>?hull" | |
| 53339 | 2665 | from as(4) obtain u1 v1 b1 where | 
| 2666 | obt1: "u1\<ge>0" "v1\<ge>0" "u1 + v1 = 1" "b1 \<in> convex hull s" "x = u1 *\<^sub>R a + v1 *\<^sub>R b1" | |
| 2667 | by auto | |
| 2668 | from as(5) obtain u2 v2 b2 where | |
| 2669 | obt2: "u2\<ge>0" "v2\<ge>0" "u2 + v2 = 1" "b2 \<in> convex hull s" "y = u2 *\<^sub>R a + v2 *\<^sub>R b2" | |
| 2670 | by auto | |
| 50804 | 2671 | have *: "\<And>(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" | 
| 2672 | by (auto simp add: algebra_simps) | |
| 2673 | have **: "\<exists>b \<in> convex hull s. u *\<^sub>R x + v *\<^sub>R y = | |
| 2674 | (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (b - (u * u1) *\<^sub>R b - (v * u2) *\<^sub>R b)" | |
| 2675 | proof (cases "u * v1 + v * v2 = 0") | |
| 2676 | case True | |
| 2677 | have *: "\<And>(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" | |
| 2678 | by (auto simp add: algebra_simps) | |
| 2679 | from True have ***: "u * v1 = 0" "v * v2 = 0" | |
| 60420 | 2680 | 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 | 2681 | by arith+ | 
| 50804 | 2682 | then have "u * u1 + v * u2 = 1" | 
| 2683 | using as(3) obt1(3) obt2(3) by auto | |
| 2684 | then show ?thesis | |
| 2685 | unfolding obt1(5) obt2(5) * | |
| 2686 | using assms hull_subset[of s convex] | |
| 2687 | by (auto simp add: *** scaleR_right_distrib) | |
| 33175 | 2688 | next | 
| 50804 | 2689 | case False | 
| 2690 | have "1 - (u * u1 + v * u2) = (u + v) - (u * u1 + v * u2)" | |
| 2691 | using as(3) obt1(3) obt2(3) by (auto simp add: field_simps) | |
| 2692 | also have "\<dots> = u * (v1 + u1 - u1) + v * (v2 + u2 - u2)" | |
| 2693 | using as(3) obt1(3) obt2(3) by (auto simp add: field_simps) | |
| 2694 | also have "\<dots> = u * v1 + v * v2" | |
| 2695 | by simp | |
| 2696 | finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" by auto | |
| 2697 | have "0 \<le> u * v1 + v * v2" "0 \<le> u * v1" "0 \<le> u * v1 + v * v2" "0 \<le> v * v2" | |
| 56536 | 2698 | using as(1,2) obt1(1,2) obt2(1,2) by auto | 
| 50804 | 2699 | then show ?thesis | 
| 2700 | unfolding obt1(5) obt2(5) | |
| 2701 | unfolding * and ** | |
| 2702 | using False | |
| 53339 | 2703 | apply (rule_tac | 
| 2704 | x = "((u * v1) / (u * v1 + v * v2)) *\<^sub>R b1 + ((v * v2) / (u * v1 + v * v2)) *\<^sub>R b2" in bexI) | |
| 50804 | 2705 | defer | 
| 53676 | 2706 | apply (rule convexD [OF convex_convex_hull]) | 
| 50804 | 2707 | using obt1(4) obt2(4) | 
| 49530 | 2708 | unfolding add_divide_distrib[symmetric] and zero_le_divide_iff | 
| 50804 | 2709 | apply (auto simp add: scaleR_left_distrib scaleR_right_distrib) | 
| 2710 | done | |
| 2711 | qed | |
| 2712 | have u1: "u1 \<le> 1" | |
| 2713 | unfolding obt1(3)[symmetric] and not_le using obt1(2) by auto | |
| 2714 | have u2: "u2 \<le> 1" | |
| 2715 | unfolding obt2(3)[symmetric] and not_le using obt2(2) by auto | |
| 53339 | 2716 | have "u1 * u + u2 * v \<le> max u1 u2 * u + max u1 u2 * v" | 
| 50804 | 2717 | apply (rule add_mono) | 
| 2718 | apply (rule_tac [!] mult_right_mono) | |
| 2719 | using as(1,2) obt1(1,2) obt2(1,2) | |
| 2720 | apply auto | |
| 2721 | done | |
| 2722 | also have "\<dots> \<le> 1" | |
| 2723 | unfolding distrib_left[symmetric] and as(3) using u1 u2 by auto | |
| 2724 | finally show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull" | |
| 2725 | unfolding mem_Collect_eq | |
| 2726 | apply (rule_tac x="u * u1 + v * u2" in exI) | |
| 2727 | apply (rule conjI) | |
| 2728 | defer | |
| 2729 | apply (rule_tac x="1 - u * u1 - v * u2" in exI) | |
| 2730 | unfolding Bex_def | |
| 2731 | using as(1,2) obt1(1,2) obt2(1,2) ** | |
| 56536 | 2732 | apply (auto simp add: algebra_simps) | 
| 50804 | 2733 | done | 
| 33175 | 2734 | qed | 
| 2735 | qed | |
| 2736 | ||
| 66287 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65719diff
changeset | 2737 | lemma convex_hull_insert_alt: | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65719diff
changeset | 2738 | "convex hull (insert a S) = | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65719diff
changeset | 2739 |       (if S = {} then {a}
 | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65719diff
changeset | 2740 |       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 | 2741 | apply (auto simp: convex_hull_insert) | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65719diff
changeset | 2742 | using diff_eq_eq apply fastforce | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65719diff
changeset | 2743 | by (metis add.group_left_neutral add_le_imp_le_diff diff_add_cancel) | 
| 33175 | 2744 | |
| 60420 | 2745 | subsubsection \<open>Explicit expression for convex hull\<close> | 
| 33175 | 2746 | |
| 2747 | lemma convex_hull_indexed: | |
| 2748 | fixes s :: "'a::real_vector set" | |
| 50804 | 2749 | shows "convex hull s = | 
| 53347 | 2750 |     {y. \<exists>k u x.
 | 
| 2751 |       (\<forall>i\<in>{1::nat .. k}. 0 \<le> u i \<and> x i \<in> s) \<and>
 | |
| 64267 | 2752 |       (sum u {1..k} = 1) \<and> (sum (\<lambda>i. u i *\<^sub>R x i) {1..k} = y)}"
 | 
| 53339 | 2753 | (is "?xyz = ?hull") | 
| 50804 | 2754 | apply (rule hull_unique) | 
| 2755 | apply rule | |
| 2756 | defer | |
| 53676 | 2757 | apply (rule convexI) | 
| 50804 | 2758 | proof - | 
| 2759 | fix x | |
| 2760 | assume "x\<in>s" | |
| 2761 | then show "x \<in> ?hull" | |
| 2762 | unfolding mem_Collect_eq | |
| 2763 | apply (rule_tac x=1 in exI, rule_tac x="\<lambda>x. 1" in exI) | |
| 2764 | apply auto | |
| 2765 | done | |
| 33175 | 2766 | next | 
| 50804 | 2767 | fix t | 
| 2768 | assume as: "s \<subseteq> t" "convex t" | |
| 2769 | show "?hull \<subseteq> t" | |
| 2770 | apply rule | |
| 2771 | unfolding mem_Collect_eq | |
| 53302 | 2772 | apply (elim exE conjE) | 
| 50804 | 2773 | proof - | 
| 2774 | fix x k u y | |
| 2775 | assume assm: | |
| 2776 |       "\<forall>i\<in>{1::nat..k}. 0 \<le> u i \<and> y i \<in> s"
 | |
| 64267 | 2777 |       "sum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R y i) = x"
 | 
| 50804 | 2778 | show "x\<in>t" | 
| 2779 | unfolding assm(3) [symmetric] | |
| 2780 | apply (rule as(2)[unfolded convex, rule_format]) | |
| 2781 | using assm(1,2) as(1) apply auto | |
| 2782 | done | |
| 2783 | qed | |
| 33175 | 2784 | next | 
| 50804 | 2785 | fix x y u v | 
| 53347 | 2786 | assume uv: "0 \<le> u" "0 \<le> v" "u + v = (1::real)" | 
| 2787 | assume xy: "x \<in> ?hull" "y \<in> ?hull" | |
| 50804 | 2788 | from xy obtain k1 u1 x1 where | 
| 64267 | 2789 |     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 | 2790 | by auto | 
| 2791 | from xy obtain k2 u2 x2 where | |
| 64267 | 2792 |     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 | 2793 | by auto | 
| 2794 | have *: "\<And>P (x1::'a) x2 s1 s2 i. | |
| 2795 | (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 | 2796 |     "{1..k1 + k2} \<inter> {1..k1} = {1..k1}" "{1..k1 + k2} \<inter> - {1..k1} = (\<lambda>i. i + k1) ` {1..k2}"
 | 
| 50804 | 2797 | prefer 3 | 
| 2798 | apply (rule, rule) | |
| 2799 | unfolding image_iff | |
| 2800 | apply (rule_tac x = "x - k1" in bexI) | |
| 2801 | apply (auto simp add: not_le) | |
| 2802 | done | |
| 2803 |   have inj: "inj_on (\<lambda>i. i + k1) {1..k2}"
 | |
| 2804 | unfolding inj_on_def by auto | |
| 2805 | show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull" | |
| 2806 | apply rule | |
| 2807 | apply (rule_tac x="k1 + k2" in exI) | |
| 2808 |     apply (rule_tac x="\<lambda>i. if i \<in> {1..k1} then u * u1 i else v * u2 (i - k1)" in exI)
 | |
| 2809 |     apply (rule_tac x="\<lambda>i. if i \<in> {1..k1} then x1 i else x2 (i - k1)" in exI)
 | |
| 2810 | apply (rule, rule) | |
| 2811 | defer | |
| 2812 | apply rule | |
| 64267 | 2813 | unfolding * and sum.If_cases[OF finite_atLeastAtMost[of 1 "k1 + k2"]] and | 
| 2814 | sum.reindex[OF inj] and o_def Collect_mem_eq | |
| 2815 | unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] sum_distrib_left[symmetric] | |
| 50804 | 2816 | proof - | 
| 2817 | fix i | |
| 2818 |     assume i: "i \<in> {1..k1+k2}"
 | |
| 2819 |     show "0 \<le> (if i \<in> {1..k1} then u * u1 i else v * u2 (i - k1)) \<and>
 | |
| 2820 |       (if i \<in> {1..k1} then x1 i else x2 (i - k1)) \<in> s"
 | |
| 2821 |     proof (cases "i\<in>{1..k1}")
 | |
| 2822 | case True | |
| 2823 | then show ?thesis | |
| 56536 | 2824 | using uv(1) x(1)[THEN bspec[where x=i]] by auto | 
| 50804 | 2825 | next | 
| 2826 | case False | |
| 63040 | 2827 | define j where "j = i - k1" | 
| 53347 | 2828 |       from i False have "j \<in> {1..k2}"
 | 
| 2829 | unfolding j_def by auto | |
| 50804 | 2830 | then show ?thesis | 
| 56536 | 2831 | using False uv(2) y(1)[THEN bspec[where x=j]] | 
| 2832 | by (auto simp: j_def[symmetric]) | |
| 50804 | 2833 | qed | 
| 2834 | qed (auto simp add: not_le x(2,3) y(2,3) uv(3)) | |
| 33175 | 2835 | qed | 
| 2836 | ||
| 2837 | lemma convex_hull_finite: | |
| 2838 | fixes s :: "'a::real_vector set" | |
| 2839 | assumes "finite s" | |
| 2840 |   shows "convex hull s = {y. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and>
 | |
| 64267 | 2841 | sum u s = 1 \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y}" | 
| 53339 | 2842 | (is "?HULL = ?set") | 
| 50804 | 2843 | proof (rule hull_unique, auto simp add: convex_def[of ?set]) | 
| 2844 | fix x | |
| 2845 | assume "x \<in> s" | |
| 64267 | 2846 | 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 | 2847 | apply (rule_tac x="\<lambda>y. if x=y then 1 else 0" in exI) | 
| 2848 | apply auto | |
| 64267 | 2849 | unfolding sum.delta'[OF assms] and sum_delta''[OF assms] | 
| 50804 | 2850 | apply auto | 
| 2851 | done | |
| 33175 | 2852 | next | 
| 50804 | 2853 | fix u v :: real | 
| 2854 | assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1" | |
| 64267 | 2855 | fix ux assume ux: "\<forall>x\<in>s. 0 \<le> ux x" "sum ux s = (1::real)" | 
| 2856 | fix uy assume uy: "\<forall>x\<in>s. 0 \<le> uy x" "sum uy s = (1::real)" | |
| 53339 | 2857 |   {
 | 
| 2858 | fix x | |
| 50804 | 2859 | assume "x\<in>s" | 
| 2860 | then have "0 \<le> u * ux x + v * uy x" | |
| 2861 | using ux(1)[THEN bspec[where x=x]] uy(1)[THEN bspec[where x=x]] and uv(1,2) | |
| 56536 | 2862 | by auto | 
| 50804 | 2863 | } | 
| 2864 | moreover | |
| 2865 | have "(\<Sum>x\<in>s. u * ux x + v * uy x) = 1" | |
| 64267 | 2866 | unfolding sum.distrib and sum_distrib_left[symmetric] and ux(2) uy(2) | 
| 53302 | 2867 | using uv(3) by auto | 
| 50804 | 2868 | moreover | 
| 2869 | 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 | 2870 | unfolding scaleR_left_distrib and sum.distrib and scaleR_scaleR[symmetric] | 
| 2871 | and scaleR_right.sum [symmetric] | |
| 50804 | 2872 | by auto | 
| 2873 | ultimately | |
| 64267 | 2874 | show "\<exists>uc. (\<forall>x\<in>s. 0 \<le> uc x) \<and> sum uc s = 1 \<and> | 
| 50804 | 2875 | (\<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)" | 
| 2876 | apply (rule_tac x="\<lambda>x. u * ux x + v * uy x" in exI) | |
| 2877 | apply auto | |
| 2878 | done | |
| 33175 | 2879 | next | 
| 50804 | 2880 | fix t | 
| 2881 | assume t: "s \<subseteq> t" "convex t" | |
| 2882 | fix u | |
| 64267 | 2883 | assume u: "\<forall>x\<in>s. 0 \<le> u x" "sum u s = (1::real)" | 
| 50804 | 2884 | then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> t" | 
| 2885 | using t(2)[unfolded convex_explicit, THEN spec[where x=s], THEN spec[where x=u]] | |
| 33175 | 2886 | using assms and t(1) by auto | 
| 2887 | qed | |
| 2888 | ||
| 50804 | 2889 | |
| 60420 | 2890 | subsubsection \<open>Another formulation from Lars Schewe\<close> | 
| 33175 | 2891 | |
| 2892 | lemma convex_hull_explicit: | |
| 2893 | fixes p :: "'a::real_vector set" | |
| 53347 | 2894 | shows "convex hull p = | 
| 64267 | 2895 |     {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 | 2896 | (is "?lhs = ?rhs") | 
| 50804 | 2897 | proof - | 
| 53302 | 2898 |   {
 | 
| 2899 | fix x | |
| 2900 | assume "x\<in>?lhs" | |
| 50804 | 2901 | then obtain k u y where | 
| 64267 | 2902 |         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 | 2903 | unfolding convex_hull_indexed by auto | 
| 2904 | ||
| 50804 | 2905 |     have fin: "finite {1..k}" by auto
 | 
| 2906 |     have fin': "\<And>v. finite {i \<in> {1..k}. y i = v}" by auto
 | |
| 53302 | 2907 |     {
 | 
| 2908 | fix j | |
| 50804 | 2909 |       assume "j\<in>{1..k}"
 | 
| 64267 | 2910 |       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 | 2911 | using obt(1)[THEN bspec[where x=j]] and obt(2) | 
| 2912 | apply simp | |
| 64267 | 2913 | apply (rule sum_nonneg) | 
| 50804 | 2914 | using obt(1) | 
| 2915 | apply auto | |
| 2916 | done | |
| 2917 | } | |
| 33175 | 2918 | moreover | 
| 64267 | 2919 |     have "(\<Sum>v\<in>y ` {1..k}. sum u {i \<in> {1..k}. y i = v}) = 1"
 | 
| 2920 | unfolding sum_image_gen[OF fin, symmetric] using obt(2) by auto | |
| 2921 |     moreover have "(\<Sum>v\<in>y ` {1..k}. sum u {i \<in> {1..k}. y i = v} *\<^sub>R v) = x"
 | |
| 2922 | using sum_image_gen[OF fin, of "\<lambda>i. u i *\<^sub>R y i" y, symmetric] | |
| 2923 | unfolding scaleR_left.sum using obt(3) by auto | |
| 50804 | 2924 | ultimately | 
| 64267 | 2925 | 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 | 2926 |       apply (rule_tac x="y ` {1..k}" in exI)
 | 
| 64267 | 2927 |       apply (rule_tac x="\<lambda>v. sum u {i\<in>{1..k}. y i = v}" in exI)
 | 
| 50804 | 2928 | apply auto | 
| 2929 | done | |
| 2930 | then have "x\<in>?rhs" by auto | |
| 2931 | } | |
| 33175 | 2932 | moreover | 
| 53302 | 2933 |   {
 | 
| 2934 | fix y | |
| 2935 | assume "y\<in>?rhs" | |
| 50804 | 2936 | then obtain s u where | 
| 64267 | 2937 | 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 | 2938 | by auto | 
| 50804 | 2939 | |
| 2940 |     obtain f where f: "inj_on f {1..card s}" "f ` {1..card s} = s"
 | |
| 2941 | using ex_bij_betw_nat_finite_1[OF obt(1)] unfolding bij_betw_def by auto | |
| 2942 | ||
| 53302 | 2943 |     {
 | 
| 2944 | fix i :: nat | |
| 50804 | 2945 |       assume "i\<in>{1..card s}"
 | 
| 2946 | then have "f i \<in> s" | |
| 2947 | apply (subst f(2)[symmetric]) | |
| 2948 | apply auto | |
| 2949 | done | |
| 2950 | then have "0 \<le> u (f i)" "f i \<in> p" using obt(2,3) by auto | |
| 2951 | } | |
| 53347 | 2952 |     moreover have *: "finite {1..card s}" by auto
 | 
| 53302 | 2953 |     {
 | 
| 2954 | fix y | |
| 50804 | 2955 | assume "y\<in>s" | 
| 53302 | 2956 |       then obtain i where "i\<in>{1..card s}" "f i = y"
 | 
| 2957 |         using f using image_iff[of y f "{1..card s}"]
 | |
| 50804 | 2958 | by auto | 
| 2959 |       then have "{x. Suc 0 \<le> x \<and> x \<le> card s \<and> f x = y} = {i}"
 | |
| 2960 | apply auto | |
| 2961 | using f(1)[unfolded inj_on_def] | |
| 2962 | apply(erule_tac x=x in ballE) | |
| 2963 | apply auto | |
| 2964 | done | |
| 2965 |       then have "card {x. Suc 0 \<le> x \<and> x \<le> card s \<and> f x = y} = 1" by auto
 | |
| 2966 |       then have "(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x)) = u y"
 | |
| 2967 |           "(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y"
 | |
| 64267 | 2968 | by (auto simp add: sum_constant_scaleR) | 
| 50804 | 2969 | } | 
| 2970 | 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 | 2971 | unfolding sum_image_gen[OF *(1), of "\<lambda>x. u (f x) *\<^sub>R f x" f] | 
| 2972 | and sum_image_gen[OF *(1), of "\<lambda>x. u (f x)" f] | |
| 53339 | 2973 | unfolding f | 
| 64267 | 2974 |       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"]
 | 
| 2975 |       using sum.cong [of s s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x))" u]
 | |
| 53302 | 2976 | unfolding obt(4,5) | 
| 2977 | by auto | |
| 50804 | 2978 | ultimately | 
| 64267 | 2979 |     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 | 2980 | (\<Sum>i::nat = 1..k. u i *\<^sub>R x i) = y" | 
| 2981 | apply (rule_tac x="card s" in exI) | |
| 2982 | apply (rule_tac x="u \<circ> f" in exI) | |
| 2983 | apply (rule_tac x=f in exI) | |
| 2984 | apply fastforce | |
| 2985 | done | |
| 53302 | 2986 | then have "y \<in> ?lhs" | 
| 2987 | unfolding convex_hull_indexed by auto | |
| 50804 | 2988 | } | 
| 53302 | 2989 | ultimately show ?thesis | 
| 2990 | unfolding set_eq_iff by blast | |
| 33175 | 2991 | qed | 
| 2992 | ||
| 50804 | 2993 | |
| 60420 | 2994 | subsubsection \<open>A stepping theorem for that expansion\<close> | 
| 33175 | 2995 | |
| 2996 | lemma convex_hull_finite_step: | |
| 50804 | 2997 | fixes s :: "'a::real_vector set" | 
| 2998 | assumes "finite s" | |
| 53302 | 2999 | shows | 
| 64267 | 3000 | "(\<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) | 
| 3001 | \<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 | 3002 | (is "?lhs = ?rhs") | 
| 50804 | 3003 | proof (rule, case_tac[!] "a\<in>s") | 
| 53302 | 3004 | assume "a \<in> s" | 
| 53339 | 3005 | then have *: "insert a s = s" by auto | 
| 50804 | 3006 | assume ?lhs | 
| 3007 | then show ?rhs | |
| 3008 | unfolding * | |
| 3009 | apply (rule_tac x=0 in exI) | |
| 3010 | apply auto | |
| 3011 | done | |
| 33175 | 3012 | next | 
| 50804 | 3013 | assume ?lhs | 
| 53302 | 3014 | then obtain u where | 
| 64267 | 3015 | 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 | 3016 | by auto | 
| 3017 | assume "a \<notin> s" | |
| 3018 | then show ?rhs | |
| 3019 | apply (rule_tac x="u a" in exI) | |
| 3020 | using u(1)[THEN bspec[where x=a]] | |
| 3021 | apply simp | |
| 3022 | apply (rule_tac x=u in exI) | |
| 64267 | 3023 | using u[unfolded sum_clauses(2)[OF assms]] and \<open>a\<notin>s\<close> | 
| 50804 | 3024 | apply auto | 
| 3025 | done | |
| 33175 | 3026 | next | 
| 50804 | 3027 | assume "a \<in> s" | 
| 3028 | then have *: "insert a s = s" by auto | |
| 3029 | have fin: "finite (insert a s)" using assms by auto | |
| 3030 | assume ?rhs | |
| 64267 | 3031 | 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 | 3032 | by auto | 
| 3033 | show ?lhs | |
| 3034 | apply (rule_tac x = "\<lambda>x. (if a = x then v else 0) + u x" in exI) | |
| 64267 | 3035 | unfolding scaleR_left_distrib and sum.distrib and sum_delta''[OF fin] and sum.delta'[OF fin] | 
| 3036 | unfolding sum_clauses(2)[OF assms] | |
| 60420 | 3037 | using uv and uv(2)[THEN bspec[where x=a]] and \<open>a\<in>s\<close> | 
| 50804 | 3038 | apply auto | 
| 3039 | done | |
| 33175 | 3040 | next | 
| 50804 | 3041 | assume ?rhs | 
| 53339 | 3042 | then obtain v u where | 
| 64267 | 3043 | 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 | 3044 | by auto | 
| 3045 | moreover | |
| 3046 | assume "a \<notin> s" | |
| 3047 | moreover | |
| 64267 | 3048 | have "(\<Sum>x\<in>s. if a = x then v else u x) = sum u s" | 
| 53302 | 3049 | 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 | 3050 | apply (rule_tac sum.cong) apply rule | 
| 50804 | 3051 | defer | 
| 64267 | 3052 | apply (rule_tac sum.cong) apply rule | 
| 60420 | 3053 | using \<open>a \<notin> s\<close> | 
| 50804 | 3054 | apply auto | 
| 3055 | done | |
| 3056 | ultimately show ?lhs | |
| 3057 | apply (rule_tac x="\<lambda>x. if a = x then v else u x" in exI) | |
| 64267 | 3058 | unfolding sum_clauses(2)[OF assms] | 
| 50804 | 3059 | apply auto | 
| 3060 | done | |
| 3061 | qed | |
| 3062 | ||
| 33175 | 3063 | |
| 60420 | 3064 | subsubsection \<open>Hence some special cases\<close> | 
| 33175 | 3065 | |
| 3066 | lemma convex_hull_2: | |
| 3067 |   "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 | 3068 | proof - | 
| 3069 |   have *: "\<And>u. (\<forall>x\<in>{a, b}. 0 \<le> u x) \<longleftrightarrow> 0 \<le> u a \<and> 0 \<le> u b"
 | |
| 3070 | by auto | |
| 3071 |   have **: "finite {b}" by auto
 | |
| 3072 | show ?thesis | |
| 3073 | apply (simp add: convex_hull_finite) | |
| 3074 | unfolding convex_hull_finite_step[OF **, of a 1, unfolded * conj_assoc] | |
| 3075 | apply auto | |
| 3076 | apply (rule_tac x=v in exI) | |
| 3077 | apply (rule_tac x="1 - v" in exI) | |
| 3078 | apply simp | |
| 3079 | apply (rule_tac x=u in exI) | |
| 3080 | apply simp | |
| 3081 | apply (rule_tac x="\<lambda>x. v" in exI) | |
| 3082 | apply simp | |
| 3083 | done | |
| 3084 | qed | |
| 33175 | 3085 | |
| 3086 | 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 | 3087 | unfolding convex_hull_2 | 
| 53302 | 3088 | proof (rule Collect_cong) | 
| 3089 | have *: "\<And>x y ::real. x + y = 1 \<longleftrightarrow> x = 1 - y" | |
| 3090 | by auto | |
| 3091 | fix x | |
| 3092 | 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> | |
| 3093 | (\<exists>u. x = a + u *\<^sub>R (b - a) \<and> 0 \<le> u \<and> u \<le> 1)" | |
| 3094 | unfolding * | |
| 3095 | apply auto | |
| 3096 | apply (rule_tac[!] x=u in exI) | |
| 3097 | apply (auto simp add: algebra_simps) | |
| 3098 | done | |
| 3099 | qed | |
| 33175 | 3100 | |
| 3101 | lemma convex_hull_3: | |
| 3102 |   "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 | 3103 | proof - | 
| 3104 |   have fin: "finite {a,b,c}" "finite {b,c}" "finite {c}"
 | |
| 3105 | by auto | |
| 3106 | 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 | 3107 | by (auto simp add: field_simps) | 
| 53302 | 3108 | show ?thesis | 
| 3109 | unfolding convex_hull_finite[OF fin(1)] and convex_hull_finite_step[OF fin(2)] and * | |
| 3110 | unfolding convex_hull_finite_step[OF fin(3)] | |
| 3111 | apply (rule Collect_cong) | |
| 3112 | apply simp | |
| 3113 | apply auto | |
| 3114 | apply (rule_tac x=va in exI) | |
| 3115 | apply (rule_tac x="u c" in exI) | |
| 3116 | apply simp | |
| 3117 | apply (rule_tac x="1 - v - w" in exI) | |
| 3118 | apply simp | |
| 3119 | apply (rule_tac x=v in exI) | |
| 3120 | apply simp | |
| 3121 | apply (rule_tac x="\<lambda>x. w" in exI) | |
| 3122 | apply simp | |
| 3123 | done | |
| 3124 | qed | |
| 33175 | 3125 | |
| 3126 | lemma convex_hull_3_alt: | |
| 3127 |   "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 | 3128 | proof - | 
| 3129 | have *: "\<And>x y z ::real. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z" | |
| 3130 | by auto | |
| 3131 | show ?thesis | |
| 3132 | unfolding convex_hull_3 | |
| 3133 | apply (auto simp add: *) | |
| 3134 | apply (rule_tac x=v in exI) | |
| 3135 | apply (rule_tac x=w in exI) | |
| 3136 | apply (simp add: algebra_simps) | |
| 3137 | apply (rule_tac x=u in exI) | |
| 3138 | apply (rule_tac x=v in exI) | |
| 3139 | apply (simp add: algebra_simps) | |
| 3140 | done | |
| 3141 | qed | |
| 3142 | ||
| 33175 | 3143 | |
| 60420 | 3144 | subsection \<open>Relations among closure notions and corresponding hulls\<close> | 
| 33175 | 3145 | |
| 3146 | lemma affine_imp_convex: "affine s \<Longrightarrow> convex s" | |
| 3147 | unfolding affine_def convex_def by auto | |
| 3148 | ||
| 64394 | 3149 | lemma convex_affine_hull [simp]: "convex (affine hull S)" | 
| 3150 | by (simp add: affine_imp_convex) | |
| 3151 | ||
| 44361 
75ec83d45303
remove unnecessary euclidean_space class constraints
 huffman parents: 
44349diff
changeset | 3152 | lemma subspace_imp_convex: "subspace s \<Longrightarrow> convex s" | 
| 33175 | 3153 | using subspace_imp_affine affine_imp_convex by auto | 
| 3154 | ||
| 44361 
75ec83d45303
remove unnecessary euclidean_space class constraints
 huffman parents: 
44349diff
changeset | 3155 | lemma affine_hull_subset_span: "(affine hull s) \<subseteq> (span s)" | 
| 53302 | 3156 | by (metis hull_minimal span_inc subspace_imp_affine subspace_span) | 
| 33175 | 3157 | |
| 44361 
75ec83d45303
remove unnecessary euclidean_space class constraints
 huffman parents: 
44349diff
changeset | 3158 | lemma convex_hull_subset_span: "(convex hull s) \<subseteq> (span s)" | 
| 53302 | 3159 | by (metis hull_minimal span_inc subspace_imp_convex subspace_span) | 
| 33175 | 3160 | |
| 3161 | lemma convex_hull_subset_affine_hull: "(convex hull s) \<subseteq> (affine hull s)" | |
| 53302 | 3162 | by (metis affine_affine_hull affine_imp_convex hull_minimal hull_subset) | 
| 3163 | ||
| 3164 | lemma affine_dependent_imp_dependent: "affine_dependent s \<Longrightarrow> dependent s" | |
| 49531 | 3165 | unfolding affine_dependent_def dependent_def | 
| 33175 | 3166 | using affine_hull_subset_span by auto | 
| 3167 | ||
| 3168 | lemma dependent_imp_affine_dependent: | |
| 53302 | 3169 |   assumes "dependent {x - a| x . x \<in> s}"
 | 
| 3170 | and "a \<notin> s" | |
| 33175 | 3171 | shows "affine_dependent (insert a s)" | 
| 53302 | 3172 | proof - | 
| 49531 | 3173 | from assms(1)[unfolded dependent_explicit] obtain S u v | 
| 53347 | 3174 |     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"
 | 
| 3175 | by auto | |
| 63040 | 3176 | define t where "t = (\<lambda>x. x + a) ` S" | 
| 33175 | 3177 | |
| 53347 | 3178 | have inj: "inj_on (\<lambda>x. x + a) S" | 
| 53302 | 3179 | unfolding inj_on_def by auto | 
| 3180 | have "0 \<notin> S" | |
| 3181 | using obt(2) assms(2) unfolding subset_eq by auto | |
| 53347 | 3182 | have fin: "finite t" and "t \<subseteq> s" | 
| 53302 | 3183 | unfolding t_def using obt(1,2) by auto | 
| 3184 | then have "finite (insert a t)" and "insert a t \<subseteq> insert a s" | |
| 3185 | by auto | |
| 3186 | 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 | 3187 | apply (rule sum.cong) | 
| 60420 | 3188 | using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close> | 
| 53302 | 3189 | apply auto | 
| 3190 | done | |
| 33175 | 3191 | have "(\<Sum>x\<in>insert a t. if x = a then - (\<Sum>x\<in>t. u (x - a)) else u (x - a)) = 0" | 
| 64267 | 3192 | unfolding sum_clauses(2)[OF fin] | 
| 60420 | 3193 | using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close> | 
| 53302 | 3194 | apply auto | 
| 3195 | unfolding * | |
| 3196 | apply auto | |
| 3197 | done | |
| 33175 | 3198 | 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 | 3199 | apply (rule_tac x="v + a" in bexI) | 
| 60420 | 3200 | using obt(3,4) and \<open>0\<notin>S\<close> | 
| 53302 | 3201 | unfolding t_def | 
| 3202 | apply auto | |
| 3203 | done | |
| 3204 | 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 | 3205 | apply (rule sum.cong) | 
| 60420 | 3206 | using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close> | 
| 53302 | 3207 | apply auto | 
| 3208 | done | |
| 49531 | 3209 | have "(\<Sum>x\<in>t. u (x - a)) *\<^sub>R a = (\<Sum>v\<in>t. u (v - a) *\<^sub>R v)" | 
| 64267 | 3210 | unfolding scaleR_left.sum | 
| 3211 | unfolding t_def and sum.reindex[OF inj] and o_def | |
| 53302 | 3212 | using obt(5) | 
| 64267 | 3213 | by (auto simp add: sum.distrib scaleR_right_distrib) | 
| 53302 | 3214 | 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 | 3215 | unfolding sum_clauses(2)[OF fin] | 
| 60420 | 3216 | using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close> | 
| 53302 | 3217 | by (auto simp add: *) | 
| 3218 | ultimately show ?thesis | |
| 3219 | unfolding affine_dependent_explicit | |
| 3220 | apply (rule_tac x="insert a t" in exI) | |
| 3221 | apply auto | |
| 3222 | done | |
| 33175 | 3223 | qed | 
| 3224 | ||
| 3225 | lemma convex_cone: | |
| 53302 | 3226 | "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)" | 
| 3227 | (is "?lhs = ?rhs") | |
| 3228 | proof - | |
| 3229 |   {
 | |
| 3230 | fix x y | |
| 3231 | assume "x\<in>s" "y\<in>s" and ?lhs | |
| 3232 | then have "2 *\<^sub>R x \<in>s" "2 *\<^sub>R y \<in> s" | |
| 3233 | unfolding cone_def by auto | |
| 3234 | then have "x + y \<in> s" | |
| 60420 | 3235 | using \<open>?lhs\<close>[unfolded convex_def, THEN conjunct1] | 
| 53302 | 3236 | apply (erule_tac x="2*\<^sub>R x" in ballE) | 
| 3237 | apply (erule_tac x="2*\<^sub>R y" in ballE) | |
| 3238 | apply (erule_tac x="1/2" in allE) | |
| 3239 | apply simp | |
| 3240 | apply (erule_tac x="1/2" in allE) | |
| 3241 | apply auto | |
| 3242 | done | |
| 3243 | } | |
| 3244 | then show ?thesis | |
| 3245 | unfolding convex_def cone_def by blast | |
| 3246 | qed | |
| 3247 | ||
| 3248 | lemma affine_dependent_biggerset: | |
| 53347 | 3249 | 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 | 3250 |   assumes "finite s" "card s \<ge> DIM('a) + 2"
 | 
| 33175 | 3251 | shows "affine_dependent s" | 
| 53302 | 3252 | proof - | 
| 3253 |   have "s \<noteq> {}" using assms by auto
 | |
| 3254 | then obtain a where "a\<in>s" by auto | |
| 3255 |   have *: "{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})"
 | |
| 3256 | by auto | |
| 3257 |   have "card {x - a |x. x \<in> s - {a}} = card (s - {a})"
 | |
| 3258 | unfolding * | |
| 3259 | apply (rule card_image) | |
| 3260 | unfolding inj_on_def | |
| 3261 | apply auto | |
| 3262 | done | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 3263 |   also have "\<dots> > DIM('a)" using assms(2)
 | 
| 60420 | 3264 | unfolding card_Diff_singleton[OF assms(1) \<open>a\<in>s\<close>] by auto | 
| 53302 | 3265 | finally show ?thesis | 
| 60420 | 3266 | apply (subst insert_Diff[OF \<open>a\<in>s\<close>, symmetric]) | 
| 53302 | 3267 | apply (rule dependent_imp_affine_dependent) | 
| 3268 | apply (rule dependent_biggerset) | |
| 3269 | apply auto | |
| 3270 | done | |
| 3271 | qed | |
| 33175 | 3272 | |
| 3273 | lemma affine_dependent_biggerset_general: | |
| 53347 | 3274 | assumes "finite (s :: 'a::euclidean_space set)" | 
| 3275 | and "card s \<ge> dim s + 2" | |
| 33175 | 3276 | shows "affine_dependent s" | 
| 53302 | 3277 | proof - | 
| 33175 | 3278 |   from assms(2) have "s \<noteq> {}" by auto
 | 
| 3279 | then obtain a where "a\<in>s" by auto | |
| 53302 | 3280 |   have *: "{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})"
 | 
| 3281 | by auto | |
| 3282 |   have **: "card {x - a |x. x \<in> s - {a}} = card (s - {a})"
 | |
| 3283 | unfolding * | |
| 3284 | apply (rule card_image) | |
| 3285 | unfolding inj_on_def | |
| 3286 | apply auto | |
| 3287 | done | |
| 33175 | 3288 |   have "dim {x - a |x. x \<in> s - {a}} \<le> dim s"
 | 
| 53302 | 3289 | apply (rule subset_le_dim) | 
| 3290 | unfolding subset_eq | |
| 60420 | 3291 | using \<open>a\<in>s\<close> | 
| 63938 | 3292 | apply (auto simp add:span_superset span_diff) | 
| 53302 | 3293 | done | 
| 33175 | 3294 | also have "\<dots> < dim s + 1" by auto | 
| 53302 | 3295 |   also have "\<dots> \<le> card (s - {a})"
 | 
| 3296 | using assms | |
| 60420 | 3297 | using card_Diff_singleton[OF assms(1) \<open>a\<in>s\<close>] | 
| 53302 | 3298 | by auto | 
| 3299 | finally show ?thesis | |
| 60420 | 3300 | apply (subst insert_Diff[OF \<open>a\<in>s\<close>, symmetric]) | 
| 53302 | 3301 | apply (rule dependent_imp_affine_dependent) | 
| 3302 | apply (rule dependent_biggerset_general) | |
| 3303 | unfolding ** | |
| 3304 | apply auto | |
| 3305 | done | |
| 3306 | qed | |
| 3307 | ||
| 33175 | 3308 | |
| 60420 | 3309 | subsection \<open>Some Properties of Affine Dependent Sets\<close> | 
| 40377 | 3310 | |
| 66287 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65719diff
changeset | 3311 | lemma affine_independent_0 [simp]: "\<not> affine_dependent {}"
 | 
| 40377 | 3312 | by (simp add: affine_dependent_def) | 
| 3313 | ||
| 66287 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65719diff
changeset | 3314 | lemma affine_independent_1 [simp]: "\<not> affine_dependent {a}"
 | 
| 53302 | 3315 | by (simp add: affine_dependent_def) | 
| 3316 | ||
| 66287 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65719diff
changeset | 3317 | 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 | 3318 | 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 | 3319 | |
| 53302 | 3320 | lemma affine_hull_translation: "affine hull ((\<lambda>x. a + x) ` S) = (\<lambda>x. a + x) ` (affine hull S)" | 
| 3321 | proof - | |
| 3322 | have "affine ((\<lambda>x. a + x) ` (affine hull S))" | |
| 60303 | 3323 | using affine_translation affine_affine_hull by blast | 
| 53347 | 3324 | moreover have "(\<lambda>x. a + x) ` S \<subseteq> (\<lambda>x. a + x) ` (affine hull S)" | 
| 53302 | 3325 | using hull_subset[of S] by auto | 
| 53347 | 3326 | ultimately have h1: "affine hull ((\<lambda>x. a + x) ` S) \<subseteq> (\<lambda>x. a + x) ` (affine hull S)" | 
| 53302 | 3327 | by (metis hull_minimal) | 
| 3328 | have "affine((\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) ` S)))" | |
| 60303 | 3329 | using affine_translation affine_affine_hull by blast | 
| 53347 | 3330 | moreover have "(\<lambda>x. -a + x) ` (\<lambda>x. a + x) ` S \<subseteq> (\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) ` S))" | 
| 53302 | 3331 | using hull_subset[of "(\<lambda>x. a + x) ` S"] by auto | 
| 53347 | 3332 | moreover have "S = (\<lambda>x. -a + x) ` (\<lambda>x. a + x) ` S" | 
| 53302 | 3333 | using translation_assoc[of "-a" a] by auto | 
| 3334 | ultimately have "(\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) ` S)) >= (affine hull S)" | |
| 3335 | by (metis hull_minimal) | |
| 3336 | then have "affine hull ((\<lambda>x. a + x) ` S) >= (\<lambda>x. a + x) ` (affine hull S)" | |
| 3337 | by auto | |
| 54465 | 3338 | then show ?thesis using h1 by auto | 
| 40377 | 3339 | qed | 
| 3340 | ||
| 3341 | lemma affine_dependent_translation: | |
| 3342 | assumes "affine_dependent S" | |
| 53339 | 3343 | shows "affine_dependent ((\<lambda>x. a + x) ` S)" | 
| 53302 | 3344 | proof - | 
| 54465 | 3345 |   obtain x where x: "x \<in> S \<and> x \<in> affine hull (S - {x})"
 | 
| 53302 | 3346 | using assms affine_dependent_def by auto | 
| 3347 |   have "op + a ` (S - {x}) = op + a ` S - {a + x}"
 | |
| 3348 | by auto | |
| 53347 | 3349 |   then have "a + x \<in> affine hull ((\<lambda>x. a + x) ` S - {a + x})"
 | 
| 54465 | 3350 |     using affine_hull_translation[of a "S - {x}"] x by auto
 | 
| 53347 | 3351 | moreover have "a + x \<in> (\<lambda>x. a + x) ` S" | 
| 54465 | 3352 | using x by auto | 
| 53302 | 3353 | ultimately show ?thesis | 
| 3354 | unfolding affine_dependent_def by auto | |
| 40377 | 3355 | qed | 
| 3356 | ||
| 3357 | lemma affine_dependent_translation_eq: | |
| 54465 | 3358 | "affine_dependent S \<longleftrightarrow> affine_dependent ((\<lambda>x. a + x) ` S)" | 
| 53302 | 3359 | proof - | 
| 3360 |   {
 | |
| 53339 | 3361 | assume "affine_dependent ((\<lambda>x. a + x) ` S)" | 
| 53302 | 3362 | then have "affine_dependent S" | 
| 53339 | 3363 | using affine_dependent_translation[of "((\<lambda>x. a + x) ` S)" "-a"] translation_assoc[of "-a" a] | 
| 53302 | 3364 | by auto | 
| 3365 | } | |
| 3366 | then show ?thesis | |
| 3367 | using affine_dependent_translation by auto | |
| 40377 | 3368 | qed | 
| 3369 | ||
| 3370 | lemma affine_hull_0_dependent: | |
| 53339 | 3371 | assumes "0 \<in> affine hull S" | 
| 40377 | 3372 | shows "dependent S" | 
| 53302 | 3373 | proof - | 
| 64267 | 3374 |   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 | 3375 | using assms affine_hull_explicit[of S] by auto | 
| 53339 | 3376 | then have "\<exists>v\<in>s. u v \<noteq> 0" | 
| 64267 | 3377 | using sum_not_0[of "u" "s"] by auto | 
| 53339 | 3378 | 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 | 3379 | using s_u by auto | 
| 53302 | 3380 | then show ?thesis | 
| 3381 | unfolding dependent_explicit[of S] by auto | |
| 40377 | 3382 | qed | 
| 3383 | ||
| 3384 | lemma affine_dependent_imp_dependent2: | |
| 3385 | assumes "affine_dependent (insert 0 S)" | |
| 3386 | shows "dependent S" | |
| 53302 | 3387 | proof - | 
| 54465 | 3388 |   obtain x where x: "x \<in> insert 0 S \<and> x \<in> affine hull (insert 0 S - {x})"
 | 
| 53302 | 3389 | using affine_dependent_def[of "(insert 0 S)"] assms by blast | 
| 3390 |   then have "x \<in> span (insert 0 S - {x})"
 | |
| 3391 | using affine_hull_subset_span by auto | |
| 3392 |   moreover have "span (insert 0 S - {x}) = span (S - {x})"
 | |
| 3393 |     using insert_Diff_if[of "0" S "{x}"] span_insert_0[of "S-{x}"] by auto
 | |
| 3394 |   ultimately have "x \<in> span (S - {x})" by auto
 | |
| 3395 | then have "x \<noteq> 0 \<Longrightarrow> dependent S" | |
| 54465 | 3396 | using x dependent_def by auto | 
| 53302 | 3397 | moreover | 
| 3398 |   {
 | |
| 3399 | assume "x = 0" | |
| 3400 | then have "0 \<in> affine hull S" | |
| 54465 | 3401 |       using x hull_mono[of "S - {0}" S] by auto
 | 
| 53302 | 3402 | then have "dependent S" | 
| 3403 | using affine_hull_0_dependent by auto | |
| 3404 | } | |
| 3405 | ultimately show ?thesis by auto | |
| 40377 | 3406 | qed | 
| 3407 | ||
| 3408 | lemma affine_dependent_iff_dependent: | |
| 53302 | 3409 | assumes "a \<notin> S" | 
| 3410 | shows "affine_dependent (insert a S) \<longleftrightarrow> dependent ((\<lambda>x. -a + x) ` S)" | |
| 3411 | proof - | |
| 3412 |   have "(op + (- a) ` S) = {x - a| x . x : S}" by auto
 | |
| 3413 | then show ?thesis | |
| 3414 | using affine_dependent_translation_eq[of "(insert a S)" "-a"] | |
| 49531 | 3415 | affine_dependent_imp_dependent2 assms | 
| 53302 | 3416 | dependent_imp_affine_dependent[of a S] | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53676diff
changeset | 3417 | by (auto simp del: uminus_add_conv_diff) | 
| 40377 | 3418 | qed | 
| 3419 | ||
| 3420 | lemma affine_dependent_iff_dependent2: | |
| 53339 | 3421 | assumes "a \<in> S" | 
| 3422 |   shows "affine_dependent S \<longleftrightarrow> dependent ((\<lambda>x. -a + x) ` (S-{a}))"
 | |
| 53302 | 3423 | proof - | 
| 53339 | 3424 |   have "insert a (S - {a}) = S"
 | 
| 53302 | 3425 | using assms by auto | 
| 3426 | then show ?thesis | |
| 3427 |     using assms affine_dependent_iff_dependent[of a "S-{a}"] by auto
 | |
| 40377 | 3428 | qed | 
| 3429 | ||
| 3430 | lemma affine_hull_insert_span_gen: | |
| 53339 | 3431 | "affine hull (insert a s) = (\<lambda>x. a + x) ` span ((\<lambda>x. - a + x) ` s)" | 
| 53302 | 3432 | proof - | 
| 53339 | 3433 |   have h1: "{x - a |x. x \<in> s} = ((\<lambda>x. -a+x) ` s)"
 | 
| 53302 | 3434 | by auto | 
| 3435 |   {
 | |
| 3436 | assume "a \<notin> s" | |
| 3437 | then have ?thesis | |
| 3438 | using affine_hull_insert_span[of a s] h1 by auto | |
| 3439 | } | |
| 3440 | moreover | |
| 3441 |   {
 | |
| 3442 | assume a1: "a \<in> s" | |
| 53339 | 3443 | have "\<exists>x. x \<in> s \<and> -a+x=0" | 
| 53302 | 3444 | apply (rule exI[of _ a]) | 
| 3445 | using a1 | |
| 3446 | apply auto | |
| 3447 | done | |
| 53339 | 3448 |     then have "insert 0 ((\<lambda>x. -a+x) ` (s - {a})) = (\<lambda>x. -a+x) ` s"
 | 
| 53302 | 3449 | by auto | 
| 53339 | 3450 |     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 | 3451 |       using span_insert_0[of "op + (- a) ` (s - {a})"] by (auto simp del: uminus_add_conv_diff)
 | 
| 53339 | 3452 |     moreover have "{x - a |x. x \<in> (s - {a})} = ((\<lambda>x. -a+x) ` (s - {a}))"
 | 
| 53302 | 3453 | by auto | 
| 53339 | 3454 |     moreover have "insert a (s - {a}) = insert a s"
 | 
| 63092 | 3455 | by auto | 
| 53302 | 3456 | ultimately have ?thesis | 
| 63092 | 3457 |       using affine_hull_insert_span[of "a" "s-{a}"] by auto
 | 
| 53302 | 3458 | } | 
| 3459 | ultimately show ?thesis by auto | |
| 40377 | 3460 | qed | 
| 3461 | ||
| 3462 | lemma affine_hull_span2: | |
| 53302 | 3463 | assumes "a \<in> s" | 
| 3464 |   shows "affine hull s = (\<lambda>x. a+x) ` span ((\<lambda>x. -a+x) ` (s-{a}))"
 | |
| 3465 |   using affine_hull_insert_span_gen[of a "s - {a}", unfolded insert_Diff[OF assms]]
 | |
| 3466 | by auto | |
| 40377 | 3467 | |
| 3468 | lemma affine_hull_span_gen: | |
| 53339 | 3469 | assumes "a \<in> affine hull s" | 
| 3470 | shows "affine hull s = (\<lambda>x. a+x) ` span ((\<lambda>x. -a+x) ` s)" | |
| 53302 | 3471 | proof - | 
| 3472 | have "affine hull (insert a s) = affine hull s" | |
| 3473 | using hull_redundant[of a affine s] assms by auto | |
| 3474 | then show ?thesis | |
| 3475 | using affine_hull_insert_span_gen[of a "s"] by auto | |
| 40377 | 3476 | qed | 
| 3477 | ||
| 3478 | lemma affine_hull_span_0: | |
| 53339 | 3479 | assumes "0 \<in> affine hull S" | 
| 40377 | 3480 | shows "affine hull S = span S" | 
| 53302 | 3481 | using affine_hull_span_gen[of "0" S] assms by auto | 
| 40377 | 3482 | |
| 63016 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3483 | lemma extend_to_affine_basis_nonempty: | 
| 53339 | 3484 | fixes S V :: "'n::euclidean_space set" | 
| 3485 |   assumes "\<not> affine_dependent S" "S \<subseteq> V" "S \<noteq> {}"
 | |
| 3486 | shows "\<exists>T. \<not> affine_dependent T \<and> S \<subseteq> T \<and> T \<subseteq> V \<and> affine hull T = affine hull V" | |
| 53302 | 3487 | proof - | 
| 54465 | 3488 | obtain a where a: "a \<in> S" | 
| 53302 | 3489 | using assms by auto | 
| 53339 | 3490 |   then have h0: "independent  ((\<lambda>x. -a + x) ` (S-{a}))"
 | 
| 53302 | 3491 | using affine_dependent_iff_dependent2 assms by auto | 
| 54465 | 3492 | then obtain B where B: | 
| 53339 | 3493 |     "(\<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"
 | 
| 3494 |      using maximal_independent_subset_extend[of "(\<lambda>x. -a+x) ` (S-{a})" "(\<lambda>x. -a + x) ` V"] assms
 | |
| 53302 | 3495 | by blast | 
| 63040 | 3496 | define T where "T = (\<lambda>x. a+x) ` insert 0 B" | 
| 53339 | 3497 | then have "T = insert a ((\<lambda>x. a+x) ` B)" | 
| 3498 | by auto | |
| 3499 | then have "affine hull T = (\<lambda>x. a+x) ` span B" | |
| 3500 | using affine_hull_insert_span_gen[of a "((\<lambda>x. a+x) ` B)"] translation_assoc[of "-a" a B] | |
| 53302 | 3501 | by auto | 
| 53347 | 3502 | then have "V \<subseteq> affine hull T" | 
| 54465 | 3503 | using B assms translation_inverse_subset[of a V "span B"] | 
| 53302 | 3504 | by auto | 
| 53339 | 3505 | moreover have "T \<subseteq> V" | 
| 54465 | 3506 | using T_def B a assms by auto | 
| 53302 | 3507 | ultimately have "affine hull T = affine hull V" | 
| 44457 
d366fa5551ef
declare euclidean_simps [simp] at the point they are proved;
 huffman parents: 
44365diff
changeset | 3508 | by (metis Int_absorb1 Int_absorb2 hull_hull hull_mono) | 
| 53347 | 3509 | moreover have "S \<subseteq> T" | 
| 54465 | 3510 |     using T_def B translation_inverse_subset[of a "S-{a}" B]
 | 
| 53302 | 3511 | by auto | 
| 3512 | moreover have "\<not> affine_dependent T" | |
| 53339 | 3513 | using T_def affine_dependent_translation_eq[of "insert 0 B"] | 
| 54465 | 3514 | affine_dependent_imp_dependent2 B | 
| 53302 | 3515 | by auto | 
| 60420 | 3516 | ultimately show ?thesis using \<open>T \<subseteq> V\<close> by auto | 
| 40377 | 3517 | qed | 
| 3518 | ||
| 49531 | 3519 | lemma affine_basis_exists: | 
| 53339 | 3520 | fixes V :: "'n::euclidean_space set" | 
| 3521 | shows "\<exists>B. B \<subseteq> V \<and> \<not> affine_dependent B \<and> affine hull V = affine hull B" | |
| 53302 | 3522 | proof (cases "V = {}")
 | 
| 3523 | case True | |
| 3524 | then show ?thesis | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 3525 | using affine_independent_0 by auto | 
| 53302 | 3526 | next | 
| 3527 | case False | |
| 3528 | then obtain x where "x \<in> V" by auto | |
| 3529 | then show ?thesis | |
| 63016 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3530 |     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 | 3531 | by auto | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3532 | qed | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3533 | |
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3534 | proposition extend_to_affine_basis: | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3535 | fixes S V :: "'n::euclidean_space set" | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3536 | assumes "\<not> affine_dependent S" "S \<subseteq> V" | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3537 | 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 | 3538 | proof (cases "S = {}")
 | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3539 | case True then show ?thesis | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3540 | 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 | 3541 | next | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3542 | case False | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3543 | then show ?thesis by (metis assms extend_to_affine_basis_nonempty that) | 
| 53302 | 3544 | qed | 
| 3545 | ||
| 40377 | 3546 | |
| 60420 | 3547 | subsection \<open>Affine Dimension of a Set\<close> | 
| 40377 | 3548 | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3549 | definition aff_dim :: "('a::euclidean_space) set \<Rightarrow> int"
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3550 | where "aff_dim V = | 
| 53339 | 3551 | (SOME d :: int. | 
| 3552 | \<exists>B. affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> of_nat (card B) = d + 1)" | |
| 40377 | 3553 | |
| 3554 | lemma aff_dim_basis_exists: | |
| 49531 | 3555 |   fixes V :: "('n::euclidean_space) set"
 | 
| 53339 | 3556 | shows "\<exists>B. affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> of_nat (card B) = aff_dim V + 1" | 
| 53302 | 3557 | proof - | 
| 53347 | 3558 | obtain B where "\<not> affine_dependent B \<and> affine hull B = affine hull V" | 
| 53302 | 3559 | using affine_basis_exists[of V] by auto | 
| 3560 | then show ?thesis | |
| 53339 | 3561 | unfolding aff_dim_def | 
| 53347 | 3562 | 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 | 3563 | apply auto | 
| 53339 | 3564 | apply (rule exI[of _ "int (card B) - (1 :: int)"]) | 
| 53302 | 3565 | apply (rule exI[of _ "B"]) | 
| 3566 | apply auto | |
| 3567 | done | |
| 3568 | qed | |
| 3569 | ||
| 3570 | lemma affine_hull_nonempty: "S \<noteq> {} \<longleftrightarrow> affine hull S \<noteq> {}"
 | |
| 3571 | proof - | |
| 3572 |   have "S = {} \<Longrightarrow> affine hull S = {}"
 | |
| 3573 | using affine_hull_empty by auto | |
| 3574 |   moreover have "affine hull S = {} \<Longrightarrow> S = {}"
 | |
| 3575 | unfolding hull_def by auto | |
| 3576 | ultimately show ?thesis by blast | |
| 40377 | 3577 | qed | 
| 3578 | ||
| 3579 | lemma aff_dim_parallel_subspace_aux: | |
| 53347 | 3580 | fixes B :: "'n::euclidean_space set" | 
| 53302 | 3581 | assumes "\<not> affine_dependent B" "a \<in> B" | 
| 53339 | 3582 |   shows "finite B \<and> ((card B) - 1 = dim (span ((\<lambda>x. -a+x) ` (B-{a}))))"
 | 
| 53302 | 3583 | proof - | 
| 53339 | 3584 |   have "independent ((\<lambda>x. -a + x) ` (B-{a}))"
 | 
| 53302 | 3585 | using affine_dependent_iff_dependent2 assms by auto | 
| 53339 | 3586 |   then have fin: "dim (span ((\<lambda>x. -a+x) ` (B-{a}))) = card ((\<lambda>x. -a + x) ` (B-{a}))"
 | 
| 3587 |     "finite ((\<lambda>x. -a + x) ` (B - {a}))"
 | |
| 53347 | 3588 |     using indep_card_eq_dim_span[of "(\<lambda>x. -a+x) ` (B-{a})"] by auto
 | 
| 53302 | 3589 | show ?thesis | 
| 53339 | 3590 |   proof (cases "(\<lambda>x. -a + x) ` (B - {a}) = {}")
 | 
| 53302 | 3591 | case True | 
| 53339 | 3592 |     have "B = insert a ((\<lambda>x. a + x) ` (\<lambda>x. -a + x) ` (B - {a}))"
 | 
| 53302 | 3593 |       using translation_assoc[of "a" "-a" "(B - {a})"] assms by auto
 | 
| 53339 | 3594 |     then have "B = {a}" using True by auto
 | 
| 53302 | 3595 | then show ?thesis using assms fin by auto | 
| 3596 | next | |
| 3597 | case False | |
| 53339 | 3598 |     then have "card ((\<lambda>x. -a + x) ` (B - {a})) > 0"
 | 
| 53302 | 3599 | using fin by auto | 
| 53339 | 3600 |     moreover have h1: "card ((\<lambda>x. -a + x) ` (B-{a})) = card (B-{a})"
 | 
| 53302 | 3601 | apply (rule card_image) | 
| 3602 | using translate_inj_on | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53676diff
changeset | 3603 | apply (auto simp del: uminus_add_conv_diff) | 
| 53302 | 3604 | done | 
| 53339 | 3605 |     ultimately have "card (B-{a}) > 0" by auto
 | 
| 3606 |     then have *: "finite (B - {a})"
 | |
| 53302 | 3607 |       using card_gt_0_iff[of "(B - {a})"] by auto
 | 
| 53339 | 3608 |     then have "card (B - {a}) = card B - 1"
 | 
| 53302 | 3609 | using card_Diff_singleton assms by auto | 
| 3610 | with * show ?thesis using fin h1 by auto | |
| 3611 | qed | |
| 40377 | 3612 | qed | 
| 3613 | ||
| 3614 | lemma aff_dim_parallel_subspace: | |
| 53339 | 3615 | fixes V L :: "'n::euclidean_space set" | 
| 53302 | 3616 |   assumes "V \<noteq> {}"
 | 
| 53339 | 3617 | and "subspace L" | 
| 3618 | and "affine_parallel (affine hull V) L" | |
| 53302 | 3619 | shows "aff_dim V = int (dim L)" | 
| 3620 | proof - | |
| 53339 | 3621 | obtain B where | 
| 54465 | 3622 | B: "affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> int (card B) = aff_dim V + 1" | 
| 53302 | 3623 | using aff_dim_basis_exists by auto | 
| 3624 |   then have "B \<noteq> {}"
 | |
| 54465 | 3625 | using assms B affine_hull_nonempty[of V] affine_hull_nonempty[of B] | 
| 53302 | 3626 | by auto | 
| 54465 | 3627 | then obtain a where a: "a \<in> B" by auto | 
| 63040 | 3628 |   define Lb where "Lb = span ((\<lambda>x. -a+x) ` (B-{a}))"
 | 
| 40377 | 3629 | moreover have "affine_parallel (affine hull B) Lb" | 
| 54465 | 3630 | using Lb_def B assms affine_hull_span2[of a B] a | 
| 53339 | 3631 | affine_parallel_commut[of "Lb" "(affine hull B)"] | 
| 3632 | unfolding affine_parallel_def | |
| 3633 | by auto | |
| 53302 | 3634 | moreover have "subspace Lb" | 
| 3635 | using Lb_def subspace_span by auto | |
| 3636 |   moreover have "affine hull B \<noteq> {}"
 | |
| 54465 | 3637 | using assms B affine_hull_nonempty[of V] by auto | 
| 53302 | 3638 | ultimately have "L = Lb" | 
| 54465 | 3639 | using assms affine_parallel_subspace[of "affine hull B"] affine_affine_hull[of B] B | 
| 53302 | 3640 | by auto | 
| 53339 | 3641 | then have "dim L = dim Lb" | 
| 3642 | by auto | |
| 3643 | moreover have "card B - 1 = dim Lb" and "finite B" | |
| 54465 | 3644 | using Lb_def aff_dim_parallel_subspace_aux a B by auto | 
| 53302 | 3645 | ultimately show ?thesis | 
| 60420 | 3646 |     using B \<open>B \<noteq> {}\<close> card_gt_0_iff[of B] by auto
 | 
| 40377 | 3647 | qed | 
| 3648 | ||
| 3649 | lemma aff_independent_finite: | |
| 53339 | 3650 | fixes B :: "'n::euclidean_space set" | 
| 3651 | assumes "\<not> affine_dependent B" | |
| 53302 | 3652 | shows "finite B" | 
| 3653 | proof - | |
| 3654 |   {
 | |
| 3655 |     assume "B \<noteq> {}"
 | |
| 3656 | then obtain a where "a \<in> B" by auto | |
| 3657 | then have ?thesis | |
| 3658 | using aff_dim_parallel_subspace_aux assms by auto | |
| 3659 | } | |
| 3660 | then show ?thesis by auto | |
| 40377 | 3661 | qed | 
| 3662 | ||
| 3663 | lemma independent_finite: | |
| 53339 | 3664 | fixes B :: "'n::euclidean_space set" | 
| 53302 | 3665 | assumes "independent B" | 
| 3666 | shows "finite B" | |
| 3667 | using affine_dependent_imp_dependent[of B] aff_independent_finite[of B] assms | |
| 3668 | by auto | |
| 40377 | 3669 | |
| 3670 | lemma subspace_dim_equal: | |
| 53339 | 3671 |   assumes "subspace (S :: ('n::euclidean_space) set)"
 | 
| 3672 | and "subspace T" | |
| 3673 | and "S \<subseteq> T" | |
| 3674 | and "dim S \<ge> dim T" | |
| 53302 | 3675 | shows "S = T" | 
| 3676 | proof - | |
| 53347 | 3677 | obtain B where B: "B \<le> S" "independent B \<and> S \<subseteq> span B" "card B = dim S" | 
| 53339 | 3678 | using basis_exists[of S] by auto | 
| 3679 | then have "span B \<subseteq> S" | |
| 3680 | using span_mono[of B S] span_eq[of S] assms by metis | |
| 3681 | then have "span B = S" | |
| 53347 | 3682 | using B by auto | 
| 53339 | 3683 | have "dim S = dim T" | 
| 3684 | using assms dim_subset[of S T] by auto | |
| 3685 | then have "T \<subseteq> span B" | |
| 53347 | 3686 | using card_eq_dim[of B T] B independent_finite assms by auto | 
| 53339 | 3687 | then show ?thesis | 
| 60420 | 3688 | using assms \<open>span B = S\<close> by auto | 
| 40377 | 3689 | qed | 
| 3690 | ||
| 63016 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3691 | corollary dim_eq_span: | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3692 | fixes S :: "'a::euclidean_space set" | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3693 | 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 | 3694 | 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 | 3695 | |
| 63075 
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
 paulson <lp15@cam.ac.uk> parents: 
63072diff
changeset | 3696 | lemma dim_eq_full: | 
| 
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
 paulson <lp15@cam.ac.uk> parents: 
63072diff
changeset | 3697 | fixes S :: "'a :: euclidean_space set" | 
| 
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
 paulson <lp15@cam.ac.uk> parents: 
63072diff
changeset | 3698 |     shows "dim S = DIM('a) \<longleftrightarrow> span S = UNIV"
 | 
| 
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
 paulson <lp15@cam.ac.uk> parents: 
63072diff
changeset | 3699 | apply (rule iffI) | 
| 
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
 paulson <lp15@cam.ac.uk> parents: 
63072diff
changeset | 3700 | 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 | 3701 | by (metis dim_UNIV dim_span) | 
| 
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
 paulson <lp15@cam.ac.uk> parents: 
63072diff
changeset | 3702 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 3703 | 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 | 3704 | assumes d: "d \<subseteq> Basis" | 
| 53347 | 3705 |   shows "span d = {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
 | 
| 3706 | (is "_ = ?B") | |
| 53339 | 3707 | proof - | 
| 3708 | have "d \<subseteq> ?B" | |
| 3709 | using d by (auto simp: inner_Basis) | |
| 3710 | moreover have s: "subspace ?B" | |
| 3711 | using subspace_substandard[of "\<lambda>i. i \<notin> d"] . | |
| 3712 | ultimately have "span d \<subseteq> ?B" | |
| 3713 | 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 | 3714 | moreover have *: "card d \<le> dim (span d)" | 
| 53339 | 3715 | using independent_card_le_dim[of d "span d"] independent_substdbasis[OF assms] span_inc[of d] | 
| 3716 | by auto | |
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53348diff
changeset | 3717 | moreover from * have "dim ?B \<le> dim (span d)" | 
| 53339 | 3718 | using dim_substandard[OF assms] by auto | 
| 3719 | ultimately show ?thesis | |
| 3720 | using s subspace_dim_equal[of "span d" "?B"] subspace_span[of d] by auto | |
| 40377 | 3721 | qed | 
| 3722 | ||
| 3723 | lemma basis_to_substdbasis_subspace_isomorphism: | |
| 53339 | 3724 | fixes B :: "'a::euclidean_space set" | 
| 3725 | assumes "independent B" | |
| 3726 | shows "\<exists>f d::'a set. card d = card B \<and> linear f \<and> f ` B = d \<and> | |
| 3727 |     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"
 | |
| 3728 | proof - | |
| 3729 | have B: "card B = dim B" | |
| 3730 | using dim_unique[of B B "card B"] assms span_inc[of B] by auto | |
| 3731 | have "dim B \<le> card (Basis :: 'a set)" | |
| 3732 | using dim_subset_UNIV[of B] by simp | |
| 3733 | from ex_card[OF this] obtain d :: "'a set" where d: "d \<subseteq> Basis" and t: "card d = dim B" | |
| 3734 | by auto | |
| 53347 | 3735 |   let ?t = "{x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
 | 
| 53339 | 3736 | 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 | 3737 | apply (rule basis_to_basis_subspace_isomorphism[of "span B" ?t B "d"]) | 
| 53339 | 3738 | apply (rule subspace_span) | 
| 3739 | apply (rule subspace_substandard) | |
| 3740 | defer | |
| 3741 | apply (rule span_inc) | |
| 3742 | apply (rule assms) | |
| 3743 | defer | |
| 3744 | unfolding dim_span[of B] | |
| 3745 | apply(rule B) | |
| 54465 | 3746 | unfolding span_substd_basis[OF d, symmetric] | 
| 53339 | 3747 | apply (rule span_inc) | 
| 3748 | apply (rule independent_substdbasis[OF d]) | |
| 3749 | apply rule | |
| 3750 | apply assumption | |
| 3751 | unfolding t[symmetric] span_substd_basis[OF d] dim_substandard[OF d] | |
| 3752 | apply auto | |
| 3753 | done | |
| 60420 | 3754 | with t \<open>card B = dim B\<close> d show ?thesis by auto | 
| 40377 | 3755 | qed | 
| 3756 | ||
| 3757 | lemma aff_dim_empty: | |
| 53339 | 3758 | fixes S :: "'n::euclidean_space set" | 
| 3759 |   shows "S = {} \<longleftrightarrow> aff_dim S = -1"
 | |
| 3760 | proof - | |
| 3761 | obtain B where *: "affine hull B = affine hull S" | |
| 3762 | and "\<not> affine_dependent B" | |
| 3763 | and "int (card B) = aff_dim S + 1" | |
| 3764 | using aff_dim_basis_exists by auto | |
| 3765 | moreover | |
| 3766 |   from * have "S = {} \<longleftrightarrow> B = {}"
 | |
| 3767 | using affine_hull_nonempty[of B] affine_hull_nonempty[of S] by auto | |
| 3768 | ultimately show ?thesis | |
| 3769 | using aff_independent_finite[of B] card_gt_0_iff[of B] by auto | |
| 3770 | qed | |
| 3771 | ||
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3772 | 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 | 3773 | by (simp add: aff_dim_empty [symmetric]) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3774 | |
| 63075 
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
 paulson <lp15@cam.ac.uk> parents: 
63072diff
changeset | 3775 | lemma aff_dim_affine_hull [simp]: "aff_dim (affine hull S) = aff_dim S" | 
| 53339 | 3776 | unfolding aff_dim_def using hull_hull[of _ S] by auto | 
| 40377 | 3777 | |
| 3778 | lemma aff_dim_affine_hull2: | |
| 53339 | 3779 | assumes "affine hull S = affine hull T" | 
| 3780 | shows "aff_dim S = aff_dim T" | |
| 3781 | unfolding aff_dim_def using assms by auto | |
| 40377 | 3782 | |
| 49531 | 3783 | lemma aff_dim_unique: | 
| 53339 | 3784 | fixes B V :: "'n::euclidean_space set" | 
| 3785 | assumes "affine hull B = affine hull V \<and> \<not> affine_dependent B" | |
| 3786 | shows "of_nat (card B) = aff_dim V + 1" | |
| 3787 | proof (cases "B = {}")
 | |
| 3788 | case True | |
| 3789 |   then have "V = {}"
 | |
| 3790 | using affine_hull_nonempty[of V] affine_hull_nonempty[of B] assms | |
| 3791 | by auto | |
| 3792 | then have "aff_dim V = (-1::int)" | |
| 3793 | using aff_dim_empty by auto | |
| 3794 | then show ?thesis | |
| 60420 | 3795 |     using \<open>B = {}\<close> by auto
 | 
| 53339 | 3796 | next | 
| 3797 | case False | |
| 54465 | 3798 | then obtain a where a: "a \<in> B" by auto | 
| 63040 | 3799 |   define Lb where "Lb = span ((\<lambda>x. -a+x) ` (B-{a}))"
 | 
| 40377 | 3800 | have "affine_parallel (affine hull B) Lb" | 
| 54465 | 3801 | using Lb_def affine_hull_span2[of a B] a | 
| 53339 | 3802 | affine_parallel_commut[of "Lb" "(affine hull B)"] | 
| 3803 | unfolding affine_parallel_def by auto | |
| 3804 | moreover have "subspace Lb" | |
| 3805 | using Lb_def subspace_span by auto | |
| 3806 | ultimately have "aff_dim B = int(dim Lb)" | |
| 60420 | 3807 |     using aff_dim_parallel_subspace[of B Lb] \<open>B \<noteq> {}\<close> by auto
 | 
| 53339 | 3808 | moreover have "(card B) - 1 = dim Lb" "finite B" | 
| 54465 | 3809 | using Lb_def aff_dim_parallel_subspace_aux a assms by auto | 
| 53339 | 3810 | ultimately have "of_nat (card B) = aff_dim B + 1" | 
| 60420 | 3811 |     using \<open>B \<noteq> {}\<close> card_gt_0_iff[of B] by auto
 | 
| 53339 | 3812 | then show ?thesis | 
| 3813 | using aff_dim_affine_hull2 assms by auto | |
| 40377 | 3814 | qed | 
| 3815 | ||
| 49531 | 3816 | lemma aff_dim_affine_independent: | 
| 53339 | 3817 | fixes B :: "'n::euclidean_space set" | 
| 3818 | assumes "\<not> affine_dependent B" | |
| 3819 | shows "of_nat (card B) = aff_dim B + 1" | |
| 40377 | 3820 | using aff_dim_unique[of B B] assms by auto | 
| 3821 | ||
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3822 | lemma affine_independent_iff_card: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3823 | fixes s :: "'a::euclidean_space set" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3824 | 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 | 3825 | apply (rule iffI) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3826 | apply (simp add: aff_dim_affine_independent aff_independent_finite) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3827 | 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 | 3828 | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 3829 | lemma aff_dim_sing [simp]: | 
| 53339 | 3830 | fixes a :: "'n::euclidean_space" | 
| 3831 |   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 | 3832 |   using aff_dim_affine_independent[of "{a}"] affine_independent_1 by auto
 | 
| 40377 | 3833 | |
| 63881 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3834 | 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 | 3835 | proof (clarsimp) | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3836 | assume "a \<noteq> b" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3837 |   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 | 3838 | 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 | 3839 | also have "... = 1" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3840 | 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 | 3841 |   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 | 3842 | qed | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3843 | |
| 40377 | 3844 | lemma aff_dim_inner_basis_exists: | 
| 49531 | 3845 |   fixes V :: "('n::euclidean_space) set"
 | 
| 53339 | 3846 | shows "\<exists>B. B \<subseteq> V \<and> affine hull B = affine hull V \<and> | 
| 3847 | \<not> affine_dependent B \<and> of_nat (card B) = aff_dim V + 1" | |
| 3848 | proof - | |
| 53347 | 3849 | obtain B where B: "\<not> affine_dependent B" "B \<subseteq> V" "affine hull B = affine hull V" | 
| 53339 | 3850 | using affine_basis_exists[of V] by auto | 
| 3851 | then have "of_nat(card B) = aff_dim V+1" using aff_dim_unique by auto | |
| 53347 | 3852 | with B show ?thesis by auto | 
| 40377 | 3853 | qed | 
| 3854 | ||
| 3855 | lemma aff_dim_le_card: | |
| 53347 | 3856 | fixes V :: "'n::euclidean_space set" | 
| 53339 | 3857 | assumes "finite V" | 
| 53347 | 3858 | shows "aff_dim V \<le> of_nat (card V) - 1" | 
| 53339 | 3859 | proof - | 
| 53347 | 3860 | obtain B where B: "B \<subseteq> V" "of_nat (card B) = aff_dim V + 1" | 
| 53339 | 3861 | using aff_dim_inner_basis_exists[of V] by auto | 
| 3862 | then have "card B \<le> card V" | |
| 3863 | using assms card_mono by auto | |
| 53347 | 3864 | with B show ?thesis by auto | 
| 40377 | 3865 | qed | 
| 3866 | ||
| 3867 | lemma aff_dim_parallel_eq: | |
| 53339 | 3868 | fixes S T :: "'n::euclidean_space set" | 
| 3869 | assumes "affine_parallel (affine hull S) (affine hull T)" | |
| 3870 | shows "aff_dim S = aff_dim T" | |
| 3871 | proof - | |
| 3872 |   {
 | |
| 3873 |     assume "T \<noteq> {}" "S \<noteq> {}"
 | |
| 53347 | 3874 | then obtain L where L: "subspace L \<and> affine_parallel (affine hull T) L" | 
| 3875 | using affine_parallel_subspace[of "affine hull T"] | |
| 3876 | affine_affine_hull[of T] affine_hull_nonempty | |
| 53339 | 3877 | by auto | 
| 3878 | then have "aff_dim T = int (dim L)" | |
| 60420 | 3879 |       using aff_dim_parallel_subspace \<open>T \<noteq> {}\<close> by auto
 | 
| 53339 | 3880 | moreover have *: "subspace L \<and> affine_parallel (affine hull S) L" | 
| 53347 | 3881 | using L affine_parallel_assoc[of "affine hull S" "affine hull T" L] assms by auto | 
| 53339 | 3882 | moreover from * have "aff_dim S = int (dim L)" | 
| 60420 | 3883 |       using aff_dim_parallel_subspace \<open>S \<noteq> {}\<close> by auto
 | 
| 53339 | 3884 | ultimately have ?thesis by auto | 
| 3885 | } | |
| 3886 | moreover | |
| 3887 |   {
 | |
| 3888 |     assume "S = {}"
 | |
| 3889 |     then have "S = {}" and "T = {}"
 | |
| 3890 | using assms affine_hull_nonempty | |
| 3891 | unfolding affine_parallel_def | |
| 3892 | by auto | |
| 3893 | then have ?thesis using aff_dim_empty by auto | |
| 3894 | } | |
| 3895 | moreover | |
| 3896 |   {
 | |
| 3897 |     assume "T = {}"
 | |
| 3898 |     then have "S = {}" and "T = {}"
 | |
| 3899 | using assms affine_hull_nonempty | |
| 3900 | unfolding affine_parallel_def | |
| 3901 | by auto | |
| 3902 | then have ?thesis | |
| 3903 | using aff_dim_empty by auto | |
| 3904 | } | |
| 3905 | ultimately show ?thesis by blast | |
| 40377 | 3906 | qed | 
| 3907 | ||
| 3908 | lemma aff_dim_translation_eq: | |
| 53339 | 3909 | fixes a :: "'n::euclidean_space" | 
| 3910 | shows "aff_dim ((\<lambda>x. a + x) ` S) = aff_dim S" | |
| 3911 | proof - | |
| 53347 | 3912 | have "affine_parallel (affine hull S) (affine hull ((\<lambda>x. a + x) ` S))" | 
| 53339 | 3913 | unfolding affine_parallel_def | 
| 3914 | apply (rule exI[of _ "a"]) | |
| 3915 | using affine_hull_translation[of a S] | |
| 3916 | apply auto | |
| 3917 | done | |
| 3918 | then show ?thesis | |
| 3919 | using aff_dim_parallel_eq[of S "(\<lambda>x. a + x) ` S"] by auto | |
| 40377 | 3920 | qed | 
| 3921 | ||
| 3922 | lemma aff_dim_affine: | |
| 53339 | 3923 | fixes S L :: "'n::euclidean_space set" | 
| 3924 |   assumes "S \<noteq> {}"
 | |
| 3925 | and "affine S" | |
| 3926 | and "subspace L" | |
| 3927 | and "affine_parallel S L" | |
| 3928 | shows "aff_dim S = int (dim L)" | |
| 3929 | proof - | |
| 3930 | have *: "affine hull S = S" | |
| 3931 | using assms affine_hull_eq[of S] by auto | |
| 3932 | then have "affine_parallel (affine hull S) L" | |
| 3933 | using assms by (simp add: *) | |
| 3934 | then show ?thesis | |
| 3935 | using assms aff_dim_parallel_subspace[of S L] by blast | |
| 40377 | 3936 | qed | 
| 3937 | ||
| 3938 | lemma dim_affine_hull: | |
| 53339 | 3939 | fixes S :: "'n::euclidean_space set" | 
| 3940 | shows "dim (affine hull S) = dim S" | |
| 3941 | proof - | |
| 3942 | have "dim (affine hull S) \<ge> dim S" | |
| 3943 | using dim_subset by auto | |
| 3944 | moreover have "dim (span S) \<ge> dim (affine hull S)" | |
| 60303 | 3945 | using dim_subset affine_hull_subset_span by blast | 
| 53339 | 3946 | moreover have "dim (span S) = dim S" | 
| 3947 | using dim_span by auto | |
| 3948 | ultimately show ?thesis by auto | |
| 40377 | 3949 | qed | 
| 3950 | ||
| 3951 | lemma aff_dim_subspace: | |
| 53339 | 3952 | fixes S :: "'n::euclidean_space set" | 
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 3953 | assumes "subspace S" | 
| 53339 | 3954 | shows "aff_dim S = int (dim S)" | 
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 3955 | proof (cases "S={}")
 | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 3956 | case True with assms show ?thesis | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 3957 | by (simp add: subspace_affine) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 3958 | next | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 3959 | case False | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 3960 | 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 | 3961 | show ?thesis by auto | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 3962 | qed | 
| 40377 | 3963 | |
| 3964 | lemma aff_dim_zero: | |
| 53339 | 3965 | fixes S :: "'n::euclidean_space set" | 
| 3966 | assumes "0 \<in> affine hull S" | |
| 3967 | shows "aff_dim S = int (dim S)" | |
| 3968 | proof - | |
| 3969 | have "subspace (affine hull S)" | |
| 3970 | using subspace_affine[of "affine hull S"] affine_affine_hull assms | |
| 3971 | by auto | |
| 3972 | then have "aff_dim (affine hull S) = int (dim (affine hull S))" | |
| 3973 | using assms aff_dim_subspace[of "affine hull S"] by auto | |
| 3974 | then show ?thesis | |
| 3975 | using aff_dim_affine_hull[of S] dim_affine_hull[of S] | |
| 3976 | by auto | |
| 40377 | 3977 | qed | 
| 3978 | ||
| 63016 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3979 | lemma aff_dim_eq_dim: | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3980 | fixes S :: "'n::euclidean_space set" | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3981 | assumes "a \<in> affine hull S" | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3982 | 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 | 3983 | proof - | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3984 | 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 | 3985 | unfolding Convex_Euclidean_Space.affine_hull_translation | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3986 | 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 | 3987 | with aff_dim_zero show ?thesis | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3988 | by (metis aff_dim_translation_eq) | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3989 | qed | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3990 | |
| 63072 | 3991 | lemma aff_dim_UNIV [simp]: "aff_dim (UNIV :: 'n::euclidean_space set) = int(DIM('n))"
 | 
| 53347 | 3992 | using aff_dim_subspace[of "(UNIV :: 'n::euclidean_space set)"] | 
| 53339 | 3993 | dim_UNIV[where 'a="'n::euclidean_space"] | 
| 3994 | by auto | |
| 40377 | 3995 | |
| 3996 | lemma aff_dim_geq: | |
| 53339 | 3997 | fixes V :: "'n::euclidean_space set" | 
| 3998 | shows "aff_dim V \<ge> -1" | |
| 3999 | proof - | |
| 53347 | 4000 | obtain B where "affine hull B = affine hull V" | 
| 4001 | and "\<not> affine_dependent B" | |
| 4002 | and "int (card B) = aff_dim V + 1" | |
| 53339 | 4003 | using aff_dim_basis_exists by auto | 
| 4004 | then show ?thesis by auto | |
| 40377 | 4005 | qed | 
| 4006 | ||
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4007 | 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 | 4008 | 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 | 4009 |   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 | 4010 | 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 | 4011 | |
| 66641 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4012 | lemma aff_lowdim_subset_hyperplane: | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4013 | fixes S :: "'a::euclidean_space set" | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4014 |   assumes "aff_dim S < DIM('a)"
 | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4015 |   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 | 4016 | proof (cases "S={}")
 | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4017 | case True | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4018 | moreover | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4019 | have "(SOME b. b \<in> Basis) \<noteq> 0" | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4020 | 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 | 4021 | ultimately show ?thesis | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4022 | using that by blast | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4023 | next | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4024 | case False | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4025 | 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 | 4026 | by (meson equals0I mk_disjoint_insert) | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4027 |   have "dim (op + (-c) ` S) < DIM('a)"
 | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4028 | 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 | 4029 |   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 | 4030 | using lowdim_subset_hyperplane by blast | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4031 | moreover | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4032 |   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 | 4033 | proof - | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4034 | have "w-c \<in> span (op + (- c) ` S)" | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4035 | 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 | 4036 |     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 | 4037 | by blast | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4038 | then show ?thesis | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4039 | by (auto simp: algebra_simps) | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4040 | qed | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4041 |   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 | 4042 | by blast | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4043 | then show ?thesis | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4044 | 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 | 4045 | qed | 
| 
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
 paulson <lp15@cam.ac.uk> parents: 
66453diff
changeset | 4046 | |
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4047 | lemma affine_independent_card_dim_diffs: | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4048 | fixes S :: "'a :: euclidean_space set" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4049 | assumes "~ affine_dependent S" "a \<in> S" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4050 |     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 | 4051 | proof - | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4052 |   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 | 4053 |   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 | 4054 | proof (cases "x = a") | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4055 | case True then show ?thesis by simp | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4056 | next | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4057 | case False then show ?thesis | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4058 | using assms by (blast intro: span_superset that) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4059 | qed | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4060 | have "\<not> affine_dependent (insert a S)" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4061 | by (simp add: assms insert_absorb) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4062 |   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 | 4063 | using dependent_imp_affine_dependent by fastforce | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4064 |   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 | 4065 | by blast | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4066 |   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 | 4067 | by simp | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4068 |   also have "... = card (S - {a})"
 | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4069 | 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 | 4070 | also have "... = card S - 1" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4071 | by (simp add: aff_independent_finite assms) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4072 |   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 | 4073 | have "finite S" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4074 | by (meson assms aff_independent_finite) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4075 | 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 | 4076 |   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 | 4077 | 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 | 4078 | ultimately show ?thesis | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4079 | by auto | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4080 | qed | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4081 | |
| 49531 | 4082 | lemma independent_card_le_aff_dim: | 
| 53347 | 4083 | fixes B :: "'n::euclidean_space set" | 
| 4084 | assumes "B \<subseteq> V" | |
| 53339 | 4085 | assumes "\<not> affine_dependent B" | 
| 4086 | 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 | 4087 | proof - | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 4088 | 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 | 4089 | by (metis assms extend_to_affine_basis[of B V]) | 
| 53339 | 4090 | then have "of_nat (card T) = aff_dim V + 1" | 
| 4091 | using aff_dim_unique by auto | |
| 4092 | then show ?thesis | |
| 53347 | 4093 | using T card_mono[of T B] aff_independent_finite[of T] by auto | 
| 40377 | 4094 | qed | 
| 4095 | ||
| 4096 | lemma aff_dim_subset: | |
| 53347 | 4097 | fixes S T :: "'n::euclidean_space set" | 
| 4098 | assumes "S \<subseteq> T" | |
| 4099 | shows "aff_dim S \<le> aff_dim T" | |
| 53339 | 4100 | proof - | 
| 53347 | 4101 | obtain B where B: "\<not> affine_dependent B" "B \<subseteq> S" "affine hull B = affine hull S" | 
| 4102 | "of_nat (card B) = aff_dim S + 1" | |
| 53339 | 4103 | using aff_dim_inner_basis_exists[of S] by auto | 
| 4104 | then have "int (card B) \<le> aff_dim T + 1" | |
| 4105 | using assms independent_card_le_aff_dim[of B T] by auto | |
| 53347 | 4106 | with B show ?thesis by auto | 
| 40377 | 4107 | qed | 
| 4108 | ||
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4109 | lemma aff_dim_le_DIM: | 
| 53339 | 4110 | fixes S :: "'n::euclidean_space set" | 
| 4111 |   shows "aff_dim S \<le> int (DIM('n))"
 | |
| 49531 | 4112 | proof - | 
| 53339 | 4113 |   have "aff_dim (UNIV :: 'n::euclidean_space set) = int(DIM('n))"
 | 
| 63072 | 4114 | using aff_dim_UNIV by auto | 
| 53339 | 4115 |   then show "aff_dim (S:: 'n::euclidean_space set) \<le> int(DIM('n))"
 | 
| 63092 | 4116 |     using aff_dim_subset[of S "(UNIV :: ('n::euclidean_space) set)"] subset_UNIV by auto
 | 
| 40377 | 4117 | qed | 
| 4118 | ||
| 4119 | lemma affine_dim_equal: | |
| 53347 | 4120 | fixes S :: "'n::euclidean_space set" | 
| 4121 |   assumes "affine S" "affine T" "S \<noteq> {}" "S \<subseteq> T" "aff_dim S = aff_dim T"
 | |
| 4122 | shows "S = T" | |
| 4123 | proof - | |
| 4124 | obtain a where "a \<in> S" using assms by auto | |
| 4125 | then have "a \<in> T" using assms by auto | |
| 63040 | 4126 |   define LS where "LS = {y. \<exists>x \<in> S. (-a) + x = y}"
 | 
| 53347 | 4127 | then have ls: "subspace LS" "affine_parallel S LS" | 
| 60420 | 4128 | using assms parallel_subspace_explicit[of S a LS] \<open>a \<in> S\<close> by auto | 
| 53347 | 4129 | then have h1: "int(dim LS) = aff_dim S" | 
| 4130 | using assms aff_dim_affine[of S LS] by auto | |
| 4131 |   have "T \<noteq> {}" using assms by auto
 | |
| 63040 | 4132 |   define LT where "LT = {y. \<exists>x \<in> T. (-a) + x = y}"
 | 
| 53347 | 4133 | then have lt: "subspace LT \<and> affine_parallel T LT" | 
| 60420 | 4134 | using assms parallel_subspace_explicit[of T a LT] \<open>a \<in> T\<close> by auto | 
| 53347 | 4135 | then have "int(dim LT) = aff_dim T" | 
| 60420 | 4136 |     using assms aff_dim_affine[of T LT] \<open>T \<noteq> {}\<close> by auto
 | 
| 53347 | 4137 | then have "dim LS = dim LT" | 
| 4138 | using h1 assms by auto | |
| 4139 | moreover have "LS \<le> LT" | |
| 4140 | using LS_def LT_def assms by auto | |
| 4141 | ultimately have "LS = LT" | |
| 4142 | using subspace_dim_equal[of LS LT] ls lt by auto | |
| 4143 |   moreover have "S = {x. \<exists>y \<in> LS. a+y=x}"
 | |
| 4144 | using LS_def by auto | |
| 4145 |   moreover have "T = {x. \<exists>y \<in> LT. a+y=x}"
 | |
| 4146 | using LT_def by auto | |
| 4147 | ultimately show ?thesis by auto | |
| 40377 | 4148 | qed | 
| 4149 | ||
| 63881 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4150 | lemma aff_dim_eq_0: | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4151 | fixes S :: "'a::euclidean_space set" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4152 |   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 | 4153 | proof (cases "S = {}")
 | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4154 | case True | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4155 | then show ?thesis | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4156 | by auto | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4157 | next | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4158 | case False | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4159 | 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 | 4160 | show ?thesis | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4161 | proof safe | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4162 | assume 0: "aff_dim S = 0" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4163 |     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 | 4164 | 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 | 4165 |     then show "\<exists>a. S = {a}"
 | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4166 | 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 | 4167 | qed auto | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4168 | qed | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4169 | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4170 | lemma affine_hull_UNIV: | 
| 53347 | 4171 | fixes S :: "'n::euclidean_space set" | 
| 4172 |   assumes "aff_dim S = int(DIM('n))"
 | |
| 4173 |   shows "affine hull S = (UNIV :: ('n::euclidean_space) set)"
 | |
| 4174 | proof - | |
| 4175 |   have "S \<noteq> {}"
 | |
| 4176 | using assms aff_dim_empty[of S] by auto | |
| 4177 | have h0: "S \<subseteq> affine hull S" | |
| 4178 | using hull_subset[of S _] by auto | |
| 4179 |   have h1: "aff_dim (UNIV :: ('n::euclidean_space) set) = aff_dim S"
 | |
| 63072 | 4180 | using aff_dim_UNIV assms by auto | 
| 53347 | 4181 |   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 | 4182 | using aff_dim_le_DIM[of "affine hull S"] assms h0 by auto | 
| 53347 | 4183 | have h3: "aff_dim S \<le> aff_dim (affine hull S)" | 
| 4184 | using h0 aff_dim_subset[of S "affine hull S"] assms by auto | |
| 4185 |   then have h4: "aff_dim (affine hull S) = aff_dim (UNIV :: ('n::euclidean_space) set)"
 | |
| 4186 | using h0 h1 h2 by auto | |
| 4187 | then show ?thesis | |
| 4188 |     using affine_dim_equal[of "affine hull S" "(UNIV :: ('n::euclidean_space) set)"]
 | |
| 60420 | 4189 |       affine_affine_hull[of S] affine_UNIV assms h4 h0 \<open>S \<noteq> {}\<close>
 | 
| 53347 | 4190 | by auto | 
| 40377 | 4191 | qed | 
| 4192 | ||
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4193 | lemma disjoint_affine_hull: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4194 | 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 | 4195 |   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 | 4196 |     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 | 4197 | proof - | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4198 | 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 | 4199 | 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 | 4200 |   { fix y
 | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4201 | 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 | 4202 | then obtain a b | 
| 64267 | 4203 | where a1 [simp]: "sum a t = 1" and [simp]: "sum (\<lambda>v. a v *\<^sub>R v) t = y" | 
| 4204 | 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 | 4205 | by (auto simp: affine_hull_finite \<open>finite t\<close> \<open>finite u\<close>) | 
| 63040 | 4206 | 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 | 4207 | have [simp]: "s \<inter> t = t" "s \<inter> - t \<inter> u = u" using assms by auto | 
| 64267 | 4208 | have "sum c s = 0" | 
| 4209 | 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 | 4210 | moreover have "~ (\<forall>v\<in>s. c v = 0)" | 
| 64267 | 4211 | 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 | 4212 | moreover have "(\<Sum>v\<in>s. c v *\<^sub>R v) = 0" | 
| 64267 | 4213 | by (simp add: c_def if_smult sum_negf | 
| 4214 | 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 | 4215 | ultimately have False | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4216 | 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 | 4217 | } | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4218 | then show ?thesis by blast | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4219 | qed | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4220 | |
| 40377 | 4221 | lemma aff_dim_convex_hull: | 
| 53347 | 4222 | fixes S :: "'n::euclidean_space set" | 
| 4223 | shows "aff_dim (convex hull S) = aff_dim S" | |
| 49531 | 4224 | using aff_dim_affine_hull[of S] convex_hull_subset_affine_hull[of S] | 
| 53347 | 4225 | hull_subset[of S "convex"] aff_dim_subset[of S "convex hull S"] | 
| 4226 | aff_dim_subset[of "convex hull S" "affine hull S"] | |
| 4227 | by auto | |
| 40377 | 4228 | |
| 4229 | lemma aff_dim_cball: | |
| 53347 | 4230 | fixes a :: "'n::euclidean_space" | 
| 4231 | assumes "e > 0" | |
| 4232 |   shows "aff_dim (cball a e) = int (DIM('n))"
 | |
| 4233 | proof - | |
| 4234 | have "(\<lambda>x. a + x) ` (cball 0 e) \<subseteq> cball a e" | |
| 4235 | unfolding cball_def dist_norm by auto | |
| 4236 | then have "aff_dim (cball (0 :: 'n::euclidean_space) e) \<le> aff_dim (cball a e)" | |
| 4237 | using aff_dim_translation_eq[of a "cball 0 e"] | |
| 4238 | aff_dim_subset[of "op + a ` cball 0 e" "cball a e"] | |
| 4239 | by auto | |
| 4240 |   moreover have "aff_dim (cball (0 :: 'n::euclidean_space) e) = int (DIM('n))"
 | |
| 4241 | using hull_inc[of "(0 :: 'n::euclidean_space)" "cball 0 e"] | |
| 4242 | centre_in_cball[of "(0 :: 'n::euclidean_space)"] assms | |
| 4243 | by (simp add: dim_cball[of e] aff_dim_zero[of "cball 0 e"]) | |
| 4244 | ultimately show ?thesis | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4245 | using aff_dim_le_DIM[of "cball a e"] by auto | 
| 40377 | 4246 | qed | 
| 4247 | ||
| 4248 | lemma aff_dim_open: | |
| 53347 | 4249 | fixes S :: "'n::euclidean_space set" | 
| 4250 | assumes "open S" | |
| 4251 |     and "S \<noteq> {}"
 | |
| 4252 |   shows "aff_dim S = int (DIM('n))"
 | |
| 4253 | proof - | |
| 4254 | obtain x where "x \<in> S" | |
| 4255 | using assms by auto | |
| 4256 | then obtain e where e: "e > 0" "cball x e \<subseteq> S" | |
| 4257 | using open_contains_cball[of S] assms by auto | |
| 4258 | then have "aff_dim (cball x e) \<le> aff_dim S" | |
| 4259 | using aff_dim_subset by auto | |
| 4260 | with e show ?thesis | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4261 | using aff_dim_cball[of e x] aff_dim_le_DIM[of S] by auto | 
| 40377 | 4262 | qed | 
| 4263 | ||
| 4264 | lemma low_dim_interior: | |
| 53347 | 4265 | fixes S :: "'n::euclidean_space set" | 
| 4266 |   assumes "\<not> aff_dim S = int (DIM('n))"
 | |
| 4267 |   shows "interior S = {}"
 | |
| 4268 | proof - | |
| 4269 | have "aff_dim(interior S) \<le> aff_dim S" | |
| 4270 | using interior_subset aff_dim_subset[of "interior S" S] by auto | |
| 4271 | then show ?thesis | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4272 | using aff_dim_open[of "interior S"] aff_dim_le_DIM[of S] assms by auto | 
| 40377 | 4273 | qed | 
| 4274 | ||
| 60307 
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
 paulson <lp15@cam.ac.uk> parents: 
60303diff
changeset | 4275 | corollary empty_interior_lowdim: | 
| 
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
 paulson <lp15@cam.ac.uk> parents: 
60303diff
changeset | 4276 | 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 | 4277 |   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 | 4278 | 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 | 4279 | |
| 63016 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 4280 | corollary aff_dim_nonempty_interior: | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 4281 | fixes S :: "'a::euclidean_space set" | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 4282 |   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 | 4283 | by (metis low_dim_interior) | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 4284 | |
| 63881 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 4285 | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4286 | subsection \<open>Caratheodory's theorem.\<close> | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4287 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4288 | lemma convex_hull_caratheodory_aff_dim: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4289 |   fixes p :: "('a::euclidean_space) set"
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4290 | shows "convex hull p = | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4291 |     {y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> aff_dim p + 1 \<and>
 | 
| 64267 | 4292 | (\<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 | 4293 | unfolding convex_hull_explicit set_eq_iff mem_Collect_eq | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4294 | proof (intro allI iffI) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4295 | fix y | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4296 | 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 | 4297 | sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y" | 
| 4298 | 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 | 4299 | then obtain N where "?P N" by auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4300 | 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 | 4301 | apply (rule_tac ex_least_nat_le) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4302 | apply auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4303 | done | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4304 | 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 | 4305 | by blast | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4306 | then obtain s u where obt: "finite s" "card s = n" "s\<subseteq>p" "\<forall>x\<in>s. 0 \<le> u x" | 
| 64267 | 4307 | "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 | 4308 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4309 | have "card s \<le> aff_dim p + 1" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4310 | proof (rule ccontr, simp only: not_le) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4311 | assume "aff_dim p + 1 < card s" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4312 | then have "affine_dependent s" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4313 | 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 | 4314 | by blast | 
| 64267 | 4315 | 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 | 4316 | using affine_dependent_explicit_finite[OF obt(1)] by auto | 
| 63040 | 4317 |     define i where "i = (\<lambda>v. (u v) / (- w v)) ` {v\<in>s. w v < 0}"
 | 
| 4318 | define t where "t = Min i" | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4319 | have "\<exists>x\<in>s. w x < 0" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4320 | proof (rule ccontr, simp add: not_less) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4321 | assume as:"\<forall>x\<in>s. 0 \<le> w x" | 
| 64267 | 4322 |       then have "sum w (s - {v}) \<ge> 0"
 | 
| 4323 | apply (rule_tac sum_nonneg) | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4324 | apply auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4325 | done | 
| 64267 | 4326 | then have "sum w s > 0" | 
| 4327 | 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 | 4328 | 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 | 4329 | then show False using wv(1) by auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4330 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4331 |     then have "i \<noteq> {}" unfolding i_def by auto
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4332 | then have "t \<ge> 0" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4333 | using Min_ge_iff[of i 0 ] and obt(1) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4334 | unfolding t_def i_def | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4335 | using obt(4)[unfolded le_less] | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4336 | by (auto simp: divide_le_0_iff) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4337 | 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 | 4338 | proof | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4339 | fix v | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4340 | assume "v \<in> s" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4341 | then have v: "0 \<le> u v" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4342 | using obt(4)[THEN bspec[where x=v]] by auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4343 | show "0 \<le> u v + t * w v" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4344 | proof (cases "w v < 0") | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4345 | case False | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4346 | 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 | 4347 | next | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4348 | case True | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4349 | then have "t \<le> u v / (- w v)" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4350 | 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 | 4351 | apply (rule_tac Min_le) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4352 | using obt(1) apply auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4353 | done | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4354 | then show ?thesis | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4355 | unfolding real_0_le_add_iff | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4356 | 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 | 4357 | by auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4358 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4359 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4360 | 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 | 4361 |       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 | 4362 | then have a: "a \<in> s" "u a + t * w a = 0" by auto | 
| 64267 | 4363 |     have *: "\<And>f. sum f (s - {a}) = sum f s - ((f a)::'b::ab_group_add)"
 | 
| 4364 | 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 | 4365 | have "(\<Sum>v\<in>s. u v + t * w v) = 1" | 
| 64267 | 4366 | 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 | 4367 | 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 | 4368 | 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 | 4369 | 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 | 4370 | ultimately have "?P (n - 1)" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4371 |       apply (rule_tac x="(s - {a})" in exI)
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4372 | 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 | 4373 | using obt(1-3) and t and a | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4374 | apply (auto simp add: * scaleR_left_distrib) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4375 | done | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4376 | then show False | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4377 | using smallest[THEN spec[where x="n - 1"]] by auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4378 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4379 | then show "\<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> aff_dim p + 1 \<and> | 
| 64267 | 4380 | (\<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 | 4381 | using obt by auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4382 | qed auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4383 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4384 | lemma caratheodory_aff_dim: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4385 |   fixes p :: "('a::euclidean_space) set"
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4386 |   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 | 4387 | (is "?lhs = ?rhs") | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4388 | proof | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4389 | show "?lhs \<subseteq> ?rhs" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4390 | apply (subst convex_hull_caratheodory_aff_dim) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4391 | apply clarify | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4392 | apply (rule_tac x="s" in exI) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4393 | 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 | 4394 | done | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4395 | next | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4396 | show "?rhs \<subseteq> ?lhs" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4397 | using hull_mono by blast | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4398 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4399 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4400 | lemma convex_hull_caratheodory: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4401 |   fixes p :: "('a::euclidean_space) set"
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4402 | shows "convex hull p = | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4403 |             {y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> DIM('a) + 1 \<and>
 | 
| 64267 | 4404 | (\<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 | 4405 | (is "?lhs = ?rhs") | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4406 | proof (intro set_eqI iffI) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4407 | fix x | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4408 | assume "x \<in> ?lhs" then show "x \<in> ?rhs" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4409 | 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 | 4410 | apply (erule ex_forward)+ | 
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4411 | using aff_dim_le_DIM [of p] | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4412 | apply simp | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4413 | done | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4414 | next | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4415 | fix x | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4416 | assume "x \<in> ?rhs" then show "x \<in> ?lhs" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4417 | by (auto simp add: convex_hull_explicit) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4418 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4419 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4420 | theorem caratheodory: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4421 | "convex hull p = | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4422 |     {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 | 4423 |       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 | 4424 | proof safe | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4425 | fix x | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4426 | assume "x \<in> convex hull p" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4427 |   then obtain s u where "finite s" "s \<subseteq> p" "card s \<le> DIM('a) + 1"
 | 
| 64267 | 4428 | "\<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 | 4429 | unfolding convex_hull_caratheodory by auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4430 |   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 | 4431 | apply (rule_tac x=s in exI) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4432 | using hull_subset[of s convex] | 
| 63170 | 4433 | using convex_convex_hull[simplified convex_explicit, of s, | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4434 | THEN spec[where x=s], THEN spec[where x=u]] | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4435 | apply auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4436 | done | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4437 | next | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4438 | fix x s | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4439 |   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 | 4440 | then show "x \<in> convex hull p" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4441 | 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 | 4442 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4443 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 4444 | |
| 60420 | 4445 | subsection \<open>Relative interior of a set\<close> | 
| 40377 | 4446 | |
| 53347 | 4447 | definition "rel_interior S = | 
| 4448 |   {x. \<exists>T. openin (subtopology euclidean (affine hull S)) T \<and> x \<in> T \<and> T \<subseteq> S}"
 | |
| 4449 | ||
| 64287 | 4450 | lemma rel_interior_mono: | 
| 4451 | "\<lbrakk>S \<subseteq> T; affine hull S = affine hull T\<rbrakk> | |
| 4452 | \<Longrightarrow> (rel_interior S) \<subseteq> (rel_interior T)" | |
| 4453 | by (auto simp: rel_interior_def) | |
| 4454 | ||
| 4455 | lemma rel_interior_maximal: | |
| 4456 | "\<lbrakk>T \<subseteq> S; openin(subtopology euclidean (affine hull S)) T\<rbrakk> \<Longrightarrow> T \<subseteq> (rel_interior S)" | |
| 4457 | by (auto simp: rel_interior_def) | |
| 4458 | ||
| 53347 | 4459 | lemma rel_interior: | 
| 4460 |   "rel_interior S = {x \<in> S. \<exists>T. open T \<and> x \<in> T \<and> T \<inter> affine hull S \<subseteq> S}"
 | |
| 4461 | unfolding rel_interior_def[of S] openin_open[of "affine hull S"] | |
| 4462 | apply auto | |
| 4463 | proof - | |
| 4464 | fix x T | |
| 4465 | assume *: "x \<in> S" "open T" "x \<in> T" "T \<inter> affine hull S \<subseteq> S" | |
| 4466 | then have **: "x \<in> T \<inter> affine hull S" | |
| 4467 | using hull_inc by auto | |
| 54465 | 4468 | show "\<exists>Tb. (\<exists>Ta. open Ta \<and> Tb = affine hull S \<inter> Ta) \<and> x \<in> Tb \<and> Tb \<subseteq> S" | 
| 4469 | apply (rule_tac x = "T \<inter> (affine hull S)" in exI) | |
| 53347 | 4470 | using * ** | 
| 4471 | apply auto | |
| 4472 | done | |
| 4473 | qed | |
| 4474 | ||
| 4475 | 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)" | |
| 4476 | by (auto simp add: rel_interior) | |
| 4477 | ||
| 4478 | lemma mem_rel_interior_ball: | |
| 4479 | "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 | 4480 | apply (simp add: rel_interior, safe) | 
| 4481 | apply (force simp add: open_contains_ball) | |
| 53347 | 4482 | 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 | 4483 | apply simp | 
| 40377 | 4484 | done | 
| 4485 | ||
| 49531 | 4486 | lemma rel_interior_ball: | 
| 53347 | 4487 |   "rel_interior S = {x \<in> S. \<exists>e. e > 0 \<and> ball x e \<inter> affine hull S \<subseteq> S}"
 | 
| 4488 | using mem_rel_interior_ball [of _ S] by auto | |
| 4489 | ||
| 4490 | lemma mem_rel_interior_cball: | |
| 4491 | "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 | 4492 | apply (simp add: rel_interior, safe) | 
| 40377 | 4493 | apply (force simp add: open_contains_cball) | 
| 53347 | 4494 | 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 | 4495 | apply (simp add: subset_trans [OF ball_subset_cball]) | 
| 40377 | 4496 | apply auto | 
| 4497 | done | |
| 4498 | ||
| 53347 | 4499 | lemma rel_interior_cball: | 
| 4500 |   "rel_interior S = {x \<in> S. \<exists>e. e > 0 \<and> cball x e \<inter> affine hull S \<subseteq> S}"
 | |
| 4501 | using mem_rel_interior_cball [of _ S] by auto | |
| 40377 | 4502 | |
| 60303 | 4503 | lemma rel_interior_empty [simp]: "rel_interior {} = {}"
 | 
| 49531 | 4504 | by (auto simp add: rel_interior_def) | 
| 40377 | 4505 | |
| 60303 | 4506 | lemma affine_hull_sing [simp]: "affine hull {a :: 'n::euclidean_space} = {a}"
 | 
| 53347 | 4507 | by (metis affine_hull_eq affine_sing) | 
| 40377 | 4508 | |
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4509 | lemma rel_interior_sing [simp]: | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4510 |     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 | 4511 | apply (auto simp: rel_interior_ball) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4512 | apply (rule_tac x=1 in exI) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4513 | apply force | 
| 53347 | 4514 | done | 
| 40377 | 4515 | |
| 4516 | lemma subset_rel_interior: | |
| 53347 | 4517 | fixes S T :: "'n::euclidean_space set" | 
| 4518 | assumes "S \<subseteq> T" | |
| 4519 | and "affine hull S = affine hull T" | |
| 4520 | shows "rel_interior S \<subseteq> rel_interior T" | |
| 49531 | 4521 | using assms by (auto simp add: rel_interior_def) | 
| 4522 | ||
| 53347 | 4523 | lemma rel_interior_subset: "rel_interior S \<subseteq> S" | 
| 4524 | by (auto simp add: rel_interior_def) | |
| 4525 | ||
| 4526 | lemma rel_interior_subset_closure: "rel_interior S \<subseteq> closure S" | |
| 4527 | using rel_interior_subset by (auto simp add: closure_def) | |
| 4528 | ||
| 4529 | lemma interior_subset_rel_interior: "interior S \<subseteq> rel_interior S" | |
| 4530 | by (auto simp add: rel_interior interior_def) | |
| 40377 | 4531 | |
| 4532 | lemma interior_rel_interior: | |
| 53347 | 4533 | fixes S :: "'n::euclidean_space set" | 
| 4534 |   assumes "aff_dim S = int(DIM('n))"
 | |
| 4535 | shows "rel_interior S = interior S" | |
| 40377 | 4536 | proof - | 
| 53347 | 4537 | 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 | 4538 | using assms affine_hull_UNIV[of S] by auto | 
| 53347 | 4539 | then show ?thesis | 
| 4540 | unfolding rel_interior interior_def by auto | |
| 40377 | 4541 | qed | 
| 4542 | ||
| 60303 | 4543 | lemma rel_interior_interior: | 
| 4544 | fixes S :: "'n::euclidean_space set" | |
| 4545 | assumes "affine hull S = UNIV" | |
| 4546 | shows "rel_interior S = interior S" | |
| 4547 | using assms unfolding rel_interior interior_def by auto | |
| 4548 | ||
| 40377 | 4549 | lemma rel_interior_open: | 
| 53347 | 4550 | fixes S :: "'n::euclidean_space set" | 
| 4551 | assumes "open S" | |
| 4552 | shows "rel_interior S = S" | |
| 4553 | by (metis assms interior_eq interior_subset_rel_interior rel_interior_subset set_eq_subset) | |
| 40377 | 4554 | |
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 4555 | 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 | 4556 | by (simp add: interior_open) | 
| 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 4557 | |
| 40377 | 4558 | lemma interior_rel_interior_gen: | 
| 53347 | 4559 | fixes S :: "'n::euclidean_space set" | 
| 4560 |   shows "interior S = (if aff_dim S = int(DIM('n)) then rel_interior S else {})"
 | |
| 4561 | by (metis interior_rel_interior low_dim_interior) | |
| 40377 | 4562 | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4563 | lemma rel_interior_nonempty_interior: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4564 | 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 | 4565 |   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 | 4566 | 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 | 4567 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4568 | lemma affine_hull_nonempty_interior: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4569 | 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 | 4570 |   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 | 4571 | 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 | 4572 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4573 | lemma rel_interior_affine_hull [simp]: | 
| 53347 | 4574 | fixes S :: "'n::euclidean_space set" | 
| 4575 | shows "rel_interior (affine hull S) = affine hull S" | |
| 4576 | proof - | |
| 4577 | have *: "rel_interior (affine hull S) \<subseteq> affine hull S" | |
| 4578 | using rel_interior_subset by auto | |
| 4579 |   {
 | |
| 4580 | fix x | |
| 4581 | assume x: "x \<in> affine hull S" | |
| 63040 | 4582 | define e :: real where "e = 1" | 
| 53347 | 4583 | then have "e > 0" "ball x e \<inter> affine hull (affine hull S) \<subseteq> affine hull S" | 
| 4584 | using hull_hull[of _ S] by auto | |
| 4585 | then have "x \<in> rel_interior (affine hull S)" | |
| 4586 | using x rel_interior_ball[of "affine hull S"] by auto | |
| 4587 | } | |
| 4588 | then show ?thesis using * by auto | |
| 40377 | 4589 | qed | 
| 4590 | ||
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4591 | lemma rel_interior_UNIV [simp]: "rel_interior (UNIV :: ('n::euclidean_space) set) = UNIV"
 | 
| 53347 | 4592 | by (metis open_UNIV rel_interior_open) | 
| 40377 | 4593 | |
| 4594 | lemma rel_interior_convex_shrink: | |
| 53347 | 4595 | fixes S :: "'a::euclidean_space set" | 
| 4596 | assumes "convex S" | |
| 4597 | and "c \<in> rel_interior S" | |
| 4598 | and "x \<in> S" | |
| 4599 | and "0 < e" | |
| 4600 | and "e \<le> 1" | |
| 4601 | shows "x - e *\<^sub>R (x - c) \<in> rel_interior S" | |
| 4602 | proof - | |
| 54465 | 4603 | obtain d where "d > 0" and d: "ball c d \<inter> affine hull S \<subseteq> S" | 
| 53347 | 4604 | using assms(2) unfolding mem_rel_interior_ball by auto | 
| 4605 |   {
 | |
| 4606 | fix y | |
| 4607 | assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d" "y \<in> affine hull S" | |
| 4608 | have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" | |
| 60420 | 4609 | using \<open>e > 0\<close> by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib) | 
| 53347 | 4610 | have "x \<in> affine hull S" | 
| 4611 | using assms hull_subset[of S] by auto | |
| 49531 | 4612 | moreover have "1 / e + - ((1 - e) / e) = 1" | 
| 60420 | 4613 | using \<open>e > 0\<close> left_diff_distrib[of "1" "(1-e)" "1/e"] by auto | 
| 53347 | 4614 | ultimately have **: "(1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x \<in> affine hull S" | 
| 4615 | using as affine_affine_hull[of S] mem_affine[of "affine hull S" y x "(1 / e)" "-((1 - e) / e)"] | |
| 4616 | by (simp add: algebra_simps) | |
| 61945 | 4617 | 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 | 4618 | unfolding dist_norm norm_scaleR[symmetric] | 
| 4619 | apply (rule arg_cong[where f=norm]) | |
| 60420 | 4620 | using \<open>e > 0\<close> | 
| 53347 | 4621 | apply (auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps) | 
| 4622 | done | |
| 61945 | 4623 | also have "\<dots> = \<bar>1/e\<bar> * norm (x - e *\<^sub>R (x - c) - y)" | 
| 53347 | 4624 | by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) | 
| 4625 | also have "\<dots> < d" | |
| 60420 | 4626 | using as[unfolded dist_norm] and \<open>e > 0\<close> | 
| 4627 | by (auto simp add:pos_divide_less_eq[OF \<open>e > 0\<close>] mult.commute) | |
| 53347 | 4628 | finally have "y \<in> S" | 
| 4629 | apply (subst *) | |
| 4630 | apply (rule assms(1)[unfolded convex_alt,rule_format]) | |
| 4631 | apply (rule d[unfolded subset_eq,rule_format]) | |
| 4632 | unfolding mem_ball | |
| 4633 | using assms(3-5) ** | |
| 4634 | apply auto | |
| 4635 | done | |
| 4636 | } | |
| 4637 | then have "ball (x - e *\<^sub>R (x - c)) (e*d) \<inter> affine hull S \<subseteq> S" | |
| 4638 | by auto | |
| 4639 | moreover have "e * d > 0" | |
| 60420 | 4640 | using \<open>e > 0\<close> \<open>d > 0\<close> by simp | 
| 53347 | 4641 | moreover have c: "c \<in> S" | 
| 4642 | using assms rel_interior_subset by auto | |
| 4643 | 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 | 4644 | using convexD_alt[of S x c e] | 
| 53347 | 4645 | apply (simp add: algebra_simps) | 
| 4646 | using assms | |
| 4647 | apply auto | |
| 4648 | done | |
| 4649 | ultimately show ?thesis | |
| 60420 | 4650 | using mem_rel_interior_ball[of "x - e *\<^sub>R (x - c)" S] \<open>e > 0\<close> by auto | 
| 40377 | 4651 | qed | 
| 4652 | ||
| 4653 | lemma interior_real_semiline: | |
| 53347 | 4654 | fixes a :: real | 
| 4655 |   shows "interior {a..} = {a<..}"
 | |
| 4656 | proof - | |
| 4657 |   {
 | |
| 4658 | fix y | |
| 4659 | assume "a < y" | |
| 4660 |     then have "y \<in> interior {a..}"
 | |
| 4661 | apply (simp add: mem_interior) | |
| 4662 | apply (rule_tac x="(y-a)" in exI) | |
| 4663 | apply (auto simp add: dist_norm) | |
| 4664 | done | |
| 4665 | } | |
| 4666 | moreover | |
| 4667 |   {
 | |
| 4668 | fix y | |
| 4669 |     assume "y \<in> interior {a..}"
 | |
| 4670 |     then obtain e where e: "e > 0" "cball y e \<subseteq> {a..}"
 | |
| 4671 |       using mem_interior_cball[of y "{a..}"] by auto
 | |
| 4672 | moreover from e have "y - e \<in> cball y e" | |
| 4673 | 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 | 4674 | ultimately have "a \<le> y - e" by blast | 
| 53347 | 4675 | then have "a < y" using e by auto | 
| 4676 | } | |
| 4677 | ultimately show ?thesis by auto | |
| 40377 | 4678 | qed | 
| 4679 | ||
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4680 | lemma continuous_ge_on_Ioo: | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4681 |   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 | 4682 | shows "g (x::real) \<ge> (a::real)" | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4683 | proof- | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4684 |   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 | 4685 |   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 | 4686 |   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 | 4687 |   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 | 4688 | by (auto simp: continuous_on_closed_vimage) | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4689 |   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 | 4690 |   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 | 4691 | qed | 
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4692 | |
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4693 | lemma interior_real_semiline': | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4694 | fixes a :: real | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4695 |   shows "interior {..a} = {..<a}"
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4696 | proof - | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4697 |   {
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4698 | fix y | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4699 | assume "a > y" | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4700 |     then have "y \<in> interior {..a}"
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4701 | apply (simp add: mem_interior) | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4702 | apply (rule_tac x="(a-y)" in exI) | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4703 | apply (auto simp add: dist_norm) | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4704 | done | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4705 | } | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4706 | moreover | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4707 |   {
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4708 | fix y | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4709 |     assume "y \<in> interior {..a}"
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4710 |     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 | 4711 |       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 | 4712 | 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 | 4713 | by (auto simp add: cball_def dist_norm) | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4714 | ultimately have "a \<ge> y + e" by auto | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4715 | then have "a > y" using e by auto | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4716 | } | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4717 | ultimately show ?thesis by auto | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4718 | qed | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4719 | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 4720 | 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 | 4721 | proof- | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4722 |   have "{a..b} = {a..} \<inter> {..b}" by auto
 | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61952diff
changeset | 4723 |   also have "interior ... = {a<..} \<inter> {..<b}"
 | 
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4724 | 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 | 4725 |   also have "... = {a<..<b}" by auto
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4726 | finally show ?thesis . | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4727 | qed | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4728 | |
| 66793 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 4729 | lemma interior_atLeastLessThan [simp]: | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 4730 |   fixes a::real shows "interior {a..<b} = {a<..<b}"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 4731 | 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 | 4732 | |
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 4733 | lemma interior_lessThanAtMost [simp]: | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 4734 |   fixes a::real shows "interior {a<..b} = {a<..<b}"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 4735 | 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 | 4736 | interior_interior interior_real_semiline) | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 4737 | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 4738 | 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 | 4739 | by (metis interior_atLeastAtMost_real interior_interior) | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 4740 | |
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 4741 | lemma frontier_real_Iic [simp]: | 
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4742 | fixes a :: real | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4743 |   shows "frontier {..a} = {a}"
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 4744 | 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 | 4745 | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 4746 | lemma rel_interior_real_box [simp]: | 
| 53347 | 4747 | fixes a b :: real | 
| 4748 | assumes "a < b" | |
| 56188 | 4749 |   shows "rel_interior {a .. b} = {a <..< b}"
 | 
| 53347 | 4750 | proof - | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54465diff
changeset | 4751 |   have "box a b \<noteq> {}"
 | 
| 53347 | 4752 | using assms | 
| 4753 | unfolding set_eq_iff | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 4754 | by (auto intro!: exI[of _ "(a + b) / 2"] simp: box_def) | 
| 40377 | 4755 | then show ?thesis | 
| 56188 | 4756 | using interior_rel_interior_gen[of "cbox a b", symmetric] | 
| 62390 | 4757 | by (simp split: if_split_asm del: box_real add: box_real[symmetric] interior_cbox) | 
| 40377 | 4758 | qed | 
| 4759 | ||
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 4760 | lemma rel_interior_real_semiline [simp]: | 
| 53347 | 4761 | fixes a :: real | 
| 4762 |   shows "rel_interior {a..} = {a<..}"
 | |
| 4763 | proof - | |
| 4764 |   have *: "{a<..} \<noteq> {}"
 | |
| 4765 | unfolding set_eq_iff by (auto intro!: exI[of _ "a + 1"]) | |
| 4766 |   then show ?thesis using interior_real_semiline interior_rel_interior_gen[of "{a..}"]
 | |
| 62390 | 4767 | by (auto split: if_split_asm) | 
| 40377 | 4768 | qed | 
| 4769 | ||
| 60420 | 4770 | subsubsection \<open>Relative open sets\<close> | 
| 40377 | 4771 | |
| 53347 | 4772 | definition "rel_open S \<longleftrightarrow> rel_interior S = S" | 
| 4773 | ||
| 4774 | lemma rel_open: "rel_open S \<longleftrightarrow> openin (subtopology euclidean (affine hull S)) S" | |
| 4775 | unfolding rel_open_def rel_interior_def | |
| 4776 | apply auto | |
| 4777 | using openin_subopen[of "subtopology euclidean (affine hull S)" S] | |
| 4778 | apply auto | |
| 4779 | done | |
| 4780 | ||
| 63072 | 4781 | lemma openin_rel_interior: "openin (subtopology euclidean (affine hull S)) (rel_interior S)" | 
| 40377 | 4782 | apply (simp add: rel_interior_def) | 
| 53347 | 4783 | apply (subst openin_subopen) | 
| 4784 | apply blast | |
| 4785 | done | |
| 40377 | 4786 | |
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4787 | lemma openin_set_rel_interior: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4788 | "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 | 4789 | 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 | 4790 | |
| 49531 | 4791 | lemma affine_rel_open: | 
| 53347 | 4792 | fixes S :: "'n::euclidean_space set" | 
| 4793 | assumes "affine S" | |
| 4794 | shows "rel_open S" | |
| 4795 | unfolding rel_open_def | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 4796 | using assms rel_interior_affine_hull[of S] affine_hull_eq[of S] | 
| 53347 | 4797 | by metis | 
| 40377 | 4798 | |
| 49531 | 4799 | lemma affine_closed: | 
| 53347 | 4800 | fixes S :: "'n::euclidean_space set" | 
| 4801 | assumes "affine S" | |
| 4802 | shows "closed S" | |
| 4803 | proof - | |
| 4804 |   {
 | |
| 4805 |     assume "S \<noteq> {}"
 | |
| 4806 | then obtain L where L: "subspace L" "affine_parallel S L" | |
| 4807 | using assms affine_parallel_subspace[of S] by auto | |
| 4808 | then obtain a where a: "S = (op + a ` L)" | |
| 4809 | using affine_parallel_def[of L S] affine_parallel_commut by auto | |
| 4810 | from L have "closed L" using closed_subspace by auto | |
| 4811 | then have "closed S" | |
| 4812 | using closed_translation a by auto | |
| 4813 | } | |
| 4814 | then show ?thesis by auto | |
| 40377 | 4815 | qed | 
| 4816 | ||
| 4817 | lemma closure_affine_hull: | |
| 53347 | 4818 | fixes S :: "'n::euclidean_space set" | 
| 4819 | shows "closure S \<subseteq> affine hull S" | |
| 44524 | 4820 | by (intro closure_minimal hull_subset affine_closed affine_affine_hull) | 
| 40377 | 4821 | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 4822 | lemma closure_same_affine_hull [simp]: | 
| 53347 | 4823 | fixes S :: "'n::euclidean_space set" | 
| 40377 | 4824 | shows "affine hull (closure S) = affine hull S" | 
| 53347 | 4825 | proof - | 
| 4826 | have "affine hull (closure S) \<subseteq> affine hull S" | |
| 4827 | using hull_mono[of "closure S" "affine hull S" "affine"] | |
| 4828 | closure_affine_hull[of S] hull_hull[of "affine" S] | |
| 4829 | by auto | |
| 4830 | moreover have "affine hull (closure S) \<supseteq> affine hull S" | |
| 4831 | using hull_mono[of "S" "closure S" "affine"] closure_subset by auto | |
| 4832 | ultimately show ?thesis by auto | |
| 49531 | 4833 | qed | 
| 4834 | ||
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 4835 | lemma closure_aff_dim [simp]: | 
| 53347 | 4836 | fixes S :: "'n::euclidean_space set" | 
| 40377 | 4837 | shows "aff_dim (closure S) = aff_dim S" | 
| 53347 | 4838 | proof - | 
| 4839 | have "aff_dim S \<le> aff_dim (closure S)" | |
| 4840 | using aff_dim_subset closure_subset by auto | |
| 4841 | 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 | 4842 | using aff_dim_subset closure_affine_hull by blast | 
| 53347 | 4843 | moreover have "aff_dim (affine hull S) = aff_dim S" | 
| 4844 | using aff_dim_affine_hull by auto | |
| 4845 | ultimately show ?thesis by auto | |
| 40377 | 4846 | qed | 
| 4847 | ||
| 4848 | lemma rel_interior_closure_convex_shrink: | |
| 53347 | 4849 | fixes S :: "_::euclidean_space set" | 
| 4850 | assumes "convex S" | |
| 4851 | and "c \<in> rel_interior S" | |
| 4852 | and "x \<in> closure S" | |
| 4853 | and "e > 0" | |
| 4854 | and "e \<le> 1" | |
| 4855 | shows "x - e *\<^sub>R (x - c) \<in> rel_interior S" | |
| 4856 | proof - | |
| 4857 | obtain d where "d > 0" and d: "ball c d \<inter> affine hull S \<subseteq> S" | |
| 4858 | using assms(2) unfolding mem_rel_interior_ball by auto | |
| 4859 | have "\<exists>y \<in> S. norm (y - x) * (1 - e) < e * d" | |
| 4860 | proof (cases "x \<in> S") | |
| 4861 | case True | |
| 60420 | 4862 | then show ?thesis using \<open>e > 0\<close> \<open>d > 0\<close> | 
| 53347 | 4863 | apply (rule_tac bexI[where x=x]) | 
| 56544 | 4864 | apply (auto) | 
| 53347 | 4865 | done | 
| 4866 | next | |
| 4867 | case False | |
| 4868 | then have x: "x islimpt S" | |
| 4869 | using assms(3)[unfolded closure_def] by auto | |
| 4870 | show ?thesis | |
| 4871 | proof (cases "e = 1") | |
| 4872 | case True | |
| 4873 | obtain y where "y \<in> S" "y \<noteq> x" "dist y x < 1" | |
| 40377 | 4874 | using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto | 
| 53347 | 4875 | then show ?thesis | 
| 4876 | apply (rule_tac x=y in bexI) | |
| 4877 | unfolding True | |
| 60420 | 4878 | using \<open>d > 0\<close> | 
| 53347 | 4879 | apply auto | 
| 4880 | done | |
| 4881 | next | |
| 4882 | case False | |
| 4883 | then have "0 < e * d / (1 - e)" and *: "1 - e > 0" | |
| 60420 | 4884 | using \<open>e \<le> 1\<close> \<open>e > 0\<close> \<open>d > 0\<close> by (auto) | 
| 53347 | 4885 | then obtain y where "y \<in> S" "y \<noteq> x" "dist y x < e * d / (1 - e)" | 
| 40377 | 4886 | using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto | 
| 53347 | 4887 | then show ?thesis | 
| 4888 | apply (rule_tac x=y in bexI) | |
| 4889 | unfolding dist_norm | |
| 4890 | using pos_less_divide_eq[OF *] | |
| 4891 | apply auto | |
| 4892 | done | |
| 4893 | qed | |
| 4894 | qed | |
| 4895 | then obtain y where "y \<in> S" and y: "norm (y - x) * (1 - e) < e * d" | |
| 4896 | by auto | |
| 63040 | 4897 | define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)" | 
| 53347 | 4898 | have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" | 
| 60420 | 4899 | unfolding z_def using \<open>e > 0\<close> | 
| 53347 | 4900 | by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) | 
| 4901 | have zball: "z \<in> ball c d" | |
| 4902 | using mem_ball z_def dist_norm[of c] | |
| 4903 | using y and assms(4,5) | |
| 4904 | by (auto simp add:field_simps norm_minus_commute) | |
| 4905 | have "x \<in> affine hull S" | |
| 4906 | using closure_affine_hull assms by auto | |
| 4907 | moreover have "y \<in> affine hull S" | |
| 60420 | 4908 | using \<open>y \<in> S\<close> hull_subset[of S] by auto | 
| 53347 | 4909 | moreover have "c \<in> affine hull S" | 
| 4910 | using assms rel_interior_subset hull_subset[of S] by auto | |
| 4911 | ultimately have "z \<in> affine hull S" | |
| 49531 | 4912 | using z_def affine_affine_hull[of S] | 
| 53347 | 4913 | mem_affine_3_minus [of "affine hull S" c x y "(1 - e) / e"] | 
| 4914 | assms | |
| 4915 | by (auto simp add: field_simps) | |
| 4916 | then have "z \<in> S" using d zball by auto | |
| 4917 | obtain d1 where "d1 > 0" and d1: "ball z d1 \<le> ball c d" | |
| 40377 | 4918 | using zball open_ball[of c d] openE[of "ball c d" z] by auto | 
| 53347 | 4919 | then have "ball z d1 \<inter> affine hull S \<subseteq> ball c d \<inter> affine hull S" | 
| 4920 | by auto | |
| 4921 | then have "ball z d1 \<inter> affine hull S \<subseteq> S" | |
| 4922 | using d by auto | |
| 4923 | then have "z \<in> rel_interior S" | |
| 60420 | 4924 | using mem_rel_interior_ball using \<open>d1 > 0\<close> \<open>z \<in> S\<close> by auto | 
| 53347 | 4925 | then have "y - e *\<^sub>R (y - z) \<in> rel_interior S" | 
| 60420 | 4926 | using rel_interior_convex_shrink[of S z y e] assms \<open>y \<in> S\<close> by auto | 
| 53347 | 4927 | then show ?thesis using * by auto | 
| 4928 | qed | |
| 4929 | ||
| 62620 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 4930 | lemma rel_interior_eq: | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 4931 | "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 | 4932 | 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 | 4933 | |
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 4934 | lemma rel_interior_openin: | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 4935 | "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 | 4936 | 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 | 4937 | |
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4938 | lemma rel_interior_affine: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4939 | 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 | 4940 | 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 | 4941 | 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 | 4942 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4943 | lemma rel_interior_eq_closure: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4944 | 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 | 4945 | 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 | 4946 | proof (cases "S = {}")
 | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4947 | case True | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4948 | then show ?thesis | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4949 | by auto | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4950 | next | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4951 | case False show ?thesis | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4952 | proof | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4953 | 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 | 4954 |     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 | 4955 | 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 | 4956 | 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 | 4957 | apply (rule conjI) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4958 | 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 | 4959 | 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 | 4960 | done | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4961 | 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 | 4962 | by auto | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4963 | then show "affine S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4964 | by (metis affine_hull_eq) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4965 | next | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4966 | assume "affine S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4967 | 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 | 4968 | 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 | 4969 | qed | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4970 | qed | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 4971 | |
| 40377 | 4972 | |
| 60420 | 4973 | subsubsection\<open>Relative interior preserves under linear transformations\<close> | 
| 40377 | 4974 | |
| 4975 | lemma rel_interior_translation_aux: | |
| 53347 | 4976 | fixes a :: "'n::euclidean_space" | 
| 4977 | shows "((\<lambda>x. a + x) ` rel_interior S) \<subseteq> rel_interior ((\<lambda>x. a + x) ` S)" | |
| 4978 | proof - | |
| 4979 |   {
 | |
| 4980 | fix x | |
| 4981 | assume x: "x \<in> rel_interior S" | |
| 4982 | then obtain T where "open T" "x \<in> T \<inter> S" "T \<inter> affine hull S \<subseteq> S" | |
| 4983 | using mem_rel_interior[of x S] by auto | |
| 4984 | then have "open ((\<lambda>x. a + x) ` T)" | |
| 4985 | and "a + x \<in> ((\<lambda>x. a + x) ` T) \<inter> ((\<lambda>x. a + x) ` S)" | |
| 4986 | and "((\<lambda>x. a + x) ` T) \<inter> affine hull ((\<lambda>x. a + x) ` S) \<subseteq> (\<lambda>x. a + x) ` S" | |
| 4987 | using affine_hull_translation[of a S] open_translation[of T a] x by auto | |
| 4988 | then have "a + x \<in> rel_interior ((\<lambda>x. a + x) ` S)" | |
| 4989 | using mem_rel_interior[of "a+x" "((\<lambda>x. a + x) ` S)"] by auto | |
| 4990 | } | |
| 4991 | then show ?thesis by auto | |
| 60809 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60800diff
changeset | 4992 | qed | 
| 40377 | 4993 | |
| 4994 | lemma rel_interior_translation: | |
| 53347 | 4995 | fixes a :: "'n::euclidean_space" | 
| 4996 | shows "rel_interior ((\<lambda>x. a + x) ` S) = (\<lambda>x. a + x) ` rel_interior S" | |
| 4997 | proof - | |
| 4998 | have "(\<lambda>x. (-a) + x) ` rel_interior ((\<lambda>x. a + x) ` S) \<subseteq> rel_interior S" | |
| 4999 | using rel_interior_translation_aux[of "-a" "(\<lambda>x. a + x) ` S"] | |
| 5000 | translation_assoc[of "-a" "a"] | |
| 5001 | by auto | |
| 5002 | then have "((\<lambda>x. a + x) ` rel_interior S) \<supseteq> rel_interior ((\<lambda>x. a + x) ` S)" | |
| 5003 | using translation_inverse_subset[of a "rel_interior (op + a ` S)" "rel_interior S"] | |
| 5004 | by auto | |
| 5005 | then show ?thesis | |
| 5006 | using rel_interior_translation_aux[of a S] by auto | |
| 40377 | 5007 | qed | 
| 5008 | ||
| 5009 | ||
| 5010 | lemma affine_hull_linear_image: | |
| 53347 | 5011 | assumes "bounded_linear f" | 
| 5012 | shows "f ` (affine hull s) = affine hull f ` s" | |
| 5013 | apply rule | |
| 5014 | unfolding subset_eq ball_simps | |
| 5015 | apply (rule_tac[!] hull_induct, rule hull_inc) | |
| 5016 | prefer 3 | |
| 5017 | apply (erule imageE) | |
| 5018 | apply (rule_tac x=xa in image_eqI) | |
| 5019 | apply assumption | |
| 5020 | apply (rule hull_subset[unfolded subset_eq, rule_format]) | |
| 5021 | apply assumption | |
| 5022 | proof - | |
| 40377 | 5023 | interpret f: bounded_linear f by fact | 
| 53347 | 5024 |   show "affine {x. f x \<in> affine hull f ` s}"
 | 
| 5025 | unfolding affine_def | |
| 5026 | by (auto simp add: f.scaleR f.add affine_affine_hull[unfolded affine_def, rule_format]) | |
| 5027 |   show "affine {x. x \<in> f ` (affine hull s)}"
 | |
| 5028 | using affine_affine_hull[unfolded affine_def, of s] | |
| 40377 | 5029 | unfolding affine_def by (auto simp add: f.scaleR [symmetric] f.add [symmetric]) | 
| 5030 | qed auto | |
| 5031 | ||
| 5032 | ||
| 5033 | lemma rel_interior_injective_on_span_linear_image: | |
| 53347 | 5034 | fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space" | 
| 5035 | and S :: "'m::euclidean_space set" | |
| 5036 | assumes "bounded_linear f" | |
| 5037 | and "inj_on f (span S)" | |
| 5038 | shows "rel_interior (f ` S) = f ` (rel_interior S)" | |
| 5039 | proof - | |
| 5040 |   {
 | |
| 5041 | fix z | |
| 5042 | assume z: "z \<in> rel_interior (f ` S)" | |
| 5043 | then have "z \<in> f ` S" | |
| 5044 | using rel_interior_subset[of "f ` S"] by auto | |
| 5045 | then obtain x where x: "x \<in> S" "f x = z" by auto | |
| 5046 | obtain e2 where e2: "e2 > 0" "cball z e2 \<inter> affine hull (f ` S) \<subseteq> (f ` S)" | |
| 5047 | using z rel_interior_cball[of "f ` S"] by auto | |
| 5048 | obtain K where K: "K > 0" "\<And>x. norm (f x) \<le> norm x * K" | |
| 5049 | using assms Real_Vector_Spaces.bounded_linear.pos_bounded[of f] by auto | |
| 63040 | 5050 | define e1 where "e1 = 1 / K" | 
| 53347 | 5051 | then have e1: "e1 > 0" "\<And>x. e1 * norm (f x) \<le> norm x" | 
| 5052 | using K pos_le_divide_eq[of e1] by auto | |
| 63040 | 5053 | define e where "e = e1 * e2" | 
| 56544 | 5054 | then have "e > 0" using e1 e2 by auto | 
| 53347 | 5055 |     {
 | 
| 5056 | fix y | |
| 5057 | assume y: "y \<in> cball x e \<inter> affine hull S" | |
| 5058 | then have h1: "f y \<in> affine hull (f ` S)" | |
| 5059 | using affine_hull_linear_image[of f S] assms by auto | |
| 5060 | from y have "norm (x-y) \<le> e1 * e2" | |
| 5061 | using cball_def[of x e] dist_norm[of x y] e_def by auto | |
| 5062 | 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 | 5063 | using assms linear_diff[of f x y] linear_conv_bounded_linear[of f] by auto | 
| 53347 | 5064 | moreover have "e1 * norm (f (x-y)) \<le> norm (x - y)" | 
| 5065 | using e1 by auto | |
| 5066 | ultimately have "e1 * norm ((f x)-(f y)) \<le> e1 * e2" | |
| 5067 | by auto | |
| 5068 | then have "f y \<in> cball z e2" | |
| 5069 | using cball_def[of "f x" e2] dist_norm[of "f x" "f y"] e1 x by auto | |
| 5070 | then have "f y \<in> f ` S" | |
| 5071 | using y e2 h1 by auto | |
| 5072 | then have "y \<in> S" | |
| 5073 | 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 | 5074 | 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 | 5075 | by (metis Int_iff span_inc subsetCE) | 
| 53347 | 5076 | } | 
| 5077 | then have "z \<in> f ` (rel_interior S)" | |
| 60420 | 5078 | using mem_rel_interior_cball[of x S] \<open>e > 0\<close> x by auto | 
| 49531 | 5079 | } | 
| 53347 | 5080 | moreover | 
| 5081 |   {
 | |
| 5082 | fix x | |
| 5083 | assume x: "x \<in> rel_interior S" | |
| 54465 | 5084 | then obtain e2 where e2: "e2 > 0" "cball x e2 \<inter> affine hull S \<subseteq> S" | 
| 53347 | 5085 | using rel_interior_cball[of S] by auto | 
| 5086 | have "x \<in> S" using x rel_interior_subset by auto | |
| 5087 | then have *: "f x \<in> f ` S" by auto | |
| 5088 | have "\<forall>x\<in>span S. f x = 0 \<longrightarrow> x = 0" | |
| 5089 | using assms subspace_span linear_conv_bounded_linear[of f] | |
| 5090 | linear_injective_on_subspace_0[of f "span S"] | |
| 5091 | by auto | |
| 5092 | then obtain e1 where e1: "e1 > 0" "\<forall>x \<in> span S. e1 * norm x \<le> norm (f x)" | |
| 5093 | using assms injective_imp_isometric[of "span S" f] | |
| 5094 | subspace_span[of S] closed_subspace[of "span S"] | |
| 5095 | by auto | |
| 63040 | 5096 | define e where "e = e1 * e2" | 
| 56544 | 5097 | hence "e > 0" using e1 e2 by auto | 
| 53347 | 5098 |     {
 | 
| 5099 | fix y | |
| 5100 | assume y: "y \<in> cball (f x) e \<inter> affine hull (f ` S)" | |
| 5101 | then have "y \<in> f ` (affine hull S)" | |
| 5102 | using affine_hull_linear_image[of f S] assms by auto | |
| 5103 | then obtain xy where xy: "xy \<in> affine hull S" "f xy = y" by auto | |
| 5104 | with y have "norm (f x - f xy) \<le> e1 * e2" | |
| 5105 | using cball_def[of "f x" e] dist_norm[of "f x" y] e_def by auto | |
| 5106 | 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 | 5107 | using assms linear_diff[of f x xy] linear_conv_bounded_linear[of f] by auto | 
| 53347 | 5108 | moreover have *: "x - xy \<in> span S" | 
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 5109 | using subspace_diff[of "span S" x xy] subspace_span \<open>x \<in> S\<close> xy | 
| 53347 | 5110 | affine_hull_subset_span[of S] span_inc | 
| 5111 | by auto | |
| 5112 | moreover from * have "e1 * norm (x - xy) \<le> norm (f (x - xy))" | |
| 5113 | using e1 by auto | |
| 5114 | ultimately have "e1 * norm (x - xy) \<le> e1 * e2" | |
| 5115 | by auto | |
| 5116 | then have "xy \<in> cball x e2" | |
| 5117 | using cball_def[of x e2] dist_norm[of x xy] e1 by auto | |
| 5118 | then have "y \<in> f ` S" | |
| 5119 | using xy e2 by auto | |
| 5120 | } | |
| 5121 | then have "f x \<in> rel_interior (f ` S)" | |
| 60420 | 5122 | using mem_rel_interior_cball[of "(f x)" "(f ` S)"] * \<open>e > 0\<close> by auto | 
| 49531 | 5123 | } | 
| 53347 | 5124 | ultimately show ?thesis by auto | 
| 40377 | 5125 | qed | 
| 5126 | ||
| 5127 | lemma rel_interior_injective_linear_image: | |
| 53347 | 5128 | fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space" | 
| 5129 | assumes "bounded_linear f" | |
| 5130 | and "inj f" | |
| 5131 | shows "rel_interior (f ` S) = f ` (rel_interior S)" | |
| 5132 | using assms rel_interior_injective_on_span_linear_image[of f S] | |
| 5133 | subset_inj_on[of f "UNIV" "span S"] | |
| 5134 | by auto | |
| 5135 | ||
| 40377 | 5136 | |
| 60420 | 5137 | subsection\<open>Some Properties of subset of standard basis\<close> | 
| 40377 | 5138 | |
| 53347 | 5139 | lemma affine_hull_substd_basis: | 
| 5140 | assumes "d \<subseteq> Basis" | |
| 5141 |   shows "affine hull (insert 0 d) = {x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
 | |
| 5142 | (is "affine hull (insert 0 ?A) = ?B") | |
| 5143 | proof - | |
| 61076 | 5144 | have *: "\<And>A. op + (0::'a) ` A = A" "\<And>A. op + (- (0::'a)) ` A = A" | 
| 53347 | 5145 | by auto | 
| 5146 | show ?thesis | |
| 5147 | unfolding affine_hull_insert_span_gen span_substd_basis[OF assms,symmetric] * .. | |
| 40377 | 5148 | qed | 
| 5149 | ||
| 60303 | 5150 | lemma affine_hull_convex_hull [simp]: "affine hull (convex hull S) = affine hull S" | 
| 53347 | 5151 | by (metis Int_absorb1 Int_absorb2 convex_hull_subset_affine_hull hull_hull hull_mono hull_subset) | 
| 5152 | ||
| 40377 | 5153 | |
| 60420 | 5154 | subsection \<open>Openness and compactness are preserved by convex hull operation.\<close> | 
| 33175 | 5155 | |
| 34964 | 5156 | lemma open_convex_hull[intro]: | 
| 33175 | 5157 | fixes s :: "'a::real_normed_vector set" | 
| 5158 | assumes "open s" | |
| 53347 | 5159 | shows "open (convex hull s)" | 
| 5160 | unfolding open_contains_cball convex_hull_explicit | |
| 5161 | unfolding mem_Collect_eq ball_simps(8) | |
| 5162 | proof (rule, rule) | |
| 5163 | fix a | |
| 64267 | 5164 | 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" | 
| 5165 | 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 | 5166 | by auto | 
| 5167 | ||
| 5168 | from assms[unfolded open_contains_cball] obtain b | |
| 5169 | where b: "\<forall>x\<in>s. 0 < b x \<and> cball x (b x) \<subseteq> s" | |
| 5170 | using bchoice[of s "\<lambda>x e. e > 0 \<and> cball x e \<subseteq> s"] by auto | |
| 5171 |   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 | 5172 | using obt by auto | 
| 63040 | 5173 | define i where "i = b ` t" | 
| 53347 | 5174 | |
| 5175 | show "\<exists>e > 0. | |
| 64267 | 5176 |     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 | 5177 | apply (rule_tac x = "Min i" in exI) | 
| 5178 | unfolding subset_eq | |
| 5179 | apply rule | |
| 5180 | defer | |
| 5181 | apply rule | |
| 5182 | unfolding mem_Collect_eq | |
| 5183 | proof - | |
| 5184 | show "0 < Min i" | |
| 60420 | 5185 |       unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] \<open>b ` t\<noteq>{}\<close>]
 | 
| 53347 | 5186 | using b | 
| 5187 | apply simp | |
| 5188 | apply rule | |
| 5189 | apply (erule_tac x=x in ballE) | |
| 60420 | 5190 | using \<open>t\<subseteq>s\<close> | 
| 53347 | 5191 | apply auto | 
| 5192 | done | |
| 5193 | next | |
| 5194 | fix y | |
| 5195 | assume "y \<in> cball a (Min i)" | |
| 5196 | then have y: "norm (a - y) \<le> Min i" | |
| 5197 | unfolding dist_norm[symmetric] by auto | |
| 5198 |     {
 | |
| 5199 | fix x | |
| 5200 | assume "x \<in> t" | |
| 5201 | then have "Min i \<le> b x" | |
| 5202 | unfolding i_def | |
| 5203 | apply (rule_tac Min_le) | |
| 5204 | using obt(1) | |
| 5205 | apply auto | |
| 5206 | done | |
| 5207 | then have "x + (y - a) \<in> cball x (b x)" | |
| 5208 | using y unfolding mem_cball dist_norm by auto | |
| 60420 | 5209 | moreover from \<open>x\<in>t\<close> have "x \<in> s" | 
| 53347 | 5210 | using obt(2) by auto | 
| 5211 | ultimately have "x + (y - a) \<in> s" | |
| 54465 | 5212 | using y and b[THEN bspec[where x=x]] unfolding subset_eq by fast | 
| 53347 | 5213 | } | 
| 33175 | 5214 | moreover | 
| 53347 | 5215 | have *: "inj_on (\<lambda>v. v + (y - a)) t" | 
| 5216 | unfolding inj_on_def by auto | |
| 33175 | 5217 | have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a))) = 1" | 
| 64267 | 5218 | unfolding sum.reindex[OF *] o_def using obt(4) by auto | 
| 33175 | 5219 | moreover have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a)) *\<^sub>R v) = y" | 
| 64267 | 5220 | unfolding sum.reindex[OF *] o_def using obt(4,5) | 
| 5221 | by (simp add: sum.distrib sum_subtractf scaleR_left.sum[symmetric] scaleR_right_distrib) | |
| 53347 | 5222 | ultimately | 
| 64267 | 5223 | 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 | 5224 | apply (rule_tac x="(\<lambda>v. v + (y - a)) ` t" in exI) | 
| 5225 | apply (rule_tac x="\<lambda>v. u (v - (y - a))" in exI) | |
| 5226 | using obt(1, 3) | |
| 5227 | apply auto | |
| 5228 | done | |
| 33175 | 5229 | qed | 
| 5230 | qed | |
| 5231 | ||
| 5232 | lemma compact_convex_combinations: | |
| 5233 | fixes s t :: "'a::real_normed_vector set" | |
| 5234 | assumes "compact s" "compact t" | |
| 5235 |   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 | 5236 | proof - | 
| 33175 | 5237 |   let ?X = "{0..1} \<times> s \<times> t"
 | 
| 5238 | let ?h = "(\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" | |
| 53347 | 5239 |   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"
 | 
| 5240 | apply (rule set_eqI) | |
| 5241 | unfolding image_iff mem_Collect_eq | |
| 5242 | apply rule | |
| 5243 | apply auto | |
| 5244 | apply (rule_tac x=u in rev_bexI) | |
| 5245 | apply simp | |
| 5246 | apply (erule rev_bexI) | |
| 5247 | apply (erule rev_bexI) | |
| 5248 | apply simp | |
| 5249 | apply auto | |
| 5250 | done | |
| 56188 | 5251 | have "continuous_on ?X (\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" | 
| 33175 | 5252 | unfolding continuous_on by (rule ballI) (intro tendsto_intros) | 
| 53347 | 5253 | then show ?thesis | 
| 5254 | unfolding * | |
| 33175 | 5255 | apply (rule compact_continuous_image) | 
| 56188 | 5256 | apply (intro compact_Times compact_Icc assms) | 
| 33175 | 5257 | done | 
| 5258 | qed | |
| 5259 | ||
| 44525 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5260 | lemma finite_imp_compact_convex_hull: | 
| 53347 | 5261 | fixes s :: "'a::real_normed_vector set" | 
| 5262 | assumes "finite s" | |
| 5263 | shows "compact (convex hull s)" | |
| 44525 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5264 | proof (cases "s = {}")
 | 
| 53347 | 5265 | case True | 
| 5266 | then show ?thesis by simp | |
| 44525 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5267 | next | 
| 53347 | 5268 | case False | 
| 5269 | with assms show ?thesis | |
| 44525 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5270 | proof (induct rule: finite_ne_induct) | 
| 53347 | 5271 | case (singleton x) | 
| 5272 | show ?case by simp | |
| 44525 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5273 | next | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5274 | case (insert x A) | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5275 | 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 | 5276 |     let ?T = "{0..1::real} \<times> (convex hull A)"
 | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5277 | have "continuous_on ?T ?f" | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5278 | 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 | 5279 | moreover have "compact ?T" | 
| 56188 | 5280 | by (intro compact_Times compact_Icc insert) | 
| 44525 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5281 | ultimately have "compact (?f ` ?T)" | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5282 | by (rule compact_continuous_image) | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5283 | also have "?f ` ?T = convex hull (insert x A)" | 
| 60420 | 5284 |       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 | 5285 | apply safe | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5286 | apply (rule_tac x=a in exI, simp) | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5287 | apply (rule_tac x="1 - a" in exI, simp) | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5288 | apply fast | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5289 | 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 | 5290 | done | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5291 | finally show "compact (convex hull (insert x A))" . | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5292 | qed | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5293 | qed | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5294 | |
| 53347 | 5295 | lemma compact_convex_hull: | 
| 5296 | fixes s :: "'a::euclidean_space set" | |
| 5297 | assumes "compact s" | |
| 5298 | shows "compact (convex hull s)" | |
| 5299 | proof (cases "s = {}")
 | |
| 5300 | case True | |
| 5301 | then show ?thesis using compact_empty by simp | |
| 33175 | 5302 | next | 
| 53347 | 5303 | case False | 
| 5304 | then obtain w where "w \<in> s" by auto | |
| 5305 | show ?thesis | |
| 5306 | unfolding caratheodory[of s] | |
| 5307 |   proof (induct ("DIM('a) + 1"))
 | |
| 5308 | case 0 | |
| 5309 |     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 | 5310 | using compact_empty by auto | 
| 53347 | 5311 | from 0 show ?case unfolding * by simp | 
| 33175 | 5312 | next | 
| 5313 | case (Suc n) | |
| 53347 | 5314 | show ?case | 
| 5315 | proof (cases "n = 0") | |
| 5316 | case True | |
| 5317 |       have "{x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t} = s"
 | |
| 5318 | unfolding set_eq_iff and mem_Collect_eq | |
| 5319 | proof (rule, rule) | |
| 5320 | fix x | |
| 5321 | assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t" | |
| 5322 | then obtain t where t: "finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t" | |
| 5323 | by auto | |
| 5324 | show "x \<in> s" | |
| 5325 | proof (cases "card t = 0") | |
| 5326 | case True | |
| 5327 | then show ?thesis | |
| 5328 | using t(4) unfolding card_0_eq[OF t(1)] by simp | |
| 33175 | 5329 | next | 
| 53347 | 5330 | case False | 
| 60420 | 5331 | then have "card t = Suc 0" using t(3) \<open>n=0\<close> by auto | 
| 33175 | 5332 |           then obtain a where "t = {a}" unfolding card_Suc_eq by auto
 | 
| 53347 | 5333 | then show ?thesis using t(2,4) by simp | 
| 33175 | 5334 | qed | 
| 5335 | next | |
| 5336 | fix x assume "x\<in>s" | |
| 53347 | 5337 | then show "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t" | 
| 5338 |           apply (rule_tac x="{x}" in exI)
 | |
| 5339 | unfolding convex_hull_singleton | |
| 5340 | apply auto | |
| 5341 | done | |
| 5342 | qed | |
| 5343 | then show ?thesis using assms by simp | |
| 33175 | 5344 | next | 
| 53347 | 5345 | case False | 
| 5346 |       have "{x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t} =
 | |
| 5347 |         {(1 - u) *\<^sub>R x + u *\<^sub>R y | x y u.
 | |
| 5348 |           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}}"
 | |
| 5349 | unfolding set_eq_iff and mem_Collect_eq | |
| 5350 | proof (rule, rule) | |
| 5351 | fix x | |
| 5352 | assume "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and> | |
| 33175 | 5353 | 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 | 5354 | then obtain u v c t where obt: "x = (1 - c) *\<^sub>R u + c *\<^sub>R v" | 
| 5355 | "0 \<le> c \<and> c \<le> 1" "u \<in> s" "finite t" "t \<subseteq> s" "card t \<le> n" "v \<in> convex hull t" | |
| 5356 | by auto | |
| 33175 | 5357 | 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 | 5358 | apply (rule convexD_alt) | 
| 53347 | 5359 | using obt(2) and convex_convex_hull and hull_subset[of "insert u t" convex] | 
| 5360 | using obt(7) and hull_mono[of t "insert u t"] | |
| 5361 | apply auto | |
| 5362 | done | |
| 33175 | 5363 | ultimately show "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t" | 
| 53347 | 5364 | apply (rule_tac x="insert u t" in exI) | 
| 5365 | apply (auto simp add: card_insert_if) | |
| 5366 | done | |
| 33175 | 5367 | next | 
| 53347 | 5368 | fix x | 
| 5369 | assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t" | |
| 5370 | then obtain t where t: "finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t" | |
| 5371 | by auto | |
| 5372 | show "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and> | |
| 33175 | 5373 | 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 | 5374 | proof (cases "card t = Suc n") | 
| 5375 | case False | |
| 5376 | then have "card t \<le> n" using t(3) by auto | |
| 5377 | then show ?thesis | |
| 5378 | apply (rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI) | |
| 60420 | 5379 | using \<open>w\<in>s\<close> and t | 
| 53347 | 5380 | apply (auto intro!: exI[where x=t]) | 
| 5381 | done | |
| 33175 | 5382 | next | 
| 53347 | 5383 | case True | 
| 5384 | then obtain a u where au: "t = insert a u" "a\<notin>u" | |
| 5385 | apply (drule_tac card_eq_SucD) | |
| 5386 | apply auto | |
| 5387 | done | |
| 5388 | show ?thesis | |
| 5389 |           proof (cases "u = {}")
 | |
| 5390 | case True | |
| 5391 | then have "x = a" using t(4)[unfolded au] by auto | |
| 60420 | 5392 | show ?thesis unfolding \<open>x = a\<close> | 
| 53347 | 5393 | apply (rule_tac x=a in exI) | 
| 5394 | apply (rule_tac x=a in exI) | |
| 5395 | apply (rule_tac x=1 in exI) | |
| 60420 | 5396 | using t and \<open>n \<noteq> 0\<close> | 
| 53347 | 5397 | unfolding au | 
| 5398 |               apply (auto intro!: exI[where x="{a}"])
 | |
| 5399 | done | |
| 33175 | 5400 | next | 
| 53347 | 5401 | case False | 
| 5402 | obtain ux vx b where obt: "ux\<ge>0" "vx\<ge>0" "ux + vx = 1" | |
| 5403 | "b \<in> convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b" | |
| 5404 | using t(4)[unfolded au convex_hull_insert[OF False]] | |
| 5405 | by auto | |
| 5406 | have *: "1 - vx = ux" using obt(3) by auto | |
| 5407 | show ?thesis | |
| 5408 | apply (rule_tac x=a in exI) | |
| 5409 | apply (rule_tac x=b in exI) | |
| 5410 | apply (rule_tac x=vx in exI) | |
| 5411 | using obt and t(1-3) | |
| 5412 | unfolding au and * using card_insert_disjoint[OF _ au(2)] | |
| 5413 | apply (auto intro!: exI[where x=u]) | |
| 5414 | done | |
| 33175 | 5415 | qed | 
| 5416 | qed | |
| 5417 | qed | |
| 53347 | 5418 | then show ?thesis | 
| 5419 | using compact_convex_combinations[OF assms Suc] by simp | |
| 33175 | 5420 | qed | 
| 36362 
06475a1547cb
fix lots of looping simp calls and other warnings
 huffman parents: 
36341diff
changeset | 5421 | qed | 
| 33175 | 5422 | qed | 
| 5423 | ||
| 53347 | 5424 | |
| 60420 | 5425 | subsection \<open>Extremal points of a simplex are some vertices.\<close> | 
| 33175 | 5426 | |
| 5427 | lemma dist_increases_online: | |
| 5428 | fixes a b d :: "'a::real_inner" | |
| 5429 | assumes "d \<noteq> 0" | |
| 5430 | shows "dist a (b + d) > dist a b \<or> dist a (b - d) > dist a b" | |
| 53347 | 5431 | proof (cases "inner a d - inner b d > 0") | 
| 5432 | case True | |
| 5433 | then have "0 < inner d d + (inner a d * 2 - inner b d * 2)" | |
| 5434 | apply (rule_tac add_pos_pos) | |
| 5435 | using assms | |
| 5436 | apply auto | |
| 5437 | done | |
| 5438 | then show ?thesis | |
| 5439 | apply (rule_tac disjI2) | |
| 5440 | unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff | |
| 5441 | apply (simp add: algebra_simps inner_commute) | |
| 5442 | done | |
| 33175 | 5443 | next | 
| 53347 | 5444 | case False | 
| 5445 | then have "0 < inner d d + (inner b d * 2 - inner a d * 2)" | |
| 5446 | apply (rule_tac add_pos_nonneg) | |
| 5447 | using assms | |
| 5448 | apply auto | |
| 5449 | done | |
| 5450 | then show ?thesis | |
| 5451 | apply (rule_tac disjI1) | |
| 5452 | unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff | |
| 5453 | apply (simp add: algebra_simps inner_commute) | |
| 5454 | done | |
| 33175 | 5455 | qed | 
| 5456 | ||
| 5457 | lemma norm_increases_online: | |
| 5458 | fixes d :: "'a::real_inner" | |
| 53347 | 5459 | shows "d \<noteq> 0 \<Longrightarrow> norm (a + d) > norm a \<or> norm(a - d) > norm a" | 
| 33175 | 5460 | using dist_increases_online[of d a 0] unfolding dist_norm by auto | 
| 5461 | ||
| 5462 | lemma simplex_furthest_lt: | |
| 53347 | 5463 | fixes s :: "'a::real_inner set" | 
| 5464 | assumes "finite s" | |
| 5465 | shows "\<forall>x \<in> convex hull s. x \<notin> s \<longrightarrow> (\<exists>y \<in> convex hull s. norm (x - a) < norm(y - a))" | |
| 5466 | using assms | |
| 5467 | proof induct | |
| 5468 | fix x s | |
| 5469 | 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))" | |
| 5470 | show "\<forall>xa\<in>convex hull insert x s. xa \<notin> insert x s \<longrightarrow> | |
| 5471 | (\<exists>y\<in>convex hull insert x s. norm (xa - a) < norm (y - a))" | |
| 5472 |   proof (rule, rule, cases "s = {}")
 | |
| 5473 | case False | |
| 5474 | fix y | |
| 5475 | assume y: "y \<in> convex hull insert x s" "y \<notin> insert x s" | |
| 5476 | 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 | 5477 | using y(1)[unfolded convex_hull_insert[OF False]] by auto | 
| 5478 | show "\<exists>z\<in>convex hull insert x s. norm (y - a) < norm (z - a)" | |
| 53347 | 5479 | proof (cases "y \<in> convex hull s") | 
| 5480 | case True | |
| 5481 | then obtain z where "z \<in> convex hull s" "norm (y - a) < norm (z - a)" | |
| 33175 | 5482 | using as(3)[THEN bspec[where x=y]] and y(2) by auto | 
| 53347 | 5483 | then show ?thesis | 
| 5484 | apply (rule_tac x=z in bexI) | |
| 5485 | unfolding convex_hull_insert[OF False] | |
| 5486 | apply auto | |
| 5487 | done | |
| 33175 | 5488 | next | 
| 53347 | 5489 | case False | 
| 5490 | show ?thesis | |
| 5491 | using obt(3) | |
| 5492 | proof (cases "u = 0", case_tac[!] "v = 0") | |
| 5493 | assume "u = 0" "v \<noteq> 0" | |
| 5494 | then have "y = b" using obt by auto | |
| 5495 | then show ?thesis using False and obt(4) by auto | |
| 33175 | 5496 | next | 
| 53347 | 5497 | assume "u \<noteq> 0" "v = 0" | 
| 5498 | then have "y = x" using obt by auto | |
| 5499 | then show ?thesis using y(2) by auto | |
| 5500 | next | |
| 5501 | assume "u \<noteq> 0" "v \<noteq> 0" | |
| 5502 | then obtain w where w: "w>0" "w<u" "w<v" | |
| 5503 | using real_lbound_gt_zero[of u v] and obt(1,2) by auto | |
| 5504 | have "x \<noteq> b" | |
| 5505 | proof | |
| 5506 | assume "x = b" | |
| 5507 | then have "y = b" unfolding obt(5) | |
| 5508 | using obt(3) by (auto simp add: scaleR_left_distrib[symmetric]) | |
| 5509 | then show False using obt(4) and False by simp | |
| 5510 | qed | |
| 5511 | then have *: "w *\<^sub>R (x - b) \<noteq> 0" using w(1) by auto | |
| 5512 | show ?thesis | |
| 5513 | using dist_increases_online[OF *, of a y] | |
| 5514 | proof (elim disjE) | |
| 33175 | 5515 | assume "dist a y < dist a (y + w *\<^sub>R (x - b))" | 
| 53347 | 5516 | then have "norm (y - a) < norm ((u + w) *\<^sub>R x + (v - w) *\<^sub>R b - a)" | 
| 5517 | unfolding dist_commute[of a] | |
| 5518 | unfolding dist_norm obt(5) | |
| 5519 | by (simp add: algebra_simps) | |
| 33175 | 5520 | moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \<in> convex hull insert x s" | 
| 60420 | 5521 |             unfolding convex_hull_insert[OF \<open>s\<noteq>{}\<close>] and mem_Collect_eq
 | 
| 53347 | 5522 | apply (rule_tac x="u + w" in exI) | 
| 5523 | apply rule | |
| 5524 | defer | |
| 5525 | apply (rule_tac x="v - w" in exI) | |
| 60420 | 5526 | using \<open>u \<ge> 0\<close> and w and obt(3,4) | 
| 53347 | 5527 | apply auto | 
| 5528 | done | |
| 33175 | 5529 | ultimately show ?thesis by auto | 
| 5530 | next | |
| 5531 | assume "dist a y < dist a (y - w *\<^sub>R (x - b))" | |
| 53347 | 5532 | then have "norm (y - a) < norm ((u - w) *\<^sub>R x + (v + w) *\<^sub>R b - a)" | 
| 5533 | unfolding dist_commute[of a] | |
| 5534 | unfolding dist_norm obt(5) | |
| 5535 | by (simp add: algebra_simps) | |
| 33175 | 5536 | moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \<in> convex hull insert x s" | 
| 60420 | 5537 |             unfolding convex_hull_insert[OF \<open>s\<noteq>{}\<close>] and mem_Collect_eq
 | 
| 53347 | 5538 | apply (rule_tac x="u - w" in exI) | 
| 5539 | apply rule | |
| 5540 | defer | |
| 5541 | apply (rule_tac x="v + w" in exI) | |
| 60420 | 5542 | using \<open>u \<ge> 0\<close> and w and obt(3,4) | 
| 53347 | 5543 | apply auto | 
| 5544 | done | |
| 33175 | 5545 | ultimately show ?thesis by auto | 
| 5546 | qed | |
| 5547 | qed auto | |
| 5548 | qed | |
| 5549 | qed auto | |
| 5550 | qed (auto simp add: assms) | |
| 5551 | ||
| 5552 | lemma simplex_furthest_le: | |
| 53347 | 5553 | fixes s :: "'a::real_inner set" | 
| 5554 | assumes "finite s" | |
| 5555 |     and "s \<noteq> {}"
 | |
| 5556 | shows "\<exists>y\<in>s. \<forall>x\<in> convex hull s. norm (x - a) \<le> norm (y - a)" | |
| 5557 | proof - | |
| 5558 |   have "convex hull s \<noteq> {}"
 | |
| 5559 | using hull_subset[of s convex] and assms(2) by auto | |
| 5560 | then obtain x where x: "x \<in> convex hull s" "\<forall>y\<in>convex hull s. norm (y - a) \<le> norm (x - a)" | |
| 33175 | 5561 | using distance_attains_sup[OF finite_imp_compact_convex_hull[OF assms(1)], of a] | 
| 53347 | 5562 | unfolding dist_commute[of a] | 
| 5563 | unfolding dist_norm | |
| 5564 | by auto | |
| 5565 | show ?thesis | |
| 5566 | proof (cases "x \<in> s") | |
| 5567 | case False | |
| 5568 | then obtain y where "y \<in> convex hull s" "norm (x - a) < norm (y - a)" | |
| 5569 | using simplex_furthest_lt[OF assms(1), THEN bspec[where x=x]] and x(1) | |
| 5570 | by auto | |
| 5571 | then show ?thesis | |
| 5572 | using x(2)[THEN bspec[where x=y]] by auto | |
| 5573 | next | |
| 5574 | case True | |
| 5575 | with x show ?thesis by auto | |
| 5576 | qed | |
| 33175 | 5577 | qed | 
| 5578 | ||
| 5579 | lemma simplex_furthest_le_exists: | |
| 44525 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 5580 |   fixes s :: "('a::real_inner) set"
 | 
| 53347 | 5581 | shows "finite s \<Longrightarrow> \<forall>x\<in>(convex hull s). \<exists>y\<in>s. norm (x - a) \<le> norm (y - a)" | 
| 5582 |   using simplex_furthest_le[of s] by (cases "s = {}") auto
 | |
| 33175 | 5583 | |
| 5584 | lemma simplex_extremal_le: | |
| 53347 | 5585 | fixes s :: "'a::real_inner set" | 
| 5586 | assumes "finite s" | |
| 5587 |     and "s \<noteq> {}"
 | |
| 5588 | 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)" | |
| 5589 | proof - | |
| 5590 |   have "convex hull s \<noteq> {}"
 | |
| 5591 | using hull_subset[of s convex] and assms(2) by auto | |
| 5592 | then obtain u v where obt: "u \<in> convex hull s" "v \<in> convex hull s" | |
| 33175 | 5593 | "\<forall>x\<in>convex hull s. \<forall>y\<in>convex hull s. norm (x - y) \<le> norm (u - v)" | 
| 53347 | 5594 | using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]] | 
| 5595 | by (auto simp: dist_norm) | |
| 5596 | then show ?thesis | |
| 5597 | proof (cases "u\<notin>s \<or> v\<notin>s", elim disjE) | |
| 5598 | assume "u \<notin> s" | |
| 5599 | then obtain y where "y \<in> convex hull s" "norm (u - v) < norm (y - v)" | |
| 5600 | using simplex_furthest_lt[OF assms(1), THEN bspec[where x=u]] and obt(1) | |
| 5601 | by auto | |
| 5602 | then show ?thesis | |
| 5603 | using obt(3)[THEN bspec[where x=y], THEN bspec[where x=v]] and obt(2) | |
| 5604 | by auto | |
| 33175 | 5605 | next | 
| 53347 | 5606 | assume "v \<notin> s" | 
| 5607 | then obtain y where "y \<in> convex hull s" "norm (v - u) < norm (y - u)" | |
| 5608 | using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2) | |
| 5609 | by auto | |
| 5610 | then show ?thesis | |
| 5611 | using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1) | |
| 33175 | 5612 | by (auto simp add: norm_minus_commute) | 
| 5613 | qed auto | |
| 49531 | 5614 | qed | 
| 33175 | 5615 | |
| 5616 | lemma simplex_extremal_le_exists: | |
| 53347 | 5617 | fixes s :: "'a::real_inner set" | 
| 5618 | shows "finite s \<Longrightarrow> x \<in> convex hull s \<Longrightarrow> y \<in> convex hull s \<Longrightarrow> | |
| 5619 | \<exists>u\<in>s. \<exists>v\<in>s. norm (x - y) \<le> norm (u - v)" | |
| 5620 | using convex_hull_empty simplex_extremal_le[of s] | |
| 5621 |   by(cases "s = {}") auto
 | |
| 5622 | ||
| 33175 | 5623 | |
| 60420 | 5624 | subsection \<open>Closest point of a convex set is unique, with a continuous projection.\<close> | 
| 33175 | 5625 | |
| 53347 | 5626 | definition closest_point :: "'a::{real_inner,heine_borel} set \<Rightarrow> 'a \<Rightarrow> 'a"
 | 
| 5627 | where "closest_point s a = (SOME x. x \<in> s \<and> (\<forall>y\<in>s. dist a x \<le> dist a y))" | |
| 33175 | 5628 | |
| 5629 | lemma closest_point_exists: | |
| 53347 | 5630 | assumes "closed s" | 
| 5631 |     and "s \<noteq> {}"
 | |
| 5632 | shows "closest_point s a \<in> s" | |
| 5633 | and "\<forall>y\<in>s. dist a (closest_point s a) \<le> dist a y" | |
| 5634 | unfolding closest_point_def | |
| 5635 | apply(rule_tac[!] someI2_ex) | |
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 5636 | apply (auto intro: distance_attains_inf[OF assms(1,2), of a]) | 
| 53347 | 5637 | done | 
| 5638 | ||
| 5639 | lemma closest_point_in_set: "closed s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> closest_point s a \<in> s"
 | |
| 5640 | by (meson closest_point_exists) | |
| 5641 | ||
| 5642 | lemma closest_point_le: "closed s \<Longrightarrow> x \<in> s \<Longrightarrow> dist a (closest_point s a) \<le> dist a x" | |
| 33175 | 5643 | using closest_point_exists[of s] by auto | 
| 5644 | ||
| 5645 | lemma closest_point_self: | |
| 53347 | 5646 | assumes "x \<in> s" | 
| 5647 | shows "closest_point s x = x" | |
| 5648 | unfolding closest_point_def | |
| 5649 | apply (rule some1_equality, rule ex1I[of _ x]) | |
| 5650 | using assms | |
| 5651 | apply auto | |
| 5652 | done | |
| 5653 | ||
| 5654 | lemma closest_point_refl: "closed s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> closest_point s x = x \<longleftrightarrow> x \<in> s"
 | |
| 5655 | using closest_point_in_set[of s x] closest_point_self[of x s] | |
| 5656 | by auto | |
| 33175 | 5657 | |
| 36337 | 5658 | lemma closer_points_lemma: | 
| 33175 | 5659 | assumes "inner y z > 0" | 
| 5660 | shows "\<exists>u>0. \<forall>v>0. v \<le> u \<longrightarrow> norm(v *\<^sub>R z - y) < norm y" | |
| 53347 | 5661 | proof - | 
| 5662 | have z: "inner z z > 0" | |
| 5663 | unfolding inner_gt_zero_iff using assms by auto | |
| 5664 | then show ?thesis | |
| 5665 | using assms | |
| 5666 | apply (rule_tac x = "inner y z / inner z z" in exI) | |
| 5667 | apply rule | |
| 5668 | defer | |
| 5669 | proof rule+ | |
| 5670 | fix v | |
| 5671 | assume "0 < v" and "v \<le> inner y z / inner z z" | |
| 5672 | then show "norm (v *\<^sub>R z - y) < norm y" | |
| 5673 | unfolding norm_lt using z and assms | |
| 60420 | 5674 | by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ \<open>0<v\<close>]) | 
| 56541 | 5675 | qed auto | 
| 53347 | 5676 | qed | 
| 33175 | 5677 | |
| 5678 | lemma closer_point_lemma: | |
| 5679 | assumes "inner (y - x) (z - x) > 0" | |
| 5680 | shows "\<exists>u>0. u \<le> 1 \<and> dist (x + u *\<^sub>R (z - x)) y < dist x y" | |
| 53347 | 5681 | proof - | 
| 5682 | obtain u where "u > 0" | |
| 5683 | and u: "\<forall>v>0. v \<le> u \<longrightarrow> norm (v *\<^sub>R (z - x) - (y - x)) < norm (y - x)" | |
| 33175 | 5684 | using closer_points_lemma[OF assms] by auto | 
| 53347 | 5685 | show ?thesis | 
| 5686 | apply (rule_tac x="min u 1" in exI) | |
| 60420 | 5687 | using u[THEN spec[where x="min u 1"]] and \<open>u > 0\<close> | 
| 53347 | 5688 | unfolding dist_norm by (auto simp add: norm_minus_commute field_simps) | 
| 5689 | qed | |
| 33175 | 5690 | |
| 5691 | lemma any_closest_point_dot: | |
| 5692 | assumes "convex s" "closed s" "x \<in> s" "y \<in> s" "\<forall>z\<in>s. dist a x \<le> dist a z" | |
| 5693 | shows "inner (a - x) (y - x) \<le> 0" | |
| 53347 | 5694 | proof (rule ccontr) | 
| 5695 | assume "\<not> ?thesis" | |
| 5696 | then obtain u where u: "u>0" "u\<le>1" "dist (x + u *\<^sub>R (y - x)) a < dist x a" | |
| 5697 | using closer_point_lemma[of a x y] by auto | |
| 5698 | let ?z = "(1 - u) *\<^sub>R x + u *\<^sub>R y" | |
| 5699 | 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 | 5700 | using convexD_alt[OF assms(1,3,4), of u] using u by auto | 
| 53347 | 5701 | then show False | 
| 5702 | using assms(5)[THEN bspec[where x="?z"]] and u(3) | |
| 5703 | by (auto simp add: dist_commute algebra_simps) | |
| 5704 | qed | |
| 33175 | 5705 | |
| 5706 | lemma any_closest_point_unique: | |
| 36337 | 5707 | fixes x :: "'a::real_inner" | 
| 33175 | 5708 | assumes "convex s" "closed s" "x \<in> s" "y \<in> s" | 
| 53347 | 5709 | "\<forall>z\<in>s. dist a x \<le> dist a z" "\<forall>z\<in>s. dist a y \<le> dist a z" | 
| 5710 | shows "x = y" | |
| 5711 | using any_closest_point_dot[OF assms(1-4,5)] and any_closest_point_dot[OF assms(1-2,4,3,6)] | |
| 33175 | 5712 | unfolding norm_pths(1) and norm_le_square | 
| 5713 | by (auto simp add: algebra_simps) | |
| 5714 | ||
| 5715 | lemma closest_point_unique: | |
| 5716 | assumes "convex s" "closed s" "x \<in> s" "\<forall>z\<in>s. dist a x \<le> dist a z" | |
| 5717 | shows "x = closest_point s a" | |
| 49531 | 5718 | using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point s a"] | 
| 33175 | 5719 | using closest_point_exists[OF assms(2)] and assms(3) by auto | 
| 5720 | ||
| 5721 | lemma closest_point_dot: | |
| 5722 | assumes "convex s" "closed s" "x \<in> s" | |
| 5723 | shows "inner (a - closest_point s a) (x - closest_point s a) \<le> 0" | |
| 53347 | 5724 | apply (rule any_closest_point_dot[OF assms(1,2) _ assms(3)]) | 
| 5725 | using closest_point_exists[OF assms(2)] and assms(3) | |
| 5726 | apply auto | |
| 5727 | done | |
| 33175 | 5728 | |
| 5729 | lemma closest_point_lt: | |
| 5730 | assumes "convex s" "closed s" "x \<in> s" "x \<noteq> closest_point s a" | |
| 5731 | shows "dist a (closest_point s a) < dist a x" | |
| 53347 | 5732 | apply (rule ccontr) | 
| 5733 | apply (rule_tac notE[OF assms(4)]) | |
| 5734 | apply (rule closest_point_unique[OF assms(1-3), of a]) | |
| 5735 | using closest_point_le[OF assms(2), of _ a] | |
| 5736 | apply fastforce | |
| 5737 | done | |
| 33175 | 5738 | |
| 5739 | lemma closest_point_lipschitz: | |
| 53347 | 5740 | assumes "convex s" | 
| 5741 |     and "closed s" "s \<noteq> {}"
 | |
| 33175 | 5742 | shows "dist (closest_point s x) (closest_point s y) \<le> dist x y" | 
| 53347 | 5743 | proof - | 
| 33175 | 5744 | have "inner (x - closest_point s x) (closest_point s y - closest_point s x) \<le> 0" | 
| 53347 | 5745 | and "inner (y - closest_point s y) (closest_point s x - closest_point s y) \<le> 0" | 
| 5746 | apply (rule_tac[!] any_closest_point_dot[OF assms(1-2)]) | |
| 5747 | using closest_point_exists[OF assms(2-3)] | |
| 5748 | apply auto | |
| 5749 | done | |
| 5750 | then show ?thesis unfolding dist_norm and norm_le | |
| 33175 | 5751 | using inner_ge_zero[of "(x - closest_point s x) - (y - closest_point s y)"] | 
| 53347 | 5752 | by (simp add: inner_add inner_diff inner_commute) | 
| 5753 | qed | |
| 33175 | 5754 | |
| 5755 | lemma continuous_at_closest_point: | |
| 53347 | 5756 | assumes "convex s" | 
| 5757 | and "closed s" | |
| 5758 |     and "s \<noteq> {}"
 | |
| 33175 | 5759 | shows "continuous (at x) (closest_point s)" | 
| 49531 | 5760 | unfolding continuous_at_eps_delta | 
| 33175 | 5761 | using le_less_trans[OF closest_point_lipschitz[OF assms]] by auto | 
| 5762 | ||
| 5763 | lemma continuous_on_closest_point: | |
| 53347 | 5764 | assumes "convex s" | 
| 5765 | and "closed s" | |
| 5766 |     and "s \<noteq> {}"
 | |
| 33175 | 5767 | shows "continuous_on t (closest_point s)" | 
| 53347 | 5768 | by (metis continuous_at_imp_continuous_on continuous_at_closest_point[OF assms]) | 
| 5769 | ||
| 63881 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5770 | proposition closest_point_in_rel_interior: | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5771 |   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 | 5772 | 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 | 5773 | proof (cases "x \<in> S") | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5774 | case True | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5775 | then show ?thesis | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5776 | by (simp add: closest_point_self) | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5777 | next | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5778 | case False | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5779 | 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 | 5780 | proof - | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5781 | 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 | 5782 | 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 | 5783 | 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 | 5784 | 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 | 5785 | 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 | 5786 | 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 | 5787 | (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 | 5788 | 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 | 5789 | 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 | 5790 | 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 | 5791 | by simp | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5792 | 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 | 5793 | 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 | 5794 | 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 | 5795 | 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 | 5796 | have "y \<in> affine hull S" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5797 | unfolding y_def | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5798 | 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 | 5799 | 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 | 5800 | 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 | 5801 | ultimately have "y \<in> S" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5802 | using subsetD [OF e] by simp | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5803 | 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 | 5804 | 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 | 5805 | with no_less show False | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5806 | by (simp add: dist_norm) | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5807 | qed | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5808 | 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 | 5809 | 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 | 5810 | ultimately show ?thesis by blast | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5811 | qed | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 5812 | |
| 33175 | 5813 | |
| 60420 | 5814 | subsubsection \<open>Various point-to-set separating/supporting hyperplane theorems.\<close> | 
| 33175 | 5815 | |
| 5816 | lemma supporting_hyperplane_closed_point: | |
| 36337 | 5817 |   fixes z :: "'a::{real_inner,heine_borel}"
 | 
| 53347 | 5818 | assumes "convex s" | 
| 5819 | and "closed s" | |
| 5820 |     and "s \<noteq> {}"
 | |
| 5821 | and "z \<notin> s" | |
| 5822 | 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)" | |
| 5823 | proof - | |
| 5824 | 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 | 5825 | by (metis distance_attains_inf[OF assms(2-3)]) | 
| 53347 | 5826 | show ?thesis | 
| 5827 | apply (rule_tac x="y - z" in exI) | |
| 5828 | apply (rule_tac x="inner (y - z) y" in exI) | |
| 5829 | apply (rule_tac x=y in bexI) | |
| 5830 | apply rule | |
| 5831 | defer | |
| 5832 | apply rule | |
| 5833 | defer | |
| 5834 | apply rule | |
| 5835 | apply (rule ccontr) | |
| 60420 | 5836 | using \<open>y \<in> s\<close> | 
| 53347 | 5837 | proof - | 
| 5838 | 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 | 5839 | apply (subst diff_gt_0_iff_gt [symmetric]) | 
| 53347 | 5840 | unfolding inner_diff_right[symmetric] and inner_gt_zero_iff | 
| 60420 | 5841 | using \<open>y\<in>s\<close> \<open>z\<notin>s\<close> | 
| 53347 | 5842 | apply auto | 
| 5843 | done | |
| 33175 | 5844 | next | 
| 53347 | 5845 | fix x | 
| 5846 | assume "x \<in> s" | |
| 5847 | 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 | 5848 | using assms(1)[unfolded convex_alt] and y and \<open>x\<in>s\<close> and \<open>y\<in>s\<close> by auto | 
| 53347 | 5849 | assume "\<not> inner (y - z) y \<le> inner (y - z) x" | 
| 5850 | then obtain v where "v > 0" "v \<le> 1" "dist (y + v *\<^sub>R (x - y)) z < dist y z" | |
| 5851 | using closer_point_lemma[of z y x] by (auto simp add: inner_diff) | |
| 5852 | then show False | |
| 5853 | using *[THEN spec[where x=v]] by (auto simp add: dist_commute algebra_simps) | |
| 33175 | 5854 | qed auto | 
| 5855 | qed | |
| 5856 | ||
| 5857 | lemma separating_hyperplane_closed_point: | |
| 36337 | 5858 |   fixes z :: "'a::{real_inner,heine_borel}"
 | 
| 53347 | 5859 | assumes "convex s" | 
| 5860 | and "closed s" | |
| 5861 | and "z \<notin> s" | |
| 33175 | 5862 | shows "\<exists>a b. inner a z < b \<and> (\<forall>x\<in>s. inner a x > b)" | 
| 53347 | 5863 | proof (cases "s = {}")
 | 
| 5864 | case True | |
| 5865 | then show ?thesis | |
| 5866 | apply (rule_tac x="-z" in exI) | |
| 5867 | apply (rule_tac x=1 in exI) | |
| 5868 | using less_le_trans[OF _ inner_ge_zero[of z]] | |
| 5869 | apply auto | |
| 5870 | done | |
| 33175 | 5871 | next | 
| 53347 | 5872 | case False | 
| 5873 | 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 | 5874 | by (metis distance_attains_inf[OF assms(2) False]) | 
| 53347 | 5875 | show ?thesis | 
| 5876 | apply (rule_tac x="y - z" in exI) | |
| 5877 | apply (rule_tac x="inner (y - z) z + (norm (y - z))\<^sup>2 / 2" in exI) | |
| 5878 | apply rule | |
| 5879 | defer | |
| 5880 | apply rule | |
| 5881 | proof - | |
| 5882 | fix x | |
| 5883 | assume "x \<in> s" | |
| 5884 | have "\<not> 0 < inner (z - y) (x - y)" | |
| 5885 | apply (rule notI) | |
| 5886 | apply (drule closer_point_lemma) | |
| 5887 | proof - | |
| 33175 | 5888 | assume "\<exists>u>0. u \<le> 1 \<and> dist (y + u *\<^sub>R (x - y)) z < dist y z" | 
| 53347 | 5889 | then obtain u where "u > 0" "u \<le> 1" "dist (y + u *\<^sub>R (x - y)) z < dist y z" | 
| 5890 | by auto | |
| 5891 | then show False using y[THEN bspec[where x="y + u *\<^sub>R (x - y)"]] | |
| 33175 | 5892 | using assms(1)[unfolded convex_alt, THEN bspec[where x=y]] | 
| 60420 | 5893 | using \<open>x\<in>s\<close> \<open>y\<in>s\<close> by (auto simp add: dist_commute algebra_simps) | 
| 53347 | 5894 | qed | 
| 5895 | moreover have "0 < (norm (y - z))\<^sup>2" | |
| 60420 | 5896 | using \<open>y\<in>s\<close> \<open>z\<notin>s\<close> by auto | 
| 53347 | 5897 | then have "0 < inner (y - z) (y - z)" | 
| 5898 | 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 | 5899 | ultimately show "inner (y - z) z + (norm (y - z))\<^sup>2 / 2 < inner (y - z) x" | 
| 53347 | 5900 | unfolding power2_norm_eq_inner and not_less | 
| 5901 | by (auto simp add: field_simps inner_commute inner_diff) | |
| 60420 | 5902 | qed (insert \<open>y\<in>s\<close> \<open>z\<notin>s\<close>, auto) | 
| 33175 | 5903 | qed | 
| 5904 | ||
| 5905 | lemma separating_hyperplane_closed_0: | |
| 53347 | 5906 |   assumes "convex (s::('a::euclidean_space) set)"
 | 
| 5907 | and "closed s" | |
| 5908 | and "0 \<notin> s" | |
| 33175 | 5909 | shows "\<exists>a b. a \<noteq> 0 \<and> 0 < b \<and> (\<forall>x\<in>s. inner a x > b)" | 
| 53347 | 5910 | 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 | 5911 | case True | 
| 53347 | 5912 | have "norm ((SOME i. i\<in>Basis)::'a) = 1" "(SOME i. i\<in>Basis) \<noteq> (0::'a)" | 
| 5913 | defer | |
| 5914 | apply (subst norm_le_zero_iff[symmetric]) | |
| 5915 | apply (auto simp: SOME_Basis) | |
| 5916 | done | |
| 5917 | then show ?thesis | |
| 5918 | apply (rule_tac x="SOME i. i\<in>Basis" in exI) | |
| 5919 | apply (rule_tac x=1 in exI) | |
| 5920 | using True using DIM_positive[where 'a='a] | |
| 5921 | apply auto | |
| 5922 | done | |
| 5923 | next | |
| 5924 | case False | |
| 5925 | then show ?thesis | |
| 5926 | using False using separating_hyperplane_closed_point[OF assms] | |
| 5927 | apply (elim exE) | |
| 5928 | unfolding inner_zero_right | |
| 5929 | apply (rule_tac x=a in exI) | |
| 5930 | apply (rule_tac x=b in exI) | |
| 5931 | apply auto | |
| 5932 | done | |
| 5933 | qed | |
| 5934 | ||
| 33175 | 5935 | |
| 60420 | 5936 | subsubsection \<open>Now set-to-set for closed/compact sets\<close> | 
| 33175 | 5937 | |
| 5938 | 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 | 5939 | 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 | 5940 | 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 | 5941 | 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 | 5942 | 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 | 5943 | 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 | 5944 |     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 | 5945 |     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 | 5946 | 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 | 5947 | proof (cases "S = {}")
 | 
| 33175 | 5948 | 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 | 5949 | obtain b where b: "b > 0" "\<forall>x\<in>T. norm x \<le> b" | 
| 53347 | 5950 | using compact_imp_bounded[OF assms(4)] unfolding bounded_pos by auto | 
| 5951 | obtain z :: 'a where z: "norm z = b + 1" | |
| 5952 | 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 | 5953 | 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 | 5954 | then obtain a b where ab: "inner a z < b" "\<forall>x\<in>T. b < inner a x" | 
| 53347 | 5955 | using separating_hyperplane_closed_point[OF assms(3) compact_imp_closed[OF assms(4)], of z] | 
| 5956 | by auto | |
| 5957 | then show ?thesis | |
| 5958 | using True by auto | |
| 33175 | 5959 | next | 
| 53347 | 5960 | 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 | 5961 | 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 | 5962 |   obtain a b where "0 < b" "\<forall>x \<in> (\<Union>x\<in> S. \<Union>y \<in> T. {x - y}). b < inner a x"
 | 
| 33175 | 5963 | using separating_hyperplane_closed_point[OF convex_differences[OF assms(1,3)], of 0] | 
| 53347 | 5964 | 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 | 5965 | 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 | 5966 | then have ab: "\<forall>x\<in>S. \<forall>y\<in>T. b + inner a y < inner a x" | 
| 53347 | 5967 | apply - | 
| 5968 | apply rule | |
| 5969 | apply rule | |
| 5970 | apply (erule_tac x="x - y" in ballE) | |
| 5971 | apply (auto simp add: inner_diff) | |
| 5972 | 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 | 5973 | define k where "k = (SUP x:T. a \<bullet> x)" | 
| 53347 | 5974 | show ?thesis | 
| 5975 | apply (rule_tac x="-a" in exI) | |
| 5976 | 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 | 5977 | apply (intro conjI ballI) | 
| 53347 | 5978 | unfolding inner_minus_left and neg_less_iff_less | 
| 5979 | 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 | 5980 | fix x assume "x \<in> T" | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54258diff
changeset | 5981 | then have "inner a x - b / 2 < k" | 
| 53347 | 5982 | unfolding k_def | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54258diff
changeset | 5983 | 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 | 5984 |       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 | 5985 | 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 | 5986 | 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 | 5987 | by (intro bdd_aboveI2[where M="inner a y - b"]) (auto simp: field_simps intro: less_imp_le) | 
| 60420 | 5988 | 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 | 5989 | 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 | 5990 | by auto | 
| 33175 | 5991 | next | 
| 53347 | 5992 | 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 | 5993 | assume "x \<in> S" | 
| 53347 | 5994 | then have "k \<le> inner a x - b" | 
| 5995 | unfolding k_def | |
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54258diff
changeset | 5996 | apply (rule_tac cSUP_least) | 
| 53347 | 5997 | using assms(5) | 
| 5998 | using ab[THEN bspec[where x=x]] | |
| 5999 | apply auto | |
| 6000 | done | |
| 6001 | then show "k + b / 2 < inner a x" | |
| 60420 | 6002 | using \<open>0 < b\<close> by auto | 
| 33175 | 6003 | qed | 
| 6004 | qed | |
| 6005 | ||
| 6006 | 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 | 6007 | 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 | 6008 | 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 | 6009 | 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 | 6010 |     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 | 6011 | 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 | 6012 | 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 | 6013 |     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 | 6014 | 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 | 6015 | 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 | 6016 | obtain a b where "(\<forall>x\<in>T. inner a x < b) \<and> (\<forall>x\<in>S. b < inner a x)" | 
| 53347 | 6017 | using separating_hyperplane_closed_compact[OF assms(4-5,1-2,3)] and assms(6) | 
| 6018 | by auto | |
| 6019 | then show ?thesis | |
| 6020 | apply (rule_tac x="-a" in exI) | |
| 6021 | apply (rule_tac x="-b" in exI) | |
| 6022 | apply auto | |
| 6023 | done | |
| 6024 | qed | |
| 6025 | ||
| 33175 | 6026 | |
| 60420 | 6027 | subsubsection \<open>General case without assuming closure and getting non-strict separation\<close> | 
| 33175 | 6028 | |
| 6029 | 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 | 6030 | assumes "convex s" "(0::'a::euclidean_space) \<notin> s" | 
| 33175 | 6031 | shows "\<exists>a. a \<noteq> 0 \<and> (\<forall>x\<in>s. 0 \<le> inner a x)" | 
| 53347 | 6032 | proof - | 
| 6033 |   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 | 6034 |   have *: "frontier (cball 0 1) \<inter> \<Inter>f \<noteq> {}" if as: "f \<subseteq> ?k ` s" "finite f" for f
 | 
| 53347 | 6035 | proof - | 
| 6036 | obtain c where c: "f = ?k ` c" "c \<subseteq> s" "finite c" | |
| 6037 | using finite_subset_image[OF as(2,1)] by auto | |
| 6038 | then obtain a b where ab: "a \<noteq> 0" "0 < b" "\<forall>x\<in>convex hull c. b < inner a x" | |
| 33175 | 6039 | using separating_hyperplane_closed_0[OF convex_convex_hull, of c] | 
| 6040 | using finite_imp_compact_convex_hull[OF c(3), THEN compact_imp_closed] and assms(2) | |
| 53347 | 6041 | 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 | 6042 | by force | 
| 53347 | 6043 | then have "\<exists>x. norm x = 1 \<and> (\<forall>y\<in>c. 0 \<le> inner y x)" | 
| 6044 | apply (rule_tac x = "inverse(norm a) *\<^sub>R a" in exI) | |
| 6045 | using hull_subset[of c convex] | |
| 6046 | unfolding subset_eq and inner_scaleR | |
| 56536 | 6047 | by (auto simp add: inner_commute del: ballE elim!: ballE) | 
| 53347 | 6048 |     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 | 6049 | 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 | 6050 | qed | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 6051 |   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 | 6052 | apply (rule compact_imp_fip) | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 6053 | apply (rule compact_frontier[OF compact_cball]) | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 6054 | using * closed_halfspace_ge | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 6055 | by auto | 
| 53347 | 6056 | 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 | 6057 | unfolding frontier_cball dist_norm sphere_def by auto | 
| 53347 | 6058 | then show ?thesis | 
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 6059 | by (metis inner_commute mem_Collect_eq norm_eq_zero zero_neq_one) | 
| 53347 | 6060 | qed | 
| 33175 | 6061 | |
| 6062 | lemma separating_hyperplane_sets: | |
| 53347 | 6063 | fixes s t :: "'a::euclidean_space set" | 
| 6064 | assumes "convex s" | |
| 6065 | and "convex t" | |
| 6066 |     and "s \<noteq> {}"
 | |
| 6067 |     and "t \<noteq> {}"
 | |
| 6068 |     and "s \<inter> t = {}"
 | |
| 33175 | 6069 | 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 | 6070 | proof - | 
| 6071 | from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]] | |
| 6072 |   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 | 6073 | using assms(3-5) by fastforce | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54258diff
changeset | 6074 | then have *: "\<And>x y. x \<in> t \<Longrightarrow> y \<in> s \<Longrightarrow> inner a y \<le> inner a x" | 
| 33270 | 6075 | by (force simp add: inner_diff) | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54258diff
changeset | 6076 | then have bdd: "bdd_above ((op \<bullet> a)`s)" | 
| 60420 | 6077 |     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 | 6078 | show ?thesis | 
| 60420 | 6079 | using \<open>a\<noteq>0\<close> | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54258diff
changeset | 6080 | by (intro exI[of _ a] exI[of _ "SUP x:s. a \<bullet> x"]) | 
| 60420 | 6081 |        (auto intro!: cSUP_upper bdd cSUP_least \<open>a \<noteq> 0\<close> \<open>s \<noteq> {}\<close> *)
 | 
| 6082 | qed | |
| 6083 | ||
| 6084 | ||
| 6085 | subsection \<open>More convexity generalities\<close> | |
| 33175 | 6086 | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 6087 | lemma convex_closure [intro,simp]: | 
| 33175 | 6088 | fixes s :: "'a::real_normed_vector set" | 
| 53347 | 6089 | assumes "convex s" | 
| 6090 | shows "convex (closure s)" | |
| 53676 | 6091 | apply (rule convexI) | 
| 6092 | apply (unfold closure_sequential, elim exE) | |
| 6093 | apply (rule_tac x="\<lambda>n. u *\<^sub>R xa n + v *\<^sub>R xb n" in exI) | |
| 53347 | 6094 | apply (rule,rule) | 
| 53676 | 6095 | apply (rule convexD [OF assms]) | 
| 53347 | 6096 | apply (auto del: tendsto_const intro!: tendsto_intros) | 
| 6097 | done | |
| 33175 | 6098 | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 6099 | lemma convex_interior [intro,simp]: | 
| 33175 | 6100 | fixes s :: "'a::real_normed_vector set" | 
| 53347 | 6101 | assumes "convex s" | 
| 6102 | shows "convex (interior s)" | |
| 6103 | unfolding convex_alt Ball_def mem_interior | |
| 6104 | apply (rule,rule,rule,rule,rule,rule) | |
| 6105 | apply (elim exE conjE) | |
| 6106 | proof - | |
| 6107 | fix x y u | |
| 6108 | assume u: "0 \<le> u" "u \<le> (1::real)" | |
| 6109 | fix e d | |
| 6110 | assume ed: "ball x e \<subseteq> s" "ball y d \<subseteq> s" "0<d" "0<e" | |
| 6111 | show "\<exists>e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \<subseteq> s" | |
| 6112 | apply (rule_tac x="min d e" in exI) | |
| 6113 | apply rule | |
| 6114 | unfolding subset_eq | |
| 6115 | defer | |
| 6116 | apply rule | |
| 6117 | proof - | |
| 6118 | fix z | |
| 6119 | assume "z \<in> ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) (min d e)" | |
| 6120 | then have "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \<in> s" | |
| 6121 | apply (rule_tac assms[unfolded convex_alt, rule_format]) | |
| 6122 | using ed(1,2) and u | |
| 6123 | unfolding subset_eq mem_ball Ball_def dist_norm | |
| 6124 | apply (auto simp add: algebra_simps) | |
| 6125 | done | |
| 6126 | then show "z \<in> s" | |
| 6127 | using u by (auto simp add: algebra_simps) | |
| 6128 | qed(insert u ed(3-4), auto) | |
| 6129 | qed | |
| 33175 | 6130 | |
| 34964 | 6131 | lemma convex_hull_eq_empty[simp]: "convex hull s = {} \<longleftrightarrow> s = {}"
 | 
| 33175 | 6132 | using hull_subset[of s convex] convex_hull_empty by auto | 
| 6133 | ||
| 53347 | 6134 | |
| 60420 | 6135 | subsection \<open>Moving and scaling convex hulls.\<close> | 
| 33175 | 6136 | |
| 53676 | 6137 | lemma convex_hull_set_plus: | 
| 6138 | "convex hull (s + t) = convex hull s + convex hull t" | |
| 6139 | unfolding set_plus_image | |
| 6140 | apply (subst convex_hull_linear_image [symmetric]) | |
| 6141 | apply (simp add: linear_iff scaleR_right_distrib) | |
| 6142 | apply (simp add: convex_hull_Times) | |
| 6143 | done | |
| 6144 | ||
| 6145 | lemma translation_eq_singleton_plus: "(\<lambda>x. a + x) ` t = {a} + t"
 | |
| 6146 | unfolding set_plus_def by auto | |
| 33175 | 6147 | |
| 6148 | lemma convex_hull_translation: | |
| 6149 | "convex hull ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (convex hull s)" | |
| 53676 | 6150 | unfolding translation_eq_singleton_plus | 
| 6151 | by (simp only: convex_hull_set_plus convex_hull_singleton) | |
| 33175 | 6152 | |
| 6153 | lemma convex_hull_scaling: | |
| 6154 | "convex hull ((\<lambda>x. c *\<^sub>R x) ` s) = (\<lambda>x. c *\<^sub>R x) ` (convex hull s)" | |
| 53676 | 6155 | using linear_scaleR by (rule convex_hull_linear_image [symmetric]) | 
| 33175 | 6156 | |
| 6157 | lemma convex_hull_affinity: | |
| 6158 | "convex hull ((\<lambda>x. a + c *\<^sub>R x) ` s) = (\<lambda>x. a + c *\<^sub>R x) ` (convex hull s)" | |
| 53347 | 6159 | by(simp only: image_image[symmetric] convex_hull_scaling convex_hull_translation) | 
| 6160 | ||
| 33175 | 6161 | |
| 60420 | 6162 | subsection \<open>Convexity of cone hulls\<close> | 
| 40377 | 6163 | |
| 6164 | lemma convex_cone_hull: | |
| 53347 | 6165 | assumes "convex S" | 
| 6166 | shows "convex (cone hull S)" | |
| 53676 | 6167 | proof (rule convexI) | 
| 6168 | fix x y | |
| 6169 | assume xy: "x \<in> cone hull S" "y \<in> cone hull S" | |
| 6170 |   then have "S \<noteq> {}"
 | |
| 6171 | using cone_hull_empty_iff[of S] by auto | |
| 6172 | fix u v :: real | |
| 6173 | assume uv: "u \<ge> 0" "v \<ge> 0" "u + v = 1" | |
| 6174 | then have *: "u *\<^sub>R x \<in> cone hull S" "v *\<^sub>R y \<in> cone hull S" | |
| 6175 | using cone_cone_hull[of S] xy cone_def[of "cone hull S"] by auto | |
| 6176 | from * obtain cx :: real and xx where x: "u *\<^sub>R x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S" | |
| 6177 | using cone_hull_expl[of S] by auto | |
| 6178 | from * obtain cy :: real and yy where y: "v *\<^sub>R y = cy *\<^sub>R yy" "cy \<ge> 0" "yy \<in> S" | |
| 6179 | using cone_hull_expl[of S] by auto | |
| 53347 | 6180 |   {
 | 
| 53676 | 6181 | assume "cx + cy \<le> 0" | 
| 6182 | then have "u *\<^sub>R x = 0" and "v *\<^sub>R y = 0" | |
| 6183 | using x y by auto | |
| 6184 | then have "u *\<^sub>R x + v *\<^sub>R y = 0" | |
| 6185 | by auto | |
| 6186 | then have "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S" | |
| 60420 | 6187 |       using cone_hull_contains_0[of S] \<open>S \<noteq> {}\<close> by auto
 | 
| 40377 | 6188 | } | 
| 53676 | 6189 | moreover | 
| 6190 |   {
 | |
| 6191 | assume "cx + cy > 0" | |
| 6192 | then have "(cx / (cx + cy)) *\<^sub>R xx + (cy / (cx + cy)) *\<^sub>R yy \<in> S" | |
| 6193 | using assms mem_convex_alt[of S xx yy cx cy] x y by auto | |
| 6194 | then have "cx *\<^sub>R xx + cy *\<^sub>R yy \<in> cone hull S" | |
| 60420 | 6195 | 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 | 6196 | by (auto simp add: scaleR_right_distrib) | 
| 6197 | then have "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S" | |
| 6198 | using x y by auto | |
| 6199 | } | |
| 6200 | moreover have "cx + cy \<le> 0 \<or> cx + cy > 0" by auto | |
| 6201 | ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S" by blast | |
| 40377 | 6202 | qed | 
| 6203 | ||
| 6204 | lemma cone_convex_hull: | |
| 53347 | 6205 | assumes "cone S" | 
| 6206 | shows "cone (convex hull S)" | |
| 6207 | proof (cases "S = {}")
 | |
| 6208 | case True | |
| 6209 | then show ?thesis by auto | |
| 6210 | next | |
| 6211 | case False | |
| 54465 | 6212 | then have *: "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` S = S)" | 
| 6213 | using cone_iff[of S] assms by auto | |
| 53347 | 6214 |   {
 | 
| 6215 | fix c :: real | |
| 6216 | assume "c > 0" | |
| 6217 | then have "op *\<^sub>R c ` (convex hull S) = convex hull (op *\<^sub>R c ` S)" | |
| 6218 | using convex_hull_scaling[of _ S] by auto | |
| 6219 | also have "\<dots> = convex hull S" | |
| 60420 | 6220 | using * \<open>c > 0\<close> by auto | 
| 53347 | 6221 | finally have "op *\<^sub>R c ` (convex hull S) = convex hull S" | 
| 6222 | by auto | |
| 40377 | 6223 | } | 
| 53347 | 6224 | then have "0 \<in> convex hull S" "\<And>c. c > 0 \<Longrightarrow> (op *\<^sub>R c ` (convex hull S)) = (convex hull S)" | 
| 6225 | using * hull_subset[of S convex] by auto | |
| 6226 | then show ?thesis | |
| 60420 | 6227 |     using \<open>S \<noteq> {}\<close> cone_iff[of "convex hull S"] by auto
 | 
| 6228 | qed | |
| 6229 | ||
| 6230 | subsection \<open>Convex set as intersection of halfspaces\<close> | |
| 33175 | 6231 | |
| 6232 | 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 | 6233 |   fixes s :: "('a::euclidean_space) set"
 | 
| 33175 | 6234 | assumes "closed s" "convex s" | 
| 60585 | 6235 |   shows "s = \<Inter>{h. s \<subseteq> h \<and> (\<exists>a b. h = {x. inner a x \<le> b})}"
 | 
| 53347 | 6236 | apply (rule set_eqI) | 
| 6237 | apply rule | |
| 6238 | unfolding Inter_iff Ball_def mem_Collect_eq | |
| 6239 | apply (rule,rule,erule conjE) | |
| 6240 | proof - | |
| 54465 | 6241 | fix x | 
| 53347 | 6242 |   assume "\<forall>xa. s \<subseteq> xa \<and> (\<exists>a b. xa = {x. inner a x \<le> b}) \<longrightarrow> x \<in> xa"
 | 
| 6243 |   then have "\<forall>a b. s \<subseteq> {x. inner a x \<le> b} \<longrightarrow> x \<in> {x. inner a x \<le> b}"
 | |
| 6244 | by blast | |
| 6245 | then show "x \<in> s" | |
| 6246 | apply (rule_tac ccontr) | |
| 6247 | apply (drule separating_hyperplane_closed_point[OF assms(2,1)]) | |
| 6248 | apply (erule exE)+ | |
| 6249 | apply (erule_tac x="-a" in allE) | |
| 6250 | apply (erule_tac x="-b" in allE) | |
| 6251 | apply auto | |
| 6252 | done | |
| 33175 | 6253 | qed auto | 
| 6254 | ||
| 53347 | 6255 | |
| 60420 | 6256 | subsection \<open>Radon's theorem (from Lars Schewe)\<close> | 
| 33175 | 6257 | |
| 6258 | lemma radon_ex_lemma: | |
| 6259 | assumes "finite c" "affine_dependent c" | |
| 64267 | 6260 | 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 | 6261 | proof - | 
| 55697 | 6262 | from assms(2)[unfolded affine_dependent_explicit] | 
| 6263 | obtain s u where | |
| 64267 | 6264 | "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 | 6265 | by blast | 
| 53347 | 6266 | then show ?thesis | 
| 6267 | apply (rule_tac x="\<lambda>v. if v\<in>s then u v else 0" in exI) | |
| 64267 | 6268 | unfolding if_smult scaleR_zero_left and sum.inter_restrict[OF assms(1), symmetric] | 
| 53347 | 6269 | apply (auto simp add: Int_absorb1) | 
| 6270 | done | |
| 6271 | qed | |
| 33175 | 6272 | |
| 6273 | lemma radon_s_lemma: | |
| 53347 | 6274 | assumes "finite s" | 
| 64267 | 6275 | and "sum f s = (0::real)" | 
| 6276 |   shows "sum f {x\<in>s. 0 < f x} = - sum f {x\<in>s. f x < 0}"
 | |
| 53347 | 6277 | proof - | 
| 6278 | have *: "\<And>x. (if f x < 0 then f x else 0) + (if 0 < f x then f x else 0) = f x" | |
| 6279 | by auto | |
| 6280 | show ?thesis | |
| 64267 | 6281 | unfolding add_eq_0_iff[symmetric] and sum.inter_filter[OF assms(1)] | 
| 6282 | and sum.distrib[symmetric] and * | |
| 53347 | 6283 | 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 | 6284 | by assumption | 
| 53347 | 6285 | qed | 
| 33175 | 6286 | |
| 6287 | lemma radon_v_lemma: | |
| 53347 | 6288 | assumes "finite s" | 
| 64267 | 6289 | and "sum f s = 0" | 
| 53347 | 6290 | and "\<forall>x. g x = (0::real) \<longrightarrow> f x = (0::'a::euclidean_space)" | 
| 64267 | 6291 |   shows "(sum f {x\<in>s. 0 < g x}) = - sum f {x\<in>s. g x < 0}"
 | 
| 53347 | 6292 | proof - | 
| 6293 | have *: "\<And>x. (if 0 < g x then f x else 0) + (if g x < 0 then f x else 0) = f x" | |
| 6294 | using assms(3) by auto | |
| 6295 | show ?thesis | |
| 64267 | 6296 | unfolding eq_neg_iff_add_eq_0 and sum.inter_filter[OF assms(1)] | 
| 6297 | and sum.distrib[symmetric] and * | |
| 53347 | 6298 | using assms(2) | 
| 6299 | apply assumption | |
| 6300 | done | |
| 6301 | qed | |
| 33175 | 6302 | |
| 6303 | lemma radon_partition: | |
| 6304 | assumes "finite c" "affine_dependent c" | |
| 53347 | 6305 |   shows "\<exists>m p. m \<inter> p = {} \<and> m \<union> p = c \<and> (convex hull m) \<inter> (convex hull p) \<noteq> {}"
 | 
| 6306 | proof - | |
| 64267 | 6307 | 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 | 6308 | using radon_ex_lemma[OF assms] by auto | 
| 6309 |   have fin: "finite {x \<in> c. 0 < u x}" "finite {x \<in> c. 0 > u x}"
 | |
| 6310 | using assms(1) by auto | |
| 64267 | 6311 |   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}"
 | 
| 6312 |   have "sum u {x \<in> c. 0 < u x} \<noteq> 0"
 | |
| 53347 | 6313 | proof (cases "u v \<ge> 0") | 
| 6314 | case False | |
| 6315 | then have "u v < 0" by auto | |
| 6316 | then show ?thesis | |
| 6317 |     proof (cases "\<exists>w\<in>{x \<in> c. 0 < u x}. u w > 0")
 | |
| 6318 | case True | |
| 6319 | then show ?thesis | |
| 64267 | 6320 | using sum_nonneg_eq_0_iff[of _ u, OF fin(1)] by auto | 
| 33175 | 6321 | next | 
| 53347 | 6322 | case False | 
| 64267 | 6323 | then have "sum u c \<le> sum (\<lambda>x. if x=v then u v else 0) c" | 
| 6324 | apply (rule_tac sum_mono) | |
| 53347 | 6325 | apply auto | 
| 6326 | done | |
| 6327 | then show ?thesis | |
| 64267 | 6328 | unfolding sum.delta[OF assms(1)] using uv(2) and \<open>u v < 0\<close> and uv(1) by auto | 
| 53347 | 6329 | qed | 
| 64267 | 6330 | qed (insert sum_nonneg_eq_0_iff[of _ u, OF fin(1)] uv(2-3), auto) | 
| 6331 | ||
| 6332 |   then have *: "sum u {x\<in>c. u x > 0} > 0"
 | |
| 53347 | 6333 | unfolding less_le | 
| 6334 | apply (rule_tac conjI) | |
| 64267 | 6335 | apply (rule_tac sum_nonneg) | 
| 6336 | apply auto | |
| 6337 | done | |
| 6338 |   moreover have "sum u ({x \<in> c. 0 < u x} \<union> {x \<in> c. u x < 0}) = sum u c"
 | |
| 33175 | 6339 |     "(\<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 | 6340 | using assms(1) | 
| 64267 | 6341 | apply (rule_tac[!] sum.mono_neutral_left) | 
| 6342 | apply auto | |
| 6343 | done | |
| 6344 |   then have "sum u {x \<in> c. 0 < u x} = - sum u {x \<in> c. 0 > u x}"
 | |
| 53347 | 6345 |     "(\<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)"
 | 
| 6346 | unfolding eq_neg_iff_add_eq_0 | |
| 6347 | using uv(1,4) | |
| 64267 | 6348 | by (auto simp add: sum.union_inter_neutral[OF fin, symmetric]) | 
| 6349 |   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 | 6350 | apply rule | 
| 6351 | apply (rule mult_nonneg_nonneg) | |
| 6352 | using * | |
| 6353 | apply auto | |
| 6354 | done | |
| 6355 |   ultimately have "z \<in> convex hull {v \<in> c. u v \<le> 0}"
 | |
| 6356 | unfolding convex_hull_explicit mem_Collect_eq | |
| 6357 |     apply (rule_tac x="{v \<in> c. u v < 0}" in exI)
 | |
| 64267 | 6358 |     apply (rule_tac x="\<lambda>y. inverse (sum u {x\<in>c. u x > 0}) * - u y" in exI)
 | 
| 6359 | using assms(1) unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] and z_def | |
| 6360 | apply (auto simp add: sum_negf sum_distrib_left[symmetric]) | |
| 6361 | done | |
| 6362 |   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 | 6363 | apply rule | 
| 6364 | apply (rule mult_nonneg_nonneg) | |
| 6365 | using * | |
| 6366 | apply auto | |
| 6367 | done | |
| 6368 |   then have "z \<in> convex hull {v \<in> c. u v > 0}"
 | |
| 6369 | unfolding convex_hull_explicit mem_Collect_eq | |
| 6370 |     apply (rule_tac x="{v \<in> c. 0 < u v}" in exI)
 | |
| 64267 | 6371 |     apply (rule_tac x="\<lambda>y. inverse (sum u {x\<in>c. u x > 0}) * u y" in exI)
 | 
| 53347 | 6372 | using assms(1) | 
| 64267 | 6373 | unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] and z_def | 
| 53347 | 6374 | using * | 
| 64267 | 6375 | apply (auto simp add: sum_negf sum_distrib_left[symmetric]) | 
| 53347 | 6376 | done | 
| 6377 | ultimately show ?thesis | |
| 6378 |     apply (rule_tac x="{v\<in>c. u v \<le> 0}" in exI)
 | |
| 6379 |     apply (rule_tac x="{v\<in>c. u v > 0}" in exI)
 | |
| 6380 | apply auto | |
| 6381 | done | |
| 6382 | qed | |
| 6383 | ||
| 6384 | lemma radon: | |
| 6385 | assumes "affine_dependent c" | |
| 6386 |   obtains m p where "m \<subseteq> c" "p \<subseteq> c" "m \<inter> p = {}" "(convex hull m) \<inter> (convex hull p) \<noteq> {}"
 | |
| 6387 | proof - | |
| 55697 | 6388 | from assms[unfolded affine_dependent_explicit] | 
| 6389 | obtain s u where | |
| 64267 | 6390 | "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 | 6391 | by blast | 
| 53347 | 6392 | then have *: "finite s" "affine_dependent s" and s: "s \<subseteq> c" | 
| 6393 | unfolding affine_dependent_explicit by auto | |
| 55697 | 6394 | from radon_partition[OF *] | 
| 6395 |   obtain m p where "m \<inter> p = {}" "m \<union> p = s" "convex hull m \<inter> convex hull p \<noteq> {}"
 | |
| 6396 | by blast | |
| 53347 | 6397 | then show ?thesis | 
| 6398 | apply (rule_tac that[of p m]) | |
| 6399 | using s | |
| 6400 | apply auto | |
| 6401 | done | |
| 6402 | qed | |
| 6403 | ||
| 33175 | 6404 | |
| 60420 | 6405 | subsection \<open>Helly's theorem\<close> | 
| 33175 | 6406 | |
| 53347 | 6407 | lemma helly_induct: | 
| 6408 | fixes f :: "'a::euclidean_space set set" | |
| 6409 | assumes "card f = n" | |
| 6410 |     and "n \<ge> DIM('a) + 1"
 | |
| 60585 | 6411 |     and "\<forall>s\<in>f. convex s" "\<forall>t\<subseteq>f. card t = DIM('a) + 1 \<longrightarrow> \<Inter>t \<noteq> {}"
 | 
| 53347 | 6412 |   shows "\<Inter>f \<noteq> {}"
 | 
| 6413 | using assms | |
| 6414 | proof (induct n arbitrary: f) | |
| 6415 | case 0 | |
| 6416 | then show ?case by auto | |
| 6417 | next | |
| 6418 | case (Suc n) | |
| 6419 | have "finite f" | |
| 60420 | 6420 | using \<open>card f = Suc n\<close> by (auto intro: card_ge_0_finite) | 
| 53347 | 6421 |   show "\<Inter>f \<noteq> {}"
 | 
| 6422 |     apply (cases "n = DIM('a)")
 | |
| 6423 | apply (rule Suc(5)[rule_format]) | |
| 60420 | 6424 | unfolding \<open>card f = Suc n\<close> | 
| 53347 | 6425 | proof - | 
| 6426 |     assume ng: "n \<noteq> DIM('a)"
 | |
| 6427 |     then have "\<exists>X. \<forall>s\<in>f. X s \<in> \<Inter>(f - {s})"
 | |
| 6428 | apply (rule_tac bchoice) | |
| 6429 | unfolding ex_in_conv | |
| 6430 | apply (rule, rule Suc(1)[rule_format]) | |
| 60420 | 6431 | unfolding card_Diff_singleton_if[OF \<open>finite f\<close>] \<open>card f = Suc n\<close> | 
| 53347 | 6432 | defer | 
| 6433 | defer | |
| 6434 | apply (rule Suc(4)[rule_format]) | |
| 6435 | defer | |
| 6436 | apply (rule Suc(5)[rule_format]) | |
| 60420 | 6437 | using Suc(3) \<open>finite f\<close> | 
| 53347 | 6438 | apply auto | 
| 6439 | done | |
| 6440 |     then obtain X where X: "\<forall>s\<in>f. X s \<in> \<Inter>(f - {s})" by auto
 | |
| 6441 | show ?thesis | |
| 6442 | proof (cases "inj_on X f") | |
| 6443 | case False | |
| 6444 | then obtain s t where st: "s\<noteq>t" "s\<in>f" "t\<in>f" "X s = X t" | |
| 6445 | unfolding inj_on_def by auto | |
| 6446 |       then have *: "\<Inter>f = \<Inter>(f - {s}) \<inter> \<Inter>(f - {t})" by auto
 | |
| 6447 | show ?thesis | |
| 6448 | unfolding * | |
| 6449 | unfolding ex_in_conv[symmetric] | |
| 6450 | apply (rule_tac x="X s" in exI) | |
| 6451 | apply rule | |
| 6452 | apply (rule X[rule_format]) | |
| 6453 | using X st | |
| 6454 | apply auto | |
| 6455 | done | |
| 6456 | next | |
| 6457 | case True | |
| 6458 |       then obtain m p where mp: "m \<inter> p = {}" "m \<union> p = X ` f" "convex hull m \<inter> convex hull p \<noteq> {}"
 | |
| 6459 | using radon_partition[of "X ` f"] and affine_dependent_biggerset[of "X ` f"] | |
| 60420 | 6460 | unfolding card_image[OF True] and \<open>card f = Suc n\<close> | 
| 6461 | using Suc(3) \<open>finite f\<close> and ng | |
| 53347 | 6462 | by auto | 
| 6463 | have "m \<subseteq> X ` f" "p \<subseteq> X ` f" | |
| 6464 | using mp(2) by auto | |
| 6465 | then obtain g h where gh:"m = X ` g" "p = X ` h" "g \<subseteq> f" "h \<subseteq> f" | |
| 6466 | unfolding subset_image_iff by auto | |
| 6467 | then have "f \<union> (g \<union> h) = f" by auto | |
| 6468 | then have f: "f = g \<union> h" | |
| 6469 | using inj_on_Un_image_eq_iff[of X f "g \<union> h"] and True | |
| 6470 | unfolding mp(2)[unfolded image_Un[symmetric] gh] | |
| 6471 | by auto | |
| 6472 |       have *: "g \<inter> h = {}"
 | |
| 6473 | using mp(1) | |
| 6474 | unfolding gh | |
| 6475 | using inj_on_image_Int[OF True gh(3,4)] | |
| 6476 | by auto | |
| 6477 | have "convex hull (X ` h) \<subseteq> \<Inter>g" "convex hull (X ` g) \<subseteq> \<Inter>h" | |
| 6478 | apply (rule_tac [!] hull_minimal) | |
| 6479 | using Suc gh(3-4) | |
| 6480 | unfolding subset_eq | |
| 6481 | apply (rule_tac [2] convex_Inter, rule_tac [4] convex_Inter) | |
| 6482 | apply rule | |
| 6483 | prefer 3 | |
| 6484 | apply rule | |
| 6485 | proof - | |
| 6486 | fix x | |
| 6487 | assume "x \<in> X ` g" | |
| 55697 | 6488 | then obtain y where "y \<in> g" "x = X y" | 
| 6489 | unfolding image_iff .. | |
| 53347 | 6490 | then show "x \<in> \<Inter>h" | 
| 6491 | using X[THEN bspec[where x=y]] using * f by auto | |
| 6492 | next | |
| 6493 | fix x | |
| 6494 | assume "x \<in> X ` h" | |
| 55697 | 6495 | then obtain y where "y \<in> h" "x = X y" | 
| 6496 | unfolding image_iff .. | |
| 53347 | 6497 | then show "x \<in> \<Inter>g" | 
| 6498 | using X[THEN bspec[where x=y]] using * f by auto | |
| 6499 | qed auto | |
| 6500 | then show ?thesis | |
| 6501 | unfolding f using mp(3)[unfolded gh] by blast | |
| 6502 | qed | |
| 6503 | qed auto | |
| 6504 | qed | |
| 6505 | ||
| 6506 | lemma helly: | |
| 6507 | 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 | 6508 |   assumes "card f \<ge> DIM('a) + 1" "\<forall>s\<in>f. convex s"
 | 
| 60585 | 6509 |     and "\<forall>t\<subseteq>f. card t = DIM('a) + 1 \<longrightarrow> \<Inter>t \<noteq> {}"
 | 
| 53347 | 6510 |   shows "\<Inter>f \<noteq> {}"
 | 
| 6511 | apply (rule helly_induct) | |
| 6512 | using assms | |
| 6513 | apply auto | |
| 6514 | done | |
| 6515 | ||
| 33175 | 6516 | |
| 60420 | 6517 | subsection \<open>Epigraphs of convex functions\<close> | 
| 33175 | 6518 | |
| 53348 | 6519 | definition "epigraph s (f :: _ \<Rightarrow> real) = {xy. fst xy \<in> s \<and> f (fst xy) \<le> snd xy}"
 | 
| 6520 | ||
| 6521 | lemma mem_epigraph: "(x, y) \<in> epigraph s f \<longleftrightarrow> x \<in> s \<and> f x \<le> y" | |
| 6522 | unfolding epigraph_def by auto | |
| 6523 | ||
| 6524 | lemma convex_epigraph: "convex (epigraph s f) \<longleftrightarrow> convex_on s f \<and> convex s" | |
| 36338 | 6525 | unfolding convex_def convex_on_def | 
| 6526 | unfolding Ball_def split_paired_All epigraph_def | |
| 6527 | unfolding mem_Collect_eq fst_conv snd_conv fst_add snd_add fst_scaleR snd_scaleR Ball_def[symmetric] | |
| 53348 | 6528 | apply safe | 
| 6529 | defer | |
| 6530 | apply (erule_tac x=x in allE) | |
| 6531 | apply (erule_tac x="f x" in allE) | |
| 6532 | apply safe | |
| 6533 | apply (erule_tac x=xa in allE) | |
| 6534 | apply (erule_tac x="f xa" in allE) | |
| 6535 | prefer 3 | |
| 6536 | apply (rule_tac y="u * f a + v * f aa" in order_trans) | |
| 6537 | defer | |
| 6538 | apply (auto intro!:mult_left_mono add_mono) | |
| 6539 | done | |
| 6540 | ||
| 6541 | lemma convex_epigraphI: "convex_on s f \<Longrightarrow> convex s \<Longrightarrow> convex (epigraph s f)" | |
| 6542 | unfolding convex_epigraph by auto | |
| 6543 | ||
| 6544 | lemma convex_epigraph_convex: "convex s \<Longrightarrow> convex_on s f \<longleftrightarrow> convex(epigraph s f)" | |
| 6545 | by (simp add: convex_epigraph) | |
| 6546 | ||
| 33175 | 6547 | |
| 60420 | 6548 | subsubsection \<open>Use this to derive general bound property of convex function\<close> | 
| 33175 | 6549 | |
| 6550 | lemma convex_on: | |
| 6551 | assumes "convex s" | |
| 53348 | 6552 | shows "convex_on s f \<longleftrightarrow> | 
| 64267 | 6553 |     (\<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>
 | 
| 6554 |       f (sum (\<lambda>i. u i *\<^sub>R x i) {1..k} ) \<le> sum (\<lambda>i. u i * f(x i)) {1..k})"
 | |
| 33175 | 6555 | unfolding convex_epigraph_convex[OF assms] convex epigraph_def Ball_def mem_Collect_eq | 
| 64267 | 6556 | unfolding fst_sum snd_sum fst_scaleR snd_scaleR | 
| 36338 | 6557 | apply safe | 
| 6558 | apply (drule_tac x=k in spec) | |
| 6559 | apply (drule_tac x=u in spec) | |
| 6560 | apply (drule_tac x="\<lambda>i. (x i, f (x i))" in spec) | |
| 6561 | apply simp | |
| 53348 | 6562 | using assms[unfolded convex] | 
| 6563 | apply simp | |
| 6564 | apply (rule_tac y="\<Sum>i = 1..k. u i * f (fst (x i))" in order_trans) | |
| 6565 | defer | |
| 64267 | 6566 | apply (rule sum_mono) | 
| 53348 | 6567 | apply (erule_tac x=i in allE) | 
| 6568 | unfolding real_scaleR_def | |
| 6569 | apply (rule mult_left_mono) | |
| 6570 | using assms[unfolded convex] | |
| 6571 | apply auto | |
| 6572 | done | |
| 33175 | 6573 | |
| 36338 | 6574 | |
| 60420 | 6575 | subsection \<open>Convexity of general and special intervals\<close> | 
| 33175 | 6576 | |
| 6577 | lemma is_interval_convex: | |
| 53348 | 6578 | fixes s :: "'a::euclidean_space set" | 
| 6579 | assumes "is_interval s" | |
| 6580 | shows "convex s" | |
| 37732 
6432bf0d7191
generalize type of is_interval to class euclidean_space
 huffman parents: 
37673diff
changeset | 6581 | proof (rule convexI) | 
| 53348 | 6582 | fix x y and u v :: real | 
| 6583 | assume as: "x \<in> s" "y \<in> s" "0 \<le> u" "0 \<le> v" "u + v = 1" | |
| 6584 | then have *: "u = 1 - v" "1 - v \<ge> 0" and **: "v = 1 - u" "1 - u \<ge> 0" | |
| 6585 | by auto | |
| 6586 |   {
 | |
| 6587 | fix a b | |
| 6588 | assume "\<not> b \<le> u * a + v * b" | |
| 6589 | then have "u * a < (1 - v) * b" | |
| 6590 | unfolding not_le using as(4) by (auto simp add: field_simps) | |
| 6591 | then have "a < b" | |
| 6592 | unfolding * using as(4) *(2) | |
| 6593 | apply (rule_tac mult_left_less_imp_less[of "1 - v"]) | |
| 6594 | apply (auto simp add: field_simps) | |
| 6595 | done | |
| 6596 | then have "a \<le> u * a + v * b" | |
| 6597 | unfolding * using as(4) | |
| 6598 | by (auto simp add: field_simps intro!:mult_right_mono) | |
| 6599 | } | |
| 6600 | moreover | |
| 6601 |   {
 | |
| 6602 | fix a b | |
| 6603 | assume "\<not> u * a + v * b \<le> a" | |
| 6604 | then have "v * b > (1 - u) * a" | |
| 6605 | unfolding not_le using as(4) by (auto simp add: field_simps) | |
| 6606 | then have "a < b" | |
| 6607 | unfolding * using as(4) | |
| 6608 | apply (rule_tac mult_left_less_imp_less) | |
| 6609 | apply (auto simp add: field_simps) | |
| 6610 | done | |
| 6611 | then have "u * a + v * b \<le> b" | |
| 6612 | unfolding ** | |
| 6613 | using **(2) as(3) | |
| 6614 | by (auto simp add: field_simps intro!:mult_right_mono) | |
| 6615 | } | |
| 6616 | ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> s" | |
| 6617 | apply - | |
| 6618 | apply (rule assms[unfolded is_interval_def, rule_format, OF as(1,2)]) | |
| 6619 | using as(3-) DIM_positive[where 'a='a] | |
| 6620 | apply (auto simp: inner_simps) | |
| 6621 | 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 | 6622 | qed | 
| 33175 | 6623 | |
| 6624 | lemma is_interval_connected: | |
| 53348 | 6625 | fixes s :: "'a::euclidean_space set" | 
| 33175 | 6626 | shows "is_interval s \<Longrightarrow> connected s" | 
| 6627 | using is_interval_convex convex_connected by auto | |
| 6628 | ||
| 62618 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 6629 | lemma convex_box [simp]: "convex (cbox a b)" "convex (box a (b::'a::euclidean_space))" | 
| 56188 | 6630 | apply (rule_tac[!] is_interval_convex)+ | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 6631 | using is_interval_box is_interval_cbox | 
| 53348 | 6632 | apply auto | 
| 6633 | done | |
| 33175 | 6634 | |
| 63928 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6635 | 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 | 6636 | lemma connected_imp_perfect: | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6637 | fixes a :: "'a::metric_space" | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6638 |   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 | 6639 | shows "a islimpt S" | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6640 | proof - | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6641 | 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 | 6642 | proof - | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6643 | 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 | 6644 | 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 | 6645 |     have "openin (subtopology euclidean S) {a}"
 | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6646 | 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 | 6647 |     moreover have "closedin (subtopology euclidean S) {a}"
 | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6648 | by (simp add: assms) | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6649 | ultimately show "False" | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6650 | 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 | 6651 | qed | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6652 | then show ?thesis | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6653 | unfolding islimpt_def by blast | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6654 | qed | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6655 | |
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6656 | lemma connected_imp_perfect_aff_dim: | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6657 | "\<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 | 6658 | using aff_dim_sing connected_imp_perfect by blast | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 6659 | |
| 61808 | 6660 | subsection \<open>On \<open>real\<close>, \<open>is_interval\<close>, \<open>convex\<close> and \<open>connected\<close> are all equivalent.\<close> | 
| 33175 | 6661 | |
| 6662 | lemma is_interval_1: | |
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6663 | "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 | 6664 | unfolding is_interval_def by auto | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6665 | |
| 54465 | 6666 | lemma is_interval_connected_1: | 
| 6667 | fixes s :: "real set" | |
| 6668 | shows "is_interval s \<longleftrightarrow> connected s" | |
| 6669 | apply rule | |
| 6670 | apply (rule is_interval_connected, assumption) | |
| 6671 | unfolding is_interval_1 | |
| 6672 | apply rule | |
| 6673 | apply rule | |
| 6674 | apply rule | |
| 6675 | apply rule | |
| 6676 | apply (erule conjE) | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 6677 | apply (rule ccontr) | 
| 54465 | 6678 | proof - | 
| 6679 | fix a b x | |
| 6680 | assume as: "connected s" "a \<in> s" "b \<in> s" "a \<le> x" "x \<le> b" "x \<notin> s" | |
| 6681 | then have *: "a < x" "x < b" | |
| 6682 | unfolding not_le [symmetric] by auto | |
| 6683 |   let ?halfl = "{..<x} "
 | |
| 6684 |   let ?halfr = "{x<..}"
 | |
| 6685 |   {
 | |
| 6686 | fix y | |
| 6687 | assume "y \<in> s" | |
| 60420 | 6688 | with \<open>x \<notin> s\<close> have "x \<noteq> y" by auto | 
| 54465 | 6689 | then have "y \<in> ?halfr \<union> ?halfl" by auto | 
| 6690 | } | |
| 6691 | moreover have "a \<in> ?halfl" "b \<in> ?halfr" using * by auto | |
| 6692 |   then have "?halfl \<inter> s \<noteq> {}" "?halfr \<inter> s \<noteq> {}"
 | |
| 6693 | using as(2-3) by auto | |
| 6694 | ultimately show False | |
| 6695 | apply (rule_tac notE[OF as(1)[unfolded connected_def]]) | |
| 6696 | apply (rule_tac x = ?halfl in exI) | |
| 6697 | apply (rule_tac x = ?halfr in exI) | |
| 6698 | apply rule | |
| 6699 | apply (rule open_lessThan) | |
| 6700 | apply rule | |
| 6701 | apply (rule open_greaterThan) | |
| 6702 | apply auto | |
| 6703 | done | |
| 6704 | qed | |
| 33175 | 6705 | |
| 6706 | lemma is_interval_convex_1: | |
| 54465 | 6707 | fixes s :: "real set" | 
| 6708 | shows "is_interval s \<longleftrightarrow> convex s" | |
| 6709 | by (metis is_interval_convex convex_connected is_interval_connected_1) | |
| 33175 | 6710 | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 6711 | lemma connected_compact_interval_1: | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 6712 |      "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 | 6713 | 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 | 6714 | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6715 | lemma connected_convex_1: | 
| 54465 | 6716 | fixes s :: "real set" | 
| 6717 | shows "connected s \<longleftrightarrow> convex s" | |
| 6718 | by (metis is_interval_convex convex_connected is_interval_connected_1) | |
| 53348 | 6719 | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6720 | lemma connected_convex_1_gen: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6721 | fixes s :: "'a :: euclidean_space set" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6722 |   assumes "DIM('a) = 1"
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6723 | shows "connected s \<longleftrightarrow> convex s" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6724 | proof - | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6725 | 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 | 6726 | using subspace_isomorphism [where 'a = 'a and 'b = real] | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6727 | by (metis DIM_real dim_UNIV subspace_UNIV assms) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6728 | then have "f -` (f ` s) = s" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6729 | by (simp add: inj_vimage_image_eq) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6730 | then show ?thesis | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6731 | 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 | 6732 | qed | 
| 53348 | 6733 | |
| 60420 | 6734 | subsection \<open>Another intermediate value theorem formulation\<close> | 
| 33175 | 6735 | |
| 53348 | 6736 | 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 | 6737 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" | 
| 53348 | 6738 | assumes "a \<le> b" | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6739 |     and "continuous_on {a..b} f"
 | 
| 53348 | 6740 | 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 | 6741 |   shows "\<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
 | 
| 56188 | 6742 | proof - | 
| 6743 | have "f a \<in> f ` cbox a b" "f b \<in> f ` cbox a b" | |
| 53348 | 6744 | apply (rule_tac[!] imageI) | 
| 6745 | using assms(1) | |
| 6746 | apply auto | |
| 6747 | done | |
| 6748 | then show ?thesis | |
| 56188 | 6749 | 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 | 6750 | by (simp add: connected_continuous_image assms) | 
| 53348 | 6751 | qed | 
| 6752 | ||
| 6753 | lemma ivt_increasing_component_1: | |
| 6754 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6755 |   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 | 6756 |     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 | 6757 | by (rule ivt_increasing_component_on_1) (auto simp add: continuous_at_imp_continuous_on) | 
| 6758 | ||
| 6759 | lemma ivt_decreasing_component_on_1: | |
| 6760 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" | |
| 6761 | assumes "a \<le> b" | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6762 |     and "continuous_on {a..b} f"
 | 
| 53348 | 6763 | and "(f b)\<bullet>k \<le> y" | 
| 6764 | and "y \<le> (f a)\<bullet>k" | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6765 |   shows "\<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
 | 
| 53348 | 6766 | apply (subst neg_equal_iff_equal[symmetric]) | 
| 44531 
1d477a2b1572
replace some continuous_on lemmas with more general versions
 huffman parents: 
44525diff
changeset | 6767 | using ivt_increasing_component_on_1[of a b "\<lambda>x. - f x" k "- y"] | 
| 53348 | 6768 | using assms using continuous_on_minus | 
| 6769 | apply auto | |
| 6770 | done | |
| 6771 | ||
| 6772 | lemma ivt_decreasing_component_1: | |
| 6773 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 6774 |   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 | 6775 |     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 | 6776 | by (rule ivt_decreasing_component_on_1) (auto simp: continuous_at_imp_continuous_on) | 
| 6777 | ||
| 33175 | 6778 | |
| 60420 | 6779 | subsection \<open>A bound within a convex hull, and so an interval\<close> | 
| 33175 | 6780 | |
| 6781 | lemma convex_on_convex_hull_bound: | |
| 53348 | 6782 | assumes "convex_on (convex hull s) f" | 
| 6783 | and "\<forall>x\<in>s. f x \<le> b" | |
| 6784 | shows "\<forall>x\<in> convex hull s. f x \<le> b" | |
| 6785 | proof | |
| 6786 | fix x | |
| 6787 | assume "x \<in> convex hull s" | |
| 6788 | then obtain k u v where | |
| 64267 | 6789 |     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 | 6790 | unfolding convex_hull_indexed mem_Collect_eq by auto | 
| 53348 | 6791 | have "(\<Sum>i = 1..k. u i * f (v i)) \<le> b" | 
| 64267 | 6792 |     using sum_mono[of "{1..k}" "\<lambda>i. u i * f (v i)" "\<lambda>i. u i * b"]
 | 
| 6793 | unfolding sum_distrib_right[symmetric] obt(2) mult_1 | |
| 53348 | 6794 | apply (drule_tac meta_mp) | 
| 6795 | apply (rule mult_left_mono) | |
| 6796 | using assms(2) obt(1) | |
| 6797 | apply auto | |
| 6798 | done | |
| 6799 | then show "f x \<le> b" | |
| 6800 | using assms(1)[unfolded convex_on[OF convex_convex_hull], rule_format, of k u v] | |
| 6801 | unfolding obt(2-3) | |
| 6802 | using obt(1) and hull_subset[unfolded subset_eq, rule_format, of _ s] | |
| 6803 | by auto | |
| 6804 | qed | |
| 6805 | ||
| 64267 | 6806 | lemma inner_sum_Basis[simp]: "i \<in> Basis \<Longrightarrow> (\<Sum>Basis) \<bullet> i = 1" | 
| 6807 | 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 | 6808 | |
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6809 | 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 | 6810 | 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 | 6811 | 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 | 6812 |   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 | 6813 | 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 | 6814 |   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 | 6815 | 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 | 6816 | finally show "convex (S + T)" . | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6817 | qed | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6818 | |
| 64267 | 6819 | lemma convex_set_sum: | 
| 55929 
91f245c23bc5
remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
 huffman parents: 
55928diff
changeset | 6820 | 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 | 6821 | 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 | 6822 | proof (cases "finite A") | 
| 
91f245c23bc5
remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
 huffman parents: 
55928diff
changeset | 6823 | 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 | 6824 | 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 | 6825 | qed auto | 
| 
91f245c23bc5
remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
 huffman parents: 
55928diff
changeset | 6826 | |
| 64267 | 6827 | lemma finite_set_sum: | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6828 | 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 | 6829 | 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 | 6830 | |
| 64267 | 6831 | lemma set_sum_eq: | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6832 |   "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 | 6833 | apply (induct set: finite) | 
| 
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 | apply simp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6836 | apply (safe elim!: set_plus_elim) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6837 | 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 | 6838 | apply simp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6839 | apply (rule_tac f="\<lambda>x. a + x" in arg_cong) | 
| 64267 | 6840 | apply (rule sum.cong [OF refl]) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6841 | apply clarsimp | 
| 57865 | 6842 | apply fast | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6843 | done | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6844 | |
| 64267 | 6845 | lemma box_eq_set_sum_Basis: | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6846 |   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 | 6847 | apply (subst set_sum_eq [OF finite_Basis]) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6848 | apply safe | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6849 | apply (fast intro: euclidean_representation [symmetric]) | 
| 64267 | 6850 | apply (subst inner_sum_left) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6851 | 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 | 6852 | apply (drule (1) bspec) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6853 | apply clarsimp | 
| 64267 | 6854 | apply (frule sum.remove [OF finite_Basis]) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6855 | apply (erule trans) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6856 | apply simp | 
| 64267 | 6857 | apply (rule sum.neutral) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6858 | apply clarsimp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6859 | apply (frule_tac x=i in bspec, assumption) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6860 | apply (drule_tac x=x in bspec, assumption) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6861 | apply clarsimp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6862 | 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 | 6863 | apply (rule ccontr) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6864 | apply simp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6865 | done | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6866 | |
| 64267 | 6867 | lemma convex_hull_set_sum: | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6868 | "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 | 6869 | proof (cases "finite A") | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6870 | assume "finite A" then show ?thesis | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6871 | 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 | 6872 | qed simp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6873 | |
| 56188 | 6874 | lemma convex_hull_eq_real_cbox: | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6875 | fixes x y :: real assumes "x \<le> y" | 
| 56188 | 6876 |   shows "convex hull {x, y} = cbox x y"
 | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6877 | proof (rule hull_unique) | 
| 60420 | 6878 |   show "{x, y} \<subseteq> cbox x y" using \<open>x \<le> y\<close> by auto
 | 
| 56188 | 6879 | show "convex (cbox x y)" | 
| 6880 | by (rule convex_box) | |
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6881 | next | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6882 |   fix s assume "{x, y} \<subseteq> s" and "convex s"
 | 
| 56188 | 6883 | then show "cbox x y \<subseteq> s" | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6884 | 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 | 6885 | by - (clarify, simp (no_asm_use), fast) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6886 | 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 | 6887 | |
| 33175 | 6888 | 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 | 6889 |   "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 | 6890 | (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 | 6891 | proof - | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 6892 | have One[simp]: "\<And>i. i \<in> Basis \<Longrightarrow> One \<bullet> i = 1" | 
| 64267 | 6893 | by (simp add: inner_sum_left sum.If_cases inner_Basis) | 
| 56188 | 6894 |   have "?int = {x. \<forall>i\<in>Basis. x \<bullet> i \<in> cbox 0 1}"
 | 
| 6895 | by (auto simp: cbox_def) | |
| 6896 | also have "\<dots> = (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` cbox 0 1)" | |
| 64267 | 6897 | by (simp only: box_eq_set_sum_Basis) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6898 |   also have "\<dots> = (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` (convex hull {0, 1}))"
 | 
| 56188 | 6899 | 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 | 6900 |   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 | 6901 | 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 | 6902 |   also have "\<dots> = convex hull (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` {0, 1})"
 | 
| 64267 | 6903 | by (simp only: convex_hull_set_sum) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6904 |   also have "\<dots> = convex hull {x. \<forall>i\<in>Basis. x\<bullet>i \<in> {0, 1}}"
 | 
| 64267 | 6905 | by (simp only: box_eq_set_sum_Basis) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6906 |   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 | 6907 | by simp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 6908 | 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 | 6909 | qed | 
| 33175 | 6910 | |
| 60420 | 6911 | text \<open>And this is a finite set of vertices.\<close> | 
| 33175 | 6912 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 6913 | lemma unit_cube_convex_hull: | 
| 56188 | 6914 | obtains s :: "'a::euclidean_space set" | 
| 6915 | where "finite s" and "cbox 0 (\<Sum>Basis) = convex hull s" | |
| 53348 | 6916 |   apply (rule that[of "{x::'a. \<forall>i\<in>Basis. x\<bullet>i=0 \<or> x\<bullet>i=1}"])
 | 
| 6917 | 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"]) | |
| 6918 | prefer 3 | |
| 6919 | apply (rule unit_interval_convex_hull) | |
| 6920 | apply rule | |
| 6921 | unfolding mem_Collect_eq | |
| 6922 | proof - | |
| 6923 | fix x :: 'a | |
| 6924 | 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 | 6925 | show "x \<in> (\<lambda>s. \<Sum>i\<in>Basis. (if i\<in>s then 1 else 0) *\<^sub>R i) ` Pow Basis" | 
| 53348 | 6926 |     apply (rule image_eqI[where x="{i. i\<in>Basis \<and> x\<bullet>i = 1}"])
 | 
| 6927 | using as | |
| 6928 | apply (subst euclidean_eq_iff) | |
| 57865 | 6929 | apply auto | 
| 53348 | 6930 | 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 | 6931 | qed auto | 
| 33175 | 6932 | |
| 60420 | 6933 | text \<open>Hence any cube (could do any nonempty interval).\<close> | 
| 33175 | 6934 | |
| 6935 | lemma cube_convex_hull: | |
| 53348 | 6936 | assumes "d > 0" | 
| 56188 | 6937 | obtains s :: "'a::euclidean_space set" where | 
| 6938 | "finite s" and "cbox (x - (\<Sum>i\<in>Basis. d*\<^sub>Ri)) (x + (\<Sum>i\<in>Basis. d*\<^sub>Ri)) = convex hull s" | |
| 53348 | 6939 | 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 | 6940 | let ?d = "(\<Sum>i\<in>Basis. d*\<^sub>Ri)::'a" | 
| 56188 | 6941 | have *: "cbox (x - ?d) (x + ?d) = (\<lambda>y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 (\<Sum>Basis)" | 
| 53348 | 6942 | apply (rule set_eqI, rule) | 
| 6943 | unfolding image_iff | |
| 6944 | defer | |
| 6945 | apply (erule bexE) | |
| 6946 | proof - | |
| 6947 | fix y | |
| 56188 | 6948 | assume as: "y\<in>cbox (x - ?d) (x + ?d)" | 
| 6949 | 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 | 6950 | using assms by (simp add: mem_box field_simps inner_simps) | 
| 60420 | 6951 | 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 | 6952 | by (intro bexI[of _ "inverse (2 * d) *\<^sub>R (y - (x - ?d))"]) auto | 
| 33175 | 6953 | next | 
| 53348 | 6954 | fix y z | 
| 56188 | 6955 | 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 | 6956 | have "\<And>i. i\<in>Basis \<Longrightarrow> 0 \<le> d * (z \<bullet> i) \<and> d * (z \<bullet> i) \<le> d" | 
| 56188 | 6957 | using assms as(1)[unfolded mem_box] | 
| 53348 | 6958 | apply (erule_tac x=i in ballE) | 
| 6959 | apply rule | |
| 56536 | 6960 | prefer 2 | 
| 53348 | 6961 | apply (rule mult_right_le_one_le) | 
| 6962 | using assms | |
| 6963 | apply auto | |
| 6964 | done | |
| 56188 | 6965 | then show "y \<in> cbox (x - ?d) (x + ?d)" | 
| 6966 | unfolding as(2) mem_box | |
| 53348 | 6967 | apply - | 
| 6968 | apply rule | |
| 56188 | 6969 | using as(1)[unfolded mem_box] | 
| 53348 | 6970 | apply (erule_tac x=i in ballE) | 
| 6971 | using assms | |
| 6972 | apply (auto simp: inner_simps) | |
| 6973 | done | |
| 6974 | qed | |
| 56188 | 6975 | obtain s where "finite s" "cbox 0 (\<Sum>Basis::'a) = convex hull s" | 
| 53348 | 6976 | using unit_cube_convex_hull by auto | 
| 6977 | then show ?thesis | |
| 6978 | apply (rule_tac that[of "(\<lambda>y. x - ?d + (2 * d) *\<^sub>R y)` s"]) | |
| 6979 | unfolding * and convex_hull_affinity | |
| 6980 | apply auto | |
| 6981 | done | |
| 6982 | qed | |
| 6983 | ||
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6984 | 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 | 6985 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6986 | lemma image_stretch_interval: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6987 | "(\<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 | 6988 |   (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 | 6989 | 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 | 6990 | (\<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 | 6991 | proof cases | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6992 |   assume *: "cbox a b \<noteq> {}"
 | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6993 | show ?thesis | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 6994 | 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 | 6995 | 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 | 6996 | 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 | 6997 | 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 | 6998 | 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 | 6999 | 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 | 7000 | 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 | 7001 | 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 | 7002 | 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 | 7003 | proof (cases "m i = 0") | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7004 | case True | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7005 | 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 | 7006 | next | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7007 | case False | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7008 | 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 | 7009 | 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 | 7010 | from False have | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7011 | "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 | 7012 | "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 | 7013 | 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 | 7014 | 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 | 7015 | 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 | 7016 | qed | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7017 | qed | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7018 | qed simp | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7019 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7020 | lemma interval_image_stretch_interval: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7021 | "\<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 | 7022 | 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 | 7023 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7024 | 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 | 7025 | 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 | 7026 | 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 | 7027 | 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 | 7028 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7029 | lemma cbox_image_unit_interval: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7030 | fixes a :: "'a::euclidean_space" | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7031 |   assumes "cbox a b \<noteq> {}"
 | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7032 | shows "cbox a b = | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7033 | 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 | 7034 | using assms | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7035 | apply (simp add: box_ne_empty image_stretch_interval cbox_translation [symmetric]) | 
| 64267 | 7036 | 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 | 7037 | done | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7038 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7039 | 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 | 7040 | fixes a :: "'a::euclidean_space" | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7041 | 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 | 7042 | proof (cases "cbox a b = {}")
 | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7043 | 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 | 7044 | by blast | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7045 | next | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7046 | case False | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7047 | 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 | 7048 | 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 | 7049 | have lin: "linear (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k)" | 
| 64267 | 7050 | 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 | 7051 | 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 | 7052 | 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 | 7053 | then show ?thesis | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7054 | apply (rule that) | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7055 | 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 | 7056 | 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 | 7057 | done | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7058 | qed | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 7059 | |
| 33175 | 7060 | |
| 60420 | 7061 | subsection \<open>Bounded convex function on open set is continuous\<close> | 
| 33175 | 7062 | |
| 7063 | lemma convex_on_bounded_continuous: | |
| 36338 | 7064 |   fixes s :: "('a::real_normed_vector) set"
 | 
| 53348 | 7065 | assumes "open s" | 
| 7066 | and "convex_on s f" | |
| 61945 | 7067 | and "\<forall>x\<in>s. \<bar>f x\<bar> \<le> b" | 
| 33175 | 7068 | shows "continuous_on s f" | 
| 53348 | 7069 | apply (rule continuous_at_imp_continuous_on) | 
| 7070 | unfolding continuous_at_real_range | |
| 7071 | proof (rule,rule,rule) | |
| 7072 | fix x and e :: real | |
| 7073 | assume "x \<in> s" "e > 0" | |
| 63040 | 7074 | define B where "B = \<bar>b\<bar> + 1" | 
| 61945 | 7075 | have B: "0 < B" "\<And>x. x\<in>s \<Longrightarrow> \<bar>f x\<bar> \<le> B" | 
| 53348 | 7076 | unfolding B_def | 
| 7077 | defer | |
| 7078 | apply (drule assms(3)[rule_format]) | |
| 7079 | apply auto | |
| 7080 | done | |
| 7081 | obtain k where "k > 0" and k: "cball x k \<subseteq> s" | |
| 7082 | using assms(1)[unfolded open_contains_cball, THEN bspec[where x=x]] | |
| 60420 | 7083 | using \<open>x\<in>s\<close> by auto | 
| 33175 | 7084 | show "\<exists>d>0. \<forall>x'. norm (x' - x) < d \<longrightarrow> \<bar>f x' - f x\<bar> < e" | 
| 53348 | 7085 | apply (rule_tac x="min (k / 2) (e / (2 * B) * k)" in exI) | 
| 7086 | apply rule | |
| 7087 | defer | |
| 7088 | proof (rule, rule) | |
| 7089 | fix y | |
| 7090 | assume as: "norm (y - x) < min (k / 2) (e / (2 * B) * k)" | |
| 7091 | show "\<bar>f y - f x\<bar> < e" | |
| 7092 | proof (cases "y = x") | |
| 7093 | case False | |
| 63040 | 7094 | define t where "t = k / norm (y - x)" | 
| 53348 | 7095 | have "2 < t" "0<t" | 
| 60420 | 7096 | unfolding t_def using as False and \<open>k>0\<close> | 
| 53348 | 7097 | by (auto simp add:field_simps) | 
| 7098 | have "y \<in> s" | |
| 7099 | apply (rule k[unfolded subset_eq,rule_format]) | |
| 7100 | unfolding mem_cball dist_norm | |
| 7101 | apply (rule order_trans[of _ "2 * norm (x - y)"]) | |
| 7102 | using as | |
| 7103 | by (auto simp add: field_simps norm_minus_commute) | |
| 7104 |       {
 | |
| 63040 | 7105 | define w where "w = x + t *\<^sub>R (y - x)" | 
| 53348 | 7106 | have "w \<in> s" | 
| 7107 | unfolding w_def | |
| 7108 | apply (rule k[unfolded subset_eq,rule_format]) | |
| 7109 | unfolding mem_cball dist_norm | |
| 7110 | unfolding t_def | |
| 60420 | 7111 | using \<open>k>0\<close> | 
| 53348 | 7112 | apply auto | 
| 7113 | done | |
| 7114 | have "(1 / t) *\<^sub>R x + - x + ((t - 1) / t) *\<^sub>R x = (1 / t - 1 + (t - 1) / t) *\<^sub>R x" | |
| 7115 | by (auto simp add: algebra_simps) | |
| 7116 | also have "\<dots> = 0" | |
| 60420 | 7117 | using \<open>t > 0\<close> by (auto simp add:field_simps) | 
| 53348 | 7118 | finally have w: "(1 / t) *\<^sub>R w + ((t - 1) / t) *\<^sub>R x = y" | 
| 60420 | 7119 | unfolding w_def using False and \<open>t > 0\<close> | 
| 53348 | 7120 | by (auto simp add: algebra_simps) | 
| 7121 | have "2 * B < e * t" | |
| 60420 | 7122 | unfolding t_def using \<open>0 < e\<close> \<open>0 < k\<close> \<open>B > 0\<close> and as and False | 
| 53348 | 7123 | by (auto simp add:field_simps) | 
| 7124 | then have "(f w - f x) / t < e" | |
| 60420 | 7125 | using B(2)[OF \<open>w\<in>s\<close>] and B(2)[OF \<open>x\<in>s\<close>] | 
| 7126 | using \<open>t > 0\<close> by (auto simp add:field_simps) | |
| 53348 | 7127 | then have th1: "f y - f x < e" | 
| 7128 | apply - | |
| 7129 | apply (rule le_less_trans) | |
| 7130 | defer | |
| 7131 | apply assumption | |
| 33175 | 7132 | using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w] | 
| 60420 | 7133 | using \<open>0 < t\<close> \<open>2 < t\<close> and \<open>x \<in> s\<close> \<open>w \<in> s\<close> | 
| 53348 | 7134 | by (auto simp add:field_simps) | 
| 7135 | } | |
| 49531 | 7136 | moreover | 
| 53348 | 7137 |       {
 | 
| 63040 | 7138 | define w where "w = x - t *\<^sub>R (y - x)" | 
| 53348 | 7139 | have "w \<in> s" | 
| 7140 | unfolding w_def | |
| 7141 | apply (rule k[unfolded subset_eq,rule_format]) | |
| 7142 | unfolding mem_cball dist_norm | |
| 7143 | unfolding t_def | |
| 60420 | 7144 | using \<open>k > 0\<close> | 
| 53348 | 7145 | apply auto | 
| 7146 | done | |
| 7147 | have "(1 / (1 + t)) *\<^sub>R x + (t / (1 + t)) *\<^sub>R x = (1 / (1 + t) + t / (1 + t)) *\<^sub>R x" | |
| 7148 | by (auto simp add: algebra_simps) | |
| 7149 | also have "\<dots> = x" | |
| 60420 | 7150 | using \<open>t > 0\<close> by (auto simp add:field_simps) | 
| 53348 | 7151 | finally have w: "(1 / (1+t)) *\<^sub>R w + (t / (1 + t)) *\<^sub>R y = x" | 
| 60420 | 7152 | unfolding w_def using False and \<open>t > 0\<close> | 
| 53348 | 7153 | by (auto simp add: algebra_simps) | 
| 7154 | have "2 * B < e * t" | |
| 7155 | unfolding t_def | |
| 60420 | 7156 | using \<open>0 < e\<close> \<open>0 < k\<close> \<open>B > 0\<close> and as and False | 
| 53348 | 7157 | by (auto simp add:field_simps) | 
| 7158 | then have *: "(f w - f y) / t < e" | |
| 60420 | 7159 | using B(2)[OF \<open>w\<in>s\<close>] and B(2)[OF \<open>y\<in>s\<close>] | 
| 7160 | using \<open>t > 0\<close> | |
| 53348 | 7161 | by (auto simp add:field_simps) | 
| 49531 | 7162 | have "f x \<le> 1 / (1 + t) * f w + (t / (1 + t)) * f y" | 
| 33175 | 7163 | using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w] | 
| 60420 | 7164 | using \<open>0 < t\<close> \<open>2 < t\<close> and \<open>y \<in> s\<close> \<open>w \<in> s\<close> | 
| 53348 | 7165 | by (auto simp add:field_simps) | 
| 7166 | also have "\<dots> = (f w + t * f y) / (1 + t)" | |
| 60420 | 7167 | using \<open>t > 0\<close> by (auto simp add: divide_simps) | 
| 53348 | 7168 | also have "\<dots> < e + f y" | 
| 60420 | 7169 | using \<open>t > 0\<close> * \<open>e > 0\<close> by (auto simp add: field_simps) | 
| 53348 | 7170 | finally have "f x - f y < e" by auto | 
| 7171 | } | |
| 49531 | 7172 | ultimately show ?thesis by auto | 
| 60420 | 7173 | qed (insert \<open>0<e\<close>, auto) | 
| 7174 | qed (insert \<open>0<e\<close> \<open>0<k\<close> \<open>0<B\<close>, auto simp: field_simps) | |
| 7175 | qed | |
| 7176 | ||
| 7177 | ||
| 7178 | subsection \<open>Upper bound on a ball implies upper and lower bounds\<close> | |
| 33175 | 7179 | |
| 7180 | lemma convex_bounds_lemma: | |
| 36338 | 7181 | fixes x :: "'a::real_normed_vector" | 
| 53348 | 7182 | assumes "convex_on (cball x e) f" | 
| 7183 | and "\<forall>y \<in> cball x e. f y \<le> b" | |
| 61945 | 7184 | shows "\<forall>y \<in> cball x e. \<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" | 
| 53348 | 7185 | apply rule | 
| 7186 | proof (cases "0 \<le> e") | |
| 7187 | case True | |
| 7188 | fix y | |
| 7189 | assume y: "y \<in> cball x e" | |
| 63040 | 7190 | define z where "z = 2 *\<^sub>R x - y" | 
| 53348 | 7191 | have *: "x - (2 *\<^sub>R x - y) = y - x" | 
| 7192 | by (simp add: scaleR_2) | |
| 7193 | have z: "z \<in> cball x e" | |
| 7194 | using y unfolding z_def mem_cball dist_norm * by (auto simp add: norm_minus_commute) | |
| 7195 | have "(1 / 2) *\<^sub>R y + (1 / 2) *\<^sub>R z = x" | |
| 7196 | unfolding z_def by (auto simp add: algebra_simps) | |
| 7197 | then show "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" | |
| 7198 | using assms(1)[unfolded convex_on_def,rule_format, OF y z, of "1/2" "1/2"] | |
| 7199 | using assms(2)[rule_format,OF y] assms(2)[rule_format,OF z] | |
| 7200 | by (auto simp add:field_simps) | |
| 7201 | next | |
| 7202 | case False | |
| 7203 | fix y | |
| 7204 | assume "y \<in> cball x e" | |
| 7205 | then have "dist x y < 0" | |
| 7206 | using False unfolding mem_cball not_le by (auto simp del: dist_not_less_zero) | |
| 7207 | then show "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" | |
| 7208 | using zero_le_dist[of x y] by auto | |
| 7209 | qed | |
| 7210 | ||
| 33175 | 7211 | |
| 60420 | 7212 | subsubsection \<open>Hence a convex function on an open set is continuous\<close> | 
| 33175 | 7213 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 7214 | 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 | 7215 | 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 | 7216 | |
| 33175 | 7217 | lemma convex_on_continuous: | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7218 |   assumes "open (s::('a::euclidean_space) set)" "convex_on s f"
 | 
| 33175 | 7219 | shows "continuous_on s f" | 
| 53348 | 7220 | unfolding continuous_on_eq_continuous_at[OF assms(1)] | 
| 7221 | proof | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 7222 | note dimge1 = DIM_positive[where 'a='a] | 
| 53348 | 7223 | fix x | 
| 7224 | assume "x \<in> s" | |
| 7225 | then obtain e where e: "cball x e \<subseteq> s" "e > 0" | |
| 7226 | using assms(1) unfolding open_contains_cball by auto | |
| 63040 | 7227 |   define d where "d = e / real DIM('a)"
 | 
| 53348 | 7228 | have "0 < d" | 
| 60420 | 7229 | 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 | 7230 | 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 | 7231 | obtain c | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7232 | where c: "finite c" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7233 | and c1: "convex hull c \<subseteq> cball x e" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7234 | and c2: "cball x d \<subseteq> convex hull c" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7235 | proof | 
| 63040 | 7236 |     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 | 7237 | show "finite c" | 
| 64267 | 7238 | unfolding c_def by (simp add: finite_set_sum) | 
| 56188 | 7239 |     have 1: "convex hull c = {a. \<forall>i\<in>Basis. a \<bullet> i \<in> cbox (x \<bullet> i - d) (x \<bullet> i + d)}"
 | 
| 64267 | 7240 | unfolding box_eq_set_sum_Basis | 
| 7241 | unfolding c_def convex_hull_set_sum | |
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7242 | apply (subst convex_hull_linear_image [symmetric]) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7243 | apply (simp add: linear_iff scaleR_add_left) | 
| 64267 | 7244 | apply (rule sum.cong [OF refl]) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7245 | apply (rule image_cong [OF _ refl]) | 
| 56188 | 7246 | apply (rule convex_hull_eq_real_cbox) | 
| 60420 | 7247 | apply (cut_tac \<open>0 < d\<close>, simp) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7248 | done | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7249 |     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 | 7250 | 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 | 7251 | show "cball x d \<subseteq> convex hull c" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7252 | unfolding 2 | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7253 | apply clarsimp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7254 | apply (simp only: dist_norm) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7255 | apply (subst inner_diff_left [symmetric]) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7256 | apply simp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7257 | apply (erule (1) order_trans [OF Basis_le_norm]) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7258 | done | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7259 | 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 | 7260 | by (simp add: d_def DIM_positive) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7261 | show "convex hull c \<subseteq> cball x e" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7262 | unfolding 2 | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7263 | apply clarsimp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7264 | apply (subst euclidean_dist_l2) | 
| 67155 | 7265 | apply (rule order_trans [OF L2_set_le_sum]) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7266 | apply (rule zero_le_dist) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7267 | unfolding e' | 
| 64267 | 7268 | apply (rule sum_mono) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7269 | apply simp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7270 | done | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7271 | qed | 
| 63040 | 7272 | define k where "k = Max (f ` c)" | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7273 | 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 | 7274 | 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 | 7275 | apply(rule subset_trans[OF _ e(1)]) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7276 | apply(rule c1) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7277 | done | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7278 | then have k: "\<forall>y\<in>convex hull c. f y \<le> k" | 
| 53348 | 7279 | apply (rule_tac convex_on_convex_hull_bound) | 
| 7280 | apply assumption | |
| 7281 | unfolding k_def | |
| 7282 | apply (rule, rule Max_ge) | |
| 7283 | using c(1) | |
| 7284 | apply auto | |
| 7285 | 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 | 7286 | 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 | 7287 | unfolding d_def | 
| 53348 | 7288 | apply (rule mult_imp_div_pos_le) | 
| 60420 | 7289 | 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 | 7290 | 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 | 7291 | 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 | 7292 | done | 
| 53348 | 7293 | 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 | 7294 | by (rule subset_cball) | 
| 53348 | 7295 | have conv: "convex_on (cball x d) f" | 
| 7296 | apply (rule convex_on_subset) | |
| 7297 | apply (rule convex_on_subset[OF assms(2)]) | |
| 7298 | apply (rule e(1)) | |
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7299 | apply (rule dsube) | 
| 53348 | 7300 | done | 
| 61945 | 7301 | 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 | 7302 | apply (rule convex_bounds_lemma) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7303 | apply (rule ballI) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7304 | apply (rule k [rule_format]) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7305 | apply (erule rev_subsetD) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7306 | apply (rule c2) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 7307 | done | 
| 53348 | 7308 | then have "continuous_on (ball x d) f" | 
| 7309 | apply (rule_tac convex_on_bounded_continuous) | |
| 7310 | apply (rule open_ball, rule convex_on_subset[OF conv]) | |
| 7311 | apply (rule ball_subset_cball) | |
| 33270 | 7312 | apply force | 
| 7313 | done | |
| 53348 | 7314 | then show "continuous (at x) f" | 
| 7315 | unfolding continuous_on_eq_continuous_at[OF open_ball] | |
| 60420 | 7316 | using \<open>d > 0\<close> by auto | 
| 7317 | qed | |
| 7318 | ||
| 33175 | 7319 | end |