src/HOL/Library/Convex.thy
author wenzelm
Tue, 26 Apr 2016 22:44:31 +0200
changeset 63060 293ede07b775
parent 63040 eb4ddd18d635
child 63092 a949b2a5f51d
permissions -rw-r--r--
some uses of 'obtain' with structure statement;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
36648
43b66dcd9266 Add Convex to Library build
hoelzl
parents: 36623
diff changeset
     1
(*  Title:      HOL/Library/Convex.thy
43b66dcd9266 Add Convex to Library build
hoelzl
parents: 36623
diff changeset
     2
    Author:     Armin Heller, TU Muenchen
43b66dcd9266 Add Convex to Library build
hoelzl
parents: 36623
diff changeset
     3
    Author:     Johannes Hoelzl, TU Muenchen
43b66dcd9266 Add Convex to Library build
hoelzl
parents: 36623
diff changeset
     4
*)
43b66dcd9266 Add Convex to Library build
hoelzl
parents: 36623
diff changeset
     5
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
     6
section \<open>Convexity in real vector spaces\<close>
36648
43b66dcd9266 Add Convex to Library build
hoelzl
parents: 36623
diff changeset
     7
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
     8
theory Convex
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
     9
imports Product_Vector
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    10
begin
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    11
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
    12
subsection \<open>Convexity\<close>
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    13
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
    14
definition convex :: "'a::real_vector set \<Rightarrow> bool"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
    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
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    16
53676
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
    17
lemma convexI:
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
    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"
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
    19
  shows "convex s"
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
    20
  using assms unfolding convex_def by fast
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
    21
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
    22
lemma convexD:
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
    23
  assumes "convex s" and "x \<in> s" and "y \<in> s" and "0 \<le> u" and "0 \<le> v" and "u + v = 1"
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
    24
  shows "u *\<^sub>R x + v *\<^sub>R y \<in> s"
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
    25
  using assms unfolding convex_def by fast
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
    26
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    27
lemma convex_alt:
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    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)"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    29
  (is "_ \<longleftrightarrow> ?alt")
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    30
proof
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    31
  assume alt[rule_format]: ?alt
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
    32
  {
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
    33
    fix x y and u v :: real
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
    34
    assume mem: "x \<in> s" "y \<in> s"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
    35
    assume "0 \<le> u" "0 \<le> v"
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
    36
    moreover
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
    37
    assume "u + v = 1"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
    38
    then have "u = 1 - v" by auto
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
    39
    ultimately have "u *\<^sub>R x + v *\<^sub>R y \<in> s"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
    40
      using alt[OF mem] by auto
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
    41
  }
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
    42
  then show "convex s"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
    43
    unfolding convex_def by auto
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    44
qed (auto simp: convex_def)
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    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
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    47
  assumes "convex s" "a \<in> s" "b \<in> s" "0 \<le> u" "u \<le> 1"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    48
  shows "((1 - u) *\<^sub>R a + u *\<^sub>R b) \<in> s"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    49
  using assms unfolding convex_alt by auto
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    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
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60178
diff changeset
    59
lemma convex_empty[intro,simp]: "convex {}"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    60
  unfolding convex_def by simp
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    61
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60178
diff changeset
    62
lemma convex_singleton[intro,simp]: "convex {a}"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    63
  unfolding convex_def by (auto simp: scaleR_left_distrib[symmetric])
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    64
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60178
diff changeset
    65
lemma convex_UNIV[intro,simp]: "convex UNIV"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    66
  unfolding convex_def by auto
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    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
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    69
  unfolding convex_def by auto
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    70
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    71
lemma convex_Int: "convex s \<Longrightarrow> convex t \<Longrightarrow> convex (s \<inter> t)"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    72
  unfolding convex_def by auto
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    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
d29d63460d84 new lemmas
huffman
parents: 51642
diff changeset
    75
  unfolding convex_def by auto
d29d63460d84 new lemmas
huffman
parents: 51642
diff changeset
    76
d29d63460d84 new lemmas
huffman
parents: 51642
diff changeset
    77
lemma convex_Times: "convex s \<Longrightarrow> convex t \<Longrightarrow> convex (s \<times> t)"
d29d63460d84 new lemmas
huffman
parents: 51642
diff changeset
    78
  unfolding convex_def by auto
d29d63460d84 new lemmas
huffman
parents: 51642
diff changeset
    79
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    80
lemma convex_halfspace_le: "convex {x. inner a x \<le> b}"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    81
  unfolding convex_def
44142
8e27e0177518 avoid warnings about duplicate rules
huffman
parents: 43337
diff changeset
    82
  by (auto simp: inner_add intro!: convex_bound_le)
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    83
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    84
lemma convex_halfspace_ge: "convex {x. inner a x \<ge> b}"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    85
proof -
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
    86
  have *: "{x. inner a x \<ge> b} = {x. inner (-a) x \<le> -b}"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
    87
    by auto
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
    88
  show ?thesis
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
    89
    unfolding * using convex_halfspace_le[of "-a" "-b"] by auto
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    90
qed
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    91
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    92
lemma convex_hyperplane: "convex {x. inner a x = b}"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
    93
