| author | wenzelm | 
| Sun, 11 Aug 2019 22:36:34 +0200 | |
| changeset 70503 | f0b2635ee17f | 
| parent 70136 | f03a01a18c6e | 
| child 70802 | 160eaf566bcb | 
| permissions | -rw-r--r-- | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 1 | (* Title: HOL/Analysis/Convex_Euclidean_Space.thy | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 2 | Author: L C Paulson, University of Cambridge | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 3 | Author: Robert Himmelmann, TU Muenchen | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 4 | Author: Bogdan Grechuk, University of Edinburgh | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 5 | Author: Armin Heller, TU Muenchen | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 6 | Author: Johannes Hoelzl, TU Muenchen | 
| 33175 | 7 | *) | 
| 8 | ||
| 69619 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 9 | section \<open>Convex Sets and Functions on (Normed) Euclidean Spaces\<close> | 
| 33175 | 10 | |
| 11 | theory Convex_Euclidean_Space | |
| 44132 | 12 | imports | 
| 69619 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 13 | Convex | 
| 69617 | 14 | Topology_Euclidean_Space | 
| 33175 | 15 | begin | 
| 16 | ||
| 70136 | 17 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Topological Properties of Convex Sets and Functions\<close> | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 18 | |
| 64267 | 19 | lemma convex_supp_sum: | 
| 20 | assumes "convex S" and 1: "supp_sum u I = 1" | |
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 21 | and "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> u i \<and> (u i = 0 \<or> f i \<in> S)" | 
| 64267 | 22 | shows "supp_sum (\<lambda>i. u i *\<^sub>R f i) I \<in> S" | 
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 23 | proof - | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 24 |   have fin: "finite {i \<in> I. u i \<noteq> 0}"
 | 
| 64267 | 25 | using 1 sum.infinite by (force simp: supp_sum_def support_on_def) | 
| 26 |   then have eq: "supp_sum (\<lambda>i. u i *\<^sub>R f i) I = sum (\<lambda>i. u i *\<^sub>R f i) {i \<in> I. u i \<noteq> 0}"
 | |
| 27 | by (force intro: sum.mono_neutral_left simp: supp_sum_def support_on_def) | |
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 28 | show ?thesis | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 29 | apply (simp add: eq) | 
| 64267 | 30 | apply (rule convex_sum [OF fin \<open>convex S\<close>]) | 
| 31 | using 1 assms apply (auto simp: supp_sum_def support_on_def) | |
| 63969 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 32 | done | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 33 | qed | 
| 
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
 hoelzl parents: 
63967diff
changeset | 34 | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 35 | lemma closure_bounded_linear_image_subset: | 
| 44524 | 36 | assumes f: "bounded_linear f" | 
| 53333 | 37 | shows "f ` closure S \<subseteq> closure (f ` S)" | 
| 44524 | 38 | using linear_continuous_on [OF f] closed_closure closure_subset | 
| 39 | by (rule image_closure_subset) | |
| 40 | ||
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 41 | lemma closure_linear_image_subset: | 
| 53339 | 42 | fixes f :: "'m::euclidean_space \<Rightarrow> 'n::real_normed_vector" | 
| 49529 | 43 | assumes "linear f" | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 44 | shows "f ` (closure S) \<subseteq> closure (f ` S)" | 
| 44524 | 45 | using assms unfolding linear_conv_bounded_linear | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 46 | by (rule closure_bounded_linear_image_subset) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 47 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 48 | lemma closed_injective_linear_image: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 49 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 50 | assumes S: "closed S" and f: "linear f" "inj f" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 51 | shows "closed (f ` S)" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 52 | proof - | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 53 | obtain g where g: "linear g" "g \<circ> f = id" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 54 | using linear_injective_left_inverse [OF f] by blast | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 55 | then have confg: "continuous_on (range f) g" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 56 | using linear_continuous_on linear_conv_bounded_linear by blast | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 57 | have [simp]: "g ` f ` S = S" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 58 | using g by (simp add: image_comp) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 59 | have cgf: "closed (g ` f ` S)" | 
| 61808 | 60 | by (simp add: \<open>g \<circ> f = id\<close> S image_comp) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 61 | have [simp]: "(range f \<inter> g -` S) = f ` S" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 62 | using g unfolding o_def id_def image_def by auto metis+ | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 63 | show ?thesis | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 64 | proof (rule closedin_closed_trans [of "range f"]) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69710diff
changeset | 65 | show "closedin (top_of_set (range f)) (f ` S)" | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 66 | using continuous_closedin_preimage [OF confg cgf] by simp | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 67 | show "closed (range f)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 68 | apply (rule closed_injective_image_subspace) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 69 | using f apply (auto simp: linear_linear linear_injective_0) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 70 | done | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 71 | qed | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 72 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 73 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 74 | lemma closed_injective_linear_image_eq: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 75 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 76 | assumes f: "linear f" "inj f" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 77 | shows "(closed(image f s) \<longleftrightarrow> closed s)" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 78 | by (metis closed_injective_linear_image closure_eq closure_linear_image_subset closure_subset_eq f(1) f(2) inj_image_subset_iff) | 
| 40377 | 79 | |
| 80 | lemma closure_injective_linear_image: | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 81 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 82 | shows "\<lbrakk>linear f; inj f\<rbrakk> \<Longrightarrow> f ` (closure S) = closure (f ` S)" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 83 | apply (rule subset_antisym) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 84 | apply (simp add: closure_linear_image_subset) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 85 | by (simp add: closure_minimal closed_injective_linear_image closure_subset image_mono) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 86 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 87 | lemma closure_bounded_linear_image: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 88 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 89 | shows "\<lbrakk>linear f; bounded S\<rbrakk> \<Longrightarrow> f ` (closure S) = closure (f ` S)" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 90 | apply (rule subset_antisym, simp add: closure_linear_image_subset) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 91 | apply (rule closure_minimal, simp add: closure_subset image_mono) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 92 | by (meson bounded_closure closed_closure compact_continuous_image compact_eq_bounded_closed linear_continuous_on linear_conv_bounded_linear) | 
| 40377 | 93 | |
| 44524 | 94 | lemma closure_scaleR: | 
| 53339 | 95 | fixes S :: "'a::real_normed_vector set" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68607diff
changeset | 96 | shows "((*\<^sub>R) c) ` (closure S) = closure (((*\<^sub>R) c) ` S)" | 
| 44524 | 97 | proof | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68607diff
changeset | 98 | show "((*\<^sub>R) c) ` (closure S) \<subseteq> closure (((*\<^sub>R) c) ` S)" | 
| 53333 | 99 | using bounded_linear_scaleR_right | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 100 | by (rule closure_bounded_linear_image_subset) | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68607diff
changeset | 101 | show "closure (((*\<^sub>R) c) ` S) \<subseteq> ((*\<^sub>R) c) ` (closure S)" | 
| 49529 | 102 | by (intro closure_minimal image_mono closure_subset closed_scaling closed_closure) | 
| 103 | qed | |
| 104 | ||
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 105 | lemma sphere_eq_empty [simp]: | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 106 |   fixes a :: "'a::{real_normed_vector, perfect_space}"
 | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 107 |   shows "sphere a r = {} \<longleftrightarrow> r < 0"
 | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 108 | by (auto simp: sphere_def dist_norm) (metis dist_norm le_less_linear vector_choose_dist) | 
| 49529 | 109 | |
| 40377 | 110 | lemma cone_closure: | 
| 53347 | 111 | fixes S :: "'a::real_normed_vector set" | 
| 49529 | 112 | assumes "cone S" | 
| 113 | shows "cone (closure S)" | |
| 114 | proof (cases "S = {}")
 | |
| 115 | case True | |
| 116 | then show ?thesis by auto | |
| 117 | next | |
| 118 | case False | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68607diff
changeset | 119 | then have "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (*\<^sub>R) c ` S = S)" | 
| 49529 | 120 | using cone_iff[of S] assms by auto | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68607diff
changeset | 121 | then have "0 \<in> closure S \<and> (\<forall>c. c > 0 \<longrightarrow> (*\<^sub>R) c ` closure S = closure S)" | 
| 68031 | 122 | using closure_subset by (auto simp: closure_scaleR) | 
| 53339 | 123 | then show ?thesis | 
| 60974 
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
 paulson <lp15@cam.ac.uk> parents: 
60809diff
changeset | 124 | using False cone_iff[of "closure S"] by auto | 
| 49529 | 125 | qed | 
| 126 | ||
| 66939 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 127 | corollary component_complement_connected: | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 128 | fixes S :: "'a::real_normed_vector set" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 129 | assumes "connected S" "C \<in> components (-S)" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 130 | shows "connected(-C)" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 131 | using component_diff_connected [of S UNIV] assms | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 132 | by (auto simp: Compl_eq_Diff_UNIV) | 
| 33175 | 133 | |
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62097diff
changeset | 134 | proposition clopen: | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 135 | fixes S :: "'a :: real_normed_vector set" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 136 |   shows "closed S \<and> open S \<longleftrightarrow> S = {} \<or> S = UNIV"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 137 | by (force intro!: connected_UNIV [unfolded connected_clopen, rule_format]) | 
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62097diff
changeset | 138 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62097diff
changeset | 139 | corollary compact_open: | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 140 | fixes S :: "'a :: euclidean_space set" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 141 |   shows "compact S \<and> open S \<longleftrightarrow> S = {}"
 | 
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62097diff
changeset | 142 | by (auto simp: compact_eq_bounded_closed clopen) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62097diff
changeset | 143 | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 144 | corollary finite_imp_not_open: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 145 |     fixes S :: "'a::{real_normed_vector, perfect_space} set"
 | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 146 |     shows "\<lbrakk>finite S; open S\<rbrakk> \<Longrightarrow> S={}"
 | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 147 | using clopen [of S] finite_imp_closed not_bounded_UNIV by blast | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 148 | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 149 | corollary empty_interior_finite: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 150 |     fixes S :: "'a::{real_normed_vector, perfect_space} set"
 | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 151 |     shows "finite S \<Longrightarrow> interior S = {}"
 | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 152 | by (metis interior_subset finite_subset open_interior [of S] finite_imp_not_open) | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 153 | |
| 60420 | 154 | text \<open>Balls, being convex, are connected.\<close> | 
| 33175 | 155 | |
| 156 | lemma convex_local_global_minimum: | |
| 157 | fixes s :: "'a::real_normed_vector set" | |
| 53347 | 158 | assumes "e > 0" | 
| 159 | and "convex_on s f" | |
| 160 | and "ball x e \<subseteq> s" | |
| 161 | and "\<forall>y\<in>ball x e. f x \<le> f y" | |
| 33175 | 162 | shows "\<forall>y\<in>s. f x \<le> f y" | 
| 53302 | 163 | proof (rule ccontr) | 
| 164 | have "x \<in> s" using assms(1,3) by auto | |
| 165 | assume "\<not> ?thesis" | |
| 166 | then obtain y where "y\<in>s" and y: "f x > f y" by auto | |
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61952diff
changeset | 167 | then have xy: "0 < dist x y" by auto | 
| 53347 | 168 | then obtain u where "0 < u" "u \<le> 1" and u: "u < e / dist x y" | 
| 68527 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
68074diff
changeset | 169 | using field_lbound_gt_zero[of 1 "e / dist x y"] xy \<open>e>0\<close> by auto | 
| 53302 | 170 | then have "f ((1-u) *\<^sub>R x + u *\<^sub>R y) \<le> (1-u) * f x + u * f y" | 
| 60420 | 171 | using \<open>x\<in>s\<close> \<open>y\<in>s\<close> | 
| 53302 | 172 | using assms(2)[unfolded convex_on_def, | 
| 173 | THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x="1-u"]] | |
| 50804 | 174 | by auto | 
| 33175 | 175 | moreover | 
| 50804 | 176 | have *: "x - ((1 - u) *\<^sub>R x + u *\<^sub>R y) = u *\<^sub>R (x - y)" | 
| 177 | by (simp add: algebra_simps) | |
| 178 | have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> ball x e" | |
| 53302 | 179 | unfolding mem_ball dist_norm | 
| 60420 | 180 | unfolding * and norm_scaleR and abs_of_pos[OF \<open>0<u\<close>] | 
| 50804 | 181 | unfolding dist_norm[symmetric] | 
| 53302 | 182 | using u | 
| 183 | unfolding pos_less_divide_eq[OF xy] | |
| 184 | by auto | |
| 185 | then have "f x \<le> f ((1 - u) *\<^sub>R x + u *\<^sub>R y)" | |
| 186 | using assms(4) by auto | |
| 50804 | 187 | ultimately show False | 
| 60420 | 188 | using mult_strict_left_mono[OF y \<open>u>0\<close>] | 
| 53302 | 189 | unfolding left_diff_distrib | 
| 190 | by auto | |
| 33175 | 191 | qed | 
| 192 | ||
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 193 | lemma convex_ball [iff]: | 
| 33175 | 194 | fixes x :: "'a::real_normed_vector" | 
| 49531 | 195 | shows "convex (ball x e)" | 
| 68031 | 196 | proof (auto simp: convex_def) | 
| 50804 | 197 | fix y z | 
| 198 | assume yz: "dist x y < e" "dist x z < e" | |
| 199 | fix u v :: real | |
| 200 | assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1" | |
| 201 | have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z" | |
| 202 | using uv yz | |
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 203 | using convex_on_dist [of "ball x e" x, unfolded convex_on_def, | 
| 53302 | 204 | THEN bspec[where x=y], THEN bspec[where x=z]] | 
| 50804 | 205 | by auto | 
| 206 | then show "dist x (u *\<^sub>R y + v *\<^sub>R z) < e" | |
| 207 | using convex_bound_lt[OF yz uv] by auto | |
| 33175 | 208 | qed | 
| 209 | ||
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 210 | lemma convex_cball [iff]: | 
| 33175 | 211 | fixes x :: "'a::real_normed_vector" | 
| 53347 | 212 | shows "convex (cball x e)" | 
| 213 | proof - | |
| 214 |   {
 | |
| 215 | fix y z | |
| 216 | assume yz: "dist x y \<le> e" "dist x z \<le> e" | |
| 217 | fix u v :: real | |
| 218 | assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1" | |
| 219 | have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z" | |
| 220 | using uv yz | |
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 221 | using convex_on_dist [of "cball x e" x, unfolded convex_on_def, | 
| 53347 | 222 | THEN bspec[where x=y], THEN bspec[where x=z]] | 
| 223 | by auto | |
| 224 | then have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> e" | |
| 225 | using convex_bound_le[OF yz uv] by auto | |
| 226 | } | |
| 68031 | 227 | then show ?thesis by (auto simp: convex_def Ball_def) | 
| 33175 | 228 | qed | 
| 229 | ||
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 230 | lemma connected_ball [iff]: | 
| 33175 | 231 | fixes x :: "'a::real_normed_vector" | 
| 232 | shows "connected (ball x e)" | |
| 233 | using convex_connected convex_ball by auto | |
| 234 | ||
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 235 | lemma connected_cball [iff]: | 
| 33175 | 236 | fixes x :: "'a::real_normed_vector" | 
| 53302 | 237 | shows "connected (cball x e)" | 
| 33175 | 238 | using convex_connected convex_cball by auto | 
| 239 | ||
| 50804 | 240 | |
| 33175 | 241 | lemma bounded_convex_hull: | 
| 242 | fixes s :: "'a::real_normed_vector set" | |
| 53347 | 243 | assumes "bounded s" | 
| 244 | shows "bounded (convex hull s)" | |
| 50804 | 245 | proof - | 
| 246 | from assms obtain B where B: "\<forall>x\<in>s. norm x \<le> B" | |
| 247 | unfolding bounded_iff by auto | |
| 248 | show ?thesis | |
| 249 | apply (rule bounded_subset[OF bounded_cball, of _ 0 B]) | |
| 44170 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 huffman parents: 
44142diff
changeset | 250 | unfolding subset_hull[of convex, OF convex_cball] | 
| 53302 | 251 | unfolding subset_eq mem_cball dist_norm using B | 
| 252 | apply auto | |
| 50804 | 253 | done | 
| 254 | qed | |
| 33175 | 255 | |
| 256 | lemma finite_imp_bounded_convex_hull: | |
| 257 | fixes s :: "'a::real_normed_vector set" | |
| 53302 | 258 | shows "finite s \<Longrightarrow> bounded (convex hull s)" | 
| 259 | using bounded_convex_hull finite_imp_bounded | |
| 260 | by auto | |
| 33175 | 261 | |
| 40377 | 262 | lemma aff_dim_cball: | 
| 53347 | 263 | fixes a :: "'n::euclidean_space" | 
| 264 | assumes "e > 0" | |
| 265 |   shows "aff_dim (cball a e) = int (DIM('n))"
 | |
| 266 | proof - | |
| 267 | have "(\<lambda>x. a + x) ` (cball 0 e) \<subseteq> cball a e" | |
| 268 | unfolding cball_def dist_norm by auto | |
| 269 | then have "aff_dim (cball (0 :: 'n::euclidean_space) e) \<le> aff_dim (cball a e)" | |
| 270 | using aff_dim_translation_eq[of a "cball 0 e"] | |
| 67399 | 271 | aff_dim_subset[of "(+) a ` cball 0 e" "cball a e"] | 
| 53347 | 272 | by auto | 
| 273 |   moreover have "aff_dim (cball (0 :: 'n::euclidean_space) e) = int (DIM('n))"
 | |
| 274 | using hull_inc[of "(0 :: 'n::euclidean_space)" "cball 0 e"] | |
| 275 | centre_in_cball[of "(0 :: 'n::euclidean_space)"] assms | |
| 276 | by (simp add: dim_cball[of e] aff_dim_zero[of "cball 0 e"]) | |
| 277 | ultimately show ?thesis | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 278 | using aff_dim_le_DIM[of "cball a e"] by auto | 
| 40377 | 279 | qed | 
| 280 | ||
| 281 | lemma aff_dim_open: | |
| 53347 | 282 | fixes S :: "'n::euclidean_space set" | 
| 283 | assumes "open S" | |
| 284 |     and "S \<noteq> {}"
 | |
| 285 |   shows "aff_dim S = int (DIM('n))"
 | |
| 286 | proof - | |
| 287 | obtain x where "x \<in> S" | |
| 288 | using assms by auto | |
| 289 | then obtain e where e: "e > 0" "cball x e \<subseteq> S" | |
| 290 | using open_contains_cball[of S] assms by auto | |
| 291 | then have "aff_dim (cball x e) \<le> aff_dim S" | |
| 292 | using aff_dim_subset by auto | |
| 293 | with e show ?thesis | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 294 | using aff_dim_cball[of e x] aff_dim_le_DIM[of S] by auto | 
| 40377 | 295 | qed | 
| 296 | ||
| 297 | lemma low_dim_interior: | |
| 53347 | 298 | fixes S :: "'n::euclidean_space set" | 
| 299 |   assumes "\<not> aff_dim S = int (DIM('n))"
 | |
| 300 |   shows "interior S = {}"
 | |
| 301 | proof - | |
| 302 | have "aff_dim(interior S) \<le> aff_dim S" | |
| 303 | using interior_subset aff_dim_subset[of "interior S" S] by auto | |
| 304 | then show ?thesis | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 305 | using aff_dim_open[of "interior S"] aff_dim_le_DIM[of S] assms by auto | 
| 40377 | 306 | qed | 
| 307 | ||
| 60307 
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
 paulson <lp15@cam.ac.uk> parents: 
60303diff
changeset | 308 | corollary empty_interior_lowdim: | 
| 
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
 paulson <lp15@cam.ac.uk> parents: 
60303diff
changeset | 309 | fixes S :: "'n::euclidean_space set" | 
| 
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
 paulson <lp15@cam.ac.uk> parents: 
60303diff
changeset | 310 |   shows "dim S < DIM ('n) \<Longrightarrow> interior S = {}"
 | 
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 311 | by (metis low_dim_interior affine_hull_UNIV dim_affine_hull less_not_refl dim_UNIV) | 
| 60307 
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
 paulson <lp15@cam.ac.uk> parents: 
60303diff
changeset | 312 | |
| 63016 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 313 | corollary aff_dim_nonempty_interior: | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 314 | fixes S :: "'a::euclidean_space set" | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 315 |   shows "interior S \<noteq> {} \<Longrightarrow> aff_dim S = DIM('a)"
 | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 316 | by (metis low_dim_interior) | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 317 | |
| 63881 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 318 | |
| 60420 | 319 | subsection \<open>Relative interior of a set\<close> | 
| 40377 | 320 | |
| 70136 | 321 | definition\<^marker>\<open>tag important\<close> "rel_interior S = | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69710diff
changeset | 322 |   {x. \<exists>T. openin (top_of_set (affine hull S)) T \<and> x \<in> T \<and> T \<subseteq> S}"
 | 
| 53347 | 323 | |
| 64287 | 324 | lemma rel_interior_mono: | 
| 325 | "\<lbrakk>S \<subseteq> T; affine hull S = affine hull T\<rbrakk> | |
| 326 | \<Longrightarrow> (rel_interior S) \<subseteq> (rel_interior T)" | |
| 327 | by (auto simp: rel_interior_def) | |
| 328 | ||
| 329 | lemma rel_interior_maximal: | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69710diff
changeset | 330 | "\<lbrakk>T \<subseteq> S; openin(top_of_set (affine hull S)) T\<rbrakk> \<Longrightarrow> T \<subseteq> (rel_interior S)" | 
| 64287 | 331 | by (auto simp: rel_interior_def) | 
| 332 | ||
| 53347 | 333 | lemma rel_interior: | 
| 334 |   "rel_interior S = {x \<in> S. \<exists>T. open T \<and> x \<in> T \<and> T \<inter> affine hull S \<subseteq> S}"
 | |
