author | nipkow |
Tue, 17 Jun 2025 06:29:55 +0200 | |
changeset 82732 | 71574900b6ba |
parent 82488 | b52e57ed7e29 |
permissions | -rw-r--r-- |
70086
72c52a897de2
First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
69802
diff
changeset
|
1 |
(* Title: HOL/Analysis/Convex.thy |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2 |
Author: L C Paulson, University of Cambridge |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
3 |
Author: Robert Himmelmann, TU Muenchen |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
4 |
Author: Bogdan Grechuk, University of Edinburgh |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
5 |
Author: Armin Heller, TU Muenchen |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
6 |
Author: Johannes Hoelzl, TU Muenchen |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
7 |
*) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
8 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
9 |
section \<open>Convex Sets and Functions\<close> |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
10 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
11 |
theory Convex |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
12 |
imports |
79583
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
13 |
Affine "HOL-Library.Set_Algebras" "HOL-Library.FuncSet" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
14 |
begin |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
15 |
|
71044 | 16 |
subsection \<open>Convex Sets\<close> |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
17 |
|
70136 | 18 |
definition\<^marker>\<open>tag important\<close> convex :: "'a::real_vector set \<Rightarrow> bool" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
19 |
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)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
20 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
21 |
lemma convexI: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
22 |
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" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
23 |
shows "convex s" |
78475 | 24 |
by (simp add: assms convex_def) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
25 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
26 |
lemma convexD: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
27 |
assumes "convex s" and "x \<in> s" and "y \<in> s" and "0 \<le> u" and "0 \<le> v" and "u + v = 1" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
28 |
shows "u *\<^sub>R x + v *\<^sub>R y \<in> s" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
29 |
using assms unfolding convex_def by fast |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
30 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
31 |
lemma convex_alt: "convex s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u. 0 \<le> u \<and> u \<le> 1 \<longrightarrow> ((1 - u) *\<^sub>R x + u *\<^sub>R y) \<in> s)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
32 |
(is "_ \<longleftrightarrow> ?alt") |
82485
12fe1e2b87e4
Generalised a lemma and added another
paulson <lp15@cam.ac.uk>
parents:
80654
diff
changeset
|
33 |
by (metis convex_def diff_eq_eq diff_ge_0_iff_ge) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
34 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
35 |
lemma convexD_alt: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
36 |
assumes "convex s" "a \<in> s" "b \<in> s" "0 \<le> u" "u \<le> 1" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
37 |
shows "((1 - u) *\<^sub>R a + u *\<^sub>R b) \<in> s" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
38 |
using assms unfolding convex_alt by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
39 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
40 |
lemma mem_convex_alt: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
41 |
assumes "convex S" "x \<in> S" "y \<in> S" "u \<ge> 0" "v \<ge> 0" "u + v > 0" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
42 |
shows "((u/(u+v)) *\<^sub>R x + (v/(u+v)) *\<^sub>R y) \<in> S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
43 |
using assms |
72385 | 44 |
by (simp add: convex_def zero_le_divide_iff add_divide_distrib [symmetric]) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
45 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
46 |
lemma convex_empty[intro,simp]: "convex {}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
47 |
unfolding convex_def by simp |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
48 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
49 |
lemma convex_singleton[intro,simp]: "convex {a}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
50 |
unfolding convex_def by (auto simp: scaleR_left_distrib[symmetric]) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
51 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
52 |
lemma convex_UNIV[intro,simp]: "convex UNIV" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
53 |
unfolding convex_def by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
54 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
55 |
lemma convex_Inter: "(\<And>s. s\<in>f \<Longrightarrow> convex s) \<Longrightarrow> convex(\<Inter>f)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
56 |
unfolding convex_def by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
57 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
58 |
lemma convex_Int: "convex s \<Longrightarrow> convex t \<Longrightarrow> convex (s \<inter> t)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
59 |
unfolding convex_def by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
60 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
61 |
lemma convex_INT: "(\<And>i. i \<in> A \<Longrightarrow> convex (B i)) \<Longrightarrow> convex (\<Inter>i\<in>A. B i)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
62 |
unfolding convex_def by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
63 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
64 |
lemma convex_Times: "convex s \<Longrightarrow> convex t \<Longrightarrow> convex (s \<times> t)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
65 |
unfolding convex_def by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
66 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
67 |
lemma convex_halfspace_le: "convex {x. inner a x \<le> b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
68 |
unfolding convex_def |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
69 |
by (auto simp: inner_add intro!: convex_bound_le) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
70 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
71 |
lemma convex_halfspace_ge: "convex {x. inner a x \<ge> b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
72 |
proof - |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
73 |
have *: "{x. inner a x \<ge> b} = {x. inner (-a) x \<le> -b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
74 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
75 |
show ?thesis |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
76 |
unfolding * using convex_halfspace_le[of "-a" "-b"] by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
77 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
78 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
79 |
lemma convex_halfspace_abs_le: "convex {x. \<bar>inner a x\<bar> \<le> b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
80 |
proof - |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
81 |
have *: "{x. \<bar>inner a x\<bar> \<le> b} = {x. inner a x \<le> b} \<inter> {x. -b \<le> inner a x}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
82 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
83 |
show ?thesis |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
84 |
unfolding * by (simp add: convex_Int convex_halfspace_ge convex_halfspace_le) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
85 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
86 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
87 |
lemma convex_hyperplane: "convex {x. inner a x = b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
88 |
proof - |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
89 |
have *: "{x. inner a x = b} = {x. inner a x \<le> b} \<inter> {x. inner a x \<ge> b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
90 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
91 |
show ?thesis using convex_halfspace_le convex_halfspace_ge |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
92 |
by (auto intro!: convex_Int simp: *) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
93 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
94 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
95 |
lemma convex_halfspace_lt: "convex {x. inner a x < b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
96 |
unfolding convex_def |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
97 |
by (auto simp: convex_bound_lt inner_add) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
98 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
99 |
lemma convex_halfspace_gt: "convex {x. inner a x > b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
100 |
using convex_halfspace_lt[of "-a" "-b"] by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
101 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
102 |
lemma convex_halfspace_Re_ge: "convex {x. Re x \<ge> b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
103 |
using convex_halfspace_ge[of b "1::complex"] by simp |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
104 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
105 |
lemma convex_halfspace_Re_le: "convex {x. Re x \<le> b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
106 |
using convex_halfspace_le[of "1::complex" b] by simp |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
107 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
108 |
lemma convex_halfspace_Im_ge: "convex {x. Im x \<ge> b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
109 |
using convex_halfspace_ge[of b \<i>] by simp |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
110 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
111 |
lemma convex_halfspace_Im_le: "convex {x. Im x \<le> b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
112 |
using convex_halfspace_le[of \<i> b] by simp |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
113 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
114 |
lemma convex_halfspace_Re_gt: "convex {x. Re x > b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
115 |
using convex_halfspace_gt[of b "1::complex"] by simp |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
116 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
117 |
lemma convex_halfspace_Re_lt: "convex {x. Re x < b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
118 |
using convex_halfspace_lt[of "1::complex" b] by simp |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
119 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
120 |
lemma convex_halfspace_Im_gt: "convex {x. Im x > b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
121 |
using convex_halfspace_gt[of b \<i>] by simp |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
122 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
123 |
lemma convex_halfspace_Im_lt: "convex {x. Im x < b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
124 |
using convex_halfspace_lt[of \<i> b] by simp |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
125 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
126 |
lemma convex_real_interval [iff]: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
127 |
fixes a b :: "real" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
128 |
shows "convex {a..}" and "convex {..b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
129 |
and "convex {a<..}" and "convex {..<b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
130 |
and "convex {a..b}" and "convex {a<..b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
131 |
and "convex {a..<b}" and "convex {a<..<b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
132 |
proof - |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
133 |
have "{a..} = {x. a \<le> inner 1 x}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
134 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
135 |
then show 1: "convex {a..}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
136 |
by (simp only: convex_halfspace_ge) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
137 |
have "{..b} = {x. inner 1 x \<le> b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
138 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
139 |
then show 2: "convex {..b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
140 |
by (simp only: convex_halfspace_le) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
141 |
have "{a<..} = {x. a < inner 1 x}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
142 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
143 |
then show 3: "convex {a<..}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
144 |
by (simp only: convex_halfspace_gt) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
145 |
have "{..<b} = {x. inner 1 x < b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
146 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
147 |
then show 4: "convex {..<b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
148 |
by (simp only: convex_halfspace_lt) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
149 |
have "{a..b} = {a..} \<inter> {..b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
150 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
151 |
then show "convex {a..b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
152 |
by (simp only: convex_Int 1 2) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
153 |
have "{a<..b} = {a<..} \<inter> {..b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
154 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
155 |
then show "convex {a<..b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
156 |
by (simp only: convex_Int 3 2) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
157 |
have "{a..<b} = {a..} \<inter> {..<b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
158 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
159 |
then show "convex {a..<b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
160 |
by (simp only: convex_Int 1 4) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
161 |
have "{a<..<b} = {a<..} \<inter> {..<b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
162 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
163 |
then show "convex {a<..<b}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
164 |
by (simp only: convex_Int 3 4) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
165 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
166 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
167 |
lemma convex_Reals: "convex \<real>" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
168 |
by (simp add: convex_def scaleR_conv_of_real) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
169 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
170 |
|
70136 | 171 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>Explicit expressions for convexity in terms of arbitrary sums\<close> |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
172 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
173 |
lemma convex_sum: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
174 |
fixes C :: "'a::real_vector set" |
72385 | 175 |
assumes "finite S" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
176 |
and "convex C" |
78475 | 177 |
and a: "(\<Sum> i \<in> S. a i) = 1" "\<And>i. i \<in> S \<Longrightarrow> a i \<ge> 0" |
178 |
and C: "\<And>i. i \<in> S \<Longrightarrow> y i \<in> C" |
|
72385 | 179 |
shows "(\<Sum> j \<in> S. a j *\<^sub>R y j) \<in> C" |
78475 | 180 |
using \<open>finite S\<close> a C |
181 |
proof (induction arbitrary: a set: finite) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
182 |
case empty |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
183 |
then show ?case by simp |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
184 |
next |
78475 | 185 |
case (insert i S) |
72385 | 186 |
then have "0 \<le> sum a S" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
187 |
by (simp add: sum_nonneg) |
72385 | 188 |
have "a i *\<^sub>R y i + (\<Sum>j\<in>S. a j *\<^sub>R y j) \<in> C" |
189 |
proof (cases "sum a S = 0") |
|
78475 | 190 |
case True with insert show ?thesis |
191 |
by (simp add: sum_nonneg_eq_0_iff) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
192 |
next |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
193 |
case False |
72385 | 194 |
with \<open>0 \<le> sum a S\<close> have "0 < sum a S" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
195 |
by simp |
72385 | 196 |
then have "(\<Sum>j\<in>S. (a j / sum a S) *\<^sub>R y j) \<in> C" |
78475 | 197 |
using insert |
198 |
by (simp add: insert.IH flip: sum_divide_distrib) |
|
199 |
with \<open>convex C\<close> insert \<open>0 \<le> sum a S\<close> |
|
72385 | 200 |
have "a i *\<^sub>R y i + sum a S *\<^sub>R (\<Sum>j\<in>S. (a j / sum a S) *\<^sub>R y j) \<in> C" |
78475 | 201 |
by (simp add: convex_def) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
202 |
then show ?thesis |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
203 |
by (simp add: scaleR_sum_right False) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
204 |
qed |
72385 | 205 |
then show ?case using \<open>finite S\<close> and \<open>i \<notin> S\<close> |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
206 |
by simp |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
207 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
208 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
209 |
lemma convex: |
78475 | 210 |
"convex S \<longleftrightarrow> |
211 |
(\<forall>(k::nat) u x. (\<forall>i. 1\<le>i \<and> i\<le>k \<longrightarrow> 0 \<le> u i \<and> x i \<in>S) \<and> (sum u {1..k} = 1) |
|
212 |
\<longrightarrow> sum (\<lambda>i. u i *\<^sub>R x i) {1..k} \<in> S)" |
|
213 |
(is "?lhs = ?rhs") |
|
214 |
proof |
|
215 |
show "?lhs \<Longrightarrow> ?rhs" |
|
216 |
by (metis (full_types) atLeastAtMost_iff convex_sum finite_atLeastAtMost) |
|
72385 | 217 |
assume *: "\<forall>k u x. (\<forall> i :: nat. 1 \<le> i \<and> i \<le> k \<longrightarrow> 0 \<le> u i \<and> x i \<in> S) \<and> sum u {1..k} = 1 |
218 |
\<longrightarrow> (\<Sum>i = 1..k. u i *\<^sub>R (x i :: 'a)) \<in> S" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
219 |
{ |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
220 |
fix \<mu> :: real |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
221 |
fix x y :: 'a |
72385 | 222 |
assume xy: "x \<in> S" "y \<in> S" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
223 |
assume mu: "\<mu> \<ge> 0" "\<mu> \<le> 1" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
224 |
let ?u = "\<lambda>i. if (i :: nat) = 1 then \<mu> else 1 - \<mu>" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
225 |
let ?x = "\<lambda>i. if (i :: nat) = 1 then x else y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
226 |
have "{1 :: nat .. 2} \<inter> - {x. x = 1} = {2}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
227 |
by auto |
78475 | 228 |
then have S: "(\<Sum>j \<in> {1..2}. ?u j *\<^sub>R ?x j) \<in> S" |
229 |
using sum.If_cases[of "{(1 :: nat) .. 2}" "\<lambda>x. x = 1" "\<lambda>x. \<mu>" "\<lambda>x. 1 - \<mu>"] |
|
230 |
using mu xy "*" by auto |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
231 |
have grarr: "(\<Sum>j \<in> {Suc (Suc 0)..2}. ?u j *\<^sub>R ?x j) = (1 - \<mu>) *\<^sub>R y" |
70097
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
70086
diff
changeset
|
232 |
using sum.atLeast_Suc_atMost[of "Suc (Suc 0)" 2 "\<lambda> j. (1 - \<mu>) *\<^sub>R y"] by auto |
78475 | 233 |
with sum.atLeast_Suc_atMost |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
234 |
have "(\<Sum>j \<in> {1..2}. ?u j *\<^sub>R ?x j) = \<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y" |
78475 | 235 |
by (smt (verit, best) Suc_1 Suc_eq_plus1 add_0 le_add1) |
72385 | 236 |
then have "(1 - \<mu>) *\<^sub>R y + \<mu> *\<^sub>R x \<in> S" |
237 |
using S by (auto simp: add.commute) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
238 |
} |
72385 | 239 |
then show "convex S" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
240 |
unfolding convex_alt by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
241 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
242 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
243 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
244 |
lemma convex_explicit: |
72385 | 245 |
fixes S :: "'a::real_vector set" |
246 |
shows "convex S \<longleftrightarrow> |
|
247 |
(\<forall>t u. finite t \<and> t \<subseteq> S \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and> sum u t = 1 \<longrightarrow> sum (\<lambda>x. u x *\<^sub>R x) t \<in> S)" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
248 |
proof safe |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
249 |
fix t |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
250 |
fix u :: "'a \<Rightarrow> real" |
72385 | 251 |
assume "convex S" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
252 |
and "finite t" |
72385 | 253 |
and "t \<subseteq> S" "\<forall>x\<in>t. 0 \<le> u x" "sum u t = 1" |
254 |
then show "(\<Sum>x\<in>t. u x *\<^sub>R x) \<in> S" |
|
78475 | 255 |
by (simp add: convex_sum subsetD) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
256 |
next |
72385 | 257 |
assume *: "\<forall>t. \<forall> u. finite t \<and> t \<subseteq> S \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and> |
258 |
sum u t = 1 \<longrightarrow> (\<Sum>x\<in>t. u x *\<^sub>R x) \<in> S" |
|
259 |
show "convex S" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
260 |
unfolding convex_alt |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
261 |
proof safe |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
262 |
fix x y |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
263 |
fix \<mu> :: real |
72385 | 264 |
assume **: "x \<in> S" "y \<in> S" "0 \<le> \<mu>" "\<mu> \<le> 1" |
265 |
show "(1 - \<mu>) *\<^sub>R x + \<mu> *\<^sub>R y \<in> S" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
266 |
proof (cases "x = y") |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
267 |
case False |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
268 |
then show ?thesis |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
269 |
using *[rule_format, of "{x, y}" "\<lambda> z. if z = x then 1 - \<mu> else \<mu>"] ** |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
270 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
271 |
next |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
272 |
case True |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
273 |
then show ?thesis |
82488 | 274 |
by (simp add: "**") |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
275 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
276 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
277 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
278 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
279 |
lemma convex_finite: |
72385 | 280 |
assumes "finite S" |
281 |
shows "convex S \<longleftrightarrow> (\<forall>u. (\<forall>x\<in>S. 0 \<le> u x) \<and> sum u S = 1 \<longrightarrow> sum (\<lambda>x. u x *\<^sub>R x) S \<in> S)" |
|
282 |
(is "?lhs = ?rhs") |
|
283 |
proof |
|
284 |
{ have if_distrib_arg: "\<And>P f g x. (if P then f else g) x = (if P then f x else g x)" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
285 |
by simp |
72385 | 286 |
fix T :: "'a set" and u :: "'a \<Rightarrow> real" |
287 |
assume sum: "\<forall>u. (\<forall>x\<in>S. 0 \<le> u x) \<and> sum u S = 1 \<longrightarrow> (\<Sum>x\<in>S. u x *\<^sub>R x) \<in> S" |
|
288 |
assume *: "\<forall>x\<in>T. 0 \<le> u x" "sum u T = 1" |
|
289 |
assume "T \<subseteq> S" |
|
290 |
then have "S \<inter> T = T" by auto |
|
78475 | 291 |
with sum[THEN spec[where x="\<lambda>x. if x\<in>T then u x else 0"]] * |
292 |
have "(\<Sum>x\<in>T. u x *\<^sub>R x) \<in> S" |
|
72385 | 293 |
by (auto simp: assms sum.If_cases if_distrib if_distrib_arg) } |
294 |
moreover assume ?rhs |
|
295 |
ultimately show ?lhs |
|
296 |
unfolding convex_explicit by auto |
|
297 |
qed (auto simp: convex_explicit assms) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
298 |
|
71044 | 299 |
subsection \<open>Convex Functions on a Set\<close> |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
300 |
|
70136 | 301 |
definition\<^marker>\<open>tag important\<close> convex_on :: "'a::real_vector set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> bool" |
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
302 |
where "convex_on S f \<longleftrightarrow> convex S \<and> |
72385 | 303 |
(\<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)" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
304 |
|
74729
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
305 |
definition\<^marker>\<open>tag important\<close> concave_on :: "'a::real_vector set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> bool" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
306 |
where "concave_on S f \<equiv> convex_on S (\<lambda>x. - f x)" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
307 |
|
79583
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
308 |
lemma convex_on_iff_concave: "convex_on S f = concave_on S (\<lambda>x. - f x)" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
309 |
by (simp add: concave_on_def) |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
310 |
|
74729
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
311 |
lemma concave_on_iff: |
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
312 |
"concave_on S f \<longleftrightarrow> convex S \<and> |
74729
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
313 |
(\<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) \<ge> u * f x + v * f y)" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
314 |
by (auto simp: concave_on_def convex_on_def algebra_simps) |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
315 |
|
79945
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
316 |
lemma concave_onD: |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
317 |
assumes "concave_on A f" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
318 |
shows "\<And>t x y. t \<ge> 0 \<Longrightarrow> t \<le> 1 \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
319 |
f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \<ge> (1 - t) * f x + t * f y" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
320 |
using assms by (auto simp: concave_on_iff) |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
321 |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
322 |
lemma convex_onI [intro?]: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
323 |
assumes "\<And>t x y. t > 0 \<Longrightarrow> t < 1 \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
324 |
f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \<le> (1 - t) * f x + t * f y" |
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
325 |
and "convex A" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
326 |
shows "convex_on A f" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
327 |
unfolding convex_on_def |
78475 | 328 |
by (smt (verit, del_insts) assms mult_cancel_right1 mult_eq_0_iff scaleR_collapse scaleR_eq_0_iff) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
329 |
|
79945
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
330 |
lemma convex_onD: |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
331 |
assumes "convex_on A f" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
332 |
shows "\<And>t x y. t \<ge> 0 \<Longrightarrow> t \<le> 1 \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
333 |
f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \<le> (1 - t) * f x + t * f y" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
334 |
using assms by (auto simp: convex_on_def) |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
335 |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
336 |
lemma convex_on_linorderI [intro?]: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
337 |
fixes A :: "('a::{linorder,real_vector}) set" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
338 |
assumes "\<And>t x y. t > 0 \<Longrightarrow> t < 1 \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x < y \<Longrightarrow> |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
339 |
f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \<le> (1 - t) * f x + t * f y" |
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
340 |
and "convex A" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
341 |
shows "convex_on A f" |
78475 | 342 |
by (smt (verit, best) add.commute assms convex_onI distrib_left linorder_cases mult.commute mult_cancel_left2 scaleR_collapse) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
343 |
|
79945
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
344 |
lemma concave_on_linorderI [intro?]: |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
345 |
fixes A :: "('a::{linorder,real_vector}) set" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
346 |
assumes "\<And>t x y. t > 0 \<Longrightarrow> t < 1 \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x < y \<Longrightarrow> |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
347 |
f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \<ge> (1 - t) * f x + t * f y" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
348 |
and "convex A" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
349 |
shows "concave_on A f" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
350 |
by (smt (verit) assms concave_on_def convex_on_linorderI mult_minus_right) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
351 |
|
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
352 |
lemma convex_on_imp_convex: "convex_on A f \<Longrightarrow> convex A" |
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
353 |
by (auto simp: convex_on_def) |
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
354 |
|
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
355 |
lemma concave_on_imp_convex: "concave_on A f \<Longrightarrow> convex A" |
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
356 |
by (simp add: concave_on_def convex_on_imp_convex) |
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
357 |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
358 |
lemma convex_onD_Icc: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
359 |
assumes "convex_on {x..y} f" "x \<le> (y :: _ :: {real_vector,preorder})" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
360 |
shows "\<And>t. t \<ge> 0 \<Longrightarrow> t \<le> 1 \<Longrightarrow> |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
361 |
f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \<le> (1 - t) * f x + t * f y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
362 |
using assms(2) by (intro convex_onD [OF assms(1)]) simp_all |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
363 |
|
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
364 |
lemma convex_on_subset: "\<lbrakk>convex_on T f; S \<subseteq> T; convex S\<rbrakk> \<Longrightarrow> convex_on S f" |
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
365 |
by (simp add: convex_on_def subset_iff) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
366 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
367 |
lemma convex_on_add [intro]: |
72385 | 368 |
assumes "convex_on S f" |
369 |
and "convex_on S g" |
|
370 |
shows "convex_on S (\<lambda>x. f x + g x)" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
371 |
proof - |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
372 |
{ |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
373 |
fix x y |
72385 | 374 |
assume "x \<in> S" "y \<in> S" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
375 |
moreover |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
376 |
fix u v :: real |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
377 |
assume "0 \<le> u" "0 \<le> v" "u + v = 1" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
378 |
ultimately |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
379 |
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)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
380 |
using assms unfolding convex_on_def by (auto simp: add_mono) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
381 |
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)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
382 |
by (simp add: field_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
383 |
} |
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
384 |
with assms show ?thesis |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
385 |
unfolding convex_on_def by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
386 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
387 |
|
79583
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
388 |
lemma convex_on_ident: "convex_on S (\<lambda>x. x) \<longleftrightarrow> convex S" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
389 |
by (simp add: convex_on_def) |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
390 |
|
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
391 |
lemma concave_on_ident: "concave_on S (\<lambda>x. x) \<longleftrightarrow> convex S" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
392 |
by (simp add: concave_on_iff) |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
393 |
|
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
394 |
lemma convex_on_const: "convex_on S (\<lambda>x. a) \<longleftrightarrow> convex S" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
395 |
by (simp add: convex_on_def flip: distrib_right) |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
396 |
|
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
397 |
lemma concave_on_const: "concave_on S (\<lambda>x. a) \<longleftrightarrow> convex S" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
398 |
by (simp add: concave_on_iff flip: distrib_right) |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
399 |
|
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
400 |
lemma convex_on_diff: |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
401 |
assumes "convex_on S f" and "concave_on S g" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
402 |
shows "convex_on S (\<lambda>x. f x - g x)" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
403 |
using assms concave_on_def convex_on_add by fastforce |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
404 |
|
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
405 |
lemma concave_on_diff: |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
406 |
assumes "concave_on S f" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
407 |
and "convex_on S g" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
408 |
shows "concave_on S (\<lambda>x. f x - g x)" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
409 |
using convex_on_diff assms concave_on_def by fastforce |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
410 |
|
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
411 |
lemma concave_on_add: |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
412 |
assumes "concave_on S f" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
413 |
and "concave_on S g" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
414 |
shows "concave_on S (\<lambda>x. f x + g x)" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
415 |
using assms convex_on_iff_concave concave_on_diff concave_on_def by fastforce |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
416 |
|
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
417 |
lemma convex_on_mul: |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
418 |
fixes S::"real set" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
419 |
assumes "convex_on S f" "convex_on S g" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
420 |
assumes "mono_on S f" "mono_on S g" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
421 |
assumes fty: "f \<in> S \<rightarrow> {0..}" and gty: "g \<in> S \<rightarrow> {0..}" |
79945
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
422 |
shows "convex_on S (\<lambda>x. f x*g x)" |
79583
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
423 |
proof (intro convex_on_linorderI) |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
424 |
show "convex S" |
79945
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
425 |
using assms convex_on_imp_convex by auto |
79583
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
426 |
fix t::real and x y |
79945
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
427 |
assume t: "0 < t" "t < 1" and xy: "x \<in> S" "y \<in> S" "x<y" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
428 |
have *: "t*(1-t) * f x * g y + t*(1-t) * f y * g x \<le> t*(1-t) * f x * g x + t*(1-t) * f y * g y" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
429 |
using t \<open>mono_on S f\<close> \<open>mono_on S g\<close> xy |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
430 |
by (smt (verit, ccfv_SIG) left_diff_distrib mono_onD mult_left_less_imp_less zero_le_mult_iff) |
79583
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
431 |
have inS: "(1-t)*x + t*y \<in> S" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
432 |
using t xy \<open>convex S\<close> by (simp add: convex_alt) |
79945
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
433 |
then have "f ((1-t)*x + t*y) * g ((1-t)*x + t*y) \<le> ((1-t) * f x + t * f y)*g ((1-t)*x + t*y)" |
79583
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
434 |
using convex_onD [OF \<open>convex_on S f\<close>, of t x y] t xy fty gty |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
435 |
by (intro mult_mono add_nonneg_nonneg) (auto simp: Pi_iff zero_le_mult_iff) |
79945
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
436 |
also have "\<dots> \<le> ((1-t) * f x + t * f y) * ((1-t)*g x + t*g y)" |
79583
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
437 |
using convex_onD [OF \<open>convex_on S g\<close>, of t x y] t xy fty gty inS |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
438 |
by (intro mult_mono add_nonneg_nonneg) (auto simp: Pi_iff zero_le_mult_iff) |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
439 |
also have "\<dots> \<le> (1-t) * (f x*g x) + t * (f y*g y)" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
440 |
using * by (simp add: algebra_simps) |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
441 |
finally show "f ((1-t) *\<^sub>R x + t *\<^sub>R y) * g ((1-t) *\<^sub>R x + t *\<^sub>R y) \<le> (1-t)*(f x*g x) + t*(f y*g y)" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
442 |
by simp |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
443 |
qed |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
444 |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
445 |
lemma convex_on_cmul [intro]: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
446 |
fixes c :: real |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
447 |
assumes "0 \<le> c" |
72385 | 448 |
and "convex_on S f" |
449 |
shows "convex_on S (\<lambda>x. c * f x)" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
450 |
proof - |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
451 |
have *: "u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
452 |
for u c fx v fy :: real |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
453 |
by (simp add: field_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
454 |
show ?thesis using assms(2) and mult_left_mono [OF _ assms(1)] |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
455 |
unfolding convex_on_def and * by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
456 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
457 |
|
79583
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
458 |
lemma convex_on_cdiv [intro]: |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
459 |
fixes c :: real |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
460 |
assumes "0 \<le> c" and "convex_on S f" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
461 |
shows "convex_on S (\<lambda>x. f x / c)" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
462 |
unfolding divide_inverse |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
463 |
using convex_on_cmul [of "inverse c" S f] by (simp add: mult.commute assms) |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
464 |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
465 |
lemma convex_lower: |
72385 | 466 |
assumes "convex_on S f" |
467 |
and "x \<in> S" |
|
468 |
and "y \<in> S" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
469 |
and "0 \<le> u" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
470 |
and "0 \<le> v" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
471 |
and "u + v = 1" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
472 |
shows "f (u *\<^sub>R x + v *\<^sub>R y) \<le> max (f x) (f y)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
473 |
proof - |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
474 |
let ?m = "max (f x) (f y)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
475 |
have "u * f x + v * f y \<le> u * max (f x) (f y) + v * max (f x) (f y)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
476 |
using assms(4,5) by (auto simp: mult_left_mono add_mono) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
477 |
also have "\<dots> = max (f x) (f y)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
478 |
using assms(6) by (simp add: distrib_right [symmetric]) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
479 |
finally show ?thesis |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
480 |
using assms unfolding convex_on_def by fastforce |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
481 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
482 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
483 |
lemma convex_on_dist [intro]: |
72385 | 484 |
fixes S :: "'a::real_normed_vector set" |
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
485 |
assumes "convex S" |
72385 | 486 |
shows "convex_on S (\<lambda>x. dist a x)" |
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
487 |
unfolding convex_on_def dist_norm |
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
488 |
proof (intro conjI strip) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
489 |
fix x y |
72385 | 490 |
assume "x \<in> S" "y \<in> S" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
491 |
fix u v :: real |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
492 |
assume "0 \<le> u" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
493 |
assume "0 \<le> v" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
494 |
assume "u + v = 1" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
495 |
have "a = u *\<^sub>R a + v *\<^sub>R a" |
78475 | 496 |
by (metis \<open>u + v = 1\<close> scaleR_left.add scaleR_one) |
497 |
then have "a - (u *\<^sub>R x + v *\<^sub>R y) = (u *\<^sub>R (a - x)) + (v *\<^sub>R (a - y))" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
498 |
by (auto simp: algebra_simps) |
78475 | 499 |
then show "norm (a - (u *\<^sub>R x + v *\<^sub>R y)) \<le> u * norm (a - x) + v * norm (a - y)" |
500 |
by (smt (verit, best) \<open>0 \<le> u\<close> \<open>0 \<le> v\<close> norm_scaleR norm_triangle_ineq) |
|
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
501 |
qed (use assms in auto) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
502 |
|
79945
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
503 |
lemma concave_on_mul: |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
504 |
fixes S::"real set" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
505 |
assumes f: "concave_on S f" and g: "concave_on S g" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
506 |
assumes "mono_on S f" "antimono_on S g" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
507 |
assumes fty: "f \<in> S \<rightarrow> {0..}" and gty: "g \<in> S \<rightarrow> {0..}" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
508 |
shows "concave_on S (\<lambda>x. f x * g x)" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
509 |
proof (intro concave_on_linorderI) |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
510 |
show "convex S" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
511 |
using concave_on_imp_convex f by blast |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
512 |
fix t::real and x y |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
513 |
assume t: "0 < t" "t < 1" and xy: "x \<in> S" "y \<in> S" "x<y" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
514 |
have inS: "(1-t)*x + t*y \<in> S" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
515 |
using t xy \<open>convex S\<close> by (simp add: convex_alt) |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
516 |
have "f x * g y + f y * g x \<ge> f x * g x + f y * g y" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
517 |
using \<open>mono_on S f\<close> \<open>antimono_on S g\<close> |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
518 |
unfolding monotone_on_def by (smt (verit, best) left_diff_distrib mult_left_mono xy) |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
519 |
with t have *: "t*(1-t) * f x * g y + t*(1-t) * f y * g x \<ge> t*(1-t) * f x * g x + t*(1-t) * f y * g y" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
520 |
by (smt (verit, ccfv_SIG) distrib_left mult_left_mono diff_ge_0_iff_ge mult.assoc) |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
521 |
have "(1 - t) * (f x * g x) + t * (f y * g y) \<le> ((1-t) * f x + t * f y) * ((1-t) * g x + t * g y)" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
522 |
using * by (simp add: algebra_simps) |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
523 |
also have "\<dots> \<le> ((1-t) * f x + t * f y)*g ((1-t)*x + t*y)" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
524 |
using concave_onD [OF \<open>concave_on S g\<close>, of t x y] t xy fty gty inS |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
525 |
by (intro mult_mono add_nonneg_nonneg) (auto simp: Pi_iff zero_le_mult_iff) |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
526 |
also have "\<dots> \<le> f ((1-t)*x + t*y) * g ((1-t)*x + t*y)" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
527 |
using concave_onD [OF \<open>concave_on S f\<close>, of t x y] t xy fty gty inS |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
528 |
by (intro mult_mono add_nonneg_nonneg) (auto simp: Pi_iff zero_le_mult_iff) |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
529 |
finally show "(1 - t) * (f x * g x) + t * (f y * g y) |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
530 |
\<le> f ((1 - t) *\<^sub>R x + t *\<^sub>R y) * g ((1 - t) *\<^sub>R x + t *\<^sub>R y)" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
531 |
by simp |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
532 |
qed |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
533 |
|
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
534 |
lemma concave_on_cmul [intro]: |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
535 |
fixes c :: real |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
536 |
assumes "0 \<le> c" and "concave_on S f" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
537 |
shows "concave_on S (\<lambda>x. c * f x)" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
538 |
using assms convex_on_cmul [of c S "\<lambda>x. - f x"] |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
539 |
by (auto simp: concave_on_def) |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
540 |
|
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
541 |
lemma concave_on_cdiv [intro]: |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
542 |
fixes c :: real |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
543 |
assumes "0 \<le> c" and "concave_on S f" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
544 |
shows "concave_on S (\<lambda>x. f x / c)" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
545 |
unfolding divide_inverse |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
546 |
using concave_on_cmul [of "inverse c" S f] by (simp add: mult.commute assms) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
547 |
|
70136 | 548 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>Arithmetic operations on sets preserve convexity\<close> |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
549 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
550 |
lemma convex_linear_image: |
78475 | 551 |
assumes "linear f" and "convex S" |
72385 | 552 |
shows "convex (f ` S)" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
553 |
proof - |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
554 |
interpret f: linear f by fact |
72385 | 555 |
from \<open>convex S\<close> show "convex (f ` S)" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
556 |
by (simp add: convex_def f.scaleR [symmetric] f.add [symmetric]) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
557 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
558 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
559 |
lemma convex_linear_vimage: |
78475 | 560 |
assumes "linear f" and "convex S" |
72385 | 561 |
shows "convex (f -` S)" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
562 |
proof - |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
563 |
interpret f: linear f by fact |
72385 | 564 |
from \<open>convex S\<close> show "convex (f -` S)" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
565 |
by (simp add: convex_def f.add f.scaleR) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
566 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
567 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
568 |
lemma convex_scaling: |
72385 | 569 |
assumes "convex S" |
570 |
shows "convex ((\<lambda>x. c *\<^sub>R x) ` S)" |
|
78475 | 571 |
by (simp add: assms convex_linear_image) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
572 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
573 |
lemma convex_scaled: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
574 |
assumes "convex S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
575 |
shows "convex ((\<lambda>x. x *\<^sub>R c) ` S)" |
78475 | 576 |
by (simp add: assms convex_linear_image) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
577 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
578 |
lemma convex_negations: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
579 |
assumes "convex S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
580 |
shows "convex ((\<lambda>x. - x) ` S)" |
78475 | 581 |
by (simp add: assms convex_linear_image module_hom_uminus) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
582 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
583 |
lemma convex_sums: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
584 |
assumes "convex S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
585 |
and "convex T" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
586 |
shows "convex (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
587 |
proof - |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
588 |
have "linear (\<lambda>(x, y). x + y)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
589 |
by (auto intro: linearI simp: scaleR_add_right) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
590 |
with assms have "convex ((\<lambda>(x, y). x + y) ` (S \<times> T))" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
591 |
by (intro convex_linear_image convex_Times) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
592 |
also have "((\<lambda>(x, y). x + y) ` (S \<times> T)) = (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
593 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
594 |
finally show ?thesis . |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
595 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
596 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
597 |
lemma convex_differences: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
598 |
assumes "convex S" "convex T" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
599 |
shows "convex (\<Union>x\<in> S. \<Union>y \<in> T. {x - y})" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
600 |
proof - |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
601 |
have "{x - y| x y. x \<in> S \<and> y \<in> T} = {x + y |x y. x \<in> S \<and> y \<in> uminus ` T}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
602 |
by (auto simp: diff_conv_add_uminus simp del: add_uminus_conv_diff) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
603 |
then show ?thesis |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
604 |
using convex_sums[OF assms(1) convex_negations[OF assms(2)]] by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
605 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
606 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
607 |
lemma convex_translation: |
69661 | 608 |
"convex ((+) a ` S)" if "convex S" |
82488 | 609 |
using convex_sums [OF convex_singleton [of a] that] |
610 |
by (simp add: UNION_singleton_eq_range) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
611 |
|
69661 | 612 |
lemma convex_translation_subtract: |
613 |
"convex ((\<lambda>b. b - a) ` S)" if "convex S" |
|
614 |
using convex_translation [of S "- a"] that by (simp cong: image_cong_simp) |
|
615 |
||
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
616 |
lemma convex_affinity: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
617 |
assumes "convex S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
618 |
shows "convex ((\<lambda>x. a + c *\<^sub>R x) ` S)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
619 |
proof - |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
620 |
have "(\<lambda>x. a + c *\<^sub>R x) ` S = (+) a ` (*\<^sub>R) c ` S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
621 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
622 |
then show ?thesis |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
623 |
using convex_translation[OF convex_scaling[OF assms], of a c] by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
624 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
625 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
626 |
lemma convex_on_sum: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
627 |
fixes a :: "'a \<Rightarrow> real" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
628 |
and y :: "'a \<Rightarrow> 'b::real_vector" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
629 |
and f :: "'b \<Rightarrow> real" |
78475 | 630 |
assumes "finite S" "S \<noteq> {}" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
631 |
and "convex_on C f" |
78475 | 632 |
and "(\<Sum> i \<in> S. a i) = 1" |
633 |
and "\<And>i. i \<in> S \<Longrightarrow> a i \<ge> 0" |
|
634 |
and "\<And>i. i \<in> S \<Longrightarrow> y i \<in> C" |
|
635 |
shows "f (\<Sum> i \<in> S. a i *\<^sub>R y i) \<le> (\<Sum> i \<in> S. a i * f (y i))" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
636 |
using assms |
78475 | 637 |
proof (induct S arbitrary: a rule: finite_ne_induct) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
638 |
case (singleton i) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
639 |
then show ?case |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
640 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
641 |
next |
78475 | 642 |
case (insert i S) |
643 |
then have yai: "y i \<in> C" "a i \<ge> 0" |
|
644 |
by auto |
|
645 |
with insert have conv: "\<And>x y \<mu>. x \<in> C \<Longrightarrow> y \<in> C \<Longrightarrow> 0 \<le> \<mu> \<Longrightarrow> \<mu> \<le> 1 \<Longrightarrow> |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
646 |
f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y" |
78475 | 647 |
by (simp add: convex_on_def) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
648 |
show ?case |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
649 |
proof (cases "a i = 1") |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
650 |
case True |
78475 | 651 |
with insert have "(\<Sum> j \<in> S. a j) = 0" |
652 |
by auto |
|
653 |
with insert show ?thesis |
|
654 |
by (simp add: sum_nonneg_eq_0_iff) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
655 |
next |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
656 |
case False |
78475 | 657 |
then have ai1: "a i < 1" |
658 |
using sum_nonneg_leq_bound[of "insert i S" a] insert by force |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
659 |
then have i0: "1 - a i > 0" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
660 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
661 |
let ?a = "\<lambda>j. a j / (1 - a i)" |
78475 | 662 |
have a_nonneg: "?a j \<ge> 0" if "j \<in> S" for j |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
663 |
using i0 insert that by fastforce |
78475 | 664 |
have "(\<Sum> j \<in> insert i S. a j) = 1" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
665 |
using insert by auto |
78475 | 666 |
then have "(\<Sum> j \<in> S. a j) = 1 - a i" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
667 |
using sum.insert insert by fastforce |
78475 | 668 |
then have "(\<Sum> j \<in> S. a j) / (1 - a i) = 1" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
669 |
using i0 by auto |
78475 | 670 |
then have a1: "(\<Sum> j \<in> S. ?a j) = 1" |
671 |
unfolding sum_divide_distrib by simp |
|
79670
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
672 |
have "convex C" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
673 |
using \<open>convex_on C f\<close> by (simp add: convex_on_def) |
78475 | 674 |
have asum: "(\<Sum> j \<in> S. ?a j *\<^sub>R y j) \<in> C" |
675 |
using insert convex_sum [OF \<open>finite S\<close> \<open>convex C\<close> a1 a_nonneg] by auto |
|
676 |
have asum_le: "f (\<Sum> j \<in> S. ?a j *\<^sub>R y j) \<le> (\<Sum> j \<in> S. ?a j * f (y j))" |
|
677 |
using a_nonneg a1 insert by blast |
|
678 |
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)" |
|
679 |
by (simp add: add.commute insert.hyps) |
|
680 |
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)" |
|
681 |
using i0 by auto |
|
682 |
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)" |
|
683 |
using scaleR_right.sum[of "inverse (1 - a i)" "\<lambda> j. a j *\<^sub>R y j" S, symmetric] |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
684 |
by (auto simp: algebra_simps) |
78475 | 685 |
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)" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
686 |
by (auto simp: divide_inverse) |
78475 | 687 |
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)" |
688 |
using ai1 by (smt (verit) asum conv real_scaleR_def yai) |
|
689 |
also have "\<dots> \<le> (1 - a i) * (\<Sum> j \<in> S. ?a j * f (y j)) + a i * f (y i)" |
|
690 |
using asum_le i0 by fastforce |
|
691 |
also have "\<dots> = (\<Sum> j \<in> S. a j * f (y j)) + a i * f (y i)" |
|
692 |
using i0 by (auto simp: sum_distrib_left) |
|
693 |
finally show ?thesis |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
694 |
using insert by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
695 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
696 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
697 |
|
79670
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
698 |
lemma concave_on_sum: |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
699 |
fixes a :: "'a \<Rightarrow> real" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
700 |
and y :: "'a \<Rightarrow> 'b::real_vector" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
701 |
and f :: "'b \<Rightarrow> real" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
702 |
assumes "finite S" "S \<noteq> {}" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
703 |
and "concave_on C f" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
704 |
and "(\<Sum>i \<in> S. a i) = 1" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
705 |
and "\<And>i. i \<in> S \<Longrightarrow> a i \<ge> 0" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
706 |
and "\<And>i. i \<in> S \<Longrightarrow> y i \<in> C" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
707 |
shows "f (\<Sum>i \<in> S. a i *\<^sub>R y i) \<ge> (\<Sum>i \<in> S. a i * f (y i))" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
708 |
proof - |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
709 |
have "(uminus \<circ> f) (\<Sum>i\<in>S. a i *\<^sub>R y i) \<le> (\<Sum>i\<in>S. a i * (uminus \<circ> f) (y i))" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
710 |
proof (intro convex_on_sum) |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
711 |
show "convex_on C (uminus \<circ> f)" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
712 |
by (simp add: assms convex_on_iff_concave) |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
713 |
qed (use assms in auto) |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
714 |
then show ?thesis |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
715 |
by (simp add: sum_negf o_def) |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
716 |
qed |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
717 |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
718 |
lemma convex_on_alt: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
719 |
fixes C :: "'a::real_vector set" |
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
720 |
shows "convex_on C f \<longleftrightarrow> convex C \<and> |
78475 | 721 |
(\<forall>x \<in> C. \<forall>y \<in> C. \<forall> \<mu> :: real. \<mu> \<ge> 0 \<and> \<mu> \<le> 1 \<longrightarrow> |
722 |
f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y)" |
|
723 |
by (smt (verit) convex_on_def) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
724 |
|
79583
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
725 |
lemma convex_on_slope_le: |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
726 |
fixes f :: "real \<Rightarrow> real" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
727 |
assumes f: "convex_on I f" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
728 |
and I: "x \<in> I" "y \<in> I" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
729 |
and t: "x < t" "t < y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
730 |
shows "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
731 |
and "(f x - f y) / (x - y) \<le> (f t - f y) / (t - y)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
732 |
proof - |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
733 |
define a where "a \<equiv> (t - y) / (x - y)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
734 |
with t have "0 \<le> a" "0 \<le> 1 - a" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
735 |
by (auto simp: field_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
736 |
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" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
737 |
by (auto simp: convex_on_def) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
738 |
have "a * x + (1 - a) * y = a * (x - y) + y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
739 |
by (simp add: field_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
740 |
also have "\<dots> = t" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
741 |
unfolding a_def using \<open>x < t\<close> \<open>t < y\<close> by simp |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
742 |
finally have "f t \<le> a * f x + (1 - a) * f y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
743 |
using cvx by simp |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
744 |
also have "\<dots> = a * (f x - f y) + f y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
745 |
by (simp add: field_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
746 |
finally have "f t - f y \<le> a * (f x - f y)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
747 |
by simp |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
748 |
with t show "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
749 |
by (simp add: le_divide_eq divide_le_eq field_simps a_def) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
750 |
with t show "(f x - f y) / (x - y) \<le> (f t - f y) / (t - y)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
751 |
by (simp add: le_divide_eq divide_le_eq field_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
752 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
753 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
754 |
lemma pos_convex_function: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
755 |
fixes f :: "real \<Rightarrow> real" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
756 |
assumes "convex C" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
757 |
and leq: "\<And>x y. x \<in> C \<Longrightarrow> y \<in> C \<Longrightarrow> f' x * (y - x) \<le> f y - f x" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
758 |
shows "convex_on C f" |
72385 | 759 |
unfolding convex_on_alt |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
760 |
using assms |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
761 |
proof safe |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
762 |
fix x y \<mu> :: real |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
763 |
let ?x = "\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
764 |
assume *: "convex C" "x \<in> C" "y \<in> C" "\<mu> \<ge> 0" "\<mu> \<le> 1" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
765 |
then have "1 - \<mu> \<ge> 0" by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
766 |
then have xpos: "?x \<in> C" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
767 |
using * unfolding convex_alt by fastforce |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
768 |
have geq: "\<mu> * (f x - f ?x) + (1 - \<mu>) * (f y - f ?x) \<ge> |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
769 |
\<mu> * f' ?x * (x - ?x) + (1 - \<mu>) * f' ?x * (y - ?x)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
770 |
using add_mono [OF mult_left_mono [OF leq [OF xpos *(2)] \<open>\<mu> \<ge> 0\<close>] |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
771 |
mult_left_mono [OF leq [OF xpos *(3)] \<open>1 - \<mu> \<ge> 0\<close>]] |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
772 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
773 |
then have "\<mu> * f x + (1 - \<mu>) * f y - f ?x \<ge> 0" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
774 |
by (auto simp: field_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
775 |
then show "f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y" |
72385 | 776 |
by auto |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
777 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
778 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
779 |
lemma atMostAtLeast_subset_convex: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
780 |
fixes C :: "real set" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
781 |
assumes "convex C" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
782 |
and "x \<in> C" "y \<in> C" "x < y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
783 |
shows "{x .. y} \<subseteq> C" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
784 |
proof safe |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
785 |
fix z assume z: "z \<in> {x .. y}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
786 |
have less: "z \<in> C" if *: "x < z" "z < y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
787 |
proof - |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
788 |
let ?\<mu> = "(y - z) / (y - x)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
789 |
have "0 \<le> ?\<mu>" "?\<mu> \<le> 1" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
790 |
using assms * by (auto simp: field_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
791 |
then have comb: "?\<mu> * x + (1 - ?\<mu>) * y \<in> C" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
792 |
using assms iffD1[OF convex_alt, rule_format, of C y x ?\<mu>] |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
793 |
by (simp add: algebra_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
794 |
have "?\<mu> * x + (1 - ?\<mu>) * y = (y - z) * x / (y - x) + (1 - (y - z) / (y - x)) * y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
795 |
by (auto simp: field_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
796 |
also have "\<dots> = ((y - z) * x + (y - x - (y - z)) * y) / (y - x)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
797 |
using assms by (simp only: add_divide_distrib) (auto simp: field_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
798 |
also have "\<dots> = z" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
799 |
using assms by (auto simp: field_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
800 |
finally show ?thesis |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
801 |
using comb by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
802 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
803 |
show "z \<in> C" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
804 |
using z less assms by (auto simp: le_less) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
805 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
806 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
807 |
lemma f''_imp_f': |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
808 |
fixes f :: "real \<Rightarrow> real" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
809 |
assumes "convex C" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
810 |
and f': "\<And>x. x \<in> C \<Longrightarrow> DERIV f x :> (f' x)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
811 |
and f'': "\<And>x. x \<in> C \<Longrightarrow> DERIV f' x :> (f'' x)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
812 |
and pos: "\<And>x. x \<in> C \<Longrightarrow> f'' x \<ge> 0" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
813 |
and x: "x \<in> C" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
814 |
and y: "y \<in> C" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
815 |
shows "f' x * (y - x) \<le> f y - f x" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
816 |
using assms |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
817 |
proof - |
78475 | 818 |
have "f y - f x \<ge> f' x * (y - x)" "f' y * (x - y) \<le> f x - f y" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
819 |
if *: "x \<in> C" "y \<in> C" "y > x" for x y :: real |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
820 |
proof - |
78475 | 821 |
from * have ge: "y - x > 0" "y - x \<ge> 0" and le: "x - y < 0" "x - y \<le> 0" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
822 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
823 |
then obtain z1 where z1: "z1 > x" "z1 < y" "f y - f x = (y - x) * f' z1" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
824 |
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>], |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
825 |
THEN f', THEN MVT2[OF \<open>x < y\<close>, rule_format, unfolded atLeastAtMost_iff[symmetric]]] |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
826 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
827 |
then have "z1 \<in> C" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
828 |
using atMostAtLeast_subset_convex \<open>convex C\<close> \<open>x \<in> C\<close> \<open>y \<in> C\<close> \<open>x < y\<close> |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
829 |
by fastforce |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
830 |
obtain z2 where z2: "z2 > x" "z2 < z1" "f' z1 - f' x = (z1 - x) * f'' z2" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
831 |
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>], |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
832 |
THEN f'', THEN MVT2[OF \<open>x < z1\<close>, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1 |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
833 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
834 |
obtain z3 where z3: "z3 > z1" "z3 < y" "f' y - f' z1 = (y - z1) * f'' z3" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
835 |
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>], |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
836 |
THEN f'', THEN MVT2[OF \<open>z1 < y\<close>, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1 |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
837 |
by auto |
78475 | 838 |
from z1 have "f x - f y = (x - y) * f' z1" |
839 |
by (simp add: field_simps) |
|
840 |
then have cool': "f' y - (f x - f y) / (x - y) = (y - z1) * f'' z3" |
|
841 |
using le(1) z3(3) by auto |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
842 |
have "z3 \<in> C" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
843 |
using z3 * atMostAtLeast_subset_convex \<open>convex C\<close> \<open>x \<in> C\<close> \<open>z1 \<in> C\<close> \<open>x < z1\<close> |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
844 |
by fastforce |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
845 |
then have B': "f'' z3 \<ge> 0" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
846 |
using assms by auto |
78475 | 847 |
with cool' have "f' y - (f x - f y) / (x - y) \<ge> 0" |
848 |
using z1 by auto |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
849 |
then have res: "f' y * (x - y) \<le> f x - f y" |
78475 | 850 |
by (meson diff_ge_0_iff_ge le(1) neg_divide_le_eq) |
851 |
have cool: "(f y - f x) / (y - x) - f' x = (z1 - x) * f'' z2" |
|
852 |
using le(1) z1(3) z2(3) by auto |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
853 |
have "z2 \<in> C" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
854 |
using z2 z1 * atMostAtLeast_subset_convex \<open>convex C\<close> \<open>z1 \<in> C\<close> \<open>y \<in> C\<close> \<open>z1 < y\<close> |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
855 |
by fastforce |
78475 | 856 |
with z1 assms have "(z1 - x) * f'' z2 \<ge> 0" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
857 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
858 |
then show "f y - f x \<ge> f' x * (y - x)" "f' y * (x - y) \<le> f x - f y" |
78475 | 859 |
using that(3) z1(3) res cool by auto |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
860 |
qed |
78475 | 861 |
then show ?thesis |
862 |
using x y by fastforce |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
863 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
864 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
865 |
lemma f''_ge0_imp_convex: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
866 |
fixes f :: "real \<Rightarrow> real" |
78475 | 867 |
assumes "convex C" |
868 |
and "\<And>x. x \<in> C \<Longrightarrow> DERIV f x :> (f' x)" |
|
869 |
and "\<And>x. x \<in> C \<Longrightarrow> DERIV f' x :> (f'' x)" |
|
870 |
and "\<And>x. x \<in> C \<Longrightarrow> f'' x \<ge> 0" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
871 |
shows "convex_on C f" |
78475 | 872 |
by (metis assms f''_imp_f' pos_convex_function) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
873 |
|
74729
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
874 |
lemma f''_le0_imp_concave: |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
875 |
fixes f :: "real \<Rightarrow> real" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
876 |
assumes "convex C" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
877 |
and "\<And>x. x \<in> C \<Longrightarrow> DERIV f x :> (f' x)" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
878 |
and "\<And>x. x \<in> C \<Longrightarrow> DERIV f' x :> (f'' x)" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
879 |
and "\<And>x. x \<in> C \<Longrightarrow> f'' x \<le> 0" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
880 |
shows "concave_on C f" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
881 |
unfolding concave_on_def |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
882 |
by (rule assms f''_ge0_imp_convex derivative_eq_intros | simp)+ |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
883 |
|
79583
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
884 |
lemma convex_power_even: |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
885 |
assumes "even n" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
886 |
shows "convex_on (UNIV::real set) (\<lambda>x. x^n)" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
887 |
proof (intro f''_ge0_imp_convex) |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
888 |
show "((\<lambda>x. x ^ n) has_real_derivative of_nat n * x^(n-1)) (at x)" for x |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
889 |
by (rule derivative_eq_intros | simp)+ |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
890 |
show "((\<lambda>x. of_nat n * x^(n-1)) has_real_derivative of_nat n * of_nat (n-1) * x^(n-2)) (at x)" for x |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
891 |
by (rule derivative_eq_intros | simp add: eval_nat_numeral)+ |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
892 |
show "\<And>x. 0 \<le> real n * real (n - 1) * x ^ (n - 2)" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
893 |
using assms by (auto simp: zero_le_mult_iff zero_le_even_power) |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
894 |
qed auto |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
895 |
|
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
896 |
lemma convex_power_odd: |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
897 |
assumes "odd n" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
898 |
shows "convex_on {0::real..} (\<lambda>x. x^n)" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
899 |
proof (intro f''_ge0_imp_convex) |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
900 |
show "((\<lambda>x. x ^ n) has_real_derivative of_nat n * x^(n-1)) (at x)" for x |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
901 |
by (rule derivative_eq_intros | simp)+ |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
902 |
show "((\<lambda>x. of_nat n * x^(n-1)) has_real_derivative of_nat n * of_nat (n-1) * x^(n-2)) (at x)" for x |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
903 |
by (rule derivative_eq_intros | simp add: eval_nat_numeral)+ |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
904 |
show "\<And>x. x \<in> {0::real..} \<Longrightarrow> 0 \<le> real n * real (n - 1) * x ^ (n - 2)" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
905 |
using assms by (auto simp: zero_le_mult_iff zero_le_even_power) |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
906 |
qed auto |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
907 |
|
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
908 |
lemma convex_power2: "convex_on (UNIV::real set) power2" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
909 |
by (simp add: convex_power_even) |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
910 |
|
74729
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
911 |
lemma log_concave: |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
912 |
fixes b :: real |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
913 |
assumes "b > 1" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
914 |
shows "concave_on {0<..} (\<lambda> x. log b x)" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
915 |
using assms |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
916 |
by (intro f''_le0_imp_concave derivative_eq_intros | simp)+ |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
917 |
|
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
918 |
lemma ln_concave: "concave_on {0<..} ln" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
919 |
unfolding log_ln by (simp add: log_concave) |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
920 |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
921 |
lemma minus_log_convex: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
922 |
fixes b :: real |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
923 |
assumes "b > 1" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
924 |
shows "convex_on {0 <..} (\<lambda> x. - log b x)" |
74729
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
925 |
using assms concave_on_def log_concave by blast |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
926 |
|
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
927 |
lemma powr_convex: |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
928 |
assumes "p \<ge> 1" shows "convex_on {0<..} (\<lambda>x. x powr p)" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
929 |
using assms |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
930 |
by (intro f''_ge0_imp_convex derivative_eq_intros | simp)+ |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
931 |
|
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
932 |
lemma exp_convex: "convex_on UNIV exp" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
933 |
by (intro f''_ge0_imp_convex derivative_eq_intros | simp)+ |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
934 |
|
79670
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
935 |
text \<open>The AM-GM inequality: the arithmetic mean exceeds the geometric mean.\<close> |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
936 |
lemma arith_geom_mean: |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
937 |
fixes x :: "'a \<Rightarrow> real" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
938 |
assumes "finite S" "S \<noteq> {}" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
939 |
and x: "\<And>i. i \<in> S \<Longrightarrow> x i \<ge> 0" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
940 |
shows "(\<Sum>i \<in> S. x i / card S) \<ge> (\<Prod>i \<in> S. x i) powr (1 / card S)" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
941 |
proof (cases "\<exists>i\<in>S. x i = 0") |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
942 |
case True |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
943 |
then have "(\<Prod>i \<in> S. x i) = 0" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
944 |
by (simp add: \<open>finite S\<close>) |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
945 |
moreover have "(\<Sum>i \<in> S. x i / card S) \<ge> 0" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
946 |
by (simp add: sum_nonneg x) |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
947 |
ultimately show ?thesis |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
948 |
by simp |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
949 |
next |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
950 |
case False |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
951 |
have "ln (\<Sum>i \<in> S. (1 / card S) *\<^sub>R x i) \<ge> (\<Sum>i \<in> S. (1 / card S) * ln (x i))" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
952 |
proof (intro concave_on_sum) |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
953 |
show "concave_on {0<..} ln" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
954 |
by (simp add: ln_concave) |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
955 |
show "\<And>i. i\<in>S \<Longrightarrow> x i \<in> {0<..}" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
956 |
using False x by fastforce |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
957 |
qed (use assms False in auto) |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
958 |
moreover have "(\<Sum>i \<in> S. (1 / card S) *\<^sub>R x i) > 0" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
959 |
using False assms by (simp add: card_gt_0_iff less_eq_real_def sum_pos) |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
960 |
ultimately have "(\<Sum>i \<in> S. (1 / card S) *\<^sub>R x i) \<ge> exp (\<Sum>i \<in> S. (1 / card S) * ln (x i))" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
961 |
using ln_ge_iff by blast |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
962 |
then have "(\<Sum>i \<in> S. x i / card S) \<ge> exp (\<Sum>i \<in> S. ln (x i) / card S)" |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
963 |
by (simp add: divide_simps) |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
964 |
then show ?thesis |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
965 |
using assms False |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
966 |
by (smt (verit, ccfv_SIG) divide_inverse exp_ln exp_powr_real exp_sum inverse_eq_divide prod.cong prod_powr_distrib) |
f471e1715fc4
A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents:
79583
diff
changeset
|
967 |
qed |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
968 |
|
70136 | 969 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>Convexity of real functions\<close> |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
970 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
971 |
lemma convex_on_realI: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
972 |
assumes "connected A" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
973 |
and "\<And>x. x \<in> A \<Longrightarrow> (f has_real_derivative f' x) (at x)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
974 |
and "\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x \<le> y \<Longrightarrow> f' x \<le> f' y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
975 |
shows "convex_on A f" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
976 |
proof (rule convex_on_linorderI) |
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
977 |
show "convex A" |
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
978 |
using \<open>connected A\<close> convex_real_interval interval_cases |
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
979 |
by (smt (verit, ccfv_SIG) connectedD_interval convex_UNIV convex_empty) |
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
980 |
\<comment> \<open>the equivalence of "connected" and "convex" for real intervals is proved later\<close> |
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
981 |
next |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
982 |
fix t x y :: real |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
983 |
assume t: "t > 0" "t < 1" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
984 |
assume xy: "x \<in> A" "y \<in> A" "x < y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
985 |
define z where "z = (1 - t) * x + t * y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
986 |
with \<open>connected A\<close> and xy have ivl: "{x..y} \<subseteq> A" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
987 |
using connected_contains_Icc by blast |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
988 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
989 |
from xy t have xz: "z > x" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
990 |
by (simp add: z_def algebra_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
991 |
have "y - z = (1 - t) * (y - x)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
992 |
by (simp add: z_def algebra_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
993 |
also from xy t have "\<dots> > 0" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
994 |
by (intro mult_pos_pos) simp_all |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
995 |
finally have yz: "z < y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
996 |
by simp |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
997 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
998 |
from assms xz yz ivl t have "\<exists>\<xi>. \<xi> > x \<and> \<xi> < z \<and> f z - f x = (z - x) * f' \<xi>" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
999 |
by (intro MVT2) (auto intro!: assms(2)) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1000 |
then obtain \<xi> where \<xi>: "\<xi> > x" "\<xi> < z" "f' \<xi> = (f z - f x) / (z - x)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1001 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1002 |
from assms xz yz ivl t have "\<exists>\<eta>. \<eta> > z \<and> \<eta> < y \<and> f y - f z = (y - z) * f' \<eta>" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1003 |
by (intro MVT2) (auto intro!: assms(2)) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1004 |
then obtain \<eta> where \<eta>: "\<eta> > z" "\<eta> < y" "f' \<eta> = (f y - f z) / (y - z)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1005 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1006 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1007 |
from \<eta>(3) have "(f y - f z) / (y - z) = f' \<eta>" .. |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1008 |
also from \<xi> \<eta> ivl have "\<xi> \<in> A" "\<eta> \<in> A" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1009 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1010 |
with \<xi> \<eta> have "f' \<eta> \<ge> f' \<xi>" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1011 |
by (intro assms(3)) auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1012 |
also from \<xi>(3) have "f' \<xi> = (f z - f x) / (z - x)" . |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1013 |
finally have "(f y - f z) * (z - x) \<ge> (f z - f x) * (y - z)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1014 |
using xz yz by (simp add: field_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1015 |
also have "z - x = t * (y - x)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1016 |
by (simp add: z_def algebra_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1017 |
also have "y - z = (1 - t) * (y - x)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1018 |
by (simp add: z_def algebra_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1019 |
finally have "(f y - f z) * t \<ge> (f z - f x) * (1 - t)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1020 |
using xy by simp |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1021 |
then show "(1 - t) * f x + t * f y \<ge> f ((1 - t) *\<^sub>R x + t *\<^sub>R y)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1022 |
by (simp add: z_def algebra_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1023 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1024 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1025 |
lemma convex_on_inverse: |
78475 | 1026 |
fixes A :: "real set" |
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
1027 |
assumes "A \<subseteq> {0<..}" "convex A" |
78475 | 1028 |
shows "convex_on A inverse" |
1029 |
proof - |
|
1030 |
have "convex_on {0::real<..} inverse" |
|
1031 |
proof (intro convex_on_realI) |
|
1032 |
fix u v :: real |
|
1033 |
assume "u \<in> {0<..}" "v \<in> {0<..}" "u \<le> v" |
|
1034 |
with assms show "-inverse (u^2) \<le> -inverse (v^2)" |
|
1035 |
by simp |
|
1036 |
next |
|
1037 |
show "\<And>x. x \<in> {0<..} \<Longrightarrow> (inverse has_real_derivative - inverse (x\<^sup>2)) (at x)" |
|
1038 |
by (rule derivative_eq_intros | simp add: power2_eq_square)+ |
|
1039 |
qed auto |
|
1040 |
then show ?thesis |
|
1041 |
using assms convex_on_subset by blast |
|
1042 |
qed |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1043 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1044 |
lemma convex_onD_Icc': |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1045 |
assumes "convex_on {x..y} f" "c \<in> {x..y}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1046 |
defines "d \<equiv> y - x" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1047 |
shows "f c \<le> (f y - f x) / d * (c - x) + f x" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1048 |
proof (cases x y rule: linorder_cases) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1049 |
case less |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1050 |
then have d: "d > 0" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1051 |
by (simp add: d_def) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1052 |
from assms(2) less have A: "0 \<le> (c - x) / d" "(c - x) / d \<le> 1" |
70817
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents:
70136
diff
changeset
|
1053 |
by (simp_all add: d_def field_split_simps) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1054 |
have "f c = f (x + (c - x) * 1)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1055 |
by simp |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1056 |
also from less have "1 = ((y - x) / d)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1057 |
by (simp add: d_def) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1058 |
also from d have "x + (c - x) * \<dots> = (1 - (c - x) / d) *\<^sub>R x + ((c - x) / d) *\<^sub>R y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1059 |
by (simp add: field_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1060 |
also have "f \<dots> \<le> (1 - (c - x) / d) * f x + (c - x) / d * f y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1061 |
using assms less by (intro convex_onD_Icc) simp_all |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1062 |
also from d have "\<dots> = (f y - f x) / d * (c - x) + f x" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1063 |
by (simp add: field_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1064 |
finally show ?thesis . |
78475 | 1065 |
qed (use assms in auto) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1066 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1067 |
lemma convex_onD_Icc'': |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1068 |
assumes "convex_on {x..y} f" "c \<in> {x..y}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1069 |
defines "d \<equiv> y - x" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1070 |
shows "f c \<le> (f x - f y) / d * (y - c) + f y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1071 |
proof (cases x y rule: linorder_cases) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1072 |
case less |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1073 |
then have d: "d > 0" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1074 |
by (simp add: d_def) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1075 |
from assms(2) less have A: "0 \<le> (y - c) / d" "(y - c) / d \<le> 1" |
70817
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents:
70136
diff
changeset
|
1076 |
by (simp_all add: d_def field_split_simps) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1077 |
have "f c = f (y - (y - c) * 1)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1078 |
by simp |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1079 |
also from less have "1 = ((y - x) / d)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1080 |
by (simp add: d_def) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1081 |
also from d have "y - (y - c) * \<dots> = (1 - (1 - (y - c) / d)) *\<^sub>R x + (1 - (y - c) / d) *\<^sub>R y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1082 |
by (simp add: field_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1083 |
also have "f \<dots> \<le> (1 - (1 - (y - c) / d)) * f x + (1 - (y - c) / d) * f y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1084 |
using assms less by (intro convex_onD_Icc) (simp_all add: field_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1085 |
also from d have "\<dots> = (f x - f y) / d * (y - c) + f y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1086 |
by (simp add: field_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1087 |
finally show ?thesis . |
78475 | 1088 |
qed (use assms in auto) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1089 |
|
79945
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1090 |
lemma concave_onD_Icc: |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1091 |
assumes "concave_on {x..y} f" "x \<le> (y :: _ :: {real_vector,preorder})" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1092 |
shows "\<And>t. t \<ge> 0 \<Longrightarrow> t \<le> 1 \<Longrightarrow> |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1093 |
f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \<ge> (1 - t) * f x + t * f y" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1094 |
using assms(2) by (intro concave_onD [OF assms(1)]) simp_all |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1095 |
|
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1096 |
lemma concave_onD_Icc': |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1097 |
assumes "concave_on {x..y} f" "c \<in> {x..y}" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1098 |
defines "d \<equiv> y - x" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1099 |
shows "f c \<ge> (f y - f x) / d * (c - x) + f x" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1100 |
proof - |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1101 |
have "- f c \<le> (f x - f y) / d * (c - x) - f x" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1102 |
using assms convex_onD_Icc' [of x y "\<lambda>x. - f x" c] |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1103 |
by (simp add: concave_on_def) |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1104 |
then show ?thesis |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1105 |
by (smt (verit, best) divide_minus_left mult_minus_left) |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1106 |
qed |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1107 |
|
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1108 |
lemma concave_onD_Icc'': |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1109 |
assumes "concave_on {x..y} f" "c \<in> {x..y}" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1110 |
defines "d \<equiv> y - x" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1111 |
shows "f c \<ge> (f x - f y) / d * (y - c) + f y" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1112 |
proof - |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1113 |
have "- f c \<le> (f y - f x) / d * (y - c) - f y" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1114 |
using assms convex_onD_Icc'' [of x y "\<lambda>x. - f x" c] |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1115 |
by (simp add: concave_on_def) |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1116 |
then show ?thesis |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1117 |
by (smt (verit, best) divide_minus_left mult_minus_left) |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1118 |
qed |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1119 |
|
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1120 |
lemma convex_on_le_max: |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1121 |
fixes a::real |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1122 |
assumes "convex_on {x..y} f" and a: "a \<in> {x..y}" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1123 |
shows "f a \<le> max (f x) (f y)" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1124 |
proof - |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1125 |
have *: "(f y - f x) * (a - x) \<le> (f y - f x) * (y - x)" if "f x \<le> f y" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1126 |
using a that by (intro mult_left_mono) auto |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1127 |
have "f a \<le> (f y - f x) / (y - x) * (a - x) + f x" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1128 |
using assms convex_onD_Icc' by blast |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1129 |
also have "\<dots> \<le> max (f x) (f y)" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1130 |
using a * |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1131 |
by (simp add: divide_le_0_iff mult_le_0_iff zero_le_mult_iff max_def add.commute mult.commute scaling_mono) |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1132 |
finally show ?thesis . |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1133 |
qed |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1134 |
|
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1135 |
lemma concave_on_ge_min: |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1136 |
fixes a::real |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1137 |
assumes "concave_on {x..y} f" and a: "a \<in> {x..y}" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1138 |
shows "f a \<ge> min (f x) (f y)" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1139 |
proof - |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1140 |
have *: "(f y - f x) * (a - x) \<ge> (f y - f x) * (y - x)" if "f x \<ge> f y" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1141 |
using a that by (intro mult_left_mono_neg) auto |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1142 |
have "min (f x) (f y) \<le> (f y - f x) / (y - x) * (a - x) + f x" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1143 |
using a * apply (simp add: zero_le_divide_iff mult_le_0_iff zero_le_mult_iff min_def) |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1144 |
by (smt (verit, best) nonzero_eq_divide_eq pos_divide_le_eq) |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1145 |
also have "\<dots> \<le> f a" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1146 |
using assms concave_onD_Icc' by blast |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1147 |
finally show ?thesis . |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1148 |
qed |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
79670
diff
changeset
|
1149 |
|
80653
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1150 |
subsection \<open>Convexity of the generalised binomial\<close> |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1151 |
|
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1152 |
lemma mono_on_mul: |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1153 |
fixes f::"'a::ord \<Rightarrow> 'b::ordered_semiring" |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1154 |
assumes "mono_on S f" "mono_on S g" |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1155 |
assumes fty: "f \<in> S \<rightarrow> {0..}" and gty: "g \<in> S \<rightarrow> {0..}" |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1156 |
shows "mono_on S (\<lambda>x. f x * g x)" |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1157 |
using assms by (auto simp: Pi_iff monotone_on_def intro!: mult_mono) |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1158 |
|
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1159 |
lemma mono_on_prod: |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1160 |
fixes f::"'i \<Rightarrow> 'a::ord \<Rightarrow> 'b::linordered_idom" |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1161 |
assumes "\<And>i. i \<in> I \<Longrightarrow> mono_on S (f i)" |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1162 |
assumes "\<And>i. i \<in> I \<Longrightarrow> f i \<in> S \<rightarrow> {0..}" |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1163 |
shows "mono_on S (\<lambda>x. prod (\<lambda>i. f i x) I)" |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1164 |
using assms |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1165 |
by (induction I rule: infinite_finite_induct) |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1166 |
(auto simp: mono_on_const Pi_iff prod_nonneg mono_on_mul mono_onI) |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1167 |
|
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1168 |
lemma convex_gchoose_aux: "convex_on {k-1..} (\<lambda>a. prod (\<lambda>i. a - of_nat i) {0..<k})" |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1169 |
proof (induction k) |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1170 |
case 0 |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1171 |
then show ?case |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1172 |
by (simp add: convex_on_def) |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1173 |
next |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1174 |
case (Suc k) |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1175 |
have "convex_on {real k..} (\<lambda>a. (\<Prod>i = 0..<k. a - real i) * (a - real k))" |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1176 |
proof (intro convex_on_mul convex_on_diff) |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1177 |
show "convex_on {real k..} (\<lambda>x. \<Prod>i = 0..<k. x - real i)" |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1178 |
using Suc convex_on_subset by fastforce |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1179 |
show "mono_on {real k..} (\<lambda>x. \<Prod>i = 0..<k. x - real i)" |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1180 |
by (force simp: monotone_on_def intro!: prod_mono) |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1181 |
next |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1182 |
show "(\<lambda>x. \<Prod>i = 0..<k. x - real i) \<in> {real k..} \<rightarrow> {0..}" |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1183 |
by (auto intro!: prod_nonneg) |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1184 |
qed (auto simp: convex_on_ident concave_on_const mono_onI) |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1185 |
then show ?case |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1186 |
by simp |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1187 |
qed |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1188 |
|
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1189 |
lemma convex_gchoose: "convex_on {k-1..} (\<lambda>x. x gchoose k)" |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1190 |
by (simp add: gbinomial_prod_rev convex_on_cdiv convex_gchoose_aux) |
b98f1057da0e
New library material; also fixed the spelling error powr_ge_pzero -> powr_ge_zero
paulson <lp15@cam.ac.uk>
parents:
79945
diff
changeset
|
1191 |
|
79583
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
1192 |
subsection \<open>Some inequalities: Applications of convexity\<close> |
74729
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1193 |
|
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1194 |
lemma Youngs_inequality_0: |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1195 |
fixes a::real |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1196 |
assumes "0 \<le> \<alpha>" "0 \<le> \<beta>" "\<alpha>+\<beta> = 1" "a>0" "b>0" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1197 |
shows "a powr \<alpha> * b powr \<beta> \<le> \<alpha>*a + \<beta>*b" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1198 |
proof - |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1199 |
have "\<alpha> * ln a + \<beta> * ln b \<le> ln (\<alpha> * a + \<beta> * b)" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1200 |
using assms ln_concave by (simp add: concave_on_iff) |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1201 |
moreover have "0 < \<alpha> * a + \<beta> * b" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1202 |
using assms by (smt (verit) mult_pos_pos split_mult_pos_le) |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1203 |
ultimately show ?thesis |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1204 |
using assms by (simp add: powr_def mult_exp_exp flip: ln_ge_iff) |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1205 |
qed |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1206 |
|
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1207 |
lemma Youngs_inequality: |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1208 |
fixes p::real |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1209 |
assumes "p>1" "q>1" "1/p + 1/q = 1" "a\<ge>0" "b\<ge>0" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1210 |
shows "a * b \<le> a powr p / p + b powr q / q" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1211 |
proof (cases "a=0 \<or> b=0") |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1212 |
case False |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1213 |
then show ?thesis |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1214 |
using Youngs_inequality_0 [of "1/p" "1/q" "a powr p" "b powr q"] assms |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1215 |
by (simp add: powr_powr) |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1216 |
qed (use assms in auto) |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1217 |
|
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1218 |
lemma Cauchy_Schwarz_ineq_sum: |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1219 |
fixes a :: "'a \<Rightarrow> 'b::linordered_field" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1220 |
shows "(\<Sum>i\<in>I. a i * b i)\<^sup>2 \<le> (\<Sum>i\<in>I. (a i)\<^sup>2) * (\<Sum>i\<in>I. (b i)\<^sup>2)" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1221 |
proof (cases "(\<Sum>i\<in>I. (b i)\<^sup>2) > 0") |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1222 |
case False |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1223 |
then consider "\<And>i. i\<in>I \<Longrightarrow> b i = 0" | "infinite I" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1224 |
by (metis (mono_tags, lifting) sum_pos2 zero_le_power2 zero_less_power2) |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1225 |
thus ?thesis |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1226 |
by fastforce |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1227 |
next |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1228 |
case True |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1229 |
define r where "r \<equiv> (\<Sum>i\<in>I. a i * b i) / (\<Sum>i\<in>I. (b i)\<^sup>2)" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1230 |
have "0 \<le> (\<Sum>i\<in>I. (a i - r * b i)\<^sup>2)" |
79532
bb5d036f3523
Type class patch suggested by Achim Brucker, plus tidied lemma
paulson <lp15@cam.ac.uk>
parents:
78656
diff
changeset
|
1231 |
by (simp add: sum_nonneg) |
74729
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1232 |
also have "... = (\<Sum>i\<in>I. (a i)\<^sup>2) - 2 * r * (\<Sum>i\<in>I. a i * b i) + r\<^sup>2 * (\<Sum>i\<in>I. (b i)\<^sup>2)" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1233 |
by (simp add: algebra_simps power2_eq_square sum_distrib_left flip: sum.distrib) |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1234 |
also have "\<dots> = (\<Sum>i\<in>I. (a i)\<^sup>2) - ((\<Sum>i\<in>I. a i * b i))\<^sup>2 / (\<Sum>i\<in>I. (b i)\<^sup>2)" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1235 |
by (simp add: r_def power2_eq_square) |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1236 |
finally have "0 \<le> (\<Sum>i\<in>I. (a i)\<^sup>2) - ((\<Sum>i\<in>I. a i * b i))\<^sup>2 / (\<Sum>i\<in>I. (b i)\<^sup>2)" . |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1237 |
hence "((\<Sum>i\<in>I. a i * b i))\<^sup>2 / (\<Sum>i\<in>I. (b i)\<^sup>2) \<le> (\<Sum>i\<in>I. (a i)\<^sup>2)" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1238 |
by (simp add: le_diff_eq) |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1239 |
thus "((\<Sum>i\<in>I. a i * b i))\<^sup>2 \<le> (\<Sum>i\<in>I. (a i)\<^sup>2) * (\<Sum>i\<in>I. (b i)\<^sup>2)" |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1240 |
by (simp add: pos_divide_le_eq True) |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1241 |
qed |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1242 |
|
82485
12fe1e2b87e4
Generalised a lemma and added another
paulson <lp15@cam.ac.uk>
parents:
80654
diff
changeset
|
1243 |
text \<open>The inequality between the arithmetic mean and the root mean square\<close> |
79583
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
1244 |
lemma sum_squared_le_sum_of_squares: |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
1245 |
fixes f :: "'a \<Rightarrow> real" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
1246 |
shows "(\<Sum>i\<in>I. f i)\<^sup>2 \<le> (\<Sum>y\<in>I. (f y)\<^sup>2) * card I" |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
1247 |
proof (cases "finite I \<and> I \<noteq> {}") |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
1248 |
case True |
82485
12fe1e2b87e4
Generalised a lemma and added another
paulson <lp15@cam.ac.uk>
parents:
80654
diff
changeset
|
1249 |
then have "(\<Sum>i\<in>I. f i / of_nat (card I))\<^sup>2 \<le> (\<Sum>i\<in>I. (f i)\<^sup>2 / of_nat (card I))" |
82488 | 1250 |
using convex_on_sum [OF _ _ convex_power2, where a = "\<lambda>x. 1 / of_nat(card I)" and S=I] |
79583
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
1251 |
by simp |
82488 | 1252 |
with True show ?thesis |
79583
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
1253 |
by (simp add: divide_simps power2_eq_square split: if_split_asm flip: sum_divide_distrib) |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
1254 |
qed auto |
a521c241e946
Further lemmas concerning complexity and measures
paulson <lp15@cam.ac.uk>
parents:
79582
diff
changeset
|
1255 |
|
82485
12fe1e2b87e4
Generalised a lemma and added another
paulson <lp15@cam.ac.uk>
parents:
80654
diff
changeset
|
1256 |
lemma sum_squared_le_sum_of_squares_2: |
12fe1e2b87e4
Generalised a lemma and added another
paulson <lp15@cam.ac.uk>
parents:
80654
diff
changeset
|
1257 |
"(x+y)/2 \<le> sqrt ((x\<^sup>2 + y\<^sup>2) / 2)" |
12fe1e2b87e4
Generalised a lemma and added another
paulson <lp15@cam.ac.uk>
parents:
80654
diff
changeset
|
1258 |
proof - |
12fe1e2b87e4
Generalised a lemma and added another
paulson <lp15@cam.ac.uk>
parents:
80654
diff
changeset
|
1259 |
have "(x + y)\<^sup>2 / 2^2 \<le> (x\<^sup>2 + y\<^sup>2) / 2" |
82488 | 1260 |
using sum_squared_le_sum_of_squares [of "\<lambda>b. if b then x else y" UNIV] |
82485
12fe1e2b87e4
Generalised a lemma and added another
paulson <lp15@cam.ac.uk>
parents:
80654
diff
changeset
|
1261 |
by (simp add: UNIV_bool add.commute) |
12fe1e2b87e4
Generalised a lemma and added another
paulson <lp15@cam.ac.uk>
parents:
80654
diff
changeset
|
1262 |
then show ?thesis |
12fe1e2b87e4
Generalised a lemma and added another
paulson <lp15@cam.ac.uk>
parents:
80654
diff
changeset
|
1263 |
by (metis power_divide real_le_rsqrt) |
12fe1e2b87e4
Generalised a lemma and added another
paulson <lp15@cam.ac.uk>
parents:
80654
diff
changeset
|
1264 |
qed |
12fe1e2b87e4
Generalised a lemma and added another
paulson <lp15@cam.ac.uk>
parents:
80654
diff
changeset
|
1265 |
|
74729
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1266 |
subsection \<open>Misc related lemmas\<close> |
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
paulson <lp15@cam.ac.uk>
parents:
72385
diff
changeset
|
1267 |
|
69661 | 1268 |
lemma convex_translation_eq [simp]: |
1269 |
"convex ((+) a ` s) \<longleftrightarrow> convex s" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1270 |
by (metis convex_translation translation_galois) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1271 |
|
69661 | 1272 |
lemma convex_translation_subtract_eq [simp]: |
1273 |
"convex ((\<lambda>b. b - a) ` s) \<longleftrightarrow> convex s" |
|
1274 |
using convex_translation_eq [of "- a"] by (simp cong: image_cong_simp) |
|
1275 |
||
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1276 |
lemma convex_linear_image_eq [simp]: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1277 |
fixes f :: "'a::real_vector \<Rightarrow> 'b::real_vector" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1278 |
shows "\<lbrakk>linear f; inj f\<rbrakk> \<Longrightarrow> convex (f ` s) \<longleftrightarrow> convex s" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1279 |
by (metis (no_types) convex_linear_image convex_linear_vimage inj_vimage_image_eq) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1280 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1281 |
lemma vector_choose_size: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1282 |
assumes "0 \<le> c" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1283 |
obtains x :: "'a::{real_normed_vector, perfect_space}" where "norm x = c" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1284 |
proof - |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1285 |
obtain a::'a where "a \<noteq> 0" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1286 |
using UNIV_not_singleton UNIV_eq_I set_zero singletonI by fastforce |
82488 | 1287 |
show ?thesis |
1288 |
proof |
|
1289 |
show "norm (scaleR (c / norm a) a) = c" |
|
1290 |
by (simp add: \<open>a \<noteq> 0\<close> assms) |
|
1291 |
qed |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1292 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1293 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1294 |
lemma vector_choose_dist: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1295 |
assumes "0 \<le> c" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1296 |
obtains y :: "'a::{real_normed_vector, perfect_space}" where "dist x y = c" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1297 |
by (metis add_diff_cancel_left' assms dist_commute dist_norm vector_choose_size) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1298 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1299 |
lemma sum_delta'': |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1300 |
fixes s::"'a::real_vector set" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1301 |
assumes "finite s" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1302 |
shows "(\<Sum>x\<in>s. (if y = x then f x else 0) *\<^sub>R x) = (if y\<in>s then (f y) *\<^sub>R y else 0)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1303 |
proof - |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1304 |
have *: "\<And>x y. (if y = x then f x else (0::real)) *\<^sub>R x = (if x=y then (f x) *\<^sub>R x else 0)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1305 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1306 |
show ?thesis |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1307 |
unfolding * using sum.delta[OF assms, of y "\<lambda>x. f x *\<^sub>R x"] by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1308 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1309 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1310 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1311 |
subsection \<open>Cones\<close> |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1312 |
|
70136 | 1313 |
definition\<^marker>\<open>tag important\<close> cone :: "'a::real_vector set \<Rightarrow> bool" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1314 |
where "cone s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>c\<ge>0. c *\<^sub>R x \<in> s)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1315 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1316 |
lemma cone_empty[intro, simp]: "cone {}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1317 |
unfolding cone_def by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1318 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1319 |
lemma cone_univ[intro, simp]: "cone UNIV" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1320 |
unfolding cone_def by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1321 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1322 |
lemma cone_Inter[intro]: "\<forall>s\<in>f. cone s \<Longrightarrow> cone (\<Inter>f)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1323 |
unfolding cone_def by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1324 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1325 |
lemma subspace_imp_cone: "subspace S \<Longrightarrow> cone S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1326 |
by (simp add: cone_def subspace_scale) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1327 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1328 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1329 |
subsubsection \<open>Conic hull\<close> |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1330 |
|
72385 | 1331 |
lemma cone_cone_hull: "cone (cone hull S)" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1332 |
unfolding hull_def by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1333 |
|
72385 | 1334 |
lemma cone_hull_eq: "cone hull S = S \<longleftrightarrow> cone S" |
1335 |
by (metis cone_cone_hull hull_same) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1336 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1337 |
lemma mem_cone: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1338 |
assumes "cone S" "x \<in> S" "c \<ge> 0" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1339 |
shows "c *\<^sub>R x \<in> S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1340 |
using assms cone_def[of S] by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1341 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1342 |
lemma cone_contains_0: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1343 |
assumes "cone S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1344 |
shows "S \<noteq> {} \<longleftrightarrow> 0 \<in> S" |
72385 | 1345 |
using assms mem_cone by fastforce |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1346 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1347 |
lemma cone_0: "cone {0}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1348 |
unfolding cone_def by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1349 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1350 |
lemma cone_Union[intro]: "(\<forall>s\<in>f. cone s) \<longrightarrow> cone (\<Union>f)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1351 |
unfolding cone_def by blast |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1352 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1353 |
lemma cone_iff: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1354 |
assumes "S \<noteq> {}" |
78475 | 1355 |
shows "cone S \<longleftrightarrow> 0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> ((*\<^sub>R) c) ` S = S)" (is "_ = ?rhs") |
1356 |
proof |
|
1357 |
assume "cone S" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1358 |
{ |
78475 | 1359 |
fix c :: real |
1360 |
assume "c > 0" |
|
1361 |
have "x \<in> ((*\<^sub>R) c) ` S" if "x \<in> S" for x |
|
1362 |
using \<open>cone S\<close> \<open>c>0\<close> mem_cone[of S x "1/c"] that |
|
1363 |
exI[of "(\<lambda>t. t \<in> S \<and> x = c *\<^sub>R t)" "(1 / c) *\<^sub>R x"] |
|
1364 |
by auto |
|
1365 |
then have "((*\<^sub>R) c) ` S = S" |
|
1366 |
using \<open>0 < c\<close> \<open>cone S\<close> mem_cone by fastforce |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1367 |
} |
78475 | 1368 |
then show "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> ((*\<^sub>R) c) ` S = S)" |
1369 |
using \<open>cone S\<close> cone_contains_0[of S] assms by auto |
|
1370 |
next |
|
1371 |
show "?rhs \<Longrightarrow> cone S" |
|
1372 |
by (metis Convex.cone_def imageI order_neq_le_trans scaleR_zero_left) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1373 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1374 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1375 |
lemma cone_hull_empty: "cone hull {} = {}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1376 |
by (metis cone_empty cone_hull_eq) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1377 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1378 |
lemma cone_hull_empty_iff: "S = {} \<longleftrightarrow> cone hull S = {}" |
78475 | 1379 |
by (metis cone_hull_empty hull_subset subset_empty) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1380 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1381 |
lemma cone_hull_contains_0: "S \<noteq> {} \<longleftrightarrow> 0 \<in> cone hull S" |
78475 | 1382 |
by (metis cone_cone_hull cone_contains_0 cone_hull_empty_iff) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1383 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1384 |
lemma mem_cone_hull: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1385 |
assumes "x \<in> S" "c \<ge> 0" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1386 |
shows "c *\<^sub>R x \<in> cone hull S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1387 |
by (metis assms cone_cone_hull hull_inc mem_cone) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1388 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1389 |
proposition cone_hull_expl: "cone hull S = {c *\<^sub>R x | c x. c \<ge> 0 \<and> x \<in> S}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1390 |
(is "?lhs = ?rhs") |
78475 | 1391 |
proof |
1392 |
have "?rhs \<in> Collect cone" |
|
1393 |
using Convex.cone_def by fastforce |
|
1394 |
moreover have "S \<subseteq> ?rhs" |
|
1395 |
by (smt (verit) mem_Collect_eq scaleR_one subsetI) |
|
1396 |
ultimately show "?lhs \<subseteq> ?rhs" |
|
1397 |
using hull_minimal by blast |
|
1398 |
qed (use mem_cone_hull in auto) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1399 |
|
71242 | 1400 |
lemma convex_cone: |
78475 | 1401 |
"convex S \<and> cone S \<longleftrightarrow> (\<forall>x\<in>S. \<forall>y\<in>S. (x + y) \<in> S) \<and> (\<forall>x\<in>S. \<forall>c\<ge>0. (c *\<^sub>R x) \<in> S)" |
71242 | 1402 |
(is "?lhs = ?rhs") |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1403 |
proof - |
71242 | 1404 |
{ |
1405 |
fix x y |
|
78475 | 1406 |
assume "x\<in>S" "y\<in>S" and ?lhs |
1407 |
then have "2 *\<^sub>R x \<in>S" "2 *\<^sub>R y \<in> S" "convex S" |
|
71242 | 1408 |
unfolding cone_def by auto |
78475 | 1409 |
then have "x + y \<in> S" |
1410 |
using convexD [OF \<open>convex S\<close>, of "2*\<^sub>R x" "2*\<^sub>R y"] |
|
1411 |
by (smt (verit, ccfv_threshold) field_sum_of_halves scaleR_2 scaleR_half_double) |
|
71242 | 1412 |
} |
1413 |
then show ?thesis |
|
1414 |
unfolding convex_def cone_def by blast |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1415 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1416 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1417 |
|
70136 | 1418 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>Connectedness of convex sets\<close> |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1419 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1420 |
lemma convex_connected: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1421 |
fixes S :: "'a::real_normed_vector set" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1422 |
assumes "convex S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1423 |
shows "connected S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1424 |
proof (rule connectedI) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1425 |
fix A B |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1426 |
assume "open A" "open B" "A \<inter> B \<inter> S = {}" "S \<subseteq> A \<union> B" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1427 |
moreover |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1428 |
assume "A \<inter> S \<noteq> {}" "B \<inter> S \<noteq> {}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1429 |
then obtain a b where a: "a \<in> A" "a \<in> S" and b: "b \<in> B" "b \<in> S" by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1430 |
define f where [abs_def]: "f u = u *\<^sub>R a + (1 - u) *\<^sub>R b" for u |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1431 |
then have "continuous_on {0 .. 1} f" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1432 |
by (auto intro!: continuous_intros) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1433 |
then have "connected (f ` {0 .. 1})" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1434 |
by (auto intro!: connected_continuous_image) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1435 |
note connectedD[OF this, of A B] |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1436 |
moreover have "a \<in> A \<inter> f ` {0 .. 1}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1437 |
using a by (auto intro!: image_eqI[of _ _ 1] simp: f_def) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1438 |
moreover have "b \<in> B \<inter> f ` {0 .. 1}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1439 |
using b by (auto intro!: image_eqI[of _ _ 0] simp: f_def) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1440 |
moreover have "f ` {0 .. 1} \<subseteq> S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1441 |
using \<open>convex S\<close> a b unfolding convex_def f_def by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1442 |
ultimately show False by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1443 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1444 |
|
71136 | 1445 |
corollary%unimportant connected_UNIV[intro]: "connected (UNIV :: 'a::real_normed_vector set)" |
78475 | 1446 |
by (simp add: convex_connected) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1447 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1448 |
lemma convex_prod: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1449 |
assumes "\<And>i. i \<in> Basis \<Longrightarrow> convex {x. P i x}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1450 |
shows "convex {x. \<forall>i\<in>Basis. P i (x\<bullet>i)}" |
78475 | 1451 |
using assms by (auto simp: inner_add_left convex_def) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1452 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1453 |
lemma convex_positive_orthant: "convex {x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i)}" |
71136 | 1454 |
by (rule convex_prod) (simp flip: atLeast_def) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1455 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1456 |
subsection \<open>Convex hull\<close> |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1457 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1458 |
lemma convex_convex_hull [iff]: "convex (convex hull s)" |
78475 | 1459 |
by (metis (mono_tags) convex_Inter hull_def mem_Collect_eq) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1460 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1461 |
lemma convex_hull_subset: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1462 |
"s \<subseteq> convex hull t \<Longrightarrow> convex hull s \<subseteq> convex hull t" |
71174 | 1463 |
by (simp add: subset_hull) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1464 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1465 |
lemma convex_hull_eq: "convex hull s = s \<longleftrightarrow> convex s" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1466 |
by (metis convex_convex_hull hull_same) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1467 |
|
70136 | 1468 |
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Convex hull is "preserved" by a linear function\<close> |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1469 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1470 |
lemma convex_hull_linear_image: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1471 |
assumes f: "linear f" |
78475 | 1472 |
shows "f ` (convex hull S) = convex hull (f ` S)" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1473 |
proof |
78475 | 1474 |
show "convex hull (f ` S) \<subseteq> f ` (convex hull S)" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1475 |
by (intro hull_minimal image_mono hull_subset convex_linear_image assms convex_convex_hull) |
78475 | 1476 |
show "f ` (convex hull S) \<subseteq> convex hull (f ` S)" |
1477 |
by (meson convex_convex_hull convex_linear_vimage f hull_minimal hull_subset image_subset_iff_subset_vimage) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1478 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1479 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1480 |
lemma in_convex_hull_linear_image: |
78475 | 1481 |
assumes "linear f" "x \<in> convex hull S" |
1482 |
shows "f x \<in> convex hull (f ` S)" |
|
1483 |
using assms convex_hull_linear_image image_eqI by blast |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1484 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1485 |
lemma convex_hull_Times: |
78475 | 1486 |
"convex hull (S \<times> T) = (convex hull S) \<times> (convex hull T)" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1487 |
proof |
78475 | 1488 |
show "convex hull (S \<times> T) \<subseteq> (convex hull S) \<times> (convex hull T)" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1489 |
by (intro hull_minimal Sigma_mono hull_subset convex_Times convex_convex_hull) |
78475 | 1490 |
have "(x, y) \<in> convex hull (S \<times> T)" if x: "x \<in> convex hull S" and y: "y \<in> convex hull T" for x y |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1491 |
proof (rule hull_induct [OF x], rule hull_induct [OF y]) |
78475 | 1492 |
fix x y assume "x \<in> S" and "y \<in> T" |
1493 |
then show "(x, y) \<in> convex hull (S \<times> T)" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1494 |
by (simp add: hull_inc) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1495 |
next |
78475 | 1496 |
fix x let ?S = "((\<lambda>y. (0, y)) -` (\<lambda>p. (- x, 0) + p) ` (convex hull S \<times> T))" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1497 |
have "convex ?S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1498 |
by (intro convex_linear_vimage convex_translation convex_convex_hull, |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1499 |
simp add: linear_iff) |
78475 | 1500 |
also have "?S = {y. (x, y) \<in> convex hull (S \<times> T)}" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1501 |
by (auto simp: image_def Bex_def) |
78475 | 1502 |
finally show "convex {y. (x, y) \<in> convex hull (S \<times> T)}" . |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1503 |
next |
78475 | 1504 |
show "convex {x. (x, y) \<in> convex hull S \<times> T}" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1505 |
proof - |
78475 | 1506 |
fix y let ?S = "((\<lambda>x. (x, 0)) -` (\<lambda>p. (0, - y) + p) ` (convex hull S \<times> T))" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1507 |
have "convex ?S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1508 |
by (intro convex_linear_vimage convex_translation convex_convex_hull, |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1509 |
simp add: linear_iff) |
78475 | 1510 |
also have "?S = {x. (x, y) \<in> convex hull (S \<times> T)}" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1511 |
by (auto simp: image_def Bex_def) |
78475 | 1512 |
finally show "convex {x. (x, y) \<in> convex hull (S \<times> T)}" . |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1513 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1514 |
qed |
78475 | 1515 |
then show "(convex hull S) \<times> (convex hull T) \<subseteq> convex hull (S \<times> T)" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1516 |
unfolding subset_eq split_paired_Ball_Sigma by blast |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1517 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1518 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1519 |
|
70136 | 1520 |
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Stepping theorems for convex hulls of finite sets\<close> |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1521 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1522 |
lemma convex_hull_empty[simp]: "convex hull {} = {}" |
78475 | 1523 |
by (simp add: hull_same) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1524 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1525 |
lemma convex_hull_singleton[simp]: "convex hull {a} = {a}" |
78475 | 1526 |
by (simp add: hull_same) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1527 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1528 |
lemma convex_hull_insert: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1529 |
fixes S :: "'a::real_vector set" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1530 |
assumes "S \<noteq> {}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1531 |
shows "convex hull (insert a S) = |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1532 |
{x. \<exists>u\<ge>0. \<exists>v\<ge>0. \<exists>b. (u + v = 1) \<and> b \<in> (convex hull S) \<and> (x = u *\<^sub>R a + v *\<^sub>R b)}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1533 |
(is "_ = ?hull") |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1534 |
proof (intro equalityI hull_minimal subsetI) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1535 |
fix x |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1536 |
assume "x \<in> insert a S" |
78475 | 1537 |
then show "x \<in> ?hull" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1538 |
unfolding insert_iff |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1539 |
proof |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1540 |
assume "x = a" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1541 |
then show ?thesis |
78475 | 1542 |
by (smt (verit, del_insts) add.right_neutral assms ex_in_conv hull_inc mem_Collect_eq scaleR_one scaleR_zero_left) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1543 |
next |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1544 |
assume "x \<in> S" |
78475 | 1545 |
with hull_subset show ?thesis |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1546 |
by force |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1547 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1548 |
next |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1549 |
fix x |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1550 |
assume "x \<in> ?hull" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1551 |
then obtain u v b where obt: "u\<ge>0" "v\<ge>0" "u + v = 1" "b \<in> convex hull S" "x = u *\<^sub>R a + v *\<^sub>R b" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1552 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1553 |
have "a \<in> convex hull insert a S" "b \<in> convex hull insert a S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1554 |
using hull_mono[of S "insert a S" convex] hull_mono[of "{a}" "insert a S" convex] and obt(4) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1555 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1556 |
then show "x \<in> convex hull insert a S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1557 |
unfolding obt(5) using obt(1-3) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1558 |
by (rule convexD [OF convex_convex_hull]) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1559 |
next |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1560 |
show "convex ?hull" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1561 |
proof (rule convexI) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1562 |
fix x y u v |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1563 |
assume as: "(0::real) \<le> u" "0 \<le> v" "u + v = 1" and x: "x \<in> ?hull" and y: "y \<in> ?hull" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1564 |
from x obtain u1 v1 b1 where |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1565 |
obt1: "u1\<ge>0" "v1\<ge>0" "u1 + v1 = 1" "b1 \<in> convex hull S" and xeq: "x = u1 *\<^sub>R a + v1 *\<^sub>R b1" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1566 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1567 |
from y obtain u2 v2 b2 where |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1568 |
obt2: "u2\<ge>0" "v2\<ge>0" "u2 + v2 = 1" "b2 \<in> convex hull S" and yeq: "y = u2 *\<^sub>R a + v2 *\<^sub>R b2" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1569 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1570 |
have *: "\<And>(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1571 |
by (auto simp: algebra_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1572 |
have "\<exists>b \<in> convex hull S. u *\<^sub>R x + v *\<^sub>R y = |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1573 |
(u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (b - (u * u1) *\<^sub>R b - (v * u2) *\<^sub>R b)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1574 |
proof (cases "u * v1 + v * v2 = 0") |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1575 |
case True |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1576 |
have *: "\<And>(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1577 |
by (auto simp: algebra_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1578 |
have eq0: "u * v1 = 0" "v * v2 = 0" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1579 |
using True mult_nonneg_nonneg[OF \<open>u\<ge>0\<close> \<open>v1\<ge>0\<close>] mult_nonneg_nonneg[OF \<open>v\<ge>0\<close> \<open>v2\<ge>0\<close>] |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1580 |
by arith+ |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1581 |
then have "u * u1 + v * u2 = 1" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1582 |
using as(3) obt1(3) obt2(3) by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1583 |
then show ?thesis |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1584 |
using "*" eq0 as obt1(4) xeq yeq by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1585 |
next |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1586 |
case False |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1587 |
have "1 - (u * u1 + v * u2) = (u + v) - (u * u1 + v * u2)" |
78475 | 1588 |
by (simp add: as(3)) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1589 |
also have "\<dots> = u * v1 + v * v2" |
78475 | 1590 |
by (smt (verit, ccfv_SIG) distrib_left mult_cancel_left1 obt1(3) obt2(3)) |
1591 |
finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" . |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1592 |
let ?b = "((u * v1) / (u * v1 + v * v2)) *\<^sub>R b1 + ((v * v2) / (u * v1 + v * v2)) *\<^sub>R b2" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1593 |
have zeroes: "0 \<le> u * v1 + v * v2" "0 \<le> u * v1" "0 \<le> u * v1 + v * v2" "0 \<le> v * v2" |
78475 | 1594 |
using as obt1 obt2 by auto |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1595 |
show ?thesis |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1596 |
proof |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1597 |
show "u *\<^sub>R x + v *\<^sub>R y = (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (?b - (u * u1) *\<^sub>R ?b - (v * u2) *\<^sub>R ?b)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1598 |
unfolding xeq yeq * ** |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1599 |
using False by (auto simp: scaleR_left_distrib scaleR_right_distrib) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1600 |
show "?b \<in> convex hull S" |
78475 | 1601 |
using False mem_convex_alt obt1(4) obt2(4) zeroes(2) zeroes(4) by fastforce |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1602 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1603 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1604 |
then obtain b where b: "b \<in> convex hull S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1605 |
"u *\<^sub>R x + v *\<^sub>R y = (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (b - (u * u1) *\<^sub>R b - (v * u2) *\<^sub>R b)" .. |
78475 | 1606 |
obtain u1: "u1 \<le> 1" and u2: "u2 \<le> 1" |
1607 |
using obt1 obt2 by auto |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1608 |
have "u1 * u + u2 * v \<le> max u1 u2 * u + max u1 u2 * v" |
78475 | 1609 |
by (smt (verit, ccfv_SIG) as mult_right_mono) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1610 |
also have "\<dots> \<le> 1" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1611 |
unfolding distrib_left[symmetric] and as(3) using u1 u2 by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1612 |
finally have le1: "u1 * u + u2 * v \<le> 1" . |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1613 |
show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1614 |
proof (intro CollectI exI conjI) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1615 |
show "0 \<le> u * u1 + v * u2" |
78475 | 1616 |
by (simp add: as obt1(1) obt2(1)) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1617 |
show "0 \<le> 1 - u * u1 - v * u2" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1618 |
by (simp add: le1 diff_diff_add mult.commute) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1619 |
qed (use b in \<open>auto simp: algebra_simps\<close>) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1620 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1621 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1622 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1623 |
lemma convex_hull_insert_alt: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1624 |
"convex hull (insert a S) = |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1625 |
(if S = {} then {a} |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1626 |
else {(1 - u) *\<^sub>R a + u *\<^sub>R x |x u. 0 \<le> u \<and> u \<le> 1 \<and> x \<in> convex hull S})" |
78475 | 1627 |
apply (simp add: convex_hull_insert) |
1628 |
using diff_add_cancel diff_ge_0_iff_ge |
|
1629 |
by (smt (verit, del_insts) Collect_cong) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1630 |
|
70136 | 1631 |
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Explicit expression for convex hull\<close> |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1632 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1633 |
proposition convex_hull_indexed: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1634 |
fixes S :: "'a::real_vector set" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1635 |
shows "convex hull S = |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1636 |
{y. \<exists>k u x. (\<forall>i\<in>{1::nat .. k}. 0 \<le> u i \<and> x i \<in> S) \<and> |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1637 |
(sum u {1..k} = 1) \<and> (\<Sum>i = 1..k. u i *\<^sub>R x i) = y}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1638 |
(is "?xyz = ?hull") |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1639 |
proof (rule hull_unique [OF _ convexI]) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1640 |
show "S \<subseteq> ?hull" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1641 |
by (clarsimp, rule_tac x=1 in exI, rule_tac x="\<lambda>x. 1" in exI, auto) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1642 |
next |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1643 |
fix T |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1644 |
assume "S \<subseteq> T" "convex T" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1645 |
then show "?hull \<subseteq> T" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1646 |
by (blast intro: convex_sum) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1647 |
next |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1648 |
fix x y u v |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1649 |
assume uv: "0 \<le> u" "0 \<le> v" "u + v = (1::real)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1650 |
assume xy: "x \<in> ?hull" "y \<in> ?hull" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1651 |
from xy obtain k1 u1 x1 where |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1652 |
x [rule_format]: "\<forall>i\<in>{1::nat..k1}. 0\<le>u1 i \<and> x1 i \<in> S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1653 |
"sum u1 {Suc 0..k1} = 1" "(\<Sum>i = Suc 0..k1. u1 i *\<^sub>R x1 i) = x" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1654 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1655 |
from xy obtain k2 u2 x2 where |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1656 |
y [rule_format]: "\<forall>i\<in>{1::nat..k2}. 0\<le>u2 i \<and> x2 i \<in> S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1657 |
"sum u2 {Suc 0..k2} = 1" "(\<Sum>i = Suc 0..k2. u2 i *\<^sub>R x2 i) = y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1658 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1659 |
have *: "\<And>P (x::'a) y s t i. (if P i then s else t) *\<^sub>R (if P i then x else y) = (if P i then s *\<^sub>R x else t *\<^sub>R y)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1660 |
"{1..k1 + k2} \<inter> {1..k1} = {1..k1}" "{1..k1 + k2} \<inter> - {1..k1} = (\<lambda>i. i + k1) ` {1..k2}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1661 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1662 |
have inj: "inj_on (\<lambda>i. i + k1) {1..k2}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1663 |
unfolding inj_on_def by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1664 |
let ?uu = "\<lambda>i. if i \<in> {1..k1} then u * u1 i else v * u2 (i - k1)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1665 |
let ?xx = "\<lambda>i. if i \<in> {1..k1} then x1 i else x2 (i - k1)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1666 |
show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1667 |
proof (intro CollectI exI conjI ballI) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1668 |
show "0 \<le> ?uu i" "?xx i \<in> S" if "i \<in> {1..k1+k2}" for i |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1669 |
using that by (auto simp add: le_diff_conv uv(1) x(1) uv(2) y(1)) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1670 |
show "(\<Sum>i = 1..k1 + k2. ?uu i) = 1" "(\<Sum>i = 1..k1 + k2. ?uu i *\<^sub>R ?xx i) = u *\<^sub>R x + v *\<^sub>R y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1671 |
unfolding * sum.If_cases[OF finite_atLeastAtMost[of 1 "k1 + k2"]] |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1672 |
sum.reindex[OF inj] Collect_mem_eq o_def |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1673 |
unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] sum_distrib_left[symmetric] |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1674 |
by (simp_all add: sum_distrib_left[symmetric] x(2,3) y(2,3) uv(3)) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1675 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1676 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1677 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1678 |
lemma convex_hull_finite: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1679 |
fixes S :: "'a::real_vector set" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1680 |
assumes "finite S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1681 |
shows "convex hull S = {y. \<exists>u. (\<forall>x\<in>S. 0 \<le> u x) \<and> sum u S = 1 \<and> sum (\<lambda>x. u x *\<^sub>R x) S = y}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1682 |
(is "?HULL = _") |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1683 |
proof (rule hull_unique [OF _ convexI]; clarify) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1684 |
fix x |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1685 |
assume "x \<in> S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1686 |
then show "\<exists>u. (\<forall>x\<in>S. 0 \<le> u x) \<and> sum u S = 1 \<and> (\<Sum>x\<in>S. u x *\<^sub>R x) = x" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1687 |
by (rule_tac x="\<lambda>y. if x=y then 1 else 0" in exI) (auto simp: sum.delta'[OF assms] sum_delta''[OF assms]) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1688 |
next |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1689 |
fix u v :: real |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1690 |
assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1691 |
fix ux assume ux [rule_format]: "\<forall>x\<in>S. 0 \<le> ux x" "sum ux S = (1::real)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1692 |
fix uy assume uy [rule_format]: "\<forall>x\<in>S. 0 \<le> uy x" "sum uy S = (1::real)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1693 |
have "0 \<le> u * ux x + v * uy x" if "x\<in>S" for x |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1694 |
by (simp add: that uv ux(1) uy(1)) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1695 |
moreover |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1696 |
have "(\<Sum>x\<in>S. u * ux x + v * uy x) = 1" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1697 |
unfolding sum.distrib and sum_distrib_left[symmetric] ux(2) uy(2) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1698 |
using uv(3) by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1699 |
moreover |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1700 |
have "(\<Sum>x\<in>S. (u * ux x + v * uy x) *\<^sub>R x) = u *\<^sub>R (\<Sum>x\<in>S. ux x *\<^sub>R x) + v *\<^sub>R (\<Sum>x\<in>S. uy x *\<^sub>R x)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1701 |
unfolding scaleR_left_distrib sum.distrib scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1702 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1703 |
ultimately |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1704 |
show "\<exists>uc. (\<forall>x\<in>S. 0 \<le> uc x) \<and> sum uc S = 1 \<and> |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1705 |
(\<Sum>x\<in>S. uc x *\<^sub>R x) = u *\<^sub>R (\<Sum>x\<in>S. ux x *\<^sub>R x) + v *\<^sub>R (\<Sum>x\<in>S. uy x *\<^sub>R x)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1706 |
by (rule_tac x="\<lambda>x. u * ux x + v * uy x" in exI, auto) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1707 |
qed (use assms in \<open>auto simp: convex_explicit\<close>) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1708 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1709 |
|
70136 | 1710 |
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Another formulation\<close> |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1711 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1712 |
text "Formalized by Lars Schewe." |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1713 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1714 |
lemma convex_hull_explicit: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1715 |
fixes p :: "'a::real_vector set" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1716 |
shows "convex hull p = |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1717 |
{y. \<exists>S u. finite S \<and> S \<subseteq> p \<and> (\<forall>x\<in>S. 0 \<le> u x) \<and> sum u S = 1 \<and> sum (\<lambda>v. u v *\<^sub>R v) S = y}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1718 |
(is "?lhs = ?rhs") |
78475 | 1719 |
proof (intro subset_antisym subsetI) |
1720 |
fix x |
|
1721 |
assume "x \<in> convex hull p" |
|
1722 |
then obtain k u y where |
|
1723 |
obt: "\<forall>i\<in>{1::nat..k}. 0 \<le> u i \<and> y i \<in> p" "sum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R y i) = x" |
|
1724 |
unfolding convex_hull_indexed by auto |
|
1725 |
have fin: "finite {1..k}" by auto |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1726 |
{ |
78475 | 1727 |
fix j |
1728 |
assume "j\<in>{1..k}" |
|
1729 |
then have "y j \<in> p \<and> 0 \<le> sum u {i. Suc 0 \<le> i \<and> i \<le> k \<and> y i = y j}" |
|
1730 |
by (metis (mono_tags, lifting) One_nat_def atLeastAtMost_iff mem_Collect_eq obt(1) sum_nonneg) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1731 |
} |
78475 | 1732 |
moreover have "(\<Sum>v\<in>y ` {1..k}. sum u {i \<in> {1..k}. y i = v}) = 1" |
1733 |
unfolding sum.image_gen[OF fin, symmetric] using obt(2) by auto |
|
1734 |
moreover have "(\<Sum>v\<in>y ` {1..k}. sum u {i \<in> {1..k}. y i = v} *\<^sub>R v) = x" |
|
1735 |
using sum.image_gen[OF fin, of "\<lambda>i. u i *\<^sub>R y i" y, symmetric] |
|
1736 |
unfolding scaleR_left.sum using obt(3) by auto |
|
1737 |
ultimately |
|
1738 |
have "\<exists>S u. finite S \<and> S \<subseteq> p \<and> (\<forall>x\<in>S. 0 \<le> u x) \<and> sum u S = 1 \<and> (\<Sum>v\<in>S. u v *\<^sub>R v) = x" |
|
1739 |
by (smt (verit, ccfv_SIG) imageE mem_Collect_eq obt(1) subsetI sum.cong sum.infinite sum_nonneg) |
|
1740 |
then show "x \<in> ?rhs" by auto |
|
1741 |
next |
|
1742 |
fix y |
|
1743 |
assume "y \<in> ?rhs" |
|
1744 |
then obtain S u where |
|
1745 |
S: "finite S" "S \<subseteq> p" "\<forall>x\<in>S. 0 \<le> u x" "sum u S = 1" "(\<Sum>v\<in>S. u v *\<^sub>R v) = y" |
|
1746 |
by auto |
|
1747 |
obtain f where f: "inj_on f {1..card S}" "f ` {1..card S} = S" |
|
1748 |
using ex_bij_betw_nat_finite_1[OF S(1)] unfolding bij_betw_def by auto |
|
1749 |
then have "0 \<le> u (f i)" "f i \<in> p" if "i \<in> {1..card S}" for i |
|
1750 |
using S \<open>i \<in> {1..card S}\<close> by blast+ |
|
1751 |
moreover |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1752 |
{ |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1753 |
fix y |
78475 | 1754 |
assume "y\<in>S" |
1755 |
then obtain i where "i\<in>{1..card S}" "f i = y" |
|
1756 |
by (metis f(2) image_iff) |
|
1757 |
then have "{x. Suc 0 \<le> x \<and> x \<le> card S \<and> f x = y} = {i}" |
|
1758 |
using f(1) inj_onD by fastforce |
|
1759 |
then have "(\<Sum>x\<in>{x \<in> {1..card S}. f x = y}. u (f x)) = u y" |
|
1760 |
"(\<Sum>x\<in>{x \<in> {1..card S}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y" |
|
1761 |
by (simp_all add: sum_constant_scaleR \<open>f i = y\<close>) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1762 |
} |
78475 | 1763 |
then have "(\<Sum>x = 1..card S. u (f x)) = 1" "(\<Sum>i = 1..card S. u (f i) *\<^sub>R f i) = y" |
1764 |
by (metis (mono_tags, lifting) S(4,5) f sum.reindex_cong)+ |
|
1765 |
ultimately |
|
1766 |
show "y \<in> convex hull p" |
|
1767 |
unfolding convex_hull_indexed |
|
1768 |
by (smt (verit, del_insts) mem_Collect_eq sum.cong) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1769 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1770 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1771 |
|
70136 | 1772 |
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>A stepping theorem for that expansion\<close> |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1773 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1774 |
lemma convex_hull_finite_step: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1775 |
fixes S :: "'a::real_vector set" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1776 |
assumes "finite S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1777 |
shows |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1778 |
"(\<exists>u. (\<forall>x\<in>insert a S. 0 \<le> u x) \<and> sum u (insert a S) = w \<and> sum (\<lambda>x. u x *\<^sub>R x) (insert a S) = y) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1779 |
\<longleftrightarrow> (\<exists>v\<ge>0. \<exists>u. (\<forall>x\<in>S. 0 \<le> u x) \<and> sum u S = w - v \<and> sum (\<lambda>x. u x *\<^sub>R x) S = y - v *\<^sub>R a)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1780 |
(is "?lhs = ?rhs") |
72385 | 1781 |
proof (cases "a \<in> S") |
1782 |
case True |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1783 |
then have *: "insert a S = S" by auto |
72385 | 1784 |
show ?thesis |
1785 |
proof |
|
1786 |
assume ?lhs |
|
1787 |
then show ?rhs |
|
1788 |
unfolding * by force |
|
1789 |
next |
|
1790 |
have fin: "finite (insert a S)" using assms by auto |
|
1791 |
assume ?rhs |
|
1792 |
then obtain v u where uv: "v\<ge>0" "\<forall>x\<in>S. 0 \<le> u x" "sum u S = w - v" "(\<Sum>x\<in>S. u x *\<^sub>R x) = y - v *\<^sub>R a" |
|
1793 |
by auto |
|
1794 |
then show ?lhs |
|
1795 |
using uv True assms |
|
1796 |
apply (rule_tac x = "\<lambda>x. (if a = x then v else 0) + u x" in exI) |
|
1797 |
apply (auto simp: sum_clauses scaleR_left_distrib sum.distrib sum_delta''[OF fin]) |
|
1798 |
done |
|
1799 |
qed |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1800 |
next |
72385 | 1801 |
case False |
1802 |
show ?thesis |
|
1803 |
proof |
|
1804 |
assume ?lhs |
|
1805 |
then obtain u where u: "\<forall>x\<in>insert a S. 0 \<le> u x" "sum u (insert a S) = w" "(\<Sum>x\<in>insert a S. u x *\<^sub>R x) = y" |
|
1806 |
by auto |
|
1807 |
then show ?rhs |
|
1808 |
using u \<open>a\<notin>S\<close> by (rule_tac x="u a" in exI) (auto simp: sum_clauses assms) |
|
1809 |
next |
|
1810 |
assume ?rhs |
|
1811 |
then obtain v u where uv: "v\<ge>0" "\<forall>x\<in>S. 0 \<le> u x" "sum u S = w - v" "(\<Sum>x\<in>S. u x *\<^sub>R x) = y - v *\<^sub>R a" |
|
1812 |
by auto |
|
1813 |
moreover |
|
1814 |
have "(\<Sum>x\<in>S. if a = x then v else u x) = sum u S" "(\<Sum>x\<in>S. (if a = x then v else u x) *\<^sub>R x) = (\<Sum>x\<in>S. u x *\<^sub>R x)" |
|
1815 |
using False by (auto intro!: sum.cong) |
|
1816 |
ultimately show ?lhs |
|
1817 |
using False by (rule_tac x="\<lambda>x. if a = x then v else u x" in exI) (auto simp: sum_clauses(2)[OF assms]) |
|
1818 |
qed |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1819 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1820 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1821 |
|
70136 | 1822 |
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Hence some special cases\<close> |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1823 |
|
72385 | 1824 |
lemma convex_hull_2: "convex hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b | u v. 0 \<le> u \<and> 0 \<le> v \<and> u + v = 1}" |
1825 |
(is "?lhs = ?rhs") |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1826 |
proof - |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1827 |
have **: "finite {b}" by auto |
72385 | 1828 |
have "\<And>x v u. \<lbrakk>0 \<le> v; v \<le> 1; (1 - v) *\<^sub>R b = x - v *\<^sub>R a\<rbrakk> |
1829 |
\<Longrightarrow> \<exists>u v. x = u *\<^sub>R a + v *\<^sub>R b \<and> 0 \<le> u \<and> 0 \<le> v \<and> u + v = 1" |
|
1830 |
by (metis add.commute diff_add_cancel diff_ge_0_iff_ge) |
|
1831 |
moreover |
|
1832 |
have "\<And>u v. \<lbrakk>0 \<le> u; 0 \<le> v; u + v = 1\<rbrakk> |
|
1833 |
\<Longrightarrow> \<exists>p\<ge>0. \<exists>q. 0 \<le> q b \<and> q b = 1 - p \<and> q b *\<^sub>R b = u *\<^sub>R a + v *\<^sub>R b - p *\<^sub>R a" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1834 |
apply (rule_tac x=u in exI, simp) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1835 |
apply (rule_tac x="\<lambda>x. v" in exI, simp) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1836 |
done |
72385 | 1837 |
ultimately show ?thesis |
1838 |
using convex_hull_finite_step[OF **, of a 1] |
|
1839 |
by (auto simp add: convex_hull_finite) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1840 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1841 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1842 |
lemma convex_hull_2_alt: "convex hull {a,b} = {a + u *\<^sub>R (b - a) | u. 0 \<le> u \<and> u \<le> 1}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1843 |
unfolding convex_hull_2 |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1844 |
proof (rule Collect_cong) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1845 |
have *: "\<And>x y ::real. x + y = 1 \<longleftrightarrow> x = 1 - y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1846 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1847 |
fix x |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1848 |
show "(\<exists>v u. x = v *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> v \<and> 0 \<le> u \<and> v + u = 1) \<longleftrightarrow> |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1849 |
(\<exists>u. x = a + u *\<^sub>R (b - a) \<and> 0 \<le> u \<and> u \<le> 1)" |
72385 | 1850 |
apply (simp add: *) |
1851 |
by (rule ex_cong1) (auto simp: algebra_simps) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1852 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1853 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1854 |
lemma convex_hull_3: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1855 |
"convex hull {a,b,c} = { u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c | u v w. 0 \<le> u \<and> 0 \<le> v \<and> 0 \<le> w \<and> u + v + w = 1}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1856 |
proof - |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1857 |
have fin: "finite {a,b,c}" "finite {b,c}" "finite {c}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1858 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1859 |
have *: "\<And>x y z ::real. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1860 |
by (auto simp: field_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1861 |
show ?thesis |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1862 |
unfolding convex_hull_finite[OF fin(1)] and convex_hull_finite_step[OF fin(2)] and * |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1863 |
unfolding convex_hull_finite_step[OF fin(3)] |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1864 |
apply (rule Collect_cong, simp) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1865 |
apply auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1866 |
apply (rule_tac x=va in exI) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1867 |
apply (rule_tac x="u c" in exI, simp) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1868 |
apply (rule_tac x="1 - v - w" in exI, simp) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1869 |
apply (rule_tac x=v in exI, simp) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1870 |
apply (rule_tac x="\<lambda>x. w" in exI, simp) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1871 |
done |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1872 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1873 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1874 |
lemma convex_hull_3_alt: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1875 |
"convex hull {a,b,c} = {a + u *\<^sub>R (b - a) + v *\<^sub>R (c - a) | u v. 0 \<le> u \<and> 0 \<le> v \<and> u + v \<le> 1}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1876 |
proof - |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1877 |
have *: "\<And>x y z ::real. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1878 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1879 |
show ?thesis |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1880 |
unfolding convex_hull_3 |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1881 |
apply (auto simp: *) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1882 |
apply (rule_tac x=v in exI) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1883 |
apply (rule_tac x=w in exI) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1884 |
apply (simp add: algebra_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1885 |
apply (rule_tac x=u in exI) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1886 |
apply (rule_tac x=v in exI) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1887 |
apply (simp add: algebra_simps) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1888 |
done |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1889 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1890 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1891 |
|
70136 | 1892 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>Relations among closure notions and corresponding hulls\<close> |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1893 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1894 |
lemma affine_imp_convex: "affine s \<Longrightarrow> convex s" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1895 |
unfolding affine_def convex_def by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1896 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1897 |
lemma convex_affine_hull [simp]: "convex (affine hull S)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1898 |
by (simp add: affine_imp_convex) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1899 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1900 |
lemma subspace_imp_convex: "subspace s \<Longrightarrow> convex s" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1901 |
using subspace_imp_affine affine_imp_convex by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1902 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1903 |
lemma convex_hull_subset_span: "(convex hull s) \<subseteq> (span s)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1904 |
by (metis hull_minimal span_superset subspace_imp_convex subspace_span) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1905 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1906 |
lemma convex_hull_subset_affine_hull: "(convex hull s) \<subseteq> (affine hull s)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1907 |
by (metis affine_affine_hull affine_imp_convex hull_minimal hull_subset) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1908 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1909 |
lemma aff_dim_convex_hull: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1910 |
fixes S :: "'n::euclidean_space set" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1911 |
shows "aff_dim (convex hull S) = aff_dim S" |
78475 | 1912 |
by (smt (verit) aff_dim_affine_hull aff_dim_subset convex_hull_subset_affine_hull hull_subset) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1913 |
|
71242 | 1914 |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1915 |
subsection \<open>Caratheodory's theorem\<close> |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1916 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1917 |
lemma convex_hull_caratheodory_aff_dim: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1918 |
fixes p :: "('a::euclidean_space) set" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1919 |
shows "convex hull p = |
72385 | 1920 |
{y. \<exists>S u. finite S \<and> S \<subseteq> p \<and> card S \<le> aff_dim p + 1 \<and> |
1921 |
(\<forall>x\<in>S. 0 \<le> u x) \<and> sum u S = 1 \<and> sum (\<lambda>v. u v *\<^sub>R v) S = y}" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1922 |
unfolding convex_hull_explicit set_eq_iff mem_Collect_eq |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1923 |
proof (intro allI iffI) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1924 |
fix y |
72385 | 1925 |
let ?P = "\<lambda>n. \<exists>S u. finite S \<and> card S = n \<and> S \<subseteq> p \<and> (\<forall>x\<in>S. 0 \<le> u x) \<and> |
1926 |
sum u S = 1 \<and> (\<Sum>v\<in>S. u v *\<^sub>R v) = y" |
|
1927 |
assume "\<exists>S u. finite S \<and> S \<subseteq> p \<and> (\<forall>x\<in>S. 0 \<le> u x) \<and> sum u S = 1 \<and> (\<Sum>v\<in>S. u v *\<^sub>R v) = y" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1928 |
then obtain N where "?P N" by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1929 |
then have "\<exists>n\<le>N. (\<forall>k<n. \<not> ?P k) \<and> ?P n" |
72385 | 1930 |
by (rule_tac ex_least_nat_le, auto) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1931 |
then obtain n where "?P n" and smallest: "\<forall>k<n. \<not> ?P k" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1932 |
by blast |
78475 | 1933 |
then obtain S u where S: "finite S" "card S = n" "S\<subseteq>p" |
1934 |
and u: "\<forall>x\<in>S. 0 \<le> u x" "sum u S = 1" "(\<Sum>v\<in>S. u v *\<^sub>R v) = y" by auto |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1935 |
|
72385 | 1936 |
have "card S \<le> aff_dim p + 1" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1937 |
proof (rule ccontr, simp only: not_le) |
72385 | 1938 |
assume "aff_dim p + 1 < card S" |
1939 |
then have "affine_dependent S" |
|
78475 | 1940 |
by (smt (verit) independent_card_le_aff_dim S(3)) |
72385 | 1941 |
then obtain w v where wv: "sum w S = 0" "v\<in>S" "w v \<noteq> 0" "(\<Sum>v\<in>S. w v *\<^sub>R v) = 0" |
78475 | 1942 |
using affine_dependent_explicit_finite[OF S(1)] by auto |
72385 | 1943 |
define i where "i = (\<lambda>v. (u v) / (- w v)) ` {v\<in>S. w v < 0}" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1944 |
define t where "t = Min i" |
72385 | 1945 |
have "\<exists>x\<in>S. w x < 0" |
78475 | 1946 |
by (smt (verit, best) S(1) sum_pos2 wv) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1947 |
then have "i \<noteq> {}" unfolding i_def by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1948 |
then have "t \<ge> 0" |
78475 | 1949 |
using Min_ge_iff[of i 0] and S(1) u[unfolded le_less] |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1950 |
unfolding t_def i_def |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1951 |
by (auto simp: divide_le_0_iff) |
72385 | 1952 |
have t: "\<forall>v\<in>S. u v + t * w v \<ge> 0" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1953 |
proof |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1954 |
fix v |
72385 | 1955 |
assume "v \<in> S" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1956 |
then have v: "0 \<le> u v" |
78475 | 1957 |
using u(1) by blast |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1958 |
show "0 \<le> u v + t * w v" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1959 |
proof (cases "w v < 0") |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1960 |
case False |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1961 |
thus ?thesis using v \<open>t\<ge>0\<close> by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1962 |
next |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1963 |
case True |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1964 |
then have "t \<le> u v / (- w v)" |
78475 | 1965 |
using \<open>v\<in>S\<close> S unfolding t_def i_def by (auto intro: Min_le) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1966 |
then show ?thesis |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1967 |
unfolding real_0_le_add_iff |
72385 | 1968 |
using True neg_le_minus_divide_eq by auto |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1969 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1970 |
qed |
72385 | 1971 |
obtain a where "a \<in> S" and "t = (\<lambda>v. (u v) / (- w v)) a" and "w a < 0" |
78475 | 1972 |
using Min_in[OF _ \<open>i\<noteq>{}\<close>] and S(1) unfolding i_def t_def by auto |
72385 | 1973 |
then have a: "a \<in> S" "u a + t * w a = 0" by auto |
1974 |
have *: "\<And>f. sum f (S - {a}) = sum f S - ((f a)::'b::ab_group_add)" |
|
78475 | 1975 |
unfolding sum.remove[OF S(1) \<open>a\<in>S\<close>] by auto |
72385 | 1976 |
have "(\<Sum>v\<in>S. u v + t * w v) = 1" |
78475 | 1977 |
by (metis add.right_neutral mult_zero_right sum.distrib sum_distrib_left u(2) wv(1)) |
72385 | 1978 |
moreover have "(\<Sum>v\<in>S. u v *\<^sub>R v + (t * w v) *\<^sub>R v) - (u a *\<^sub>R a + (t * w a) *\<^sub>R a) = y" |
78475 | 1979 |
unfolding sum.distrib u(3) scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] wv(4) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1980 |
using a(2) [THEN eq_neg_iff_add_eq_0 [THEN iffD2]] by simp |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1981 |
ultimately have "?P (n - 1)" |
72385 | 1982 |
apply (rule_tac x="(S - {a})" in exI) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1983 |
apply (rule_tac x="\<lambda>v. u v + t * w v" in exI) |
78475 | 1984 |
using S t a |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1985 |
apply (auto simp: * scaleR_left_distrib) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1986 |
done |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1987 |
then show False |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1988 |
using smallest[THEN spec[where x="n - 1"]] by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1989 |
qed |
72385 | 1990 |
then show "\<exists>S u. finite S \<and> S \<subseteq> p \<and> card S \<le> aff_dim p + 1 \<and> |
1991 |
(\<forall>x\<in>S. 0 \<le> u x) \<and> sum u S = 1 \<and> (\<Sum>v\<in>S. u v *\<^sub>R v) = y" |
|
78475 | 1992 |
using S u by auto |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1993 |
qed auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1994 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1995 |
lemma caratheodory_aff_dim: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1996 |
fixes p :: "('a::euclidean_space) set" |
72385 | 1997 |
shows "convex hull p = {x. \<exists>S. finite S \<and> S \<subseteq> p \<and> card S \<le> aff_dim p + 1 \<and> x \<in> convex hull S}" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1998 |
(is "?lhs = ?rhs") |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
1999 |
proof |
72385 | 2000 |
have "\<And>x S u. \<lbrakk>finite S; S \<subseteq> p; int (card S) \<le> aff_dim p + 1; \<forall>x\<in>S. 0 \<le> u x; sum u S = 1\<rbrakk> |
2001 |
\<Longrightarrow> (\<Sum>v\<in>S. u v *\<^sub>R v) \<in> convex hull S" |
|
78475 | 2002 |
by (metis (mono_tags, lifting) convex_convex_hull convex_explicit hull_subset) |
72385 | 2003 |
then show "?lhs \<subseteq> ?rhs" |
2004 |
by (subst convex_hull_caratheodory_aff_dim, auto) |
|
2005 |
qed (use hull_mono in auto) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2006 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2007 |
lemma convex_hull_caratheodory: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2008 |
fixes p :: "('a::euclidean_space) set" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2009 |
shows "convex hull p = |
72385 | 2010 |
{y. \<exists>S u. finite S \<and> S \<subseteq> p \<and> card S \<le> DIM('a) + 1 \<and> |
2011 |
(\<forall>x\<in>S. 0 \<le> u x) \<and> sum u S = 1 \<and> sum (\<lambda>v. u v *\<^sub>R v) S = y}" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2012 |
(is "?lhs = ?rhs") |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2013 |
proof (intro set_eqI iffI) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2014 |
fix x |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2015 |
assume "x \<in> ?lhs" then show "x \<in> ?rhs" |
72385 | 2016 |
unfolding convex_hull_caratheodory_aff_dim |
2017 |
using aff_dim_le_DIM [of p] by fastforce |
|
2018 |
qed (auto simp: convex_hull_explicit) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2019 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2020 |
theorem caratheodory: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2021 |
"convex hull p = |
72385 | 2022 |
{x::'a::euclidean_space. \<exists>S. finite S \<and> S \<subseteq> p \<and> card S \<le> DIM('a) + 1 \<and> x \<in> convex hull S}" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2023 |
proof safe |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2024 |
fix x |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2025 |
assume "x \<in> convex hull p" |
72385 | 2026 |
then obtain S u where "finite S" "S \<subseteq> p" "card S \<le> DIM('a) + 1" |
2027 |
"\<forall>x\<in>S. 0 \<le> u x" "sum u S = 1" "(\<Sum>v\<in>S. u v *\<^sub>R v) = x" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2028 |
unfolding convex_hull_caratheodory by auto |
72385 | 2029 |
then show "\<exists>S. finite S \<and> S \<subseteq> p \<and> card S \<le> DIM('a) + 1 \<and> x \<in> convex hull S" |
2030 |
using convex_hull_finite by fastforce |
|
2031 |
qed (use hull_mono in force) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2032 |
|
70136 | 2033 |
subsection\<^marker>\<open>tag unimportant\<close>\<open>Some Properties of subset of standard basis\<close> |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2034 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2035 |
lemma affine_hull_substd_basis: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2036 |
assumes "d \<subseteq> Basis" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2037 |
shows "affine hull (insert 0 d) = {x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2038 |
(is "affine hull (insert 0 ?A) = ?B") |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2039 |
proof - |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2040 |
have *: "\<And>A. (+) (0::'a) ` A = A" "\<And>A. (+) (- (0::'a)) ` A = A" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2041 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2042 |
show ?thesis |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2043 |
unfolding affine_hull_insert_span_gen span_substd_basis[OF assms,symmetric] * .. |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2044 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2045 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2046 |
lemma affine_hull_convex_hull [simp]: "affine hull (convex hull S) = affine hull S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2047 |
by (metis Int_absorb1 Int_absorb2 convex_hull_subset_affine_hull hull_hull hull_mono hull_subset) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2048 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2049 |
|
70136 | 2050 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>Moving and scaling convex hulls\<close> |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2051 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2052 |
lemma convex_hull_set_plus: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2053 |
"convex hull (S + T) = convex hull S + convex hull T" |
78475 | 2054 |
by (simp add: set_plus_image linear_iff scaleR_right_distrib convex_hull_Times |
2055 |
flip: convex_hull_linear_image) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2056 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2057 |
lemma translation_eq_singleton_plus: "(\<lambda>x. a + x) ` T = {a} + T" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2058 |
unfolding set_plus_def by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2059 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2060 |
lemma convex_hull_translation: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2061 |
"convex hull ((\<lambda>x. a + x) ` S) = (\<lambda>x. a + x) ` (convex hull S)" |
78475 | 2062 |
by (simp add: convex_hull_set_plus translation_eq_singleton_plus) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2063 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2064 |
lemma convex_hull_scaling: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2065 |
"convex hull ((\<lambda>x. c *\<^sub>R x) ` S) = (\<lambda>x. c *\<^sub>R x) ` (convex hull S)" |
78475 | 2066 |
by (simp add: convex_hull_linear_image) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2067 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2068 |
lemma convex_hull_affinity: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2069 |
"convex hull ((\<lambda>x. a + c *\<^sub>R x) ` S) = (\<lambda>x. a + c *\<^sub>R x) ` (convex hull S)" |
72385 | 2070 |
by (metis convex_hull_scaling convex_hull_translation image_image) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2071 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2072 |
|
70136 | 2073 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>Convexity of cone hulls\<close> |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2074 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2075 |
lemma convex_cone_hull: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2076 |
assumes "convex S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2077 |
shows "convex (cone hull S)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2078 |
proof (rule convexI) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2079 |
fix x y |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2080 |
assume xy: "x \<in> cone hull S" "y \<in> cone hull S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2081 |
then have "S \<noteq> {}" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2082 |
using cone_hull_empty_iff[of S] by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2083 |
fix u v :: real |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2084 |
assume uv: "u \<ge> 0" "v \<ge> 0" "u + v = 1" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2085 |
then have *: "u *\<^sub>R x \<in> cone hull S" "v *\<^sub>R y \<in> cone hull S" |
78475 | 2086 |
by (simp_all add: cone_cone_hull mem_cone uv xy) |
2087 |
then obtain cx :: real and xx |
|
2088 |
and cy :: real and yy where x: "u *\<^sub>R x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S" |
|
2089 |
and y: "v *\<^sub>R y = cy *\<^sub>R yy" "cy \<ge> 0" "yy \<in> S" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2090 |
using cone_hull_expl[of S] by auto |
78475 | 2091 |
|
2092 |
have "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S" if "cx + cy \<le> 0" |
|
2093 |
using "*"(1) nless_le that x(2) y by fastforce |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2094 |
moreover |
78475 | 2095 |
have "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S" if "cx + cy > 0" |
2096 |
proof - |
|
2097 |
have "(cx / (cx + cy)) *\<^sub>R xx + (cy / (cx + cy)) *\<^sub>R yy \<in> S" |
|
2098 |
using assms mem_convex_alt[of S xx yy cx cy] x y that by auto |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2099 |
then have "cx *\<^sub>R xx + cy *\<^sub>R yy \<in> cone hull S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2100 |
using mem_cone_hull[of "(cx/(cx+cy)) *\<^sub>R xx + (cy/(cx+cy)) *\<^sub>R yy" S "cx+cy"] \<open>cx+cy>0\<close> |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2101 |
by (auto simp: scaleR_right_distrib) |
78475 | 2102 |
then show ?thesis |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2103 |
using x y by auto |
78475 | 2104 |
qed |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2105 |
moreover have "cx + cy \<le> 0 \<or> cx + cy > 0" by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2106 |
ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S" by blast |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2107 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2108 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2109 |
lemma cone_convex_hull: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2110 |
assumes "cone S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2111 |
shows "cone (convex hull S)" |
78475 | 2112 |
by (metis (no_types, lifting) affine_hull_convex_hull affine_hull_eq_empty assms cone_iff convex_hull_scaling hull_inc) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2113 |
|
78656
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2114 |
section \<open>Conic sets and conic hull\<close> |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2115 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2116 |
definition conic :: "'a::real_vector set \<Rightarrow> bool" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2117 |
where "conic S \<equiv> \<forall>x c. x \<in> S \<longrightarrow> 0 \<le> c \<longrightarrow> (c *\<^sub>R x) \<in> S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2118 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2119 |
lemma conicD: "\<lbrakk>conic S; x \<in> S; 0 \<le> c\<rbrakk> \<Longrightarrow> (c *\<^sub>R x) \<in> S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2120 |
by (meson conic_def) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2121 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2122 |
lemma subspace_imp_conic: "subspace S \<Longrightarrow> conic S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2123 |
by (simp add: conic_def subspace_def) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2124 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2125 |
lemma conic_empty [simp]: "conic {}" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2126 |
using conic_def by blast |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2127 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2128 |
lemma conic_UNIV: "conic UNIV" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2129 |
by (simp add: conic_def) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2130 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2131 |
lemma conic_Inter: "(\<And>S. S \<in> \<F> \<Longrightarrow> conic S) \<Longrightarrow> conic(\<Inter>\<F>)" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2132 |
by (simp add: conic_def) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2133 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2134 |
lemma conic_linear_image: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2135 |
"\<lbrakk>conic S; linear f\<rbrakk> \<Longrightarrow> conic(f ` S)" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2136 |
by (smt (verit) conic_def image_iff linear.scaleR) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2137 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2138 |
lemma conic_linear_image_eq: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2139 |
"\<lbrakk>linear f; inj f\<rbrakk> \<Longrightarrow> conic (f ` S) \<longleftrightarrow> conic S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2140 |
by (smt (verit) conic_def conic_linear_image inj_image_mem_iff linear_cmul) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2141 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2142 |
lemma conic_mul: "\<lbrakk>conic S; x \<in> S; 0 \<le> c\<rbrakk> \<Longrightarrow> (c *\<^sub>R x) \<in> S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2143 |
using conic_def by blast |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2144 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2145 |
lemma conic_conic_hull: "conic(conic hull S)" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2146 |
by (metis (no_types, lifting) conic_Inter hull_def mem_Collect_eq) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2147 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2148 |
lemma conic_hull_eq: "(conic hull S = S) \<longleftrightarrow> conic S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2149 |
by (metis conic_conic_hull hull_same) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2150 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2151 |
lemma conic_hull_UNIV [simp]: "conic hull UNIV = UNIV" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2152 |
by simp |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2153 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2154 |
lemma conic_negations: "conic S \<Longrightarrow> conic (image uminus S)" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2155 |
by (auto simp: conic_def image_iff) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2156 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2157 |
lemma conic_span [iff]: "conic(span S)" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2158 |
by (simp add: subspace_imp_conic) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2159 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2160 |
lemma conic_hull_explicit: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2161 |
"conic hull S = {c *\<^sub>R x| c x. 0 \<le> c \<and> x \<in> S}" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2162 |
proof (rule hull_unique) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2163 |
show "S \<subseteq> {c *\<^sub>R x |c x. 0 \<le> c \<and> x \<in> S}" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2164 |
by (metis (no_types) cone_hull_expl hull_subset) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2165 |
show "conic {c *\<^sub>R x |c x. 0 \<le> c \<and> x \<in> S}" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2166 |
using mult_nonneg_nonneg by (force simp: conic_def) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2167 |
qed (auto simp: conic_def) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2168 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2169 |
lemma conic_hull_as_image: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2170 |
"conic hull S = (\<lambda>z. fst z *\<^sub>R snd z) ` ({0..} \<times> S)" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2171 |
by (force simp: conic_hull_explicit) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2172 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2173 |
lemma conic_hull_linear_image: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2174 |
"linear f \<Longrightarrow> conic hull f ` S = f ` (conic hull S)" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2175 |
by (force simp: conic_hull_explicit image_iff set_eq_iff linear_scale) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2176 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2177 |
lemma conic_hull_image_scale: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2178 |
assumes "\<And>x. x \<in> S \<Longrightarrow> 0 < c x" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2179 |
shows "conic hull (\<lambda>x. c x *\<^sub>R x) ` S = conic hull S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2180 |
proof |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2181 |
show "conic hull (\<lambda>x. c x *\<^sub>R x) ` S \<subseteq> conic hull S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2182 |
proof (rule hull_minimal) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2183 |
show "(\<lambda>x. c x *\<^sub>R x) ` S \<subseteq> conic hull S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2184 |
using assms conic_hull_explicit by fastforce |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2185 |
qed (simp add: conic_conic_hull) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2186 |
show "conic hull S \<subseteq> conic hull (\<lambda>x. c x *\<^sub>R x) ` S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2187 |
proof (rule hull_minimal) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2188 |
show "S \<subseteq> conic hull (\<lambda>x. c x *\<^sub>R x) ` S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2189 |
proof clarsimp |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2190 |
fix x |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2191 |
assume "x \<in> S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2192 |
then have "x = inverse(c x) *\<^sub>R c x *\<^sub>R x" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2193 |
using assms by fastforce |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2194 |
then show "x \<in> conic hull (\<lambda>x. c x *\<^sub>R x) ` S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2195 |
by (smt (verit, best) \<open>x \<in> S\<close> assms conic_conic_hull conic_mul hull_inc image_eqI inverse_nonpositive_iff_nonpositive) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2196 |
qed |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2197 |
qed (simp add: conic_conic_hull) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2198 |
qed |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2199 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2200 |
lemma convex_conic_hull: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2201 |
assumes "convex S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2202 |
shows "convex (conic hull S)" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2203 |
proof (clarsimp simp add: conic_hull_explicit convex_alt) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2204 |
fix c x d y and u :: real |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2205 |
assume \<section>: "(0::real) \<le> c" "x \<in> S" "(0::real) \<le> d" "y \<in> S" "0 \<le> u" "u \<le> 1" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2206 |
show "\<exists>c'' x''. ((1 - u) * c) *\<^sub>R x + (u * d) *\<^sub>R y = c'' *\<^sub>R x'' \<and> 0 \<le> c'' \<and> x'' \<in> S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2207 |
proof (cases "(1 - u) * c = 0") |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2208 |
case True |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2209 |
with \<open>0 \<le> d\<close> \<open>y \<in> S\<close>\<open>0 \<le> u\<close> |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2210 |
show ?thesis by force |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2211 |
next |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2212 |
case False |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2213 |
define \<xi> where "\<xi> \<equiv> (1 - u) * c + u * d" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2214 |
have *: "c * u \<le> c" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2215 |
by (simp add: "\<section>" mult_left_le) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2216 |
have "\<xi> > 0" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2217 |
using False \<section> by (smt (verit, best) \<xi>_def split_mult_pos_le) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2218 |
then have **: "c + d * u = \<xi> + c * u" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2219 |
by (simp add: \<xi>_def mult.commute right_diff_distrib') |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2220 |
show ?thesis |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2221 |
proof (intro exI conjI) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2222 |
show "0 \<le> \<xi>" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2223 |
using \<open>0 < \<xi>\<close> by auto |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2224 |
show "((1 - u) * c) *\<^sub>R x + (u * d) *\<^sub>R y = \<xi> *\<^sub>R (((1 - u) * c / \<xi>) *\<^sub>R x + (u * d / \<xi>) *\<^sub>R y)" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2225 |
using \<open>\<xi> > 0\<close> by (simp add: algebra_simps diff_divide_distrib) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2226 |
show "((1 - u) * c / \<xi>) *\<^sub>R x + (u * d / \<xi>) *\<^sub>R y \<in> S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2227 |
using \<open>0 < \<xi>\<close> |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2228 |
by (intro convexD [OF assms]) (auto simp: \<section> field_split_simps * **) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2229 |
qed |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2230 |
qed |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2231 |
qed |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2232 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2233 |
lemma conic_halfspace_le: "conic {x. a \<bullet> x \<le> 0}" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2234 |
by (auto simp: conic_def mult_le_0_iff) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2235 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2236 |
lemma conic_halfspace_ge: "conic {x. a \<bullet> x \<ge> 0}" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2237 |
by (auto simp: conic_def mult_le_0_iff) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2238 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2239 |
lemma conic_hull_empty [simp]: "conic hull {} = {}" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2240 |
by (simp add: conic_hull_eq) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2241 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2242 |
lemma conic_contains_0: "conic S \<Longrightarrow> (0 \<in> S \<longleftrightarrow> S \<noteq> {})" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2243 |
by (simp add: Convex.cone_def cone_contains_0 conic_def) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2244 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2245 |
lemma conic_hull_eq_empty: "conic hull S = {} \<longleftrightarrow> (S = {})" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2246 |
using conic_hull_explicit by fastforce |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2247 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2248 |
lemma conic_sums: "\<lbrakk>conic S; conic T\<rbrakk> \<Longrightarrow> conic (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2249 |
by (simp add: conic_def) (metis scaleR_right_distrib) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2250 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2251 |
lemma conic_Times: "\<lbrakk>conic S; conic T\<rbrakk> \<Longrightarrow> conic(S \<times> T)" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2252 |
by (auto simp: conic_def) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2253 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2254 |
lemma conic_Times_eq: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2255 |
"conic(S \<times> T) \<longleftrightarrow> S = {} \<or> T = {} \<or> conic S \<and> conic T" (is "?lhs = ?rhs") |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2256 |
proof |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2257 |
show "?lhs \<Longrightarrow> ?rhs" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2258 |
by (force simp: conic_def) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2259 |
show "?rhs \<Longrightarrow> ?lhs" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2260 |
by (force simp: conic_Times) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2261 |
qed |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2262 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2263 |
lemma conic_hull_0 [simp]: "conic hull {0} = {0}" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2264 |
by (simp add: conic_hull_eq subspace_imp_conic) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2265 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2266 |
lemma conic_hull_contains_0 [simp]: "0 \<in> conic hull S \<longleftrightarrow> (S \<noteq> {})" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2267 |
by (simp add: conic_conic_hull conic_contains_0 conic_hull_eq_empty) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2268 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2269 |
lemma conic_hull_eq_sing: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2270 |
"conic hull S = {x} \<longleftrightarrow> S = {0} \<and> x = 0" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2271 |
proof |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2272 |
show "conic hull S = {x} \<Longrightarrow> S = {0} \<and> x = 0" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2273 |
by (metis conic_conic_hull conic_contains_0 conic_def conic_hull_eq hull_inc insert_not_empty singleton_iff) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2274 |
qed simp |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2275 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2276 |
lemma conic_hull_Int_affine_hull: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2277 |
assumes "T \<subseteq> S" "0 \<notin> affine hull S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2278 |
shows "(conic hull T) \<inter> (affine hull S) = T" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2279 |
proof - |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2280 |
have TaffS: "T \<subseteq> affine hull S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2281 |
using \<open>T \<subseteq> S\<close> hull_subset by fastforce |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2282 |
moreover |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2283 |
have "conic hull T \<inter> affine hull S \<subseteq> T" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2284 |
proof (clarsimp simp: conic_hull_explicit) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2285 |
fix c x |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2286 |
assume "c *\<^sub>R x \<in> affine hull S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2287 |
and "0 \<le> c" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2288 |
and "x \<in> T" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2289 |
show "c *\<^sub>R x \<in> T" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2290 |
proof (cases "c=1") |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2291 |
case True |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2292 |
then show ?thesis |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2293 |
by (simp add: \<open>x \<in> T\<close>) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2294 |
next |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2295 |
case False |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2296 |
then have "x /\<^sub>R (1 - c) = x + (c * inverse (1 - c)) *\<^sub>R x" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2297 |
by (smt (verit, ccfv_SIG) diff_add_cancel mult.commute real_vector_affinity_eq scaleR_collapse scaleR_scaleR) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2298 |
then have "0 = inverse(1 - c) *\<^sub>R c *\<^sub>R x + (1 - inverse(1 - c)) *\<^sub>R x" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2299 |
by (simp add: algebra_simps) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2300 |
then have "0 \<in> affine hull S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2301 |
by (smt (verit) \<open>c *\<^sub>R x \<in> affine hull S\<close> \<open>x \<in> T\<close> affine_affine_hull TaffS in_mono mem_affine) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2302 |
then show ?thesis |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2303 |
using assms by auto |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2304 |
qed |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2305 |
qed |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2306 |
ultimately show ?thesis |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2307 |
by (auto simp: hull_inc) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2308 |
qed |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2309 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2310 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2311 |
section \<open>Convex cones and corresponding hulls\<close> |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2312 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2313 |
definition convex_cone :: "'a::real_vector set \<Rightarrow> bool" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2314 |
where "convex_cone \<equiv> \<lambda>S. S \<noteq> {} \<and> convex S \<and> conic S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2315 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2316 |
lemma convex_cone_iff: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2317 |
"convex_cone S \<longleftrightarrow> |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2318 |
0 \<in> S \<and> (\<forall>x \<in> S. \<forall>y \<in> S. x + y \<in> S) \<and> (\<forall>x \<in> S. \<forall>c\<ge>0. c *\<^sub>R x \<in> S)" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2319 |
by (metis cone_def conic_contains_0 conic_def convex_cone convex_cone_def) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2320 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2321 |
lemma convex_cone_add: "\<lbrakk>convex_cone S; x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> x+y \<in> S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2322 |
by (simp add: convex_cone_iff) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2323 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2324 |
lemma convex_cone_scaleR: "\<lbrakk>convex_cone S; 0 \<le> c; x \<in> S\<rbrakk> \<Longrightarrow> c *\<^sub>R x \<in> S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2325 |
by (simp add: convex_cone_iff) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2326 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2327 |
lemma convex_cone_nonempty: "convex_cone S \<Longrightarrow> S \<noteq> {}" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2328 |
by (simp add: convex_cone_def) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2329 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2330 |
lemma convex_cone_linear_image: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2331 |
"convex_cone S \<and> linear f \<Longrightarrow> convex_cone(f ` S)" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2332 |
by (simp add: conic_linear_image convex_cone_def convex_linear_image) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2333 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2334 |
lemma convex_cone_linear_image_eq: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2335 |
"\<lbrakk>linear f; inj f\<rbrakk> \<Longrightarrow> (convex_cone(f ` S) \<longleftrightarrow> convex_cone S)" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2336 |
by (simp add: conic_linear_image_eq convex_cone_def) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2337 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2338 |
lemma convex_cone_halfspace_ge: "convex_cone {x. a \<bullet> x \<ge> 0}" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2339 |
by (simp add: convex_cone_iff inner_simps(2)) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2340 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2341 |
lemma convex_cone_halfspace_le: "convex_cone {x. a \<bullet> x \<le> 0}" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2342 |
by (simp add: convex_cone_iff inner_right_distrib mult_nonneg_nonpos) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2343 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2344 |
lemma convex_cone_contains_0: "convex_cone S \<Longrightarrow> 0 \<in> S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2345 |
using convex_cone_iff by blast |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2346 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2347 |
lemma convex_cone_Inter: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2348 |
"(\<And>S. S \<in> f \<Longrightarrow> convex_cone S) \<Longrightarrow> convex_cone(\<Inter> f)" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2349 |
by (simp add: convex_cone_iff) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2350 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2351 |
lemma convex_cone_convex_cone_hull: "convex_cone(convex_cone hull S)" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2352 |
by (metis (no_types, lifting) convex_cone_Inter hull_def mem_Collect_eq) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2353 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2354 |
lemma convex_convex_cone_hull: "convex(convex_cone hull S)" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2355 |
by (meson convex_cone_convex_cone_hull convex_cone_def) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2356 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2357 |
lemma conic_convex_cone_hull: "conic(convex_cone hull S)" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2358 |
by (metis convex_cone_convex_cone_hull convex_cone_def) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2359 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2360 |
lemma convex_cone_hull_nonempty: "convex_cone hull S \<noteq> {}" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2361 |
by (simp add: convex_cone_convex_cone_hull convex_cone_nonempty) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2362 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2363 |
lemma convex_cone_hull_contains_0: "0 \<in> convex_cone hull S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2364 |
by (simp add: convex_cone_contains_0 convex_cone_convex_cone_hull) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2365 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2366 |
lemma convex_cone_hull_add: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2367 |
"\<lbrakk>x \<in> convex_cone hull S; y \<in> convex_cone hull S\<rbrakk> \<Longrightarrow> x + y \<in> convex_cone hull S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2368 |
by (simp add: convex_cone_add convex_cone_convex_cone_hull) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2369 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2370 |
lemma convex_cone_hull_mul: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2371 |
"\<lbrakk>x \<in> convex_cone hull S; 0 \<le> c\<rbrakk> \<Longrightarrow> (c *\<^sub>R x) \<in> convex_cone hull S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2372 |
by (simp add: conic_convex_cone_hull conic_mul) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2373 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2374 |
thm convex_sums |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2375 |
lemma convex_cone_sums: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2376 |
"\<lbrakk>convex_cone S; convex_cone T\<rbrakk> \<Longrightarrow> convex_cone (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2377 |
by (simp add: convex_cone_def conic_sums convex_sums) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2378 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2379 |
lemma convex_cone_Times: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2380 |
"\<lbrakk>convex_cone S; convex_cone T\<rbrakk> \<Longrightarrow> convex_cone(S \<times> T)" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2381 |
by (simp add: conic_Times convex_Times convex_cone_def) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2382 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2383 |
lemma convex_cone_Times_D1: "convex_cone (S \<times> T) \<Longrightarrow> convex_cone S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2384 |
by (metis Times_empty conic_Times_eq convex_cone_def convex_convex_hull convex_hull_Times hull_same times_eq_iff) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2385 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2386 |
lemma convex_cone_Times_eq: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2387 |
"convex_cone(S \<times> T) \<longleftrightarrow> convex_cone S \<and> convex_cone T" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2388 |
proof (cases "S={} \<or> T={}") |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2389 |
case True |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2390 |
then show ?thesis |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2391 |
by (auto dest: convex_cone_nonempty) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2392 |
next |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2393 |
case False |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2394 |
then have "convex_cone (S \<times> T) \<Longrightarrow> convex_cone T" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2395 |
by (metis conic_Times_eq convex_cone_def convex_convex_hull convex_hull_Times hull_same times_eq_iff) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2396 |
then show ?thesis |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2397 |
using convex_cone_Times convex_cone_Times_D1 by blast |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2398 |
qed |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2399 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2400 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2401 |
lemma convex_cone_hull_Un: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2402 |
"convex_cone hull(S \<union> T) = (\<Union>x \<in> convex_cone hull S. \<Union>y \<in> convex_cone hull T. {x + y})" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2403 |
(is "?lhs = ?rhs") |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2404 |
proof |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2405 |
show "?lhs \<subseteq> ?rhs" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2406 |
proof (rule hull_minimal) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2407 |
show "S \<union> T \<subseteq> (\<Union>x\<in>convex_cone hull S. \<Union>y\<in>convex_cone hull T. {x + y})" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2408 |
apply (clarsimp simp: subset_iff) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2409 |
by (metis add_0 convex_cone_hull_contains_0 group_cancel.rule0 hull_inc) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2410 |
show "convex_cone (\<Union>x\<in>convex_cone hull S. \<Union>y\<in>convex_cone hull T. {x + y})" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2411 |
by (simp add: convex_cone_convex_cone_hull convex_cone_sums) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2412 |
qed |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2413 |
next |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2414 |
show "?rhs \<subseteq> ?lhs" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2415 |
by clarify (metis convex_cone_hull_add hull_mono le_sup_iff subsetD subsetI) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2416 |
qed |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2417 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2418 |
lemma convex_cone_singleton [iff]: "convex_cone {0}" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2419 |
by (simp add: convex_cone_iff) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2420 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2421 |
lemma convex_hull_subset_convex_cone_hull: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2422 |
"convex hull S \<subseteq> convex_cone hull S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2423 |
by (simp add: convex_convex_cone_hull hull_minimal hull_subset) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2424 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2425 |
lemma conic_hull_subset_convex_cone_hull: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2426 |
"conic hull S \<subseteq> convex_cone hull S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2427 |
by (simp add: conic_convex_cone_hull hull_minimal hull_subset) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2428 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2429 |
lemma subspace_imp_convex_cone: "subspace S \<Longrightarrow> convex_cone S" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2430 |
by (simp add: convex_cone_iff subspace_def) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2431 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2432 |
lemma convex_cone_span: "convex_cone(span S)" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2433 |
by (simp add: subspace_imp_convex_cone) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2434 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2435 |
lemma convex_cone_negations: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2436 |
"convex_cone S \<Longrightarrow> convex_cone (image uminus S)" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2437 |
by (simp add: convex_cone_linear_image module_hom_uminus) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2438 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2439 |
lemma subspace_convex_cone_symmetric: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2440 |
"subspace S \<longleftrightarrow> convex_cone S \<and> (\<forall>x \<in> S. -x \<in> S)" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2441 |
by (smt (verit) convex_cone_iff scaleR_left.minus subspace_def subspace_neg) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2442 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2443 |
lemma convex_cone_hull_separate_nonempty: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2444 |
assumes "S \<noteq> {}" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2445 |
shows "convex_cone hull S = conic hull (convex hull S)" (is "?lhs = ?rhs") |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2446 |
proof |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2447 |
show "?lhs \<subseteq> ?rhs" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2448 |
by (metis assms conic_conic_hull convex_cone_def convex_conic_hull convex_convex_hull hull_subset subset_empty subset_hull) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2449 |
show "?rhs \<subseteq> ?lhs" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2450 |
by (simp add: conic_convex_cone_hull convex_hull_subset_convex_cone_hull subset_hull) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2451 |
qed |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2452 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2453 |
lemma convex_cone_hull_empty [simp]: "convex_cone hull {} = {0}" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2454 |
by (metis convex_cone_hull_contains_0 convex_cone_singleton hull_redundant hull_same) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2455 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2456 |
lemma convex_cone_hull_separate: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2457 |
"convex_cone hull S = insert 0 (conic hull (convex hull S))" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2458 |
proof(cases "S={}") |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2459 |
case False |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2460 |
then show ?thesis |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2461 |
using convex_cone_hull_contains_0 convex_cone_hull_separate_nonempty by blast |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2462 |
qed auto |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2463 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2464 |
lemma convex_cone_hull_convex_hull_nonempty: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2465 |
"S \<noteq> {} \<Longrightarrow> convex_cone hull S = (\<Union>x \<in> convex hull S. \<Union>c\<in>{0..}. {c *\<^sub>R x})" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2466 |
by (force simp: convex_cone_hull_separate_nonempty conic_hull_as_image) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2467 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2468 |
lemma convex_cone_hull_convex_hull: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2469 |
"convex_cone hull S = insert 0 (\<Union>x \<in> convex hull S. \<Union>c\<in>{0..}. {c *\<^sub>R x})" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2470 |
by (force simp: convex_cone_hull_separate conic_hull_as_image) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2471 |
|
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2472 |
lemma convex_cone_hull_linear_image: |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2473 |
"linear f \<Longrightarrow> convex_cone hull (f ` S) = image f (convex_cone hull S)" |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2474 |
by (metis (no_types, lifting) conic_hull_linear_image convex_cone_hull_separate convex_hull_linear_image image_insert linear_0) |
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents:
78475
diff
changeset
|
2475 |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2476 |
subsection \<open>Radon's theorem\<close> |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2477 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2478 |
text "Formalized by Lars Schewe." |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2479 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2480 |
lemma Radon_ex_lemma: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2481 |
assumes "finite c" "affine_dependent c" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2482 |
shows "\<exists>u. sum u c = 0 \<and> (\<exists>v\<in>c. u v \<noteq> 0) \<and> sum (\<lambda>v. u v *\<^sub>R v) c = 0" |
78475 | 2483 |
using affine_dependent_explicit_finite assms by blast |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2484 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2485 |
lemma Radon_s_lemma: |
72385 | 2486 |
assumes "finite S" |
2487 |
and "sum f S = (0::real)" |
|
2488 |
shows "sum f {x\<in>S. 0 < f x} = - sum f {x\<in>S. f x < 0}" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2489 |
proof - |
78475 | 2490 |
have "\<And>x. (if f x < 0 then f x else 0) + (if 0 < f x then f x else 0) = f x" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2491 |
by auto |
78475 | 2492 |
then show ?thesis |
2493 |
using assms by (simp add: sum.inter_filter flip: sum.distrib add_eq_0_iff) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2494 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2495 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2496 |
lemma Radon_v_lemma: |
72385 | 2497 |
assumes "finite S" |
2498 |
and "sum f S = 0" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2499 |
and "\<forall>x. g x = (0::real) \<longrightarrow> f x = (0::'a::euclidean_space)" |
72385 | 2500 |
shows "(sum f {x\<in>S. 0 < g x}) = - sum f {x\<in>S. g x < 0}" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2501 |
proof - |
78475 | 2502 |
have "\<And>x. (if 0 < g x then f x else 0) + (if g x < 0 then f x else 0) = f x" |
2503 |
using assms by auto |
|
2504 |
then show ?thesis |
|
2505 |
using assms by (simp add: sum.inter_filter eq_neg_iff_add_eq_0 flip: sum.distrib add_eq_0_iff) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2506 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2507 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2508 |
lemma Radon_partition: |
72385 | 2509 |
assumes "finite C" "affine_dependent C" |
78475 | 2510 |
shows "\<exists>M P. M \<inter> P = {} \<and> M \<union> P = C \<and> (convex hull M) \<inter> (convex hull P) \<noteq> {}" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2511 |
proof - |
72385 | 2512 |
obtain u v where uv: "sum u C = 0" "v\<in>C" "u v \<noteq> 0" "(\<Sum>v\<in>C. u v *\<^sub>R v) = 0" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2513 |
using Radon_ex_lemma[OF assms] by auto |
72385 | 2514 |
have fin: "finite {x \<in> C. 0 < u x}" "finite {x \<in> C. 0 > u x}" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2515 |
using assms(1) by auto |
72385 | 2516 |
define z where "z = inverse (sum u {x\<in>C. u x > 0}) *\<^sub>R sum (\<lambda>x. u x *\<^sub>R x) {x\<in>C. u x > 0}" |
2517 |
have "sum u {x \<in> C. 0 < u x} \<noteq> 0" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2518 |
proof (cases "u v \<ge> 0") |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2519 |
case False |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2520 |
then have "u v < 0" by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2521 |
then show ?thesis |
78475 | 2522 |
by (smt (verit) assms(1) fin(1) mem_Collect_eq sum.neutral_const sum_mono_inv uv) |
2523 |
next |
|
2524 |
case True |
|
2525 |
with fin uv show "sum u {x \<in> C. 0 < u x} \<noteq> 0" |
|
2526 |
by (smt (verit) fin(1) mem_Collect_eq sum_nonneg_eq_0_iff uv) |
|
2527 |
qed |
|
72385 | 2528 |
then have *: "sum u {x\<in>C. u x > 0} > 0" |
2529 |
unfolding less_le by (metis (no_types, lifting) mem_Collect_eq sum_nonneg) |
|
2530 |
moreover have "sum u ({x \<in> C. 0 < u x} \<union> {x \<in> C. u x < 0}) = sum u C" |
|
2531 |
"(\<Sum>x\<in>{x \<in> C. 0 < u x} \<union> {x \<in> C. u x < 0}. u x *\<^sub>R x) = (\<Sum>x\<in>C. u x *\<^sub>R x)" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2532 |
using assms(1) |
72385 | 2533 |
by (rule_tac[!] sum.mono_neutral_left, auto) |
2534 |
then have "sum u {x \<in> C. 0 < u x} = - sum u {x \<in> C. 0 > u x}" |
|
2535 |
"(\<Sum>x\<in>{x \<in> C. 0 < u x}. u x *\<^sub>R x) = - (\<Sum>x\<in>{x \<in> C. 0 > u x}. u x *\<^sub>R x)" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2536 |
unfolding eq_neg_iff_add_eq_0 |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2537 |
using uv(1,4) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2538 |
by (auto simp: sum.union_inter_neutral[OF fin, symmetric]) |
72385 | 2539 |
moreover have "\<forall>x\<in>{v \<in> C. u v < 0}. 0 \<le> inverse (sum u {x \<in> C. 0 < u x}) * - u x" |
2540 |
using * by (fastforce intro: mult_nonneg_nonneg) |
|
2541 |
ultimately have "z \<in> convex hull {v \<in> C. u v \<le> 0}" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2542 |
unfolding convex_hull_explicit mem_Collect_eq |
72385 | 2543 |
apply (rule_tac x="{v \<in> C. u v < 0}" in exI) |
2544 |
apply (rule_tac x="\<lambda>y. inverse (sum u {x\<in>C. u x > 0}) * - u y" in exI) |
|
2545 |
using assms(1) unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] |
|
2546 |
by (auto simp: z_def sum_negf sum_distrib_left[symmetric]) |
|
2547 |
moreover have "\<forall>x\<in>{v \<in> C. 0 < u v}. 0 \<le> inverse (sum u {x \<in> C. 0 < u x}) * u x" |
|
2548 |
using * by (fastforce intro: mult_nonneg_nonneg) |
|
2549 |
then have "z \<in> convex hull {v \<in> C. u v > 0}" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2550 |
unfolding convex_hull_explicit mem_Collect_eq |
72385 | 2551 |
apply (rule_tac x="{v \<in> C. 0 < u v}" in exI) |
2552 |
apply (rule_tac x="\<lambda>y. inverse (sum u {x\<in>C. u x > 0}) * u y" in exI) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2553 |
using assms(1) |
72385 | 2554 |
unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] |
2555 |
using * by (auto simp: z_def sum_negf sum_distrib_left[symmetric]) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2556 |
ultimately show ?thesis |
72385 | 2557 |
apply (rule_tac x="{v\<in>C. u v \<le> 0}" in exI) |
2558 |
apply (rule_tac x="{v\<in>C. u v > 0}" in exI, auto) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2559 |
done |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2560 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2561 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2562 |
theorem Radon: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2563 |
assumes "affine_dependent c" |
78475 | 2564 |
obtains M P where "M \<subseteq> c" "P \<subseteq> c" "M \<inter> P = {}" "(convex hull M) \<inter> (convex hull P) \<noteq> {}" |
2565 |
by (smt (verit) Radon_partition affine_dependent_explicit affine_dependent_explicit_finite assms le_sup_iff) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2566 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2567 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2568 |
subsection \<open>Helly's theorem\<close> |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2569 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2570 |
lemma Helly_induct: |
78475 | 2571 |
fixes \<F> :: "'a::euclidean_space set set" |
2572 |
assumes "card \<F> = n" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2573 |
and "n \<ge> DIM('a) + 1" |
78475 | 2574 |
and "\<forall>S\<in>\<F>. convex S" "\<forall>T\<subseteq>\<F>. card T = DIM('a) + 1 \<longrightarrow> \<Inter>T \<noteq> {}" |
2575 |
shows "\<Inter>\<F> \<noteq> {}" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2576 |
using assms |
78475 | 2577 |
proof (induction n arbitrary: \<F>) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2578 |
case 0 |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2579 |
then show ?case by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2580 |
next |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2581 |
case (Suc n) |
78475 | 2582 |
have "finite \<F>" |
2583 |
using \<open>card \<F> = Suc n\<close> by (auto intro: card_ge_0_finite) |
|
2584 |
show "\<Inter>\<F> \<noteq> {}" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2585 |
proof (cases "n = DIM('a)") |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2586 |
case True |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2587 |
then show ?thesis |
78475 | 2588 |
by (simp add: Suc.prems) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2589 |
next |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2590 |
case False |
78475 | 2591 |
have "\<Inter>(\<F> - {S}) \<noteq> {}" if "S \<in> \<F>" for S |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2592 |
proof (rule Suc.IH[rule_format]) |
78475 | 2593 |
show "card (\<F> - {S}) = n" |
2594 |
by (simp add: Suc.prems(1) \<open>finite \<F>\<close> that) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2595 |
show "DIM('a) + 1 \<le> n" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2596 |
using False Suc.prems(2) by linarith |
78475 | 2597 |
show "\<And>t. \<lbrakk>t \<subseteq> \<F> - {S}; card t = DIM('a) + 1\<rbrakk> \<Longrightarrow> \<Inter>t \<noteq> {}" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2598 |
by (simp add: Suc.prems(4) subset_Diff_insert) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2599 |
qed (use Suc in auto) |
78475 | 2600 |
then have "\<forall>S\<in>\<F>. \<exists>x. x \<in> \<Inter>(\<F> - {S})" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2601 |
by blast |
78475 | 2602 |
then obtain X where X: "\<And>S. S\<in>\<F> \<Longrightarrow> X S \<in> \<Inter>(\<F> - {S})" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2603 |
by metis |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2604 |
show ?thesis |
78475 | 2605 |
proof (cases "inj_on X \<F>") |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2606 |
case False |
78475 | 2607 |
then obtain S T where "S\<noteq>T" and st: "S\<in>\<F>" "T\<in>\<F>" "X S = X T" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2608 |
unfolding inj_on_def by auto |
78475 | 2609 |
then have *: "\<Inter>\<F> = \<Inter>(\<F> - {S}) \<inter> \<Inter>(\<F> - {T})" by auto |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2610 |
show ?thesis |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2611 |
by (metis "*" X disjoint_iff_not_equal st) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2612 |
next |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2613 |
case True |
78475 | 2614 |
then obtain M P where mp: "M \<inter> P = {}" "M \<union> P = X ` \<F>" "convex hull M \<inter> convex hull P \<noteq> {}" |
2615 |
using Radon_partition[of "X ` \<F>"] and affine_dependent_biggerset[of "X ` \<F>"] |
|
2616 |
unfolding card_image[OF True] and \<open>card \<F> = Suc n\<close> |
|
2617 |
using Suc(3) \<open>finite \<F>\<close> and False |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2618 |
by auto |
78475 | 2619 |
have "M \<subseteq> X ` \<F>" "P \<subseteq> X ` \<F>" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2620 |
using mp(2) by auto |
78475 | 2621 |
then obtain \<G> \<H> where gh:"M = X ` \<G>" "P = X ` \<H>" "\<G> \<subseteq> \<F>" "\<H> \<subseteq> \<F>" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2622 |
unfolding subset_image_iff by auto |
78475 | 2623 |
then have "\<F> \<union> (\<G> \<union> \<H>) = \<F>" by auto |
2624 |
then have \<F>: "\<F> = \<G> \<union> \<H>" |
|
2625 |
using inj_on_Un_image_eq_iff[of X \<F> "\<G> \<union> \<H>"] and True |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2626 |
unfolding mp(2)[unfolded image_Un[symmetric] gh] |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2627 |
by auto |
78475 | 2628 |
have *: "\<G> \<inter> \<H> = {}" |
2629 |
using gh local.mp(1) by blast |
|
2630 |
have "convex hull (X ` \<H>) \<subseteq> \<Inter>\<G>" "convex hull (X ` \<G>) \<subseteq> \<Inter>\<H>" |
|
2631 |
by (rule hull_minimal; use X * \<F> in \<open>auto simp: Suc.prems(3) convex_Inter\<close>)+ |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2632 |
then show ?thesis |
78475 | 2633 |
unfolding \<F> using mp(3)[unfolded gh] by blast |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2634 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2635 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2636 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2637 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2638 |
theorem Helly: |
78475 | 2639 |
fixes \<F> :: "'a::euclidean_space set set" |
2640 |
assumes "card \<F> \<ge> DIM('a) + 1" "\<forall>s\<in>\<F>. convex s" |
|
2641 |
and "\<And>t. \<lbrakk>t\<subseteq>\<F>; card t = DIM('a) + 1\<rbrakk> \<Longrightarrow> \<Inter>t \<noteq> {}" |
|
2642 |
shows "\<Inter>\<F> \<noteq> {}" |
|
72385 | 2643 |
using Helly_induct assms by blast |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2644 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2645 |
subsection \<open>Epigraphs of convex functions\<close> |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2646 |
|
70136 | 2647 |
definition\<^marker>\<open>tag important\<close> "epigraph S (f :: _ \<Rightarrow> real) = {xy. fst xy \<in> S \<and> f (fst xy) \<le> snd xy}" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2648 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2649 |
lemma mem_epigraph: "(x, y) \<in> epigraph S f \<longleftrightarrow> x \<in> S \<and> f x \<le> y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2650 |
unfolding epigraph_def by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2651 |
|
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
2652 |
lemma convex_epigraph: "convex (epigraph S f) \<longleftrightarrow> convex_on S f" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2653 |
proof safe |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2654 |
assume L: "convex (epigraph S f)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2655 |
then show "convex_on S f" |
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
2656 |
by (fastforce simp: convex_def convex_on_def epigraph_def) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2657 |
next |
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
2658 |
assume "convex_on S f" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2659 |
then show "convex (epigraph S f)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2660 |
unfolding convex_def convex_on_def epigraph_def |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2661 |
apply safe |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2662 |
apply (rule_tac [2] y="u * f a + v * f aa" in order_trans) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2663 |
apply (auto intro!:mult_left_mono add_mono) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2664 |
done |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2665 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2666 |
|
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
2667 |
lemma convex_epigraphI: "convex_on S f \<Longrightarrow> convex (epigraph S f)" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2668 |
unfolding convex_epigraph by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2669 |
|
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
2670 |
lemma convex_epigraph_convex: "convex_on S f \<longleftrightarrow> convex(epigraph S f)" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2671 |
by (simp add: convex_epigraph) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2672 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2673 |
|
70136 | 2674 |
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Use this to derive general bound property of convex function\<close> |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2675 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2676 |
lemma convex_on: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2677 |
assumes "convex S" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2678 |
shows "convex_on S f \<longleftrightarrow> |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2679 |
(\<forall>k u x. (\<forall>i\<in>{1..k::nat}. 0 \<le> u i \<and> x i \<in> S) \<and> sum u {1..k} = 1 \<longrightarrow> |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2680 |
f (sum (\<lambda>i. u i *\<^sub>R x i) {1..k}) \<le> sum (\<lambda>i. u i * f(x i)) {1..k})" |
72385 | 2681 |
(is "?lhs = (\<forall>k u x. ?rhs k u x)") |
2682 |
proof |
|
2683 |
assume ?lhs |
|
2684 |
then have \<section>: "convex {xy. fst xy \<in> S \<and> f (fst xy) \<le> snd xy}" |
|
2685 |
by (metis assms convex_epigraph epigraph_def) |
|
2686 |
show "\<forall>k u x. ?rhs k u x" |
|
2687 |
proof (intro allI) |
|
2688 |
fix k u x |
|
2689 |
show "?rhs k u x" |
|
2690 |
using \<section> |
|
2691 |
unfolding convex mem_Collect_eq fst_sum snd_sum |
|
2692 |
apply safe |
|
2693 |
apply (drule_tac x=k in spec) |
|
2694 |
apply (drule_tac x=u in spec) |
|
2695 |
apply (drule_tac x="\<lambda>i. (x i, f (x i))" in spec) |
|
2696 |
apply simp |
|
2697 |
done |
|
2698 |
qed |
|
2699 |
next |
|
2700 |
assume "\<forall>k u x. ?rhs k u x" |
|
2701 |
then show ?lhs |
|
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
79532
diff
changeset
|
2702 |
unfolding convex_epigraph_convex convex epigraph_def Ball_def mem_Collect_eq fst_sum snd_sum |
72385 | 2703 |
using assms[unfolded convex] apply clarsimp |
2704 |
apply (rule_tac y="\<Sum>i = 1..k. u i * f (fst (x i))" in order_trans) |
|
2705 |
by (auto simp add: mult_left_mono intro: sum_mono) |
|
2706 |
qed |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2707 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2708 |
|
70136 | 2709 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>A bound within a convex hull\<close> |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2710 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2711 |
lemma convex_on_convex_hull_bound: |
72385 | 2712 |
assumes "convex_on (convex hull S) f" |
2713 |
and "\<forall>x\<in>S. f x \<le> b" |
|
2714 |
shows "\<forall>x\<in> convex hull S. f x \<le> b" |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2715 |
proof |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2716 |
fix x |
72385 | 2717 |
assume "x \<in> convex hull S" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2718 |
then obtain k u v where |
72385 | 2719 |
u: "\<forall>i\<in>{1..k::nat}. 0 \<le> u i \<and> v i \<in> S" "sum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R v i) = x" |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2720 |
unfolding convex_hull_indexed mem_Collect_eq by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2721 |
have "(\<Sum>i = 1..k. u i * f (v i)) \<le> b" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2722 |
using sum_mono[of "{1..k}" "\<lambda>i. u i * f (v i)" "\<lambda>i. u i * b"] |
72385 | 2723 |
unfolding sum_distrib_right[symmetric] u(2) mult_1 |
2724 |
using assms(2) mult_left_mono u(1) by blast |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2725 |
then show "f x \<le> b" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2726 |
using assms(1)[unfolded convex_on[OF convex_convex_hull], rule_format, of k u v] |
72385 | 2727 |
using hull_inc u by fastforce |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2728 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2729 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2730 |
lemma convex_set_plus: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2731 |
assumes "convex S" and "convex T" shows "convex (S + T)" |
78475 | 2732 |
by (metis assms convex_hull_eq convex_hull_set_plus) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2733 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2734 |
lemma convex_set_sum: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2735 |
assumes "\<And>i. i \<in> A \<Longrightarrow> convex (B i)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2736 |
shows "convex (\<Sum>i\<in>A. B i)" |
78475 | 2737 |
using assms |
2738 |
by (induction A rule: infinite_finite_induct) (auto simp: convex_set_plus) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2739 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2740 |
lemma finite_set_sum: |
78475 | 2741 |
assumes "\<forall>i\<in>A. finite (B i)" shows "finite (\<Sum>i\<in>A. B i)" |
2742 |
using assms |
|
2743 |
by (induction A rule: infinite_finite_induct) (auto simp: finite_set_plus) |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2744 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2745 |
lemma box_eq_set_sum_Basis: |
72385 | 2746 |
"{x. \<forall>i\<in>Basis. x\<bullet>i \<in> B i} = (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` (B i))" (is "?lhs = ?rhs") |
2747 |
proof - |
|
2748 |
have "\<And>x. \<forall>i\<in>Basis. x \<bullet> i \<in> B i \<Longrightarrow> |
|
2749 |
\<exists>s. x = sum s Basis \<and> (\<forall>i\<in>Basis. s i \<in> (\<lambda>x. x *\<^sub>R i) ` B i)" |
|
2750 |
by (metis (mono_tags, lifting) euclidean_representation image_iff) |
|
2751 |
moreover |
|
2752 |
have "sum f Basis \<bullet> i \<in> B i" if "i \<in> Basis" and f: "\<forall>i\<in>Basis. f i \<in> (\<lambda>x. x *\<^sub>R i) ` B i" for i f |
|
2753 |
proof - |
|
2754 |
have "(\<Sum>x\<in>Basis - {i}. f x \<bullet> i) = 0" |
|
78475 | 2755 |
proof (intro strip sum.neutral) |
72385 | 2756 |
show "f x \<bullet> i = 0" if "x \<in> Basis - {i}" for x |
2757 |
using that f \<open>i \<in> Basis\<close> inner_Basis that by fastforce |
|
2758 |
qed |
|
2759 |
then have "(\<Sum>x\<in>Basis. f x \<bullet> i) = f i \<bullet> i" |
|
2760 |
by (metis (no_types) \<open>i \<in> Basis\<close> add.right_neutral sum.remove [OF finite_Basis]) |
|
2761 |
then have "(\<Sum>x\<in>Basis. f x \<bullet> i) \<in> B i" |
|
2762 |
using f that(1) by auto |
|
2763 |
then show ?thesis |
|
2764 |
by (simp add: inner_sum_left) |
|
2765 |
qed |
|
2766 |
ultimately show ?thesis |
|
2767 |
by (subst set_sum_alt [OF finite_Basis]) auto |
|
2768 |
qed |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2769 |
|
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2770 |
lemma convex_hull_set_sum: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2771 |
"convex hull (\<Sum>i\<in>A. B i) = (\<Sum>i\<in>A. convex hull (B i))" |
78475 | 2772 |
by (induction A rule: infinite_finite_induct) (auto simp: convex_hull_set_plus) |
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
diff
changeset
|
2773 |
|
79532
bb5d036f3523
Type class patch suggested by Achim Brucker, plus tidied lemma
paulson <lp15@cam.ac.uk>
parents:
78656
diff
changeset
|
2774 |
end |