proof -
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
    94
  have *: "{x. inner a x = b} = {x. inner a x \<le> b} \<inter> {x. inner a x \<ge> b}"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
    95
    by auto
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    96
  show ?thesis using convex_halfspace_le convex_halfspace_ge
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    97
    by (auto intro!: convex_Int simp: *)
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    98
qed
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
    99
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   100
lemma convex_halfspace_lt: "convex {x. inner a x < b}"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   101
  unfolding convex_def
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   102
  by (auto simp: convex_bound_lt inner_add)
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   103
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   104
lemma convex_halfspace_gt: "convex {x. inner a x > b}"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   105
   using convex_halfspace_lt[of "-a" "-b"] by auto
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   106
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   107
lemma convex_real_interval [iff]:
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   108
  fixes a b :: "real"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   109
  shows "convex {a..}" and "convex {..b}"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   110
    and "convex {a<..}" and "convex {..<b}"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   111
    and "convex {a..b}" and "convex {a<..b}"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   112
    and "convex {a..<b}" and "convex {a<..<b}"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   113
proof -
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   114
  have "{a..} = {x. a \<le> inner 1 x}"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   115
    by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   116
  then show 1: "convex {a..}"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   117
    by (simp only: convex_halfspace_ge)
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   118
  have "{..b} = {x. inner 1 x \<le> b}"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   119
    by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   120
  then show 2: "convex {..b}"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   121
    by (simp only: convex_halfspace_le)
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   122
  have "{a<..} = {x. a < inner 1 x}"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   123
    by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   124
  then show 3: "convex {a<..}"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   125
    by (simp only: convex_halfspace_gt)
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   126
  have "{..<b} = {x. inner 1 x < b}"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   127
    by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   128
  then show 4: "convex {..<b}"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   129
    by (simp only: convex_halfspace_lt)
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   130
  have "{a..b} = {a..} \<inter> {..b}"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   131
    by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   132
  then show "convex {a..b}"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   133
    by (simp only: convex_Int 1 2)
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   134
  have "{a<..b} = {a<..} \<inter> {..b}"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   135
    by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   136
  then show "convex {a<..b}"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   137
    by (simp only: convex_Int 3 2)
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   138
  have "{a..<b} = {a..} \<inter> {..<b}"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   139
    by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   140
  then show "convex {a..<b}"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   141
    by (simp only: convex_Int 1 4)
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   142
  have "{a<..<b} = {a<..} \<inter> {..<b}"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   143
    by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   144
  then show "convex {a<..<b}"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   145
    by (simp only: convex_Int 3 4)
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   146
qed
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   147
61070
b72a990adfe2 prefer symbols;
wenzelm
parents: 60449
diff changeset
   148
lemma convex_Reals: "convex \<real>"
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59557
diff changeset
   149
  by (simp add: convex_def scaleR_conv_of_real)
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   150
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   151
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   152
subsection \<open>Explicit expressions for convexity in terms of arbitrary sums\<close>
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   153
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   154
lemma convex_setsum:
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   155
  fixes C :: "'a::real_vector set"
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   156
  assumes "finite s"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   157
    and "convex C"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   158
    and "(\<Sum> i \<in> s. a i) = 1"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   159
  assumes "\<And>i. i \<in> s \<Longrightarrow> a i \<ge> 0"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   160
    and "\<And>i. i \<in> s \<Longrightarrow> y i \<in> C"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   161
  shows "(\<Sum> j \<in> s. a j *\<^sub>R y j) \<in> C"
55909
df6133adb63f tuned proof script
huffman
parents: 54230
diff changeset
   162
  using assms(1,3,4,5)
df6133adb63f tuned proof script
huffman
parents: 54230
diff changeset
   163
proof (induct arbitrary: a set: finite)
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   164
  case empty
55909
df6133adb63f tuned proof script
huffman
parents: 54230
diff changeset
   165
  then show ?case by simp
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   166
next
55909
df6133adb63f tuned proof script
huffman
parents: 54230
diff changeset
   167
  case (insert i s) note IH = this(3)
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   168
  have "a i + setsum a s = 1"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   169
    and "0 \<le> a i"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   170
    and "\<forall>j\<in>s. 0 \<le> a j"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   171
    and "y i \<in> C"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   172
    and "\<forall>j\<in>s. y j \<in> C"
55909
df6133adb63f tuned proof script
huffman
parents: 54230
diff changeset
   173
    using insert.hyps(1,2) insert.prems by simp_all
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   174
  then have "0 \<le> setsum a s"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   175
    by (simp add: setsum_nonneg)
55909
df6133adb63f tuned proof script
huffman
parents: 54230
diff changeset
   176
  have "a i *\<^sub>R y i + (\<Sum>j\<in>s. a j *\<^sub>R y j) \<in> C"
df6133adb63f tuned proof script
huffman
parents: 54230
diff changeset
   177
  proof (cases)
df6133adb63f tuned proof script
huffman
parents: 54230
diff changeset
   178
    assume z: "setsum a s = 0"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   179
    with \<open>a i + setsum a s = 1\<close> have "a i = 1"
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   180
      by simp
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   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
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   182
      by simp
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   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
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   184
      by simp
55909
df6133adb63f tuned proof script
huffman
parents: 54230
diff changeset
   185
  next
df6133adb63f tuned proof script
huffman
parents: 54230
diff changeset
   186
    assume nz: "setsum a s \<noteq> 0"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   187
    with \<open>0 \<le> setsum a s\<close> have "0 < setsum a s"
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   188
      by simp
55909
df6133adb63f tuned proof script
huffman
parents: 54230
diff changeset
   189
    then have "(\<Sum>j\<in>s. (a j / setsum a s) *\<^sub>R y j) \<in> C"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   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
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   192
    from \<open>convex C\<close> and \<open>y i \<in> C\<close> and this and \<open>0 \<le> a i\<close>
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   193
      and \<open>0 \<le> setsum a s\<close> and \<open>a i + setsum a s = 1\<close>
55909
df6133adb63f tuned proof script
huffman
parents: 54230
diff changeset
   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"