| 335 | unfolding rel_interior_def[of S] openin_open[of "affine hull S"] | |
| 336 | apply auto | |
| 337 | proof - | |
| 338 | fix x T | |
| 339 | assume *: "x \<in> S" "open T" "x \<in> T" "T \<inter> affine hull S \<subseteq> S" | |
| 340 | then have **: "x \<in> T \<inter> affine hull S" | |
| 341 | using hull_inc by auto | |
| 54465 | 342 | show "\<exists>Tb. (\<exists>Ta. open Ta \<and> Tb = affine hull S \<inter> Ta) \<and> x \<in> Tb \<and> Tb \<subseteq> S" | 
| 343 | apply (rule_tac x = "T \<inter> (affine hull S)" in exI) | |
| 53347 | 344 | using * ** | 
| 345 | apply auto | |
| 346 | done | |
| 347 | qed | |
| 348 | ||
| 349 | lemma mem_rel_interior: "x \<in> rel_interior S \<longleftrightarrow> (\<exists>T. open T \<and> x \<in> T \<inter> S \<and> T \<inter> affine hull S \<subseteq> S)" | |
| 68031 | 350 | by (auto simp: rel_interior) | 
| 53347 | 351 | |
| 352 | lemma mem_rel_interior_ball: | |
| 353 | "x \<in> rel_interior S \<longleftrightarrow> x \<in> S \<and> (\<exists>e. e > 0 \<and> ball x e \<inter> affine hull S \<subseteq> S)" | |
| 40377 | 354 | apply (simp add: rel_interior, safe) | 
| 68031 | 355 | apply (force simp: open_contains_ball) | 
| 356 | apply (rule_tac x = "ball x e" in exI, simp) | |
| 40377 | 357 | done | 
| 358 | ||
| 49531 | 359 | lemma rel_interior_ball: | 
| 53347 | 360 |   "rel_interior S = {x \<in> S. \<exists>e. e > 0 \<and> ball x e \<inter> affine hull S \<subseteq> S}"
 | 
| 361 | using mem_rel_interior_ball [of _ S] by auto | |
| 362 | ||
| 363 | lemma mem_rel_interior_cball: | |
| 364 | "x \<in> rel_interior S \<longleftrightarrow> x \<in> S \<and> (\<exists>e. e > 0 \<and> cball x e \<inter> affine hull S \<subseteq> S)" | |
| 49531 | 365 | apply (simp add: rel_interior, safe) | 
| 68031 | 366 | apply (force simp: open_contains_cball) | 
| 53347 | 367 | apply (rule_tac x = "ball x e" in exI) | 
| 68031 | 368 | apply (simp add: subset_trans [OF ball_subset_cball], auto) | 
| 40377 | 369 | done | 
| 370 | ||
| 53347 | 371 | lemma rel_interior_cball: | 
| 372 |   "rel_interior S = {x \<in> S. \<exists>e. e > 0 \<and> cball x e \<inter> affine hull S \<subseteq> S}"
 | |
| 373 | using mem_rel_interior_cball [of _ S] by auto | |
| 40377 | 374 | |
| 60303 | 375 | lemma rel_interior_empty [simp]: "rel_interior {} = {}"
 | 
| 68031 | 376 | by (auto simp: rel_interior_def) | 
| 40377 | 377 | |
| 60303 | 378 | lemma affine_hull_sing [simp]: "affine hull {a :: 'n::euclidean_space} = {a}"
 | 
| 53347 | 379 | by (metis affine_hull_eq affine_sing) | 
| 40377 | 380 | |
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 381 | lemma rel_interior_sing [simp]: | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 382 |     fixes a :: "'n::euclidean_space"  shows "rel_interior {a} = {a}"
 | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 383 | apply (auto simp: rel_interior_ball) | 
| 68031 | 384 | apply (rule_tac x=1 in exI, force) | 
| 53347 | 385 | done | 
| 40377 | 386 | |
| 387 | lemma subset_rel_interior: | |
| 53347 | 388 | fixes S T :: "'n::euclidean_space set" | 
| 389 | assumes "S \<subseteq> T" | |
| 390 | and "affine hull S = affine hull T" | |
| 391 | shows "rel_interior S \<subseteq> rel_interior T" | |
| 68031 | 392 | using assms by (auto simp: rel_interior_def) | 
| 49531 | 393 | |
| 53347 | 394 | lemma rel_interior_subset: "rel_interior S \<subseteq> S" | 
| 68031 | 395 | by (auto simp: rel_interior_def) | 
| 53347 | 396 | |
| 397 | lemma rel_interior_subset_closure: "rel_interior S \<subseteq> closure S" | |
| 68031 | 398 | using rel_interior_subset by (auto simp: closure_def) | 
| 53347 | 399 | |
| 400 | lemma interior_subset_rel_interior: "interior S \<subseteq> rel_interior S" | |
| 68031 | 401 | by (auto simp: rel_interior interior_def) | 
| 40377 | 402 | |
| 403 | lemma interior_rel_interior: | |
| 53347 | 404 | fixes S :: "'n::euclidean_space set" | 
| 405 |   assumes "aff_dim S = int(DIM('n))"
 | |
| 406 | shows "rel_interior S = interior S" | |
| 40377 | 407 | proof - | 
| 53347 | 408 | have "affine hull S = UNIV" | 
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 409 | using assms affine_hull_UNIV[of S] by auto | 
| 53347 | 410 | then show ?thesis | 
| 411 | unfolding rel_interior interior_def by auto | |
| 40377 | 412 | qed | 
| 413 | ||
| 60303 | 414 | lemma rel_interior_interior: | 
| 415 | fixes S :: "'n::euclidean_space set" | |
| 416 | assumes "affine hull S = UNIV" | |
| 417 | shows "rel_interior S = interior S" | |
| 418 | using assms unfolding rel_interior interior_def by auto | |
| 419 | ||
| 40377 | 420 | lemma rel_interior_open: | 
| 53347 | 421 | fixes S :: "'n::euclidean_space set" | 
| 422 | assumes "open S" | |
| 423 | shows "rel_interior S = S" | |
| 424 | by (metis assms interior_eq interior_subset_rel_interior rel_interior_subset set_eq_subset) | |
| 40377 | 425 | |
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 426 | lemma interior_ball [simp]: "interior (ball x e) = ball x e" | 
| 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 427 | by (simp add: interior_open) | 
| 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 428 | |
| 40377 | 429 | lemma interior_rel_interior_gen: | 
| 53347 | 430 | fixes S :: "'n::euclidean_space set" | 
| 431 |   shows "interior S = (if aff_dim S = int(DIM('n)) then rel_interior S else {})"
 | |
| 432 | by (metis interior_rel_interior low_dim_interior) | |
| 40377 | 433 | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 434 | lemma rel_interior_nonempty_interior: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 435 | fixes S :: "'n::euclidean_space set" | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 436 |   shows "interior S \<noteq> {} \<Longrightarrow> rel_interior S = interior S"
 | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 437 | by (metis interior_rel_interior_gen) | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 438 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 439 | lemma affine_hull_nonempty_interior: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 440 | fixes S :: "'n::euclidean_space set" | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 441 |   shows "interior S \<noteq> {} \<Longrightarrow> affine hull S = UNIV"
 | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 442 | by (metis affine_hull_UNIV interior_rel_interior_gen) | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 443 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 444 | lemma rel_interior_affine_hull [simp]: | 
| 53347 | 445 | fixes S :: "'n::euclidean_space set" | 
| 446 | shows "rel_interior (affine hull S) = affine hull S" | |
| 447 | proof - | |
| 448 | have *: "rel_interior (affine hull S) \<subseteq> affine hull S" | |
| 449 | using rel_interior_subset by auto | |
| 450 |   {
 | |
| 451 | fix x | |
| 452 | assume x: "x \<in> affine hull S" | |
| 63040 | 453 | define e :: real where "e = 1" | 
| 53347 | 454 | then have "e > 0" "ball x e \<inter> affine hull (affine hull S) \<subseteq> affine hull S" | 
| 455 | using hull_hull[of _ S] by auto | |
| 456 | then have "x \<in> rel_interior (affine hull S)" | |
| 457 | using x rel_interior_ball[of "affine hull S"] by auto | |
| 458 | } | |
| 459 | then show ?thesis using * by auto | |
| 40377 | 460 | qed | 
| 461 | ||
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 462 | lemma rel_interior_UNIV [simp]: "rel_interior (UNIV :: ('n::euclidean_space) set) = UNIV"
 | 
| 53347 | 463 | by (metis open_UNIV rel_interior_open) | 
| 40377 | 464 | |
| 465 | lemma rel_interior_convex_shrink: | |
| 53347 | 466 | fixes S :: "'a::euclidean_space set" | 
| 467 | assumes "convex S" | |
| 468 | and "c \<in> rel_interior S" | |
| 469 | and "x \<in> S" | |
| 470 | and "0 < e" | |
| 471 | and "e \<le> 1" | |
| 472 | shows "x - e *\<^sub>R (x - c) \<in> rel_interior S" | |
| 473 | proof - | |
| 54465 | 474 | obtain d where "d > 0" and d: "ball c d \<inter> affine hull S \<subseteq> S" | 
| 53347 | 475 | using assms(2) unfolding mem_rel_interior_ball by auto | 
| 476 |   {
 | |
| 477 | fix y | |
| 478 | assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d" "y \<in> affine hull S" | |
| 479 | have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" | |
| 68031 | 480 | using \<open>e > 0\<close> by (auto simp: scaleR_left_diff_distrib scaleR_right_diff_distrib) | 
| 53347 | 481 | have "x \<in> affine hull S" | 
| 482 | using assms hull_subset[of S] by auto | |
| 49531 | 483 | moreover have "1 / e + - ((1 - e) / e) = 1" | 
| 60420 | 484 | using \<open>e > 0\<close> left_diff_distrib[of "1" "(1-e)" "1/e"] by auto | 
| 53347 | 485 | ultimately have **: "(1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x \<in> affine hull S" | 
| 486 | using as affine_affine_hull[of S] mem_affine[of "affine hull S" y x "(1 / e)" "-((1 - e) / e)"] | |
| 487 | by (simp add: algebra_simps) | |
| 61945 | 488 | have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = \<bar>1/e\<bar> * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" | 
| 53347 | 489 | unfolding dist_norm norm_scaleR[symmetric] | 
| 490 | apply (rule arg_cong[where f=norm]) | |
| 60420 | 491 | using \<open>e > 0\<close> | 
| 68031 | 492 | apply (auto simp: euclidean_eq_iff[where 'a='a] field_simps inner_simps) | 
| 53347 | 493 | done | 
| 61945 | 494 | also have "\<dots> = \<bar>1/e\<bar> * norm (x - e *\<^sub>R (x - c) - y)" | 
| 53347 | 495 | by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) | 
| 496 | also have "\<dots> < d" | |
| 60420 | 497 | using as[unfolded dist_norm] and \<open>e > 0\<close> | 
| 68031 | 498 | by (auto simp:pos_divide_less_eq[OF \<open>e > 0\<close>] mult.commute) | 
| 53347 | 499 | finally have "y \<in> S" | 
| 500 | apply (subst *) | |
| 501 | apply (rule assms(1)[unfolded convex_alt,rule_format]) | |
| 68058 | 502 | apply (rule d[THEN subsetD]) | 
| 53347 | 503 | unfolding mem_ball | 
| 504 | using assms(3-5) ** | |
| 505 | apply auto | |
| 506 | done | |
| 507 | } | |
| 508 | then have "ball (x - e *\<^sub>R (x - c)) (e*d) \<inter> affine hull S \<subseteq> S" | |
| 509 | by auto | |
| 510 | moreover have "e * d > 0" | |
| 60420 | 511 | using \<open>e > 0\<close> \<open>d > 0\<close> by simp | 
| 53347 | 512 | moreover have c: "c \<in> S" | 
| 513 | using assms rel_interior_subset by auto | |
| 514 | moreover from c have "x - e *\<^sub>R (x - c) \<in> S" | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61222diff
changeset | 515 | using convexD_alt[of S x c e] | 
| 53347 | 516 | apply (simp add: algebra_simps) | 
| 517 | using assms | |
| 518 | apply auto | |
| 519 | done | |
| 520 | ultimately show ?thesis | |
| 60420 | 521 | using mem_rel_interior_ball[of "x - e *\<^sub>R (x - c)" S] \<open>e > 0\<close> by auto | 
| 40377 | 522 | qed | 
| 523 | ||
| 69710 
61372780515b
some renamings and a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
69619diff
changeset | 524 | lemma interior_real_atLeast [simp]: | 
| 53347 | 525 | fixes a :: real | 
| 526 |   shows "interior {a..} = {a<..}"
 | |
| 527 | proof - | |
| 528 |   {
 | |
| 529 | fix y | |
| 530 | assume "a < y" | |
| 531 |     then have "y \<in> interior {a..}"
 | |
| 532 | apply (simp add: mem_interior) | |
| 533 | apply (rule_tac x="(y-a)" in exI) | |
| 68031 | 534 | apply (auto simp: dist_norm) | 
| 53347 | 535 | done | 
| 536 | } | |
| 537 | moreover | |
| 538 |   {
 | |
| 539 | fix y | |
| 540 |     assume "y \<in> interior {a..}"
 | |
| 541 |     then obtain e where e: "e > 0" "cball y e \<subseteq> {a..}"
 | |
| 542 |       using mem_interior_cball[of y "{a..}"] by auto
 | |
| 543 | moreover from e have "y - e \<in> cball y e" | |
| 68031 | 544 | by (auto simp: cball_def dist_norm) | 
| 60307 
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
 paulson <lp15@cam.ac.uk> parents: 
60303diff
changeset | 545 | ultimately have "a \<le> y - e" by blast | 
| 53347 | 546 | then have "a < y" using e by auto | 
| 547 | } | |
| 548 | ultimately show ?thesis by auto | |
| 40377 | 549 | qed | 
| 550 | ||
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 551 | lemma continuous_ge_on_Ioo: | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 552 |   assumes "continuous_on {c..d} g" "\<And>x. x \<in> {c<..<d} \<Longrightarrow> g x \<ge> a" "c < d" "x \<in> {c..d}"
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 553 | shows "g (x::real) \<ge> (a::real)" | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 554 | proof- | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 555 |   from assms(3) have "{c..d} = closure {c<..<d}" by (rule closure_greaterThanLessThan[symmetric])
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 556 |   also from assms(2) have "{c<..<d} \<subseteq> (g -` {a..} \<inter> {c..d})" by auto
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 557 |   hence "closure {c<..<d} \<subseteq> closure (g -` {a..} \<inter> {c..d})" by (rule closure_mono)
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 558 |   also from assms(1) have "closed (g -` {a..} \<inter> {c..d})"
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 559 | by (auto simp: continuous_on_closed_vimage) | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 560 |   hence "closure (g -` {a..} \<inter> {c..d}) = g -` {a..} \<inter> {c..d}" by simp
 | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61952diff
changeset | 561 |   finally show ?thesis using \<open>x \<in> {c..d}\<close> by auto
 | 
| 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61952diff
changeset | 562 | qed | 
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 563 | |
| 69710 
61372780515b
some renamings and a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
69619diff
changeset | 564 | lemma interior_real_atMost [simp]: | 
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 565 | fixes a :: real | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 566 |   shows "interior {..a} = {..<a}"
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 567 | proof - | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 568 |   {
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 569 | fix y | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 570 | assume "a > y" | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 571 |     then have "y \<in> interior {..a}"
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 572 | apply (simp add: mem_interior) | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 573 | apply (rule_tac x="(a-y)" in exI) | 
| 68031 | 574 | apply (auto simp: dist_norm) | 
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 575 | done | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 576 | } | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 577 | moreover | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 578 |   {
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 579 | fix y | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 580 |     assume "y \<in> interior {..a}"
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 581 |     then obtain e where e: "e > 0" "cball y e \<subseteq> {..a}"
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 582 |       using mem_interior_cball[of y "{..a}"] by auto
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 583 | moreover from e have "y + e \<in> cball y e" | 
| 68031 | 584 | by (auto simp: cball_def dist_norm) | 
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 585 | ultimately have "a \<ge> y + e" by auto | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 586 | then have "a > y" using e by auto | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 587 | } | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 588 | ultimately show ?thesis by auto | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 589 | qed | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 590 | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 591 | lemma interior_atLeastAtMost_real [simp]: "interior {a..b} = {a<..<b :: real}"
 | 
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 592 | proof- | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 593 |   have "{a..b} = {a..} \<inter> {..b}" by auto
 | 
| 68031 | 594 |   also have "interior \<dots> = {a<..} \<inter> {..<b}"
 | 
| 69710 
61372780515b
some renamings and a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
69619diff
changeset | 595 | by (simp add: interior_real_atLeast interior_real_atMost) | 
| 68031 | 596 |   also have "\<dots> = {a<..<b}" by auto
 | 
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 597 | finally show ?thesis . | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 598 | qed | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 599 | |
| 66793 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 600 | lemma interior_atLeastLessThan [simp]: | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 601 |   fixes a::real shows "interior {a..<b} = {a<..<b}"
 | 
| 69710 
61372780515b
some renamings and a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
69619diff
changeset | 602 | by (metis atLeastLessThan_def greaterThanLessThan_def interior_atLeastAtMost_real interior_Int interior_interior interior_real_atLeast) | 
| 66793 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 603 | |
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 604 | lemma interior_lessThanAtMost [simp]: | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 605 |   fixes a::real shows "interior {a<..b} = {a<..<b}"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 606 | by (metis atLeastAtMost_def greaterThanAtMost_def interior_atLeastAtMost_real interior_Int | 
| 69710 
61372780515b
some renamings and a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
69619diff
changeset | 607 | interior_interior interior_real_atLeast) | 
| 66793 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 608 | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 609 | lemma interior_greaterThanLessThan_real [simp]: "interior {a<..<b} = {a<..<b :: real}"
 | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 610 | by (metis interior_atLeastAtMost_real interior_interior) | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 611 | |
| 69710 
61372780515b
some renamings and a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
69619diff
changeset | 612 | lemma frontier_real_atMost [simp]: | 
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 613 | fixes a :: real | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 614 |   shows "frontier {..a} = {a}"
 | 
| 69710 
61372780515b
some renamings and a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
69619diff
changeset | 615 | unfolding frontier_def by (auto simp: interior_real_atMost) | 
| 
61372780515b
some renamings and a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
69619diff
changeset | 616 | |
| 
61372780515b
some renamings and a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
69619diff
changeset | 617 | lemma frontier_real_atLeast [simp]: "frontier {a..} = {a::real}"
 | 
| 
61372780515b
some renamings and a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
69619diff
changeset | 618 | by (auto simp: frontier_def) | 
| 
61372780515b
some renamings and a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
69619diff
changeset | 619 | |
| 
61372780515b
some renamings and a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
69619diff
changeset | 620 | lemma frontier_real_greaterThan [simp]: "frontier {a<..} = {a::real}"
 | 
| 
61372780515b
some renamings and a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
69619diff
changeset | 621 | by (auto simp: interior_open frontier_def) | 
| 
61372780515b
some renamings and a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
69619diff
changeset | 622 | |
| 
61372780515b
some renamings and a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
69619diff
changeset | 623 | lemma frontier_real_lessThan [simp]: "frontier {..<a} = {a::real}"
 | 
| 
61372780515b
some renamings and a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
69619diff
changeset | 624 | by (auto simp: interior_open frontier_def) | 
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61848diff
changeset | 625 | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 626 | lemma rel_interior_real_box [simp]: | 
| 53347 | 627 | fixes a b :: real | 
| 628 | assumes "a < b" | |
| 56188 | 629 |   shows "rel_interior {a .. b} = {a <..< b}"
 | 
| 53347 | 630 | proof - | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54465diff
changeset | 631 |   have "box a b \<noteq> {}"
 | 
| 53347 | 632 | using assms | 
| 633 | unfolding set_eq_iff | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 634 | by (auto intro!: exI[of _ "(a + b) / 2"] simp: box_def) | 
| 40377 | 635 | then show ?thesis | 
| 56188 | 636 | using interior_rel_interior_gen[of "cbox a b", symmetric] | 
| 62390 | 637 | by (simp split: if_split_asm del: box_real add: box_real[symmetric] interior_cbox) | 
| 40377 | 638 | qed | 
| 639 | ||
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 640 | lemma rel_interior_real_semiline [simp]: | 
| 53347 | 641 | fixes a :: real | 
| 642 |   shows "rel_interior {a..} = {a<..}"
 | |
| 643 | proof - | |
| 644 |   have *: "{a<..} \<noteq> {}"
 | |
| 645 | unfolding set_eq_iff by (auto intro!: exI[of _ "a + 1"]) | |
| 69710 
61372780515b
some renamings and a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
69619diff
changeset | 646 |   then show ?thesis using interior_real_atLeast interior_rel_interior_gen[of "{a..}"]
 | 
