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