df6133adb63f tuned proof script
huffman
parents: 54230
diff changeset
   195
      by (rule convexD)
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   196
    then show ?thesis
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   197
      by (simp add: scaleR_setsum_right nz)
55909
df6133adb63f tuned proof script
huffman
parents: 54230
diff changeset
   198
  qed
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   199
  then show ?case using \<open>finite s\<close> and \<open>i \<notin> s\<close>
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   200
    by simp
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   201
qed
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   202
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   203
lemma convex:
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   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)
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   205
      \<longrightarrow> setsum (\<lambda>i. u i *\<^sub>R x i) {1..k} \<in> s)"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   206
proof safe
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   207
  fix k :: nat
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   208
  fix u :: "nat \<Rightarrow> real"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   209
  fix x
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   210
  assume "convex s"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   211
    "\<forall>i. 1 \<le> i \<and> i \<le> k \<longrightarrow> 0 \<le> u i \<and> x i \<in> s"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   212
    "setsum u {1..k} = 1"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   213
  with convex_setsum[of "{1 .. k}" s] show "(\<Sum>j\<in>{1 .. k}. u j *\<^sub>R x j) \<in> s"
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   214
    by auto
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   215
next
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   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
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   217
    \<longrightarrow> (\<Sum>i = 1..k. u i *\<^sub>R (x i :: 'a)) \<in> s"
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   218
  {
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   219
    fix \<mu> :: real
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   220
    fix x y :: 'a
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   221
    assume xy: "x \<in> s" "y \<in> s"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   222
    assume mu: "\<mu> \<ge> 0" "\<mu> \<le> 1"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   223
    let ?u = "\<lambda>i. if (i :: nat) = 1 then \<mu> else 1 - \<mu>"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   224
    let ?x = "\<lambda>i. if (i :: nat) = 1 then x else y"
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   225
    have "{1 :: nat .. 2} \<inter> - {x. x = 1} = {2}"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   226
      by auto
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   227
    then have card: "card ({1 :: nat .. 2} \<inter> - {x. x = 1}) = 1"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   228
      by simp
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   229
    then have "setsum ?u {1 .. 2} = 1"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56796
diff changeset
   230
      using setsum.If_cases[of "{(1 :: nat) .. 2}" "\<lambda> x. x = 1" "\<lambda> x. \<mu>" "\<lambda> x. 1 - \<mu>"]
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   231
      by auto
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   232
    with *[rule_format, of "2" ?u ?x] have s: "(\<Sum>j \<in> {1..2}. ?u j *\<^sub>R ?x j) \<in> s"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   233
      using mu xy by auto
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   234
    have grarr: "(\<Sum>j \<in> {Suc (Suc 0)..2}. ?u j *\<^sub>R ?x j) = (1 - \<mu>) *\<^sub>R y"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   235
      using setsum_head_Suc[of "Suc (Suc 0)" 2 "\<lambda> j. (1 - \<mu>) *\<^sub>R y"] by auto
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   236
    from setsum_head_Suc[of "Suc 0" 2 "\<lambda> j. ?u j *\<^sub>R ?x j", simplified this]
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   237
    have "(\<Sum>j \<in> {1..2}. ?u j *\<^sub>R ?x j) = \<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   238
      by auto
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   239
    then have "(1 - \<mu>) *\<^sub>R y + \<mu> *\<^sub>R x \<in> s"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   240
      using s by (auto simp: add.commute)
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   241
  }
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   242
  then show "convex s"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   243
    unfolding convex_alt by auto
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   244
qed
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   245
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   246
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   247
lemma convex_explicit:
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   248
  fixes s :: "'a::real_vector set"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   249
  shows "convex s \<longleftrightarrow>
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   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
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   251
proof safe
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   252
  fix t
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   253
  fix u :: "'a \<Rightarrow> real"
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   254
  assume "convex s"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   255
    and "finite t"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   256
    and "t \<subseteq> s" "\<forall>x\<in>t. 0 \<le> u x" "setsum u t = 1"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   257
  then show "(\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   258
    using convex_setsum[of t s u "\<lambda> x. x"] by auto
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   259
next
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   260
  assume *: "\<forall>t. \<forall> u. finite t \<and> t \<subseteq> s \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and>
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   261
    setsum u t = 1 \<longrightarrow> (\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   262
  show "convex s"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   263
    unfolding convex_alt
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   264
  proof safe
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   265
    fix x y
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   266
    fix \<mu> :: real
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   267
    assume **: "x \<in> s" "y \<in> s" "0 \<le> \<mu>" "\<mu> \<le> 1"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   268
    show "(1 - \<mu>) *\<^sub>R x + \<mu> *\<^sub>R y \<in> s"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   269
    proof (cases "x = y")
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   270
      case False
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   271
      then show ?thesis
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   272
        using *[rule_format, of "{x, y}" "\<lambda> z. if z = x then 1 - \<mu> else \<mu>"] **
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   273
          by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   274
    next
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   275
      case True
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   276
      then show ?thesis
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   277
        using *[rule_format, of "{x, y}" "\<lambda> z. 1"] **
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   278
          by (auto simp: field_simps real_vector.scale_left_diff_distrib)
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   279
    qed
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   280
  qed
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   281
qed
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   282
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   283
lemma convex_finite:
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   284
  assumes "finite s"
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   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
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   286
  unfolding convex_explicit
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   287
proof safe
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   288
  fix t u
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   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
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   290
    and as: "finite t" "t \<subseteq> s" "\<forall>x\<in>t. 0 \<le> u x" "setsum u t = (1::real)"
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   291
  have *: "s \<inter> t = t"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   292
    using as(2) by auto
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   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)"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   294
    by simp
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   295
  show "(\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   296
   using sum[THEN spec[where x="\<lambda>x. if x\<in>t then u x else 0"]] as *
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56796
diff changeset
   297
   by (auto simp: assms setsum.If_cases if_distrib if_distrib_arg)
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   298
qed (erule_tac x=s in allE, erule_tac x=u in allE, auto)
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   299
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   300
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   301
subsection \<open>Functions that are convex on a set\<close>
55909
df6133adb63f tuned proof script
huffman
parents: 54230
diff changeset
   302
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   303
definition convex_on :: "'a::real_vector set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> bool"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   304
  where "convex_on s f \<longleftrightarrow>
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   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
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   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
f1b0908028cf tuned proof;
wenzelm
parents: 62376
diff changeset
   326
  with assms[of t x y] assms[of "1 - t" y x]
f1b0908028cf tuned proof;
wenzelm
parents: 62376
diff changeset
   327
  show "f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \<le> (1 - t) * f x + t * f y"
f1b0908028cf tuned proof;
wenzelm
parents: 62376
diff changeset
   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
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   343
lemma convex_on_subset: "convex_on t f \<Longrightarrow> s \<subseteq> t \<Longrightarrow> convex_on s f"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   344
  unfolding convex_on_def by auto
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   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
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   347
  assumes "convex_on s f"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   348
    and "convex_on s g"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   349
  shows "convex_on s (\<lambda>x. f x + g x)"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   350
proof -
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   351
  {
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   352
    fix x y
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   353
    assume "x \<in> s" "y \<in> s"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   354
    moreover
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   355
    fix u v :: real
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   356
    assume "0 \<le> u" "0 \<le> v" "u + v = 1"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   357
    ultimately
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   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
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   359
      using assms unfolding convex_on_def by (auto simp: add_mono)
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   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)"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   361
      by (simp add: field_simps)
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   362
  }
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   363
  then show ?thesis
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   364
    unfolding convex_on_def by auto
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   365
qed
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   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
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   368
  fixes c :: real
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   369
  assumes "0 \<le> c"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   370
    and "convex_on s f"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   371
  shows "convex_on s (\<lambda>x. c * f x)"
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   372
proof -
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   373
  have *: "\<And>u c fx v fy :: real. u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   374
    by (simp add: field_simps)
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   375
  show ?thesis using assms(2) and mult_left_mono [OF _ assms(1)]
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   376
    unfolding convex_on_def and * by auto
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   377
qed
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   378
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   379
lemma convex_lower:
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   380
  assumes "convex_on s f"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   381
    and "x \<in> s"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   382
    and "y \<in> s"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   383
    and "0 \<le> u"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   384
    and "0 \<le> v"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   385
    and "u + v = 1"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   386
  shows "f (u *\<^sub>R x + v *\<^sub>R y) \<le> max (f x) (f y)"
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   387
proof -
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   388
  let ?m = "max (f x) (f y)"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   389
  have "u * f x + v * f y \<le> u * max (f x) (f y) + v * max (f x) (f y)"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   390
    using assms(4,5) by (auto simp: mult_left_mono add_mono)
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   391
  also have "\<dots> = max (f x) (f y)"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   392
    using assms(6) by (simp add: distrib_right [symmetric])
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   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
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   395
qed
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   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
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   398
  fixes s :: "'a::real_normed_vector set"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   399
  shows "convex_on s (\<lambda>x. dist a x)"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   400
proof (auto simp: convex_on_def dist_norm)
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   401
  fix x y
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   402
  assume "x \<in> s" "y \<in> s"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   403
  fix u v :: real
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   404
  assume "0 \<le> u"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   405
  assume "0 \<le> v"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   406
  assume "u + v = 1"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   407
  have "a = u *\<^sub>R a + v *\<^sub>R a"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   408
    unfolding scaleR_left_distrib[symmetric] and \<open>u + v = 1\<close> by simp
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   409
  then have *: "a - (u *\<^sub>R x + v *\<^sub>R y) = (u *\<^sub>R (a - x)) + (v *\<^sub>R (a - y))"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   410
    by (auto simp: algebra_simps)
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   411
  show "norm (a - (u *\<^sub>R x + v *\<^sub>R y)) \<le> u * norm (a - x) + v * norm (a - y)"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   412
    unfolding * using norm_triangle_ineq[of "u *\<^sub>R (a - x)" "v *\<^sub>R (a - y)"]
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   413
    using \<open>0 \<le> u\<close> \<open>0 \<le> v\<close> by auto
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   414
qed
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   415
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   416
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   417
subsection \<open>Arithmetic operations on sets preserve convexity\<close>
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   418
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53596
diff changeset
   419
lemma convex_linear_image:
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   420
  assumes "linear f"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   421
    and "convex s"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   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
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   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
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   427
qed
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   428
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53596
diff changeset
   429
lemma convex_linear_vimage:
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   430
  assumes "linear f"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   431
    and "convex s"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   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
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   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
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   440
  assumes "convex s"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   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
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   443
  have "linear (\<lambda>x. c *\<^sub>R x)"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   444
    by (simp add: linearI scaleR_add_right)
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   445
  then show ?thesis
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   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
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   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
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   460
  assumes "convex s"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   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
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   463
  have "linear (\<lambda>x. - x)"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   464
    by (simp add: linearI)
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   465
  then show ?thesis
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   466
    using \<open>convex s\<close> by (rule convex_linear_image)
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   467
qed
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   468
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   469
lemma convex_sums:
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   470
  assumes "convex s"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   471
    and "convex t"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   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
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   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
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   481
qed
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   482
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   483
lemma convex_differences:
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   484
  assumes "convex s" "convex t"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   485
  shows "convex {x - y| x y. x \<in> s \<and> y \<in> t}"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   486
proof -
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   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
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   488
    by (auto simp: diff_conv_add_uminus simp del: add_uminus_conv_diff)
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   489
  then show ?thesis
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   490
    using convex_sums[OF assms(1) convex_negations[OF assms(2)]] by auto
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   491
qed
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   492
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   493
lemma convex_translation:
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   494
  assumes "convex s"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   495
  shows "convex ((\<lambda>x. a + x) ` s)"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   496
proof -
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   497
  have "{a + y |y. y \<in> s} = (\<lambda>x. a + x) ` s"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   498
    by auto
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   499
  then show ?thesis
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   500
    using convex_sums[OF convex_singleton[of a] assms] by auto
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   501
qed
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   502
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   503
lemma convex_affinity:
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   504
  assumes "convex s"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   505
  shows "convex ((\<lambda>x. a + c *\<^sub>R x) ` s)"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   506
proof -
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   507
  have "(\<lambda>x. a + c *\<^sub>R x) ` s = op + a ` op *\<^sub>R c ` s"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   508
    by auto
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   509
  then show ?thesis
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   510
    using convex_translation[OF convex_scaling[OF assms], of a c] by auto
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   511
qed
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   512
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   513
lemma pos_is_convex: "convex {0 :: real <..}"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   514
  unfolding convex_alt
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   515
proof safe
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   516
  fix y x \<mu> :: real
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   517
  assume *: "y > 0" "x > 0" "\<mu> \<ge> 0" "\<mu> \<le> 1"
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   518
  {
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   519
    assume "\<mu> = 0"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   520
    then have "\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y = y" by simp
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   521
    then have "\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y > 0" using * by simp
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   522
  }
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   523
  moreover
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   524
  {
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   525
    assume "\<mu> = 1"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   526
    then have "\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y > 0" using * by simp
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   527
  }
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   528
  moreover
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   529
  {
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   530
    assume "\<mu> \<noteq> 1" "\<mu> \<noteq> 0"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   531
    then have "\<mu> > 0" "(1 - \<mu>) > 0" using * by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   532
    then have "\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y > 0" using *
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   533
      by (auto simp: add_pos_pos)
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   534
  }
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   535
  ultimately show "(1 - \<mu>) *\<^sub>R y + \<mu> *\<^sub>R x > 0"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   536
    using assms by fastforce
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   537
qed
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   538
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   539
lemma convex_on_setsum:
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   540
  fixes a :: "'a \<Rightarrow> real"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   541
    and y :: "'a \<Rightarrow> 'b::real_vector"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   542
    and f :: "'b \<Rightarrow> real"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   543
  assumes "finite s" "s \<noteq> {}"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   544
    and "convex_on C f"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   545
    and "convex C"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   546
    and "(\<Sum> i \<in> s. a i) = 1"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   547
    and "\<And>i. i \<in> s \<Longrightarrow> a i \<ge> 0"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   548
    and "\<And>i. i \<in> s \<Longrightarrow> y i \<in> C"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   549
  shows "f (\<Sum> i \<in> s. a i *\<^sub>R y i) \<le> (\<Sum> i \<in> s. a i * f (y i))"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   550
  using assms
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   551
proof (induct s arbitrary: a rule: finite_ne_induct)
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   552
  case (singleton i)
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   553
  then have ai: "a i = 1" by auto
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   554
  then show ?case by auto
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   555
next
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   556
  case (insert i s)
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   557
  then have "convex_on C f" by simp
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   558
  from this[unfolded convex_on_def, rule_format]
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   559
  have conv: "\<And>x y \<mu>. x \<in> C \<Longrightarrow> y \<in> C \<Longrightarrow> 0 \<le> \<mu> \<Longrightarrow> \<mu> \<le> 1 \<Longrightarrow>
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   560
      f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   561
    by simp
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   562
  show ?case
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   563
  proof (cases "a i = 1")
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   564
    case True
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   565
    then have "(\<Sum> j \<in> s. a j) = 0"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   566
      using insert by auto
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   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
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   569
    then show ?thesis
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   570
      using insert by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   571
  next
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   572
    case False
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   573
    from insert have yai: "y i \<in> C" "a i \<ge> 0"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   574
      by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   575
    have fis: "finite (insert i s)"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   576
      using insert by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   577
    then have ai1: "a i \<le> 1"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   578
      using setsum_nonneg_leq_bound[of "insert i s" a] insert by simp
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   579
    then have "a i < 1"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   580
      using False by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   581
    then have i0: "1 - a i > 0"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   582
      by auto
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   583
    let ?a = "\<lambda>j. a j / (1 - a i)"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   584
    have a_nonneg: "?a j \<ge> 0" if "j \<in> s" for j
60449
229bad93377e renamed "prems" to "that";
wenzelm
parents: 60423
diff changeset
   585
      using i0 insert that by fastforce
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   586
    have "(\<Sum> j \<in> insert i s. a j) = 1"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   587
      using insert by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   588
    then have "(\<Sum> j \<in> s. a j) = 1 - a i"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   589
      using setsum.insert insert by fastforce
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   590
    then have "(\<Sum> j \<in> s. a j) / (1 - a i) = 1"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   591
      using i0 by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   592
    then have a1: "(\<Sum> j \<in> s. ?a j) = 1"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   593
      unfolding setsum_divide_distrib by simp
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   594
    have "convex C" using insert by auto
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   595
    then have asum: "(\<Sum> j \<in> s. ?a j *\<^sub>R y j) \<in> C"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   596
      using insert convex_setsum[OF \<open>finite s\<close>
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   597
        \<open>convex C\<close> a1 a_nonneg] by auto
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   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
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   599
      using a_nonneg a1 insert by blast
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   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
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   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
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   602
      by (auto simp only: add.commute)
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   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)"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   604
      using i0 by auto
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   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
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   606
      using scaleR_right.setsum[of "inverse (1 - a i)" "\<lambda> j. a j *\<^sub>R y j" s, symmetric]
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   607
      by (auto simp: algebra_simps)
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   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
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   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)"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   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
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   612
      by (auto simp: add.commute)
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   613
    also have "\<dots> \<le> (1 - a i) * (\<Sum> j \<in> s. ?a j * f (y j)) + a i * f (y i)"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   614
      using add_right_mono[OF mult_left_mono[of _ _ "1 - a i",
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   615
        OF asum_le less_imp_le[OF i0]], of "a i * f (y i)"] by simp
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   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
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   618
    also have "\<dots> = (\<Sum> j \<in> s. a j * f (y j)) + a i * f (y i)"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   619
      using i0 by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   620
    also have "\<dots> = (\<Sum> j \<in> insert i s. a j * f (y j))"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   621
      using insert by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   622
    finally show ?thesis
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   623
      by simp
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   624
  qed
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   625
qed
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   626
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   627
lemma convex_on_alt:
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   628
  fixes C :: "'a::real_vector set"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   629
  assumes "convex C"
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   630
  shows "convex_on C f \<longleftrightarrow>
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   631
    (\<forall>x \<in> C. \<forall> y \<in> C. \<forall> \<mu> :: real. \<mu> \<ge> 0 \<and> \<mu> \<le> 1 \<longrightarrow>
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   632
      f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y)"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   633
