| author | wenzelm |
| Thu, 01 Sep 2016 20:34:43 +0200 | |
| changeset 63761 | 2ca536d0163e |
| parent 63649 | e690d6f2185b |
| child 63918 | 6bf55e6e0b75 |
| permissions | -rw-r--r-- |
| 36648 | 1 |
(* Title: HOL/Library/Convex.thy |
2 |
Author: Armin Heller, TU Muenchen |
|
3 |
Author: Johannes Hoelzl, TU Muenchen |
|
4 |
*) |
|
5 |
||
| 60423 | 6 |
section \<open>Convexity in real vector spaces\<close> |
| 36648 | 7 |
|
| 36623 | 8 |
theory Convex |
| 63485 | 9 |
imports Product_Vector |
| 36623 | 10 |
begin |
11 |
||
| 60423 | 12 |
subsection \<open>Convexity\<close> |
| 36623 | 13 |
|
| 49609 | 14 |
definition convex :: "'a::real_vector set \<Rightarrow> bool" |
15 |
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)" |
|
| 36623 | 16 |
|
| 53676 | 17 |
lemma convexI: |
18 |
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" |
|
19 |
shows "convex s" |
|
20 |
using assms unfolding convex_def by fast |
|
21 |
||
22 |
lemma convexD: |
|
23 |
assumes "convex s" and "x \<in> s" and "y \<in> s" and "0 \<le> u" and "0 \<le> v" and "u + v = 1" |
|
24 |
shows "u *\<^sub>R x + v *\<^sub>R y \<in> s" |
|
25 |
using assms unfolding convex_def by fast |
|
26 |
||
| 63485 | 27 |
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)" |
| 36623 | 28 |
(is "_ \<longleftrightarrow> ?alt") |
29 |
proof |
|
| 63485 | 30 |
show "convex s" if alt: ?alt |
31 |
proof - |
|
32 |
{
|
|
33 |
fix x y and u v :: real |
|
34 |
assume mem: "x \<in> s" "y \<in> s" |
|
35 |
assume "0 \<le> u" "0 \<le> v" |
|
36 |
moreover |
|
37 |
assume "u + v = 1" |
|
38 |
then have "u = 1 - v" by auto |
|
39 |
ultimately have "u *\<^sub>R x + v *\<^sub>R y \<in> s" |
|
40 |
using alt [rule_format, OF mem] by auto |
|
41 |
} |
|
42 |
then show ?thesis |
|
43 |
unfolding convex_def by auto |
|
44 |
qed |
|
45 |
show ?alt if "convex s" |
|
46 |
using that by (auto simp: convex_def) |
|
47 |
qed |
|
| 36623 | 48 |
|
|
61426
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents:
61070
diff
changeset
|
49 |
lemma convexD_alt: |
| 36623 | 50 |
assumes "convex s" "a \<in> s" "b \<in> s" "0 \<le> u" "u \<le> 1" |
51 |
shows "((1 - u) *\<^sub>R a + u *\<^sub>R b) \<in> s" |
|
52 |
using assms unfolding convex_alt by auto |
|
53 |
||
|
61520
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
54 |
lemma mem_convex_alt: |
|
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
55 |
assumes "convex S" "x \<in> S" "y \<in> S" "u \<ge> 0" "v \<ge> 0" "u + v > 0" |
|
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
56 |
shows "((u/(u+v)) *\<^sub>R x + (v/(u+v)) *\<^sub>R y) \<in> S" |
|
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
57 |
apply (rule convexD) |
|
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
58 |
using assms |
| 63485 | 59 |
apply (simp_all add: zero_le_divide_iff add_divide_distrib [symmetric]) |
|
61520
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
60 |
done |
|
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
61 |
|
| 60303 | 62 |
lemma convex_empty[intro,simp]: "convex {}"
|
| 36623 | 63 |
unfolding convex_def by simp |
64 |
||
| 60303 | 65 |
lemma convex_singleton[intro,simp]: "convex {a}"
|
| 36623 | 66 |
unfolding convex_def by (auto simp: scaleR_left_distrib[symmetric]) |
67 |
||
| 60303 | 68 |
lemma convex_UNIV[intro,simp]: "convex UNIV" |
| 36623 | 69 |
unfolding convex_def by auto |
70 |
||
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62418
diff
changeset
|
71 |
lemma convex_Inter: "(\<And>s. s\<in>f \<Longrightarrow> convex s) \<Longrightarrow> convex(\<Inter>f)" |
| 36623 | 72 |
unfolding convex_def by auto |
73 |
||
74 |
lemma convex_Int: "convex s \<Longrightarrow> convex t \<Longrightarrow> convex (s \<inter> t)" |
|
75 |
unfolding convex_def by auto |
|
76 |
||
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62418
diff
changeset
|
77 |
lemma convex_INT: "(\<And>i. i \<in> A \<Longrightarrow> convex (B i)) \<Longrightarrow> convex (\<Inter>i\<in>A. B i)" |
| 53596 | 78 |
unfolding convex_def by auto |
79 |
||
80 |
lemma convex_Times: "convex s \<Longrightarrow> convex t \<Longrightarrow> convex (s \<times> t)" |
|
81 |
unfolding convex_def by auto |
|
82 |
||
| 36623 | 83 |
lemma convex_halfspace_le: "convex {x. inner a x \<le> b}"
|
84 |
unfolding convex_def |
|
| 44142 | 85 |
by (auto simp: inner_add intro!: convex_bound_le) |
| 36623 | 86 |
|
87 |
lemma convex_halfspace_ge: "convex {x. inner a x \<ge> b}"
|
|
88 |
proof - |
|
| 56796 | 89 |
have *: "{x. inner a x \<ge> b} = {x. inner (-a) x \<le> -b}"
|
90 |
by auto |
|
91 |
show ?thesis |
|
92 |
unfolding * using convex_halfspace_le[of "-a" "-b"] by auto |
|
| 36623 | 93 |
qed |
94 |
||
95 |
lemma convex_hyperplane: "convex {x. inner a x = b}"
|
|
| 49609 | 96 |
proof - |
| 56796 | 97 |
have *: "{x. inner a x = b} = {x. inner a x \<le> b} \<inter> {x. inner a x \<ge> b}"
|
98 |
by auto |
|
| 36623 | 99 |
show ?thesis using convex_halfspace_le convex_halfspace_ge |
100 |
by (auto intro!: convex_Int simp: *) |
|
101 |
qed |
|
102 |
||
103 |
lemma convex_halfspace_lt: "convex {x. inner a x < b}"
|
|
104 |
unfolding convex_def |
|
105 |
by (auto simp: convex_bound_lt inner_add) |
|
106 |
||
107 |
lemma convex_halfspace_gt: "convex {x. inner a x > b}"
|
|
108 |
using convex_halfspace_lt[of "-a" "-b"] by auto |
|
109 |
||
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
110 |
lemma convex_real_interval [iff]: |
| 36623 | 111 |
fixes a b :: "real" |
112 |
shows "convex {a..}" and "convex {..b}"
|
|
| 49609 | 113 |
and "convex {a<..}" and "convex {..<b}"
|
114 |
and "convex {a..b}" and "convex {a<..b}"
|
|
115 |
and "convex {a..<b}" and "convex {a<..<b}"
|
|
| 36623 | 116 |
proof - |
| 60423 | 117 |
have "{a..} = {x. a \<le> inner 1 x}"
|
118 |
by auto |
|
119 |
then show 1: "convex {a..}"
|
|
120 |
by (simp only: convex_halfspace_ge) |
|
121 |
have "{..b} = {x. inner 1 x \<le> b}"
|
|
122 |
by auto |
|
123 |
then show 2: "convex {..b}"
|
|
124 |
by (simp only: convex_halfspace_le) |
|
125 |
have "{a<..} = {x. a < inner 1 x}"
|
|
126 |
by auto |
|
127 |
then show 3: "convex {a<..}"
|
|
128 |
by (simp only: convex_halfspace_gt) |
|
129 |
have "{..<b} = {x. inner 1 x < b}"
|
|
130 |
by auto |
|
131 |
then show 4: "convex {..<b}"
|
|
132 |
by (simp only: convex_halfspace_lt) |
|
133 |
have "{a..b} = {a..} \<inter> {..b}"
|
|
134 |
by auto |
|
135 |
then show "convex {a..b}"
|
|
136 |
by (simp only: convex_Int 1 2) |
|
137 |
have "{a<..b} = {a<..} \<inter> {..b}"
|
|
138 |
by auto |
|
139 |
then show "convex {a<..b}"
|
|
140 |
by (simp only: convex_Int 3 2) |
|
141 |
have "{a..<b} = {a..} \<inter> {..<b}"
|
|
142 |
by auto |
|
143 |
then show "convex {a..<b}"
|
|
144 |
by (simp only: convex_Int 1 4) |
|
145 |
have "{a<..<b} = {a<..} \<inter> {..<b}"
|
|
146 |
by auto |
|
147 |
then show "convex {a<..<b}"
|
|
148 |
by (simp only: convex_Int 3 4) |
|
| 36623 | 149 |
qed |
150 |
||
| 61070 | 151 |
lemma convex_Reals: "convex \<real>" |
| 59862 | 152 |
by (simp add: convex_def scaleR_conv_of_real) |
| 60423 | 153 |
|
154 |
||
155 |
subsection \<open>Explicit expressions for convexity in terms of arbitrary sums\<close> |
|
| 36623 | 156 |
|
157 |
lemma convex_setsum: |
|
158 |
fixes C :: "'a::real_vector set" |
|
| 56796 | 159 |
assumes "finite s" |
160 |
and "convex C" |
|
161 |
and "(\<Sum> i \<in> s. a i) = 1" |
|
162 |
assumes "\<And>i. i \<in> s \<Longrightarrow> a i \<ge> 0" |
|
163 |
and "\<And>i. i \<in> s \<Longrightarrow> y i \<in> C" |
|
| 36623 | 164 |
shows "(\<Sum> j \<in> s. a j *\<^sub>R y j) \<in> C" |
| 55909 | 165 |
using assms(1,3,4,5) |
166 |
proof (induct arbitrary: a set: finite) |
|
| 49609 | 167 |
case empty |
| 55909 | 168 |
then show ?case by simp |
| 36623 | 169 |
next |
| 55909 | 170 |
case (insert i s) note IH = this(3) |
| 56796 | 171 |
have "a i + setsum a s = 1" |
172 |
and "0 \<le> a i" |
|
173 |
and "\<forall>j\<in>s. 0 \<le> a j" |
|
174 |
and "y i \<in> C" |
|
175 |
and "\<forall>j\<in>s. y j \<in> C" |
|
| 55909 | 176 |
using insert.hyps(1,2) insert.prems by simp_all |
| 56796 | 177 |
then have "0 \<le> setsum a s" |
178 |
by (simp add: setsum_nonneg) |
|
| 55909 | 179 |
have "a i *\<^sub>R y i + (\<Sum>j\<in>s. a j *\<^sub>R y j) \<in> C" |
| 63516 | 180 |
proof (cases "setsum a s = 0") |
181 |
case True |
|
| 60423 | 182 |
with \<open>a i + setsum a s = 1\<close> have "a i = 1" |
| 56796 | 183 |
by simp |
| 63516 | 184 |
from setsum_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" |
| 56796 | 185 |
by simp |
| 60423 | 186 |
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> |
| 56796 | 187 |
by simp |
| 55909 | 188 |
next |
| 63516 | 189 |
case False |
| 60423 | 190 |
with \<open>0 \<le> setsum a s\<close> have "0 < setsum a s" |
| 56796 | 191 |
by simp |
| 55909 | 192 |
then have "(\<Sum>j\<in>s. (a j / setsum a s) *\<^sub>R y j) \<in> C" |
| 60423 | 193 |
using \<open>\<forall>j\<in>s. 0 \<le> a j\<close> and \<open>\<forall>j\<in>s. y j \<in> C\<close> |
|
56571
f4635657d66f
added divide_nonneg_nonneg and co; made it a simp rule
hoelzl
parents:
56544
diff
changeset
|
194 |
by (simp add: IH setsum_divide_distrib [symmetric]) |
| 60423 | 195 |
from \<open>convex C\<close> and \<open>y i \<in> C\<close> and this and \<open>0 \<le> a i\<close> |
196 |
and \<open>0 \<le> setsum a s\<close> and \<open>a i + setsum a s = 1\<close> |
|
| 55909 | 197 |
have "a i *\<^sub>R y i + setsum a s *\<^sub>R (\<Sum>j\<in>s. (a j / setsum a s) *\<^sub>R y j) \<in> C" |
198 |
by (rule convexD) |
|
| 56796 | 199 |
then show ?thesis |
| 63516 | 200 |
by (simp add: scaleR_setsum_right False) |
| 55909 | 201 |
qed |
| 60423 | 202 |
then show ?case using \<open>finite s\<close> and \<open>i \<notin> s\<close> |
| 56796 | 203 |
by simp |
| 36623 | 204 |
qed |
205 |
||
206 |
lemma convex: |
|
| 49609 | 207 |
"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> (setsum u {1..k} = 1)
|
208 |
\<longrightarrow> setsum (\<lambda>i. u i *\<^sub>R x i) {1..k} \<in> s)"
|
|
| 36623 | 209 |
proof safe |
| 49609 | 210 |
fix k :: nat |
211 |
fix u :: "nat \<Rightarrow> real" |
|
212 |
fix x |
|
| 36623 | 213 |
assume "convex s" |
214 |
"\<forall>i. 1 \<le> i \<and> i \<le> k \<longrightarrow> 0 \<le> u i \<and> x i \<in> s" |
|
215 |
"setsum u {1..k} = 1"
|
|
| 60423 | 216 |
with convex_setsum[of "{1 .. k}" s] show "(\<Sum>j\<in>{1 .. k}. u j *\<^sub>R x j) \<in> s"
|
| 56796 | 217 |
by auto |
| 36623 | 218 |
next |
| 60423 | 219 |
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> setsum u {1..k} = 1
|
| 36623 | 220 |
\<longrightarrow> (\<Sum>i = 1..k. u i *\<^sub>R (x i :: 'a)) \<in> s" |
| 56796 | 221 |
{
|
222 |
fix \<mu> :: real |
|
| 49609 | 223 |
fix x y :: 'a |
224 |
assume xy: "x \<in> s" "y \<in> s" |
|
225 |
assume mu: "\<mu> \<ge> 0" "\<mu> \<le> 1" |
|
226 |
let ?u = "\<lambda>i. if (i :: nat) = 1 then \<mu> else 1 - \<mu>" |
|
227 |
let ?x = "\<lambda>i. if (i :: nat) = 1 then x else y" |
|
| 56796 | 228 |
have "{1 :: nat .. 2} \<inter> - {x. x = 1} = {2}"
|
229 |
by auto |
|
230 |
then have card: "card ({1 :: nat .. 2} \<inter> - {x. x = 1}) = 1"
|
|
231 |
by simp |
|
| 49609 | 232 |
then have "setsum ?u {1 .. 2} = 1"
|
| 57418 | 233 |
using setsum.If_cases[of "{(1 :: nat) .. 2}" "\<lambda> x. x = 1" "\<lambda> x. \<mu>" "\<lambda> x. 1 - \<mu>"]
|
| 36623 | 234 |
by auto |
| 60423 | 235 |
with *[rule_format, of "2" ?u ?x] have s: "(\<Sum>j \<in> {1..2}. ?u j *\<^sub>R ?x j) \<in> s"
|
| 36623 | 236 |
using mu xy by auto |
237 |
have grarr: "(\<Sum>j \<in> {Suc (Suc 0)..2}. ?u j *\<^sub>R ?x j) = (1 - \<mu>) *\<^sub>R y"
|
|
238 |
using setsum_head_Suc[of "Suc (Suc 0)" 2 "\<lambda> j. (1 - \<mu>) *\<^sub>R y"] by auto |
|
239 |
from setsum_head_Suc[of "Suc 0" 2 "\<lambda> j. ?u j *\<^sub>R ?x j", simplified this] |
|
| 56796 | 240 |
have "(\<Sum>j \<in> {1..2}. ?u j *\<^sub>R ?x j) = \<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y"
|
241 |
by auto |
|
242 |
then have "(1 - \<mu>) *\<^sub>R y + \<mu> *\<^sub>R x \<in> s" |
|
| 60423 | 243 |
using s by (auto simp: add.commute) |
| 49609 | 244 |
} |
| 56796 | 245 |
then show "convex s" |
246 |
unfolding convex_alt by auto |
|
| 36623 | 247 |
qed |
248 |
||
249 |
||
250 |
lemma convex_explicit: |
|
251 |
fixes s :: "'a::real_vector set" |
|
252 |
shows "convex s \<longleftrightarrow> |
|
| 49609 | 253 |
(\<forall>t u. finite t \<and> t \<subseteq> s \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and> setsum u t = 1 \<longrightarrow> setsum (\<lambda>x. u x *\<^sub>R x) t \<in> s)" |
| 36623 | 254 |
proof safe |
| 49609 | 255 |
fix t |
256 |
fix u :: "'a \<Rightarrow> real" |
|
| 56796 | 257 |
assume "convex s" |
258 |
and "finite t" |
|
259 |
and "t \<subseteq> s" "\<forall>x\<in>t. 0 \<le> u x" "setsum u t = 1" |
|
| 49609 | 260 |
then show "(\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s" |
| 36623 | 261 |
using convex_setsum[of t s u "\<lambda> x. x"] by auto |
262 |
next |
|
| 60423 | 263 |
assume *: "\<forall>t. \<forall> u. finite t \<and> t \<subseteq> s \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and> |
| 56796 | 264 |
setsum u t = 1 \<longrightarrow> (\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s" |
| 36623 | 265 |
show "convex s" |
266 |
unfolding convex_alt |
|
267 |
proof safe |
|
| 49609 | 268 |
fix x y |
269 |
fix \<mu> :: real |
|
| 60423 | 270 |
assume **: "x \<in> s" "y \<in> s" "0 \<le> \<mu>" "\<mu> \<le> 1" |
271 |
show "(1 - \<mu>) *\<^sub>R x + \<mu> *\<^sub>R y \<in> s" |
|
272 |
proof (cases "x = y") |
|
273 |
case False |
|
274 |
then show ?thesis |
|
275 |
using *[rule_format, of "{x, y}" "\<lambda> z. if z = x then 1 - \<mu> else \<mu>"] **
|
|
| 63485 | 276 |
by auto |
| 60423 | 277 |
next |
278 |
case True |
|
279 |
then show ?thesis |
|
280 |
using *[rule_format, of "{x, y}" "\<lambda> z. 1"] **
|
|
| 63485 | 281 |
by (auto simp: field_simps real_vector.scale_left_diff_distrib) |
| 60423 | 282 |
qed |
| 36623 | 283 |
qed |
284 |
qed |
|
285 |
||
| 49609 | 286 |
lemma convex_finite: |
287 |
assumes "finite s" |
|
| 56796 | 288 |
shows "convex s \<longleftrightarrow> (\<forall>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<longrightarrow> setsum (\<lambda>x. u x *\<^sub>R x) s \<in> s)" |
| 36623 | 289 |
unfolding convex_explicit |
| 63649 | 290 |
apply safe |
291 |
subgoal for u by (erule allE [where x=s], erule allE [where x=u]) auto |
|
292 |
subgoal for t u |
|
293 |
proof - |
|
294 |
have if_distrib_arg: "\<And>P f g x. (if P then f else g) x = (if P then f x else g x)" |
|
295 |
by simp |
|
296 |
assume sum: "\<forall>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> s" |
|
297 |
assume *: "\<forall>x\<in>t. 0 \<le> u x" "setsum u t = 1" |
|
298 |
assume "t \<subseteq> s" |
|
299 |
then have "s \<inter> t = t" by auto |
|
300 |
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" |
|
301 |
by (auto simp: assms setsum.If_cases if_distrib if_distrib_arg) |
|
302 |
qed |
|
303 |
done |
|
| 36623 | 304 |
|
| 56796 | 305 |
|
| 60423 | 306 |
subsection \<open>Functions that are convex on a set\<close> |
| 55909 | 307 |
|
| 49609 | 308 |
definition convex_on :: "'a::real_vector set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> bool"
|
309 |
where "convex_on s f \<longleftrightarrow> |
|
310 |
(\<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)" |
|
| 36623 | 311 |
|
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
312 |
lemma convex_onI [intro?]: |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
313 |
assumes "\<And>t x y. t > 0 \<Longrightarrow> t < 1 \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> |
| 63485 | 314 |
f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \<le> (1 - t) * f x + t * f y" |
315 |
shows "convex_on A f" |
|
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
316 |
unfolding convex_on_def |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
317 |
proof clarify |
| 63485 | 318 |
fix x y |
319 |
fix u v :: real |
|
320 |
assume A: "x \<in> A" "y \<in> A" "u \<ge> 0" "v \<ge> 0" "u + v = 1" |
|
321 |
from A(5) have [simp]: "v = 1 - u" |
|
322 |
by (simp add: algebra_simps) |
|
323 |
from A(1-4) show "f (u *\<^sub>R x + v *\<^sub>R y) \<le> u * f x + v * f y" |
|
324 |
using assms[of u y x] |
|
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
325 |
by (cases "u = 0 \<or> u = 1") (auto simp: algebra_simps) |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
326 |
qed |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
327 |
|
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
328 |
lemma convex_on_linorderI [intro?]: |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
329 |
fixes A :: "('a::{linorder,real_vector}) set"
|
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
330 |
assumes "\<And>t x y. t > 0 \<Longrightarrow> t < 1 \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x < y \<Longrightarrow> |
| 63485 | 331 |
f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \<le> (1 - t) * f x + t * f y" |
332 |
shows "convex_on A f" |
|
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
333 |
proof |
| 63485 | 334 |
fix x y |
335 |
fix t :: real |
|
336 |
assume A: "x \<in> A" "y \<in> A" "t > 0" "t < 1" |
|
337 |
with assms [of t x y] assms [of "1 - t" y x] |
|
| 62418 | 338 |
show "f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \<le> (1 - t) * f x + t * f y" |
339 |
by (cases x y rule: linorder_cases) (auto simp: algebra_simps) |
|
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
340 |
qed |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
341 |
|
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
342 |
lemma convex_onD: |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
343 |
assumes "convex_on A f" |
| 63485 | 344 |
shows "\<And>t x y. t \<ge> 0 \<Longrightarrow> t \<le> 1 \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> |
345 |
f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \<le> (1 - t) * f x + t * f y" |
|
346 |
using assms by (auto simp: convex_on_def) |
|
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
347 |
|
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
348 |
lemma convex_onD_Icc: |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
349 |
assumes "convex_on {x..y} f" "x \<le> (y :: _ :: {real_vector,preorder})"
|
| 63485 | 350 |
shows "\<And>t. t \<ge> 0 \<Longrightarrow> t \<le> 1 \<Longrightarrow> |
351 |
f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \<le> (1 - t) * f x + t * f y" |
|
352 |
using assms(2) by (intro convex_onD [OF assms(1)]) simp_all |
|
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
353 |
|
| 36623 | 354 |
lemma convex_on_subset: "convex_on t f \<Longrightarrow> s \<subseteq> t \<Longrightarrow> convex_on s f" |
355 |
unfolding convex_on_def by auto |
|
356 |
||
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
357 |
lemma convex_on_add [intro]: |
| 56796 | 358 |
assumes "convex_on s f" |
359 |
and "convex_on s g" |
|
| 36623 | 360 |
shows "convex_on s (\<lambda>x. f x + g x)" |
| 49609 | 361 |
proof - |
| 56796 | 362 |
{
|
363 |
fix x y |
|
364 |
assume "x \<in> s" "y \<in> s" |
|
| 49609 | 365 |
moreover |
366 |
fix u v :: real |
|
367 |
assume "0 \<le> u" "0 \<le> v" "u + v = 1" |
|
368 |
ultimately |
|
369 |
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)" |
|
| 60423 | 370 |
using assms unfolding convex_on_def by (auto simp: add_mono) |
| 49609 | 371 |
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)" |
372 |
by (simp add: field_simps) |
|
373 |
} |
|
| 56796 | 374 |
then show ?thesis |
375 |
unfolding convex_on_def by auto |
|
| 36623 | 376 |
qed |
377 |
||
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
378 |
lemma convex_on_cmul [intro]: |
| 56796 | 379 |
fixes c :: real |
380 |
assumes "0 \<le> c" |
|
381 |
and "convex_on s f" |
|
| 36623 | 382 |
shows "convex_on s (\<lambda>x. c * f x)" |
| 56796 | 383 |
proof - |
| 63485 | 384 |
have *: "u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" |
385 |
for u c fx v fy :: real |
|
| 49609 | 386 |
by (simp add: field_simps) |
387 |
show ?thesis using assms(2) and mult_left_mono [OF _ assms(1)] |
|
388 |
unfolding convex_on_def and * by auto |
|
| 36623 | 389 |
qed |
390 |
||
391 |
lemma convex_lower: |
|
| 56796 | 392 |
assumes "convex_on s f" |
393 |
and "x \<in> s" |
|
394 |
and "y \<in> s" |
|
395 |
and "0 \<le> u" |
|
396 |
and "0 \<le> v" |
|
397 |
and "u + v = 1" |
|
| 36623 | 398 |
shows "f (u *\<^sub>R x + v *\<^sub>R y) \<le> max (f x) (f y)" |
| 56796 | 399 |
proof - |
| 36623 | 400 |
let ?m = "max (f x) (f y)" |
401 |
have "u * f x + v * f y \<le> u * max (f x) (f y) + v * max (f x) (f y)" |
|
| 60423 | 402 |
using assms(4,5) by (auto simp: mult_left_mono add_mono) |
| 56796 | 403 |
also have "\<dots> = max (f x) (f y)" |
| 60423 | 404 |
using assms(6) by (simp add: distrib_right [symmetric]) |
| 36623 | 405 |
finally show ?thesis |
|
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44282
diff
changeset
|
406 |
using assms unfolding convex_on_def by fastforce |
| 36623 | 407 |
qed |
408 |
||
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
409 |
lemma convex_on_dist [intro]: |
| 36623 | 410 |
fixes s :: "'a::real_normed_vector set" |
411 |
shows "convex_on s (\<lambda>x. dist a x)" |
|
| 60423 | 412 |
proof (auto simp: convex_on_def dist_norm) |
| 49609 | 413 |
fix x y |
| 56796 | 414 |
assume "x \<in> s" "y \<in> s" |
| 49609 | 415 |
fix u v :: real |
| 56796 | 416 |
assume "0 \<le> u" |
417 |
assume "0 \<le> v" |
|
418 |
assume "u + v = 1" |
|
| 49609 | 419 |
have "a = u *\<^sub>R a + v *\<^sub>R a" |
| 60423 | 420 |
unfolding scaleR_left_distrib[symmetric] and \<open>u + v = 1\<close> by simp |
| 49609 | 421 |
then have *: "a - (u *\<^sub>R x + v *\<^sub>R y) = (u *\<^sub>R (a - x)) + (v *\<^sub>R (a - y))" |
| 60423 | 422 |
by (auto simp: algebra_simps) |
| 36623 | 423 |
show "norm (a - (u *\<^sub>R x + v *\<^sub>R y)) \<le> u * norm (a - x) + v * norm (a - y)" |
424 |
unfolding * using norm_triangle_ineq[of "u *\<^sub>R (a - x)" "v *\<^sub>R (a - y)"] |
|
| 60423 | 425 |
using \<open>0 \<le> u\<close> \<open>0 \<le> v\<close> by auto |
| 36623 | 426 |
qed |
427 |
||
| 49609 | 428 |
|
| 60423 | 429 |
subsection \<open>Arithmetic operations on sets preserve convexity\<close> |
| 49609 | 430 |
|
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
431 |
lemma convex_linear_image: |
| 56796 | 432 |
assumes "linear f" |
433 |
and "convex s" |
|
434 |
shows "convex (f ` s)" |
|
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
435 |
proof - |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
436 |
interpret f: linear f by fact |
| 60423 | 437 |
from \<open>convex s\<close> show "convex (f ` s)" |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
438 |
by (simp add: convex_def f.scaleR [symmetric] f.add [symmetric]) |
| 36623 | 439 |
qed |
440 |
||
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
441 |
lemma convex_linear_vimage: |
| 56796 | 442 |
assumes "linear f" |
443 |
and "convex s" |
|
444 |
shows "convex (f -` s)" |
|
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
445 |
proof - |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
446 |
interpret f: linear f by fact |
| 60423 | 447 |
from \<open>convex s\<close> show "convex (f -` s)" |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
448 |
by (simp add: convex_def f.add f.scaleR) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
449 |
qed |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
450 |
|
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
451 |
lemma convex_scaling: |
| 56796 | 452 |
assumes "convex s" |
453 |
shows "convex ((\<lambda>x. c *\<^sub>R x) ` s)" |
|
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
454 |
proof - |
| 56796 | 455 |
have "linear (\<lambda>x. c *\<^sub>R x)" |
456 |
by (simp add: linearI scaleR_add_right) |
|
457 |
then show ?thesis |
|
| 60423 | 458 |
using \<open>convex s\<close> by (rule convex_linear_image) |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
459 |
qed |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
460 |
|
|
60178
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
immler
parents:
59862
diff
changeset
|
461 |
lemma convex_scaled: |
|
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
immler
parents:
59862
diff
changeset
|
462 |
assumes "convex s" |
|
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
immler
parents:
59862
diff
changeset
|
463 |
shows "convex ((\<lambda>x. x *\<^sub>R c) ` s)" |
|
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
immler
parents:
59862
diff
changeset
|
464 |
proof - |
|
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
immler
parents:
59862
diff
changeset
|
465 |
have "linear (\<lambda>x. x *\<^sub>R c)" |
|
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
immler
parents:
59862
diff
changeset
|
466 |
by (simp add: linearI scaleR_add_left) |
|
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
immler
parents:
59862
diff
changeset
|
467 |
then show ?thesis |
| 60423 | 468 |
using \<open>convex s\<close> by (rule convex_linear_image) |
|
60178
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
immler
parents:
59862
diff
changeset
|
469 |
qed |
|
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
immler
parents:
59862
diff
changeset
|
470 |
|
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
471 |
lemma convex_negations: |
| 56796 | 472 |
assumes "convex s" |
473 |
shows "convex ((\<lambda>x. - x) ` s)" |
|
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
474 |
proof - |
| 56796 | 475 |
have "linear (\<lambda>x. - x)" |
476 |
by (simp add: linearI) |
|
477 |
then show ?thesis |
|
| 60423 | 478 |
using \<open>convex s\<close> by (rule convex_linear_image) |
| 36623 | 479 |
qed |
480 |
||
481 |
lemma convex_sums: |
|
| 56796 | 482 |
assumes "convex s" |
483 |
and "convex t" |
|
| 36623 | 484 |
shows "convex {x + y| x y. x \<in> s \<and> y \<in> t}"
|
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
485 |
proof - |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
486 |
have "linear (\<lambda>(x, y). x + y)" |
| 60423 | 487 |
by (auto intro: linearI simp: scaleR_add_right) |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
488 |
with assms have "convex ((\<lambda>(x, y). x + y) ` (s \<times> t))" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
489 |
by (intro convex_linear_image convex_Times) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
490 |
also have "((\<lambda>(x, y). x + y) ` (s \<times> t)) = {x + y| x y. x \<in> s \<and> y \<in> t}"
|
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
491 |
by auto |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53596
diff
changeset
|
492 |
finally show ?thesis . |
| 36623 | 493 |
qed |
494 |
||
495 |
lemma convex_differences: |
|
496 |
assumes "convex s" "convex t" |
|
497 |
shows "convex {x - y| x y. x \<in> s \<and> y \<in> t}"
|
|
498 |
proof - |
|
499 |
have "{x - y| x y. x \<in> s \<and> y \<in> t} = {x + y |x y. x \<in> s \<and> y \<in> uminus ` t}"
|
|
| 60423 | 500 |
by (auto simp: diff_conv_add_uminus simp del: add_uminus_conv_diff) |
| 49609 | 501 |
then show ?thesis |
502 |
using convex_sums[OF assms(1) convex_negations[OF assms(2)]] by auto |
|
| 36623 | 503 |
qed |
504 |
||
| 49609 | 505 |
lemma convex_translation: |
506 |
assumes "convex s" |
|
507 |
shows "convex ((\<lambda>x. a + x) ` s)" |
|
508 |
proof - |
|
| 56796 | 509 |
have "{a + y |y. y \<in> s} = (\<lambda>x. a + x) ` s"
|
510 |
by auto |
|
| 49609 | 511 |
then show ?thesis |
512 |
using convex_sums[OF convex_singleton[of a] assms] by auto |
|
513 |
qed |
|
| 36623 | 514 |
|
| 49609 | 515 |
lemma convex_affinity: |
516 |
assumes "convex s" |
|
517 |
shows "convex ((\<lambda>x. a + c *\<^sub>R x) ` s)" |
|
518 |
proof - |
|
| 56796 | 519 |
have "(\<lambda>x. a + c *\<^sub>R x) ` s = op + a ` op *\<^sub>R c ` s" |
520 |
by auto |
|
| 49609 | 521 |
then show ?thesis |
522 |
using convex_translation[OF convex_scaling[OF assms], of a c] by auto |
|
523 |
qed |
|
| 36623 | 524 |
|
| 49609 | 525 |
lemma pos_is_convex: "convex {0 :: real <..}"
|
526 |
unfolding convex_alt |
|
| 36623 | 527 |
proof safe |
528 |
fix y x \<mu> :: real |
|
| 60423 | 529 |
assume *: "y > 0" "x > 0" "\<mu> \<ge> 0" "\<mu> \<le> 1" |
| 56796 | 530 |
{
|
531 |
assume "\<mu> = 0" |
|
| 63485 | 532 |
then have "\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y = y" |
533 |
by simp |
|
534 |
then have "\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y > 0" |
|
535 |
using * by simp |
|
| 56796 | 536 |
} |
| 36623 | 537 |
moreover |
| 56796 | 538 |
{
|
539 |
assume "\<mu> = 1" |
|
| 63485 | 540 |
then have "\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y > 0" |
541 |
using * by simp |
|
| 56796 | 542 |
} |
| 36623 | 543 |
moreover |
| 56796 | 544 |
{
|
545 |
assume "\<mu> \<noteq> 1" "\<mu> \<noteq> 0" |
|
| 63485 | 546 |
then have "\<mu> > 0" "(1 - \<mu>) > 0" |
547 |
using * by auto |
|
548 |
then have "\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y > 0" |
|
549 |
using * by (auto simp: add_pos_pos) |
|
| 56796 | 550 |
} |
551 |
ultimately show "(1 - \<mu>) *\<^sub>R y + \<mu> *\<^sub>R x > 0" |
|
| 63092 | 552 |
by fastforce |
| 36623 | 553 |
qed |
554 |
||
555 |
lemma convex_on_setsum: |
|
556 |
fixes a :: "'a \<Rightarrow> real" |
|
| 49609 | 557 |
and y :: "'a \<Rightarrow> 'b::real_vector" |
558 |
and f :: "'b \<Rightarrow> real" |
|
| 36623 | 559 |
assumes "finite s" "s \<noteq> {}"
|
| 49609 | 560 |
and "convex_on C f" |
561 |
and "convex C" |
|
562 |
and "(\<Sum> i \<in> s. a i) = 1" |
|
563 |
and "\<And>i. i \<in> s \<Longrightarrow> a i \<ge> 0" |
|
564 |
and "\<And>i. i \<in> s \<Longrightarrow> y i \<in> C" |
|
| 36623 | 565 |
shows "f (\<Sum> i \<in> s. a i *\<^sub>R y i) \<le> (\<Sum> i \<in> s. a i * f (y i))" |
| 49609 | 566 |
using assms |
567 |
proof (induct s arbitrary: a rule: finite_ne_induct) |
|
| 36623 | 568 |
case (singleton i) |
| 63485 | 569 |
then have ai: "a i = 1" |
570 |
by auto |
|
571 |
then show ?case |
|
572 |
by auto |
|
| 36623 | 573 |
next |
| 60423 | 574 |
case (insert i s) |
| 63485 | 575 |
then have "convex_on C f" |
576 |
by simp |
|
| 36623 | 577 |
from this[unfolded convex_on_def, rule_format] |
| 56796 | 578 |
have conv: "\<And>x y \<mu>. x \<in> C \<Longrightarrow> y \<in> C \<Longrightarrow> 0 \<le> \<mu> \<Longrightarrow> \<mu> \<le> 1 \<Longrightarrow> |
579 |
f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y" |
|
| 36623 | 580 |
by simp |
| 60423 | 581 |
show ?case |
582 |
proof (cases "a i = 1") |
|
583 |
case True |
|
| 49609 | 584 |
then have "(\<Sum> j \<in> s. a j) = 0" |
| 60423 | 585 |
using insert by auto |
| 49609 | 586 |
then have "\<And>j. j \<in> s \<Longrightarrow> a j = 0" |
|
62376
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents:
61694
diff
changeset
|
587 |
using insert by (fastforce simp: setsum_nonneg_eq_0_iff) |
| 60423 | 588 |
then show ?thesis |
589 |
using insert by auto |
|
590 |
next |
|
591 |
case False |
|
592 |
from insert have yai: "y i \<in> C" "a i \<ge> 0" |
|
593 |
by auto |
|
594 |
have fis: "finite (insert i s)" |
|
595 |
using insert by auto |
|
596 |
then have ai1: "a i \<le> 1" |
|
597 |
using setsum_nonneg_leq_bound[of "insert i s" a] insert by simp |
|
598 |
then have "a i < 1" |
|
599 |
using False by auto |
|
600 |
then have i0: "1 - a i > 0" |
|
601 |
by auto |
|
| 49609 | 602 |
let ?a = "\<lambda>j. a j / (1 - a i)" |
| 60423 | 603 |
have a_nonneg: "?a j \<ge> 0" if "j \<in> s" for j |
| 60449 | 604 |
using i0 insert that by fastforce |
| 60423 | 605 |
have "(\<Sum> j \<in> insert i s. a j) = 1" |
606 |
using insert by auto |
|
607 |
then have "(\<Sum> j \<in> s. a j) = 1 - a i" |
|
608 |
using setsum.insert insert by fastforce |
|
609 |
then have "(\<Sum> j \<in> s. a j) / (1 - a i) = 1" |
|
610 |
using i0 by auto |
|
611 |
then have a1: "(\<Sum> j \<in> s. ?a j) = 1" |
|
612 |
unfolding setsum_divide_distrib by simp |
|
613 |
have "convex C" using insert by auto |
|
| 49609 | 614 |
then have asum: "(\<Sum> j \<in> s. ?a j *\<^sub>R y j) \<in> C" |
| 63485 | 615 |
using insert convex_setsum [OF \<open>finite s\<close> \<open>convex C\<close> a1 a_nonneg] by auto |
| 36623 | 616 |
have asum_le: "f (\<Sum> j \<in> s. ?a j *\<^sub>R y j) \<le> (\<Sum> j \<in> s. ?a j * f (y j))" |
| 60423 | 617 |
using a_nonneg a1 insert by blast |
| 36623 | 618 |
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)" |
| 60423 | 619 |
using setsum.insert[of s i "\<lambda> j. a j *\<^sub>R y j", OF \<open>finite s\<close> \<open>i \<notin> s\<close>] insert |
620 |
by (auto simp only: add.commute) |
|
| 36623 | 621 |
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)" |
622 |
using i0 by auto |
|
623 |
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)" |
|
| 49609 | 624 |
using scaleR_right.setsum[of "inverse (1 - a i)" "\<lambda> j. a j *\<^sub>R y j" s, symmetric] |
| 60423 | 625 |
by (auto simp: algebra_simps) |
| 36623 | 626 |
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)" |
|
36778
739a9379e29b
avoid using real-specific versions of generic lemmas
huffman
parents:
36648
diff
changeset
|
627 |
by (auto simp: divide_inverse) |
| 36623 | 628 |
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)" |
629 |
using conv[of "y i" "(\<Sum> j \<in> s. ?a j *\<^sub>R y j)" "a i", OF yai(1) asum yai(2) ai1] |
|
| 60423 | 630 |
by (auto simp: add.commute) |
| 36623 | 631 |
also have "\<dots> \<le> (1 - a i) * (\<Sum> j \<in> s. ?a j * f (y j)) + a i * f (y i)" |
| 63485 | 632 |
using add_right_mono [OF mult_left_mono [of _ _ "1 - a i", |
633 |
OF asum_le less_imp_le[OF i0]], of "a i * f (y i)"] |
|
634 |
by simp |
|
| 36623 | 635 |
also have "\<dots> = (\<Sum> j \<in> s. (1 - a i) * ?a j * f (y j)) + a i * f (y i)" |
| 63485 | 636 |
unfolding setsum_right_distrib[of "1 - a i" "\<lambda> j. ?a j * f (y j)"] |
637 |
using i0 by auto |
|
| 60423 | 638 |
also have "\<dots> = (\<Sum> j \<in> s. a j * f (y j)) + a i * f (y i)" |
639 |
using i0 by auto |
|
640 |
also have "\<dots> = (\<Sum> j \<in> insert i s. a j * f (y j))" |
|
641 |
using insert by auto |
|
642 |
finally show ?thesis |
|
| 56796 | 643 |
by simp |
| 60423 | 644 |
qed |
| 36623 | 645 |
qed |
646 |
||
647 |
lemma convex_on_alt: |
|
648 |
fixes C :: "'a::real_vector set" |
|
649 |
assumes "convex C" |
|
| 56796 | 650 |
shows "convex_on C f \<longleftrightarrow> |
651 |
(\<forall>x \<in> C. \<forall> y \<in> C. \<forall> \<mu> :: real. \<mu> \<ge> 0 \<and> \<mu> \<le> 1 \<longrightarrow> |
|
652 |
f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y)" |
|
| 36623 | 653 |
proof safe |
| 49609 | 654 |
fix x y |
655 |
fix \<mu> :: real |
|
| 60423 | 656 |
assume *: "convex_on C f" "x \<in> C" "y \<in> C" "0 \<le> \<mu>" "\<mu> \<le> 1" |
| 36623 | 657 |
from this[unfolded convex_on_def, rule_format] |
| 63485 | 658 |
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 |
| 56796 | 659 |
by auto |
| 63485 | 660 |
from this [of "\<mu>" "1 - \<mu>", simplified] * |
| 56796 | 661 |
show "f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y" |
662 |
by auto |
|
| 36623 | 663 |
next |
| 60423 | 664 |
assume *: "\<forall>x\<in>C. \<forall>y\<in>C. \<forall>\<mu>. 0 \<le> \<mu> \<and> \<mu> \<le> 1 \<longrightarrow> |
| 56796 | 665 |
f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y" |
666 |
{
|
|
667 |
fix x y |
|
| 49609 | 668 |
fix u v :: real |
| 60423 | 669 |
assume **: "x \<in> C" "y \<in> C" "u \<ge> 0" "v \<ge> 0" "u + v = 1" |
| 49609 | 670 |
then have[simp]: "1 - u = v" by auto |
| 60423 | 671 |
from *[rule_format, of x y u] |
| 56796 | 672 |
have "f (u *\<^sub>R x + v *\<^sub>R y) \<le> u * f x + v * f y" |
| 60423 | 673 |
using ** by auto |
| 49609 | 674 |
} |
| 56796 | 675 |
then show "convex_on C f" |
676 |
unfolding convex_on_def by auto |
|
| 36623 | 677 |
qed |
678 |
||
| 43337 | 679 |
lemma convex_on_diff: |
680 |
fixes f :: "real \<Rightarrow> real" |
|
| 56796 | 681 |
assumes f: "convex_on I f" |
682 |
and I: "x \<in> I" "y \<in> I" |
|
683 |
and t: "x < t" "t < y" |
|
| 49609 | 684 |
shows "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)" |
| 56796 | 685 |
and "(f x - f y) / (x - y) \<le> (f t - f y) / (t - y)" |
| 43337 | 686 |
proof - |
| 63040 | 687 |
define a where "a \<equiv> (t - y) / (x - y)" |
| 56796 | 688 |
with t have "0 \<le> a" "0 \<le> 1 - a" |
689 |
by (auto simp: field_simps) |
|
| 60423 | 690 |
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" |
| 43337 | 691 |
by (auto simp: convex_on_def) |
| 56796 | 692 |
have "a * x + (1 - a) * y = a * (x - y) + y" |
693 |
by (simp add: field_simps) |
|
694 |
also have "\<dots> = t" |
|
| 60423 | 695 |
unfolding a_def using \<open>x < t\<close> \<open>t < y\<close> by simp |
| 56796 | 696 |
finally have "f t \<le> a * f x + (1 - a) * f y" |
697 |
using cvx by simp |
|
698 |
also have "\<dots> = a * (f x - f y) + f y" |
|
699 |
by (simp add: field_simps) |
|
700 |
finally have "f t - f y \<le> a * (f x - f y)" |
|
701 |
by simp |
|
| 43337 | 702 |
with t show "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)" |
| 44142 | 703 |
by (simp add: le_divide_eq divide_le_eq field_simps a_def) |
| 43337 | 704 |
with t show "(f x - f y) / (x - y) \<le> (f t - f y) / (t - y)" |
| 44142 | 705 |
by (simp add: le_divide_eq divide_le_eq field_simps) |
| 43337 | 706 |
qed |
| 36623 | 707 |
|
708 |
lemma pos_convex_function: |
|
709 |
fixes f :: "real \<Rightarrow> real" |
|
710 |
assumes "convex C" |
|
| 56796 | 711 |
and leq: "\<And>x y. x \<in> C \<Longrightarrow> y \<in> C \<Longrightarrow> f' x * (y - x) \<le> f y - f x" |
| 36623 | 712 |
shows "convex_on C f" |
| 49609 | 713 |
unfolding convex_on_alt[OF assms(1)] |
714 |
using assms |
|
| 36623 | 715 |
proof safe |
716 |
fix x y \<mu> :: real |
|
717 |
let ?x = "\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y" |
|
| 60423 | 718 |
assume *: "convex C" "x \<in> C" "y \<in> C" "\<mu> \<ge> 0" "\<mu> \<le> 1" |
| 49609 | 719 |
then have "1 - \<mu> \<ge> 0" by auto |
| 56796 | 720 |
then have xpos: "?x \<in> C" |
| 60423 | 721 |
using * unfolding convex_alt by fastforce |
| 56796 | 722 |
have geq: "\<mu> * (f x - f ?x) + (1 - \<mu>) * (f y - f ?x) \<ge> |
723 |
\<mu> * f' ?x * (x - ?x) + (1 - \<mu>) * f' ?x * (y - ?x)" |
|
| 63485 | 724 |
using add_mono [OF mult_left_mono [OF leq [OF xpos *(2)] \<open>\<mu> \<ge> 0\<close>] |
725 |
mult_left_mono [OF leq [OF xpos *(3)] \<open>1 - \<mu> \<ge> 0\<close>]] |
|
| 56796 | 726 |
by auto |
| 49609 | 727 |
then have "\<mu> * f x + (1 - \<mu>) * f y - f ?x \<ge> 0" |
| 60423 | 728 |
by (auto simp: field_simps) |
| 49609 | 729 |
then show "f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y" |
| 36623 | 730 |
using convex_on_alt by auto |
731 |
qed |
|
732 |
||
733 |
lemma atMostAtLeast_subset_convex: |
|
734 |
fixes C :: "real set" |
|
735 |
assumes "convex C" |
|
| 49609 | 736 |
and "x \<in> C" "y \<in> C" "x < y" |
| 36623 | 737 |
shows "{x .. y} \<subseteq> C"
|
738 |
proof safe |
|
| 60423 | 739 |
fix z assume z: "z \<in> {x .. y}"
|
740 |
have less: "z \<in> C" if *: "x < z" "z < y" |
|
741 |
proof - |
|
| 49609 | 742 |
let ?\<mu> = "(y - z) / (y - x)" |
| 56796 | 743 |
have "0 \<le> ?\<mu>" "?\<mu> \<le> 1" |
| 60423 | 744 |
using assms * by (auto simp: field_simps) |
| 49609 | 745 |
then have comb: "?\<mu> * x + (1 - ?\<mu>) * y \<in> C" |
746 |
using assms iffD1[OF convex_alt, rule_format, of C y x ?\<mu>] |
|
747 |
by (simp add: algebra_simps) |
|
| 36623 | 748 |
have "?\<mu> * x + (1 - ?\<mu>) * y = (y - z) * x / (y - x) + (1 - (y - z) / (y - x)) * y" |
| 60423 | 749 |
by (auto simp: field_simps) |
| 36623 | 750 |
also have "\<dots> = ((y - z) * x + (y - x - (y - z)) * y) / (y - x)" |
| 63485 | 751 |
using assms by (simp only: add_divide_distrib) (auto simp: field_simps) |
| 36623 | 752 |
also have "\<dots> = z" |
| 49609 | 753 |
using assms by (auto simp: field_simps) |
| 60423 | 754 |
finally show ?thesis |
| 56796 | 755 |
using comb by auto |
| 60423 | 756 |
qed |
| 63485 | 757 |
show "z \<in> C" |
758 |
using z less assms by (auto simp: le_less) |
|
| 36623 | 759 |
qed |
760 |
||
761 |
lemma f''_imp_f': |
|
762 |
fixes f :: "real \<Rightarrow> real" |
|
763 |
assumes "convex C" |
|
| 49609 | 764 |
and f': "\<And>x. x \<in> C \<Longrightarrow> DERIV f x :> (f' x)" |
765 |
and f'': "\<And>x. x \<in> C \<Longrightarrow> DERIV f' x :> (f'' x)" |
|
766 |
and pos: "\<And>x. x \<in> C \<Longrightarrow> f'' x \<ge> 0" |
|
| 63485 | 767 |
and x: "x \<in> C" |
768 |
and y: "y \<in> C" |
|
| 36623 | 769 |
shows "f' x * (y - x) \<le> f y - f x" |
| 49609 | 770 |
using assms |
| 36623 | 771 |
proof - |
| 63485 | 772 |
have less_imp: "f y - f x \<ge> f' x * (y - x)" "f' y * (x - y) \<le> f x - f y" |
773 |
if *: "x \<in> C" "y \<in> C" "y > x" for x y :: real |
|
774 |
proof - |
|
775 |
from * have ge: "y - x > 0" "y - x \<ge> 0" |
|
| 60423 | 776 |
by auto |
777 |
from * have le: "x - y < 0" "x - y \<le> 0" |
|
778 |
by auto |
|
| 36623 | 779 |
then obtain z1 where z1: "z1 > x" "z1 < y" "f y - f x = (y - x) * f' z1" |
| 60423 | 780 |
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>], |
| 63485 | 781 |
THEN f', THEN MVT2[OF \<open>x < y\<close>, rule_format, unfolded atLeastAtMost_iff[symmetric]]] |
| 36623 | 782 |
by auto |
| 60423 | 783 |
then have "z1 \<in> C" |
784 |
using atMostAtLeast_subset_convex \<open>convex C\<close> \<open>x \<in> C\<close> \<open>y \<in> C\<close> \<open>x < y\<close> |
|
785 |
by fastforce |
|
| 36623 | 786 |
from z1 have z1': "f x - f y = (x - y) * f' z1" |
| 60423 | 787 |
by (simp add: field_simps) |
| 36623 | 788 |
obtain z2 where z2: "z2 > x" "z2 < z1" "f' z1 - f' x = (z1 - x) * f'' z2" |
| 60423 | 789 |
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>], |
| 63485 | 790 |
THEN f'', THEN MVT2[OF \<open>x < z1\<close>, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1 |
| 36623 | 791 |
by auto |
792 |
obtain z3 where z3: "z3 > z1" "z3 < y" "f' y - f' z1 = (y - z1) * f'' z3" |
|
| 60423 | 793 |
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>], |
| 63485 | 794 |
THEN f'', THEN MVT2[OF \<open>z1 < y\<close>, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1 |
| 36623 | 795 |
by auto |
796 |
have "f' y - (f x - f y) / (x - y) = f' y - f' z1" |
|
| 60423 | 797 |
using * z1' by auto |
798 |
also have "\<dots> = (y - z1) * f'' z3" |
|
799 |
using z3 by auto |
|
800 |
finally have cool': "f' y - (f x - f y) / (x - y) = (y - z1) * f'' z3" |
|
801 |
by simp |
|
802 |
have A': "y - z1 \<ge> 0" |
|
803 |
using z1 by auto |
|
804 |
have "z3 \<in> C" |
|
805 |
using z3 * atMostAtLeast_subset_convex \<open>convex C\<close> \<open>x \<in> C\<close> \<open>z1 \<in> C\<close> \<open>x < z1\<close> |
|
806 |
by fastforce |
|
807 |
then have B': "f'' z3 \<ge> 0" |
|
808 |
using assms by auto |
|
809 |
from A' B' have "(y - z1) * f'' z3 \<ge> 0" |
|
810 |
by auto |
|
811 |
from cool' this have "f' y - (f x - f y) / (x - y) \<ge> 0" |
|
812 |
by auto |
|
| 36623 | 813 |
from mult_right_mono_neg[OF this le(2)] |
814 |
have "f' y * (x - y) - (f x - f y) / (x - y) * (x - y) \<le> 0 * (x - y)" |
|
|
36778
739a9379e29b
avoid using real-specific versions of generic lemmas
huffman
parents:
36648
diff
changeset
|
815 |
by (simp add: algebra_simps) |
| 60423 | 816 |
then have "f' y * (x - y) - (f x - f y) \<le> 0" |
817 |
using le by auto |
|
818 |
then have res: "f' y * (x - y) \<le> f x - f y" |
|
819 |
by auto |
|
| 36623 | 820 |
have "(f y - f x) / (y - x) - f' x = f' z1 - f' x" |
| 60423 | 821 |
using * z1 by auto |
822 |
also have "\<dots> = (z1 - x) * f'' z2" |
|
823 |
using z2 by auto |
|
824 |
finally have cool: "(f y - f x) / (y - x) - f' x = (z1 - x) * f'' z2" |
|
825 |
by simp |
|
826 |
have A: "z1 - x \<ge> 0" |
|
827 |
using z1 by auto |
|
828 |
have "z2 \<in> C" |
|
829 |
using z2 z1 * atMostAtLeast_subset_convex \<open>convex C\<close> \<open>z1 \<in> C\<close> \<open>y \<in> C\<close> \<open>z1 < y\<close> |
|
830 |
by fastforce |
|
831 |
then have B: "f'' z2 \<ge> 0" |
|
832 |
using assms by auto |
|
833 |
from A B have "(z1 - x) * f'' z2 \<ge> 0" |
|
834 |
by auto |
|
835 |
with cool have "(f y - f x) / (y - x) - f' x \<ge> 0" |
|
836 |
by auto |
|
| 36623 | 837 |
from mult_right_mono[OF this ge(2)] |
838 |
have "(f y - f x) / (y - x) * (y - x) - f' x * (y - x) \<ge> 0 * (y - x)" |
|
|
36778
739a9379e29b
avoid using real-specific versions of generic lemmas
huffman
parents:
36648
diff
changeset
|
839 |
by (simp add: algebra_simps) |
| 60423 | 840 |
then have "f y - f x - f' x * (y - x) \<ge> 0" |
841 |
using ge by auto |
|
| 63485 | 842 |
then show "f y - f x \<ge> f' x * (y - x)" "f' y * (x - y) \<le> f x - f y" |
| 60423 | 843 |
using res by auto |
| 63485 | 844 |
qed |
845 |
show ?thesis |
|
846 |
proof (cases "x = y") |
|
847 |
case True |
|
848 |
with x y show ?thesis by auto |
|
849 |
next |
|
850 |
case False |
|
851 |
with less_imp x y show ?thesis |
|
852 |
by (auto simp: neq_iff) |
|
853 |
qed |
|
| 36623 | 854 |
qed |
855 |
||
856 |
lemma f''_ge0_imp_convex: |
|
857 |
fixes f :: "real \<Rightarrow> real" |
|
858 |
assumes conv: "convex C" |
|
| 49609 | 859 |
and f': "\<And>x. x \<in> C \<Longrightarrow> DERIV f x :> (f' x)" |
860 |
and f'': "\<And>x. x \<in> C \<Longrightarrow> DERIV f' x :> (f'' x)" |
|
861 |
and pos: "\<And>x. x \<in> C \<Longrightarrow> f'' x \<ge> 0" |
|
| 36623 | 862 |
shows "convex_on C f" |
| 56796 | 863 |
using f''_imp_f'[OF conv f' f'' pos] assms pos_convex_function |
864 |
by fastforce |
|
| 36623 | 865 |
|
866 |
lemma minus_log_convex: |
|
867 |
fixes b :: real |
|
868 |
assumes "b > 1" |
|
869 |
shows "convex_on {0 <..} (\<lambda> x. - log b x)"
|
|
870 |
proof - |
|
| 56796 | 871 |
have "\<And>z. z > 0 \<Longrightarrow> DERIV (log b) z :> 1 / (ln b * z)" |
872 |
using DERIV_log by auto |
|
| 49609 | 873 |
then have f': "\<And>z. z > 0 \<Longrightarrow> DERIV (\<lambda> z. - log b z) z :> - 1 / (ln b * z)" |
|
56479
91958d4b30f7
revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents:
56409
diff
changeset
|
874 |
by (auto simp: DERIV_minus) |
| 63485 | 875 |
have "\<And>z::real. z > 0 \<Longrightarrow> DERIV inverse z :> - (inverse z ^ Suc (Suc 0))" |
| 36623 | 876 |
using less_imp_neq[THEN not_sym, THEN DERIV_inverse] by auto |
877 |
from this[THEN DERIV_cmult, of _ "- 1 / ln b"] |
|
| 63485 | 878 |
have "\<And>z::real. z > 0 \<Longrightarrow> |
| 49609 | 879 |
DERIV (\<lambda> z. (- 1 / ln b) * inverse z) z :> (- 1 / ln b) * (- (inverse z ^ Suc (Suc 0)))" |
| 36623 | 880 |
by auto |
| 56796 | 881 |
then have f''0: "\<And>z::real. z > 0 \<Longrightarrow> |
882 |
DERIV (\<lambda> z. - 1 / (ln b * z)) z :> 1 / (ln b * z * z)" |
|
| 60423 | 883 |
unfolding inverse_eq_divide by (auto simp: mult.assoc) |
| 56796 | 884 |
have f''_ge0: "\<And>z::real. z > 0 \<Longrightarrow> 1 / (ln b * z * z) \<ge> 0" |
| 60423 | 885 |
using \<open>b > 1\<close> by (auto intro!: less_imp_le) |
| 63485 | 886 |
from f''_ge0_imp_convex[OF pos_is_convex, unfolded greaterThan_iff, OF f' f''0 f''_ge0] |
887 |
show ?thesis |
|
888 |
by auto |
|
| 36623 | 889 |
qed |
890 |
||
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
891 |
|
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
892 |
subsection \<open>Convexity of real functions\<close> |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
893 |
|
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
894 |
lemma convex_on_realI: |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
895 |
assumes "connected A" |
| 63485 | 896 |
and "\<And>x. x \<in> A \<Longrightarrow> (f has_real_derivative f' x) (at x)" |
897 |
and "\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x \<le> y \<Longrightarrow> f' x \<le> f' y" |
|
898 |
shows "convex_on A f" |
|
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
899 |
proof (rule convex_on_linorderI) |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
900 |
fix t x y :: real |
| 63485 | 901 |
assume t: "t > 0" "t < 1" |
902 |
assume xy: "x \<in> A" "y \<in> A" "x < y" |
|
| 63040 | 903 |
define z where "z = (1 - t) * x + t * y" |
| 63485 | 904 |
with \<open>connected A\<close> and xy have ivl: "{x..y} \<subseteq> A"
|
905 |
using connected_contains_Icc by blast |
|
|
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
906 |
|
| 63485 | 907 |
from xy t have xz: "z > x" |
908 |
by (simp add: z_def algebra_simps) |
|
909 |
have "y - z = (1 - t) * (y - x)" |
|
910 |
by (simp add: z_def algebra_simps) |
|
911 |
also from xy t have "\<dots> > 0" |
|
912 |
by (intro mult_pos_pos) simp_all |
|
913 |
finally have yz: "z < y" |
|
914 |
by simp |
|
|
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
915 |
|
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
916 |
from assms xz yz ivl t have "\<exists>\<xi>. \<xi> > x \<and> \<xi> < z \<and> f z - f x = (z - x) * f' \<xi>" |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
917 |
by (intro MVT2) (auto intro!: assms(2)) |
| 63485 | 918 |
then obtain \<xi> where \<xi>: "\<xi> > x" "\<xi> < z" "f' \<xi> = (f z - f x) / (z - x)" |
919 |
by auto |
|
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
920 |
from assms xz yz ivl t have "\<exists>\<eta>. \<eta> > z \<and> \<eta> < y \<and> f y - f z = (y - z) * f' \<eta>" |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
921 |
by (intro MVT2) (auto intro!: assms(2)) |
| 63485 | 922 |
then obtain \<eta> where \<eta>: "\<eta> > z" "\<eta> < y" "f' \<eta> = (f y - f z) / (y - z)" |
923 |
by auto |
|
|
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
924 |
|
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
925 |
from \<eta>(3) have "(f y - f z) / (y - z) = f' \<eta>" .. |
| 63485 | 926 |
also from \<xi> \<eta> ivl have "\<xi> \<in> A" "\<eta> \<in> A" |
927 |
by auto |
|
928 |
with \<xi> \<eta> have "f' \<eta> \<ge> f' \<xi>" |
|
929 |
by (intro assms(3)) auto |
|
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
930 |
also from \<xi>(3) have "f' \<xi> = (f z - f x) / (z - x)" . |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
931 |
finally have "(f y - f z) * (z - x) \<ge> (f z - f x) * (y - z)" |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
932 |
using xz yz by (simp add: field_simps) |
| 63485 | 933 |
also have "z - x = t * (y - x)" |
934 |
by (simp add: z_def algebra_simps) |
|
935 |
also have "y - z = (1 - t) * (y - x)" |
|
936 |
by (simp add: z_def algebra_simps) |
|
937 |
finally have "(f y - f z) * t \<ge> (f z - f x) * (1 - t)" |
|
938 |
using xy by simp |
|
939 |
then show "(1 - t) * f x + t * f y \<ge> f ((1 - t) *\<^sub>R x + t *\<^sub>R y)" |
|
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
940 |
by (simp add: z_def algebra_simps) |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
941 |
qed |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
942 |
|
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
943 |
lemma convex_on_inverse: |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
944 |
assumes "A \<subseteq> {0<..}"
|
| 63485 | 945 |
shows "convex_on A (inverse :: real \<Rightarrow> real)" |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
946 |
proof (rule convex_on_subset[OF _ assms], intro convex_on_realI[of _ _ "\<lambda>x. -inverse (x^2)"]) |
| 63485 | 947 |
fix u v :: real |
948 |
assume "u \<in> {0<..}" "v \<in> {0<..}" "u \<le> v"
|
|
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
949 |
with assms show "-inverse (u^2) \<le> -inverse (v^2)" |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
950 |
by (intro le_imp_neg_le le_imp_inverse_le power_mono) (simp_all) |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
951 |
qed (insert assms, auto intro!: derivative_eq_intros simp: divide_simps power2_eq_square) |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
952 |
|
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
953 |
lemma convex_onD_Icc': |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
954 |
assumes "convex_on {x..y} f" "c \<in> {x..y}"
|
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
955 |
defines "d \<equiv> y - x" |
| 63485 | 956 |
shows "f c \<le> (f y - f x) / d * (c - x) + f x" |
957 |
proof (cases x y rule: linorder_cases) |
|
958 |
case less |
|
959 |
then have d: "d > 0" |
|
960 |
by (simp add: d_def) |
|
|
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
961 |
from assms(2) less have A: "0 \<le> (c - x) / d" "(c - x) / d \<le> 1" |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
962 |
by (simp_all add: d_def divide_simps) |
| 63485 | 963 |
have "f c = f (x + (c - x) * 1)" |
964 |
by simp |
|
965 |
also from less have "1 = ((y - x) / d)" |
|
966 |
by (simp add: d_def) |
|
|
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
967 |
also from d have "x + (c - x) * \<dots> = (1 - (c - x) / d) *\<^sub>R x + ((c - x) / d) *\<^sub>R y" |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
968 |
by (simp add: field_simps) |
| 63485 | 969 |
also have "f \<dots> \<le> (1 - (c - x) / d) * f x + (c - x) / d * f y" |
970 |
using assms less by (intro convex_onD_Icc) simp_all |
|
971 |
also from d have "\<dots> = (f y - f x) / d * (c - x) + f x" |
|
972 |
by (simp add: field_simps) |
|
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
973 |
finally show ?thesis . |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
974 |
qed (insert assms(2), simp_all) |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
975 |
|
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
976 |
lemma convex_onD_Icc'': |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
977 |
assumes "convex_on {x..y} f" "c \<in> {x..y}"
|
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
978 |
defines "d \<equiv> y - x" |
| 63485 | 979 |
shows "f c \<le> (f x - f y) / d * (y - c) + f y" |
980 |
proof (cases x y rule: linorder_cases) |
|
981 |
case less |
|
982 |
then have d: "d > 0" |
|
983 |
by (simp add: d_def) |
|
|
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
984 |
from assms(2) less have A: "0 \<le> (y - c) / d" "(y - c) / d \<le> 1" |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
985 |
by (simp_all add: d_def divide_simps) |
| 63485 | 986 |
have "f c = f (y - (y - c) * 1)" |
987 |
by simp |
|
988 |
also from less have "1 = ((y - x) / d)" |
|
989 |
by (simp add: d_def) |
|
|
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
990 |
also from d have "y - (y - c) * \<dots> = (1 - (1 - (y - c) / d)) *\<^sub>R x + (1 - (y - c) / d) *\<^sub>R y" |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
991 |
by (simp add: field_simps) |
| 63485 | 992 |
also have "f \<dots> \<le> (1 - (1 - (y - c) / d)) * f x + (1 - (y - c) / d) * f y" |
993 |
using assms less by (intro convex_onD_Icc) (simp_all add: field_simps) |
|
994 |
also from d have "\<dots> = (f x - f y) / d * (y - c) + f y" |
|
995 |
by (simp add: field_simps) |
|
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
996 |
finally show ?thesis . |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
997 |
qed (insert assms(2), simp_all) |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61520
diff
changeset
|
998 |
|
| 36623 | 999 |
end |