| 62390 | 647 | by (auto split: if_split_asm) | 
| 40377 | 648 | qed | 
| 649 | ||
| 60420 | 650 | subsubsection \<open>Relative open sets\<close> | 
| 40377 | 651 | |
| 70136 | 652 | definition\<^marker>\<open>tag important\<close> "rel_open S \<longleftrightarrow> rel_interior S = S" | 
| 53347 | 653 | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69710diff
changeset | 654 | lemma rel_open: "rel_open S \<longleftrightarrow> openin (top_of_set (affine hull S)) S" | 
| 53347 | 655 | unfolding rel_open_def rel_interior_def | 
| 656 | apply auto | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69710diff
changeset | 657 | using openin_subopen[of "top_of_set (affine hull S)" S] | 
| 53347 | 658 | apply auto | 
| 659 | done | |
| 660 | ||
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69710diff
changeset | 661 | lemma openin_rel_interior: "openin (top_of_set (affine hull S)) (rel_interior S)" | 
| 40377 | 662 | apply (simp add: rel_interior_def) | 
| 68031 | 663 | apply (subst openin_subopen, blast) | 
| 53347 | 664 | done | 
| 40377 | 665 | |
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 666 | lemma openin_set_rel_interior: | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69710diff
changeset | 667 | "openin (top_of_set S) (rel_interior S)" | 
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 668 | by (rule openin_subset_trans [OF openin_rel_interior rel_interior_subset hull_subset]) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 669 | |
| 49531 | 670 | lemma affine_rel_open: | 
| 53347 | 671 | fixes S :: "'n::euclidean_space set" | 
| 672 | assumes "affine S" | |
| 673 | shows "rel_open S" | |
| 674 | unfolding rel_open_def | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 675 | using assms rel_interior_affine_hull[of S] affine_hull_eq[of S] | 
| 53347 | 676 | by metis | 
| 40377 | 677 | |
| 49531 | 678 | lemma affine_closed: | 
| 53347 | 679 | fixes S :: "'n::euclidean_space set" | 
| 680 | assumes "affine S" | |
| 681 | shows "closed S" | |
| 682 | proof - | |
| 683 |   {
 | |
| 684 |     assume "S \<noteq> {}"
 | |
| 685 | then obtain L where L: "subspace L" "affine_parallel S L" | |
| 686 | using assms affine_parallel_subspace[of S] by auto | |
| 67399 | 687 | then obtain a where a: "S = ((+) a ` L)" | 
| 53347 | 688 | using affine_parallel_def[of L S] affine_parallel_commut by auto | 
| 689 | from L have "closed L" using closed_subspace by auto | |
| 690 | then have "closed S" | |
| 691 | using closed_translation a by auto | |
| 692 | } | |
| 693 | then show ?thesis by auto | |
| 40377 | 694 | qed | 
| 695 | ||
| 696 | lemma closure_affine_hull: | |
| 53347 | 697 | fixes S :: "'n::euclidean_space set" | 
| 698 | shows "closure S \<subseteq> affine hull S" | |
| 44524 | 699 | by (intro closure_minimal hull_subset affine_closed affine_affine_hull) | 
| 40377 | 700 | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 701 | lemma closure_same_affine_hull [simp]: | 
| 53347 | 702 | fixes S :: "'n::euclidean_space set" | 
| 40377 | 703 | shows "affine hull (closure S) = affine hull S" | 
| 53347 | 704 | proof - | 
| 705 | have "affine hull (closure S) \<subseteq> affine hull S" | |
| 706 | using hull_mono[of "closure S" "affine hull S" "affine"] | |
| 707 | closure_affine_hull[of S] hull_hull[of "affine" S] | |
| 708 | by auto | |
| 709 | moreover have "affine hull (closure S) \<supseteq> affine hull S" | |
| 710 | using hull_mono[of "S" "closure S" "affine"] closure_subset by auto | |
| 711 | ultimately show ?thesis by auto | |
| 49531 | 712 | qed | 
| 713 | ||
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 714 | lemma closure_aff_dim [simp]: | 
| 53347 | 715 | fixes S :: "'n::euclidean_space set" | 
| 40377 | 716 | shows "aff_dim (closure S) = aff_dim S" | 
| 53347 | 717 | proof - | 
| 718 | have "aff_dim S \<le> aff_dim (closure S)" | |
| 719 | using aff_dim_subset closure_subset by auto | |
| 720 | moreover have "aff_dim (closure S) \<le> aff_dim (affine hull S)" | |
| 63075 
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
 paulson <lp15@cam.ac.uk> parents: 
63072diff
changeset | 721 | using aff_dim_subset closure_affine_hull by blast | 
| 53347 | 722 | moreover have "aff_dim (affine hull S) = aff_dim S" | 
| 723 | using aff_dim_affine_hull by auto | |
| 724 | ultimately show ?thesis by auto | |
| 40377 | 725 | qed | 
| 726 | ||
| 727 | lemma rel_interior_closure_convex_shrink: | |
| 53347 | 728 | fixes S :: "_::euclidean_space set" | 
| 729 | assumes "convex S" | |
| 730 | and "c \<in> rel_interior S" | |
| 731 | and "x \<in> closure S" | |
| 732 | and "e > 0" | |
| 733 | and "e \<le> 1" | |
| 734 | shows "x - e *\<^sub>R (x - c) \<in> rel_interior S" | |
| 735 | proof - | |
| 736 | obtain d where "d > 0" and d: "ball c d \<inter> affine hull S \<subseteq> S" | |
| 737 | using assms(2) unfolding mem_rel_interior_ball by auto | |
| 738 | have "\<exists>y \<in> S. norm (y - x) * (1 - e) < e * d" | |
| 739 | proof (cases "x \<in> S") | |
| 740 | case True | |
| 60420 | 741 | then show ?thesis using \<open>e > 0\<close> \<open>d > 0\<close> | 
| 68031 | 742 | apply (rule_tac bexI[where x=x], auto) | 
| 53347 | 743 | done | 
| 744 | next | |
| 745 | case False | |
| 746 | then have x: "x islimpt S" | |
| 747 | using assms(3)[unfolded closure_def] by auto | |
| 748 | show ?thesis | |
| 749 | proof (cases "e = 1") | |
| 750 | case True | |
| 751 | obtain y where "y \<in> S" "y \<noteq> x" "dist y x < 1" | |
| 40377 | 752 | using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto | 
| 53347 | 753 | then show ?thesis | 
| 754 | apply (rule_tac x=y in bexI) | |
| 755 | unfolding True | |
| 60420 | 756 | using \<open>d > 0\<close> | 
| 53347 | 757 | apply auto | 
| 758 | done | |
| 759 | next | |
| 760 | case False | |
| 761 | then have "0 < e * d / (1 - e)" and *: "1 - e > 0" | |
| 68031 | 762 | using \<open>e \<le> 1\<close> \<open>e > 0\<close> \<open>d > 0\<close> by auto | 
| 53347 | 763 | then obtain y where "y \<in> S" "y \<noteq> x" "dist y x < e * d / (1 - e)" | 
| 40377 | 764 | using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto | 
| 53347 | 765 | then show ?thesis | 
| 766 | apply (rule_tac x=y in bexI) | |
| 767 | unfolding dist_norm | |
| 768 | using pos_less_divide_eq[OF *] | |
| 769 | apply auto | |
| 770 | done | |
| 771 | qed | |
| 772 | qed | |
| 773 | then obtain y where "y \<in> S" and y: "norm (y - x) * (1 - e) < e * d" | |
| 774 | by auto | |
| 63040 | 775 | define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)" | 
| 53347 | 776 | have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" | 
| 60420 | 777 | unfolding z_def using \<open>e > 0\<close> | 
| 68031 | 778 | by (auto simp: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) | 
| 53347 | 779 | have zball: "z \<in> ball c d" | 
| 780 | using mem_ball z_def dist_norm[of c] | |
| 781 | using y and assms(4,5) | |
| 68031 | 782 | by (auto simp:field_simps norm_minus_commute) | 
| 53347 | 783 | have "x \<in> affine hull S" | 
| 784 | using closure_affine_hull assms by auto | |
| 785 | moreover have "y \<in> affine hull S" | |
| 60420 | 786 | using \<open>y \<in> S\<close> hull_subset[of S] by auto | 
| 53347 | 787 | moreover have "c \<in> affine hull S" | 
| 788 | using assms rel_interior_subset hull_subset[of S] by auto | |
| 789 | ultimately have "z \<in> affine hull S" | |
| 49531 | 790 | using z_def affine_affine_hull[of S] | 
| 53347 | 791 | mem_affine_3_minus [of "affine hull S" c x y "(1 - e) / e"] | 
| 792 | assms | |
| 68031 | 793 | by (auto simp: field_simps) | 
| 53347 | 794 | then have "z \<in> S" using d zball by auto | 
| 795 | obtain d1 where "d1 > 0" and d1: "ball z d1 \<le> ball c d" | |
| 40377 | 796 | using zball open_ball[of c d] openE[of "ball c d" z] by auto | 
| 53347 | 797 | then have "ball z d1 \<inter> affine hull S \<subseteq> ball c d \<inter> affine hull S" | 
| 798 | by auto | |
| 799 | then have "ball z d1 \<inter> affine hull S \<subseteq> S" | |
| 800 | using d by auto | |
| 801 | then have "z \<in> rel_interior S" | |
| 60420 | 802 | using mem_rel_interior_ball using \<open>d1 > 0\<close> \<open>z \<in> S\<close> by auto | 
| 53347 | 803 | then have "y - e *\<^sub>R (y - z) \<in> rel_interior S" | 
| 60420 | 804 | using rel_interior_convex_shrink[of S z y e] assms \<open>y \<in> S\<close> by auto | 
| 53347 | 805 | then show ?thesis using * by auto | 
| 806 | qed | |
| 807 | ||
| 62620 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 808 | lemma rel_interior_eq: | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69710diff
changeset | 809 | "rel_interior s = s \<longleftrightarrow> openin(top_of_set (affine hull s)) s" | 
| 62620 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 810 | using rel_open rel_open_def by blast | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 811 | |
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 812 | lemma rel_interior_openin: | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69710diff
changeset | 813 | "openin(top_of_set (affine hull s)) s \<Longrightarrow> rel_interior s = s" | 
| 62620 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 814 | by (simp add: rel_interior_eq) | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 815 | |
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 816 | lemma rel_interior_affine: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 817 | fixes S :: "'n::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 818 | shows "affine S \<Longrightarrow> rel_interior S = S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 819 | using affine_rel_open rel_open_def by auto | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 820 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 821 | lemma rel_interior_eq_closure: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 822 | fixes S :: "'n::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 823 | shows "rel_interior S = closure S \<longleftrightarrow> affine S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 824 | proof (cases "S = {}")
 | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 825 | case True | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 826 | then show ?thesis | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 827 | by auto | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 828 | next | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 829 | case False show ?thesis | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 830 | proof | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 831 | assume eq: "rel_interior S = closure S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 832 |     have "S = {} \<or> S = affine hull S"
 | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 833 | apply (rule connected_clopen [THEN iffD1, rule_format]) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 834 | apply (simp add: affine_imp_convex convex_connected) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 835 | apply (rule conjI) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 836 | apply (metis eq closure_subset openin_rel_interior rel_interior_subset subset_antisym) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 837 | apply (metis closed_subset closure_subset_eq eq hull_subset rel_interior_subset) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 838 | done | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 839 | with False have "affine hull S = S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 840 | by auto | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 841 | then show "affine S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 842 | by (metis affine_hull_eq) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 843 | next | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 844 | assume "affine S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 845 | then show "rel_interior S = closure S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 846 | by (simp add: rel_interior_affine affine_closed) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 847 | qed | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 848 | qed | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 849 | |
| 40377 | 850 | |
| 70136 | 851 | subsubsection\<^marker>\<open>tag unimportant\<close>\<open>Relative interior preserves under linear transformations\<close> | 
| 40377 | 852 | |
| 853 | lemma rel_interior_translation_aux: | |
| 53347 | 854 | fixes a :: "'n::euclidean_space" | 
| 855 | shows "((\<lambda>x. a + x) ` rel_interior S) \<subseteq> rel_interior ((\<lambda>x. a + x) ` S)" | |
| 856 | proof - | |
| 857 |   {
 | |
| 858 | fix x | |
| 859 | assume x: "x \<in> rel_interior S" | |
| 860 | then obtain T where "open T" "x \<in> T \<inter> S" "T \<inter> affine hull S \<subseteq> S" | |
| 861 | using mem_rel_interior[of x S] by auto | |
| 862 | then have "open ((\<lambda>x. a + x) ` T)" | |
| 863 | and "a + x \<in> ((\<lambda>x. a + x) ` T) \<inter> ((\<lambda>x. a + x) ` S)" | |
| 864 | and "((\<lambda>x. a + x) ` T) \<inter> affine hull ((\<lambda>x. a + x) ` S) \<subseteq> (\<lambda>x. a + x) ` S" | |
| 865 | using affine_hull_translation[of a S] open_translation[of T a] x by auto | |
| 866 | then have "a + x \<in> rel_interior ((\<lambda>x. a + x) ` S)" | |
| 867 | using mem_rel_interior[of "a+x" "((\<lambda>x. a + x) ` S)"] by auto | |
| 868 | } | |
| 869 | then show ?thesis by auto | |
| 60809 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60800diff
changeset | 870 | qed | 
| 40377 | 871 | |
| 872 | lemma rel_interior_translation: | |
| 53347 | 873 | fixes a :: "'n::euclidean_space" | 
| 874 | shows "rel_interior ((\<lambda>x. a + x) ` S) = (\<lambda>x. a + x) ` rel_interior S" | |
| 875 | proof - | |
| 876 | have "(\<lambda>x. (-a) + x) ` rel_interior ((\<lambda>x. a + x) ` S) \<subseteq> rel_interior S" | |
| 877 | using rel_interior_translation_aux[of "-a" "(\<lambda>x. a + x) ` S"] | |
| 878 | translation_assoc[of "-a" "a"] | |
| 879 | by auto | |
| 880 | then have "((\<lambda>x. a + x) ` rel_interior S) \<supseteq> rel_interior ((\<lambda>x. a + x) ` S)" | |
| 67399 | 881 | using translation_inverse_subset[of a "rel_interior ((+) a ` S)" "rel_interior S"] | 
| 53347 | 882 | by auto | 
| 883 | then show ?thesis | |
| 884 | using rel_interior_translation_aux[of a S] by auto | |
| 40377 | 885 | qed | 
| 886 | ||
| 887 | ||
| 888 | lemma affine_hull_linear_image: | |
| 53347 | 889 | assumes "bounded_linear f" | 
| 890 | shows "f ` (affine hull s) = affine hull f ` s" | |
| 891 | proof - | |
| 40377 | 892 | interpret f: bounded_linear f by fact | 
| 68058 | 893 |   have "affine {x. f x \<in> affine hull f ` s}"
 | 
| 53347 | 894 | unfolding affine_def | 
| 68031 | 895 | by (auto simp: f.scaleR f.add affine_affine_hull[unfolded affine_def, rule_format]) | 
| 68058 | 896 |   moreover have "affine {x. x \<in> f ` (affine hull s)}"
 | 
| 53347 | 897 | using affine_affine_hull[unfolded affine_def, of s] | 
| 68031 | 898 | unfolding affine_def by (auto simp: f.scaleR [symmetric] f.add [symmetric]) | 
| 68058 | 899 | ultimately show ?thesis | 
| 900 | by (auto simp: hull_inc elim!: hull_induct) | |
| 901 | qed | |
| 40377 | 902 | |
| 903 | ||
| 904 | lemma rel_interior_injective_on_span_linear_image: | |
| 53347 | 905 | fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space" | 
| 906 | and S :: "'m::euclidean_space set" | |
| 907 | assumes "bounded_linear f" | |
| 908 | and "inj_on f (span S)" | |
| 909 | shows "rel_interior (f ` S) = f ` (rel_interior S)" | |
| 910 | proof - | |
| 911 |   {
 | |
| 912 | fix z | |
| 913 | assume z: "z \<in> rel_interior (f ` S)" | |
| 914 | then have "z \<in> f ` S" | |
| 915 | using rel_interior_subset[of "f ` S"] by auto | |
| 916 | then obtain x where x: "x \<in> S" "f x = z" by auto | |
| 917 | obtain e2 where e2: "e2 > 0" "cball z e2 \<inter> affine hull (f ` S) \<subseteq> (f ` S)" | |
| 918 | using z rel_interior_cball[of "f ` S"] by auto | |
| 919 | obtain K where K: "K > 0" "\<And>x. norm (f x) \<le> norm x * K" | |
| 920 | using assms Real_Vector_Spaces.bounded_linear.pos_bounded[of f] by auto | |
| 63040 | 921 | define e1 where "e1 = 1 / K" | 
| 53347 | 922 | then have e1: "e1 > 0" "\<And>x. e1 * norm (f x) \<le> norm x" | 
| 923 | using K pos_le_divide_eq[of e1] by auto | |
| 63040 | 924 | define e where "e = e1 * e2" | 
| 56544 | 925 | then have "e > 0" using e1 e2 by auto | 
| 53347 | 926 |     {
 | 
| 927 | fix y | |
| 928 | assume y: "y \<in> cball x e \<inter> affine hull S" | |
| 929 | then have h1: "f y \<in> affine hull (f ` S)" | |
| 930 | using affine_hull_linear_image[of f S] assms by auto | |
| 931 | from y have "norm (x-y) \<le> e1 * e2" | |
| 932 | using cball_def[of x e] dist_norm[of x y] e_def by auto | |
| 933 | moreover have "f x - f y = f (x - y)" | |
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 934 | using assms linear_diff[of f x y] linear_conv_bounded_linear[of f] by auto | 
| 53347 | 935 | moreover have "e1 * norm (f (x-y)) \<le> norm (x - y)" | 
| 936 | using e1 by auto | |
| 937 | ultimately have "e1 * norm ((f x)-(f y)) \<le> e1 * e2" | |
| 938 | by auto | |
| 939 | then have "f y \<in> cball z e2" | |
| 940 | using cball_def[of "f x" e2] dist_norm[of "f x" "f y"] e1 x by auto | |
| 941 | then have "f y \<in> f ` S" | |
| 942 | using y e2 h1 by auto | |
| 943 | then have "y \<in> S" | |
| 944 | using assms y hull_subset[of S] affine_hull_subset_span | |
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 945 | inj_on_image_mem_iff [OF \<open>inj_on f (span S)\<close>] | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 946 | by (metis Int_iff span_superset subsetCE) | 
| 53347 | 947 | } | 
| 948 | then have "z \<in> f ` (rel_interior S)" | |
| 60420 | 949 | using mem_rel_interior_cball[of x S] \<open>e > 0\<close> x by auto | 
| 49531 | 950 | } | 
| 53347 | 951 | moreover | 
| 952 |   {
 | |
| 953 | fix x | |
| 954 | assume x: "x \<in> rel_interior S" | |
| 54465 | 955 | then obtain e2 where e2: "e2 > 0" "cball x e2 \<inter> affine hull S \<subseteq> S" | 
| 53347 | 956 | using rel_interior_cball[of S] by auto | 
| 957 | have "x \<in> S" using x rel_interior_subset by auto | |
| 958 | then have *: "f x \<in> f ` S" by auto | |
| 959 | have "\<forall>x\<in>span S. f x = 0 \<longrightarrow> x = 0" | |
| 960 | using assms subspace_span linear_conv_bounded_linear[of f] | |
| 961 | linear_injective_on_subspace_0[of f "span S"] | |
| 962 | by auto | |
| 963 | then obtain e1 where e1: "e1 > 0" "\<forall>x \<in> span S. e1 * norm x \<le> norm (f x)" | |
| 964 | using assms injective_imp_isometric[of "span S" f] | |
| 965 | subspace_span[of S] closed_subspace[of "span S"] | |
| 966 | by auto | |
| 63040 | 967 | define e where "e = e1 * e2" | 
| 56544 | 968 | hence "e > 0" using e1 e2 by auto | 
| 53347 | 969 |     {
 | 
| 970 | fix y | |
| 971 | assume y: "y \<in> cball (f x) e \<inter> affine hull (f ` S)" | |
| 972 | then have "y \<in> f ` (affine hull S)" | |
| 973 | using affine_hull_linear_image[of f S] assms by auto | |
| 974 | then obtain xy where xy: "xy \<in> affine hull S" "f xy = y" by auto | |
| 975 | with y have "norm (f x - f xy) \<le> e1 * e2" | |
| 976 | using cball_def[of "f x" e] dist_norm[of "f x" y] e_def by auto | |
| 977 | moreover have "f x - f xy = f (x - xy)" | |
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63332diff
changeset | 978 | using assms linear_diff[of f x xy] linear_conv_bounded_linear[of f] by auto | 
| 53347 | 979 | moreover have *: "x - xy \<in> span S" | 
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 980 | using subspace_diff[of "span S" x xy] subspace_span \<open>x \<in> S\<close> xy | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 981 | affine_hull_subset_span[of S] span_superset | 
| 53347 | 982 | by auto | 
| 983 | moreover from * have "e1 * norm (x - xy) \<le> norm (f (x - xy))" | |
| 984 | using e1 by auto | |
| 985 | ultimately have "e1 * norm (x - xy) \<le> e1 * e2" | |
| 986 | by auto | |
| 987 | then have "xy \<in> cball x e2" | |
| 988 | using cball_def[of x e2] dist_norm[of x xy] e1 by auto | |
| 989 | then have "y \<in> f ` S" | |
| 990 | using xy e2 by auto | |
| 991 | } | |
| 992 | then have "f x \<in> rel_interior (f ` S)" | |
| 60420 | 993 | using mem_rel_interior_cball[of "(f x)" "(f ` S)"] * \<open>e > 0\<close> by auto | 
| 49531 | 994 | } | 
| 53347 | 995 | ultimately show ?thesis by auto | 
| 40377 | 996 | qed | 
| 997 | ||
| 998 | lemma rel_interior_injective_linear_image: | |
| 53347 | 999 | fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space" | 
| 1000 | assumes "bounded_linear f" | |
| 1001 | and "inj f" | |
| 1002 | shows "rel_interior (f ` S) = f ` (rel_interior S)" | |
| 1003 | using assms rel_interior_injective_on_span_linear_image[of f S] | |
| 1004 | subset_inj_on[of f "UNIV" "span S"] | |
| 1005 | by auto | |
| 1006 | ||
| 40377 | 1007 | |
| 70136 | 1008 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Openness and compactness are preserved by convex hull operation\<close> | 
| 33175 | 1009 | |
| 34964 | 1010 | lemma open_convex_hull[intro]: | 
| 68052 | 1011 | fixes S :: "'a::real_normed_vector set" | 
| 1012 | assumes "open S" | |
| 1013 | shows "open (convex hull S)" | |
| 1014 | proof (clarsimp simp: open_contains_cball convex_hull_explicit) | |
| 1015 | fix T and u :: "'a\<Rightarrow>real" | |
| 1016 | assume obt: "finite T" "T\<subseteq>S" "\<forall>x\<in>T. 0 \<le> u x" "sum u T = 1" | |
| 53347 | 1017 | |
| 1018 | from assms[unfolded open_contains_cball] obtain b | |
| 68052 | 1019 | where b: "\<And>x. x\<in>S \<Longrightarrow> 0 < b x \<and> cball x (b x) \<subseteq> S" by metis | 
| 1020 |   have "b ` T \<noteq> {}"
 | |