proof safe
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   634
  fix x y
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   635
  fix \<mu> :: real
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   636
  assume *: "convex_on C f" "x \<in> C" "y \<in> C" "0 \<le> \<mu>" "\<mu> \<le> 1"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   637
  from this[unfolded convex_on_def, rule_format]
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   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"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   639
    by auto
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   640
  from this[of "\<mu>" "1 - \<mu>", simplified] *
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   641
  show "f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   642
    by auto
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   643
next
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   644
  assume *: "\<forall>x\<in>C. \<forall>y\<in>C. \<forall>\<mu>. 0 \<le> \<mu> \<and> \<mu> \<le> 1 \<longrightarrow>
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   645
    f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   646
  {
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   647
    fix x y
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   648
    fix u v :: real
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   649
    assume **: "x \<in> C" "y \<in> C" "u \<ge> 0" "v \<ge> 0" "u + v = 1"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   650
    then have[simp]: "1 - u = v" by auto
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   651
    from *[rule_format, of x y u]
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   652
    have "f (u *\<^sub>R x + v *\<^sub>R y) \<le> u * f x + v * f y"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   653
      using ** by auto
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   654
  }
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   655
  then show "convex_on C f"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   656
    unfolding convex_on_def by auto
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   657
qed
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   658
43337
57a1c19f8e3b lemma about differences of convex functions
hoelzl
parents: 38642
diff changeset
   659
lemma convex_on_diff:
57a1c19f8e3b lemma about differences of convex functions
hoelzl
parents: 38642
diff changeset
   660
  fixes f :: "real \<Rightarrow> real"
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   661
  assumes f: "convex_on I f"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   662
    and I: "x \<in> I" "y \<in> I"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   663
    and t: "x < t" "t < y"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   664
  shows "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)"
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   665
    and "(f x - f y) / (x - y) \<le> (f t - f y) / (t - y)"
43337
57a1c19f8e3b lemma about differences of convex functions
hoelzl
parents: 38642
diff changeset
   666
proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63007
diff changeset
   667
  define a where "a \<equiv> (t - y) / (x - y)"
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   668
  with t have "0 \<le> a" "0 \<le> 1 - a"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   669
    by (auto simp: field_simps)
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   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
57a1c19f8e3b lemma about differences of convex functions
hoelzl
parents: 38642
diff changeset
   671
    by (auto simp: convex_on_def)
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   672
  have "a * x + (1 - a) * y = a * (x - y) + y"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   673
    by (simp add: field_simps)
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   674
  also have "\<dots> = t"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   675
    unfolding a_def using \<open>x < t\<close> \<open>t < y\<close> by simp
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   676
  finally have "f t \<le> a * f x + (1 - a) * f y"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   677
    using cvx by simp
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   678
  also have "\<dots> = a * (f x - f y) + f y"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   679
    by (simp add: field_simps)
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   680
  finally have "f t - f y \<le> a * (f x - f y)"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   681
    by simp