| 56889 
48a745e1bde7
avoid the Complex constructor, use the more natural Re/Im view; moved csqrt to Complex.
 hoelzl parents: 
56571diff
changeset | 1021 | using obt by auto | 
| 68052 | 1022 | define i where "i = b ` T" | 
| 1023 | let ?\<Phi> = "\<lambda>y. \<exists>F. finite F \<and> F \<subseteq> S \<and> (\<exists>u. (\<forall>x\<in>F. 0 \<le> u x) \<and> sum u F = 1 \<and> (\<Sum>v\<in>F. u v *\<^sub>R v) = y)" | |
| 1024 | let ?a = "\<Sum>v\<in>T. u v *\<^sub>R v" | |
| 1025 |   show "\<exists>e > 0. cball ?a e \<subseteq> {y. ?\<Phi> y}"
 | |
| 1026 | proof (intro exI subsetI conjI) | |
| 53347 | 1027 | show "0 < Min i" | 
| 68052 | 1028 |       unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] \<open>b ` T\<noteq>{}\<close>]
 | 
| 1029 | using b \<open>T\<subseteq>S\<close> by auto | |
| 53347 | 1030 | next | 
| 1031 | fix y | |
| 68052 | 1032 | assume "y \<in> cball ?a (Min i)" | 
| 1033 | then have y: "norm (?a - y) \<le> Min i" | |
| 53347 | 1034 | unfolding dist_norm[symmetric] by auto | 
| 68052 | 1035 |     { fix x
 | 
| 1036 | assume "x \<in> T" | |
| 53347 | 1037 | then have "Min i \<le> b x" | 
| 68052 | 1038 | by (simp add: i_def obt(1)) | 
| 1039 | then have "x + (y - ?a) \<in> cball x (b x)" | |
| 53347 | 1040 | using y unfolding mem_cball dist_norm by auto | 
| 68052 | 1041 | moreover have "x \<in> S" | 
| 1042 | using \<open>x\<in>T\<close> \<open>T\<subseteq>S\<close> by auto | |
| 1043 | ultimately have "x + (y - ?a) \<in> S" | |
| 1044 | using y b by blast | |
| 53347 | 1045 | } | 
| 33175 | 1046 | moreover | 
| 68052 | 1047 | have *: "inj_on (\<lambda>v. v + (y - ?a)) T" | 
| 53347 | 1048 | unfolding inj_on_def by auto | 
| 68052 | 1049 | have "(\<Sum>v\<in>(\<lambda>v. v + (y - ?a)) ` T. u (v - (y - ?a)) *\<^sub>R v) = y" | 
| 1050 | unfolding sum.reindex[OF *] o_def using obt(4) | |
| 64267 | 1051 | by (simp add: sum.distrib sum_subtractf scaleR_left.sum[symmetric] scaleR_right_distrib) | 
| 68052 | 1052 |     ultimately show "y \<in> {y. ?\<Phi> y}"
 | 
| 1053 | proof (intro CollectI exI conjI) | |
| 1054 | show "finite ((\<lambda>v. v + (y - ?a)) ` T)" | |
| 1055 | by (simp add: obt(1)) | |
| 1056 | show "sum (\<lambda>v. u (v - (y - ?a))) ((\<lambda>v. v + (y - ?a)) ` T) = 1" | |
| 1057 | unfolding sum.reindex[OF *] o_def using obt(4) by auto | |
| 1058 | qed (use obt(1, 3) in auto) | |
| 33175 | 1059 | qed | 
| 1060 | qed | |
| 1061 | ||
| 1062 | lemma compact_convex_combinations: | |
| 68052 | 1063 | fixes S T :: "'a::real_normed_vector set" | 
| 1064 | assumes "compact S" "compact T" | |
| 1065 |   shows "compact { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \<le> u \<and> u \<le> 1 \<and> x \<in> S \<and> y \<in> T}"
 | |
| 53347 | 1066 | proof - | 
| 68052 | 1067 |   let ?X = "{0..1} \<times> S \<times> T"
 | 
| 33175 | 1068 | let ?h = "(\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" | 
| 68052 | 1069 |   have *: "{ (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \<le> u \<and> u \<le> 1 \<and> x \<in> S \<and> y \<in> T} = ?h ` ?X"
 | 
| 1070 | by force | |
| 56188 | 1071 | have "continuous_on ?X (\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" | 
| 33175 | 1072 | unfolding continuous_on by (rule ballI) (intro tendsto_intros) | 
| 68052 | 1073 | with assms show ?thesis | 
| 1074 | by (simp add: * compact_Times compact_continuous_image) | |
| 33175 | 1075 | qed | 
| 1076 | ||
| 44525 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 1077 | lemma finite_imp_compact_convex_hull: | 
| 68052 | 1078 | fixes S :: "'a::real_normed_vector set" | 
| 1079 | assumes "finite S" | |
| 1080 | shows "compact (convex hull S)" | |
| 1081 | proof (cases "S = {}")
 | |
| 53347 | 1082 | case True | 
| 1083 | then show ?thesis by simp | |
| 44525 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 1084 | next | 
| 53347 | 1085 | case False | 
| 1086 | with assms show ?thesis | |
| 44525 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 1087 | proof (induct rule: finite_ne_induct) | 
| 53347 | 1088 | case (singleton x) | 
| 1089 | show ?case by simp | |
| 44525 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 1090 | next | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 1091 | case (insert x A) | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 1092 | let ?f = "\<lambda>(u, y::'a). u *\<^sub>R x + (1 - u) *\<^sub>R y" | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 1093 |     let ?T = "{0..1::real} \<times> (convex hull A)"
 | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 1094 | have "continuous_on ?T ?f" | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 1095 | unfolding split_def continuous_on by (intro ballI tendsto_intros) | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 1096 | moreover have "compact ?T" | 
| 56188 | 1097 | by (intro compact_Times compact_Icc insert) | 
| 44525 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 1098 | ultimately have "compact (?f ` ?T)" | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 1099 | by (rule compact_continuous_image) | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 1100 | also have "?f ` ?T = convex hull (insert x A)" | 
| 60420 | 1101 |       unfolding convex_hull_insert [OF \<open>A \<noteq> {}\<close>]
 | 
| 44525 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 1102 | apply safe | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 1103 | apply (rule_tac x=a in exI, simp) | 
| 68031 | 1104 | apply (rule_tac x="1 - a" in exI, simp, fast) | 
| 44525 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 1105 | apply (rule_tac x="(u, b)" in image_eqI, simp_all) | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 1106 | done | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 1107 | finally show "compact (convex hull (insert x A))" . | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 1108 | qed | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 1109 | qed | 
| 
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
 huffman parents: 
44524diff
changeset | 1110 | |
| 53347 | 1111 | lemma compact_convex_hull: | 
| 68052 | 1112 | fixes S :: "'a::euclidean_space set" | 
| 1113 | assumes "compact S" | |
| 1114 | shows "compact (convex hull S)" | |
| 1115 | proof (cases "S = {}")
 | |
| 53347 | 1116 | case True | 
| 1117 | then show ?thesis using compact_empty by simp | |
| 33175 | 1118 | next | 
| 53347 | 1119 | case False | 
| 68052 | 1120 | then obtain w where "w \<in> S" by auto | 
| 53347 | 1121 | show ?thesis | 
| 68052 | 1122 | unfolding caratheodory[of S] | 
| 53347 | 1123 |   proof (induct ("DIM('a) + 1"))
 | 
| 1124 | case 0 | |
| 68052 | 1125 |     have *: "{x.\<exists>sa. finite sa \<and> sa \<subseteq> S \<and> card sa \<le> 0 \<and> x \<in> convex hull sa} = {}"
 | 
| 36362 
06475a1547cb
fix lots of looping simp calls and other warnings
 huffman parents: 
36341diff
changeset | 1126 | using compact_empty by auto | 
| 53347 | 1127 | from 0 show ?case unfolding * by simp | 
| 33175 | 1128 | next | 
| 1129 | case (Suc n) | |
| 53347 | 1130 | show ?case | 
| 1131 | proof (cases "n = 0") | |
| 1132 | case True | |
| 68052 | 1133 |       have "{x. \<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T} = S"
 | 
| 53347 | 1134 | unfolding set_eq_iff and mem_Collect_eq | 
| 1135 | proof (rule, rule) | |
| 1136 | fix x | |
| 68052 | 1137 | assume "\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T" | 
| 1138 | then obtain T where T: "finite T" "T \<subseteq> S" "card T \<le> Suc n" "x \<in> convex hull T" | |
| 53347 | 1139 | by auto | 
| 68052 | 1140 | show "x \<in> S" | 
| 1141 | proof (cases "card T = 0") | |
| 53347 | 1142 | case True | 
| 1143 | then show ?thesis | |
| 68052 | 1144 | using T(4) unfolding card_0_eq[OF T(1)] by simp | 
| 33175 | 1145 | next | 
| 53347 | 1146 | case False | 
| 68052 | 1147 | then have "card T = Suc 0" using T(3) \<open>n=0\<close> by auto | 
| 1148 |           then obtain a where "T = {a}" unfolding card_Suc_eq by auto
 | |
| 1149 | then show ?thesis using T(2,4) by simp | |
| 33175 | 1150 | qed | 
| 1151 | next | |
| 68052 | 1152 | fix x assume "x\<in>S" | 
| 1153 | then show "\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T" | |
| 53347 | 1154 |           apply (rule_tac x="{x}" in exI)
 | 
| 1155 | unfolding convex_hull_singleton | |
| 1156 | apply auto | |
| 1157 | done | |
| 1158 | qed | |
| 1159 | then show ?thesis using assms by simp | |
| 33175 | 1160 | next | 
| 53347 | 1161 | case False | 
| 68052 | 1162 |       have "{x. \<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T} =
 | 
| 53347 | 1163 |         {(1 - u) *\<^sub>R x + u *\<^sub>R y | x y u.
 | 
| 68052 | 1164 |           0 \<le> u \<and> u \<le> 1 \<and> x \<in> S \<and> y \<in> {x. \<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> n \<and> x \<in> convex hull T}}"
 | 
| 53347 | 1165 | unfolding set_eq_iff and mem_Collect_eq | 
| 1166 | proof (rule, rule) | |
| 1167 | fix x | |
| 1168 | assume "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and> | |
| 68052 | 1169 | 0 \<le> c \<and> c \<le> 1 \<and> u \<in> S \<and> (\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> n \<and> v \<in> convex hull T)" | 
| 1170 | then obtain u v c T where obt: "x = (1 - c) *\<^sub>R u + c *\<^sub>R v" | |
| 1171 | "0 \<le> c \<and> c \<le> 1" "u \<in> S" "finite T" "T \<subseteq> S" "card T \<le> n" "v \<in> convex hull T" | |
| 53347 | 1172 | by auto | 
| 68052 | 1173 | moreover have "(1 - c) *\<^sub>R u + c *\<^sub>R v \<in> convex hull insert u T" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61222diff
changeset | 1174 | apply (rule convexD_alt) | 
| 68052 | 1175 | using obt(2) and convex_convex_hull and hull_subset[of "insert u T" convex] | 
| 1176 | using obt(7) and hull_mono[of T "insert u T"] | |
| 53347 | 1177 | apply auto | 
| 1178 | done | |
| 68052 | 1179 | ultimately show "\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T" | 
| 1180 | apply (rule_tac x="insert u T" in exI) | |
| 68031 | 1181 | apply (auto simp: card_insert_if) | 
| 53347 | 1182 | done | 
| 33175 | 1183 | next | 
| 53347 | 1184 | fix x | 
| 68052 | 1185 | assume "\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T" | 
| 1186 | then obtain T where T: "finite T" "T \<subseteq> S" "card T \<le> Suc n" "x \<in> convex hull T" | |
| 53347 | 1187 | by auto | 
| 1188 | show "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and> | |
| 68052 | 1189 | 0 \<le> c \<and> c \<le> 1 \<and> u \<in> S \<and> (\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> n \<and> v \<in> convex hull T)" | 
| 1190 | proof (cases "card T = Suc n") | |
| 53347 | 1191 | case False | 
| 68052 | 1192 | then have "card T \<le> n" using T(3) by auto | 
| 53347 | 1193 | then show ?thesis | 
| 1194 | apply (rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI) | |
| 68052 | 1195 | using \<open>w\<in>S\<close> and T | 
| 1196 | apply (auto intro!: exI[where x=T]) | |
| 53347 | 1197 | done | 
| 33175 | 1198 | next | 
| 53347 | 1199 | case True | 
| 68052 | 1200 | then obtain a u where au: "T = insert a u" "a\<notin>u" | 
| 68031 | 1201 | apply (drule_tac card_eq_SucD, auto) | 
| 53347 | 1202 | done | 
| 1203 | show ?thesis | |
| 1204 |           proof (cases "u = {}")
 | |
| 1205 | case True | |
| 68052 | 1206 | then have "x = a" using T(4)[unfolded au] by auto | 
| 60420 | 1207 | show ?thesis unfolding \<open>x = a\<close> | 
| 53347 | 1208 | apply (rule_tac x=a in exI) | 
| 1209 | apply (rule_tac x=a in exI) | |
| 1210 | apply (rule_tac x=1 in exI) | |
| 68052 | 1211 | using T and \<open>n \<noteq> 0\<close> | 
| 53347 | 1212 | unfolding au | 
| 1213 |               apply (auto intro!: exI[where x="{a}"])
 | |
| 1214 | done | |
| 33175 | 1215 | next | 
| 53347 | 1216 | case False | 
| 1217 | obtain ux vx b where obt: "ux\<ge>0" "vx\<ge>0" "ux + vx = 1" | |
| 1218 | "b \<in> convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b" | |
| 68052 | 1219 | using T(4)[unfolded au convex_hull_insert[OF False]] | 
| 53347 | 1220 | by auto | 
| 1221 | have *: "1 - vx = ux" using obt(3) by auto | |
| 1222 | show ?thesis | |
| 1223 | apply (rule_tac x=a in exI) | |
| 1224 | apply (rule_tac x=b in exI) | |
| 1225 | apply (rule_tac x=vx in exI) | |
| 68052 | 1226 | using obt and T(1-3) | 
| 53347 | 1227 | unfolding au and * using card_insert_disjoint[OF _ au(2)] | 
| 1228 | apply (auto intro!: exI[where x=u]) | |
| 1229 | done | |
| 33175 | 1230 | qed | 
| 1231 | qed | |
| 1232 | qed | |
| 53347 | 1233 | then show ?thesis | 
| 1234 | using compact_convex_combinations[OF assms Suc] by simp | |
| 33175 | 1235 | qed | 
| 36362 
06475a1547cb
fix lots of looping simp calls and other warnings
 huffman parents: 
36341diff
changeset | 1236 | qed | 
| 33175 | 1237 | qed | 
| 1238 | ||
| 53347 | 1239 | |
| 70136 | 1240 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Extremal points of a simplex are some vertices\<close> | 
| 33175 | 1241 | |
| 1242 | lemma dist_increases_online: | |
| 1243 | fixes a b d :: "'a::real_inner" | |
| 1244 | assumes "d \<noteq> 0" | |
| 1245 | shows "dist a (b + d) > dist a b \<or> dist a (b - d) > dist a b" | |
| 53347 | 1246 | proof (cases "inner a d - inner b d > 0") | 
| 1247 | case True | |
| 1248 | then have "0 < inner d d + (inner a d * 2 - inner b d * 2)" | |
| 1249 | apply (rule_tac add_pos_pos) | |
| 1250 | using assms | |
| 1251 | apply auto | |
| 1252 | done | |
| 1253 | then show ?thesis | |
| 1254 | apply (rule_tac disjI2) | |
| 1255 | unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff | |
| 1256 | apply (simp add: algebra_simps inner_commute) | |
| 1257 | done | |
| 33175 | 1258 | next | 
| 53347 | 1259 | case False | 
| 1260 | then have "0 < inner d d + (inner b d * 2 - inner a d * 2)" | |
| 1261 | apply (rule_tac add_pos_nonneg) | |
| 1262 | using assms | |
| 1263 | apply auto | |
| 1264 | done | |
| 1265 | then show ?thesis | |
| 1266 | apply (rule_tac disjI1) | |
| 1267 | unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff | |
| 1268 | apply (simp add: algebra_simps inner_commute) | |
| 1269 | done | |
| 33175 | 1270 | qed | 
| 1271 | ||
| 1272 | lemma norm_increases_online: | |
| 1273 | fixes d :: "'a::real_inner" | |
| 53347 | 1274 | shows "d \<noteq> 0 \<Longrightarrow> norm (a + d) > norm a \<or> norm(a - d) > norm a" | 
| 33175 | 1275 | using dist_increases_online[of d a 0] unfolding dist_norm by auto | 
| 1276 | ||
| 1277 | lemma simplex_furthest_lt: | |
| 68052 | 1278 | fixes S :: "'a::real_inner set" | 
| 1279 | assumes "finite S" | |
| 1280 | shows "\<forall>x \<in> convex hull S. x \<notin> S \<longrightarrow> (\<exists>y \<in> convex hull S. norm (x - a) < norm(y - a))" | |
| 53347 | 1281 | using assms | 
| 1282 | proof induct | |
| 68052 | 1283 | fix x S | 
| 1284 | assume as: "finite S" "x\<notin>S" "\<forall>x\<in>convex hull S. x \<notin> S \<longrightarrow> (\<exists>y\<in>convex hull S. norm (x - a) < norm (y - a))" | |
| 1285 | show "\<forall>xa\<in>convex hull insert x S. xa \<notin> insert x S \<longrightarrow> | |
| 1286 | (\<exists>y\<in>convex hull insert x S. norm (xa - a) < norm (y - a))" | |
| 1287 |   proof (intro impI ballI, cases "S = {}")
 | |