43337
57a1c19f8e3b lemma about differences of convex functions
hoelzl
parents: 38642
diff changeset
   682
  with t show "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)"
44142
8e27e0177518 avoid warnings about duplicate rules
huffman
parents: 43337
diff changeset
   683
    by (simp add: le_divide_eq divide_le_eq field_simps a_def)
43337
57a1c19f8e3b lemma about differences of convex functions
hoelzl
parents: 38642
diff changeset
   684
  with t show "(f x - f y) / (x - y) \<le> (f t - f y) / (t - y)"
44142
8e27e0177518 avoid warnings about duplicate rules
huffman
parents: 43337
diff changeset
   685
    by (simp add: le_divide_eq divide_le_eq field_simps)
43337
57a1c19f8e3b lemma about differences of convex functions
hoelzl
parents: 38642
diff changeset
   686
qed
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   687
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   688
lemma pos_convex_function:
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   689
  fixes f :: "real \<Rightarrow> real"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   690
  assumes "convex C"
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   691
    and leq: "\<And>x y. x \<in> C \<Longrightarrow> y \<in> C \<Longrightarrow> f' x * (y - x) \<le> f y - f x"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   692
  shows "convex_on C f"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   693
  unfolding convex_on_alt[OF assms(1)]
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   694
  using assms
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   695
proof safe
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   696
  fix x y \<mu> :: real
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   697
  let ?x = "\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   698
  assume *: "convex C" "x \<in> C" "y \<in> C" "\<mu> \<ge> 0" "\<mu> \<le> 1"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   699
  then have "1 - \<mu> \<ge> 0" by auto
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   700
  then have xpos: "?x \<in> C"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   701
    using * unfolding convex_alt by fastforce
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   702
  have geq: "\<mu> * (f x - f ?x) + (1 - \<mu>) * (f y - f ?x) \<ge>
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   703
      \<mu> * f' ?x * (x - ?x) + (1 - \<mu>) * f' ?x * (y - ?x)"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   704
    using add_mono[OF mult_left_mono[OF leq[OF xpos *(2)] \<open>\<mu> \<ge> 0\<close>]
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   705
      mult_left_mono[OF leq[OF xpos *(3)] \<open>1 - \<mu> \<ge> 0\<close>]]
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   706
    by auto
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   707
  then have "\<mu> * f x + (1 - \<mu>) * f y - f ?x \<ge> 0"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   708
    by (auto simp: field_simps)
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   709
  then show "f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   710
    using convex_on_alt by auto
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   711
qed
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   712
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   713
lemma atMostAtLeast_subset_convex:
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   714
  fixes C :: "real set"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   715
  assumes "convex C"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   716
    and "x \<in> C" "y \<in> C" "x < y"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   717
  shows "{x .. y} \<subseteq> C"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   718
proof safe
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   719
  fix z assume z: "z \<in> {x .. y}"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   720
  have less: "z \<in> C" if *: "x < z" "z < y"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   721
  proof -
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   722
    let ?\<mu> = "(y - z) / (y - x)"
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   723
    have "0 \<le> ?\<mu>" "?\<mu> \<le> 1"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   724
      using assms * by (auto simp: field_simps)
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   725
    then have comb: "?\<mu> * x + (1 - ?\<mu>) * y \<in> C"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   726
      using assms iffD1[OF convex_alt, rule_format, of C y x ?\<mu>]
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   727
      by (simp add: algebra_simps)
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   728
    have "?\<mu> * x + (1 - ?\<mu>) * y = (y - z) * x / (y - x) + (1 - (y - z) / (y - x)) * y"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   729
      by (auto simp: field_simps)
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   730
    also have "\<dots> = ((y - z) * x + (y - x - (y - z)) * y) / (y - x)"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   731
      using assms unfolding add_divide_distrib by (auto simp: field_simps)
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   732
    also have "\<dots> = z"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   733
      using assms by (auto simp: field_simps)
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   734
    finally show ?thesis
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   735
      using comb by auto
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   736
  qed
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   737
  show "z \<in> C" using z less assms
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   738
    unfolding atLeastAtMost_iff le_less by auto
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   739
qed
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   740
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   741
lemma f''_imp_f':
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   742
  fixes f :: "real \<Rightarrow> real"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   743
  assumes "convex C"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   744
    and f': "\<And>x. x \<in> C \<Longrightarrow> DERIV f x :> (f' x)"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   745
    and f'': "\<And>x. x \<in> C \<Longrightarrow> DERIV f' x :> (f'' x)"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   746
    and pos: "\<And>x. x \<in> C \<Longrightarrow> f'' x \<ge> 0"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   747
    and "x \<in> C" "y \<in> C"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   748
  shows "f' x * (y - x) \<le> f y - f x"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   749
  using assms
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   750
proof -
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   751
  {
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   752
    fix x y :: real
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   753
    assume *: "x \<in> C" "y \<in> C" "y > x"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   754
    then have ge: "y - x > 0" "y - x \<ge> 0"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   755
      by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   756
    from * have le: "x - y < 0" "x - y \<le> 0"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   757
      by auto
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   758
    then obtain z1 where z1: "z1 > x" "z1 < y" "f y - f x = (y - x) * f' z1"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   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>],
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   760
        THEN f', THEN MVT2[OF \<open>x < y\<close>, rule_format, unfolded atLeastAtMost_iff[symmetric]]]
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   761
      by auto
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   762
    then have "z1 \<in> C"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   763
      using atMostAtLeast_subset_convex \<open>convex C\<close> \<open>x \<in> C\<close> \<open>y \<in> C\<close> \<open>x < y\<close>
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   764
      by fastforce
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   765
    from z1 have z1': "f x - f y = (x - y) * f' z1"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   766
      by (simp add: field_simps)
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   767
    obtain z2 where z2: "z2 > x" "z2 < z1" "f' z1 - f' x = (z1 - x) * f'' z2"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   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>],
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   769
        THEN f'', THEN MVT2[OF \<open>x < z1\<close>, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   770
      by auto
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   771
    obtain z3 where z3: "z3 > z1" "z3 < y" "f' y - f' z1 = (y - z1) * f'' z3"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   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>],
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   773
        THEN f'', THEN MVT2[OF \<open>z1 < y\<close>, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   774
      by auto
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   775
    have "f' y - (f x - f y) / (x - y) = f' y - f' z1"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   776
      using * z1' by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   777
    also have "\<dots> = (y - z1) * f'' z3"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   778
      using z3 by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   779
    finally have cool': "f' y - (f x - f y) / (x - y) = (y - z1) * f'' z3"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   780
      by simp
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   781
    have A': "y - z1 \<ge> 0"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   782
      using z1 by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   783
    have "z3 \<in> C"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   784
      using z3 * atMostAtLeast_subset_convex \<open>convex C\<close> \<open>x \<in> C\<close> \<open>z1 \<in> C\<close> \<open>x < z1\<close>
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   785
      by fastforce
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   786
    then have B': "f'' z3 \<ge> 0"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   787
      using assms by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   788
    from A' B' have "(y - z1) * f'' z3 \<ge> 0"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   789
      by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   790
    from cool' this have "f' y - (f x - f y) / (x - y) \<ge> 0"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   791
      by auto
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   792
    from mult_right_mono_neg[OF this le(2)]
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   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
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   795
    then have "f' y * (x - y) - (f x - f y) \<le> 0"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   796
      using le by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   797
    then have res: "f' y * (x - y) \<le> f x - f y"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   798
      by auto
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   799
    have "(f y - f x) / (y - x) - f' x = f' z1 - f' x"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   800
      using * z1 by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   801
    also have "\<dots> = (z1 - x) * f'' z2"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   802
      using z2 by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   803
    finally have cool: "(f y - f x) / (y - x) - f' x = (z1 - x) * f'' z2"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   804
      by simp
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   805
    have A: "z1 - x \<ge> 0"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   806
      using z1 by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   807
    have "z2 \<in> C"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   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>
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   809
      by fastforce
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   810
    then have B: "f'' z2 \<ge> 0"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   811
      using assms by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   812
    from A B have "(z1 - x) * f'' z2 \<ge> 0"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   813
      by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   814
    with cool have "(f y - f x) / (y - x) - f' x \<ge> 0"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   815
      by auto
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   816
    from mult_right_mono[OF this ge(2)]
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   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
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   819
    then have "f y - f x - f' x * (y - x) \<ge> 0"
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   820
      using ge by auto
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   821
    then have "f y - f x \<ge> f' x * (y - x)" "f' y * (x - y) \<le> f x - f y"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   822
      using res by auto
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   823
  } note less_imp = this
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   824
  {
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   825
    fix x y :: real
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   826
    assume "x \<in> C" "y \<in> C" "x \<noteq> y"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   827
    then have"f y - f x \<ge> f' x * (y - x)"
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   828
    unfolding neq_iff using less_imp by auto
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   829
  }
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   830
  moreover
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   831
  {
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   832
    fix x y :: real
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   833
    assume "x \<in> C" "y \<in> C" "x = y"
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   834
    then have "f y - f x \<ge> f' x * (y - x)" by auto
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   835
  }
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   836
  ultimately show ?thesis using assms by blast
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   837
qed
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   838
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   839
lemma f''_ge0_imp_convex:
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   840
  fixes f :: "real \<Rightarrow> real"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   841
  assumes conv: "convex C"
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   842
    and f': "\<And>x. x \<in> C \<Longrightarrow> DERIV f x :> (f' x)"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   843
    and f'': "\<And>x. x \<in> C \<Longrightarrow> DERIV f' x :> (f'' x)"
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   844
    and pos: "\<And>x. x \<in> C \<Longrightarrow> f'' x \<ge> 0"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   845
  shows "convex_on C f"
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   846
  using f''_imp_f'[OF conv f' f'' pos] assms pos_convex_function
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   847
  by fastforce
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   848
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   849
lemma minus_log_convex:
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   850
  fixes b :: real
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   851
  assumes "b > 1"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   852
  shows "convex_on {0 <..} (\<lambda> x. - log b x)"
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   853
proof -
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   854
  have "\<And>z. z > 0 \<Longrightarrow> DERIV (log b) z :> 1 / (ln b * z)"
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   855
    using DERIV_log by auto
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   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
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   858
  have "\<And>z :: real. z > 0 \<Longrightarrow> DERIV inverse z :> - (inverse z ^ Suc (Suc 0))"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   859
    using less_imp_neq[THEN not_sym, THEN DERIV_inverse] by auto
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   860
  from this[THEN DERIV_cmult, of _ "- 1 / ln b"]