| 53347 | 1288 | case False | 
| 1289 | fix y | |
| 68052 | 1290 | assume y: "y \<in> convex hull insert x S" "y \<notin> insert x S" | 
| 1291 | obtain u v b where obt: "u\<ge>0" "v\<ge>0" "u + v = 1" "b \<in> convex hull S" "y = u *\<^sub>R x + v *\<^sub>R b" | |
| 33175 | 1292 | using y(1)[unfolded convex_hull_insert[OF False]] by auto | 
| 68052 | 1293 | show "\<exists>z\<in>convex hull insert x S. norm (y - a) < norm (z - a)" | 
| 1294 | proof (cases "y \<in> convex hull S") | |
| 53347 | 1295 | case True | 
| 68052 | 1296 | then obtain z where "z \<in> convex hull S" "norm (y - a) < norm (z - a)" | 
| 33175 | 1297 | using as(3)[THEN bspec[where x=y]] and y(2) by auto | 
| 53347 | 1298 | then show ?thesis | 
| 1299 | apply (rule_tac x=z in bexI) | |
| 1300 | unfolding convex_hull_insert[OF False] | |
| 1301 | apply auto | |
| 1302 | done | |
| 33175 | 1303 | next | 
| 53347 | 1304 | case False | 
| 1305 | show ?thesis | |
| 1306 | using obt(3) | |
| 1307 | proof (cases "u = 0", case_tac[!] "v = 0") | |
| 1308 | assume "u = 0" "v \<noteq> 0" | |
| 1309 | then have "y = b" using obt by auto | |
| 1310 | then show ?thesis using False and obt(4) by auto | |
| 33175 | 1311 | next | 
| 53347 | 1312 | assume "u \<noteq> 0" "v = 0" | 
| 1313 | then have "y = x" using obt by auto | |
| 1314 | then show ?thesis using y(2) by auto | |
| 1315 | next | |
| 1316 | assume "u \<noteq> 0" "v \<noteq> 0" | |
| 1317 | then obtain w where w: "w>0" "w<u" "w<v" | |
| 68527 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
68074diff
changeset | 1318 | using field_lbound_gt_zero[of u v] and obt(1,2) by auto | 
| 53347 | 1319 | have "x \<noteq> b" | 
| 1320 | proof | |
| 1321 | assume "x = b" | |
| 1322 | then have "y = b" unfolding obt(5) | |
| 68031 | 1323 | using obt(3) by (auto simp: scaleR_left_distrib[symmetric]) | 
| 53347 | 1324 | then show False using obt(4) and False by simp | 
| 1325 | qed | |
| 1326 | then have *: "w *\<^sub>R (x - b) \<noteq> 0" using w(1) by auto | |
| 1327 | show ?thesis | |
| 1328 | using dist_increases_online[OF *, of a y] | |
| 1329 | proof (elim disjE) | |
| 33175 | 1330 | assume "dist a y < dist a (y + w *\<^sub>R (x - b))" | 
| 53347 | 1331 | then have "norm (y - a) < norm ((u + w) *\<^sub>R x + (v - w) *\<^sub>R b - a)" | 
| 1332 | unfolding dist_commute[of a] | |
| 1333 | unfolding dist_norm obt(5) | |
| 1334 | by (simp add: algebra_simps) | |
| 68052 | 1335 | moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \<in> convex hull insert x S" | 
| 1336 |             unfolding convex_hull_insert[OF \<open>S\<noteq>{}\<close>]
 | |
| 1337 | proof (intro CollectI conjI exI) | |
| 1338 | show "u + w \<ge> 0" "v - w \<ge> 0" | |
| 1339 | using obt(1) w by auto | |
| 1340 | qed (use obt in auto) | |
| 33175 | 1341 | ultimately show ?thesis by auto | 
| 1342 | next | |
| 1343 | assume "dist a y < dist a (y - w *\<^sub>R (x - b))" | |
| 53347 | 1344 | then have "norm (y - a) < norm ((u - w) *\<^sub>R x + (v + w) *\<^sub>R b - a)" | 
| 1345 | unfolding dist_commute[of a] | |
| 1346 | unfolding dist_norm obt(5) | |
| 1347 | by (simp add: algebra_simps) | |
| 68052 | 1348 | moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \<in> convex hull insert x S" | 
| 1349 |             unfolding convex_hull_insert[OF \<open>S\<noteq>{}\<close>]
 | |
| 1350 | proof (intro CollectI conjI exI) | |
| 1351 | show "u - w \<ge> 0" "v + w \<ge> 0" | |
| 1352 | using obt(1) w by auto | |
| 1353 | qed (use obt in auto) | |
| 33175 | 1354 | ultimately show ?thesis by auto | 
| 1355 | qed | |
| 1356 | qed auto | |
| 1357 | qed | |
| 1358 | qed auto | |
| 68031 | 1359 | qed (auto simp: assms) | 
| 33175 | 1360 | |
| 1361 | lemma simplex_furthest_le: | |
| 68052 | 1362 | fixes S :: "'a::real_inner set" | 
| 1363 | assumes "finite S" | |
| 1364 |     and "S \<noteq> {}"
 | |
| 1365 | shows "\<exists>y\<in>S. \<forall>x\<in> convex hull S. norm (x - a) \<le> norm (y - a)" | |
| 53347 | 1366 | proof - | 
| 68052 | 1367 |   have "convex hull S \<noteq> {}"
 | 
| 1368 | using hull_subset[of S convex] and assms(2) by auto | |
| 1369 | then obtain x where x: "x \<in> convex hull S" "\<forall>y\<in>convex hull S. norm (y - a) \<le> norm (x - a)" | |
| 1370 | using distance_attains_sup[OF finite_imp_compact_convex_hull[OF \<open>finite S\<close>], of a] | |
| 53347 | 1371 | unfolding dist_commute[of a] | 
| 1372 | unfolding dist_norm | |
| 1373 | by auto | |
| 1374 | show ?thesis | |
| 68052 | 1375 | proof (cases "x \<in> S") | 
| 53347 | 1376 | case False | 
| 68052 | 1377 | then obtain y where "y \<in> convex hull S" "norm (x - a) < norm (y - a)" | 
| 53347 | 1378 | using simplex_furthest_lt[OF assms(1), THEN bspec[where x=x]] and x(1) | 
| 1379 | by auto | |
| 1380 | then show ?thesis | |
| 1381 | using x(2)[THEN bspec[where x=y]] by auto | |
| 1382 | next | |
| 1383 | case True | |
| 1384 | with x show ?thesis by auto | |
| 1385 | qed | |
| 33175 | 1386 | qed | 
| 1387 | ||
| 1388 | lemma simplex_furthest_le_exists: | |
| 68052 | 1389 |   fixes S :: "('a::real_inner) set"
 | 
| 1390 | shows "finite S \<Longrightarrow> \<forall>x\<in>(convex hull S). \<exists>y\<in>S. norm (x - a) \<le> norm (y - a)" | |
| 1391 |   using simplex_furthest_le[of S] by (cases "S = {}") auto
 | |
| 33175 | 1392 | |
| 1393 | lemma simplex_extremal_le: | |
| 68052 | 1394 | fixes S :: "'a::real_inner set" | 
| 1395 | assumes "finite S" | |
| 1396 |     and "S \<noteq> {}"
 | |
| 1397 | shows "\<exists>u\<in>S. \<exists>v\<in>S. \<forall>x\<in>convex hull S. \<forall>y \<in> convex hull S. norm (x - y) \<le> norm (u - v)" | |
| 53347 | 1398 | proof - | 
| 68052 | 1399 |   have "convex hull S \<noteq> {}"
 | 
| 1400 | using hull_subset[of S convex] and assms(2) by auto | |
| 1401 | then obtain u v where obt: "u \<in> convex hull S" "v \<in> convex hull S" | |
| 1402 | "\<forall>x\<in>convex hull S. \<forall>y\<in>convex hull S. norm (x - y) \<le> norm (u - v)" | |
| 53347 | 1403 | using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]] | 
| 1404 | by (auto simp: dist_norm) | |
| 1405 | then show ?thesis | |
| 68052 | 1406 | proof (cases "u\<notin>S \<or> v\<notin>S", elim disjE) | 
| 1407 | assume "u \<notin> S" | |
| 1408 | then obtain y where "y \<in> convex hull S" "norm (u - v) < norm (y - v)" | |
| 53347 | 1409 | using simplex_furthest_lt[OF assms(1), THEN bspec[where x=u]] and obt(1) | 
| 1410 | by auto | |
| 1411 | then show ?thesis | |
| 1412 | using obt(3)[THEN bspec[where x=y], THEN bspec[where x=v]] and obt(2) | |
| 1413 | by auto | |
| 33175 | 1414 | next | 
| 68052 | 1415 | assume "v \<notin> S" | 
| 1416 | then obtain y where "y \<in> convex hull S" "norm (v - u) < norm (y - u)" | |
| 53347 | 1417 | using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2) | 
| 1418 | by auto | |
| 1419 | then show ?thesis | |
| 1420 | using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1) | |
| 68031 | 1421 | by (auto simp: norm_minus_commute) | 
| 33175 | 1422 | qed auto | 
| 49531 | 1423 | qed | 
| 33175 | 1424 | |
| 1425 | lemma simplex_extremal_le_exists: | |
| 68052 | 1426 | fixes S :: "'a::real_inner set" | 
| 1427 | shows "finite S \<Longrightarrow> x \<in> convex hull S \<Longrightarrow> y \<in> convex hull S \<Longrightarrow> | |
| 1428 | \<exists>u\<in>S. \<exists>v\<in>S. norm (x - y) \<le> norm (u - v)" | |
| 1429 | using convex_hull_empty simplex_extremal_le[of S] | |
| 1430 |   by(cases "S = {}") auto
 | |
| 53347 | 1431 | |
| 33175 | 1432 | |
| 67968 | 1433 | subsection \<open>Closest point of a convex set is unique, with a continuous projection\<close> | 
| 33175 | 1434 | |
| 70136 | 1435 | definition\<^marker>\<open>tag important\<close> closest_point :: "'a::{real_inner,heine_borel} set \<Rightarrow> 'a \<Rightarrow> 'a"
 | 
| 68052 | 1436 | where "closest_point S a = (SOME x. x \<in> S \<and> (\<forall>y\<in>S. dist a x \<le> dist a y))" | 
| 33175 | 1437 | |
| 1438 | lemma closest_point_exists: | |
| 68052 | 1439 | assumes "closed S" | 
| 1440 |     and "S \<noteq> {}"
 | |
| 1441 | shows "closest_point S a \<in> S" | |
| 1442 | and "\<forall>y\<in>S. dist a (closest_point S a) \<le> dist a y" | |
| 53347 | 1443 | unfolding closest_point_def | 
| 1444 | apply(rule_tac[!] someI2_ex) | |
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1445 | apply (auto intro: distance_attains_inf[OF assms(1,2), of a]) | 
| 53347 | 1446 | done | 
| 1447 | ||
| 68052 | 1448 | lemma closest_point_in_set: "closed S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> closest_point S a \<in> S"
 | 
| 53347 | 1449 | by (meson closest_point_exists) | 
| 1450 | ||
| 68052 | 1451 | lemma closest_point_le: "closed S \<Longrightarrow> x \<in> S \<Longrightarrow> dist a (closest_point S a) \<le> dist a x" | 
| 1452 | using closest_point_exists[of S] by auto | |
| 33175 | 1453 | |
| 1454 | lemma closest_point_self: | |
| 68052 | 1455 | assumes "x \<in> S" | 
| 1456 | shows "closest_point S x = x" | |
| 53347 | 1457 | unfolding closest_point_def | 
| 1458 | apply (rule some1_equality, rule ex1I[of _ x]) | |
| 1459 | using assms | |
| 1460 | apply auto | |
| 1461 | done | |
| 1462 | ||
| 68052 | 1463 | lemma closest_point_refl: "closed S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> closest_point S x = x \<longleftrightarrow> x \<in> S"
 | 
| 1464 | using closest_point_in_set[of S x] closest_point_self[of x S] | |
| 53347 | 1465 | by auto | 
| 33175 | 1466 | |
| 36337 | 1467 | lemma closer_points_lemma: | 
| 33175 | 1468 | assumes "inner y z > 0" | 
| 1469 | shows "\<exists>u>0. \<forall>v>0. v \<le> u \<longrightarrow> norm(v *\<^sub>R z - y) < norm y" | |
| 53347 | 1470 | proof - | 
| 1471 | have z: "inner z z > 0" | |
| 1472 | unfolding inner_gt_zero_iff using assms by auto | |
| 68031 | 1473 | have "norm (v *\<^sub>R z - y) < norm y" | 
| 1474 | if "0 < v" and "v \<le> inner y z / inner z z" for v | |
| 1475 | unfolding norm_lt using z assms that | |
| 1476 | by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ \<open>0<v\<close>]) | |
| 53347 | 1477 | then show ?thesis | 
| 68031 | 1478 | using assms z | 
| 1479 | by (rule_tac x = "inner y z / inner z z" in exI) auto | |
| 53347 | 1480 | qed | 
| 33175 | 1481 | |
| 1482 | lemma closer_point_lemma: | |
| 1483 | assumes "inner (y - x) (z - x) > 0" | |
| 1484 | shows "\<exists>u>0. u \<le> 1 \<and> dist (x + u *\<^sub>R (z - x)) y < dist x y" | |
| 53347 | 1485 | proof - | 
| 1486 | obtain u where "u > 0" | |
| 1487 | and u: "\<forall>v>0. v \<le> u \<longrightarrow> norm (v *\<^sub>R (z - x) - (y - x)) < norm (y - x)" | |
| 33175 | 1488 | using closer_points_lemma[OF assms] by auto | 
| 53347 | 1489 | show ?thesis | 
| 1490 | apply (rule_tac x="min u 1" in exI) | |
| 60420 | 1491 | using u[THEN spec[where x="min u 1"]] and \<open>u > 0\<close> | 
| 68031 | 1492 | unfolding dist_norm by (auto simp: norm_minus_commute field_simps) | 
| 53347 | 1493 | qed | 
| 33175 | 1494 | |
| 1495 | lemma any_closest_point_dot: | |
| 68052 | 1496 | assumes "convex S" "closed S" "x \<in> S" "y \<in> S" "\<forall>z\<in>S. dist a x \<le> dist a z" | 
| 33175 | 1497 | shows "inner (a - x) (y - x) \<le> 0" | 
| 53347 | 1498 | proof (rule ccontr) | 
| 1499 | assume "\<not> ?thesis" | |
| 1500 | then obtain u where u: "u>0" "u\<le>1" "dist (x + u *\<^sub>R (y - x)) a < dist x a" | |
| 1501 | using closer_point_lemma[of a x y] by auto | |
| 1502 | let ?z = "(1 - u) *\<^sub>R x + u *\<^sub>R y" | |
| 68052 | 1503 | have "?z \<in> S" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61222diff
changeset | 1504 | using convexD_alt[OF assms(1,3,4), of u] using u by auto | 
| 53347 | 1505 | then show False | 
| 1506 | using assms(5)[THEN bspec[where x="?z"]] and u(3) | |
| 68031 | 1507 | by (auto simp: dist_commute algebra_simps) | 
| 53347 | 1508 | qed | 
| 33175 | 1509 | |
| 1510 | lemma any_closest_point_unique: | |
| 36337 | 1511 | fixes x :: "'a::real_inner" | 
| 68052 | 1512 | assumes "convex S" "closed S" "x \<in> S" "y \<in> S" | 
| 1513 | "\<forall>z\<in>S. dist a x \<le> dist a z" "\<forall>z\<in>S. dist a y \<le> dist a z" | |
| 53347 | 1514 | shows "x = y" | 
| 1515 | using any_closest_point_dot[OF assms(1-4,5)] and any_closest_point_dot[OF assms(1-2,4,3,6)] | |
| 33175 | 1516 | unfolding norm_pths(1) and norm_le_square | 
| 68031 | 1517 | by (auto simp: algebra_simps) | 
| 33175 | 1518 | |
| 1519 | lemma closest_point_unique: | |
| 68052 | 1520 | assumes "convex S" "closed S" "x \<in> S" "\<forall>z\<in>S. dist a x \<le> dist a z" | 
| 1521 | shows "x = closest_point S a" | |
| 1522 | using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point S a"] | |
| 33175 | 1523 | using closest_point_exists[OF assms(2)] and assms(3) by auto | 
| 1524 | ||
| 1525 | lemma closest_point_dot: | |
| 68052 | 1526 | assumes "convex S" "closed S" "x \<in> S" | 
| 1527 | shows "inner (a - closest_point S a) (x - closest_point S a) \<le> 0" | |
| 53347 | 1528 | apply (rule any_closest_point_dot[OF assms(1,2) _ assms(3)]) | 
| 1529 | using closest_point_exists[OF assms(2)] and assms(3) | |
| 1530 | apply auto | |
| 1531 | done | |
| 33175 | 1532 | |
| 1533 | lemma closest_point_lt: | |
| 68052 | 1534 | assumes "convex S" "closed S" "x \<in> S" "x \<noteq> closest_point S a" | 
| 1535 | shows "dist a (closest_point S a) < dist a x" | |
| 53347 | 1536 | apply (rule ccontr) | 
| 1537 | apply (rule_tac notE[OF assms(4)]) | |
| 1538 | apply (rule closest_point_unique[OF assms(1-3), of a]) | |
| 1539 | using closest_point_le[OF assms(2), of _ a] | |
| 1540 | apply fastforce | |
| 1541 | done | |
| 33175 | 1542 | |
| 69618 | 1543 | lemma setdist_closest_point: | 
| 1544 |     "\<lbrakk>closed S; S \<noteq> {}\<rbrakk> \<Longrightarrow> setdist {a} S = dist a (closest_point S a)"
 | |
| 1545 | apply (rule setdist_unique) | |
| 1546 | using closest_point_le | |
| 1547 | apply (auto simp: closest_point_in_set) | |
| 1548 | done | |
| 1549 | ||
| 33175 | 1550 | lemma closest_point_lipschitz: | 
| 68052 | 1551 | assumes "convex S" | 
| 1552 |     and "closed S" "S \<noteq> {}"
 | |
| 1553 | shows "dist (closest_point S x) (closest_point S y) \<le> dist x y" | |
| 53347 | 1554 | proof - | 
| 68052 | 1555 | have "inner (x - closest_point S x) (closest_point S y - closest_point S x) \<le> 0" | 
| 1556 | and "inner (y - closest_point S y) (closest_point S x - closest_point S y) \<le> 0" | |
| 53347 | 1557 | apply (rule_tac[!] any_closest_point_dot[OF assms(1-2)]) | 
| 1558 | using closest_point_exists[OF assms(2-3)] | |
| 1559 | apply auto | |
| 1560 | done | |
| 1561 | then show ?thesis unfolding dist_norm and norm_le | |
| 68052 | 1562 | using inner_ge_zero[of "(x - closest_point S x) - (y - closest_point S y)"] | 
| 53347 | 1563 | by (simp add: inner_add inner_diff inner_commute) | 
| 1564 | qed | |
| 33175 | 1565 | |
| 1566 | lemma continuous_at_closest_point: | |
| 68052 | 1567 | assumes "convex S" | 
| 1568 | and "closed S" | |
| 1569 |     and "S \<noteq> {}"
 | |
| 1570 | shows "continuous (at x) (closest_point S)" | |
| 49531 | 1571 | unfolding continuous_at_eps_delta | 
| 33175 | 1572 | using le_less_trans[OF closest_point_lipschitz[OF assms]] by auto | 
| 1573 | ||
| 1574 | lemma continuous_on_closest_point: | |
| 68052 | 1575 | assumes "convex S" | 
| 1576 | and "closed S" | |
| 1577 |     and "S \<noteq> {}"
 | |
| 1578 | shows "continuous_on t (closest_point S)" | |
| 53347 | 1579 | by (metis continuous_at_imp_continuous_on continuous_at_closest_point[OF assms]) | 
| 1580 | ||
| 63881 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1581 | proposition closest_point_in_rel_interior: | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1582 |   assumes "closed S" "S \<noteq> {}" and x: "x \<in> affine hull S"
 | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1583 | shows "closest_point S x \<in> rel_interior S \<longleftrightarrow> x \<in> rel_interior S" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1584 | proof (cases "x \<in> S") | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1585 | case True | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1586 | then show ?thesis | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1587 | by (simp add: closest_point_self) | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1588 | next | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1589 | case False | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1590 | then have "False" if asm: "closest_point S x \<in> rel_interior S" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1591 | proof - | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1592 | obtain e where "e > 0" and clox: "closest_point S x \<in> S" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1593 | and e: "cball (closest_point S x) e \<inter> affine hull S \<subseteq> S" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1594 | using asm mem_rel_interior_cball by blast | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1595 | then have clo_notx: "closest_point S x \<noteq> x" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1596 | using \<open>x \<notin> S\<close> by auto | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1597 | define y where "y \<equiv> closest_point S x - | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1598 | (min 1 (e / norm(closest_point S x - x))) *\<^sub>R (closest_point S x - x)" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1599 | have "x - y = (1 - min 1 (e / norm (closest_point S x - x))) *\<^sub>R (x - closest_point S x)" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1600 | by (simp add: y_def algebra_simps) | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1601 | then have "norm (x - y) = abs ((1 - min 1 (e / norm (closest_point S x - x)))) * norm(x - closest_point S x)" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1602 | by simp | 
| 68031 | 1603 | also have "\<dots> < norm(x - closest_point S x)" | 
| 63881 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1604 | using clo_notx \<open>e > 0\<close> | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1605 | by (auto simp: mult_less_cancel_right2 divide_simps) | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1606 | finally have no_less: "norm (x - y) < norm (x - closest_point S x)" . | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1607 | have "y \<in> affine hull S" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1608 | unfolding y_def | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1609 | by (meson affine_affine_hull clox hull_subset mem_affine_3_minus2 subsetD x) | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1610 | moreover have "dist (closest_point S x) y \<le> e" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1611 | using \<open>e > 0\<close> by (auto simp: y_def min_mult_distrib_right) | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1612 | ultimately have "y \<in> S" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1613 | using subsetD [OF e] by simp | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1614 | then have "dist x (closest_point S x) \<le> dist x y" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1615 | by (simp add: closest_point_le \<open>closed S\<close>) | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1616 | with no_less show False | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1617 | by (simp add: dist_norm) | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1618 | qed | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1619 | moreover have "x \<notin> rel_interior S" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1620 | using rel_interior_subset False by blast | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1621 | ultimately show ?thesis by blast | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1622 | qed | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1623 | |
| 33175 | 1624 | |
| 70136 | 1625 | subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Various point-to-set separating/supporting hyperplane theorems\<close> | 
| 33175 | 1626 | |
| 1627 | lemma supporting_hyperplane_closed_point: | |
| 36337 | 1628 |   fixes z :: "'a::{real_inner,heine_borel}"
 | 
| 68052 | 1629 | assumes "convex S" | 
| 1630 | and "closed S" | |
| 1631 |     and "S \<noteq> {}"
 | |
| 1632 | and "z \<notin> S" | |
| 1633 | shows "\<exists>a b. \<exists>y\<in>S. inner a z < b \<and> inner a y = b \<and> (\<forall>x\<in>S. inner a x \<ge> b)" | |
| 53347 | 1634 | proof - | 
| 68052 | 1635 | obtain y where "y \<in> S" and y: "\<forall>x\<in>S. dist z y \<le> dist z x" | 
| 63075 
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
 paulson <lp15@cam.ac.uk> parents: 
63072diff
changeset | 1636 | by (metis distance_attains_inf[OF assms(2-3)]) | 
| 53347 | 1637 | show ?thesis | 
| 68052 | 1638 | proof (intro exI bexI conjI ballI) | 
| 1639 | show "(y - z) \<bullet> z < (y - z) \<bullet> y" | |
| 1640 | by (metis \<open>y \<in> S\<close> assms(4) diff_gt_0_iff_gt inner_commute inner_diff_left inner_gt_zero_iff right_minus_eq) | |
| 1641 | show "(y - z) \<bullet> y \<le> (y - z) \<bullet> x" if "x \<in> S" for x | |
| 1642 | proof (rule ccontr) | |
| 1643 | have *: "\<And>u. 0 \<le> u \<and> u \<le> 1 \<longrightarrow> dist z y \<le> dist z ((1 - u) *\<^sub>R y + u *\<^sub>R x)" | |
| 1644 | using assms(1)[unfolded convex_alt] and y and \<open>x\<in>S\<close> and \<open>y\<in>S\<close> by auto | |
| 1645 | assume "\<not> (y - z) \<bullet> y \<le> (y - z) \<bullet> x" | |
| 1646 | then obtain v where "v > 0" "v \<le> 1" "dist (y + v *\<^sub>R (x - y)) z < dist y z" | |
| 1647 | using closer_point_lemma[of z y x] by (auto simp: inner_diff) | |
| 1648 | then show False | |
| 1649 | using *[of v] by (auto simp: dist_commute algebra_simps) | |
| 1650 | qed | |
| 1651 | qed (use \<open>y \<in> S\<close> in auto) | |
| 33175 | 1652 | qed | 
| 1653 | ||
| 1654 | lemma separating_hyperplane_closed_point: | |
| 36337 | 1655 |   fixes z :: "'a::{real_inner,heine_borel}"
 | 
| 68052 | 1656 | assumes "convex S" | 
| 1657 | and "closed S" | |
| 1658 | and "z \<notin> S" | |
| 1659 | shows "\<exists>a b. inner a z < b \<and> (\<forall>x\<in>S. inner a x > b)" | |
| 1660 | proof (cases "S = {}")
 | |
| 53347 | 1661 | case True | 
| 1662 | then show ?thesis | |
| 68052 | 1663 | by (simp add: gt_ex) | 
| 33175 | 1664 | next | 
| 53347 | 1665 | case False | 
| 68052 | 1666 | obtain y where "y \<in> S" and y: "\<And>x. x \<in> S \<Longrightarrow> dist z y \<le> dist z x" | 
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1667 | by (metis distance_attains_inf[OF assms(2) False]) | 
| 53347 | 1668 | show ?thesis | 
| 68052 | 1669 | proof (intro exI conjI ballI) | 
| 1670 | show "(y - z) \<bullet> z < inner (y - z) z + (norm (y - z))\<^sup>2 / 2" | |
| 1671 | using \<open>y\<in>S\<close> \<open>z\<notin>S\<close> by auto | |
| 1672 | next | |
| 53347 | 1673 | fix x | 
| 68052 | 1674 | assume "x \<in> S" | 
| 1675 | have "False" if *: "0 < inner (z - y) (x - y)" | |
| 53347 | 1676 | proof - | 
| 68052 | 1677 | obtain u where "u > 0" "u \<le> 1" "dist (y + u *\<^sub>R (x - y)) z < dist y z" | 
| 1678 | using * closer_point_lemma by blast | |
| 1679 | then show False using y[of "y + u *\<^sub>R (x - y)"] convexD_alt [OF \<open>convex S\<close>] | |
| 1680 | using \<open>x\<in>S\<close> \<open>y\<in>S\<close> by (auto simp: dist_commute algebra_simps) | |
| 53347 | 1681 | qed | 
| 1682 | moreover have "0 < (norm (y - z))\<^sup>2" | |
| 68052 | 1683 | using \<open>y\<in>S\<close> \<open>z\<notin>S\<close> by auto | 
| 53347 | 1684 | then have "0 < inner (y - z) (y - z)" | 
| 1685 | unfolding power2_norm_eq_inner by simp | |
| 68052 | 1686 | ultimately show "(y - z) \<bullet> z + (norm (y - z))\<^sup>2 / 2 < (y - z) \<bullet> x" | 
| 1687 | by (force simp: field_simps power2_norm_eq_inner inner_commute inner_diff) | |
| 1688 | qed | |
| 33175 | 1689 | qed | 
| 1690 | ||
| 1691 | lemma separating_hyperplane_closed_0: | |
| 68052 | 1692 |   assumes "convex (S::('a::euclidean_space) set)"
 | 
| 1693 | and "closed S" | |
| 1694 | and "0 \<notin> S" | |
| 1695 | shows "\<exists>a b. a \<noteq> 0 \<and> 0 < b \<and> (\<forall>x\<in>S. inner a x > b)" | |
| 1696 | proof (cases "S = {}")
 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 1697 | case True | 
| 68052 | 1698 | have "(SOME i. i\<in>Basis) \<noteq> (0::'a)" | 
| 1699 | by (metis Basis_zero SOME_Basis) | |
| 53347 | 1700 | then show ?thesis | 
| 68052 | 1701 | using True zero_less_one by blast | 
| 53347 | 1702 | next | 
| 1703 | case False | |
| 1704 | then show ?thesis | |
| 1705 | using False using separating_hyperplane_closed_point[OF assms] | |
| 68052 | 1706 | by (metis all_not_in_conv inner_zero_left inner_zero_right less_eq_real_def not_le) | 
| 53347 | 1707 | qed | 
| 1708 | ||
| 33175 | 1709 | |
| 70136 | 1710 | subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Now set-to-set for closed/compact sets\<close> | 
| 33175 | 1711 | |
| 1712 | lemma separating_hyperplane_closed_compact: | |
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1713 | fixes S :: "'a::euclidean_space set" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1714 | assumes "convex S" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1715 | and "closed S" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1716 | and "convex T" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1717 | and "compact T" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1718 |     and "T \<noteq> {}"
 | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1719 |     and "S \<inter> T = {}"
 | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1720 | shows "\<exists>a b. (\<forall>x\<in>S. inner a x < b) \<and> (\<forall>x\<in>T. inner a x > b)" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1721 | proof (cases "S = {}")
 | 
| 33175 | 1722 | case True | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1723 | obtain b where b: "b > 0" "\<forall>x\<in>T. norm x \<le> b" | 
| 53347 | 1724 | using compact_imp_bounded[OF assms(4)] unfolding bounded_pos by auto | 
| 1725 | obtain z :: 'a where z: "norm z = b + 1" | |
| 1726 | using vector_choose_size[of "b + 1"] and b(1) by auto | |
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1727 | then have "z \<notin> T" using b(2)[THEN bspec[where x=z]] by auto | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1728 | then obtain a b where ab: "inner a z < b" "\<forall>x\<in>T. b < inner a x" | 
| 53347 | 1729 | using separating_hyperplane_closed_point[OF assms(3) compact_imp_closed[OF assms(4)], of z] | 
| 1730 | by auto | |
| 1731 | then show ?thesis | |
| 1732 | using True by auto | |
| 33175 | 1733 | next | 
| 53347 | 1734 | case False | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1735 | then obtain y where "y \<in> S" by auto | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1736 |   obtain a b where "0 < b" "\<forall>x \<in> (\<Union>x\<in> S. \<Union>y \<in> T. {x - y}). b < inner a x"
 | 
| 33175 | 1737 | using separating_hyperplane_closed_point[OF convex_differences[OF assms(1,3)], of 0] | 
| 53347 | 1738 | using closed_compact_differences[OF assms(2,4)] | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1739 | using assms(6) by auto | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1740 | then have ab: "\<forall>x\<in>S. \<forall>y\<in>T. b + inner a y < inner a x" | 
| 53347 | 1741 | apply - | 
| 1742 | apply rule | |
| 1743 | apply rule | |
| 1744 | apply (erule_tac x="x - y" in ballE) | |
| 68031 | 1745 | apply (auto simp: inner_diff) | 
| 53347 | 1746 | done | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69064diff
changeset | 1747 | define k where "k = (SUP x\<in>T. a \<bullet> x)" | 
| 53347 | 1748 | show ?thesis | 
| 1749 | apply (rule_tac x="-a" in exI) | |
| 1750 | apply (rule_tac x="-(k + b / 2)" in exI) | |
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54258diff
changeset | 1751 | apply (intro conjI ballI) | 
| 53347 | 1752 | unfolding inner_minus_left and neg_less_iff_less | 
| 1753 | proof - | |
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1754 | fix x assume "x \<in> T" | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54258diff
changeset | 1755 | then have "inner a x - b / 2 < k" | 
| 53347 | 1756 | unfolding k_def | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54258diff
changeset | 1757 | proof (subst less_cSUP_iff) | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1758 |       show "T \<noteq> {}" by fact
 | 
| 67399 | 1759 | show "bdd_above ((\<bullet>) a ` T)" | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1760 | using ab[rule_format, of y] \<open>y \<in> S\<close> | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54258diff
changeset | 1761 | by (intro bdd_aboveI2[where M="inner a y - b"]) (auto simp: field_simps intro: less_imp_le) | 
| 60420 | 1762 | qed (auto intro!: bexI[of _ x] \<open>0<b\<close>) | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54258diff
changeset | 1763 | then show "inner a x < k + b / 2" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54258diff
changeset | 1764 | by auto | 
| 33175 | 1765 | next | 
| 53347 | 1766 | fix x | 
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1767 | assume "x \<in> S" | 
| 53347 | 1768 | then have "k \<le> inner a x - b" | 
| 1769 | unfolding k_def | |
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54258diff
changeset | 1770 | apply (rule_tac cSUP_least) | 
| 53347 | 1771 | using assms(5) | 
| 1772 | using ab[THEN bspec[where x=x]] | |
| 1773 | apply auto | |
| 1774 | done | |
| 1775 | then show "k + b / 2 < inner a x" | |
| 60420 | 1776 | using \<open>0 < b\<close> by auto | 
| 33175 | 1777 | qed | 
| 1778 | qed | |
| 1779 | ||
| 1780 | lemma separating_hyperplane_compact_closed: | |
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1781 | fixes S :: "'a::euclidean_space set" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1782 | assumes "convex S" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1783 | and "compact S" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1784 |     and "S \<noteq> {}"
 | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1785 | and "convex T" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1786 | and "closed T" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1787 |     and "S \<inter> T = {}"
 | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1788 | shows "\<exists>a b. (\<forall>x\<in>S. inner a x < b) \<and> (\<forall>x\<in>T. inner a x > b)" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1789 | proof - | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 1790 | obtain a b where "(\<forall>x\<in>T. inner a x < b) \<and> (\<forall>x\<in>S. b < inner a x)" | 
| 53347 | 1791 | using separating_hyperplane_closed_compact[OF assms(4-5,1-2,3)] and assms(6) | 
| 1792 | by auto | |
| 1793 | then show ?thesis | |
| 1794 | apply (rule_tac x="-a" in exI) | |
| 68031 | 1795 | apply (rule_tac x="-b" in exI, auto) | 
| 53347 | 1796 | done | 
| 1797 | qed | |
| 1798 | ||
| 33175 | 1799 | |
| 70136 | 1800 | subsubsection\<^marker>\<open>tag unimportant\<close> \<open>General case without assuming closure and getting non-strict separation\<close> | 
| 33175 | 1801 | |
| 1802 | lemma separating_hyperplane_set_0: | |
| 68031 | 1803 | assumes "convex S" "(0::'a::euclidean_space) \<notin> S" | 
| 1804 | shows "\<exists>a. a \<noteq> 0 \<and> (\<forall>x\<in>S. 0 \<le> inner a x)" | |
| 53347 | 1805 | proof - | 
| 1806 |   let ?k = "\<lambda>c. {x::'a. 0 \<le> inner c x}"
 | |
| 68031 | 1807 |   have *: "frontier (cball 0 1) \<inter> \<Inter>f \<noteq> {}" if as: "f \<subseteq> ?k ` S" "finite f" for f
 | 
| 53347 | 1808 | proof - | 
| 68031 | 1809 | obtain c where c: "f = ?k ` c" "c \<subseteq> S" "finite c" | 
| 53347 | 1810 | using finite_subset_image[OF as(2,1)] by auto | 
| 1811 | then obtain a b where ab: "a \<noteq> 0" "0 < b" "\<forall>x\<in>convex hull c. b < inner a x" | |
| 33175 | 1812 | using separating_hyperplane_closed_0[OF convex_convex_hull, of c] | 
| 1813 | using finite_imp_compact_convex_hull[OF c(3), THEN compact_imp_closed] and assms(2) | |
| 53347 | 1814 | using subset_hull[of convex, OF assms(1), symmetric, of c] | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 1815 | by force | 
| 53347 | 1816 | then have "\<exists>x. norm x = 1 \<and> (\<forall>y\<in>c. 0 \<le> inner y x)" | 
| 1817 | apply (rule_tac x = "inverse(norm a) *\<^sub>R a" in exI) | |
| 1818 | using hull_subset[of c convex] | |
| 1819 | unfolding subset_eq and inner_scaleR | |
| 68031 | 1820 | by (auto simp: inner_commute del: ballE elim!: ballE) | 
| 53347 | 1821 |     then show "frontier (cball 0 1) \<inter> \<Inter>f \<noteq> {}"
 | 
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1822 | unfolding c(1) frontier_cball sphere_def dist_norm by auto | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1823 | qed | 
| 68031 | 1824 |   have "frontier (cball 0 1) \<inter> (\<Inter>(?k ` S)) \<noteq> {}"
 | 
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1825 | apply (rule compact_imp_fip) | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1826 | apply (rule compact_frontier[OF compact_cball]) | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1827 | using * closed_halfspace_ge | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1828 | by auto | 
| 68031 | 1829 | then obtain x where "norm x = 1" "\<forall>y\<in>S. x\<in>?k y" | 
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1830 | unfolding frontier_cball dist_norm sphere_def by auto | 
| 53347 | 1831 | then show ?thesis | 
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62131diff
changeset | 1832 | by (metis inner_commute mem_Collect_eq norm_eq_zero zero_neq_one) | 
| 53347 | 1833 | qed | 
| 33175 | 1834 | |
| 1835 | lemma separating_hyperplane_sets: | |
| 68031 | 1836 | fixes S T :: "'a::euclidean_space set" | 
| 1837 | assumes "convex S" | |
| 1838 | and "convex T" | |
| 1839 |     and "S \<noteq> {}"
 | |
| 1840 |     and "T \<noteq> {}"
 | |
| 1841 |     and "S \<inter> T = {}"
 | |
| 1842 | shows "\<exists>a b. a \<noteq> 0 \<and> (\<forall>x\<in>S. inner a x \<le> b) \<and> (\<forall>x\<in>T. inner a x \<ge> b)" | |
| 53347 | 1843 | proof - | 
| 1844 | from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]] | |
| 68031 | 1845 |   obtain a where "a \<noteq> 0" "\<forall>x\<in>{x - y |x y. x \<in> T \<and> y \<in> S}. 0 \<le> inner a x"
 | 
| 1846 | using assms(3-5) by force | |
| 1847 | then have *: "\<And>x y. x \<in> T \<Longrightarrow> y \<in> S \<Longrightarrow> inner a y \<le> inner a x" | |
| 1848 | by (force simp: inner_diff) | |
| 1849 | then have bdd: "bdd_above (((\<bullet>) a)`S)" | |
| 1850 |     using \<open>T \<noteq> {}\<close> by (auto intro: bdd_aboveI2[OF *])
 | |
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54258diff
changeset | 1851 | show ?thesis | 
| 60420 | 1852 | using \<open>a\<noteq>0\<close> | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69064diff
changeset | 1853 | by (intro exI[of _ a] exI[of _ "SUP x\<in>S. a \<bullet> x"]) | 
| 68031 | 1854 |        (auto intro!: cSUP_upper bdd cSUP_least \<open>a \<noteq> 0\<close> \<open>S \<noteq> {}\<close> *)
 | 
| 60420 | 1855 | qed | 
| 1856 | ||
| 1857 | ||
| 70136 | 1858 | subsection\<^marker>\<open>tag unimportant\<close> \<open>More convexity generalities\<close> | 
| 33175 | 1859 | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1860 | lemma convex_closure [intro,simp]: | 
| 68031 | 1861 | fixes S :: "'a::real_normed_vector set" | 
| 1862 | assumes "convex S" | |
| 1863 | shows "convex (closure S)" | |
| 53676 | 1864 | apply (rule convexI) | 
| 1865 | apply (unfold closure_sequential, elim exE) | |
| 1866 | apply (rule_tac x="\<lambda>n. u *\<^sub>R xa n + v *\<^sub>R xb n" in exI) | |
| 53347 | 1867 | apply (rule,rule) | 
| 53676 | 1868 | apply (rule convexD [OF assms]) | 
| 53347 | 1869 | apply (auto del: tendsto_const intro!: tendsto_intros) | 
| 1870 | done | |
| 33175 | 1871 | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1872 | lemma convex_interior [intro,simp]: | 
| 68031 | 1873 | fixes S :: "'a::real_normed_vector set" | 
| 1874 | assumes "convex S" | |
| 1875 | shows "convex (interior S)" | |
| 53347 | 1876 | unfolding convex_alt Ball_def mem_interior | 
| 68052 | 1877 | proof clarify | 
| 53347 | 1878 | fix x y u | 
| 1879 | assume u: "0 \<le> u" "u \<le> (1::real)" | |
| 1880 | fix e d | |
| 68031 | 1881 | assume ed: "ball x e \<subseteq> S" "ball y d \<subseteq> S" "0<d" "0<e" | 
| 1882 | show "\<exists>e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \<subseteq> S" | |
| 68052 | 1883 | proof (intro exI conjI subsetI) | 
| 53347 | 1884 | fix z | 
| 1885 | assume "z \<in> ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) (min d e)" | |
| 68031 | 1886 | then have "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \<in> S" | 
| 53347 | 1887 | apply (rule_tac assms[unfolded convex_alt, rule_format]) | 
| 1888 | using ed(1,2) and u | |
| 1889 | unfolding subset_eq mem_ball Ball_def dist_norm | |
| 68031 | 1890 | apply (auto simp: algebra_simps) | 
| 53347 | 1891 | done | 
| 68031 | 1892 | then show "z \<in> S" | 
| 1893 | using u by (auto simp: algebra_simps) | |
| 53347 | 1894 | qed(insert u ed(3-4), auto) | 
| 1895 | qed | |
| 33175 | 1896 | |
| 68031 | 1897 | lemma convex_hull_eq_empty[simp]: "convex hull S = {} \<longleftrightarrow> S = {}"
 | 
| 1898 | using hull_subset[of S convex] convex_hull_empty by auto | |
| 33175 | 1899 | |
| 53347 | 1900 | |
| 70136 | 1901 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Convex set as intersection of halfspaces\<close> | 
| 33175 | 1902 | |
| 1903 | lemma convex_halfspace_intersection: | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 1904 |   fixes s :: "('a::euclidean_space) set"
 | 
| 33175 | 1905 | assumes "closed s" "convex s" | 
| 60585 | 1906 |   shows "s = \<Inter>{h. s \<subseteq> h \<and> (\<exists>a b. h = {x. inner a x \<le> b})}"
 | 
| 68031 | 1907 | apply (rule set_eqI, rule) | 
| 53347 | 1908 | unfolding Inter_iff Ball_def mem_Collect_eq | 
| 1909 | apply (rule,rule,erule conjE) | |
| 1910 | proof - | |
| 54465 | 1911 | fix x | 
| 53347 | 1912 |   assume "\<forall>xa. s \<subseteq> xa \<and> (\<exists>a b. xa = {x. inner a x \<le> b}) \<longrightarrow> x \<in> xa"
 | 
| 1913 |   then have "\<forall>a b. s \<subseteq> {x. inner a x \<le> b} \<longrightarrow> x \<in> {x. inner a x \<le> b}"
 | |