49609
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   861
  have "\<And>z :: real. z > 0 \<Longrightarrow>
89e10ed7668b tuned proofs;
wenzelm
parents: 44890
diff changeset
   862
    DERIV (\<lambda> z. (- 1 / ln b) * inverse z) z :> (- 1 / ln b) * (- (inverse z ^ Suc (Suc 0)))"
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   863
    by auto
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   864
  then have f''0: "\<And>z::real. z > 0 \<Longrightarrow>
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   865
    DERIV (\<lambda> z. - 1 / (ln b * z)) z :> 1 / (ln b * z * z)"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   866
    unfolding inverse_eq_divide by (auto simp: mult.assoc)
56796
9f84219715a7 tuned proofs;
wenzelm
parents: 56571
diff changeset
   867
  have f''_ge0: "\<And>z::real. z > 0 \<Longrightarrow> 1 / (ln b * z * z) \<ge> 0"
60423
5035a2af185b misc tuning;
wenzelm
parents: 60303
diff changeset
   868
    using \<open>b > 1\<close> by (auto intro!: less_imp_le)
36623
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   869
  from f''_ge0_imp_convex[OF pos_is_convex,
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   870
    unfolded greaterThan_iff, OF f' f''0 f''_ge0]
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   871
  show ?thesis by auto
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   872
qed
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   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
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63007
diff changeset
   885
  define z where "z = (1 - t) * x + t * y"
61585
a9599d3d7610 isabelle update_cartouches -c -t;
wenzelm
parents: 61531
diff changeset
   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
d26348b667f2 Moved Convex theory to library.
hoelzl
parents:
diff changeset
   961
end