| 1914 | by blast | |
| 1915 | then show "x \<in> s" | |
| 1916 | apply (rule_tac ccontr) | |
| 1917 | apply (drule separating_hyperplane_closed_point[OF assms(2,1)]) | |
| 1918 | apply (erule exE)+ | |
| 1919 | apply (erule_tac x="-a" in allE) | |
| 68031 | 1920 | apply (erule_tac x="-b" in allE, auto) | 
| 53347 | 1921 | done | 
| 33175 | 1922 | qed auto | 
| 1923 | ||
| 53347 | 1924 | |
| 70136 | 1925 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Convexity of general and special intervals\<close> | 
| 33175 | 1926 | |
| 1927 | lemma is_interval_convex: | |
| 68052 | 1928 | fixes S :: "'a::euclidean_space set" | 
| 1929 | assumes "is_interval S" | |
| 1930 | shows "convex S" | |
| 37732 
6432bf0d7191
generalize type of is_interval to class euclidean_space
 huffman parents: 
37673diff
changeset | 1931 | proof (rule convexI) | 
| 53348 | 1932 | fix x y and u v :: real | 
| 68052 | 1933 | assume as: "x \<in> S" "y \<in> S" "0 \<le> u" "0 \<le> v" "u + v = 1" | 
| 53348 | 1934 | then have *: "u = 1 - v" "1 - v \<ge> 0" and **: "v = 1 - u" "1 - u \<ge> 0" | 
| 1935 | by auto | |
| 1936 |   {
 | |
| 1937 | fix a b | |
| 1938 | assume "\<not> b \<le> u * a + v * b" | |
| 1939 | then have "u * a < (1 - v) * b" | |
| 68031 | 1940 | unfolding not_le using as(4) by (auto simp: field_simps) | 
| 53348 | 1941 | then have "a < b" | 
| 1942 | unfolding * using as(4) *(2) | |
| 1943 | apply (rule_tac mult_left_less_imp_less[of "1 - v"]) | |
| 68031 | 1944 | apply (auto simp: field_simps) | 
| 53348 | 1945 | done | 
| 1946 | then have "a \<le> u * a + v * b" | |
| 1947 | unfolding * using as(4) | |
| 68031 | 1948 | by (auto simp: field_simps intro!:mult_right_mono) | 
| 53348 | 1949 | } | 
| 1950 | moreover | |
| 1951 |   {
 | |
| 1952 | fix a b | |
| 1953 | assume "\<not> u * a + v * b \<le> a" | |
| 1954 | then have "v * b > (1 - u) * a" | |
| 68031 | 1955 | unfolding not_le using as(4) by (auto simp: field_simps) | 
| 53348 | 1956 | then have "a < b" | 
| 1957 | unfolding * using as(4) | |
| 1958 | apply (rule_tac mult_left_less_imp_less) | |
| 68031 | 1959 | apply (auto simp: field_simps) | 
| 53348 | 1960 | done | 
| 1961 | then have "u * a + v * b \<le> b" | |
| 1962 | unfolding ** | |
| 1963 | using **(2) as(3) | |
| 68031 | 1964 | by (auto simp: field_simps intro!:mult_right_mono) | 
| 53348 | 1965 | } | 
| 68052 | 1966 | ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> S" | 
| 53348 | 1967 | apply - | 
| 1968 | apply (rule assms[unfolded is_interval_def, rule_format, OF as(1,2)]) | |
| 1969 | using as(3-) DIM_positive[where 'a='a] | |
| 1970 | apply (auto simp: inner_simps) | |
| 1971 | done | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 1972 | qed | 
| 33175 | 1973 | |
| 1974 | lemma is_interval_connected: | |
| 68052 | 1975 | fixes S :: "'a::euclidean_space set" | 
| 1976 | shows "is_interval S \<Longrightarrow> connected S" | |
| 33175 | 1977 | using is_interval_convex convex_connected by auto | 
| 1978 | ||
| 62618 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1979 | lemma convex_box [simp]: "convex (cbox a b)" "convex (box a (b::'a::euclidean_space))" | 
| 56188 | 1980 | apply (rule_tac[!] is_interval_convex)+ | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1981 | using is_interval_box is_interval_cbox | 
| 53348 | 1982 | apply auto | 
| 1983 | done | |
| 33175 | 1984 | |
| 63928 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 1985 | text\<open>A non-singleton connected set is perfect (i.e. has no isolated points). \<close> | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 1986 | lemma connected_imp_perfect: | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 1987 | fixes a :: "'a::metric_space" | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 1988 |   assumes "connected S" "a \<in> S" and S: "\<And>x. S \<noteq> {x}"
 | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 1989 | shows "a islimpt S" | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 1990 | proof - | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 1991 | have False if "a \<in> T" "open T" "\<And>y. \<lbrakk>y \<in> S; y \<in> T\<rbrakk> \<Longrightarrow> y = a" for T | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 1992 | proof - | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 1993 | obtain e where "e > 0" and e: "cball a e \<subseteq> T" | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 1994 | using \<open>open T\<close> \<open>a \<in> T\<close> by (auto simp: open_contains_cball) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69710diff
changeset | 1995 |     have "openin (top_of_set S) {a}"
 | 
| 63928 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 1996 | unfolding openin_open using that \<open>a \<in> S\<close> by blast | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69710diff
changeset | 1997 |     moreover have "closedin (top_of_set S) {a}"
 | 
| 63928 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 1998 | by (simp add: assms) | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 1999 | ultimately show "False" | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 2000 | using \<open>connected S\<close> connected_clopen S by blast | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 2001 | qed | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 2002 | then show ?thesis | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 2003 | unfolding islimpt_def by blast | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 2004 | qed | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 2005 | |
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 2006 | lemma connected_imp_perfect_aff_dim: | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 2007 | "\<lbrakk>connected S; aff_dim S \<noteq> 0; a \<in> S\<rbrakk> \<Longrightarrow> a islimpt S" | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 2008 | using aff_dim_sing connected_imp_perfect by blast | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 2009 | |
| 70136 | 2010 | subsection\<^marker>\<open>tag unimportant\<close> \<open>On \<open>real\<close>, \<open>is_interval\<close>, \<open>convex\<close> and \<open>connected\<close> are all equivalent\<close> | 
| 33175 | 2011 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 2012 | lemma mem_is_interval_1_I: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 2013 | fixes a b c::real | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 2014 | assumes "is_interval S" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 2015 | assumes "a \<in> S" "c \<in> S" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 2016 | assumes "a \<le> b" "b \<le> c" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 2017 | shows "b \<in> S" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 2018 | using assms is_interval_1 by blast | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2019 | |
| 54465 | 2020 | lemma is_interval_connected_1: | 
| 2021 | fixes s :: "real set" | |
| 2022 | shows "is_interval s \<longleftrightarrow> connected s" | |
| 2023 | apply rule | |
| 2024 | apply (rule is_interval_connected, assumption) | |
| 2025 | unfolding is_interval_1 | |
| 2026 | apply rule | |
| 2027 | apply rule | |
| 2028 | apply rule | |
| 2029 | apply rule | |
| 2030 | apply (erule conjE) | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2031 | apply (rule ccontr) | 
| 54465 | 2032 | proof - | 
| 2033 | fix a b x | |
| 2034 | assume as: "connected s" "a \<in> s" "b \<in> s" "a \<le> x" "x \<le> b" "x \<notin> s" | |
| 2035 | then have *: "a < x" "x < b" | |
| 2036 | unfolding not_le [symmetric] by auto | |
| 2037 |   let ?halfl = "{..<x} "
 | |
| 2038 |   let ?halfr = "{x<..}"
 | |
| 2039 |   {
 | |
| 2040 | fix y | |
| 2041 | assume "y \<in> s" | |
| 60420 | 2042 | with \<open>x \<notin> s\<close> have "x \<noteq> y" by auto | 
| 54465 | 2043 | then have "y \<in> ?halfr \<union> ?halfl" by auto | 
| 2044 | } | |
| 2045 | moreover have "a \<in> ?halfl" "b \<in> ?halfr" using * by auto | |
| 2046 |   then have "?halfl \<inter> s \<noteq> {}" "?halfr \<inter> s \<noteq> {}"
 | |
| 2047 | using as(2-3) by auto | |
| 2048 | ultimately show False | |
| 2049 | apply (rule_tac notE[OF as(1)[unfolded connected_def]]) | |
| 2050 | apply (rule_tac x = ?halfl in exI) | |
| 68031 | 2051 | apply (rule_tac x = ?halfr in exI, rule) | 
| 2052 | apply (rule open_lessThan, rule) | |
| 2053 | apply (rule open_greaterThan, auto) | |
| 54465 | 2054 | done | 
| 2055 | qed | |
| 33175 | 2056 | |
| 2057 | lemma is_interval_convex_1: | |
| 54465 | 2058 | fixes s :: "real set" | 
| 2059 | shows "is_interval s \<longleftrightarrow> convex s" | |
| 2060 | by (metis is_interval_convex convex_connected is_interval_connected_1) | |
| 33175 | 2061 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 2062 | lemma is_interval_ball_real: "is_interval (ball a b)" for a b::real | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 2063 | by (metis connected_ball is_interval_connected_1) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 2064 | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2065 | lemma connected_compact_interval_1: | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2066 |      "connected S \<and> compact S \<longleftrightarrow> (\<exists>a b. S = {a..b::real})"
 | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2067 | by (auto simp: is_interval_connected_1 [symmetric] is_interval_compact) | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2068 | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2069 | lemma connected_convex_1: | 
| 54465 | 2070 | fixes s :: "real set" | 
| 2071 | shows "connected s \<longleftrightarrow> convex s" | |
| 2072 | by (metis is_interval_convex convex_connected is_interval_connected_1) | |
| 53348 | 2073 | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2074 | lemma connected_convex_1_gen: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2075 | fixes s :: "'a :: euclidean_space set" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2076 |   assumes "DIM('a) = 1"
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2077 | shows "connected s \<longleftrightarrow> convex s" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2078 | proof - | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2079 | obtain f:: "'a \<Rightarrow> real" where linf: "linear f" and "inj f" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 2080 | using subspace_isomorphism[OF subspace_UNIV subspace_UNIV, | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 2081 | where 'a='a and 'b=real] | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 2082 | unfolding Euclidean_Space.dim_UNIV | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 2083 | by (auto simp: assms) | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2084 | then have "f -` (f ` s) = s" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2085 | by (simp add: inj_vimage_image_eq) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2086 | then show ?thesis | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2087 | by (metis connected_convex_1 convex_linear_vimage linf convex_connected connected_linear_image) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2088 | qed | 
| 53348 | 2089 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 2090 | lemma is_interval_cball_1[intro, simp]: "is_interval (cball a b)" for a b::real | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 2091 | by (simp add: is_interval_convex_1) | 
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 2092 | |
| 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: 
69922diff
changeset | 2093 | lemma [simp]: | 
| 
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: 
69922diff
changeset | 2094 | fixes r s::real | 
| 
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: 
69922diff
changeset | 2095 |   shows is_interval_io: "is_interval {..<r}"
 | 
| 
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: 
69922diff
changeset | 2096 |     and is_interval_oi: "is_interval {r<..}"
 | 
| 
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: 
69922diff
changeset | 2097 |     and is_interval_oo: "is_interval {r<..<s}"
 | 
| 
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: 
69922diff
changeset | 2098 |     and is_interval_oc: "is_interval {r<..s}"
 | 
| 
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: 
69922diff
changeset | 2099 |     and is_interval_co: "is_interval {r..<s}"
 | 
| 
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: 
69922diff
changeset | 2100 | by (simp_all add: is_interval_convex_1) | 
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 2101 | |
| 70136 | 2102 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Another intermediate value theorem formulation\<close> | 
| 33175 | 2103 | |
| 53348 | 2104 | lemma ivt_increasing_component_on_1: | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 2105 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" | 
| 53348 | 2106 | assumes "a \<le> b" | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2107 |     and "continuous_on {a..b} f"
 | 
| 53348 | 2108 | and "(f a)\<bullet>k \<le> y" "y \<le> (f b)\<bullet>k" | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2109 |   shows "\<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
 | 
| 56188 | 2110 | proof - | 
| 2111 | have "f a \<in> f ` cbox a b" "f b \<in> f ` cbox a b" | |
| 53348 | 2112 | apply (rule_tac[!] imageI) | 
| 2113 | using assms(1) | |
| 2114 | apply auto | |
| 2115 | done | |
| 2116 | then show ?thesis | |
| 56188 | 2117 | using connected_ivt_component[of "f ` cbox a b" "f a" "f b" k y] | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: 
66793diff
changeset | 2118 | by (simp add: connected_continuous_image assms) | 
| 53348 | 2119 | qed | 
| 2120 | ||
| 2121 | lemma ivt_increasing_component_1: | |
| 2122 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2123 |   shows "a \<le> b \<Longrightarrow> \<forall>x\<in>{a..b}. continuous (at x) f \<Longrightarrow>
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2124 |     f a\<bullet>k \<le> y \<Longrightarrow> y \<le> f b\<bullet>k \<Longrightarrow> \<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
 | 
| 68031 | 2125 | by (rule ivt_increasing_component_on_1) (auto simp: continuous_at_imp_continuous_on) | 
| 53348 | 2126 | |
| 2127 | lemma ivt_decreasing_component_on_1: | |
| 2128 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" | |
| 2129 | assumes "a \<le> b" | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2130 |     and "continuous_on {a..b} f"
 | 
| 53348 | 2131 | and "(f b)\<bullet>k \<le> y" | 
| 2132 | and "y \<le> (f a)\<bullet>k" | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2133 |   shows "\<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
 | 
| 53348 | 2134 | apply (subst neg_equal_iff_equal[symmetric]) | 
| 44531 
1d477a2b1572
replace some continuous_on lemmas with more general versions
 huffman parents: 
44525diff
changeset | 2135 | using ivt_increasing_component_on_1[of a b "\<lambda>x. - f x" k "- y"] | 
| 53348 | 2136 | using assms using continuous_on_minus | 
| 2137 | apply auto | |
| 2138 | done | |
| 2139 | ||
| 2140 | lemma ivt_decreasing_component_1: | |
| 2141 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2142 |   shows "a \<le> b \<Longrightarrow> \<forall>x\<in>{a..b}. continuous (at x) f \<Longrightarrow>
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 2143 |     f b\<bullet>k \<le> y \<Longrightarrow> y \<le> f a\<bullet>k \<Longrightarrow> \<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
 | 
| 53348 | 2144 | by (rule ivt_decreasing_component_on_1) (auto simp: continuous_at_imp_continuous_on) | 
| 2145 | ||
| 33175 | 2146 | |
| 70136 | 2147 | subsection\<^marker>\<open>tag unimportant\<close> \<open>A bound within an interval\<close> | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2148 | |
| 56188 | 2149 | lemma convex_hull_eq_real_cbox: | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2150 | fixes x y :: real assumes "x \<le> y" | 
| 56188 | 2151 |   shows "convex hull {x, y} = cbox x y"
 | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2152 | proof (rule hull_unique) | 
| 60420 | 2153 |   show "{x, y} \<subseteq> cbox x y" using \<open>x \<le> y\<close> by auto
 | 
| 56188 | 2154 | show "convex (cbox x y)" | 
| 2155 | by (rule convex_box) | |
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2156 | next | 
| 68058 | 2157 |   fix S assume "{x, y} \<subseteq> S" and "convex S"
 | 
| 2158 | then show "cbox x y \<subseteq> S" | |
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2159 | unfolding is_interval_convex_1 [symmetric] is_interval_def Basis_real_def | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2160 | by - (clarify, simp (no_asm_use), fast) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2161 | qed | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 2162 | |
| 33175 | 2163 | lemma unit_interval_convex_hull: | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 2164 |   "cbox (0::'a::euclidean_space) One = convex hull {x. \<forall>i\<in>Basis. (x\<bullet>i = 0) \<or> (x\<bullet>i = 1)}"
 | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 2165 | (is "?int = convex hull ?points") | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 2166 | proof - | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 2167 | have One[simp]: "\<And>i. i \<in> Basis \<Longrightarrow> One \<bullet> i = 1" | 
| 64267 | 2168 | by (simp add: inner_sum_left sum.If_cases inner_Basis) | 
| 56188 | 2169 |   have "?int = {x. \<forall>i\<in>Basis. x \<bullet> i \<in> cbox 0 1}"
 | 
| 2170 | by (auto simp: cbox_def) | |
| 2171 | also have "\<dots> = (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` cbox 0 1)" | |
| 64267 | 2172 | by (simp only: box_eq_set_sum_Basis) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2173 |   also have "\<dots> = (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` (convex hull {0, 1}))"
 | 
| 56188 | 2174 | by (simp only: convex_hull_eq_real_cbox zero_le_one) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2175 |   also have "\<dots> = (\<Sum>i\<in>Basis. convex hull ((\<lambda>x. x *\<^sub>R i) ` {0, 1}))"
 | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 2176 | by (simp add: convex_hull_linear_image) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2177 |   also have "\<dots> = convex hull (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` {0, 1})"
 | 
| 64267 | 2178 | by (simp only: convex_hull_set_sum) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2179 |   also have "\<dots> = convex hull {x. \<forall>i\<in>Basis. x\<bullet>i \<in> {0, 1}}"
 | 
| 64267 | 2180 | by (simp only: box_eq_set_sum_Basis) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2181 |   also have "convex hull {x. \<forall>i\<in>Basis. x\<bullet>i \<in> {0, 1}} = convex hull ?points"
 | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2182 | by simp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2183 | finally show ?thesis . | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 2184 | qed | 
| 33175 | 2185 | |
| 60420 | 2186 | text \<open>And this is a finite set of vertices.\<close> | 
| 33175 | 2187 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 2188 | lemma unit_cube_convex_hull: | 
| 68058 | 2189 | obtains S :: "'a::euclidean_space set" | 
| 2190 | where "finite S" and "cbox 0 (\<Sum>Basis) = convex hull S" | |
| 2191 | proof | |
| 2192 |   show "finite {x::'a. \<forall>i\<in>Basis. x \<bullet> i = 0 \<or> x \<bullet> i = 1}"
 | |
| 2193 | proof (rule finite_subset, clarify) | |
| 2194 | show "finite ((\<lambda>S. \<Sum>i\<in>Basis. (if i \<in> S then 1 else 0) *\<^sub>R i) ` Pow Basis)" | |
| 2195 | using finite_Basis by blast | |
| 2196 | fix x :: 'a | |
| 2197 | assume as: "\<forall>i\<in>Basis. x \<bullet> i = 0 \<or> x \<bullet> i = 1" | |
| 2198 | show "x \<in> (\<lambda>S. \<Sum>i\<in>Basis. (if i\<in>S then 1 else 0) *\<^sub>R i) ` Pow Basis" | |
| 2199 |       apply (rule image_eqI[where x="{i. i\<in>Basis \<and> x\<bullet>i = 1}"])
 | |
| 2200 | using as | |
| 2201 | apply (subst euclidean_eq_iff, auto) | |
| 2202 | done | |
| 2203 | qed | |
| 2204 |   show "cbox 0 One = convex hull {x. \<forall>i\<in>Basis. x \<bullet> i = 0 \<or> x \<bullet> i = 1}"
 | |
| 2205 | using unit_interval_convex_hull by blast | |
| 2206 | qed | |
| 33175 | 2207 | |
| 60420 | 2208 | text \<open>Hence any cube (could do any nonempty interval).\<close> | 
| 33175 | 2209 | |
| 2210 | lemma cube_convex_hull: | |
| 53348 | 2211 | assumes "d > 0" | 
| 68058 | 2212 | obtains S :: "'a::euclidean_space set" where | 
| 2213 | "finite S" and "cbox (x - (\<Sum>i\<in>Basis. d*\<^sub>Ri)) (x + (\<Sum>i\<in>Basis. d*\<^sub>Ri)) = convex hull S" | |
| 53348 | 2214 | proof - | 
| 68058 | 2215 | let ?d = "(\<Sum>i\<in>Basis. d *\<^sub>R i)::'a" | 
| 56188 | 2216 | have *: "cbox (x - ?d) (x + ?d) = (\<lambda>y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 (\<Sum>Basis)" | 
| 68058 | 2217 | proof (intro set_eqI iffI) | 
| 53348 | 2218 | fix y | 
| 68058 | 2219 | assume "y \<in> cbox (x - ?d) (x + ?d)" | 
| 56188 | 2220 | then have "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \<in> cbox 0 (\<Sum>Basis)" | 
| 58776 
95e58e04e534
use NO_MATCH-simproc for distribution rules in field_simps, otherwise field_simps on '(a / (c + d)) * (e + f)' can be non-terminating
 hoelzl parents: 
57865diff
changeset | 2221 | using assms by (simp add: mem_box field_simps inner_simps) | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68607diff
changeset | 2222 | with \<open>0 < d\<close> show "y \<in> (\<lambda>y. x - sum ((*\<^sub>R) d) Basis + (2 * d) *\<^sub>R y) ` cbox 0 One" | 
| 68058 | 2223 | by (auto intro: image_eqI[where x= "inverse (2 * d) *\<^sub>R (y - (x - ?d))"]) | 
| 33175 | 2224 | next | 
| 68058 | 2225 | fix y | 
| 2226 | assume "y \<in> (\<lambda>y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 One" | |
| 2227 | then obtain z where z: "z \<in> cbox 0 One" "y = x - ?d + (2*d) *\<^sub>R z" | |
| 68031 | 2228 | by auto | 
| 56188 | 2229 | then show "y \<in> cbox (x - ?d) (x + ?d)" | 
| 68058 | 2230 | using z assms by (auto simp: mem_box inner_simps) | 
| 53348 | 2231 | qed | 
| 68058 | 2232 | obtain S where "finite S" "cbox 0 (\<Sum>Basis::'a) = convex hull S" | 
| 53348 | 2233 | using unit_cube_convex_hull by auto | 
| 2234 | then show ?thesis | |
| 68058 | 2235 | by (rule_tac that[of "(\<lambda>y. x - ?d + (2 * d) *\<^sub>R y)` S"]) (auto simp: convex_hull_affinity *) | 
| 53348 | 2236 | qed | 
| 2237 | ||
| 70136 | 2238 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Representation of any interval as a finite convex hull\<close> | 
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2239 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2240 | lemma image_stretch_interval: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2241 | "(\<lambda>x. \<Sum>k\<in>Basis. (m k * (x\<bullet>k)) *\<^sub>R k) ` cbox a (b::'a::euclidean_space) = | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2242 |   (if (cbox a b) = {} then {} else
 | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2243 | cbox (\<Sum>k\<in>Basis. (min (m k * (a\<bullet>k)) (m k * (b\<bullet>k))) *\<^sub>R k::'a) | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2244 | (\<Sum>k\<in>Basis. (max (m k * (a\<bullet>k)) (m k * (b\<bullet>k))) *\<^sub>R k))" | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2245 | proof cases | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2246 |   assume *: "cbox a b \<noteq> {}"
 | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2247 | show ?thesis | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2248 | unfolding box_ne_empty if_not_P[OF *] | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2249 | apply (simp add: cbox_def image_Collect set_eq_iff euclidean_eq_iff[where 'a='a] ball_conj_distrib[symmetric]) | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2250 | apply (subst choice_Basis_iff[symmetric]) | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2251 | proof (intro allI ball_cong refl) | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2252 | fix x i :: 'a assume "i \<in> Basis" | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2253 | with * have a_le_b: "a \<bullet> i \<le> b \<bullet> i" | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2254 | unfolding box_ne_empty by auto | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2255 | show "(\<exists>xa. x \<bullet> i = m i * xa \<and> a \<bullet> i \<le> xa \<and> xa \<le> b \<bullet> i) \<longleftrightarrow> | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2256 | min (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) \<le> x \<bullet> i \<and> x \<bullet> i \<le> max (m i * (a \<bullet> i)) (m i * (b \<bullet> i))" | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2257 | proof (cases "m i = 0") | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2258 | case True | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2259 | with a_le_b show ?thesis by auto | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2260 | next | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2261 | case False | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2262 | then have *: "\<And>a b. a = m i * b \<longleftrightarrow> b = a / m i" | 
| 68031 | 2263 | by (auto simp: field_simps) | 
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2264 | from False have | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2265 | "min (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) = (if 0 < m i then m i * (a \<bullet> i) else m i * (b \<bullet> i))" | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2266 | "max (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) = (if 0 < m i then m i * (b \<bullet> i) else m i * (a \<bullet> i))" | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2267 | using a_le_b by (auto simp: min_def max_def mult_le_cancel_left) | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2268 | with False show ?thesis using a_le_b | 
| 68031 | 2269 | unfolding * by (auto simp: le_divide_eq divide_le_eq ac_simps) | 
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2270 | qed | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2271 | qed | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2272 | qed simp | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2273 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2274 | lemma interval_image_stretch_interval: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2275 | "\<exists>u v. (\<lambda>x. \<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k) ` cbox a (b::'a::euclidean_space) = cbox u (v::'a::euclidean_space)" | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2276 | unfolding image_stretch_interval by auto | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2277 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2278 | lemma cbox_translation: "cbox (c + a) (c + b) = image (\<lambda>x. c + x) (cbox a b)" | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2279 | using image_affinity_cbox [of 1 c a b] | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2280 | using box_ne_empty [of "a+c" "b+c"] box_ne_empty [of a b] | 
| 68031 | 2281 | by (auto simp: inner_left_distrib add.commute) | 
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2282 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2283 | lemma cbox_image_unit_interval: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2284 | fixes a :: "'a::euclidean_space" | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2285 |   assumes "cbox a b \<noteq> {}"
 | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2286 | shows "cbox a b = | 
| 67399 | 2287 | (+) a ` (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k) ` cbox 0 One" | 
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2288 | using assms | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2289 | apply (simp add: box_ne_empty image_stretch_interval cbox_translation [symmetric]) | 
| 64267 | 2290 | apply (simp add: min_def max_def algebra_simps sum_subtractf euclidean_representation) | 
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2291 | done | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2292 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2293 | lemma closed_interval_as_convex_hull: | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2294 | fixes a :: "'a::euclidean_space" | 
| 68058 | 2295 | obtains S where "finite S" "cbox a b = convex hull S" | 
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2296 | proof (cases "cbox a b = {}")
 | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2297 | case True with convex_hull_empty that show ?thesis | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2298 | by blast | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2299 | next | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2300 | case False | 
| 68058 | 2301 | obtain S::"'a set" where "finite S" and eq: "cbox 0 One = convex hull S" | 
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2302 | by (blast intro: unit_cube_convex_hull) | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2303 | have lin: "linear (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k)" | 
| 64267 | 2304 | by (rule linear_compose_sum) (auto simp: algebra_simps linearI) | 
| 68058 | 2305 | have "finite ((+) a ` (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k) ` S)" | 
| 2306 | by (rule finite_imageI \<open>finite S\<close>)+ | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2307 | then show ?thesis | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2308 | apply (rule that) | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2309 | apply (simp add: convex_hull_translation convex_hull_linear_image [OF lin, symmetric]) | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2310 | apply (simp add: eq [symmetric] cbox_image_unit_interval [OF False]) | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2311 | done | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2312 | qed | 
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62950diff
changeset | 2313 | |
| 33175 | 2314 | |
| 70136 | 2315 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Bounded convex function on open set is continuous\<close> | 
| 33175 | 2316 | |
| 2317 | lemma convex_on_bounded_continuous: | |
| 68058 | 2318 |   fixes S :: "('a::real_normed_vector) set"
 | 
| 2319 | assumes "open S" | |
| 2320 | and "convex_on S f" | |
| 2321 | and "\<forall>x\<in>S. \<bar>f x\<bar> \<le> b" | |
| 2322 | shows "continuous_on S f" | |
| 53348 | 2323 | apply (rule continuous_at_imp_continuous_on) | 
| 2324 | unfolding continuous_at_real_range | |
| 2325 | proof (rule,rule,rule) | |
| 2326 | fix x and e :: real | |
| 68058 | 2327 | assume "x \<in> S" "e > 0" | 
| 63040 | 2328 | define B where "B = \<bar>b\<bar> + 1" | 
| 68058 | 2329 | then have B: "0 < B""\<And>x. x\<in>S \<Longrightarrow> \<bar>f x\<bar> \<le> B" | 
| 2330 | using assms(3) by auto | |
| 2331 | obtain k where "k > 0" and k: "cball x k \<subseteq> S" | |
| 2332 | using \<open>x \<in> S\<close> assms(1) open_contains_cball_eq by blast | |
| 33175 | 2333 | show "\<exists>d>0. \<forall>x'. norm (x' - x) < d \<longrightarrow> \<bar>f x' - f x\<bar> < e" | 
| 68058 | 2334 | proof (intro exI conjI allI impI) | 
| 53348 | 2335 | fix y | 
| 2336 | assume as: "norm (y - x) < min (k / 2) (e / (2 * B) * k)" | |
| 2337 | show "\<bar>f y - f x\<bar> < e" | |
| 2338 | proof (cases "y = x") | |
| 2339 | case False | |
| 63040 | 2340 | define t where "t = k / norm (y - x)" | 
| 53348 | 2341 | have "2 < t" "0<t" | 
| 60420 | 2342 | unfolding t_def using as False and \<open>k>0\<close> | 
| 68031 | 2343 | by (auto simp:field_simps) | 
| 68058 | 2344 | have "y \<in> S" | 
| 2345 | apply (rule k[THEN subsetD]) | |
| 53348 | 2346 | unfolding mem_cball dist_norm | 
| 2347 | apply (rule order_trans[of _ "2 * norm (x - y)"]) | |
| 2348 | using as | |
| 68031 | 2349 | by (auto simp: field_simps norm_minus_commute) | 
| 53348 | 2350 |       {
 | 
| 63040 | 2351 | define w where "w = x + t *\<^sub>R (y - x)" | 
| 68058 | 2352 | have "w \<in> S" | 
| 2353 | using \<open>k>0\<close> by (auto simp: dist_norm t_def w_def k[THEN subsetD]) | |
| 53348 | 2354 | have "(1 / t) *\<^sub>R x + - x + ((t - 1) / t) *\<^sub>R x = (1 / t - 1 + (t - 1) / t) *\<^sub>R x" | 
| 68031 | 2355 | by (auto simp: algebra_simps) | 
| 53348 | 2356 | also have "\<dots> = 0" | 
| 68031 | 2357 | using \<open>t > 0\<close> by (auto simp:field_simps) | 
| 53348 | 2358 | finally have w: "(1 / t) *\<^sub>R w + ((t - 1) / t) *\<^sub>R x = y" | 
| 60420 | 2359 | unfolding w_def using False and \<open>t > 0\<close> | 
| 68031 | 2360 | by (auto simp: algebra_simps) | 
| 68052 | 2361 | have 2: "2 * B < e * t" | 
| 60420 | 2362 | unfolding t_def using \<open>0 < e\<close> \<open>0 < k\<close> \<open>B > 0\<close> and as and False | 
| 68031 | 2363 | by (auto simp:field_simps) | 
| 68052 | 2364 | have "f y - f x \<le> (f w - f x) / t" | 
| 33175 | 2365 | using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w] | 
| 68058 | 2366 | using \<open>0 < t\<close> \<open>2 < t\<close> and \<open>x \<in> S\<close> \<open>w \<in> S\<close> | 
| 68031 | 2367 | by (auto simp:field_simps) | 
| 68052 | 2368 | also have "... < e" | 
| 68058 | 2369 | using B(2)[OF \<open>w\<in>S\<close>] and B(2)[OF \<open>x\<in>S\<close>] 2 \<open>t > 0\<close> by (auto simp: field_simps) | 
| 68052 | 2370 | finally have th1: "f y - f x < e" . | 
| 53348 | 2371 | } | 
| 49531 | 2372 | moreover | 
| 53348 | 2373 |       {
 | 
| 63040 | 2374 | define w where "w = x - t *\<^sub>R (y - x)" | 
| 68058 | 2375 | have "w \<in> S" | 
| 2376 | using \<open>k > 0\<close> by (auto simp: dist_norm t_def w_def k[THEN subsetD]) | |
| 53348 | 2377 | have "(1 / (1 + t)) *\<^sub>R x + (t / (1 + t)) *\<^sub>R x = (1 / (1 + t) + t / (1 + t)) *\<^sub>R x" | 
| 68031 | 2378 | by (auto simp: algebra_simps) | 
| 53348 | 2379 | also have "\<dots> = x" | 
| 68031 | 2380 | using \<open>t > 0\<close> by (auto simp:field_simps) | 
| 53348 | 2381 | finally have w: "(1 / (1+t)) *\<^sub>R w + (t / (1 + t)) *\<^sub>R y = x" | 
| 60420 | 2382 | unfolding w_def using False and \<open>t > 0\<close> | 
| 68031 | 2383 | by (auto simp: algebra_simps) | 
| 53348 | 2384 | have "2 * B < e * t" | 
| 2385 | unfolding t_def | |
| 60420 | 2386 | using \<open>0 < e\<close> \<open>0 < k\<close> \<open>B > 0\<close> and as and False | 
| 68031 | 2387 | by (auto simp:field_simps) | 
| 53348 | 2388 | then have *: "(f w - f y) / t < e" | 
| 68058 | 2389 | using B(2)[OF \<open>w\<in>S\<close>] and B(2)[OF \<open>y\<in>S\<close>] | 
| 60420 | 2390 | using \<open>t > 0\<close> | 
| 68031 | 2391 | by (auto simp:field_simps) | 
| 49531 | 2392 | have "f x \<le> 1 / (1 + t) * f w + (t / (1 + t)) * f y" | 
| 33175 | 2393 | using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w] | 
| 68058 | 2394 | using \<open>0 < t\<close> \<open>2 < t\<close> and \<open>y \<in> S\<close> \<open>w \<in> S\<close> | 
| 68031 | 2395 | by (auto simp:field_simps) | 
| 53348 | 2396 | also have "\<dots> = (f w + t * f y) / (1 + t)" | 
| 68031 | 2397 | using \<open>t > 0\<close> by (auto simp: divide_simps) | 
| 53348 | 2398 | also have "\<dots> < e + f y" | 
| 68031 | 2399 | using \<open>t > 0\<close> * \<open>e > 0\<close> by (auto simp: field_simps) | 
| 53348 | 2400 | finally have "f x - f y < e" by auto | 
| 2401 | } | |
| 49531 | 2402 | ultimately show ?thesis by auto | 
| 60420 | 2403 | qed (insert \<open>0<e\<close>, auto) | 
| 2404 | qed (insert \<open>0<e\<close> \<open>0<k\<close> \<open>0<B\<close>, auto simp: field_simps) | |
| 2405 | qed | |
| 2406 | ||
| 2407 | ||
| 70136 | 2408 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Upper bound on a ball implies upper and lower bounds\<close> | 
| 33175 | 2409 | |
| 2410 | lemma convex_bounds_lemma: | |
| 36338 | 2411 | fixes x :: "'a::real_normed_vector" | 
| 53348 | 2412 | assumes "convex_on (cball x e) f" | 
| 2413 | and "\<forall>y \<in> cball x e. f y \<le> b" | |
| 61945 | 2414 | shows "\<forall>y \<in> cball x e. \<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" | 
| 53348 | 2415 | apply rule | 
| 2416 | proof (cases "0 \<le> e") | |
| 2417 | case True | |
| 2418 | fix y | |
| 2419 | assume y: "y \<in> cball x e" | |
| 63040 | 2420 | define z where "z = 2 *\<^sub>R x - y" | 
| 53348 | 2421 | have *: "x - (2 *\<^sub>R x - y) = y - x" | 
| 2422 | by (simp add: scaleR_2) | |
| 2423 | have z: "z \<in> cball x e" | |
| 68031 | 2424 | using y unfolding z_def mem_cball dist_norm * by (auto simp: norm_minus_commute) | 
| 53348 | 2425 | have "(1 / 2) *\<^sub>R y + (1 / 2) *\<^sub>R z = x" | 
| 68031 | 2426 | unfolding z_def by (auto simp: algebra_simps) | 
| 53348 | 2427 | then show "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" | 
| 2428 | using assms(1)[unfolded convex_on_def,rule_format, OF y z, of "1/2" "1/2"] | |
| 2429 | using assms(2)[rule_format,OF y] assms(2)[rule_format,OF z] | |
| 68031 | 2430 | by (auto simp:field_simps) | 
| 53348 | 2431 | next | 
| 2432 | case False | |
| 2433 | fix y | |
| 2434 | assume "y \<in> cball x e" | |
| 2435 | then have "dist x y < 0" | |
| 2436 | using False unfolding mem_cball not_le by (auto simp del: dist_not_less_zero) | |
| 2437 | then show "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" | |
| 2438 | using zero_le_dist[of x y] by auto | |
| 2439 | qed | |
| 2440 | ||
| 33175 | 2441 | |
| 70136 | 2442 | subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Hence a convex function on an open set is continuous\<close> | 
| 33175 | 2443 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 2444 | lemma real_of_nat_ge_one_iff: "1 \<le> real (n::nat) \<longleftrightarrow> 1 \<le> n" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 2445 | by auto | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 2446 | |
| 33175 | 2447 | lemma convex_on_continuous: | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2448 |   assumes "open (s::('a::euclidean_space) set)" "convex_on s f"
 | 
| 33175 | 2449 | shows "continuous_on s f" | 
| 53348 | 2450 | unfolding continuous_on_eq_continuous_at[OF assms(1)] | 
| 2451 | proof | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 2452 | note dimge1 = DIM_positive[where 'a='a] | 
| 53348 | 2453 | fix x | 
| 2454 | assume "x \<in> s" | |
| 2455 | then obtain e where e: "cball x e \<subseteq> s" "e > 0" | |
| 2456 | using assms(1) unfolding open_contains_cball by auto | |
| 63040 | 2457 |   define d where "d = e / real DIM('a)"
 | 
| 53348 | 2458 | have "0 < d" | 
| 60420 | 2459 | unfolding d_def using \<open>e > 0\<close> dimge1 by auto | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 2460 | let ?d = "(\<Sum>i\<in>Basis. d *\<^sub>R i)::'a" | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2461 | obtain c | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2462 | where c: "finite c" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2463 | and c1: "convex hull c \<subseteq> cball x e" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2464 | and c2: "cball x d \<subseteq> convex hull c" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2465 | proof | 
| 63040 | 2466 |     define c where "c = (\<Sum>i\<in>Basis. (\<lambda>a. a *\<^sub>R i) ` {x\<bullet>i - d, x\<bullet>i + d})"
 | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2467 | show "finite c" | 
| 64267 | 2468 | unfolding c_def by (simp add: finite_set_sum) | 
| 56188 | 2469 |     have 1: "convex hull c = {a. \<forall>i\<in>Basis. a \<bullet> i \<in> cbox (x \<bullet> i - d) (x \<bullet> i + d)}"
 | 
| 64267 | 2470 | unfolding box_eq_set_sum_Basis | 
| 2471 | unfolding c_def convex_hull_set_sum | |
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2472 | apply (subst convex_hull_linear_image [symmetric]) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2473 | apply (simp add: linear_iff scaleR_add_left) | 
| 64267 | 2474 | apply (rule sum.cong [OF refl]) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2475 | apply (rule image_cong [OF _ refl]) | 
| 56188 | 2476 | apply (rule convex_hull_eq_real_cbox) | 
| 60420 | 2477 | apply (cut_tac \<open>0 < d\<close>, simp) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2478 | done | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2479 |     then have 2: "convex hull c = {a. \<forall>i\<in>Basis. a \<bullet> i \<in> cball (x \<bullet> i) d}"
 | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2480 | by (simp add: dist_norm abs_le_iff algebra_simps) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2481 | show "cball x d \<subseteq> convex hull c" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2482 | unfolding 2 | 
| 68058 | 2483 | by (clarsimp simp: dist_norm) (metis inner_commute inner_diff_right norm_bound_Basis_le) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2484 | have e': "e = (\<Sum>(i::'a)\<in>Basis. d)" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 2485 | by (simp add: d_def DIM_positive) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2486 | show "convex hull c \<subseteq> cball x e" | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2487 | unfolding 2 | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2488 | apply clarsimp | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2489 | apply (subst euclidean_dist_l2) | 
| 67155 | 2490 | apply (rule order_trans [OF L2_set_le_sum]) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2491 | apply (rule zero_le_dist) | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2492 | unfolding e' | 
| 68031 | 2493 | apply (rule sum_mono, simp) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2494 | done | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2495 | qed | 
| 63040 | 2496 | define k where "k = Max (f ` c)" | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2497 | have "convex_on (convex hull c) f" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50104diff
changeset | 2498 | apply(rule convex_on_subset[OF assms(2)]) | 
| 68069 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68058diff
changeset | 2499 | apply(rule subset_trans[OF c1 e(1)]) | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2500 | done | 
| 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2501 | then have k: "\<forall>y\<in>convex hull c. f y \<le> k" | 
| 68031 | 2502 | apply (rule_tac convex_on_convex_hull_bound, assumption) | 
| 68048 | 2503 | by (simp add: k_def c) | 
| 2504 |   have "e \<le> e * real DIM('a)"
 | |
| 2505 | using e(2) real_of_nat_ge_one_iff by auto | |
| 2506 | then have "d \<le> e" | |
| 2507 | by (simp add: d_def divide_simps) | |
| 53348 | 2508 | then have dsube: "cball x d \<subseteq> cball x e" | 
| 53620 
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
 huffman parents: 
53600diff
changeset | 2509 | by (rule subset_cball) | 
| 53348 | 2510 | have conv: "convex_on (cball x d) f" | 
| 68031 | 2511 | using \<open>convex_on (convex hull c) f\<close> c2 convex_on_subset by blast | 
| 61945 | 2512 | then have "\<forall>y\<in>cball x d. \<bar>f y\<bar> \<le> k + 2 * \<bar>f x\<bar>" | 
| 68048 | 2513 | by (rule convex_bounds_lemma) (use c2 k in blast) | 
| 53348 | 2514 | then have "continuous_on (ball x d) f" | 
| 2515 | apply (rule_tac convex_on_bounded_continuous) | |
| 2516 | apply (rule open_ball, rule convex_on_subset[OF conv]) | |
| 68031 | 2517 | apply (rule ball_subset_cball, force) | 
| 33270 | 2518 | done | 
| 53348 | 2519 | then show "continuous (at x) f" | 
| 2520 | unfolding continuous_on_eq_continuous_at[OF open_ball] | |
| 60420 | 2521 | using \<open>d > 0\<close> by auto | 
| 2522 | qed | |
| 2523 | ||
| 33175 | 2524 | end |