src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy
changeset 63627 6ddb43c6b711
parent 63626 44ce6b524ff3
child 63631 2edc8da89edc
child 63633 2accfb71e33b
equal deleted inserted replaced
63626:44ce6b524ff3 63627:6ddb43c6b711
     1 (*  Title:      HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy
       
     2     Authors: Robert Himmelmann, TU Muenchen; Bogdan Grechuk, University of Edinburgh; LCP
       
     3 *)
       
     4 
       
     5 section \<open>Convex sets, functions and related things\<close>
       
     6 
       
     7 theory Convex_Euclidean_Space
       
     8 imports
       
     9   Topology_Euclidean_Space
       
    10   "~~/src/HOL/Library/Convex"
       
    11   "~~/src/HOL/Library/Set_Algebras"
       
    12 begin
       
    13 
       
    14 (*FIXME move up e.g. to Library/Convex.thy, but requires the "support" primitive*)
       
    15 lemma convex_supp_setsum:
       
    16   assumes "convex S" and 1: "supp_setsum u I = 1"
       
    17       and "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> u i \<and> (u i = 0 \<or> f i \<in> S)"
       
    18     shows "supp_setsum (\<lambda>i. u i *\<^sub>R f i) I \<in> S"
       
    19 proof -
       
    20   have fin: "finite {i \<in> I. u i \<noteq> 0}"
       
    21     using 1 setsum.infinite by (force simp: supp_setsum_def support_on_def)
       
    22   then have eq: "supp_setsum (\<lambda>i. u i *\<^sub>R f i) I = setsum (\<lambda>i. u i *\<^sub>R f i) {i \<in> I. u i \<noteq> 0}"
       
    23     by (force intro: setsum.mono_neutral_left simp: supp_setsum_def support_on_def)
       
    24   show ?thesis
       
    25     apply (simp add: eq)
       
    26     apply (rule convex_setsum [OF fin \<open>convex S\<close>])
       
    27     using 1 assms apply (auto simp: supp_setsum_def support_on_def)
       
    28     done
       
    29 qed
       
    30 
       
    31 
       
    32 lemma dim_image_eq:
       
    33   fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
       
    34   assumes lf: "linear f"
       
    35     and fi: "inj_on f (span S)"
       
    36   shows "dim (f ` S) = dim (S::'n::euclidean_space set)"
       
    37 proof -
       
    38   obtain B where B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "card B = dim S"
       
    39     using basis_exists[of S] by auto
       
    40   then have "span S = span B"
       
    41     using span_mono[of B S] span_mono[of S "span B"] span_span[of B] by auto
       
    42   then have "independent (f ` B)"
       
    43     using independent_inj_on_image[of B f] B assms by auto
       
    44   moreover have "card (f ` B) = card B"
       
    45     using assms card_image[of f B] subset_inj_on[of f "span S" B] B span_inc by auto
       
    46   moreover have "(f ` B) \<subseteq> (f ` S)"
       
    47     using B by auto
       
    48   ultimately have "dim (f ` S) \<ge> dim S"
       
    49     using independent_card_le_dim[of "f ` B" "f ` S"] B by auto
       
    50   then show ?thesis
       
    51     using dim_image_le[of f S] assms by auto
       
    52 qed
       
    53 
       
    54 lemma linear_injective_on_subspace_0:
       
    55   assumes lf: "linear f"
       
    56     and "subspace S"
       
    57   shows "inj_on f S \<longleftrightarrow> (\<forall>x \<in> S. f x = 0 \<longrightarrow> x = 0)"
       
    58 proof -
       
    59   have "inj_on f S \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y \<in> S. f x = f y \<longrightarrow> x = y)"
       
    60     by (simp add: inj_on_def)
       
    61   also have "\<dots> \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y \<in> S. f x - f y = 0 \<longrightarrow> x - y = 0)"
       
    62     by simp
       
    63   also have "\<dots> \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y \<in> S. f (x - y) = 0 \<longrightarrow> x - y = 0)"
       
    64     by (simp add: linear_diff[OF lf])
       
    65   also have "\<dots> \<longleftrightarrow> (\<forall>x \<in> S. f x = 0 \<longrightarrow> x = 0)"
       
    66     using \<open>subspace S\<close> subspace_def[of S] subspace_diff[of S] by auto
       
    67   finally show ?thesis .
       
    68 qed
       
    69 
       
    70 lemma subspace_Inter: "\<forall>s \<in> f. subspace s \<Longrightarrow> subspace (\<Inter>f)"
       
    71   unfolding subspace_def by auto
       
    72 
       
    73 lemma span_eq[simp]: "span s = s \<longleftrightarrow> subspace s"
       
    74   unfolding span_def by (rule hull_eq) (rule subspace_Inter)
       
    75 
       
    76 lemma substdbasis_expansion_unique:
       
    77   assumes d: "d \<subseteq> Basis"
       
    78   shows "(\<Sum>i\<in>d. f i *\<^sub>R i) = (x::'a::euclidean_space) \<longleftrightarrow>
       
    79     (\<forall>i\<in>Basis. (i \<in> d \<longrightarrow> f i = x \<bullet> i) \<and> (i \<notin> d \<longrightarrow> x \<bullet> i = 0))"
       
    80 proof -
       
    81   have *: "\<And>x a b P. x * (if P then a else b) = (if P then x * a else x * b)"
       
    82     by auto
       
    83   have **: "finite d"
       
    84     by (auto intro: finite_subset[OF assms])
       
    85   have ***: "\<And>i. i \<in> Basis \<Longrightarrow> (\<Sum>i\<in>d. f i *\<^sub>R i) \<bullet> i = (\<Sum>x\<in>d. if x = i then f x else 0)"
       
    86     using d
       
    87     by (auto intro!: setsum.cong simp: inner_Basis inner_setsum_left)
       
    88   show ?thesis
       
    89     unfolding euclidean_eq_iff[where 'a='a] by (auto simp: setsum.delta[OF **] ***)
       
    90 qed
       
    91 
       
    92 lemma independent_substdbasis: "d \<subseteq> Basis \<Longrightarrow> independent d"
       
    93   by (rule independent_mono[OF independent_Basis])
       
    94 
       
    95 lemma dim_cball:
       
    96   assumes "e > 0"
       
    97   shows "dim (cball (0 :: 'n::euclidean_space) e) = DIM('n)"
       
    98 proof -
       
    99   {
       
   100     fix x :: "'n::euclidean_space"
       
   101     define y where "y = (e / norm x) *\<^sub>R x"
       
   102     then have "y \<in> cball 0 e"
       
   103       using assms by auto
       
   104     moreover have *: "x = (norm x / e) *\<^sub>R y"
       
   105       using y_def assms by simp
       
   106     moreover from * have "x = (norm x/e) *\<^sub>R y"
       
   107       by auto
       
   108     ultimately have "x \<in> span (cball 0 e)"
       
   109       using span_mul[of y "cball 0 e" "norm x/e"] span_inc[of "cball 0 e"]
       
   110       by (simp add: span_superset)
       
   111   }
       
   112   then have "span (cball 0 e) = (UNIV :: 'n::euclidean_space set)"
       
   113     by auto
       
   114   then show ?thesis
       
   115     using dim_span[of "cball (0 :: 'n::euclidean_space) e"] by (auto simp add: dim_UNIV)
       
   116 qed
       
   117 
       
   118 lemma indep_card_eq_dim_span:
       
   119   fixes B :: "'n::euclidean_space set"
       
   120   assumes "independent B"
       
   121   shows "finite B \<and> card B = dim (span B)"
       
   122   using assms basis_card_eq_dim[of B "span B"] span_inc by auto
       
   123 
       
   124 lemma setsum_not_0: "setsum f A \<noteq> 0 \<Longrightarrow> \<exists>a \<in> A. f a \<noteq> 0"
       
   125   by (rule ccontr) auto
       
   126 
       
   127 lemma subset_translation_eq [simp]:
       
   128     fixes a :: "'a::real_vector" shows "op + a ` s \<subseteq> op + a ` t \<longleftrightarrow> s \<subseteq> t"
       
   129   by auto
       
   130 
       
   131 lemma translate_inj_on:
       
   132   fixes A :: "'a::ab_group_add set"
       
   133   shows "inj_on (\<lambda>x. a + x) A"
       
   134   unfolding inj_on_def by auto
       
   135 
       
   136 lemma translation_assoc:
       
   137   fixes a b :: "'a::ab_group_add"
       
   138   shows "(\<lambda>x. b + x) ` ((\<lambda>x. a + x) ` S) = (\<lambda>x. (a + b) + x) ` S"
       
   139   by auto
       
   140 
       
   141 lemma translation_invert:
       
   142   fixes a :: "'a::ab_group_add"
       
   143   assumes "(\<lambda>x. a + x) ` A = (\<lambda>x. a + x) ` B"
       
   144   shows "A = B"
       
   145 proof -
       
   146   have "(\<lambda>x. -a + x) ` ((\<lambda>x. a + x) ` A) = (\<lambda>x. - a + x) ` ((\<lambda>x. a + x) ` B)"
       
   147     using assms by auto
       
   148   then show ?thesis
       
   149     using translation_assoc[of "-a" a A] translation_assoc[of "-a" a B] by auto
       
   150 qed
       
   151 
       
   152 lemma translation_galois:
       
   153   fixes a :: "'a::ab_group_add"
       
   154   shows "T = ((\<lambda>x. a + x) ` S) \<longleftrightarrow> S = ((\<lambda>x. (- a) + x) ` T)"
       
   155   using translation_assoc[of "-a" a S]
       
   156   apply auto
       
   157   using translation_assoc[of a "-a" T]
       
   158   apply auto
       
   159   done
       
   160 
       
   161 lemma convex_translation_eq [simp]: "convex ((\<lambda>x. a + x) ` s) \<longleftrightarrow> convex s"
       
   162   by (metis convex_translation translation_galois)
       
   163 
       
   164 lemma translation_inverse_subset:
       
   165   assumes "((\<lambda>x. - a + x) ` V) \<le> (S :: 'n::ab_group_add set)"
       
   166   shows "V \<le> ((\<lambda>x. a + x) ` S)"
       
   167 proof -
       
   168   {
       
   169     fix x
       
   170     assume "x \<in> V"
       
   171     then have "x-a \<in> S" using assms by auto
       
   172     then have "x \<in> {a + v |v. v \<in> S}"
       
   173       apply auto
       
   174       apply (rule exI[of _ "x-a"])
       
   175       apply simp
       
   176       done
       
   177     then have "x \<in> ((\<lambda>x. a+x) ` S)" by auto
       
   178   }
       
   179   then show ?thesis by auto
       
   180 qed
       
   181 
       
   182 lemma convex_linear_image_eq [simp]:
       
   183     fixes f :: "'a::real_vector \<Rightarrow> 'b::real_vector"
       
   184     shows "\<lbrakk>linear f; inj f\<rbrakk> \<Longrightarrow> convex (f ` s) \<longleftrightarrow> convex s"
       
   185     by (metis (no_types) convex_linear_image convex_linear_vimage inj_vimage_image_eq)
       
   186 
       
   187 lemma basis_to_basis_subspace_isomorphism:
       
   188   assumes s: "subspace (S:: ('n::euclidean_space) set)"
       
   189     and t: "subspace (T :: ('m::euclidean_space) set)"
       
   190     and d: "dim S = dim T"
       
   191     and B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "card B = dim S"
       
   192     and C: "C \<subseteq> T" "independent C" "T \<subseteq> span C" "card C = dim T"
       
   193   shows "\<exists>f. linear f \<and> f ` B = C \<and> f ` S = T \<and> inj_on f S"
       
   194 proof -
       
   195   from B independent_bound have fB: "finite B"
       
   196     by blast
       
   197   from C independent_bound have fC: "finite C"
       
   198     by blast
       
   199   from B(4) C(4) card_le_inj[of B C] d obtain f where
       
   200     f: "f ` B \<subseteq> C" "inj_on f B" using \<open>finite B\<close> \<open>finite C\<close> by auto
       
   201   from linear_independent_extend[OF B(2)] obtain g where
       
   202     g: "linear g" "\<forall>x \<in> B. g x = f x" by blast
       
   203   from inj_on_iff_eq_card[OF fB, of f] f(2)
       
   204   have "card (f ` B) = card B" by simp
       
   205   with B(4) C(4) have ceq: "card (f ` B) = card C" using d
       
   206     by simp
       
   207   have "g ` B = f ` B" using g(2)
       
   208     by (auto simp add: image_iff)
       
   209   also have "\<dots> = C" using card_subset_eq[OF fC f(1) ceq] .
       
   210   finally have gBC: "g ` B = C" .
       
   211   have gi: "inj_on g B" using f(2) g(2)
       
   212     by (auto simp add: inj_on_def)
       
   213   note g0 = linear_indep_image_lemma[OF g(1) fB, unfolded gBC, OF C(2) gi]
       
   214   {
       
   215     fix x y
       
   216     assume x: "x \<in> S" and y: "y \<in> S" and gxy: "g x = g y"
       
   217     from B(3) x y have x': "x \<in> span B" and y': "y \<in> span B"
       
   218       by blast+
       
   219     from gxy have th0: "g (x - y) = 0"
       
   220       by (simp add: linear_diff[OF g(1)])
       
   221     have th1: "x - y \<in> span B" using x' y'
       
   222       by (metis span_sub)
       
   223     have "x = y" using g0[OF th1 th0] by simp
       
   224   }
       
   225   then have giS: "inj_on g S" unfolding inj_on_def by blast
       
   226   from span_subspace[OF B(1,3) s]
       
   227   have "g ` S = span (g ` B)"
       
   228     by (simp add: span_linear_image[OF g(1)])
       
   229   also have "\<dots> = span C"
       
   230     unfolding gBC ..
       
   231   also have "\<dots> = T"
       
   232     using span_subspace[OF C(1,3) t] .
       
   233   finally have gS: "g ` S = T" .
       
   234   from g(1) gS giS gBC show ?thesis
       
   235     by blast
       
   236 qed
       
   237 
       
   238 lemma closure_bounded_linear_image_subset:
       
   239   assumes f: "bounded_linear f"
       
   240   shows "f ` closure S \<subseteq> closure (f ` S)"
       
   241   using linear_continuous_on [OF f] closed_closure closure_subset
       
   242   by (rule image_closure_subset)
       
   243 
       
   244 lemma closure_linear_image_subset:
       
   245   fixes f :: "'m::euclidean_space \<Rightarrow> 'n::real_normed_vector"
       
   246   assumes "linear f"
       
   247   shows "f ` (closure S) \<subseteq> closure (f ` S)"
       
   248   using assms unfolding linear_conv_bounded_linear
       
   249   by (rule closure_bounded_linear_image_subset)
       
   250 
       
   251 lemma closed_injective_linear_image:
       
   252     fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
   253     assumes S: "closed S" and f: "linear f" "inj f"
       
   254     shows "closed (f ` S)"
       
   255 proof -
       
   256   obtain g where g: "linear g" "g \<circ> f = id"
       
   257     using linear_injective_left_inverse [OF f] by blast
       
   258   then have confg: "continuous_on (range f) g"
       
   259     using linear_continuous_on linear_conv_bounded_linear by blast
       
   260   have [simp]: "g ` f ` S = S"
       
   261     using g by (simp add: image_comp)
       
   262   have cgf: "closed (g ` f ` S)"
       
   263     by (simp add: \<open>g \<circ> f = id\<close> S image_comp)
       
   264   have [simp]: "{x \<in> range f. g x \<in> S} = f ` S"
       
   265     using g by (simp add: o_def id_def image_def) metis
       
   266   show ?thesis
       
   267     apply (rule closedin_closed_trans [of "range f"])
       
   268     apply (rule continuous_closedin_preimage [OF confg cgf, simplified])
       
   269     apply (rule closed_injective_image_subspace)
       
   270     using f
       
   271     apply (auto simp: linear_linear linear_injective_0)
       
   272     done
       
   273 qed
       
   274 
       
   275 lemma closed_injective_linear_image_eq:
       
   276     fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
   277     assumes f: "linear f" "inj f"
       
   278       shows "(closed(image f s) \<longleftrightarrow> closed s)"
       
   279   by (metis closed_injective_linear_image closure_eq closure_linear_image_subset closure_subset_eq f(1) f(2) inj_image_subset_iff)
       
   280 
       
   281 lemma closure_injective_linear_image:
       
   282     fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
   283     shows "\<lbrakk>linear f; inj f\<rbrakk> \<Longrightarrow> f ` (closure S) = closure (f ` S)"
       
   284   apply (rule subset_antisym)
       
   285   apply (simp add: closure_linear_image_subset)
       
   286   by (simp add: closure_minimal closed_injective_linear_image closure_subset image_mono)
       
   287 
       
   288 lemma closure_bounded_linear_image:
       
   289     fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
   290     shows "\<lbrakk>linear f; bounded S\<rbrakk> \<Longrightarrow> f ` (closure S) = closure (f ` S)"
       
   291   apply (rule subset_antisym, simp add: closure_linear_image_subset)
       
   292   apply (rule closure_minimal, simp add: closure_subset image_mono)
       
   293   by (meson bounded_closure closed_closure compact_continuous_image compact_eq_bounded_closed linear_continuous_on linear_conv_bounded_linear)
       
   294 
       
   295 lemma closure_scaleR:
       
   296   fixes S :: "'a::real_normed_vector set"
       
   297   shows "(op *\<^sub>R c) ` (closure S) = closure ((op *\<^sub>R c) ` S)"
       
   298 proof
       
   299   show "(op *\<^sub>R c) ` (closure S) \<subseteq> closure ((op *\<^sub>R c) ` S)"
       
   300     using bounded_linear_scaleR_right
       
   301     by (rule closure_bounded_linear_image_subset)
       
   302   show "closure ((op *\<^sub>R c) ` S) \<subseteq> (op *\<^sub>R c) ` (closure S)"
       
   303     by (intro closure_minimal image_mono closure_subset closed_scaling closed_closure)
       
   304 qed
       
   305 
       
   306 lemma fst_linear: "linear fst"
       
   307   unfolding linear_iff by (simp add: algebra_simps)
       
   308 
       
   309 lemma snd_linear: "linear snd"
       
   310   unfolding linear_iff by (simp add: algebra_simps)
       
   311 
       
   312 lemma fst_snd_linear: "linear (\<lambda>(x,y). x + y)"
       
   313   unfolding linear_iff by (simp add: algebra_simps)
       
   314 
       
   315 lemma scaleR_2:
       
   316   fixes x :: "'a::real_vector"
       
   317   shows "scaleR 2 x = x + x"
       
   318   unfolding one_add_one [symmetric] scaleR_left_distrib by simp
       
   319 
       
   320 lemma scaleR_half_double [simp]:
       
   321   fixes a :: "'a::real_normed_vector"
       
   322   shows "(1 / 2) *\<^sub>R (a + a) = a"
       
   323 proof -
       
   324   have "\<And>r. r *\<^sub>R (a + a) = (r * 2) *\<^sub>R a"
       
   325     by (metis scaleR_2 scaleR_scaleR)
       
   326   then show ?thesis
       
   327     by simp
       
   328 qed
       
   329 
       
   330 lemma vector_choose_size:
       
   331   assumes "0 \<le> c"
       
   332   obtains x :: "'a::{real_normed_vector, perfect_space}" where "norm x = c"
       
   333 proof -
       
   334   obtain a::'a where "a \<noteq> 0"
       
   335     using UNIV_not_singleton UNIV_eq_I set_zero singletonI by fastforce
       
   336   then show ?thesis
       
   337     by (rule_tac x="scaleR (c / norm a) a" in that) (simp add: assms)
       
   338 qed
       
   339 
       
   340 lemma vector_choose_dist:
       
   341   assumes "0 \<le> c"
       
   342   obtains y :: "'a::{real_normed_vector, perfect_space}" where "dist x y = c"
       
   343 by (metis add_diff_cancel_left' assms dist_commute dist_norm vector_choose_size)
       
   344 
       
   345 lemma sphere_eq_empty [simp]:
       
   346   fixes a :: "'a::{real_normed_vector, perfect_space}"
       
   347   shows "sphere a r = {} \<longleftrightarrow> r < 0"
       
   348 by (auto simp: sphere_def dist_norm) (metis dist_norm le_less_linear vector_choose_dist)
       
   349 
       
   350 lemma setsum_delta_notmem:
       
   351   assumes "x \<notin> s"
       
   352   shows "setsum (\<lambda>y. if (y = x) then P x else Q y) s = setsum Q s"
       
   353     and "setsum (\<lambda>y. if (x = y) then P x else Q y) s = setsum Q s"
       
   354     and "setsum (\<lambda>y. if (y = x) then P y else Q y) s = setsum Q s"
       
   355     and "setsum (\<lambda>y. if (x = y) then P y else Q y) s = setsum Q s"
       
   356   apply (rule_tac [!] setsum.cong)
       
   357   using assms
       
   358   apply auto
       
   359   done
       
   360 
       
   361 lemma setsum_delta'':
       
   362   fixes s::"'a::real_vector set"
       
   363   assumes "finite s"
       
   364   shows "(\<Sum>x\<in>s. (if y = x then f x else 0) *\<^sub>R x) = (if y\<in>s then (f y) *\<^sub>R y else 0)"
       
   365 proof -
       
   366   have *: "\<And>x y. (if y = x then f x else (0::real)) *\<^sub>R x = (if x=y then (f x) *\<^sub>R x else 0)"
       
   367     by auto
       
   368   show ?thesis
       
   369     unfolding * using setsum.delta[OF assms, of y "\<lambda>x. f x *\<^sub>R x"] by auto
       
   370 qed
       
   371 
       
   372 lemma if_smult: "(if P then x else (y::real)) *\<^sub>R v = (if P then x *\<^sub>R v else y *\<^sub>R v)"
       
   373   by (fact if_distrib)
       
   374 
       
   375 lemma dist_triangle_eq:
       
   376   fixes x y z :: "'a::real_inner"
       
   377   shows "dist x z = dist x y + dist y z \<longleftrightarrow>
       
   378     norm (x - y) *\<^sub>R (y - z) = norm (y - z) *\<^sub>R (x - y)"
       
   379 proof -
       
   380   have *: "x - y + (y - z) = x - z" by auto
       
   381   show ?thesis unfolding dist_norm norm_triangle_eq[of "x - y" "y - z", unfolded *]
       
   382     by (auto simp add:norm_minus_commute)
       
   383 qed
       
   384 
       
   385 lemma norm_minus_eqI: "x = - y \<Longrightarrow> norm x = norm y" by auto
       
   386 
       
   387 lemma Min_grI:
       
   388   assumes "finite A" "A \<noteq> {}" "\<forall>a\<in>A. x < a"
       
   389   shows "x < Min A"
       
   390   unfolding Min_gr_iff[OF assms(1,2)] using assms(3) by auto
       
   391 
       
   392 lemma norm_lt: "norm x < norm y \<longleftrightarrow> inner x x < inner y y"
       
   393   unfolding norm_eq_sqrt_inner by simp
       
   394 
       
   395 lemma norm_le: "norm x \<le> norm y \<longleftrightarrow> inner x x \<le> inner y y"
       
   396   unfolding norm_eq_sqrt_inner by simp
       
   397 
       
   398 
       
   399 subsection \<open>Affine set and affine hull\<close>
       
   400 
       
   401 definition affine :: "'a::real_vector set \<Rightarrow> bool"
       
   402   where "affine s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s)"
       
   403 
       
   404 lemma affine_alt: "affine s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u::real. (1 - u) *\<^sub>R x + u *\<^sub>R y \<in> s)"
       
   405   unfolding affine_def by (metis eq_diff_eq')
       
   406 
       
   407 lemma affine_empty [iff]: "affine {}"
       
   408   unfolding affine_def by auto
       
   409 
       
   410 lemma affine_sing [iff]: "affine {x}"
       
   411   unfolding affine_alt by (auto simp add: scaleR_left_distrib [symmetric])
       
   412 
       
   413 lemma affine_UNIV [iff]: "affine UNIV"
       
   414   unfolding affine_def by auto
       
   415 
       
   416 lemma affine_Inter [intro]: "(\<And>s. s\<in>f \<Longrightarrow> affine s) \<Longrightarrow> affine (\<Inter>f)"
       
   417   unfolding affine_def by auto
       
   418 
       
   419 lemma affine_Int[intro]: "affine s \<Longrightarrow> affine t \<Longrightarrow> affine (s \<inter> t)"
       
   420   unfolding affine_def by auto
       
   421 
       
   422 lemma affine_scaling: "affine s \<Longrightarrow> affine (image (\<lambda>x. c *\<^sub>R x) s)"
       
   423   apply (clarsimp simp add: affine_def)
       
   424   apply (rule_tac x="u *\<^sub>R x + v *\<^sub>R y" in image_eqI)
       
   425   apply (auto simp: algebra_simps)
       
   426   done
       
   427 
       
   428 lemma affine_affine_hull [simp]: "affine(affine hull s)"
       
   429   unfolding hull_def
       
   430   using affine_Inter[of "{t. affine t \<and> s \<subseteq> t}"] by auto
       
   431 
       
   432 lemma affine_hull_eq[simp]: "(affine hull s = s) \<longleftrightarrow> affine s"
       
   433   by (metis affine_affine_hull hull_same)
       
   434 
       
   435 lemma affine_hyperplane: "affine {x. a \<bullet> x = b}"
       
   436   by (simp add: affine_def algebra_simps) (metis distrib_right mult.left_neutral)
       
   437 
       
   438 
       
   439 subsubsection \<open>Some explicit formulations (from Lars Schewe)\<close>
       
   440 
       
   441 lemma affine:
       
   442   fixes V::"'a::real_vector set"
       
   443   shows "affine V \<longleftrightarrow>
       
   444     (\<forall>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> V \<and> setsum u s = 1 \<longrightarrow> (setsum (\<lambda>x. (u x) *\<^sub>R x)) s \<in> V)"
       
   445   unfolding affine_def
       
   446   apply rule
       
   447   apply(rule, rule, rule)
       
   448   apply(erule conjE)+
       
   449   defer
       
   450   apply (rule, rule, rule, rule, rule)
       
   451 proof -
       
   452   fix x y u v
       
   453   assume as: "x \<in> V" "y \<in> V" "u + v = (1::real)"
       
   454     "\<forall>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> V \<and> setsum u s = 1 \<longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
       
   455   then show "u *\<^sub>R x + v *\<^sub>R y \<in> V"
       
   456     apply (cases "x = y")
       
   457     using as(4)[THEN spec[where x="{x,y}"], THEN spec[where x="\<lambda>w. if w = x then u else v"]]
       
   458       and as(1-3)
       
   459     apply (auto simp add: scaleR_left_distrib[symmetric])
       
   460     done
       
   461 next
       
   462   fix s u
       
   463   assume as: "\<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V"
       
   464     "finite s" "s \<noteq> {}" "s \<subseteq> V" "setsum u s = (1::real)"
       
   465   define n where "n = card s"
       
   466   have "card s = 0 \<or> card s = 1 \<or> card s = 2 \<or> card s > 2" by auto
       
   467   then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
       
   468   proof (auto simp only: disjE)
       
   469     assume "card s = 2"
       
   470     then have "card s = Suc (Suc 0)"
       
   471       by auto
       
   472     then obtain a b where "s = {a, b}"
       
   473       unfolding card_Suc_eq by auto
       
   474     then show ?thesis
       
   475       using as(1)[THEN bspec[where x=a], THEN bspec[where x=b]] using as(4,5)
       
   476       by (auto simp add: setsum_clauses(2))
       
   477   next
       
   478     assume "card s > 2"
       
   479     then show ?thesis using as and n_def
       
   480     proof (induct n arbitrary: u s)
       
   481       case 0
       
   482       then show ?case by auto
       
   483     next
       
   484       case (Suc n)
       
   485       fix s :: "'a set" and u :: "'a \<Rightarrow> real"
       
   486       assume IA:
       
   487         "\<And>u s.  \<lbrakk>2 < card s; \<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V; finite s;
       
   488           s \<noteq> {}; s \<subseteq> V; setsum u s = 1; n = card s \<rbrakk> \<Longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
       
   489         and as:
       
   490           "Suc n = card s" "2 < card s" "\<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V"
       
   491            "finite s" "s \<noteq> {}" "s \<subseteq> V" "setsum u s = 1"
       
   492       have "\<exists>x\<in>s. u x \<noteq> 1"
       
   493       proof (rule ccontr)
       
   494         assume "\<not> ?thesis"
       
   495         then have "setsum u s = real_of_nat (card s)"
       
   496           unfolding card_eq_setsum by auto
       
   497         then show False
       
   498           using as(7) and \<open>card s > 2\<close>
       
   499           by (metis One_nat_def less_Suc0 Zero_not_Suc of_nat_1 of_nat_eq_iff numeral_2_eq_2)
       
   500       qed
       
   501       then obtain x where x:"x \<in> s" "u x \<noteq> 1" by auto
       
   502 
       
   503       have c: "card (s - {x}) = card s - 1"
       
   504         apply (rule card_Diff_singleton)
       
   505         using \<open>x\<in>s\<close> as(4)
       
   506         apply auto
       
   507         done
       
   508       have *: "s = insert x (s - {x})" "finite (s - {x})"
       
   509         using \<open>x\<in>s\<close> and as(4) by auto
       
   510       have **: "setsum u (s - {x}) = 1 - u x"
       
   511         using setsum_clauses(2)[OF *(2), of u x, unfolded *(1)[symmetric] as(7)] by auto
       
   512       have ***: "inverse (1 - u x) * setsum u (s - {x}) = 1"
       
   513         unfolding ** using \<open>u x \<noteq> 1\<close> by auto
       
   514       have "(\<Sum>xa\<in>s - {x}. (inverse (1 - u x) * u xa) *\<^sub>R xa) \<in> V"
       
   515       proof (cases "card (s - {x}) > 2")
       
   516         case True
       
   517         then have "s - {x} \<noteq> {}" "card (s - {x}) = n"
       
   518           unfolding c and as(1)[symmetric]
       
   519         proof (rule_tac ccontr)
       
   520           assume "\<not> s - {x} \<noteq> {}"
       
   521           then have "card (s - {x}) = 0" unfolding card_0_eq[OF *(2)] by simp
       
   522           then show False using True by auto
       
   523         qed auto
       
   524         then show ?thesis
       
   525           apply (rule_tac IA[of "s - {x}" "\<lambda>y. (inverse (1 - u x) * u y)"])
       
   526           unfolding setsum_right_distrib[symmetric]
       
   527           using as and *** and True
       
   528           apply auto
       
   529           done
       
   530       next
       
   531         case False
       
   532         then have "card (s - {x}) = Suc (Suc 0)"
       
   533           using as(2) and c by auto
       
   534         then obtain a b where "(s - {x}) = {a, b}" "a\<noteq>b"
       
   535           unfolding card_Suc_eq by auto
       
   536         then show ?thesis
       
   537           using as(3)[THEN bspec[where x=a], THEN bspec[where x=b]]
       
   538           using *** *(2) and \<open>s \<subseteq> V\<close>
       
   539           unfolding setsum_right_distrib
       
   540           by (auto simp add: setsum_clauses(2))
       
   541       qed
       
   542       then have "u x + (1 - u x) = 1 \<Longrightarrow>
       
   543           u x *\<^sub>R x + (1 - u x) *\<^sub>R ((\<Sum>xa\<in>s - {x}. u xa *\<^sub>R xa) /\<^sub>R (1 - u x)) \<in> V"
       
   544         apply -
       
   545         apply (rule as(3)[rule_format])
       
   546         unfolding  Real_Vector_Spaces.scaleR_right.setsum
       
   547         using x(1) as(6)
       
   548         apply auto
       
   549         done
       
   550       then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
       
   551         unfolding scaleR_scaleR[symmetric] and scaleR_right.setsum [symmetric]
       
   552         apply (subst *)
       
   553         unfolding setsum_clauses(2)[OF *(2)]
       
   554         using \<open>u x \<noteq> 1\<close>
       
   555         apply auto
       
   556         done
       
   557     qed
       
   558   next
       
   559     assume "card s = 1"
       
   560     then obtain a where "s={a}"
       
   561       by (auto simp add: card_Suc_eq)
       
   562     then show ?thesis
       
   563       using as(4,5) by simp
       
   564   qed (insert \<open>s\<noteq>{}\<close> \<open>finite s\<close>, auto)
       
   565 qed
       
   566 
       
   567 lemma affine_hull_explicit:
       
   568   "affine hull p =
       
   569     {y. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> setsum (\<lambda>v. (u v) *\<^sub>R v) s = y}"
       
   570   apply (rule hull_unique)
       
   571   apply (subst subset_eq)
       
   572   prefer 3
       
   573   apply rule
       
   574   unfolding mem_Collect_eq
       
   575   apply (erule exE)+
       
   576   apply (erule conjE)+
       
   577   prefer 2
       
   578   apply rule
       
   579 proof -
       
   580   fix x
       
   581   assume "x\<in>p"
       
   582   then show "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
       
   583     apply (rule_tac x="{x}" in exI)
       
   584     apply (rule_tac x="\<lambda>x. 1" in exI)
       
   585     apply auto
       
   586     done
       
   587 next
       
   588   fix t x s u
       
   589   assume as: "p \<subseteq> t" "affine t" "finite s" "s \<noteq> {}"
       
   590     "s \<subseteq> p" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
       
   591   then show "x \<in> t"
       
   592     using as(2)[unfolded affine, THEN spec[where x=s], THEN spec[where x=u]]
       
   593     by auto
       
   594 next
       
   595   show "affine {y. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y}"
       
   596     unfolding affine_def
       
   597     apply (rule, rule, rule, rule, rule)
       
   598     unfolding mem_Collect_eq
       
   599   proof -
       
   600     fix u v :: real
       
   601     assume uv: "u + v = 1"
       
   602     fix x
       
   603     assume "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
       
   604     then obtain sx ux where
       
   605       x: "finite sx" "sx \<noteq> {}" "sx \<subseteq> p" "setsum ux sx = 1" "(\<Sum>v\<in>sx. ux v *\<^sub>R v) = x"
       
   606       by auto
       
   607     fix y
       
   608     assume "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
       
   609     then obtain sy uy where
       
   610       y: "finite sy" "sy \<noteq> {}" "sy \<subseteq> p" "setsum uy sy = 1" "(\<Sum>v\<in>sy. uy v *\<^sub>R v) = y" by auto
       
   611     have xy: "finite (sx \<union> sy)"
       
   612       using x(1) y(1) by auto
       
   613     have **: "(sx \<union> sy) \<inter> sx = sx" "(sx \<union> sy) \<inter> sy = sy"
       
   614       by auto
       
   615     show "\<exists>s ua. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and>
       
   616         setsum ua s = 1 \<and> (\<Sum>v\<in>s. ua v *\<^sub>R v) = u *\<^sub>R x + v *\<^sub>R y"
       
   617       apply (rule_tac x="sx \<union> sy" in exI)
       
   618       apply (rule_tac x="\<lambda>a. (if a\<in>sx then u * ux a else 0) + (if a\<in>sy then v * uy a else 0)" in exI)
       
   619       unfolding scaleR_left_distrib setsum.distrib if_smult scaleR_zero_left
       
   620         ** setsum.inter_restrict[OF xy, symmetric]
       
   621       unfolding scaleR_scaleR[symmetric] Real_Vector_Spaces.scaleR_right.setsum [symmetric]
       
   622         and setsum_right_distrib[symmetric]
       
   623       unfolding x y
       
   624       using x(1-3) y(1-3) uv
       
   625       apply simp
       
   626       done
       
   627   qed
       
   628 qed
       
   629 
       
   630 lemma affine_hull_finite:
       
   631   assumes "finite s"
       
   632   shows "affine hull s = {y. \<exists>u. setsum u s = 1 \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = y}"
       
   633   unfolding affine_hull_explicit and set_eq_iff and mem_Collect_eq
       
   634   apply (rule, rule)
       
   635   apply (erule exE)+
       
   636   apply (erule conjE)+
       
   637   defer
       
   638   apply (erule exE)
       
   639   apply (erule conjE)
       
   640 proof -
       
   641   fix x u
       
   642   assume "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
       
   643   then show "\<exists>sa u. finite sa \<and>
       
   644       \<not> (\<forall>x. (x \<in> sa) = (x \<in> {})) \<and> sa \<subseteq> s \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = x"
       
   645     apply (rule_tac x=s in exI, rule_tac x=u in exI)
       
   646     using assms
       
   647     apply auto
       
   648     done
       
   649 next
       
   650   fix x t u
       
   651   assume "t \<subseteq> s"
       
   652   then have *: "s \<inter> t = t"
       
   653     by auto
       
   654   assume "finite t" "\<not> (\<forall>x. (x \<in> t) = (x \<in> {}))" "setsum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = x"
       
   655   then show "\<exists>u. setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
       
   656     apply (rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI)
       
   657     unfolding if_smult scaleR_zero_left and setsum.inter_restrict[OF assms, symmetric] and *
       
   658     apply auto
       
   659     done
       
   660 qed
       
   661 
       
   662 
       
   663 subsubsection \<open>Stepping theorems and hence small special cases\<close>
       
   664 
       
   665 lemma affine_hull_empty[simp]: "affine hull {} = {}"
       
   666   by (rule hull_unique) auto
       
   667 
       
   668 lemma affine_hull_finite_step:
       
   669   fixes y :: "'a::real_vector"
       
   670   shows
       
   671     "(\<exists>u. setsum u {} = w \<and> setsum (\<lambda>x. u x *\<^sub>R x) {} = y) \<longleftrightarrow> w = 0 \<and> y = 0" (is ?th1)
       
   672     and
       
   673     "finite s \<Longrightarrow>
       
   674       (\<exists>u. setsum u (insert a s) = w \<and> setsum (\<lambda>x. u x *\<^sub>R x) (insert a s) = y) \<longleftrightarrow>
       
   675       (\<exists>v u. setsum u s = w - v \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y - v *\<^sub>R a)" (is "_ \<Longrightarrow> ?lhs = ?rhs")
       
   676 proof -
       
   677   show ?th1 by simp
       
   678   assume fin: "finite s"
       
   679   show "?lhs = ?rhs"
       
   680   proof
       
   681     assume ?lhs
       
   682     then obtain u where u: "setsum u (insert a s) = w \<and> (\<Sum>x\<in>insert a s. u x *\<^sub>R x) = y"
       
   683       by auto
       
   684     show ?rhs
       
   685     proof (cases "a \<in> s")
       
   686       case True
       
   687       then have *: "insert a s = s" by auto
       
   688       show ?thesis
       
   689         using u[unfolded *]
       
   690         apply(rule_tac x=0 in exI)
       
   691         apply auto
       
   692         done
       
   693     next
       
   694       case False
       
   695       then show ?thesis
       
   696         apply (rule_tac x="u a" in exI)
       
   697         using u and fin
       
   698         apply auto
       
   699         done
       
   700     qed
       
   701   next
       
   702     assume ?rhs
       
   703     then obtain v u where vu: "setsum u s = w - v"  "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a"
       
   704       by auto
       
   705     have *: "\<And>x M. (if x = a then v else M) *\<^sub>R x = (if x = a then v *\<^sub>R x else M *\<^sub>R x)"
       
   706       by auto
       
   707     show ?lhs
       
   708     proof (cases "a \<in> s")
       
   709       case True
       
   710       then show ?thesis
       
   711         apply (rule_tac x="\<lambda>x. (if x=a then v else 0) + u x" in exI)
       
   712         unfolding setsum_clauses(2)[OF fin]
       
   713         apply simp
       
   714         unfolding scaleR_left_distrib and setsum.distrib
       
   715         unfolding vu and * and scaleR_zero_left
       
   716         apply (auto simp add: setsum.delta[OF fin])
       
   717         done
       
   718     next
       
   719       case False
       
   720       then have **:
       
   721         "\<And>x. x \<in> s \<Longrightarrow> u x = (if x = a then v else u x)"
       
   722         "\<And>x. x \<in> s \<Longrightarrow> u x *\<^sub>R x = (if x = a then v *\<^sub>R x else u x *\<^sub>R x)" by auto
       
   723       from False show ?thesis
       
   724         apply (rule_tac x="\<lambda>x. if x=a then v else u x" in exI)
       
   725         unfolding setsum_clauses(2)[OF fin] and * using vu
       
   726         using setsum.cong [of s _ "\<lambda>x. u x *\<^sub>R x" "\<lambda>x. if x = a then v *\<^sub>R x else u x *\<^sub>R x", OF _ **(2)]
       
   727         using setsum.cong [of s _ u "\<lambda>x. if x = a then v else u x", OF _ **(1)]
       
   728         apply auto
       
   729         done
       
   730     qed
       
   731   qed
       
   732 qed
       
   733 
       
   734 lemma affine_hull_2:
       
   735   fixes a b :: "'a::real_vector"
       
   736   shows "affine hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b| u v. (u + v = 1)}"
       
   737   (is "?lhs = ?rhs")
       
   738 proof -
       
   739   have *:
       
   740     "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)"
       
   741     "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto
       
   742   have "?lhs = {y. \<exists>u. setsum u {a, b} = 1 \<and> (\<Sum>v\<in>{a, b}. u v *\<^sub>R v) = y}"
       
   743     using affine_hull_finite[of "{a,b}"] by auto
       
   744   also have "\<dots> = {y. \<exists>v u. u b = 1 - v \<and> u b *\<^sub>R b = y - v *\<^sub>R a}"
       
   745     by (simp add: affine_hull_finite_step(2)[of "{b}" a])
       
   746   also have "\<dots> = ?rhs" unfolding * by auto
       
   747   finally show ?thesis by auto
       
   748 qed
       
   749 
       
   750 lemma affine_hull_3:
       
   751   fixes a b c :: "'a::real_vector"
       
   752   shows "affine hull {a,b,c} = { u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c| u v w. u + v + w = 1}"
       
   753 proof -
       
   754   have *:
       
   755     "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)"
       
   756     "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto
       
   757   show ?thesis
       
   758     apply (simp add: affine_hull_finite affine_hull_finite_step)
       
   759     unfolding *
       
   760     apply auto
       
   761     apply (rule_tac x=v in exI)
       
   762     apply (rule_tac x=va in exI)
       
   763     apply auto
       
   764     apply (rule_tac x=u in exI)
       
   765     apply force
       
   766     done
       
   767 qed
       
   768 
       
   769 lemma mem_affine:
       
   770   assumes "affine S" "x \<in> S" "y \<in> S" "u + v = 1"
       
   771   shows "u *\<^sub>R x + v *\<^sub>R y \<in> S"
       
   772   using assms affine_def[of S] by auto
       
   773 
       
   774 lemma mem_affine_3:
       
   775   assumes "affine S" "x \<in> S" "y \<in> S" "z \<in> S" "u + v + w = 1"
       
   776   shows "u *\<^sub>R x + v *\<^sub>R y + w *\<^sub>R z \<in> S"
       
   777 proof -
       
   778   have "u *\<^sub>R x + v *\<^sub>R y + w *\<^sub>R z \<in> affine hull {x, y, z}"
       
   779     using affine_hull_3[of x y z] assms by auto
       
   780   moreover
       
   781   have "affine hull {x, y, z} \<subseteq> affine hull S"
       
   782     using hull_mono[of "{x, y, z}" "S"] assms by auto
       
   783   moreover
       
   784   have "affine hull S = S"
       
   785     using assms affine_hull_eq[of S] by auto
       
   786   ultimately show ?thesis by auto
       
   787 qed
       
   788 
       
   789 lemma mem_affine_3_minus:
       
   790   assumes "affine S" "x \<in> S" "y \<in> S" "z \<in> S"
       
   791   shows "x + v *\<^sub>R (y-z) \<in> S"
       
   792   using mem_affine_3[of S x y z 1 v "-v"] assms
       
   793   by (simp add: algebra_simps)
       
   794 
       
   795 corollary mem_affine_3_minus2:
       
   796     "\<lbrakk>affine S; x \<in> S; y \<in> S; z \<in> S\<rbrakk> \<Longrightarrow> x - v *\<^sub>R (y-z) \<in> S"
       
   797   by (metis add_uminus_conv_diff mem_affine_3_minus real_vector.scale_minus_left)
       
   798 
       
   799 
       
   800 subsubsection \<open>Some relations between affine hull and subspaces\<close>
       
   801 
       
   802 lemma affine_hull_insert_subset_span:
       
   803   "affine hull (insert a s) \<subseteq> {a + v| v . v \<in> span {x - a | x . x \<in> s}}"
       
   804   unfolding subset_eq Ball_def
       
   805   unfolding affine_hull_explicit span_explicit mem_Collect_eq
       
   806   apply (rule, rule)
       
   807   apply (erule exE)+
       
   808   apply (erule conjE)+
       
   809 proof -
       
   810   fix x t u
       
   811   assume as: "finite t" "t \<noteq> {}" "t \<subseteq> insert a s" "setsum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = x"
       
   812   have "(\<lambda>x. x - a) ` (t - {a}) \<subseteq> {x - a |x. x \<in> s}"
       
   813     using as(3) by auto
       
   814   then show "\<exists>v. x = a + v \<and> (\<exists>S u. finite S \<and> S \<subseteq> {x - a |x. x \<in> s} \<and> (\<Sum>v\<in>S. u v *\<^sub>R v) = v)"
       
   815     apply (rule_tac x="x - a" in exI)
       
   816     apply (rule conjI, simp)
       
   817     apply (rule_tac x="(\<lambda>x. x - a) ` (t - {a})" in exI)
       
   818     apply (rule_tac x="\<lambda>x. u (x + a)" in exI)
       
   819     apply (rule conjI) using as(1) apply simp
       
   820     apply (erule conjI)
       
   821     using as(1)
       
   822     apply (simp add: setsum.reindex[unfolded inj_on_def] scaleR_right_diff_distrib
       
   823       setsum_subtractf scaleR_left.setsum[symmetric] setsum_diff1 scaleR_left_diff_distrib)
       
   824     unfolding as
       
   825     apply simp
       
   826     done
       
   827 qed
       
   828 
       
   829 lemma affine_hull_insert_span:
       
   830   assumes "a \<notin> s"
       
   831   shows "affine hull (insert a s) = {a + v | v . v \<in> span {x - a | x.  x \<in> s}}"
       
   832   apply (rule, rule affine_hull_insert_subset_span)
       
   833   unfolding subset_eq Ball_def
       
   834   unfolding affine_hull_explicit and mem_Collect_eq
       
   835 proof (rule, rule, erule exE, erule conjE)
       
   836   fix y v
       
   837   assume "y = a + v" "v \<in> span {x - a |x. x \<in> s}"
       
   838   then obtain t u where obt: "finite t" "t \<subseteq> {x - a |x. x \<in> s}" "a + (\<Sum>v\<in>t. u v *\<^sub>R v) = y"
       
   839     unfolding span_explicit by auto
       
   840   define f where "f = (\<lambda>x. x + a) ` t"
       
   841   have f: "finite f" "f \<subseteq> s" "(\<Sum>v\<in>f. u (v - a) *\<^sub>R (v - a)) = y - a"
       
   842     unfolding f_def using obt by (auto simp add: setsum.reindex[unfolded inj_on_def])
       
   843   have *: "f \<inter> {a} = {}" "f \<inter> - {a} = f"
       
   844     using f(2) assms by auto
       
   845   show "\<exists>sa u. finite sa \<and> sa \<noteq> {} \<and> sa \<subseteq> insert a s \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y"
       
   846     apply (rule_tac x = "insert a f" in exI)
       
   847     apply (rule_tac x = "\<lambda>x. if x=a then 1 - setsum (\<lambda>x. u (x - a)) f else u (x - a)" in exI)
       
   848     using assms and f
       
   849     unfolding setsum_clauses(2)[OF f(1)] and if_smult
       
   850     unfolding setsum.If_cases[OF f(1), of "\<lambda>x. x = a"]
       
   851     apply (auto simp add: setsum_subtractf scaleR_left.setsum algebra_simps *)
       
   852     done
       
   853 qed
       
   854 
       
   855 lemma affine_hull_span:
       
   856   assumes "a \<in> s"
       
   857   shows "affine hull s = {a + v | v. v \<in> span {x - a | x. x \<in> s - {a}}}"
       
   858   using affine_hull_insert_span[of a "s - {a}", unfolded insert_Diff[OF assms]] by auto
       
   859 
       
   860 
       
   861 subsubsection \<open>Parallel affine sets\<close>
       
   862 
       
   863 definition affine_parallel :: "'a::real_vector set \<Rightarrow> 'a::real_vector set \<Rightarrow> bool"
       
   864   where "affine_parallel S T \<longleftrightarrow> (\<exists>a. T = (\<lambda>x. a + x) ` S)"
       
   865 
       
   866 lemma affine_parallel_expl_aux:
       
   867   fixes S T :: "'a::real_vector set"
       
   868   assumes "\<forall>x. x \<in> S \<longleftrightarrow> a + x \<in> T"
       
   869   shows "T = (\<lambda>x. a + x) ` S"
       
   870 proof -
       
   871   {
       
   872     fix x
       
   873     assume "x \<in> T"
       
   874     then have "( - a) + x \<in> S"
       
   875       using assms by auto
       
   876     then have "x \<in> ((\<lambda>x. a + x) ` S)"
       
   877       using imageI[of "-a+x" S "(\<lambda>x. a+x)"] by auto
       
   878   }
       
   879   moreover have "T \<ge> (\<lambda>x. a + x) ` S"
       
   880     using assms by auto
       
   881   ultimately show ?thesis by auto
       
   882 qed
       
   883 
       
   884 lemma affine_parallel_expl: "affine_parallel S T \<longleftrightarrow> (\<exists>a. \<forall>x. x \<in> S \<longleftrightarrow> a + x \<in> T)"
       
   885   unfolding affine_parallel_def
       
   886   using affine_parallel_expl_aux[of S _ T] by auto
       
   887 
       
   888 lemma affine_parallel_reflex: "affine_parallel S S"
       
   889   unfolding affine_parallel_def
       
   890   apply (rule exI[of _ "0"])
       
   891   apply auto
       
   892   done
       
   893 
       
   894 lemma affine_parallel_commut:
       
   895   assumes "affine_parallel A B"
       
   896   shows "affine_parallel B A"
       
   897 proof -
       
   898   from assms obtain a where B: "B = (\<lambda>x. a + x) ` A"
       
   899     unfolding affine_parallel_def by auto
       
   900   have [simp]: "(\<lambda>x. x - a) = plus (- a)" by (simp add: fun_eq_iff)
       
   901   from B show ?thesis
       
   902     using translation_galois [of B a A]
       
   903     unfolding affine_parallel_def by auto
       
   904 qed
       
   905 
       
   906 lemma affine_parallel_assoc:
       
   907   assumes "affine_parallel A B"
       
   908     and "affine_parallel B C"
       
   909   shows "affine_parallel A C"
       
   910 proof -
       
   911   from assms obtain ab where "B = (\<lambda>x. ab + x) ` A"
       
   912     unfolding affine_parallel_def by auto
       
   913   moreover
       
   914   from assms obtain bc where "C = (\<lambda>x. bc + x) ` B"
       
   915     unfolding affine_parallel_def by auto
       
   916   ultimately show ?thesis
       
   917     using translation_assoc[of bc ab A] unfolding affine_parallel_def by auto
       
   918 qed
       
   919 
       
   920 lemma affine_translation_aux:
       
   921   fixes a :: "'a::real_vector"
       
   922   assumes "affine ((\<lambda>x. a + x) ` S)"
       
   923   shows "affine S"
       
   924 proof -
       
   925   {
       
   926     fix x y u v
       
   927     assume xy: "x \<in> S" "y \<in> S" "(u :: real) + v = 1"
       
   928     then have "(a + x) \<in> ((\<lambda>x. a + x) ` S)" "(a + y) \<in> ((\<lambda>x. a + x) ` S)"
       
   929       by auto
       
   930     then have h1: "u *\<^sub>R  (a + x) + v *\<^sub>R (a + y) \<in> (\<lambda>x. a + x) ` S"
       
   931       using xy assms unfolding affine_def by auto
       
   932     have "u *\<^sub>R (a + x) + v *\<^sub>R (a + y) = (u + v) *\<^sub>R a + (u *\<^sub>R x + v *\<^sub>R y)"
       
   933       by (simp add: algebra_simps)
       
   934     also have "\<dots> = a + (u *\<^sub>R x + v *\<^sub>R y)"
       
   935       using \<open>u + v = 1\<close> by auto
       
   936     ultimately have "a + (u *\<^sub>R x + v *\<^sub>R y) \<in> (\<lambda>x. a + x) ` S"
       
   937       using h1 by auto
       
   938     then have "u *\<^sub>R x + v *\<^sub>R y : S" by auto
       
   939   }
       
   940   then show ?thesis unfolding affine_def by auto
       
   941 qed
       
   942 
       
   943 lemma affine_translation:
       
   944   fixes a :: "'a::real_vector"
       
   945   shows "affine S \<longleftrightarrow> affine ((\<lambda>x. a + x) ` S)"
       
   946 proof -
       
   947   have "affine S \<Longrightarrow> affine ((\<lambda>x. a + x) ` S)"
       
   948     using affine_translation_aux[of "-a" "((\<lambda>x. a + x) ` S)"]
       
   949     using translation_assoc[of "-a" a S] by auto
       
   950   then show ?thesis using affine_translation_aux by auto
       
   951 qed
       
   952 
       
   953 lemma parallel_is_affine:
       
   954   fixes S T :: "'a::real_vector set"
       
   955   assumes "affine S" "affine_parallel S T"
       
   956   shows "affine T"
       
   957 proof -
       
   958   from assms obtain a where "T = (\<lambda>x. a + x) ` S"
       
   959     unfolding affine_parallel_def by auto
       
   960   then show ?thesis
       
   961     using affine_translation assms by auto
       
   962 qed
       
   963 
       
   964 lemma subspace_imp_affine: "subspace s \<Longrightarrow> affine s"
       
   965   unfolding subspace_def affine_def by auto
       
   966 
       
   967 
       
   968 subsubsection \<open>Subspace parallel to an affine set\<close>
       
   969 
       
   970 lemma subspace_affine: "subspace S \<longleftrightarrow> affine S \<and> 0 \<in> S"
       
   971 proof -
       
   972   have h0: "subspace S \<Longrightarrow> affine S \<and> 0 \<in> S"
       
   973     using subspace_imp_affine[of S] subspace_0 by auto
       
   974   {
       
   975     assume assm: "affine S \<and> 0 \<in> S"
       
   976     {
       
   977       fix c :: real
       
   978       fix x
       
   979       assume x: "x \<in> S"
       
   980       have "c *\<^sub>R x = (1-c) *\<^sub>R 0 + c *\<^sub>R x" by auto
       
   981       moreover
       
   982       have "(1 - c) *\<^sub>R 0 + c *\<^sub>R x \<in> S"
       
   983         using affine_alt[of S] assm x by auto
       
   984       ultimately have "c *\<^sub>R x \<in> S" by auto
       
   985     }
       
   986     then have h1: "\<forall>c. \<forall>x \<in> S. c *\<^sub>R x \<in> S" by auto
       
   987 
       
   988     {
       
   989       fix x y
       
   990       assume xy: "x \<in> S" "y \<in> S"
       
   991       define u where "u = (1 :: real)/2"
       
   992       have "(1/2) *\<^sub>R (x+y) = (1/2) *\<^sub>R (x+y)"
       
   993         by auto
       
   994       moreover
       
   995       have "(1/2) *\<^sub>R (x+y)=(1/2) *\<^sub>R x + (1-(1/2)) *\<^sub>R y"
       
   996         by (simp add: algebra_simps)
       
   997       moreover
       
   998       have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> S"
       
   999         using affine_alt[of S] assm xy by auto
       
  1000       ultimately
       
  1001       have "(1/2) *\<^sub>R (x+y) \<in> S"
       
  1002         using u_def by auto
       
  1003       moreover
       
  1004       have "x + y = 2 *\<^sub>R ((1/2) *\<^sub>R (x+y))"
       
  1005         by auto
       
  1006       ultimately
       
  1007       have "x + y \<in> S"
       
  1008         using h1[rule_format, of "(1/2) *\<^sub>R (x+y)" "2"] by auto
       
  1009     }
       
  1010     then have "\<forall>x \<in> S. \<forall>y \<in> S. x + y \<in> S"
       
  1011       by auto
       
  1012     then have "subspace S"
       
  1013       using h1 assm unfolding subspace_def by auto
       
  1014   }
       
  1015   then show ?thesis using h0 by metis
       
  1016 qed
       
  1017 
       
  1018 lemma affine_diffs_subspace:
       
  1019   assumes "affine S" "a \<in> S"
       
  1020   shows "subspace ((\<lambda>x. (-a)+x) ` S)"
       
  1021 proof -
       
  1022   have [simp]: "(\<lambda>x. x - a) = plus (- a)" by (simp add: fun_eq_iff)
       
  1023   have "affine ((\<lambda>x. (-a)+x) ` S)"
       
  1024     using  affine_translation assms by auto
       
  1025   moreover have "0 : ((\<lambda>x. (-a)+x) ` S)"
       
  1026     using assms exI[of "(\<lambda>x. x\<in>S \<and> -a+x = 0)" a] by auto
       
  1027   ultimately show ?thesis using subspace_affine by auto
       
  1028 qed
       
  1029 
       
  1030 lemma parallel_subspace_explicit:
       
  1031   assumes "affine S"
       
  1032     and "a \<in> S"
       
  1033   assumes "L \<equiv> {y. \<exists>x \<in> S. (-a) + x = y}"
       
  1034   shows "subspace L \<and> affine_parallel S L"
       
  1035 proof -
       
  1036   from assms have "L = plus (- a) ` S" by auto
       
  1037   then have par: "affine_parallel S L"
       
  1038     unfolding affine_parallel_def ..
       
  1039   then have "affine L" using assms parallel_is_affine by auto
       
  1040   moreover have "0 \<in> L"
       
  1041     using assms by auto
       
  1042   ultimately show ?thesis
       
  1043     using subspace_affine par by auto
       
  1044 qed
       
  1045 
       
  1046 lemma parallel_subspace_aux:
       
  1047   assumes "subspace A"
       
  1048     and "subspace B"
       
  1049     and "affine_parallel A B"
       
  1050   shows "A \<supseteq> B"
       
  1051 proof -
       
  1052   from assms obtain a where a: "\<forall>x. x \<in> A \<longleftrightarrow> a + x \<in> B"
       
  1053     using affine_parallel_expl[of A B] by auto
       
  1054   then have "-a \<in> A"
       
  1055     using assms subspace_0[of B] by auto
       
  1056   then have "a \<in> A"
       
  1057     using assms subspace_neg[of A "-a"] by auto
       
  1058   then show ?thesis
       
  1059     using assms a unfolding subspace_def by auto
       
  1060 qed
       
  1061 
       
  1062 lemma parallel_subspace:
       
  1063   assumes "subspace A"
       
  1064     and "subspace B"
       
  1065     and "affine_parallel A B"
       
  1066   shows "A = B"
       
  1067 proof
       
  1068   show "A \<supseteq> B"
       
  1069     using assms parallel_subspace_aux by auto
       
  1070   show "A \<subseteq> B"
       
  1071     using assms parallel_subspace_aux[of B A] affine_parallel_commut by auto
       
  1072 qed
       
  1073 
       
  1074 lemma affine_parallel_subspace:
       
  1075   assumes "affine S" "S \<noteq> {}"
       
  1076   shows "\<exists>!L. subspace L \<and> affine_parallel S L"
       
  1077 proof -
       
  1078   have ex: "\<exists>L. subspace L \<and> affine_parallel S L"
       
  1079     using assms parallel_subspace_explicit by auto
       
  1080   {
       
  1081     fix L1 L2
       
  1082     assume ass: "subspace L1 \<and> affine_parallel S L1" "subspace L2 \<and> affine_parallel S L2"
       
  1083     then have "affine_parallel L1 L2"
       
  1084       using affine_parallel_commut[of S L1] affine_parallel_assoc[of L1 S L2] by auto
       
  1085     then have "L1 = L2"
       
  1086       using ass parallel_subspace by auto
       
  1087   }
       
  1088   then show ?thesis using ex by auto
       
  1089 qed
       
  1090 
       
  1091 
       
  1092 subsection \<open>Cones\<close>
       
  1093 
       
  1094 definition cone :: "'a::real_vector set \<Rightarrow> bool"
       
  1095   where "cone s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>c\<ge>0. c *\<^sub>R x \<in> s)"
       
  1096 
       
  1097 lemma cone_empty[intro, simp]: "cone {}"
       
  1098   unfolding cone_def by auto
       
  1099 
       
  1100 lemma cone_univ[intro, simp]: "cone UNIV"
       
  1101   unfolding cone_def by auto
       
  1102 
       
  1103 lemma cone_Inter[intro]: "\<forall>s\<in>f. cone s \<Longrightarrow> cone (\<Inter>f)"
       
  1104   unfolding cone_def by auto
       
  1105 
       
  1106 lemma subspace_imp_cone: "subspace S \<Longrightarrow> cone S"
       
  1107   by (simp add: cone_def subspace_mul)
       
  1108 
       
  1109 
       
  1110 subsubsection \<open>Conic hull\<close>
       
  1111 
       
  1112 lemma cone_cone_hull: "cone (cone hull s)"
       
  1113   unfolding hull_def by auto
       
  1114 
       
  1115 lemma cone_hull_eq: "cone hull s = s \<longleftrightarrow> cone s"
       
  1116   apply (rule hull_eq)
       
  1117   using cone_Inter
       
  1118   unfolding subset_eq
       
  1119   apply auto
       
  1120   done
       
  1121 
       
  1122 lemma mem_cone:
       
  1123   assumes "cone S" "x \<in> S" "c \<ge> 0"
       
  1124   shows "c *\<^sub>R x : S"
       
  1125   using assms cone_def[of S] by auto
       
  1126 
       
  1127 lemma cone_contains_0:
       
  1128   assumes "cone S"
       
  1129   shows "S \<noteq> {} \<longleftrightarrow> 0 \<in> S"
       
  1130 proof -
       
  1131   {
       
  1132     assume "S \<noteq> {}"
       
  1133     then obtain a where "a \<in> S" by auto
       
  1134     then have "0 \<in> S"
       
  1135       using assms mem_cone[of S a 0] by auto
       
  1136   }
       
  1137   then show ?thesis by auto
       
  1138 qed
       
  1139 
       
  1140 lemma cone_0: "cone {0}"
       
  1141   unfolding cone_def by auto
       
  1142 
       
  1143 lemma cone_Union[intro]: "(\<forall>s\<in>f. cone s) \<longrightarrow> cone (\<Union>f)"
       
  1144   unfolding cone_def by blast
       
  1145 
       
  1146 lemma cone_iff:
       
  1147   assumes "S \<noteq> {}"
       
  1148   shows "cone S \<longleftrightarrow> 0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (op *\<^sub>R c) ` S = S)"
       
  1149 proof -
       
  1150   {
       
  1151     assume "cone S"
       
  1152     {
       
  1153       fix c :: real
       
  1154       assume "c > 0"
       
  1155       {
       
  1156         fix x
       
  1157         assume "x \<in> S"
       
  1158         then have "x \<in> (op *\<^sub>R c) ` S"
       
  1159           unfolding image_def
       
  1160           using \<open>cone S\<close> \<open>c>0\<close> mem_cone[of S x "1/c"]
       
  1161             exI[of "(\<lambda>t. t \<in> S \<and> x = c *\<^sub>R t)" "(1 / c) *\<^sub>R x"]
       
  1162           by auto
       
  1163       }
       
  1164       moreover
       
  1165       {
       
  1166         fix x
       
  1167         assume "x \<in> (op *\<^sub>R c) ` S"
       
  1168         then have "x \<in> S"
       
  1169           using \<open>cone S\<close> \<open>c > 0\<close>
       
  1170           unfolding cone_def image_def \<open>c > 0\<close> by auto
       
  1171       }
       
  1172       ultimately have "(op *\<^sub>R c) ` S = S" by auto
       
  1173     }
       
  1174     then have "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (op *\<^sub>R c) ` S = S)"
       
  1175       using \<open>cone S\<close> cone_contains_0[of S] assms by auto
       
  1176   }
       
  1177   moreover
       
  1178   {
       
  1179     assume a: "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (op *\<^sub>R c) ` S = S)"
       
  1180     {
       
  1181       fix x
       
  1182       assume "x \<in> S"
       
  1183       fix c1 :: real
       
  1184       assume "c1 \<ge> 0"
       
  1185       then have "c1 = 0 \<or> c1 > 0" by auto
       
  1186       then have "c1 *\<^sub>R x \<in> S" using a \<open>x \<in> S\<close> by auto
       
  1187     }
       
  1188     then have "cone S" unfolding cone_def by auto
       
  1189   }
       
  1190   ultimately show ?thesis by blast
       
  1191 qed
       
  1192 
       
  1193 lemma cone_hull_empty: "cone hull {} = {}"
       
  1194   by (metis cone_empty cone_hull_eq)
       
  1195 
       
  1196 lemma cone_hull_empty_iff: "S = {} \<longleftrightarrow> cone hull S = {}"
       
  1197   by (metis bot_least cone_hull_empty hull_subset xtrans(5))
       
  1198 
       
  1199 lemma cone_hull_contains_0: "S \<noteq> {} \<longleftrightarrow> 0 \<in> cone hull S"
       
  1200   using cone_cone_hull[of S] cone_contains_0[of "cone hull S"] cone_hull_empty_iff[of S]
       
  1201   by auto
       
  1202 
       
  1203 lemma mem_cone_hull:
       
  1204   assumes "x \<in> S" "c \<ge> 0"
       
  1205   shows "c *\<^sub>R x \<in> cone hull S"
       
  1206   by (metis assms cone_cone_hull hull_inc mem_cone)
       
  1207 
       
  1208 lemma cone_hull_expl: "cone hull S = {c *\<^sub>R x | c x. c \<ge> 0 \<and> x \<in> S}"
       
  1209   (is "?lhs = ?rhs")
       
  1210 proof -
       
  1211   {
       
  1212     fix x
       
  1213     assume "x \<in> ?rhs"
       
  1214     then obtain cx :: real and xx where x: "x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S"
       
  1215       by auto
       
  1216     fix c :: real
       
  1217     assume c: "c \<ge> 0"
       
  1218     then have "c *\<^sub>R x = (c * cx) *\<^sub>R xx"
       
  1219       using x by (simp add: algebra_simps)
       
  1220     moreover
       
  1221     have "c * cx \<ge> 0" using c x by auto
       
  1222     ultimately
       
  1223     have "c *\<^sub>R x \<in> ?rhs" using x by auto
       
  1224   }
       
  1225   then have "cone ?rhs"
       
  1226     unfolding cone_def by auto
       
  1227   then have "?rhs \<in> Collect cone"
       
  1228     unfolding mem_Collect_eq by auto
       
  1229   {
       
  1230     fix x
       
  1231     assume "x \<in> S"
       
  1232     then have "1 *\<^sub>R x \<in> ?rhs"
       
  1233       apply auto
       
  1234       apply (rule_tac x = 1 in exI)
       
  1235       apply auto
       
  1236       done
       
  1237     then have "x \<in> ?rhs" by auto
       
  1238   }
       
  1239   then have "S \<subseteq> ?rhs" by auto
       
  1240   then have "?lhs \<subseteq> ?rhs"
       
  1241     using \<open>?rhs \<in> Collect cone\<close> hull_minimal[of S "?rhs" "cone"] by auto
       
  1242   moreover
       
  1243   {
       
  1244     fix x
       
  1245     assume "x \<in> ?rhs"
       
  1246     then obtain cx :: real and xx where x: "x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S"
       
  1247       by auto
       
  1248     then have "xx \<in> cone hull S"
       
  1249       using hull_subset[of S] by auto
       
  1250     then have "x \<in> ?lhs"
       
  1251       using x cone_cone_hull[of S] cone_def[of "cone hull S"] by auto
       
  1252   }
       
  1253   ultimately show ?thesis by auto
       
  1254 qed
       
  1255 
       
  1256 lemma cone_closure:
       
  1257   fixes S :: "'a::real_normed_vector set"
       
  1258   assumes "cone S"
       
  1259   shows "cone (closure S)"
       
  1260 proof (cases "S = {}")
       
  1261   case True
       
  1262   then show ?thesis by auto
       
  1263 next
       
  1264   case False
       
  1265   then have "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` S = S)"
       
  1266     using cone_iff[of S] assms by auto
       
  1267   then have "0 \<in> closure S \<and> (\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` closure S = closure S)"
       
  1268     using closure_subset by (auto simp add: closure_scaleR)
       
  1269   then show ?thesis
       
  1270     using False cone_iff[of "closure S"] by auto
       
  1271 qed
       
  1272 
       
  1273 
       
  1274 subsection \<open>Affine dependence and consequential theorems (from Lars Schewe)\<close>
       
  1275 
       
  1276 definition affine_dependent :: "'a::real_vector set \<Rightarrow> bool"
       
  1277   where "affine_dependent s \<longleftrightarrow> (\<exists>x\<in>s. x \<in> affine hull (s - {x}))"
       
  1278 
       
  1279 lemma affine_dependent_subset:
       
  1280    "\<lbrakk>affine_dependent s; s \<subseteq> t\<rbrakk> \<Longrightarrow> affine_dependent t"
       
  1281 apply (simp add: affine_dependent_def Bex_def)
       
  1282 apply (blast dest: hull_mono [OF Diff_mono [OF _ subset_refl]])
       
  1283 done
       
  1284 
       
  1285 lemma affine_independent_subset:
       
  1286   shows "\<lbrakk>~ affine_dependent t; s \<subseteq> t\<rbrakk> \<Longrightarrow> ~ affine_dependent s"
       
  1287 by (metis affine_dependent_subset)
       
  1288 
       
  1289 lemma affine_independent_Diff:
       
  1290    "~ affine_dependent s \<Longrightarrow> ~ affine_dependent(s - t)"
       
  1291 by (meson Diff_subset affine_dependent_subset)
       
  1292 
       
  1293 lemma affine_dependent_explicit:
       
  1294   "affine_dependent p \<longleftrightarrow>
       
  1295     (\<exists>s u. finite s \<and> s \<subseteq> p \<and> setsum u s = 0 \<and>
       
  1296       (\<exists>v\<in>s. u v \<noteq> 0) \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = 0)"
       
  1297   unfolding affine_dependent_def affine_hull_explicit mem_Collect_eq
       
  1298   apply rule
       
  1299   apply (erule bexE, erule exE, erule exE)
       
  1300   apply (erule conjE)+
       
  1301   defer
       
  1302   apply (erule exE, erule exE)
       
  1303   apply (erule conjE)+
       
  1304   apply (erule bexE)
       
  1305 proof -
       
  1306   fix x s u
       
  1307   assume as: "x \<in> p" "finite s" "s \<noteq> {}" "s \<subseteq> p - {x}" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
       
  1308   have "x \<notin> s" using as(1,4) by auto
       
  1309   show "\<exists>s u. finite s \<and> s \<subseteq> p \<and> setsum u s = 0 \<and> (\<exists>v\<in>s. u v \<noteq> 0) \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
       
  1310     apply (rule_tac x="insert x s" in exI, rule_tac x="\<lambda>v. if v = x then - 1 else u v" in exI)
       
  1311     unfolding if_smult and setsum_clauses(2)[OF as(2)] and setsum_delta_notmem[OF \<open>x\<notin>s\<close>] and as
       
  1312     using as
       
  1313     apply auto
       
  1314     done
       
  1315 next
       
  1316   fix s u v
       
  1317   assume as: "finite s" "s \<subseteq> p" "setsum u s = 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0" "v \<in> s" "u v \<noteq> 0"
       
  1318   have "s \<noteq> {v}"
       
  1319     using as(3,6) by auto
       
  1320   then show "\<exists>x\<in>p. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p - {x} \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
       
  1321     apply (rule_tac x=v in bexI)
       
  1322     apply (rule_tac x="s - {v}" in exI)
       
  1323     apply (rule_tac x="\<lambda>x. - (1 / u v) * u x" in exI)
       
  1324     unfolding scaleR_scaleR[symmetric] and scaleR_right.setsum [symmetric]
       
  1325     unfolding setsum_right_distrib[symmetric] and setsum_diff1[OF as(1)]
       
  1326     using as
       
  1327     apply auto
       
  1328     done
       
  1329 qed
       
  1330 
       
  1331 lemma affine_dependent_explicit_finite:
       
  1332   fixes s :: "'a::real_vector set"
       
  1333   assumes "finite s"
       
  1334   shows "affine_dependent s \<longleftrightarrow>
       
  1335     (\<exists>u. setsum u s = 0 \<and> (\<exists>v\<in>s. u v \<noteq> 0) \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = 0)"
       
  1336   (is "?lhs = ?rhs")
       
  1337 proof
       
  1338   have *: "\<And>vt u v. (if vt then u v else 0) *\<^sub>R v = (if vt then (u v) *\<^sub>R v else 0::'a)"
       
  1339     by auto
       
  1340   assume ?lhs
       
  1341   then obtain t u v where
       
  1342     "finite t" "t \<subseteq> s" "setsum u t = 0" "v\<in>t" "u v \<noteq> 0"  "(\<Sum>v\<in>t. u v *\<^sub>R v) = 0"
       
  1343     unfolding affine_dependent_explicit by auto
       
  1344   then show ?rhs
       
  1345     apply (rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI)
       
  1346     apply auto unfolding * and setsum.inter_restrict[OF assms, symmetric]
       
  1347     unfolding Int_absorb1[OF \<open>t\<subseteq>s\<close>]
       
  1348     apply auto
       
  1349     done
       
  1350 next
       
  1351   assume ?rhs
       
  1352   then obtain u v where "setsum u s = 0"  "v\<in>s" "u v \<noteq> 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
       
  1353     by auto
       
  1354   then show ?lhs unfolding affine_dependent_explicit
       
  1355     using assms by auto
       
  1356 qed
       
  1357 
       
  1358 
       
  1359 subsection \<open>Connectedness of convex sets\<close>
       
  1360 
       
  1361 lemma connectedD:
       
  1362   "connected S \<Longrightarrow> open A \<Longrightarrow> open B \<Longrightarrow> S \<subseteq> A \<union> B \<Longrightarrow> A \<inter> B \<inter> S = {} \<Longrightarrow> A \<inter> S = {} \<or> B \<inter> S = {}"
       
  1363   by (rule Topological_Spaces.topological_space_class.connectedD)
       
  1364 
       
  1365 lemma convex_connected:
       
  1366   fixes s :: "'a::real_normed_vector set"
       
  1367   assumes "convex s"
       
  1368   shows "connected s"
       
  1369 proof (rule connectedI)
       
  1370   fix A B
       
  1371   assume "open A" "open B" "A \<inter> B \<inter> s = {}" "s \<subseteq> A \<union> B"
       
  1372   moreover
       
  1373   assume "A \<inter> s \<noteq> {}" "B \<inter> s \<noteq> {}"
       
  1374   then obtain a b where a: "a \<in> A" "a \<in> s" and b: "b \<in> B" "b \<in> s" by auto
       
  1375   define f where [abs_def]: "f u = u *\<^sub>R a + (1 - u) *\<^sub>R b" for u
       
  1376   then have "continuous_on {0 .. 1} f"
       
  1377     by (auto intro!: continuous_intros)
       
  1378   then have "connected (f ` {0 .. 1})"
       
  1379     by (auto intro!: connected_continuous_image)
       
  1380   note connectedD[OF this, of A B]
       
  1381   moreover have "a \<in> A \<inter> f ` {0 .. 1}"
       
  1382     using a by (auto intro!: image_eqI[of _ _ 1] simp: f_def)
       
  1383   moreover have "b \<in> B \<inter> f ` {0 .. 1}"
       
  1384     using b by (auto intro!: image_eqI[of _ _ 0] simp: f_def)
       
  1385   moreover have "f ` {0 .. 1} \<subseteq> s"
       
  1386     using \<open>convex s\<close> a b unfolding convex_def f_def by auto
       
  1387   ultimately show False by auto
       
  1388 qed
       
  1389 
       
  1390 corollary connected_UNIV[intro]: "connected (UNIV :: 'a::real_normed_vector set)"
       
  1391   by(simp add: convex_connected)
       
  1392 
       
  1393 proposition clopen:
       
  1394   fixes s :: "'a :: real_normed_vector set"
       
  1395   shows "closed s \<and> open s \<longleftrightarrow> s = {} \<or> s = UNIV"
       
  1396 apply (rule iffI)
       
  1397  apply (rule connected_UNIV [unfolded connected_clopen, rule_format])
       
  1398  apply (force simp add: open_openin closed_closedin, force)
       
  1399 done
       
  1400 
       
  1401 corollary compact_open:
       
  1402   fixes s :: "'a :: euclidean_space set"
       
  1403   shows "compact s \<and> open s \<longleftrightarrow> s = {}"
       
  1404   by (auto simp: compact_eq_bounded_closed clopen)
       
  1405 
       
  1406 corollary finite_imp_not_open:
       
  1407     fixes S :: "'a::{real_normed_vector, perfect_space} set"
       
  1408     shows "\<lbrakk>finite S; open S\<rbrakk> \<Longrightarrow> S={}"
       
  1409   using clopen [of S] finite_imp_closed not_bounded_UNIV by blast
       
  1410 
       
  1411 corollary empty_interior_finite:
       
  1412     fixes S :: "'a::{real_normed_vector, perfect_space} set"
       
  1413     shows "finite S \<Longrightarrow> interior S = {}"
       
  1414   by (metis interior_subset finite_subset open_interior [of S] finite_imp_not_open)
       
  1415 
       
  1416 text \<open>Balls, being convex, are connected.\<close>
       
  1417 
       
  1418 lemma convex_prod:
       
  1419   assumes "\<And>i. i \<in> Basis \<Longrightarrow> convex {x. P i x}"
       
  1420   shows "convex {x. \<forall>i\<in>Basis. P i (x\<bullet>i)}"
       
  1421   using assms unfolding convex_def
       
  1422   by (auto simp: inner_add_left)
       
  1423 
       
  1424 lemma convex_positive_orthant: "convex {x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i)}"
       
  1425   by (rule convex_prod) (simp add: atLeast_def[symmetric] convex_real_interval)
       
  1426 
       
  1427 lemma convex_local_global_minimum:
       
  1428   fixes s :: "'a::real_normed_vector set"
       
  1429   assumes "e > 0"
       
  1430     and "convex_on s f"
       
  1431     and "ball x e \<subseteq> s"
       
  1432     and "\<forall>y\<in>ball x e. f x \<le> f y"
       
  1433   shows "\<forall>y\<in>s. f x \<le> f y"
       
  1434 proof (rule ccontr)
       
  1435   have "x \<in> s" using assms(1,3) by auto
       
  1436   assume "\<not> ?thesis"
       
  1437   then obtain y where "y\<in>s" and y: "f x > f y" by auto
       
  1438   then have xy: "0 < dist x y"  by auto
       
  1439   then obtain u where "0 < u" "u \<le> 1" and u: "u < e / dist x y"
       
  1440     using real_lbound_gt_zero[of 1 "e / dist x y"] xy \<open>e>0\<close> by auto
       
  1441   then have "f ((1-u) *\<^sub>R x + u *\<^sub>R y) \<le> (1-u) * f x + u * f y"
       
  1442     using \<open>x\<in>s\<close> \<open>y\<in>s\<close>
       
  1443     using assms(2)[unfolded convex_on_def,
       
  1444       THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x="1-u"]]
       
  1445     by auto
       
  1446   moreover
       
  1447   have *: "x - ((1 - u) *\<^sub>R x + u *\<^sub>R y) = u *\<^sub>R (x - y)"
       
  1448     by (simp add: algebra_simps)
       
  1449   have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> ball x e"
       
  1450     unfolding mem_ball dist_norm
       
  1451     unfolding * and norm_scaleR and abs_of_pos[OF \<open>0<u\<close>]
       
  1452     unfolding dist_norm[symmetric]
       
  1453     using u
       
  1454     unfolding pos_less_divide_eq[OF xy]
       
  1455     by auto
       
  1456   then have "f x \<le> f ((1 - u) *\<^sub>R x + u *\<^sub>R y)"
       
  1457     using assms(4) by auto
       
  1458   ultimately show False
       
  1459     using mult_strict_left_mono[OF y \<open>u>0\<close>]
       
  1460     unfolding left_diff_distrib
       
  1461     by auto
       
  1462 qed
       
  1463 
       
  1464 lemma convex_ball [iff]:
       
  1465   fixes x :: "'a::real_normed_vector"
       
  1466   shows "convex (ball x e)"
       
  1467 proof (auto simp add: convex_def)
       
  1468   fix y z
       
  1469   assume yz: "dist x y < e" "dist x z < e"
       
  1470   fix u v :: real
       
  1471   assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1"
       
  1472   have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z"
       
  1473     using uv yz
       
  1474     using convex_on_dist [of "ball x e" x, unfolded convex_on_def,
       
  1475       THEN bspec[where x=y], THEN bspec[where x=z]]
       
  1476     by auto
       
  1477   then show "dist x (u *\<^sub>R y + v *\<^sub>R z) < e"
       
  1478     using convex_bound_lt[OF yz uv] by auto
       
  1479 qed
       
  1480 
       
  1481 lemma convex_cball [iff]:
       
  1482   fixes x :: "'a::real_normed_vector"
       
  1483   shows "convex (cball x e)"
       
  1484 proof -
       
  1485   {
       
  1486     fix y z
       
  1487     assume yz: "dist x y \<le> e" "dist x z \<le> e"
       
  1488     fix u v :: real
       
  1489     assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1"
       
  1490     have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z"
       
  1491       using uv yz
       
  1492       using convex_on_dist [of "cball x e" x, unfolded convex_on_def,
       
  1493         THEN bspec[where x=y], THEN bspec[where x=z]]
       
  1494       by auto
       
  1495     then have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> e"
       
  1496       using convex_bound_le[OF yz uv] by auto
       
  1497   }
       
  1498   then show ?thesis by (auto simp add: convex_def Ball_def)
       
  1499 qed
       
  1500 
       
  1501 lemma connected_ball [iff]:
       
  1502   fixes x :: "'a::real_normed_vector"
       
  1503   shows "connected (ball x e)"
       
  1504   using convex_connected convex_ball by auto
       
  1505 
       
  1506 lemma connected_cball [iff]:
       
  1507   fixes x :: "'a::real_normed_vector"
       
  1508   shows "connected (cball x e)"
       
  1509   using convex_connected convex_cball by auto
       
  1510 
       
  1511 
       
  1512 subsection \<open>Convex hull\<close>
       
  1513 
       
  1514 lemma convex_convex_hull [iff]: "convex (convex hull s)"
       
  1515   unfolding hull_def
       
  1516   using convex_Inter[of "{t. convex t \<and> s \<subseteq> t}"]
       
  1517   by auto
       
  1518 
       
  1519 lemma convex_hull_subset:
       
  1520     "s \<subseteq> convex hull t \<Longrightarrow> convex hull s \<subseteq> convex hull t"
       
  1521   by (simp add: convex_convex_hull subset_hull)
       
  1522 
       
  1523 lemma convex_hull_eq: "convex hull s = s \<longleftrightarrow> convex s"
       
  1524   by (metis convex_convex_hull hull_same)
       
  1525 
       
  1526 lemma bounded_convex_hull:
       
  1527   fixes s :: "'a::real_normed_vector set"
       
  1528   assumes "bounded s"
       
  1529   shows "bounded (convex hull s)"
       
  1530 proof -
       
  1531   from assms obtain B where B: "\<forall>x\<in>s. norm x \<le> B"
       
  1532     unfolding bounded_iff by auto
       
  1533   show ?thesis
       
  1534     apply (rule bounded_subset[OF bounded_cball, of _ 0 B])
       
  1535     unfolding subset_hull[of convex, OF convex_cball]
       
  1536     unfolding subset_eq mem_cball dist_norm using B
       
  1537     apply auto
       
  1538     done
       
  1539 qed
       
  1540 
       
  1541 lemma finite_imp_bounded_convex_hull:
       
  1542   fixes s :: "'a::real_normed_vector set"
       
  1543   shows "finite s \<Longrightarrow> bounded (convex hull s)"
       
  1544   using bounded_convex_hull finite_imp_bounded
       
  1545   by auto
       
  1546 
       
  1547 
       
  1548 subsubsection \<open>Convex hull is "preserved" by a linear function\<close>
       
  1549 
       
  1550 lemma convex_hull_linear_image:
       
  1551   assumes f: "linear f"
       
  1552   shows "f ` (convex hull s) = convex hull (f ` s)"
       
  1553 proof
       
  1554   show "convex hull (f ` s) \<subseteq> f ` (convex hull s)"
       
  1555     by (intro hull_minimal image_mono hull_subset convex_linear_image assms convex_convex_hull)
       
  1556   show "f ` (convex hull s) \<subseteq> convex hull (f ` s)"
       
  1557   proof (unfold image_subset_iff_subset_vimage, rule hull_minimal)
       
  1558     show "s \<subseteq> f -` (convex hull (f ` s))"
       
  1559       by (fast intro: hull_inc)
       
  1560     show "convex (f -` (convex hull (f ` s)))"
       
  1561       by (intro convex_linear_vimage [OF f] convex_convex_hull)
       
  1562   qed
       
  1563 qed
       
  1564 
       
  1565 lemma in_convex_hull_linear_image:
       
  1566   assumes "linear f"
       
  1567     and "x \<in> convex hull s"
       
  1568   shows "f x \<in> convex hull (f ` s)"
       
  1569   using convex_hull_linear_image[OF assms(1)] assms(2) by auto
       
  1570 
       
  1571 lemma convex_hull_Times:
       
  1572   "convex hull (s \<times> t) = (convex hull s) \<times> (convex hull t)"
       
  1573 proof
       
  1574   show "convex hull (s \<times> t) \<subseteq> (convex hull s) \<times> (convex hull t)"
       
  1575     by (intro hull_minimal Sigma_mono hull_subset convex_Times convex_convex_hull)
       
  1576   have "\<forall>x\<in>convex hull s. \<forall>y\<in>convex hull t. (x, y) \<in> convex hull (s \<times> t)"
       
  1577   proof (intro hull_induct)
       
  1578     fix x y assume "x \<in> s" and "y \<in> t"
       
  1579     then show "(x, y) \<in> convex hull (s \<times> t)"
       
  1580       by (simp add: hull_inc)
       
  1581   next
       
  1582     fix x let ?S = "((\<lambda>y. (0, y)) -` (\<lambda>p. (- x, 0) + p) ` (convex hull s \<times> t))"
       
  1583     have "convex ?S"
       
  1584       by (intro convex_linear_vimage convex_translation convex_convex_hull,
       
  1585         simp add: linear_iff)
       
  1586     also have "?S = {y. (x, y) \<in> convex hull (s \<times> t)}"
       
  1587       by (auto simp add: image_def Bex_def)
       
  1588     finally show "convex {y. (x, y) \<in> convex hull (s \<times> t)}" .
       
  1589   next
       
  1590     show "convex {x. \<forall>y\<in>convex hull t. (x, y) \<in> convex hull (s \<times> t)}"
       
  1591     proof (unfold Collect_ball_eq, rule convex_INT [rule_format])
       
  1592       fix y let ?S = "((\<lambda>x. (x, 0)) -` (\<lambda>p. (0, - y) + p) ` (convex hull s \<times> t))"
       
  1593       have "convex ?S"
       
  1594       by (intro convex_linear_vimage convex_translation convex_convex_hull,
       
  1595         simp add: linear_iff)
       
  1596       also have "?S = {x. (x, y) \<in> convex hull (s \<times> t)}"
       
  1597         by (auto simp add: image_def Bex_def)
       
  1598       finally show "convex {x. (x, y) \<in> convex hull (s \<times> t)}" .
       
  1599     qed
       
  1600   qed
       
  1601   then show "(convex hull s) \<times> (convex hull t) \<subseteq> convex hull (s \<times> t)"
       
  1602     unfolding subset_eq split_paired_Ball_Sigma .
       
  1603 qed
       
  1604 
       
  1605 
       
  1606 subsubsection \<open>Stepping theorems for convex hulls of finite sets\<close>
       
  1607 
       
  1608 lemma convex_hull_empty[simp]: "convex hull {} = {}"
       
  1609   by (rule hull_unique) auto
       
  1610 
       
  1611 lemma convex_hull_singleton[simp]: "convex hull {a} = {a}"
       
  1612   by (rule hull_unique) auto
       
  1613 
       
  1614 lemma convex_hull_insert:
       
  1615   fixes s :: "'a::real_vector set"
       
  1616   assumes "s \<noteq> {}"
       
  1617   shows "convex hull (insert a s) =
       
  1618     {x. \<exists>u\<ge>0. \<exists>v\<ge>0. \<exists>b. (u + v = 1) \<and> b \<in> (convex hull s) \<and> (x = u *\<^sub>R a + v *\<^sub>R b)}"
       
  1619   (is "_ = ?hull")
       
  1620   apply (rule, rule hull_minimal, rule)
       
  1621   unfolding insert_iff
       
  1622   prefer 3
       
  1623   apply rule
       
  1624 proof -
       
  1625   fix x
       
  1626   assume x: "x = a \<or> x \<in> s"
       
  1627   then show "x \<in> ?hull"
       
  1628     apply rule
       
  1629     unfolding mem_Collect_eq
       
  1630     apply (rule_tac x=1 in exI)
       
  1631     defer
       
  1632     apply (rule_tac x=0 in exI)
       
  1633     using assms hull_subset[of s convex]
       
  1634     apply auto
       
  1635     done
       
  1636 next
       
  1637   fix x
       
  1638   assume "x \<in> ?hull"
       
  1639   then obtain u v b where obt: "u\<ge>0" "v\<ge>0" "u + v = 1" "b \<in> convex hull s" "x = u *\<^sub>R a + v *\<^sub>R b"
       
  1640     by auto
       
  1641   have "a \<in> convex hull insert a s" "b \<in> convex hull insert a s"
       
  1642     using hull_mono[of s "insert a s" convex] hull_mono[of "{a}" "insert a s" convex] and obt(4)
       
  1643     by auto
       
  1644   then show "x \<in> convex hull insert a s"
       
  1645     unfolding obt(5) using obt(1-3)
       
  1646     by (rule convexD [OF convex_convex_hull])
       
  1647 next
       
  1648   show "convex ?hull"
       
  1649   proof (rule convexI)
       
  1650     fix x y u v
       
  1651     assume as: "(0::real) \<le> u" "0 \<le> v" "u + v = 1" "x\<in>?hull" "y\<in>?hull"
       
  1652     from as(4) obtain u1 v1 b1 where
       
  1653       obt1: "u1\<ge>0" "v1\<ge>0" "u1 + v1 = 1" "b1 \<in> convex hull s" "x = u1 *\<^sub>R a + v1 *\<^sub>R b1"
       
  1654       by auto
       
  1655     from as(5) obtain u2 v2 b2 where
       
  1656       obt2: "u2\<ge>0" "v2\<ge>0" "u2 + v2 = 1" "b2 \<in> convex hull s" "y = u2 *\<^sub>R a + v2 *\<^sub>R b2"
       
  1657       by auto
       
  1658     have *: "\<And>(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x"
       
  1659       by (auto simp add: algebra_simps)
       
  1660     have **: "\<exists>b \<in> convex hull s. u *\<^sub>R x + v *\<^sub>R y =
       
  1661       (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (b - (u * u1) *\<^sub>R b - (v * u2) *\<^sub>R b)"
       
  1662     proof (cases "u * v1 + v * v2 = 0")
       
  1663       case True
       
  1664       have *: "\<And>(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x"
       
  1665         by (auto simp add: algebra_simps)
       
  1666       from True have ***: "u * v1 = 0" "v * v2 = 0"
       
  1667         using mult_nonneg_nonneg[OF \<open>u\<ge>0\<close> \<open>v1\<ge>0\<close>] mult_nonneg_nonneg[OF \<open>v\<ge>0\<close> \<open>v2\<ge>0\<close>]
       
  1668         by arith+
       
  1669       then have "u * u1 + v * u2 = 1"
       
  1670         using as(3) obt1(3) obt2(3) by auto
       
  1671       then show ?thesis
       
  1672         unfolding obt1(5) obt2(5) *
       
  1673         using assms hull_subset[of s convex]
       
  1674         by (auto simp add: *** scaleR_right_distrib)
       
  1675     next
       
  1676       case False
       
  1677       have "1 - (u * u1 + v * u2) = (u + v) - (u * u1 + v * u2)"
       
  1678         using as(3) obt1(3) obt2(3) by (auto simp add: field_simps)
       
  1679       also have "\<dots> = u * (v1 + u1 - u1) + v * (v2 + u2 - u2)"
       
  1680         using as(3) obt1(3) obt2(3) by (auto simp add: field_simps)
       
  1681       also have "\<dots> = u * v1 + v * v2"
       
  1682         by simp
       
  1683       finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" by auto
       
  1684       have "0 \<le> u * v1 + v * v2" "0 \<le> u * v1" "0 \<le> u * v1 + v * v2" "0 \<le> v * v2"
       
  1685         using as(1,2) obt1(1,2) obt2(1,2) by auto
       
  1686       then show ?thesis
       
  1687         unfolding obt1(5) obt2(5)
       
  1688         unfolding * and **
       
  1689         using False
       
  1690         apply (rule_tac
       
  1691           x = "((u * v1) / (u * v1 + v * v2)) *\<^sub>R b1 + ((v * v2) / (u * v1 + v * v2)) *\<^sub>R b2" in bexI)
       
  1692         defer
       
  1693         apply (rule convexD [OF convex_convex_hull])
       
  1694         using obt1(4) obt2(4)
       
  1695         unfolding add_divide_distrib[symmetric] and zero_le_divide_iff
       
  1696         apply (auto simp add: scaleR_left_distrib scaleR_right_distrib)
       
  1697         done
       
  1698     qed
       
  1699     have u1: "u1 \<le> 1"
       
  1700       unfolding obt1(3)[symmetric] and not_le using obt1(2) by auto
       
  1701     have u2: "u2 \<le> 1"
       
  1702       unfolding obt2(3)[symmetric] and not_le using obt2(2) by auto
       
  1703     have "u1 * u + u2 * v \<le> max u1 u2 * u + max u1 u2 * v"
       
  1704       apply (rule add_mono)
       
  1705       apply (rule_tac [!] mult_right_mono)
       
  1706       using as(1,2) obt1(1,2) obt2(1,2)
       
  1707       apply auto
       
  1708       done
       
  1709     also have "\<dots> \<le> 1"
       
  1710       unfolding distrib_left[symmetric] and as(3) using u1 u2 by auto
       
  1711     finally show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull"
       
  1712       unfolding mem_Collect_eq
       
  1713       apply (rule_tac x="u * u1 + v * u2" in exI)
       
  1714       apply (rule conjI)
       
  1715       defer
       
  1716       apply (rule_tac x="1 - u * u1 - v * u2" in exI)
       
  1717       unfolding Bex_def
       
  1718       using as(1,2) obt1(1,2) obt2(1,2) **
       
  1719       apply (auto simp add: algebra_simps)
       
  1720       done
       
  1721   qed
       
  1722 qed
       
  1723 
       
  1724 
       
  1725 subsubsection \<open>Explicit expression for convex hull\<close>
       
  1726 
       
  1727 lemma convex_hull_indexed:
       
  1728   fixes s :: "'a::real_vector set"
       
  1729   shows "convex hull s =
       
  1730     {y. \<exists>k u x.
       
  1731       (\<forall>i\<in>{1::nat .. k}. 0 \<le> u i \<and> x i \<in> s) \<and>
       
  1732       (setsum u {1..k} = 1) \<and> (setsum (\<lambda>i. u i *\<^sub>R x i) {1..k} = y)}"
       
  1733   (is "?xyz = ?hull")
       
  1734   apply (rule hull_unique)
       
  1735   apply rule
       
  1736   defer
       
  1737   apply (rule convexI)
       
  1738 proof -
       
  1739   fix x
       
  1740   assume "x\<in>s"
       
  1741   then show "x \<in> ?hull"
       
  1742     unfolding mem_Collect_eq
       
  1743     apply (rule_tac x=1 in exI, rule_tac x="\<lambda>x. 1" in exI)
       
  1744     apply auto
       
  1745     done
       
  1746 next
       
  1747   fix t
       
  1748   assume as: "s \<subseteq> t" "convex t"
       
  1749   show "?hull \<subseteq> t"
       
  1750     apply rule
       
  1751     unfolding mem_Collect_eq
       
  1752     apply (elim exE conjE)
       
  1753   proof -
       
  1754     fix x k u y
       
  1755     assume assm:
       
  1756       "\<forall>i\<in>{1::nat..k}. 0 \<le> u i \<and> y i \<in> s"
       
  1757       "setsum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R y i) = x"
       
  1758     show "x\<in>t"
       
  1759       unfolding assm(3) [symmetric]
       
  1760       apply (rule as(2)[unfolded convex, rule_format])
       
  1761       using assm(1,2) as(1) apply auto
       
  1762       done
       
  1763   qed
       
  1764 next
       
  1765   fix x y u v
       
  1766   assume uv: "0 \<le> u" "0 \<le> v" "u + v = (1::real)"
       
  1767   assume xy: "x \<in> ?hull" "y \<in> ?hull"
       
  1768   from xy obtain k1 u1 x1 where
       
  1769     x: "\<forall>i\<in>{1::nat..k1}. 0\<le>u1 i \<and> x1 i \<in> s" "setsum u1 {Suc 0..k1} = 1" "(\<Sum>i = Suc 0..k1. u1 i *\<^sub>R x1 i) = x"
       
  1770     by auto
       
  1771   from xy obtain k2 u2 x2 where
       
  1772     y: "\<forall>i\<in>{1::nat..k2}. 0\<le>u2 i \<and> x2 i \<in> s" "setsum u2 {Suc 0..k2} = 1" "(\<Sum>i = Suc 0..k2. u2 i *\<^sub>R x2 i) = y"
       
  1773     by auto
       
  1774   have *: "\<And>P (x1::'a) x2 s1 s2 i.
       
  1775     (if P i then s1 else s2) *\<^sub>R (if P i then x1 else x2) = (if P i then s1 *\<^sub>R x1 else s2 *\<^sub>R x2)"
       
  1776     "{1..k1 + k2} \<inter> {1..k1} = {1..k1}" "{1..k1 + k2} \<inter> - {1..k1} = (\<lambda>i. i + k1) ` {1..k2}"
       
  1777     prefer 3
       
  1778     apply (rule, rule)
       
  1779     unfolding image_iff
       
  1780     apply (rule_tac x = "x - k1" in bexI)
       
  1781     apply (auto simp add: not_le)
       
  1782     done
       
  1783   have inj: "inj_on (\<lambda>i. i + k1) {1..k2}"
       
  1784     unfolding inj_on_def by auto
       
  1785   show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull"
       
  1786     apply rule
       
  1787     apply (rule_tac x="k1 + k2" in exI)
       
  1788     apply (rule_tac x="\<lambda>i. if i \<in> {1..k1} then u * u1 i else v * u2 (i - k1)" in exI)
       
  1789     apply (rule_tac x="\<lambda>i. if i \<in> {1..k1} then x1 i else x2 (i - k1)" in exI)
       
  1790     apply (rule, rule)
       
  1791     defer
       
  1792     apply rule
       
  1793     unfolding * and setsum.If_cases[OF finite_atLeastAtMost[of 1 "k1 + k2"]] and
       
  1794       setsum.reindex[OF inj] and o_def Collect_mem_eq
       
  1795     unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] setsum_right_distrib[symmetric]
       
  1796   proof -
       
  1797     fix i
       
  1798     assume i: "i \<in> {1..k1+k2}"
       
  1799     show "0 \<le> (if i \<in> {1..k1} then u * u1 i else v * u2 (i - k1)) \<and>
       
  1800       (if i \<in> {1..k1} then x1 i else x2 (i - k1)) \<in> s"
       
  1801     proof (cases "i\<in>{1..k1}")
       
  1802       case True
       
  1803       then show ?thesis
       
  1804         using uv(1) x(1)[THEN bspec[where x=i]] by auto
       
  1805     next
       
  1806       case False
       
  1807       define j where "j = i - k1"
       
  1808       from i False have "j \<in> {1..k2}"
       
  1809         unfolding j_def by auto
       
  1810       then show ?thesis
       
  1811         using False uv(2) y(1)[THEN bspec[where x=j]]
       
  1812         by (auto simp: j_def[symmetric])
       
  1813     qed
       
  1814   qed (auto simp add: not_le x(2,3) y(2,3) uv(3))
       
  1815 qed
       
  1816 
       
  1817 lemma convex_hull_finite:
       
  1818   fixes s :: "'a::real_vector set"
       
  1819   assumes "finite s"
       
  1820   shows "convex hull s = {y. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and>
       
  1821     setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}"
       
  1822   (is "?HULL = ?set")
       
  1823 proof (rule hull_unique, auto simp add: convex_def[of ?set])
       
  1824   fix x
       
  1825   assume "x \<in> s"
       
  1826   then show "\<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = x"
       
  1827     apply (rule_tac x="\<lambda>y. if x=y then 1 else 0" in exI)
       
  1828     apply auto
       
  1829     unfolding setsum.delta'[OF assms] and setsum_delta''[OF assms]
       
  1830     apply auto
       
  1831     done
       
  1832 next
       
  1833   fix u v :: real
       
  1834   assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1"
       
  1835   fix ux assume ux: "\<forall>x\<in>s. 0 \<le> ux x" "setsum ux s = (1::real)"
       
  1836   fix uy assume uy: "\<forall>x\<in>s. 0 \<le> uy x" "setsum uy s = (1::real)"
       
  1837   {
       
  1838     fix x
       
  1839     assume "x\<in>s"
       
  1840     then have "0 \<le> u * ux x + v * uy x"
       
  1841       using ux(1)[THEN bspec[where x=x]] uy(1)[THEN bspec[where x=x]] and uv(1,2)
       
  1842       by auto
       
  1843   }
       
  1844   moreover
       
  1845   have "(\<Sum>x\<in>s. u * ux x + v * uy x) = 1"
       
  1846     unfolding setsum.distrib and setsum_right_distrib[symmetric] and ux(2) uy(2)
       
  1847     using uv(3) by auto
       
  1848   moreover
       
  1849   have "(\<Sum>x\<in>s. (u * ux x + v * uy x) *\<^sub>R x) = u *\<^sub>R (\<Sum>x\<in>s. ux x *\<^sub>R x) + v *\<^sub>R (\<Sum>x\<in>s. uy x *\<^sub>R x)"
       
  1850     unfolding scaleR_left_distrib and setsum.distrib and scaleR_scaleR[symmetric]
       
  1851       and scaleR_right.setsum [symmetric]
       
  1852     by auto
       
  1853   ultimately
       
  1854   show "\<exists>uc. (\<forall>x\<in>s. 0 \<le> uc x) \<and> setsum uc s = 1 \<and>
       
  1855       (\<Sum>x\<in>s. uc x *\<^sub>R x) = u *\<^sub>R (\<Sum>x\<in>s. ux x *\<^sub>R x) + v *\<^sub>R (\<Sum>x\<in>s. uy x *\<^sub>R x)"
       
  1856     apply (rule_tac x="\<lambda>x. u * ux x + v * uy x" in exI)
       
  1857     apply auto
       
  1858     done
       
  1859 next
       
  1860   fix t
       
  1861   assume t: "s \<subseteq> t" "convex t"
       
  1862   fix u
       
  1863   assume u: "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = (1::real)"
       
  1864   then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> t"
       
  1865     using t(2)[unfolded convex_explicit, THEN spec[where x=s], THEN spec[where x=u]]
       
  1866     using assms and t(1) by auto
       
  1867 qed
       
  1868 
       
  1869 
       
  1870 subsubsection \<open>Another formulation from Lars Schewe\<close>
       
  1871 
       
  1872 lemma convex_hull_explicit:
       
  1873   fixes p :: "'a::real_vector set"
       
  1874   shows "convex hull p =
       
  1875     {y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = y}"
       
  1876   (is "?lhs = ?rhs")
       
  1877 proof -
       
  1878   {
       
  1879     fix x
       
  1880     assume "x\<in>?lhs"
       
  1881     then obtain k u y where
       
  1882         obt: "\<forall>i\<in>{1::nat..k}. 0 \<le> u i \<and> y i \<in> p" "setsum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R y i) = x"
       
  1883       unfolding convex_hull_indexed by auto
       
  1884 
       
  1885     have fin: "finite {1..k}" by auto
       
  1886     have fin': "\<And>v. finite {i \<in> {1..k}. y i = v}" by auto
       
  1887     {
       
  1888       fix j
       
  1889       assume "j\<in>{1..k}"
       
  1890       then have "y j \<in> p" "0 \<le> setsum u {i. Suc 0 \<le> i \<and> i \<le> k \<and> y i = y j}"
       
  1891         using obt(1)[THEN bspec[where x=j]] and obt(2)
       
  1892         apply simp
       
  1893         apply (rule setsum_nonneg)
       
  1894         using obt(1)
       
  1895         apply auto
       
  1896         done
       
  1897     }
       
  1898     moreover
       
  1899     have "(\<Sum>v\<in>y ` {1..k}. setsum u {i \<in> {1..k}. y i = v}) = 1"
       
  1900       unfolding setsum_image_gen[OF fin, symmetric] using obt(2) by auto
       
  1901     moreover have "(\<Sum>v\<in>y ` {1..k}. setsum u {i \<in> {1..k}. y i = v} *\<^sub>R v) = x"
       
  1902       using setsum_image_gen[OF fin, of "\<lambda>i. u i *\<^sub>R y i" y, symmetric]
       
  1903       unfolding scaleR_left.setsum using obt(3) by auto
       
  1904     ultimately
       
  1905     have "\<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
       
  1906       apply (rule_tac x="y ` {1..k}" in exI)
       
  1907       apply (rule_tac x="\<lambda>v. setsum u {i\<in>{1..k}. y i = v}" in exI)
       
  1908       apply auto
       
  1909       done
       
  1910     then have "x\<in>?rhs" by auto
       
  1911   }
       
  1912   moreover
       
  1913   {
       
  1914     fix y
       
  1915     assume "y\<in>?rhs"
       
  1916     then obtain s u where
       
  1917       obt: "finite s" "s \<subseteq> p" "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = y"
       
  1918       by auto
       
  1919 
       
  1920     obtain f where f: "inj_on f {1..card s}" "f ` {1..card s} = s"
       
  1921       using ex_bij_betw_nat_finite_1[OF obt(1)] unfolding bij_betw_def by auto
       
  1922 
       
  1923     {
       
  1924       fix i :: nat
       
  1925       assume "i\<in>{1..card s}"
       
  1926       then have "f i \<in> s"
       
  1927         apply (subst f(2)[symmetric])
       
  1928         apply auto
       
  1929         done
       
  1930       then have "0 \<le> u (f i)" "f i \<in> p" using obt(2,3) by auto
       
  1931     }
       
  1932     moreover have *: "finite {1..card s}" by auto
       
  1933     {
       
  1934       fix y
       
  1935       assume "y\<in>s"
       
  1936       then obtain i where "i\<in>{1..card s}" "f i = y"
       
  1937         using f using image_iff[of y f "{1..card s}"]
       
  1938         by auto
       
  1939       then have "{x. Suc 0 \<le> x \<and> x \<le> card s \<and> f x = y} = {i}"
       
  1940         apply auto
       
  1941         using f(1)[unfolded inj_on_def]
       
  1942         apply(erule_tac x=x in ballE)
       
  1943         apply auto
       
  1944         done
       
  1945       then have "card {x. Suc 0 \<le> x \<and> x \<le> card s \<and> f x = y} = 1" by auto
       
  1946       then have "(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x)) = u y"
       
  1947           "(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y"
       
  1948         by (auto simp add: setsum_constant_scaleR)
       
  1949     }
       
  1950     then have "(\<Sum>x = 1..card s. u (f x)) = 1" "(\<Sum>i = 1..card s. u (f i) *\<^sub>R f i) = y"
       
  1951       unfolding setsum_image_gen[OF *(1), of "\<lambda>x. u (f x) *\<^sub>R f x" f]
       
  1952         and setsum_image_gen[OF *(1), of "\<lambda>x. u (f x)" f]
       
  1953       unfolding f
       
  1954       using setsum.cong [of s s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x) *\<^sub>R f x)" "\<lambda>v. u v *\<^sub>R v"]
       
  1955       using setsum.cong [of s s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x))" u]
       
  1956       unfolding obt(4,5)
       
  1957       by auto
       
  1958     ultimately
       
  1959     have "\<exists>k u x. (\<forall>i\<in>{1..k}. 0 \<le> u i \<and> x i \<in> p) \<and> setsum u {1..k} = 1 \<and>
       
  1960         (\<Sum>i::nat = 1..k. u i *\<^sub>R x i) = y"
       
  1961       apply (rule_tac x="card s" in exI)
       
  1962       apply (rule_tac x="u \<circ> f" in exI)
       
  1963       apply (rule_tac x=f in exI)
       
  1964       apply fastforce
       
  1965       done
       
  1966     then have "y \<in> ?lhs"
       
  1967       unfolding convex_hull_indexed by auto
       
  1968   }
       
  1969   ultimately show ?thesis
       
  1970     unfolding set_eq_iff by blast
       
  1971 qed
       
  1972 
       
  1973 
       
  1974 subsubsection \<open>A stepping theorem for that expansion\<close>
       
  1975 
       
  1976 lemma convex_hull_finite_step:
       
  1977   fixes s :: "'a::real_vector set"
       
  1978   assumes "finite s"
       
  1979   shows
       
  1980     "(\<exists>u. (\<forall>x\<in>insert a s. 0 \<le> u x) \<and> setsum u (insert a s) = w \<and> setsum (\<lambda>x. u x *\<^sub>R x) (insert a s) = y)
       
  1981       \<longleftrightarrow> (\<exists>v\<ge>0. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = w - v \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y - v *\<^sub>R a)"
       
  1982   (is "?lhs = ?rhs")
       
  1983 proof (rule, case_tac[!] "a\<in>s")
       
  1984   assume "a \<in> s"
       
  1985   then have *: "insert a s = s" by auto
       
  1986   assume ?lhs
       
  1987   then show ?rhs
       
  1988     unfolding *
       
  1989     apply (rule_tac x=0 in exI)
       
  1990     apply auto
       
  1991     done
       
  1992 next
       
  1993   assume ?lhs
       
  1994   then obtain u where
       
  1995       u: "\<forall>x\<in>insert a s. 0 \<le> u x" "setsum u (insert a s) = w" "(\<Sum>x\<in>insert a s. u x *\<^sub>R x) = y"
       
  1996     by auto
       
  1997   assume "a \<notin> s"
       
  1998   then show ?rhs
       
  1999     apply (rule_tac x="u a" in exI)
       
  2000     using u(1)[THEN bspec[where x=a]]
       
  2001     apply simp
       
  2002     apply (rule_tac x=u in exI)
       
  2003     using u[unfolded setsum_clauses(2)[OF assms]] and \<open>a\<notin>s\<close>
       
  2004     apply auto
       
  2005     done
       
  2006 next
       
  2007   assume "a \<in> s"
       
  2008   then have *: "insert a s = s" by auto
       
  2009   have fin: "finite (insert a s)" using assms by auto
       
  2010   assume ?rhs
       
  2011   then obtain v u where uv: "v\<ge>0" "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = w - v" "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a"
       
  2012     by auto
       
  2013   show ?lhs
       
  2014     apply (rule_tac x = "\<lambda>x. (if a = x then v else 0) + u x" in exI)
       
  2015     unfolding scaleR_left_distrib and setsum.distrib and setsum_delta''[OF fin] and setsum.delta'[OF fin]
       
  2016     unfolding setsum_clauses(2)[OF assms]
       
  2017     using uv and uv(2)[THEN bspec[where x=a]] and \<open>a\<in>s\<close>
       
  2018     apply auto
       
  2019     done
       
  2020 next
       
  2021   assume ?rhs
       
  2022   then obtain v u where
       
  2023     uv: "v\<ge>0" "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = w - v" "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a"
       
  2024     by auto
       
  2025   moreover
       
  2026   assume "a \<notin> s"
       
  2027   moreover
       
  2028   have "(\<Sum>x\<in>s. if a = x then v else u x) = setsum u s"
       
  2029     and "(\<Sum>x\<in>s. (if a = x then v else u x) *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x)"
       
  2030     apply (rule_tac setsum.cong) apply rule
       
  2031     defer
       
  2032     apply (rule_tac setsum.cong) apply rule
       
  2033     using \<open>a \<notin> s\<close>
       
  2034     apply auto
       
  2035     done
       
  2036   ultimately show ?lhs
       
  2037     apply (rule_tac x="\<lambda>x. if a = x then v else u x" in exI)
       
  2038     unfolding setsum_clauses(2)[OF assms]
       
  2039     apply auto
       
  2040     done
       
  2041 qed
       
  2042 
       
  2043 
       
  2044 subsubsection \<open>Hence some special cases\<close>
       
  2045 
       
  2046 lemma convex_hull_2:
       
  2047   "convex hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b | u v. 0 \<le> u \<and> 0 \<le> v \<and> u + v = 1}"
       
  2048 proof -
       
  2049   have *: "\<And>u. (\<forall>x\<in>{a, b}. 0 \<le> u x) \<longleftrightarrow> 0 \<le> u a \<and> 0 \<le> u b"
       
  2050     by auto
       
  2051   have **: "finite {b}" by auto
       
  2052   show ?thesis
       
  2053     apply (simp add: convex_hull_finite)
       
  2054     unfolding convex_hull_finite_step[OF **, of a 1, unfolded * conj_assoc]
       
  2055     apply auto
       
  2056     apply (rule_tac x=v in exI)
       
  2057     apply (rule_tac x="1 - v" in exI)
       
  2058     apply simp
       
  2059     apply (rule_tac x=u in exI)
       
  2060     apply simp
       
  2061     apply (rule_tac x="\<lambda>x. v" in exI)
       
  2062     apply simp
       
  2063     done
       
  2064 qed
       
  2065 
       
  2066 lemma convex_hull_2_alt: "convex hull {a,b} = {a + u *\<^sub>R (b - a) | u.  0 \<le> u \<and> u \<le> 1}"
       
  2067   unfolding convex_hull_2
       
  2068 proof (rule Collect_cong)
       
  2069   have *: "\<And>x y ::real. x + y = 1 \<longleftrightarrow> x = 1 - y"
       
  2070     by auto
       
  2071   fix x
       
  2072   show "(\<exists>v u. x = v *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> v \<and> 0 \<le> u \<and> v + u = 1) \<longleftrightarrow>
       
  2073     (\<exists>u. x = a + u *\<^sub>R (b - a) \<and> 0 \<le> u \<and> u \<le> 1)"
       
  2074     unfolding *
       
  2075     apply auto
       
  2076     apply (rule_tac[!] x=u in exI)
       
  2077     apply (auto simp add: algebra_simps)
       
  2078     done
       
  2079 qed
       
  2080 
       
  2081 lemma convex_hull_3:
       
  2082   "convex hull {a,b,c} = { u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c | u v w. 0 \<le> u \<and> 0 \<le> v \<and> 0 \<le> w \<and> u + v + w = 1}"
       
  2083 proof -
       
  2084   have fin: "finite {a,b,c}" "finite {b,c}" "finite {c}"
       
  2085     by auto
       
  2086   have *: "\<And>x y z ::real. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z"
       
  2087     by (auto simp add: field_simps)
       
  2088   show ?thesis
       
  2089     unfolding convex_hull_finite[OF fin(1)] and convex_hull_finite_step[OF fin(2)] and *
       
  2090     unfolding convex_hull_finite_step[OF fin(3)]
       
  2091     apply (rule Collect_cong)
       
  2092     apply simp
       
  2093     apply auto
       
  2094     apply (rule_tac x=va in exI)
       
  2095     apply (rule_tac x="u c" in exI)
       
  2096     apply simp
       
  2097     apply (rule_tac x="1 - v - w" in exI)
       
  2098     apply simp
       
  2099     apply (rule_tac x=v in exI)
       
  2100     apply simp
       
  2101     apply (rule_tac x="\<lambda>x. w" in exI)
       
  2102     apply simp
       
  2103     done
       
  2104 qed
       
  2105 
       
  2106 lemma convex_hull_3_alt:
       
  2107   "convex hull {a,b,c} = {a + u *\<^sub>R (b - a) + v *\<^sub>R (c - a) | u v.  0 \<le> u \<and> 0 \<le> v \<and> u + v \<le> 1}"
       
  2108 proof -
       
  2109   have *: "\<And>x y z ::real. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z"
       
  2110     by auto
       
  2111   show ?thesis
       
  2112     unfolding convex_hull_3
       
  2113     apply (auto simp add: *)
       
  2114     apply (rule_tac x=v in exI)
       
  2115     apply (rule_tac x=w in exI)
       
  2116     apply (simp add: algebra_simps)
       
  2117     apply (rule_tac x=u in exI)
       
  2118     apply (rule_tac x=v in exI)
       
  2119     apply (simp add: algebra_simps)
       
  2120     done
       
  2121 qed
       
  2122 
       
  2123 
       
  2124 subsection \<open>Relations among closure notions and corresponding hulls\<close>
       
  2125 
       
  2126 lemma affine_imp_convex: "affine s \<Longrightarrow> convex s"
       
  2127   unfolding affine_def convex_def by auto
       
  2128 
       
  2129 lemma subspace_imp_convex: "subspace s \<Longrightarrow> convex s"
       
  2130   using subspace_imp_affine affine_imp_convex by auto
       
  2131 
       
  2132 lemma affine_hull_subset_span: "(affine hull s) \<subseteq> (span s)"
       
  2133   by (metis hull_minimal span_inc subspace_imp_affine subspace_span)
       
  2134 
       
  2135 lemma convex_hull_subset_span: "(convex hull s) \<subseteq> (span s)"
       
  2136   by (metis hull_minimal span_inc subspace_imp_convex subspace_span)
       
  2137 
       
  2138 lemma convex_hull_subset_affine_hull: "(convex hull s) \<subseteq> (affine hull s)"
       
  2139   by (metis affine_affine_hull affine_imp_convex hull_minimal hull_subset)
       
  2140 
       
  2141 lemma affine_dependent_imp_dependent: "affine_dependent s \<Longrightarrow> dependent s"
       
  2142   unfolding affine_dependent_def dependent_def
       
  2143   using affine_hull_subset_span by auto
       
  2144 
       
  2145 lemma dependent_imp_affine_dependent:
       
  2146   assumes "dependent {x - a| x . x \<in> s}"
       
  2147     and "a \<notin> s"
       
  2148   shows "affine_dependent (insert a s)"
       
  2149 proof -
       
  2150   from assms(1)[unfolded dependent_explicit] obtain S u v
       
  2151     where obt: "finite S" "S \<subseteq> {x - a |x. x \<in> s}" "v\<in>S" "u v  \<noteq> 0" "(\<Sum>v\<in>S. u v *\<^sub>R v) = 0"
       
  2152     by auto
       
  2153   define t where "t = (\<lambda>x. x + a) ` S"
       
  2154 
       
  2155   have inj: "inj_on (\<lambda>x. x + a) S"
       
  2156     unfolding inj_on_def by auto
       
  2157   have "0 \<notin> S"
       
  2158     using obt(2) assms(2) unfolding subset_eq by auto
       
  2159   have fin: "finite t" and "t \<subseteq> s"
       
  2160     unfolding t_def using obt(1,2) by auto
       
  2161   then have "finite (insert a t)" and "insert a t \<subseteq> insert a s"
       
  2162     by auto
       
  2163   moreover have *: "\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x)) = (\<Sum>x\<in>t. Q x)"
       
  2164     apply (rule setsum.cong)
       
  2165     using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
       
  2166     apply auto
       
  2167     done
       
  2168   have "(\<Sum>x\<in>insert a t. if x = a then - (\<Sum>x\<in>t. u (x - a)) else u (x - a)) = 0"
       
  2169     unfolding setsum_clauses(2)[OF fin]
       
  2170     using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
       
  2171     apply auto
       
  2172     unfolding *
       
  2173     apply auto
       
  2174     done
       
  2175   moreover have "\<exists>v\<in>insert a t. (if v = a then - (\<Sum>x\<in>t. u (x - a)) else u (v - a)) \<noteq> 0"
       
  2176     apply (rule_tac x="v + a" in bexI)
       
  2177     using obt(3,4) and \<open>0\<notin>S\<close>
       
  2178     unfolding t_def
       
  2179     apply auto
       
  2180     done
       
  2181   moreover have *: "\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x) *\<^sub>R x) = (\<Sum>x\<in>t. Q x *\<^sub>R x)"
       
  2182     apply (rule setsum.cong)
       
  2183     using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
       
  2184     apply auto
       
  2185     done
       
  2186   have "(\<Sum>x\<in>t. u (x - a)) *\<^sub>R a = (\<Sum>v\<in>t. u (v - a) *\<^sub>R v)"
       
  2187     unfolding scaleR_left.setsum
       
  2188     unfolding t_def and setsum.reindex[OF inj] and o_def
       
  2189     using obt(5)
       
  2190     by (auto simp add: setsum.distrib scaleR_right_distrib)
       
  2191   then have "(\<Sum>v\<in>insert a t. (if v = a then - (\<Sum>x\<in>t. u (x - a)) else u (v - a)) *\<^sub>R v) = 0"
       
  2192     unfolding setsum_clauses(2)[OF fin]
       
  2193     using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
       
  2194     by (auto simp add: *)
       
  2195   ultimately show ?thesis
       
  2196     unfolding affine_dependent_explicit
       
  2197     apply (rule_tac x="insert a t" in exI)
       
  2198     apply auto
       
  2199     done
       
  2200 qed
       
  2201 
       
  2202 lemma convex_cone:
       
  2203   "convex s \<and> cone s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. (x + y) \<in> s) \<and> (\<forall>x\<in>s. \<forall>c\<ge>0. (c *\<^sub>R x) \<in> s)"
       
  2204   (is "?lhs = ?rhs")
       
  2205 proof -
       
  2206   {
       
  2207     fix x y
       
  2208     assume "x\<in>s" "y\<in>s" and ?lhs
       
  2209     then have "2 *\<^sub>R x \<in>s" "2 *\<^sub>R y \<in> s"
       
  2210       unfolding cone_def by auto
       
  2211     then have "x + y \<in> s"
       
  2212       using \<open>?lhs\<close>[unfolded convex_def, THEN conjunct1]
       
  2213       apply (erule_tac x="2*\<^sub>R x" in ballE)
       
  2214       apply (erule_tac x="2*\<^sub>R y" in ballE)
       
  2215       apply (erule_tac x="1/2" in allE)
       
  2216       apply simp
       
  2217       apply (erule_tac x="1/2" in allE)
       
  2218       apply auto
       
  2219       done
       
  2220   }
       
  2221   then show ?thesis
       
  2222     unfolding convex_def cone_def by blast
       
  2223 qed
       
  2224 
       
  2225 lemma affine_dependent_biggerset:
       
  2226   fixes s :: "'a::euclidean_space set"
       
  2227   assumes "finite s" "card s \<ge> DIM('a) + 2"
       
  2228   shows "affine_dependent s"
       
  2229 proof -
       
  2230   have "s \<noteq> {}" using assms by auto
       
  2231   then obtain a where "a\<in>s" by auto
       
  2232   have *: "{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})"
       
  2233     by auto
       
  2234   have "card {x - a |x. x \<in> s - {a}} = card (s - {a})"
       
  2235     unfolding *
       
  2236     apply (rule card_image)
       
  2237     unfolding inj_on_def
       
  2238     apply auto
       
  2239     done
       
  2240   also have "\<dots> > DIM('a)" using assms(2)
       
  2241     unfolding card_Diff_singleton[OF assms(1) \<open>a\<in>s\<close>] by auto
       
  2242   finally show ?thesis
       
  2243     apply (subst insert_Diff[OF \<open>a\<in>s\<close>, symmetric])
       
  2244     apply (rule dependent_imp_affine_dependent)
       
  2245     apply (rule dependent_biggerset)
       
  2246     apply auto
       
  2247     done
       
  2248 qed
       
  2249 
       
  2250 lemma affine_dependent_biggerset_general:
       
  2251   assumes "finite (s :: 'a::euclidean_space set)"
       
  2252     and "card s \<ge> dim s + 2"
       
  2253   shows "affine_dependent s"
       
  2254 proof -
       
  2255   from assms(2) have "s \<noteq> {}" by auto
       
  2256   then obtain a where "a\<in>s" by auto
       
  2257   have *: "{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})"
       
  2258     by auto
       
  2259   have **: "card {x - a |x. x \<in> s - {a}} = card (s - {a})"
       
  2260     unfolding *
       
  2261     apply (rule card_image)
       
  2262     unfolding inj_on_def
       
  2263     apply auto
       
  2264     done
       
  2265   have "dim {x - a |x. x \<in> s - {a}} \<le> dim s"
       
  2266     apply (rule subset_le_dim)
       
  2267     unfolding subset_eq
       
  2268     using \<open>a\<in>s\<close>
       
  2269     apply (auto simp add:span_superset span_sub)
       
  2270     done
       
  2271   also have "\<dots> < dim s + 1" by auto
       
  2272   also have "\<dots> \<le> card (s - {a})"
       
  2273     using assms
       
  2274     using card_Diff_singleton[OF assms(1) \<open>a\<in>s\<close>]
       
  2275     by auto
       
  2276   finally show ?thesis
       
  2277     apply (subst insert_Diff[OF \<open>a\<in>s\<close>, symmetric])
       
  2278     apply (rule dependent_imp_affine_dependent)
       
  2279     apply (rule dependent_biggerset_general)
       
  2280     unfolding **
       
  2281     apply auto
       
  2282     done
       
  2283 qed
       
  2284 
       
  2285 
       
  2286 subsection \<open>Some Properties of Affine Dependent Sets\<close>
       
  2287 
       
  2288 lemma affine_independent_0: "\<not> affine_dependent {}"
       
  2289   by (simp add: affine_dependent_def)
       
  2290 
       
  2291 lemma affine_independent_1: "\<not> affine_dependent {a}"
       
  2292   by (simp add: affine_dependent_def)
       
  2293 
       
  2294 lemma affine_independent_2: "\<not> affine_dependent {a,b}"
       
  2295   by (simp add: affine_dependent_def insert_Diff_if hull_same)
       
  2296 
       
  2297 lemma affine_hull_translation: "affine hull ((\<lambda>x. a + x) `  S) = (\<lambda>x. a + x) ` (affine hull S)"
       
  2298 proof -
       
  2299   have "affine ((\<lambda>x. a + x) ` (affine hull S))"
       
  2300     using affine_translation affine_affine_hull by blast
       
  2301   moreover have "(\<lambda>x. a + x) `  S \<subseteq> (\<lambda>x. a + x) ` (affine hull S)"
       
  2302     using hull_subset[of S] by auto
       
  2303   ultimately have h1: "affine hull ((\<lambda>x. a + x) `  S) \<subseteq> (\<lambda>x. a + x) ` (affine hull S)"
       
  2304     by (metis hull_minimal)
       
  2305   have "affine((\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) `  S)))"
       
  2306     using affine_translation affine_affine_hull by blast
       
  2307   moreover have "(\<lambda>x. -a + x) ` (\<lambda>x. a + x) `  S \<subseteq> (\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) `  S))"
       
  2308     using hull_subset[of "(\<lambda>x. a + x) `  S"] by auto
       
  2309   moreover have "S = (\<lambda>x. -a + x) ` (\<lambda>x. a + x) `  S"
       
  2310     using translation_assoc[of "-a" a] by auto
       
  2311   ultimately have "(\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) `  S)) >= (affine hull S)"
       
  2312     by (metis hull_minimal)
       
  2313   then have "affine hull ((\<lambda>x. a + x) ` S) >= (\<lambda>x. a + x) ` (affine hull S)"
       
  2314     by auto
       
  2315   then show ?thesis using h1 by auto
       
  2316 qed
       
  2317 
       
  2318 lemma affine_dependent_translation:
       
  2319   assumes "affine_dependent S"
       
  2320   shows "affine_dependent ((\<lambda>x. a + x) ` S)"
       
  2321 proof -
       
  2322   obtain x where x: "x \<in> S \<and> x \<in> affine hull (S - {x})"
       
  2323     using assms affine_dependent_def by auto
       
  2324   have "op + a ` (S - {x}) = op + a ` S - {a + x}"
       
  2325     by auto
       
  2326   then have "a + x \<in> affine hull ((\<lambda>x. a + x) ` S - {a + x})"
       
  2327     using affine_hull_translation[of a "S - {x}"] x by auto
       
  2328   moreover have "a + x \<in> (\<lambda>x. a + x) ` S"
       
  2329     using x by auto
       
  2330   ultimately show ?thesis
       
  2331     unfolding affine_dependent_def by auto
       
  2332 qed
       
  2333 
       
  2334 lemma affine_dependent_translation_eq:
       
  2335   "affine_dependent S \<longleftrightarrow> affine_dependent ((\<lambda>x. a + x) ` S)"
       
  2336 proof -
       
  2337   {
       
  2338     assume "affine_dependent ((\<lambda>x. a + x) ` S)"
       
  2339     then have "affine_dependent S"
       
  2340       using affine_dependent_translation[of "((\<lambda>x. a + x) ` S)" "-a"] translation_assoc[of "-a" a]
       
  2341       by auto
       
  2342   }
       
  2343   then show ?thesis
       
  2344     using affine_dependent_translation by auto
       
  2345 qed
       
  2346 
       
  2347 lemma affine_hull_0_dependent:
       
  2348   assumes "0 \<in> affine hull S"
       
  2349   shows "dependent S"
       
  2350 proof -
       
  2351   obtain s u where s_u: "finite s \<and> s \<noteq> {} \<and> s \<subseteq> S \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
       
  2352     using assms affine_hull_explicit[of S] by auto
       
  2353   then have "\<exists>v\<in>s. u v \<noteq> 0"
       
  2354     using setsum_not_0[of "u" "s"] by auto
       
  2355   then have "finite s \<and> s \<subseteq> S \<and> (\<exists>v\<in>s. u v \<noteq> 0 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0)"
       
  2356     using s_u by auto
       
  2357   then show ?thesis
       
  2358     unfolding dependent_explicit[of S] by auto
       
  2359 qed
       
  2360 
       
  2361 lemma affine_dependent_imp_dependent2:
       
  2362   assumes "affine_dependent (insert 0 S)"
       
  2363   shows "dependent S"
       
  2364 proof -
       
  2365   obtain x where x: "x \<in> insert 0 S \<and> x \<in> affine hull (insert 0 S - {x})"
       
  2366     using affine_dependent_def[of "(insert 0 S)"] assms by blast
       
  2367   then have "x \<in> span (insert 0 S - {x})"
       
  2368     using affine_hull_subset_span by auto
       
  2369   moreover have "span (insert 0 S - {x}) = span (S - {x})"
       
  2370     using insert_Diff_if[of "0" S "{x}"] span_insert_0[of "S-{x}"] by auto
       
  2371   ultimately have "x \<in> span (S - {x})" by auto
       
  2372   then have "x \<noteq> 0 \<Longrightarrow> dependent S"
       
  2373     using x dependent_def by auto
       
  2374   moreover
       
  2375   {
       
  2376     assume "x = 0"
       
  2377     then have "0 \<in> affine hull S"
       
  2378       using x hull_mono[of "S - {0}" S] by auto
       
  2379     then have "dependent S"
       
  2380       using affine_hull_0_dependent by auto
       
  2381   }
       
  2382   ultimately show ?thesis by auto
       
  2383 qed
       
  2384 
       
  2385 lemma affine_dependent_iff_dependent:
       
  2386   assumes "a \<notin> S"
       
  2387   shows "affine_dependent (insert a S) \<longleftrightarrow> dependent ((\<lambda>x. -a + x) ` S)"
       
  2388 proof -
       
  2389   have "(op + (- a) ` S) = {x - a| x . x : S}" by auto
       
  2390   then show ?thesis
       
  2391     using affine_dependent_translation_eq[of "(insert a S)" "-a"]
       
  2392       affine_dependent_imp_dependent2 assms
       
  2393       dependent_imp_affine_dependent[of a S]
       
  2394     by (auto simp del: uminus_add_conv_diff)
       
  2395 qed
       
  2396 
       
  2397 lemma affine_dependent_iff_dependent2:
       
  2398   assumes "a \<in> S"
       
  2399   shows "affine_dependent S \<longleftrightarrow> dependent ((\<lambda>x. -a + x) ` (S-{a}))"
       
  2400 proof -
       
  2401   have "insert a (S - {a}) = S"
       
  2402     using assms by auto
       
  2403   then show ?thesis
       
  2404     using assms affine_dependent_iff_dependent[of a "S-{a}"] by auto
       
  2405 qed
       
  2406 
       
  2407 lemma affine_hull_insert_span_gen:
       
  2408   "affine hull (insert a s) = (\<lambda>x. a + x) ` span ((\<lambda>x. - a + x) ` s)"
       
  2409 proof -
       
  2410   have h1: "{x - a |x. x \<in> s} = ((\<lambda>x. -a+x) ` s)"
       
  2411     by auto
       
  2412   {
       
  2413     assume "a \<notin> s"
       
  2414     then have ?thesis
       
  2415       using affine_hull_insert_span[of a s] h1 by auto
       
  2416   }
       
  2417   moreover
       
  2418   {
       
  2419     assume a1: "a \<in> s"
       
  2420     have "\<exists>x. x \<in> s \<and> -a+x=0"
       
  2421       apply (rule exI[of _ a])
       
  2422       using a1
       
  2423       apply auto
       
  2424       done
       
  2425     then have "insert 0 ((\<lambda>x. -a+x) ` (s - {a})) = (\<lambda>x. -a+x) ` s"
       
  2426       by auto
       
  2427     then have "span ((\<lambda>x. -a+x) ` (s - {a}))=span ((\<lambda>x. -a+x) ` s)"
       
  2428       using span_insert_0[of "op + (- a) ` (s - {a})"] by (auto simp del: uminus_add_conv_diff)
       
  2429     moreover have "{x - a |x. x \<in> (s - {a})} = ((\<lambda>x. -a+x) ` (s - {a}))"
       
  2430       by auto
       
  2431     moreover have "insert a (s - {a}) = insert a s"
       
  2432       by auto
       
  2433     ultimately have ?thesis
       
  2434       using affine_hull_insert_span[of "a" "s-{a}"] by auto
       
  2435   }
       
  2436   ultimately show ?thesis by auto
       
  2437 qed
       
  2438 
       
  2439 lemma affine_hull_span2:
       
  2440   assumes "a \<in> s"
       
  2441   shows "affine hull s = (\<lambda>x. a+x) ` span ((\<lambda>x. -a+x) ` (s-{a}))"
       
  2442   using affine_hull_insert_span_gen[of a "s - {a}", unfolded insert_Diff[OF assms]]
       
  2443   by auto
       
  2444 
       
  2445 lemma affine_hull_span_gen:
       
  2446   assumes "a \<in> affine hull s"
       
  2447   shows "affine hull s = (\<lambda>x. a+x) ` span ((\<lambda>x. -a+x) ` s)"
       
  2448 proof -
       
  2449   have "affine hull (insert a s) = affine hull s"
       
  2450     using hull_redundant[of a affine s] assms by auto
       
  2451   then show ?thesis
       
  2452     using affine_hull_insert_span_gen[of a "s"] by auto
       
  2453 qed
       
  2454 
       
  2455 lemma affine_hull_span_0:
       
  2456   assumes "0 \<in> affine hull S"
       
  2457   shows "affine hull S = span S"
       
  2458   using affine_hull_span_gen[of "0" S] assms by auto
       
  2459 
       
  2460 lemma extend_to_affine_basis_nonempty:
       
  2461   fixes S V :: "'n::euclidean_space set"
       
  2462   assumes "\<not> affine_dependent S" "S \<subseteq> V" "S \<noteq> {}"
       
  2463   shows "\<exists>T. \<not> affine_dependent T \<and> S \<subseteq> T \<and> T \<subseteq> V \<and> affine hull T = affine hull V"
       
  2464 proof -
       
  2465   obtain a where a: "a \<in> S"
       
  2466     using assms by auto
       
  2467   then have h0: "independent  ((\<lambda>x. -a + x) ` (S-{a}))"
       
  2468     using affine_dependent_iff_dependent2 assms by auto
       
  2469   then obtain B where B:
       
  2470     "(\<lambda>x. -a+x) ` (S - {a}) \<subseteq> B \<and> B \<subseteq> (\<lambda>x. -a+x) ` V \<and> independent B \<and> (\<lambda>x. -a+x) ` V \<subseteq> span B"
       
  2471      using maximal_independent_subset_extend[of "(\<lambda>x. -a+x) ` (S-{a})" "(\<lambda>x. -a + x) ` V"] assms
       
  2472      by blast
       
  2473   define T where "T = (\<lambda>x. a+x) ` insert 0 B"
       
  2474   then have "T = insert a ((\<lambda>x. a+x) ` B)"
       
  2475     by auto
       
  2476   then have "affine hull T = (\<lambda>x. a+x) ` span B"
       
  2477     using affine_hull_insert_span_gen[of a "((\<lambda>x. a+x) ` B)"] translation_assoc[of "-a" a B]
       
  2478     by auto
       
  2479   then have "V \<subseteq> affine hull T"
       
  2480     using B assms translation_inverse_subset[of a V "span B"]
       
  2481     by auto
       
  2482   moreover have "T \<subseteq> V"
       
  2483     using T_def B a assms by auto
       
  2484   ultimately have "affine hull T = affine hull V"
       
  2485     by (metis Int_absorb1 Int_absorb2 hull_hull hull_mono)
       
  2486   moreover have "S \<subseteq> T"
       
  2487     using T_def B translation_inverse_subset[of a "S-{a}" B]
       
  2488     by auto
       
  2489   moreover have "\<not> affine_dependent T"
       
  2490     using T_def affine_dependent_translation_eq[of "insert 0 B"]
       
  2491       affine_dependent_imp_dependent2 B
       
  2492     by auto
       
  2493   ultimately show ?thesis using \<open>T \<subseteq> V\<close> by auto
       
  2494 qed
       
  2495 
       
  2496 lemma affine_basis_exists:
       
  2497   fixes V :: "'n::euclidean_space set"
       
  2498   shows "\<exists>B. B \<subseteq> V \<and> \<not> affine_dependent B \<and> affine hull V = affine hull B"
       
  2499 proof (cases "V = {}")
       
  2500   case True
       
  2501   then show ?thesis
       
  2502     using affine_independent_0 by auto
       
  2503 next
       
  2504   case False
       
  2505   then obtain x where "x \<in> V" by auto
       
  2506   then show ?thesis
       
  2507     using affine_dependent_def[of "{x}"] extend_to_affine_basis_nonempty[of "{x}" V]
       
  2508     by auto
       
  2509 qed
       
  2510 
       
  2511 proposition extend_to_affine_basis:
       
  2512   fixes S V :: "'n::euclidean_space set"
       
  2513   assumes "\<not> affine_dependent S" "S \<subseteq> V"
       
  2514   obtains T where "\<not> affine_dependent T" "S \<subseteq> T" "T \<subseteq> V" "affine hull T = affine hull V"
       
  2515 proof (cases "S = {}")
       
  2516   case True then show ?thesis
       
  2517     using affine_basis_exists by (metis empty_subsetI that)
       
  2518 next
       
  2519   case False
       
  2520   then show ?thesis by (metis assms extend_to_affine_basis_nonempty that)
       
  2521 qed
       
  2522 
       
  2523 
       
  2524 subsection \<open>Affine Dimension of a Set\<close>
       
  2525 
       
  2526 definition aff_dim :: "('a::euclidean_space) set \<Rightarrow> int"
       
  2527   where "aff_dim V =
       
  2528   (SOME d :: int.
       
  2529     \<exists>B. affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> of_nat (card B) = d + 1)"
       
  2530 
       
  2531 lemma aff_dim_basis_exists:
       
  2532   fixes V :: "('n::euclidean_space) set"
       
  2533   shows "\<exists>B. affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> of_nat (card B) = aff_dim V + 1"
       
  2534 proof -
       
  2535   obtain B where "\<not> affine_dependent B \<and> affine hull B = affine hull V"
       
  2536     using affine_basis_exists[of V] by auto
       
  2537   then show ?thesis
       
  2538     unfolding aff_dim_def
       
  2539       some_eq_ex[of "\<lambda>d. \<exists>B. affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> of_nat (card B) = d + 1"]
       
  2540     apply auto
       
  2541     apply (rule exI[of _ "int (card B) - (1 :: int)"])
       
  2542     apply (rule exI[of _ "B"])
       
  2543     apply auto
       
  2544     done
       
  2545 qed
       
  2546 
       
  2547 lemma affine_hull_nonempty: "S \<noteq> {} \<longleftrightarrow> affine hull S \<noteq> {}"
       
  2548 proof -
       
  2549   have "S = {} \<Longrightarrow> affine hull S = {}"
       
  2550     using affine_hull_empty by auto
       
  2551   moreover have "affine hull S = {} \<Longrightarrow> S = {}"
       
  2552     unfolding hull_def by auto
       
  2553   ultimately show ?thesis by blast
       
  2554 qed
       
  2555 
       
  2556 lemma aff_dim_parallel_subspace_aux:
       
  2557   fixes B :: "'n::euclidean_space set"
       
  2558   assumes "\<not> affine_dependent B" "a \<in> B"
       
  2559   shows "finite B \<and> ((card B) - 1 = dim (span ((\<lambda>x. -a+x) ` (B-{a}))))"
       
  2560 proof -
       
  2561   have "independent ((\<lambda>x. -a + x) ` (B-{a}))"
       
  2562     using affine_dependent_iff_dependent2 assms by auto
       
  2563   then have fin: "dim (span ((\<lambda>x. -a+x) ` (B-{a}))) = card ((\<lambda>x. -a + x) ` (B-{a}))"
       
  2564     "finite ((\<lambda>x. -a + x) ` (B - {a}))"
       
  2565     using indep_card_eq_dim_span[of "(\<lambda>x. -a+x) ` (B-{a})"] by auto
       
  2566   show ?thesis
       
  2567   proof (cases "(\<lambda>x. -a + x) ` (B - {a}) = {}")
       
  2568     case True
       
  2569     have "B = insert a ((\<lambda>x. a + x) ` (\<lambda>x. -a + x) ` (B - {a}))"
       
  2570       using translation_assoc[of "a" "-a" "(B - {a})"] assms by auto
       
  2571     then have "B = {a}" using True by auto
       
  2572     then show ?thesis using assms fin by auto
       
  2573   next
       
  2574     case False
       
  2575     then have "card ((\<lambda>x. -a + x) ` (B - {a})) > 0"
       
  2576       using fin by auto
       
  2577     moreover have h1: "card ((\<lambda>x. -a + x) ` (B-{a})) = card (B-{a})"
       
  2578        apply (rule card_image)
       
  2579        using translate_inj_on
       
  2580        apply (auto simp del: uminus_add_conv_diff)
       
  2581        done
       
  2582     ultimately have "card (B-{a}) > 0" by auto
       
  2583     then have *: "finite (B - {a})"
       
  2584       using card_gt_0_iff[of "(B - {a})"] by auto
       
  2585     then have "card (B - {a}) = card B - 1"
       
  2586       using card_Diff_singleton assms by auto
       
  2587     with * show ?thesis using fin h1 by auto
       
  2588   qed
       
  2589 qed
       
  2590 
       
  2591 lemma aff_dim_parallel_subspace:
       
  2592   fixes V L :: "'n::euclidean_space set"
       
  2593   assumes "V \<noteq> {}"
       
  2594     and "subspace L"
       
  2595     and "affine_parallel (affine hull V) L"
       
  2596   shows "aff_dim V = int (dim L)"
       
  2597 proof -
       
  2598   obtain B where
       
  2599     B: "affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> int (card B) = aff_dim V + 1"
       
  2600     using aff_dim_basis_exists by auto
       
  2601   then have "B \<noteq> {}"
       
  2602     using assms B affine_hull_nonempty[of V] affine_hull_nonempty[of B]
       
  2603     by auto
       
  2604   then obtain a where a: "a \<in> B" by auto
       
  2605   define Lb where "Lb = span ((\<lambda>x. -a+x) ` (B-{a}))"
       
  2606   moreover have "affine_parallel (affine hull B) Lb"
       
  2607     using Lb_def B assms affine_hull_span2[of a B] a
       
  2608       affine_parallel_commut[of "Lb" "(affine hull B)"]
       
  2609     unfolding affine_parallel_def
       
  2610     by auto
       
  2611   moreover have "subspace Lb"
       
  2612     using Lb_def subspace_span by auto
       
  2613   moreover have "affine hull B \<noteq> {}"
       
  2614     using assms B affine_hull_nonempty[of V] by auto
       
  2615   ultimately have "L = Lb"
       
  2616     using assms affine_parallel_subspace[of "affine hull B"] affine_affine_hull[of B] B
       
  2617     by auto
       
  2618   then have "dim L = dim Lb"
       
  2619     by auto
       
  2620   moreover have "card B - 1 = dim Lb" and "finite B"
       
  2621     using Lb_def aff_dim_parallel_subspace_aux a B by auto
       
  2622   ultimately show ?thesis
       
  2623     using B \<open>B \<noteq> {}\<close> card_gt_0_iff[of B] by auto
       
  2624 qed
       
  2625 
       
  2626 lemma aff_independent_finite:
       
  2627   fixes B :: "'n::euclidean_space set"
       
  2628   assumes "\<not> affine_dependent B"
       
  2629   shows "finite B"
       
  2630 proof -
       
  2631   {
       
  2632     assume "B \<noteq> {}"
       
  2633     then obtain a where "a \<in> B" by auto
       
  2634     then have ?thesis
       
  2635       using aff_dim_parallel_subspace_aux assms by auto
       
  2636   }
       
  2637   then show ?thesis by auto
       
  2638 qed
       
  2639 
       
  2640 lemma independent_finite:
       
  2641   fixes B :: "'n::euclidean_space set"
       
  2642   assumes "independent B"
       
  2643   shows "finite B"
       
  2644   using affine_dependent_imp_dependent[of B] aff_independent_finite[of B] assms
       
  2645   by auto
       
  2646 
       
  2647 lemma subspace_dim_equal:
       
  2648   assumes "subspace (S :: ('n::euclidean_space) set)"
       
  2649     and "subspace T"
       
  2650     and "S \<subseteq> T"
       
  2651     and "dim S \<ge> dim T"
       
  2652   shows "S = T"
       
  2653 proof -
       
  2654   obtain B where B: "B \<le> S" "independent B \<and> S \<subseteq> span B" "card B = dim S"
       
  2655     using basis_exists[of S] by auto
       
  2656   then have "span B \<subseteq> S"
       
  2657     using span_mono[of B S] span_eq[of S] assms by metis
       
  2658   then have "span B = S"
       
  2659     using B by auto
       
  2660   have "dim S = dim T"
       
  2661     using assms dim_subset[of S T] by auto
       
  2662   then have "T \<subseteq> span B"
       
  2663     using card_eq_dim[of B T] B independent_finite assms by auto
       
  2664   then show ?thesis
       
  2665     using assms \<open>span B = S\<close> by auto
       
  2666 qed
       
  2667 
       
  2668 corollary dim_eq_span:
       
  2669   fixes S :: "'a::euclidean_space set"
       
  2670   shows "\<lbrakk>S \<subseteq> T; dim T \<le> dim S\<rbrakk> \<Longrightarrow> span S = span T"
       
  2671 by (simp add: span_mono subspace_dim_equal subspace_span)
       
  2672 
       
  2673 lemma dim_eq_full:
       
  2674     fixes S :: "'a :: euclidean_space set"
       
  2675     shows "dim S = DIM('a) \<longleftrightarrow> span S = UNIV"
       
  2676 apply (rule iffI)
       
  2677  apply (metis dim_eq_span dim_subset_UNIV span_Basis span_span subset_UNIV)
       
  2678 by (metis dim_UNIV dim_span)
       
  2679 
       
  2680 lemma span_substd_basis:
       
  2681   assumes d: "d \<subseteq> Basis"
       
  2682   shows "span d = {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
       
  2683   (is "_ = ?B")
       
  2684 proof -
       
  2685   have "d \<subseteq> ?B"
       
  2686     using d by (auto simp: inner_Basis)
       
  2687   moreover have s: "subspace ?B"
       
  2688     using subspace_substandard[of "\<lambda>i. i \<notin> d"] .
       
  2689   ultimately have "span d \<subseteq> ?B"
       
  2690     using span_mono[of d "?B"] span_eq[of "?B"] by blast
       
  2691   moreover have *: "card d \<le> dim (span d)"
       
  2692     using independent_card_le_dim[of d "span d"] independent_substdbasis[OF assms] span_inc[of d]
       
  2693     by auto
       
  2694   moreover from * have "dim ?B \<le> dim (span d)"
       
  2695     using dim_substandard[OF assms] by auto
       
  2696   ultimately show ?thesis
       
  2697     using s subspace_dim_equal[of "span d" "?B"] subspace_span[of d] by auto
       
  2698 qed
       
  2699 
       
  2700 lemma basis_to_substdbasis_subspace_isomorphism:
       
  2701   fixes B :: "'a::euclidean_space set"
       
  2702   assumes "independent B"
       
  2703   shows "\<exists>f d::'a set. card d = card B \<and> linear f \<and> f ` B = d \<and>
       
  2704     f ` span B = {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0} \<and> inj_on f (span B) \<and> d \<subseteq> Basis"
       
  2705 proof -
       
  2706   have B: "card B = dim B"
       
  2707     using dim_unique[of B B "card B"] assms span_inc[of B] by auto
       
  2708   have "dim B \<le> card (Basis :: 'a set)"
       
  2709     using dim_subset_UNIV[of B] by simp
       
  2710   from ex_card[OF this] obtain d :: "'a set" where d: "d \<subseteq> Basis" and t: "card d = dim B"
       
  2711     by auto
       
  2712   let ?t = "{x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
       
  2713   have "\<exists>f. linear f \<and> f ` B = d \<and> f ` span B = ?t \<and> inj_on f (span B)"
       
  2714     apply (rule basis_to_basis_subspace_isomorphism[of "span B" ?t B "d"])
       
  2715     apply (rule subspace_span)
       
  2716     apply (rule subspace_substandard)
       
  2717     defer
       
  2718     apply (rule span_inc)
       
  2719     apply (rule assms)
       
  2720     defer
       
  2721     unfolding dim_span[of B]
       
  2722     apply(rule B)
       
  2723     unfolding span_substd_basis[OF d, symmetric]
       
  2724     apply (rule span_inc)
       
  2725     apply (rule independent_substdbasis[OF d])
       
  2726     apply rule
       
  2727     apply assumption
       
  2728     unfolding t[symmetric] span_substd_basis[OF d] dim_substandard[OF d]
       
  2729     apply auto
       
  2730     done
       
  2731   with t \<open>card B = dim B\<close> d show ?thesis by auto
       
  2732 qed
       
  2733 
       
  2734 lemma aff_dim_empty:
       
  2735   fixes S :: "'n::euclidean_space set"
       
  2736   shows "S = {} \<longleftrightarrow> aff_dim S = -1"
       
  2737 proof -
       
  2738   obtain B where *: "affine hull B = affine hull S"
       
  2739     and "\<not> affine_dependent B"
       
  2740     and "int (card B) = aff_dim S + 1"
       
  2741     using aff_dim_basis_exists by auto
       
  2742   moreover
       
  2743   from * have "S = {} \<longleftrightarrow> B = {}"
       
  2744     using affine_hull_nonempty[of B] affine_hull_nonempty[of S] by auto
       
  2745   ultimately show ?thesis
       
  2746     using aff_independent_finite[of B] card_gt_0_iff[of B] by auto
       
  2747 qed
       
  2748 
       
  2749 lemma aff_dim_empty_eq [simp]: "aff_dim ({}::'a::euclidean_space set) = -1"
       
  2750   by (simp add: aff_dim_empty [symmetric])
       
  2751 
       
  2752 lemma aff_dim_affine_hull [simp]: "aff_dim (affine hull S) = aff_dim S"
       
  2753   unfolding aff_dim_def using hull_hull[of _ S] by auto
       
  2754 
       
  2755 lemma aff_dim_affine_hull2:
       
  2756   assumes "affine hull S = affine hull T"
       
  2757   shows "aff_dim S = aff_dim T"
       
  2758   unfolding aff_dim_def using assms by auto
       
  2759 
       
  2760 lemma aff_dim_unique:
       
  2761   fixes B V :: "'n::euclidean_space set"
       
  2762   assumes "affine hull B = affine hull V \<and> \<not> affine_dependent B"
       
  2763   shows "of_nat (card B) = aff_dim V + 1"
       
  2764 proof (cases "B = {}")
       
  2765   case True
       
  2766   then have "V = {}"
       
  2767     using affine_hull_nonempty[of V] affine_hull_nonempty[of B] assms
       
  2768     by auto
       
  2769   then have "aff_dim V = (-1::int)"
       
  2770     using aff_dim_empty by auto
       
  2771   then show ?thesis
       
  2772     using \<open>B = {}\<close> by auto
       
  2773 next
       
  2774   case False
       
  2775   then obtain a where a: "a \<in> B" by auto
       
  2776   define Lb where "Lb = span ((\<lambda>x. -a+x) ` (B-{a}))"
       
  2777   have "affine_parallel (affine hull B) Lb"
       
  2778     using Lb_def affine_hull_span2[of a B] a
       
  2779       affine_parallel_commut[of "Lb" "(affine hull B)"]
       
  2780     unfolding affine_parallel_def by auto
       
  2781   moreover have "subspace Lb"
       
  2782     using Lb_def subspace_span by auto
       
  2783   ultimately have "aff_dim B = int(dim Lb)"
       
  2784     using aff_dim_parallel_subspace[of B Lb] \<open>B \<noteq> {}\<close> by auto
       
  2785   moreover have "(card B) - 1 = dim Lb" "finite B"
       
  2786     using Lb_def aff_dim_parallel_subspace_aux a assms by auto
       
  2787   ultimately have "of_nat (card B) = aff_dim B + 1"
       
  2788     using \<open>B \<noteq> {}\<close> card_gt_0_iff[of B] by auto
       
  2789   then show ?thesis
       
  2790     using aff_dim_affine_hull2 assms by auto
       
  2791 qed
       
  2792 
       
  2793 lemma aff_dim_affine_independent:
       
  2794   fixes B :: "'n::euclidean_space set"
       
  2795   assumes "\<not> affine_dependent B"
       
  2796   shows "of_nat (card B) = aff_dim B + 1"
       
  2797   using aff_dim_unique[of B B] assms by auto
       
  2798 
       
  2799 lemma affine_independent_iff_card:
       
  2800     fixes s :: "'a::euclidean_space set"
       
  2801     shows "~ affine_dependent s \<longleftrightarrow> finite s \<and> aff_dim s = int(card s) - 1"
       
  2802   apply (rule iffI)
       
  2803   apply (simp add: aff_dim_affine_independent aff_independent_finite)
       
  2804   by (metis affine_basis_exists [of s] aff_dim_unique card_subset_eq diff_add_cancel of_nat_eq_iff)
       
  2805 
       
  2806 lemma aff_dim_sing [simp]:
       
  2807   fixes a :: "'n::euclidean_space"
       
  2808   shows "aff_dim {a} = 0"
       
  2809   using aff_dim_affine_independent[of "{a}"] affine_independent_1 by auto
       
  2810 
       
  2811 lemma aff_dim_inner_basis_exists:
       
  2812   fixes V :: "('n::euclidean_space) set"
       
  2813   shows "\<exists>B. B \<subseteq> V \<and> affine hull B = affine hull V \<and>
       
  2814     \<not> affine_dependent B \<and> of_nat (card B) = aff_dim V + 1"
       
  2815 proof -
       
  2816   obtain B where B: "\<not> affine_dependent B" "B \<subseteq> V" "affine hull B = affine hull V"
       
  2817     using affine_basis_exists[of V] by auto
       
  2818   then have "of_nat(card B) = aff_dim V+1" using aff_dim_unique by auto
       
  2819   with B show ?thesis by auto
       
  2820 qed
       
  2821 
       
  2822 lemma aff_dim_le_card:
       
  2823   fixes V :: "'n::euclidean_space set"
       
  2824   assumes "finite V"
       
  2825   shows "aff_dim V \<le> of_nat (card V) - 1"
       
  2826 proof -
       
  2827   obtain B where B: "B \<subseteq> V" "of_nat (card B) = aff_dim V + 1"
       
  2828     using aff_dim_inner_basis_exists[of V] by auto
       
  2829   then have "card B \<le> card V"
       
  2830     using assms card_mono by auto
       
  2831   with B show ?thesis by auto
       
  2832 qed
       
  2833 
       
  2834 lemma aff_dim_parallel_eq:
       
  2835   fixes S T :: "'n::euclidean_space set"
       
  2836   assumes "affine_parallel (affine hull S) (affine hull T)"
       
  2837   shows "aff_dim S = aff_dim T"
       
  2838 proof -
       
  2839   {
       
  2840     assume "T \<noteq> {}" "S \<noteq> {}"
       
  2841     then obtain L where L: "subspace L \<and> affine_parallel (affine hull T) L"
       
  2842       using affine_parallel_subspace[of "affine hull T"]
       
  2843         affine_affine_hull[of T] affine_hull_nonempty
       
  2844       by auto
       
  2845     then have "aff_dim T = int (dim L)"
       
  2846       using aff_dim_parallel_subspace \<open>T \<noteq> {}\<close> by auto
       
  2847     moreover have *: "subspace L \<and> affine_parallel (affine hull S) L"
       
  2848        using L affine_parallel_assoc[of "affine hull S" "affine hull T" L] assms by auto
       
  2849     moreover from * have "aff_dim S = int (dim L)"
       
  2850       using aff_dim_parallel_subspace \<open>S \<noteq> {}\<close> by auto
       
  2851     ultimately have ?thesis by auto
       
  2852   }
       
  2853   moreover
       
  2854   {
       
  2855     assume "S = {}"
       
  2856     then have "S = {}" and "T = {}"
       
  2857       using assms affine_hull_nonempty
       
  2858       unfolding affine_parallel_def
       
  2859       by auto
       
  2860     then have ?thesis using aff_dim_empty by auto
       
  2861   }
       
  2862   moreover
       
  2863   {
       
  2864     assume "T = {}"
       
  2865     then have "S = {}" and "T = {}"
       
  2866       using assms affine_hull_nonempty
       
  2867       unfolding affine_parallel_def
       
  2868       by auto
       
  2869     then have ?thesis
       
  2870       using aff_dim_empty by auto
       
  2871   }
       
  2872   ultimately show ?thesis by blast
       
  2873 qed
       
  2874 
       
  2875 lemma aff_dim_translation_eq:
       
  2876   fixes a :: "'n::euclidean_space"
       
  2877   shows "aff_dim ((\<lambda>x. a + x) ` S) = aff_dim S"
       
  2878 proof -
       
  2879   have "affine_parallel (affine hull S) (affine hull ((\<lambda>x. a + x) ` S))"
       
  2880     unfolding affine_parallel_def
       
  2881     apply (rule exI[of _ "a"])
       
  2882     using affine_hull_translation[of a S]
       
  2883     apply auto
       
  2884     done
       
  2885   then show ?thesis
       
  2886     using aff_dim_parallel_eq[of S "(\<lambda>x. a + x) ` S"] by auto
       
  2887 qed
       
  2888 
       
  2889 lemma aff_dim_affine:
       
  2890   fixes S L :: "'n::euclidean_space set"
       
  2891   assumes "S \<noteq> {}"
       
  2892     and "affine S"
       
  2893     and "subspace L"
       
  2894     and "affine_parallel S L"
       
  2895   shows "aff_dim S = int (dim L)"
       
  2896 proof -
       
  2897   have *: "affine hull S = S"
       
  2898     using assms affine_hull_eq[of S] by auto
       
  2899   then have "affine_parallel (affine hull S) L"
       
  2900     using assms by (simp add: *)
       
  2901   then show ?thesis
       
  2902     using assms aff_dim_parallel_subspace[of S L] by blast
       
  2903 qed
       
  2904 
       
  2905 lemma dim_affine_hull:
       
  2906   fixes S :: "'n::euclidean_space set"
       
  2907   shows "dim (affine hull S) = dim S"
       
  2908 proof -
       
  2909   have "dim (affine hull S) \<ge> dim S"
       
  2910     using dim_subset by auto
       
  2911   moreover have "dim (span S) \<ge> dim (affine hull S)"
       
  2912     using dim_subset affine_hull_subset_span by blast
       
  2913   moreover have "dim (span S) = dim S"
       
  2914     using dim_span by auto
       
  2915   ultimately show ?thesis by auto
       
  2916 qed
       
  2917 
       
  2918 lemma aff_dim_subspace:
       
  2919   fixes S :: "'n::euclidean_space set"
       
  2920   assumes "subspace S"
       
  2921   shows "aff_dim S = int (dim S)"
       
  2922 proof (cases "S={}")
       
  2923   case True with assms show ?thesis
       
  2924     by (simp add: subspace_affine)
       
  2925 next
       
  2926   case False
       
  2927   with aff_dim_affine[of S S] assms subspace_imp_affine[of S] affine_parallel_reflex[of S] subspace_affine
       
  2928   show ?thesis by auto
       
  2929 qed
       
  2930 
       
  2931 lemma aff_dim_zero:
       
  2932   fixes S :: "'n::euclidean_space set"
       
  2933   assumes "0 \<in> affine hull S"
       
  2934   shows "aff_dim S = int (dim S)"
       
  2935 proof -
       
  2936   have "subspace (affine hull S)"
       
  2937     using subspace_affine[of "affine hull S"] affine_affine_hull assms
       
  2938     by auto
       
  2939   then have "aff_dim (affine hull S) = int (dim (affine hull S))"
       
  2940     using assms aff_dim_subspace[of "affine hull S"] by auto
       
  2941   then show ?thesis
       
  2942     using aff_dim_affine_hull[of S] dim_affine_hull[of S]
       
  2943     by auto
       
  2944 qed
       
  2945 
       
  2946 lemma aff_dim_eq_dim:
       
  2947   fixes S :: "'n::euclidean_space set"
       
  2948   assumes "a \<in> affine hull S"
       
  2949   shows "aff_dim S = int (dim ((\<lambda>x. -a+x) ` S))"
       
  2950 proof -
       
  2951   have "0 \<in> affine hull ((\<lambda>x. -a+x) ` S)"
       
  2952     unfolding Convex_Euclidean_Space.affine_hull_translation
       
  2953     using assms by (simp add: ab_group_add_class.ab_left_minus image_iff)
       
  2954   with aff_dim_zero show ?thesis
       
  2955     by (metis aff_dim_translation_eq)
       
  2956 qed
       
  2957 
       
  2958 lemma aff_dim_UNIV [simp]: "aff_dim (UNIV :: 'n::euclidean_space set) = int(DIM('n))"
       
  2959   using aff_dim_subspace[of "(UNIV :: 'n::euclidean_space set)"]
       
  2960     dim_UNIV[where 'a="'n::euclidean_space"]
       
  2961   by auto
       
  2962 
       
  2963 lemma aff_dim_geq:
       
  2964   fixes V :: "'n::euclidean_space set"
       
  2965   shows "aff_dim V \<ge> -1"
       
  2966 proof -
       
  2967   obtain B where "affine hull B = affine hull V"
       
  2968     and "\<not> affine_dependent B"
       
  2969     and "int (card B) = aff_dim V + 1"
       
  2970     using aff_dim_basis_exists by auto
       
  2971   then show ?thesis by auto
       
  2972 qed
       
  2973 
       
  2974 lemma aff_dim_negative_iff [simp]:
       
  2975   fixes S :: "'n::euclidean_space set"
       
  2976   shows "aff_dim S < 0 \<longleftrightarrow>S = {}"
       
  2977 by (metis aff_dim_empty aff_dim_geq diff_0 eq_iff zle_diff1_eq)
       
  2978 
       
  2979 lemma affine_independent_card_dim_diffs:
       
  2980   fixes S :: "'a :: euclidean_space set"
       
  2981   assumes "~ affine_dependent S" "a \<in> S"
       
  2982     shows "card S = dim {x - a|x. x \<in> S} + 1"
       
  2983 proof -
       
  2984   have 1: "{b - a|b. b \<in> (S - {a})} \<subseteq> {x - a|x. x \<in> S}" by auto
       
  2985   have 2: "x - a \<in> span {b - a |b. b \<in> S - {a}}" if "x \<in> S" for x
       
  2986   proof (cases "x = a")
       
  2987     case True then show ?thesis by simp
       
  2988   next
       
  2989     case False then show ?thesis
       
  2990       using assms by (blast intro: span_superset that)
       
  2991   qed
       
  2992   have "\<not> affine_dependent (insert a S)"
       
  2993     by (simp add: assms insert_absorb)
       
  2994   then have 3: "independent {b - a |b. b \<in> S - {a}}"
       
  2995       using dependent_imp_affine_dependent by fastforce
       
  2996   have "{b - a |b. b \<in> S - {a}} = (\<lambda>b. b-a) ` (S - {a})"
       
  2997     by blast
       
  2998   then have "card {b - a |b. b \<in> S - {a}} = card ((\<lambda>b. b-a) ` (S - {a}))"
       
  2999     by simp
       
  3000   also have "... = card (S - {a})"
       
  3001     by (metis (no_types, lifting) card_image diff_add_cancel inj_onI)
       
  3002   also have "... = card S - 1"
       
  3003     by (simp add: aff_independent_finite assms)
       
  3004   finally have 4: "card {b - a |b. b \<in> S - {a}} = card S - 1" .
       
  3005   have "finite S"
       
  3006     by (meson assms aff_independent_finite)
       
  3007   with \<open>a \<in> S\<close> have "card S \<noteq> 0" by auto
       
  3008   moreover have "dim {x - a |x. x \<in> S} = card S - 1"
       
  3009     using 2 by (blast intro: dim_unique [OF 1 _ 3 4])
       
  3010   ultimately show ?thesis
       
  3011     by auto
       
  3012 qed
       
  3013 
       
  3014 lemma independent_card_le_aff_dim:
       
  3015   fixes B :: "'n::euclidean_space set"
       
  3016   assumes "B \<subseteq> V"
       
  3017   assumes "\<not> affine_dependent B"
       
  3018   shows "int (card B) \<le> aff_dim V + 1"
       
  3019 proof -
       
  3020   obtain T where T: "\<not> affine_dependent T \<and> B \<subseteq> T \<and> T \<subseteq> V \<and> affine hull T = affine hull V"
       
  3021     by (metis assms extend_to_affine_basis[of B V])
       
  3022   then have "of_nat (card T) = aff_dim V + 1"
       
  3023     using aff_dim_unique by auto
       
  3024   then show ?thesis
       
  3025     using T card_mono[of T B] aff_independent_finite[of T] by auto
       
  3026 qed
       
  3027 
       
  3028 lemma aff_dim_subset:
       
  3029   fixes S T :: "'n::euclidean_space set"
       
  3030   assumes "S \<subseteq> T"
       
  3031   shows "aff_dim S \<le> aff_dim T"
       
  3032 proof -
       
  3033   obtain B where B: "\<not> affine_dependent B" "B \<subseteq> S" "affine hull B = affine hull S"
       
  3034     "of_nat (card B) = aff_dim S + 1"
       
  3035     using aff_dim_inner_basis_exists[of S] by auto
       
  3036   then have "int (card B) \<le> aff_dim T + 1"
       
  3037     using assms independent_card_le_aff_dim[of B T] by auto
       
  3038   with B show ?thesis by auto
       
  3039 qed
       
  3040 
       
  3041 lemma aff_dim_le_DIM:
       
  3042   fixes S :: "'n::euclidean_space set"
       
  3043   shows "aff_dim S \<le> int (DIM('n))"
       
  3044 proof -
       
  3045   have "aff_dim (UNIV :: 'n::euclidean_space set) = int(DIM('n))"
       
  3046     using aff_dim_UNIV by auto
       
  3047   then show "aff_dim (S:: 'n::euclidean_space set) \<le> int(DIM('n))"
       
  3048     using aff_dim_subset[of S "(UNIV :: ('n::euclidean_space) set)"] subset_UNIV by auto
       
  3049 qed
       
  3050 
       
  3051 lemma affine_dim_equal:
       
  3052   fixes S :: "'n::euclidean_space set"
       
  3053   assumes "affine S" "affine T" "S \<noteq> {}" "S \<subseteq> T" "aff_dim S = aff_dim T"
       
  3054   shows "S = T"
       
  3055 proof -
       
  3056   obtain a where "a \<in> S" using assms by auto
       
  3057   then have "a \<in> T" using assms by auto
       
  3058   define LS where "LS = {y. \<exists>x \<in> S. (-a) + x = y}"
       
  3059   then have ls: "subspace LS" "affine_parallel S LS"
       
  3060     using assms parallel_subspace_explicit[of S a LS] \<open>a \<in> S\<close> by auto
       
  3061   then have h1: "int(dim LS) = aff_dim S"
       
  3062     using assms aff_dim_affine[of S LS] by auto
       
  3063   have "T \<noteq> {}" using assms by auto
       
  3064   define LT where "LT = {y. \<exists>x \<in> T. (-a) + x = y}"
       
  3065   then have lt: "subspace LT \<and> affine_parallel T LT"
       
  3066     using assms parallel_subspace_explicit[of T a LT] \<open>a \<in> T\<close> by auto
       
  3067   then have "int(dim LT) = aff_dim T"
       
  3068     using assms aff_dim_affine[of T LT] \<open>T \<noteq> {}\<close> by auto
       
  3069   then have "dim LS = dim LT"
       
  3070     using h1 assms by auto
       
  3071   moreover have "LS \<le> LT"
       
  3072     using LS_def LT_def assms by auto
       
  3073   ultimately have "LS = LT"
       
  3074     using subspace_dim_equal[of LS LT] ls lt by auto
       
  3075   moreover have "S = {x. \<exists>y \<in> LS. a+y=x}"
       
  3076     using LS_def by auto
       
  3077   moreover have "T = {x. \<exists>y \<in> LT. a+y=x}"
       
  3078     using LT_def by auto
       
  3079   ultimately show ?thesis by auto
       
  3080 qed
       
  3081 
       
  3082 lemma affine_hull_UNIV:
       
  3083   fixes S :: "'n::euclidean_space set"
       
  3084   assumes "aff_dim S = int(DIM('n))"
       
  3085   shows "affine hull S = (UNIV :: ('n::euclidean_space) set)"
       
  3086 proof -
       
  3087   have "S \<noteq> {}"
       
  3088     using assms aff_dim_empty[of S] by auto
       
  3089   have h0: "S \<subseteq> affine hull S"
       
  3090     using hull_subset[of S _] by auto
       
  3091   have h1: "aff_dim (UNIV :: ('n::euclidean_space) set) = aff_dim S"
       
  3092     using aff_dim_UNIV assms by auto
       
  3093   then have h2: "aff_dim (affine hull S) \<le> aff_dim (UNIV :: ('n::euclidean_space) set)"
       
  3094     using aff_dim_le_DIM[of "affine hull S"] assms h0 by auto
       
  3095   have h3: "aff_dim S \<le> aff_dim (affine hull S)"
       
  3096     using h0 aff_dim_subset[of S "affine hull S"] assms by auto
       
  3097   then have h4: "aff_dim (affine hull S) = aff_dim (UNIV :: ('n::euclidean_space) set)"
       
  3098     using h0 h1 h2 by auto
       
  3099   then show ?thesis
       
  3100     using affine_dim_equal[of "affine hull S" "(UNIV :: ('n::euclidean_space) set)"]
       
  3101       affine_affine_hull[of S] affine_UNIV assms h4 h0 \<open>S \<noteq> {}\<close>
       
  3102     by auto
       
  3103 qed
       
  3104 
       
  3105 lemma disjoint_affine_hull:
       
  3106   fixes s :: "'n::euclidean_space set"
       
  3107   assumes "~ affine_dependent s" "t \<subseteq> s" "u \<subseteq> s" "t \<inter> u = {}"
       
  3108     shows "(affine hull t) \<inter> (affine hull u) = {}"
       
  3109 proof -
       
  3110   have "finite s" using assms by (simp add: aff_independent_finite)
       
  3111   then have "finite t" "finite u" using assms finite_subset by blast+
       
  3112   { fix y
       
  3113     assume yt: "y \<in> affine hull t" and yu: "y \<in> affine hull u"
       
  3114     then obtain a b
       
  3115            where a1 [simp]: "setsum a t = 1" and [simp]: "setsum (\<lambda>v. a v *\<^sub>R v) t = y"
       
  3116              and [simp]: "setsum b u = 1" "setsum (\<lambda>v. b v *\<^sub>R v) u = y"
       
  3117       by (auto simp: affine_hull_finite \<open>finite t\<close> \<open>finite u\<close>)
       
  3118     define c where "c x = (if x \<in> t then a x else if x \<in> u then -(b x) else 0)" for x
       
  3119     have [simp]: "s \<inter> t = t" "s \<inter> - t \<inter> u = u" using assms by auto
       
  3120     have "setsum c s = 0"
       
  3121       by (simp add: c_def comm_monoid_add_class.setsum.If_cases \<open>finite s\<close> setsum_negf)
       
  3122     moreover have "~ (\<forall>v\<in>s. c v = 0)"
       
  3123       by (metis (no_types) IntD1 \<open>s \<inter> t = t\<close> a1 c_def setsum_not_0 zero_neq_one)
       
  3124     moreover have "(\<Sum>v\<in>s. c v *\<^sub>R v) = 0"
       
  3125       by (simp add: c_def if_smult setsum_negf
       
  3126              comm_monoid_add_class.setsum.If_cases \<open>finite s\<close>)
       
  3127     ultimately have False
       
  3128       using assms \<open>finite s\<close> by (auto simp: affine_dependent_explicit)
       
  3129   }
       
  3130   then show ?thesis by blast
       
  3131 qed
       
  3132 
       
  3133 lemma aff_dim_convex_hull:
       
  3134   fixes S :: "'n::euclidean_space set"
       
  3135   shows "aff_dim (convex hull S) = aff_dim S"
       
  3136   using aff_dim_affine_hull[of S] convex_hull_subset_affine_hull[of S]
       
  3137     hull_subset[of S "convex"] aff_dim_subset[of S "convex hull S"]
       
  3138     aff_dim_subset[of "convex hull S" "affine hull S"]
       
  3139   by auto
       
  3140 
       
  3141 lemma aff_dim_cball:
       
  3142   fixes a :: "'n::euclidean_space"
       
  3143   assumes "e > 0"
       
  3144   shows "aff_dim (cball a e) = int (DIM('n))"
       
  3145 proof -
       
  3146   have "(\<lambda>x. a + x) ` (cball 0 e) \<subseteq> cball a e"
       
  3147     unfolding cball_def dist_norm by auto
       
  3148   then have "aff_dim (cball (0 :: 'n::euclidean_space) e) \<le> aff_dim (cball a e)"
       
  3149     using aff_dim_translation_eq[of a "cball 0 e"]
       
  3150           aff_dim_subset[of "op + a ` cball 0 e" "cball a e"]
       
  3151     by auto
       
  3152   moreover have "aff_dim (cball (0 :: 'n::euclidean_space) e) = int (DIM('n))"
       
  3153     using hull_inc[of "(0 :: 'n::euclidean_space)" "cball 0 e"]
       
  3154       centre_in_cball[of "(0 :: 'n::euclidean_space)"] assms
       
  3155     by (simp add: dim_cball[of e] aff_dim_zero[of "cball 0 e"])
       
  3156   ultimately show ?thesis
       
  3157     using aff_dim_le_DIM[of "cball a e"] by auto
       
  3158 qed
       
  3159 
       
  3160 lemma aff_dim_open:
       
  3161   fixes S :: "'n::euclidean_space set"
       
  3162   assumes "open S"
       
  3163     and "S \<noteq> {}"
       
  3164   shows "aff_dim S = int (DIM('n))"
       
  3165 proof -
       
  3166   obtain x where "x \<in> S"
       
  3167     using assms by auto
       
  3168   then obtain e where e: "e > 0" "cball x e \<subseteq> S"
       
  3169     using open_contains_cball[of S] assms by auto
       
  3170   then have "aff_dim (cball x e) \<le> aff_dim S"
       
  3171     using aff_dim_subset by auto
       
  3172   with e show ?thesis
       
  3173     using aff_dim_cball[of e x] aff_dim_le_DIM[of S] by auto
       
  3174 qed
       
  3175 
       
  3176 lemma low_dim_interior:
       
  3177   fixes S :: "'n::euclidean_space set"
       
  3178   assumes "\<not> aff_dim S = int (DIM('n))"
       
  3179   shows "interior S = {}"
       
  3180 proof -
       
  3181   have "aff_dim(interior S) \<le> aff_dim S"
       
  3182     using interior_subset aff_dim_subset[of "interior S" S] by auto
       
  3183   then show ?thesis
       
  3184     using aff_dim_open[of "interior S"] aff_dim_le_DIM[of S] assms by auto
       
  3185 qed
       
  3186 
       
  3187 corollary empty_interior_lowdim:
       
  3188   fixes S :: "'n::euclidean_space set"
       
  3189   shows "dim S < DIM ('n) \<Longrightarrow> interior S = {}"
       
  3190 by (metis low_dim_interior affine_hull_UNIV dim_affine_hull less_not_refl dim_UNIV)
       
  3191 
       
  3192 corollary aff_dim_nonempty_interior:
       
  3193   fixes S :: "'a::euclidean_space set"
       
  3194   shows "interior S \<noteq> {} \<Longrightarrow> aff_dim S = DIM('a)"
       
  3195 by (metis low_dim_interior)
       
  3196 
       
  3197 subsection \<open>Caratheodory's theorem.\<close>
       
  3198 
       
  3199 lemma convex_hull_caratheodory_aff_dim:
       
  3200   fixes p :: "('a::euclidean_space) set"
       
  3201   shows "convex hull p =
       
  3202     {y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> aff_dim p + 1 \<and>
       
  3203       (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = y}"
       
  3204   unfolding convex_hull_explicit set_eq_iff mem_Collect_eq
       
  3205 proof (intro allI iffI)
       
  3206   fix y
       
  3207   let ?P = "\<lambda>n. \<exists>s u. finite s \<and> card s = n \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and>
       
  3208     setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
       
  3209   assume "\<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
       
  3210   then obtain N where "?P N" by auto
       
  3211   then have "\<exists>n\<le>N. (\<forall>k<n. \<not> ?P k) \<and> ?P n"
       
  3212     apply (rule_tac ex_least_nat_le)
       
  3213     apply auto
       
  3214     done
       
  3215   then obtain n where "?P n" and smallest: "\<forall>k<n. \<not> ?P k"
       
  3216     by blast
       
  3217   then obtain s u where obt: "finite s" "card s = n" "s\<subseteq>p" "\<forall>x\<in>s. 0 \<le> u x"
       
  3218     "setsum u s = 1"  "(\<Sum>v\<in>s. u v *\<^sub>R v) = y" by auto
       
  3219 
       
  3220   have "card s \<le> aff_dim p + 1"
       
  3221   proof (rule ccontr, simp only: not_le)
       
  3222     assume "aff_dim p + 1 < card s"
       
  3223     then have "affine_dependent s"
       
  3224       using affine_dependent_biggerset[OF obt(1)] independent_card_le_aff_dim not_less obt(3)
       
  3225       by blast
       
  3226     then obtain w v where wv: "setsum w s = 0" "v\<in>s" "w v \<noteq> 0" "(\<Sum>v\<in>s. w v *\<^sub>R v) = 0"
       
  3227       using affine_dependent_explicit_finite[OF obt(1)] by auto
       
  3228     define i where "i = (\<lambda>v. (u v) / (- w v)) ` {v\<in>s. w v < 0}"
       
  3229     define t where "t = Min i"
       
  3230     have "\<exists>x\<in>s. w x < 0"
       
  3231     proof (rule ccontr, simp add: not_less)
       
  3232       assume as:"\<forall>x\<in>s. 0 \<le> w x"
       
  3233       then have "setsum w (s - {v}) \<ge> 0"
       
  3234         apply (rule_tac setsum_nonneg)
       
  3235         apply auto
       
  3236         done
       
  3237       then have "setsum w s > 0"
       
  3238         unfolding setsum.remove[OF obt(1) \<open>v\<in>s\<close>]
       
  3239         using as[THEN bspec[where x=v]]  \<open>v\<in>s\<close>  \<open>w v \<noteq> 0\<close> by auto
       
  3240       then show False using wv(1) by auto
       
  3241     qed
       
  3242     then have "i \<noteq> {}" unfolding i_def by auto
       
  3243     then have "t \<ge> 0"
       
  3244       using Min_ge_iff[of i 0 ] and obt(1)
       
  3245       unfolding t_def i_def
       
  3246       using obt(4)[unfolded le_less]
       
  3247       by (auto simp: divide_le_0_iff)
       
  3248     have t: "\<forall>v\<in>s. u v + t * w v \<ge> 0"
       
  3249     proof
       
  3250       fix v
       
  3251       assume "v \<in> s"
       
  3252       then have v: "0 \<le> u v"
       
  3253         using obt(4)[THEN bspec[where x=v]] by auto
       
  3254       show "0 \<le> u v + t * w v"
       
  3255       proof (cases "w v < 0")
       
  3256         case False
       
  3257         thus ?thesis using v \<open>t\<ge>0\<close> by auto
       
  3258       next
       
  3259         case True
       
  3260         then have "t \<le> u v / (- w v)"
       
  3261           using \<open>v\<in>s\<close> unfolding t_def i_def
       
  3262           apply (rule_tac Min_le)
       
  3263           using obt(1) apply auto
       
  3264           done
       
  3265         then show ?thesis
       
  3266           unfolding real_0_le_add_iff
       
  3267           using pos_le_divide_eq[OF True[unfolded neg_0_less_iff_less[symmetric]]]
       
  3268           by auto
       
  3269       qed
       
  3270     qed
       
  3271     obtain a where "a \<in> s" and "t = (\<lambda>v. (u v) / (- w v)) a" and "w a < 0"
       
  3272       using Min_in[OF _ \<open>i\<noteq>{}\<close>] and obt(1) unfolding i_def t_def by auto
       
  3273     then have a: "a \<in> s" "u a + t * w a = 0" by auto
       
  3274     have *: "\<And>f. setsum f (s - {a}) = setsum f s - ((f a)::'b::ab_group_add)"
       
  3275       unfolding setsum.remove[OF obt(1) \<open>a\<in>s\<close>] by auto
       
  3276     have "(\<Sum>v\<in>s. u v + t * w v) = 1"
       
  3277       unfolding setsum.distrib wv(1) setsum_right_distrib[symmetric] obt(5) by auto
       
  3278     moreover have "(\<Sum>v\<in>s. u v *\<^sub>R v + (t * w v) *\<^sub>R v) - (u a *\<^sub>R a + (t * w a) *\<^sub>R a) = y"
       
  3279       unfolding setsum.distrib obt(6) scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] wv(4)
       
  3280       using a(2) [THEN eq_neg_iff_add_eq_0 [THEN iffD2]] by simp
       
  3281     ultimately have "?P (n - 1)"
       
  3282       apply (rule_tac x="(s - {a})" in exI)
       
  3283       apply (rule_tac x="\<lambda>v. u v + t * w v" in exI)
       
  3284       using obt(1-3) and t and a
       
  3285       apply (auto simp add: * scaleR_left_distrib)
       
  3286       done
       
  3287     then show False
       
  3288       using smallest[THEN spec[where x="n - 1"]] by auto
       
  3289   qed
       
  3290   then show "\<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> aff_dim p + 1 \<and>
       
  3291       (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
       
  3292     using obt by auto
       
  3293 qed auto
       
  3294 
       
  3295 lemma caratheodory_aff_dim:
       
  3296   fixes p :: "('a::euclidean_space) set"
       
  3297   shows "convex hull p = {x. \<exists>s. finite s \<and> s \<subseteq> p \<and> card s \<le> aff_dim p + 1 \<and> x \<in> convex hull s}"
       
  3298         (is "?lhs = ?rhs")
       
  3299 proof
       
  3300   show "?lhs \<subseteq> ?rhs"
       
  3301     apply (subst convex_hull_caratheodory_aff_dim)
       
  3302     apply clarify
       
  3303     apply (rule_tac x="s" in exI)
       
  3304     apply (simp add: hull_subset convex_explicit [THEN iffD1, OF convex_convex_hull])
       
  3305     done
       
  3306 next
       
  3307   show "?rhs \<subseteq> ?lhs"
       
  3308     using hull_mono by blast
       
  3309 qed
       
  3310 
       
  3311 lemma convex_hull_caratheodory:
       
  3312   fixes p :: "('a::euclidean_space) set"
       
  3313   shows "convex hull p =
       
  3314             {y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> DIM('a) + 1 \<and>
       
  3315               (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = y}"
       
  3316         (is "?lhs = ?rhs")
       
  3317 proof (intro set_eqI iffI)
       
  3318   fix x
       
  3319   assume "x \<in> ?lhs" then show "x \<in> ?rhs"
       
  3320     apply (simp only: convex_hull_caratheodory_aff_dim Set.mem_Collect_eq)
       
  3321     apply (erule ex_forward)+
       
  3322     using aff_dim_le_DIM [of p]
       
  3323     apply simp
       
  3324     done
       
  3325 next
       
  3326   fix x
       
  3327   assume "x \<in> ?rhs" then show "x \<in> ?lhs"
       
  3328     by (auto simp add: convex_hull_explicit)
       
  3329 qed
       
  3330 
       
  3331 theorem caratheodory:
       
  3332   "convex hull p =
       
  3333     {x::'a::euclidean_space. \<exists>s. finite s \<and> s \<subseteq> p \<and>
       
  3334       card s \<le> DIM('a) + 1 \<and> x \<in> convex hull s}"
       
  3335 proof safe
       
  3336   fix x
       
  3337   assume "x \<in> convex hull p"
       
  3338   then obtain s u where "finite s" "s \<subseteq> p" "card s \<le> DIM('a) + 1"
       
  3339     "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
       
  3340     unfolding convex_hull_caratheodory by auto
       
  3341   then show "\<exists>s. finite s \<and> s \<subseteq> p \<and> card s \<le> DIM('a) + 1 \<and> x \<in> convex hull s"
       
  3342     apply (rule_tac x=s in exI)
       
  3343     using hull_subset[of s convex]
       
  3344     using convex_convex_hull[simplified convex_explicit, of s,
       
  3345       THEN spec[where x=s], THEN spec[where x=u]]
       
  3346     apply auto
       
  3347     done
       
  3348 next
       
  3349   fix x s
       
  3350   assume  "finite s" "s \<subseteq> p" "card s \<le> DIM('a) + 1" "x \<in> convex hull s"
       
  3351   then show "x \<in> convex hull p"
       
  3352     using hull_mono[OF \<open>s\<subseteq>p\<close>] by auto
       
  3353 qed
       
  3354 
       
  3355 
       
  3356 subsection \<open>Relative interior of a set\<close>
       
  3357 
       
  3358 definition "rel_interior S =
       
  3359   {x. \<exists>T. openin (subtopology euclidean (affine hull S)) T \<and> x \<in> T \<and> T \<subseteq> S}"
       
  3360 
       
  3361 lemma rel_interior:
       
  3362   "rel_interior S = {x \<in> S. \<exists>T. open T \<and> x \<in> T \<and> T \<inter> affine hull S \<subseteq> S}"
       
  3363   unfolding rel_interior_def[of S] openin_open[of "affine hull S"]
       
  3364   apply auto
       
  3365 proof -
       
  3366   fix x T
       
  3367   assume *: "x \<in> S" "open T" "x \<in> T" "T \<inter> affine hull S \<subseteq> S"
       
  3368   then have **: "x \<in> T \<inter> affine hull S"
       
  3369     using hull_inc by auto
       
  3370   show "\<exists>Tb. (\<exists>Ta. open Ta \<and> Tb = affine hull S \<inter> Ta) \<and> x \<in> Tb \<and> Tb \<subseteq> S"
       
  3371     apply (rule_tac x = "T \<inter> (affine hull S)" in exI)
       
  3372     using * **
       
  3373     apply auto
       
  3374     done
       
  3375 qed
       
  3376 
       
  3377 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)"
       
  3378   by (auto simp add: rel_interior)
       
  3379 
       
  3380 lemma mem_rel_interior_ball:
       
  3381   "x \<in> rel_interior S \<longleftrightarrow> x \<in> S \<and> (\<exists>e. e > 0 \<and> ball x e \<inter> affine hull S \<subseteq> S)"
       
  3382   apply (simp add: rel_interior, safe)
       
  3383   apply (force simp add: open_contains_ball)
       
  3384   apply (rule_tac x = "ball x e" in exI)
       
  3385   apply simp
       
  3386   done
       
  3387 
       
  3388 lemma rel_interior_ball:
       
  3389   "rel_interior S = {x \<in> S. \<exists>e. e > 0 \<and> ball x e \<inter> affine hull S \<subseteq> S}"
       
  3390   using mem_rel_interior_ball [of _ S] by auto
       
  3391 
       
  3392 lemma mem_rel_interior_cball:
       
  3393   "x \<in> rel_interior S \<longleftrightarrow> x \<in> S \<and> (\<exists>e. e > 0 \<and> cball x e \<inter> affine hull S \<subseteq> S)"
       
  3394   apply (simp add: rel_interior, safe)
       
  3395   apply (force simp add: open_contains_cball)
       
  3396   apply (rule_tac x = "ball x e" in exI)
       
  3397   apply (simp add: subset_trans [OF ball_subset_cball])
       
  3398   apply auto
       
  3399   done
       
  3400 
       
  3401 lemma rel_interior_cball:
       
  3402   "rel_interior S = {x \<in> S. \<exists>e. e > 0 \<and> cball x e \<inter> affine hull S \<subseteq> S}"
       
  3403   using mem_rel_interior_cball [of _ S] by auto
       
  3404 
       
  3405 lemma rel_interior_empty [simp]: "rel_interior {} = {}"
       
  3406    by (auto simp add: rel_interior_def)
       
  3407 
       
  3408 lemma affine_hull_sing [simp]: "affine hull {a :: 'n::euclidean_space} = {a}"
       
  3409   by (metis affine_hull_eq affine_sing)
       
  3410 
       
  3411 lemma rel_interior_sing [simp]:
       
  3412     fixes a :: "'n::euclidean_space"  shows "rel_interior {a} = {a}"
       
  3413   apply (auto simp: rel_interior_ball)
       
  3414   apply (rule_tac x=1 in exI)
       
  3415   apply force
       
  3416   done
       
  3417 
       
  3418 lemma subset_rel_interior:
       
  3419   fixes S T :: "'n::euclidean_space set"
       
  3420   assumes "S \<subseteq> T"
       
  3421     and "affine hull S = affine hull T"
       
  3422   shows "rel_interior S \<subseteq> rel_interior T"
       
  3423   using assms by (auto simp add: rel_interior_def)
       
  3424 
       
  3425 lemma rel_interior_subset: "rel_interior S \<subseteq> S"
       
  3426   by (auto simp add: rel_interior_def)
       
  3427 
       
  3428 lemma rel_interior_subset_closure: "rel_interior S \<subseteq> closure S"
       
  3429   using rel_interior_subset by (auto simp add: closure_def)
       
  3430 
       
  3431 lemma interior_subset_rel_interior: "interior S \<subseteq> rel_interior S"
       
  3432   by (auto simp add: rel_interior interior_def)
       
  3433 
       
  3434 lemma interior_rel_interior:
       
  3435   fixes S :: "'n::euclidean_space set"
       
  3436   assumes "aff_dim S = int(DIM('n))"
       
  3437   shows "rel_interior S = interior S"
       
  3438 proof -
       
  3439   have "affine hull S = UNIV"
       
  3440     using assms affine_hull_UNIV[of S] by auto
       
  3441   then show ?thesis
       
  3442     unfolding rel_interior interior_def by auto
       
  3443 qed
       
  3444 
       
  3445 lemma rel_interior_interior:
       
  3446   fixes S :: "'n::euclidean_space set"
       
  3447   assumes "affine hull S = UNIV"
       
  3448   shows "rel_interior S = interior S"
       
  3449   using assms unfolding rel_interior interior_def by auto
       
  3450 
       
  3451 lemma rel_interior_open:
       
  3452   fixes S :: "'n::euclidean_space set"
       
  3453   assumes "open S"
       
  3454   shows "rel_interior S = S"
       
  3455   by (metis assms interior_eq interior_subset_rel_interior rel_interior_subset set_eq_subset)
       
  3456 
       
  3457 lemma interior_ball [simp]: "interior (ball x e) = ball x e"
       
  3458   by (simp add: interior_open)
       
  3459 
       
  3460 lemma interior_rel_interior_gen:
       
  3461   fixes S :: "'n::euclidean_space set"
       
  3462   shows "interior S = (if aff_dim S = int(DIM('n)) then rel_interior S else {})"
       
  3463   by (metis interior_rel_interior low_dim_interior)
       
  3464 
       
  3465 lemma rel_interior_nonempty_interior:
       
  3466   fixes S :: "'n::euclidean_space set"
       
  3467   shows "interior S \<noteq> {} \<Longrightarrow> rel_interior S = interior S"
       
  3468 by (metis interior_rel_interior_gen)
       
  3469 
       
  3470 lemma affine_hull_nonempty_interior:
       
  3471   fixes S :: "'n::euclidean_space set"
       
  3472   shows "interior S \<noteq> {} \<Longrightarrow> affine hull S = UNIV"
       
  3473 by (metis affine_hull_UNIV interior_rel_interior_gen)
       
  3474 
       
  3475 lemma rel_interior_affine_hull [simp]:
       
  3476   fixes S :: "'n::euclidean_space set"
       
  3477   shows "rel_interior (affine hull S) = affine hull S"
       
  3478 proof -
       
  3479   have *: "rel_interior (affine hull S) \<subseteq> affine hull S"
       
  3480     using rel_interior_subset by auto
       
  3481   {
       
  3482     fix x
       
  3483     assume x: "x \<in> affine hull S"
       
  3484     define e :: real where "e = 1"
       
  3485     then have "e > 0" "ball x e \<inter> affine hull (affine hull S) \<subseteq> affine hull S"
       
  3486       using hull_hull[of _ S] by auto
       
  3487     then have "x \<in> rel_interior (affine hull S)"
       
  3488       using x rel_interior_ball[of "affine hull S"] by auto
       
  3489   }
       
  3490   then show ?thesis using * by auto
       
  3491 qed
       
  3492 
       
  3493 lemma rel_interior_UNIV [simp]: "rel_interior (UNIV :: ('n::euclidean_space) set) = UNIV"
       
  3494   by (metis open_UNIV rel_interior_open)
       
  3495 
       
  3496 lemma rel_interior_convex_shrink:
       
  3497   fixes S :: "'a::euclidean_space set"
       
  3498   assumes "convex S"
       
  3499     and "c \<in> rel_interior S"
       
  3500     and "x \<in> S"
       
  3501     and "0 < e"
       
  3502     and "e \<le> 1"
       
  3503   shows "x - e *\<^sub>R (x - c) \<in> rel_interior S"
       
  3504 proof -
       
  3505   obtain d where "d > 0" and d: "ball c d \<inter> affine hull S \<subseteq> S"
       
  3506     using assms(2) unfolding  mem_rel_interior_ball by auto
       
  3507   {
       
  3508     fix y
       
  3509     assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d" "y \<in> affine hull S"
       
  3510     have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x"
       
  3511       using \<open>e > 0\<close> by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib)
       
  3512     have "x \<in> affine hull S"
       
  3513       using assms hull_subset[of S] by auto
       
  3514     moreover have "1 / e + - ((1 - e) / e) = 1"
       
  3515       using \<open>e > 0\<close> left_diff_distrib[of "1" "(1-e)" "1/e"] by auto
       
  3516     ultimately have **: "(1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x \<in> affine hull S"
       
  3517       using as affine_affine_hull[of S] mem_affine[of "affine hull S" y x "(1 / e)" "-((1 - e) / e)"]
       
  3518       by (simp add: algebra_simps)
       
  3519     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)"
       
  3520       unfolding dist_norm norm_scaleR[symmetric]
       
  3521       apply (rule arg_cong[where f=norm])
       
  3522       using \<open>e > 0\<close>
       
  3523       apply (auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps)
       
  3524       done
       
  3525     also have "\<dots> = \<bar>1/e\<bar> * norm (x - e *\<^sub>R (x - c) - y)"
       
  3526       by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps)
       
  3527     also have "\<dots> < d"
       
  3528       using as[unfolded dist_norm] and \<open>e > 0\<close>
       
  3529       by (auto simp add:pos_divide_less_eq[OF \<open>e > 0\<close>] mult.commute)
       
  3530     finally have "y \<in> S"
       
  3531       apply (subst *)
       
  3532       apply (rule assms(1)[unfolded convex_alt,rule_format])
       
  3533       apply (rule d[unfolded subset_eq,rule_format])
       
  3534       unfolding mem_ball
       
  3535       using assms(3-5) **
       
  3536       apply auto
       
  3537       done
       
  3538   }
       
  3539   then have "ball (x - e *\<^sub>R (x - c)) (e*d) \<inter> affine hull S \<subseteq> S"
       
  3540     by auto
       
  3541   moreover have "e * d > 0"
       
  3542     using \<open>e > 0\<close> \<open>d > 0\<close> by simp
       
  3543   moreover have c: "c \<in> S"
       
  3544     using assms rel_interior_subset by auto
       
  3545   moreover from c have "x - e *\<^sub>R (x - c) \<in> S"
       
  3546     using convexD_alt[of S x c e]
       
  3547     apply (simp add: algebra_simps)
       
  3548     using assms
       
  3549     apply auto
       
  3550     done
       
  3551   ultimately show ?thesis
       
  3552     using mem_rel_interior_ball[of "x - e *\<^sub>R (x - c)" S] \<open>e > 0\<close> by auto
       
  3553 qed
       
  3554 
       
  3555 lemma interior_real_semiline:
       
  3556   fixes a :: real
       
  3557   shows "interior {a..} = {a<..}"
       
  3558 proof -
       
  3559   {
       
  3560     fix y
       
  3561     assume "a < y"
       
  3562     then have "y \<in> interior {a..}"
       
  3563       apply (simp add: mem_interior)
       
  3564       apply (rule_tac x="(y-a)" in exI)
       
  3565       apply (auto simp add: dist_norm)
       
  3566       done
       
  3567   }
       
  3568   moreover
       
  3569   {
       
  3570     fix y
       
  3571     assume "y \<in> interior {a..}"
       
  3572     then obtain e where e: "e > 0" "cball y e \<subseteq> {a..}"
       
  3573       using mem_interior_cball[of y "{a..}"] by auto
       
  3574     moreover from e have "y - e \<in> cball y e"
       
  3575       by (auto simp add: cball_def dist_norm)
       
  3576     ultimately have "a \<le> y - e" by blast
       
  3577     then have "a < y" using e by auto
       
  3578   }
       
  3579   ultimately show ?thesis by auto
       
  3580 qed
       
  3581 
       
  3582 lemma continuous_ge_on_Ioo:
       
  3583   assumes "continuous_on {c..d} g" "\<And>x. x \<in> {c<..<d} \<Longrightarrow> g x \<ge> a" "c < d" "x \<in> {c..d}"
       
  3584   shows "g (x::real) \<ge> (a::real)"
       
  3585 proof-
       
  3586   from assms(3) have "{c..d} = closure {c<..<d}" by (rule closure_greaterThanLessThan[symmetric])
       
  3587   also from assms(2) have "{c<..<d} \<subseteq> (g -` {a..} \<inter> {c..d})" by auto
       
  3588   hence "closure {c<..<d} \<subseteq> closure (g -` {a..} \<inter> {c..d})" by (rule closure_mono)
       
  3589   also from assms(1) have "closed (g -` {a..} \<inter> {c..d})"
       
  3590     by (auto simp: continuous_on_closed_vimage)
       
  3591   hence "closure (g -` {a..} \<inter> {c..d}) = g -` {a..} \<inter> {c..d}" by simp
       
  3592   finally show ?thesis using \<open>x \<in> {c..d}\<close> by auto
       
  3593 qed
       
  3594 
       
  3595 lemma interior_real_semiline':
       
  3596   fixes a :: real
       
  3597   shows "interior {..a} = {..<a}"
       
  3598 proof -
       
  3599   {
       
  3600     fix y
       
  3601     assume "a > y"
       
  3602     then have "y \<in> interior {..a}"
       
  3603       apply (simp add: mem_interior)
       
  3604       apply (rule_tac x="(a-y)" in exI)
       
  3605       apply (auto simp add: dist_norm)
       
  3606       done
       
  3607   }
       
  3608   moreover
       
  3609   {
       
  3610     fix y
       
  3611     assume "y \<in> interior {..a}"
       
  3612     then obtain e where e: "e > 0" "cball y e \<subseteq> {..a}"
       
  3613       using mem_interior_cball[of y "{..a}"] by auto
       
  3614     moreover from e have "y + e \<in> cball y e"
       
  3615       by (auto simp add: cball_def dist_norm)
       
  3616     ultimately have "a \<ge> y + e" by auto
       
  3617     then have "a > y" using e by auto
       
  3618   }
       
  3619   ultimately show ?thesis by auto
       
  3620 qed
       
  3621 
       
  3622 lemma interior_atLeastAtMost_real: "interior {a..b} = {a<..<b :: real}"
       
  3623 proof-
       
  3624   have "{a..b} = {a..} \<inter> {..b}" by auto
       
  3625   also have "interior ... = {a<..} \<inter> {..<b}"
       
  3626     by (simp add: interior_real_semiline interior_real_semiline')
       
  3627   also have "... = {a<..<b}" by auto
       
  3628   finally show ?thesis .
       
  3629 qed
       
  3630 
       
  3631 lemma frontier_real_Iic:
       
  3632   fixes a :: real
       
  3633   shows "frontier {..a} = {a}"
       
  3634   unfolding frontier_def by (auto simp add: interior_real_semiline')
       
  3635 
       
  3636 lemma rel_interior_real_box:
       
  3637   fixes a b :: real
       
  3638   assumes "a < b"
       
  3639   shows "rel_interior {a .. b} = {a <..< b}"
       
  3640 proof -
       
  3641   have "box a b \<noteq> {}"
       
  3642     using assms
       
  3643     unfolding set_eq_iff
       
  3644     by (auto intro!: exI[of _ "(a + b) / 2"] simp: box_def)
       
  3645   then show ?thesis
       
  3646     using interior_rel_interior_gen[of "cbox a b", symmetric]
       
  3647     by (simp split: if_split_asm del: box_real add: box_real[symmetric] interior_cbox)
       
  3648 qed
       
  3649 
       
  3650 lemma rel_interior_real_semiline:
       
  3651   fixes a :: real
       
  3652   shows "rel_interior {a..} = {a<..}"
       
  3653 proof -
       
  3654   have *: "{a<..} \<noteq> {}"
       
  3655     unfolding set_eq_iff by (auto intro!: exI[of _ "a + 1"])
       
  3656   then show ?thesis using interior_real_semiline interior_rel_interior_gen[of "{a..}"]
       
  3657     by (auto split: if_split_asm)
       
  3658 qed
       
  3659 
       
  3660 subsubsection \<open>Relative open sets\<close>
       
  3661 
       
  3662 definition "rel_open S \<longleftrightarrow> rel_interior S = S"
       
  3663 
       
  3664 lemma rel_open: "rel_open S \<longleftrightarrow> openin (subtopology euclidean (affine hull S)) S"
       
  3665   unfolding rel_open_def rel_interior_def
       
  3666   apply auto
       
  3667   using openin_subopen[of "subtopology euclidean (affine hull S)" S]
       
  3668   apply auto
       
  3669   done
       
  3670 
       
  3671 lemma openin_rel_interior: "openin (subtopology euclidean (affine hull S)) (rel_interior S)"
       
  3672   apply (simp add: rel_interior_def)
       
  3673   apply (subst openin_subopen)
       
  3674   apply blast
       
  3675   done
       
  3676 
       
  3677 lemma openin_set_rel_interior:
       
  3678    "openin (subtopology euclidean S) (rel_interior S)"
       
  3679 by (rule openin_subset_trans [OF openin_rel_interior rel_interior_subset hull_subset])
       
  3680 
       
  3681 lemma affine_rel_open:
       
  3682   fixes S :: "'n::euclidean_space set"
       
  3683   assumes "affine S"
       
  3684   shows "rel_open S"
       
  3685   unfolding rel_open_def
       
  3686   using assms rel_interior_affine_hull[of S] affine_hull_eq[of S]
       
  3687   by metis
       
  3688 
       
  3689 lemma affine_closed:
       
  3690   fixes S :: "'n::euclidean_space set"
       
  3691   assumes "affine S"
       
  3692   shows "closed S"
       
  3693 proof -
       
  3694   {
       
  3695     assume "S \<noteq> {}"
       
  3696     then obtain L where L: "subspace L" "affine_parallel S L"
       
  3697       using assms affine_parallel_subspace[of S] by auto
       
  3698     then obtain a where a: "S = (op + a ` L)"
       
  3699       using affine_parallel_def[of L S] affine_parallel_commut by auto
       
  3700     from L have "closed L" using closed_subspace by auto
       
  3701     then have "closed S"
       
  3702       using closed_translation a by auto
       
  3703   }
       
  3704   then show ?thesis by auto
       
  3705 qed
       
  3706 
       
  3707 lemma closure_affine_hull:
       
  3708   fixes S :: "'n::euclidean_space set"
       
  3709   shows "closure S \<subseteq> affine hull S"
       
  3710   by (intro closure_minimal hull_subset affine_closed affine_affine_hull)
       
  3711 
       
  3712 lemma closure_same_affine_hull [simp]:
       
  3713   fixes S :: "'n::euclidean_space set"
       
  3714   shows "affine hull (closure S) = affine hull S"
       
  3715 proof -
       
  3716   have "affine hull (closure S) \<subseteq> affine hull S"
       
  3717     using hull_mono[of "closure S" "affine hull S" "affine"]
       
  3718       closure_affine_hull[of S] hull_hull[of "affine" S]
       
  3719     by auto
       
  3720   moreover have "affine hull (closure S) \<supseteq> affine hull S"
       
  3721     using hull_mono[of "S" "closure S" "affine"] closure_subset by auto
       
  3722   ultimately show ?thesis by auto
       
  3723 qed
       
  3724 
       
  3725 lemma closure_aff_dim [simp]:
       
  3726   fixes S :: "'n::euclidean_space set"
       
  3727   shows "aff_dim (closure S) = aff_dim S"
       
  3728 proof -
       
  3729   have "aff_dim S \<le> aff_dim (closure S)"
       
  3730     using aff_dim_subset closure_subset by auto
       
  3731   moreover have "aff_dim (closure S) \<le> aff_dim (affine hull S)"
       
  3732     using aff_dim_subset closure_affine_hull by blast
       
  3733   moreover have "aff_dim (affine hull S) = aff_dim S"
       
  3734     using aff_dim_affine_hull by auto
       
  3735   ultimately show ?thesis by auto
       
  3736 qed
       
  3737 
       
  3738 lemma rel_interior_closure_convex_shrink:
       
  3739   fixes S :: "_::euclidean_space set"
       
  3740   assumes "convex S"
       
  3741     and "c \<in> rel_interior S"
       
  3742     and "x \<in> closure S"
       
  3743     and "e > 0"
       
  3744     and "e \<le> 1"
       
  3745   shows "x - e *\<^sub>R (x - c) \<in> rel_interior S"
       
  3746 proof -
       
  3747   obtain d where "d > 0" and d: "ball c d \<inter> affine hull S \<subseteq> S"
       
  3748     using assms(2) unfolding mem_rel_interior_ball by auto
       
  3749   have "\<exists>y \<in> S. norm (y - x) * (1 - e) < e * d"
       
  3750   proof (cases "x \<in> S")
       
  3751     case True
       
  3752     then show ?thesis using \<open>e > 0\<close> \<open>d > 0\<close>
       
  3753       apply (rule_tac bexI[where x=x])
       
  3754       apply (auto)
       
  3755       done
       
  3756   next
       
  3757     case False
       
  3758     then have x: "x islimpt S"
       
  3759       using assms(3)[unfolded closure_def] by auto
       
  3760     show ?thesis
       
  3761     proof (cases "e = 1")
       
  3762       case True
       
  3763       obtain y where "y \<in> S" "y \<noteq> x" "dist y x < 1"
       
  3764         using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto
       
  3765       then show ?thesis
       
  3766         apply (rule_tac x=y in bexI)
       
  3767         unfolding True
       
  3768         using \<open>d > 0\<close>
       
  3769         apply auto
       
  3770         done
       
  3771     next
       
  3772       case False
       
  3773       then have "0 < e * d / (1 - e)" and *: "1 - e > 0"
       
  3774         using \<open>e \<le> 1\<close> \<open>e > 0\<close> \<open>d > 0\<close> by (auto)
       
  3775       then obtain y where "y \<in> S" "y \<noteq> x" "dist y x < e * d / (1 - e)"
       
  3776         using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto
       
  3777       then show ?thesis
       
  3778         apply (rule_tac x=y in bexI)
       
  3779         unfolding dist_norm
       
  3780         using pos_less_divide_eq[OF *]
       
  3781         apply auto
       
  3782         done
       
  3783     qed
       
  3784   qed
       
  3785   then obtain y where "y \<in> S" and y: "norm (y - x) * (1 - e) < e * d"
       
  3786     by auto
       
  3787   define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)"
       
  3788   have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)"
       
  3789     unfolding z_def using \<open>e > 0\<close>
       
  3790     by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib)
       
  3791   have zball: "z \<in> ball c d"
       
  3792     using mem_ball z_def dist_norm[of c]
       
  3793     using y and assms(4,5)
       
  3794     by (auto simp add:field_simps norm_minus_commute)
       
  3795   have "x \<in> affine hull S"
       
  3796     using closure_affine_hull assms by auto
       
  3797   moreover have "y \<in> affine hull S"
       
  3798     using \<open>y \<in> S\<close> hull_subset[of S] by auto
       
  3799   moreover have "c \<in> affine hull S"
       
  3800     using assms rel_interior_subset hull_subset[of S] by auto
       
  3801   ultimately have "z \<in> affine hull S"
       
  3802     using z_def affine_affine_hull[of S]
       
  3803       mem_affine_3_minus [of "affine hull S" c x y "(1 - e) / e"]
       
  3804       assms
       
  3805     by (auto simp add: field_simps)
       
  3806   then have "z \<in> S" using d zball by auto
       
  3807   obtain d1 where "d1 > 0" and d1: "ball z d1 \<le> ball c d"
       
  3808     using zball open_ball[of c d] openE[of "ball c d" z] by auto
       
  3809   then have "ball z d1 \<inter> affine hull S \<subseteq> ball c d \<inter> affine hull S"
       
  3810     by auto
       
  3811   then have "ball z d1 \<inter> affine hull S \<subseteq> S"
       
  3812     using d by auto
       
  3813   then have "z \<in> rel_interior S"
       
  3814     using mem_rel_interior_ball using \<open>d1 > 0\<close> \<open>z \<in> S\<close> by auto
       
  3815   then have "y - e *\<^sub>R (y - z) \<in> rel_interior S"
       
  3816     using rel_interior_convex_shrink[of S z y e] assms \<open>y \<in> S\<close> by auto
       
  3817   then show ?thesis using * by auto
       
  3818 qed
       
  3819 
       
  3820 lemma rel_interior_eq:
       
  3821    "rel_interior s = s \<longleftrightarrow> openin(subtopology euclidean (affine hull s)) s"
       
  3822 using rel_open rel_open_def by blast
       
  3823 
       
  3824 lemma rel_interior_openin:
       
  3825    "openin(subtopology euclidean (affine hull s)) s \<Longrightarrow> rel_interior s = s"
       
  3826 by (simp add: rel_interior_eq)
       
  3827 
       
  3828 lemma rel_interior_affine:
       
  3829   fixes S :: "'n::euclidean_space set"
       
  3830   shows  "affine S \<Longrightarrow> rel_interior S = S"
       
  3831 using affine_rel_open rel_open_def by auto
       
  3832 
       
  3833 lemma rel_interior_eq_closure:
       
  3834   fixes S :: "'n::euclidean_space set"
       
  3835   shows "rel_interior S = closure S \<longleftrightarrow> affine S"
       
  3836 proof (cases "S = {}")
       
  3837   case True
       
  3838  then show ?thesis
       
  3839     by auto
       
  3840 next
       
  3841   case False show ?thesis
       
  3842   proof
       
  3843     assume eq: "rel_interior S = closure S"
       
  3844     have "S = {} \<or> S = affine hull S"
       
  3845       apply (rule connected_clopen [THEN iffD1, rule_format])
       
  3846        apply (simp add: affine_imp_convex convex_connected)
       
  3847       apply (rule conjI)
       
  3848        apply (metis eq closure_subset openin_rel_interior rel_interior_subset subset_antisym)
       
  3849       apply (metis closed_subset closure_subset_eq eq hull_subset rel_interior_subset)
       
  3850       done
       
  3851     with False have "affine hull S = S"
       
  3852       by auto
       
  3853     then show "affine S"
       
  3854       by (metis affine_hull_eq)
       
  3855   next
       
  3856     assume "affine S"
       
  3857     then show "rel_interior S = closure S"
       
  3858       by (simp add: rel_interior_affine affine_closed)
       
  3859   qed
       
  3860 qed
       
  3861 
       
  3862 
       
  3863 subsubsection\<open>Relative interior preserves under linear transformations\<close>
       
  3864 
       
  3865 lemma rel_interior_translation_aux:
       
  3866   fixes a :: "'n::euclidean_space"
       
  3867   shows "((\<lambda>x. a + x) ` rel_interior S) \<subseteq> rel_interior ((\<lambda>x. a + x) ` S)"
       
  3868 proof -
       
  3869   {
       
  3870     fix x
       
  3871     assume x: "x \<in> rel_interior S"
       
  3872     then obtain T where "open T" "x \<in> T \<inter> S" "T \<inter> affine hull S \<subseteq> S"
       
  3873       using mem_rel_interior[of x S] by auto
       
  3874     then have "open ((\<lambda>x. a + x) ` T)"
       
  3875       and "a + x \<in> ((\<lambda>x. a + x) ` T) \<inter> ((\<lambda>x. a + x) ` S)"
       
  3876       and "((\<lambda>x. a + x) ` T) \<inter> affine hull ((\<lambda>x. a + x) ` S) \<subseteq> (\<lambda>x. a + x) ` S"
       
  3877       using affine_hull_translation[of a S] open_translation[of T a] x by auto
       
  3878     then have "a + x \<in> rel_interior ((\<lambda>x. a + x) ` S)"
       
  3879       using mem_rel_interior[of "a+x" "((\<lambda>x. a + x) ` S)"] by auto
       
  3880   }
       
  3881   then show ?thesis by auto
       
  3882 qed
       
  3883 
       
  3884 lemma rel_interior_translation:
       
  3885   fixes a :: "'n::euclidean_space"
       
  3886   shows "rel_interior ((\<lambda>x. a + x) ` S) = (\<lambda>x. a + x) ` rel_interior S"
       
  3887 proof -
       
  3888   have "(\<lambda>x. (-a) + x) ` rel_interior ((\<lambda>x. a + x) ` S) \<subseteq> rel_interior S"
       
  3889     using rel_interior_translation_aux[of "-a" "(\<lambda>x. a + x) ` S"]
       
  3890       translation_assoc[of "-a" "a"]
       
  3891     by auto
       
  3892   then have "((\<lambda>x. a + x) ` rel_interior S) \<supseteq> rel_interior ((\<lambda>x. a + x) ` S)"
       
  3893     using translation_inverse_subset[of a "rel_interior (op + a ` S)" "rel_interior S"]
       
  3894     by auto
       
  3895   then show ?thesis
       
  3896     using rel_interior_translation_aux[of a S] by auto
       
  3897 qed
       
  3898 
       
  3899 
       
  3900 lemma affine_hull_linear_image:
       
  3901   assumes "bounded_linear f"
       
  3902   shows "f ` (affine hull s) = affine hull f ` s"
       
  3903   apply rule
       
  3904   unfolding subset_eq ball_simps
       
  3905   apply (rule_tac[!] hull_induct, rule hull_inc)
       
  3906   prefer 3
       
  3907   apply (erule imageE)
       
  3908   apply (rule_tac x=xa in image_eqI)
       
  3909   apply assumption
       
  3910   apply (rule hull_subset[unfolded subset_eq, rule_format])
       
  3911   apply assumption
       
  3912 proof -
       
  3913   interpret f: bounded_linear f by fact
       
  3914   show "affine {x. f x \<in> affine hull f ` s}"
       
  3915     unfolding affine_def
       
  3916     by (auto simp add: f.scaleR f.add affine_affine_hull[unfolded affine_def, rule_format])
       
  3917   show "affine {x. x \<in> f ` (affine hull s)}"
       
  3918     using affine_affine_hull[unfolded affine_def, of s]
       
  3919     unfolding affine_def by (auto simp add: f.scaleR [symmetric] f.add [symmetric])
       
  3920 qed auto
       
  3921 
       
  3922 
       
  3923 lemma rel_interior_injective_on_span_linear_image:
       
  3924   fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
       
  3925     and S :: "'m::euclidean_space set"
       
  3926   assumes "bounded_linear f"
       
  3927     and "inj_on f (span S)"
       
  3928   shows "rel_interior (f ` S) = f ` (rel_interior S)"
       
  3929 proof -
       
  3930   {
       
  3931     fix z
       
  3932     assume z: "z \<in> rel_interior (f ` S)"
       
  3933     then have "z \<in> f ` S"
       
  3934       using rel_interior_subset[of "f ` S"] by auto
       
  3935     then obtain x where x: "x \<in> S" "f x = z" by auto
       
  3936     obtain e2 where e2: "e2 > 0" "cball z e2 \<inter> affine hull (f ` S) \<subseteq> (f ` S)"
       
  3937       using z rel_interior_cball[of "f ` S"] by auto
       
  3938     obtain K where K: "K > 0" "\<And>x. norm (f x) \<le> norm x * K"
       
  3939      using assms Real_Vector_Spaces.bounded_linear.pos_bounded[of f] by auto
       
  3940     define e1 where "e1 = 1 / K"
       
  3941     then have e1: "e1 > 0" "\<And>x. e1 * norm (f x) \<le> norm x"
       
  3942       using K pos_le_divide_eq[of e1] by auto
       
  3943     define e where "e = e1 * e2"
       
  3944     then have "e > 0" using e1 e2 by auto
       
  3945     {
       
  3946       fix y
       
  3947       assume y: "y \<in> cball x e \<inter> affine hull S"
       
  3948       then have h1: "f y \<in> affine hull (f ` S)"
       
  3949         using affine_hull_linear_image[of f S] assms by auto
       
  3950       from y have "norm (x-y) \<le> e1 * e2"
       
  3951         using cball_def[of x e] dist_norm[of x y] e_def by auto
       
  3952       moreover have "f x - f y = f (x - y)"
       
  3953         using assms linear_diff[of f x y] linear_conv_bounded_linear[of f] by auto
       
  3954       moreover have "e1 * norm (f (x-y)) \<le> norm (x - y)"
       
  3955         using e1 by auto
       
  3956       ultimately have "e1 * norm ((f x)-(f y)) \<le> e1 * e2"
       
  3957         by auto
       
  3958       then have "f y \<in> cball z e2"
       
  3959         using cball_def[of "f x" e2] dist_norm[of "f x" "f y"] e1 x by auto
       
  3960       then have "f y \<in> f ` S"
       
  3961         using y e2 h1 by auto
       
  3962       then have "y \<in> S"
       
  3963         using assms y hull_subset[of S] affine_hull_subset_span
       
  3964           inj_on_image_mem_iff [OF \<open>inj_on f (span S)\<close>]
       
  3965         by (metis Int_iff span_inc subsetCE)
       
  3966     }
       
  3967     then have "z \<in> f ` (rel_interior S)"
       
  3968       using mem_rel_interior_cball[of x S] \<open>e > 0\<close> x by auto
       
  3969   }
       
  3970   moreover
       
  3971   {
       
  3972     fix x
       
  3973     assume x: "x \<in> rel_interior S"
       
  3974     then obtain e2 where e2: "e2 > 0" "cball x e2 \<inter> affine hull S \<subseteq> S"
       
  3975       using rel_interior_cball[of S] by auto
       
  3976     have "x \<in> S" using x rel_interior_subset by auto
       
  3977     then have *: "f x \<in> f ` S" by auto
       
  3978     have "\<forall>x\<in>span S. f x = 0 \<longrightarrow> x = 0"
       
  3979       using assms subspace_span linear_conv_bounded_linear[of f]
       
  3980         linear_injective_on_subspace_0[of f "span S"]
       
  3981       by auto
       
  3982     then obtain e1 where e1: "e1 > 0" "\<forall>x \<in> span S. e1 * norm x \<le> norm (f x)"
       
  3983       using assms injective_imp_isometric[of "span S" f]
       
  3984         subspace_span[of S] closed_subspace[of "span S"]
       
  3985       by auto
       
  3986     define e where "e = e1 * e2"
       
  3987     hence "e > 0" using e1 e2 by auto
       
  3988     {
       
  3989       fix y
       
  3990       assume y: "y \<in> cball (f x) e \<inter> affine hull (f ` S)"
       
  3991       then have "y \<in> f ` (affine hull S)"
       
  3992         using affine_hull_linear_image[of f S] assms by auto
       
  3993       then obtain xy where xy: "xy \<in> affine hull S" "f xy = y" by auto
       
  3994       with y have "norm (f x - f xy) \<le> e1 * e2"
       
  3995         using cball_def[of "f x" e] dist_norm[of "f x" y] e_def by auto
       
  3996       moreover have "f x - f xy = f (x - xy)"
       
  3997         using assms linear_diff[of f x xy] linear_conv_bounded_linear[of f] by auto
       
  3998       moreover have *: "x - xy \<in> span S"
       
  3999         using subspace_diff[of "span S" x xy] subspace_span \<open>x \<in> S\<close> xy
       
  4000           affine_hull_subset_span[of S] span_inc
       
  4001         by auto
       
  4002       moreover from * have "e1 * norm (x - xy) \<le> norm (f (x - xy))"
       
  4003         using e1 by auto
       
  4004       ultimately have "e1 * norm (x - xy) \<le> e1 * e2"
       
  4005         by auto
       
  4006       then have "xy \<in> cball x e2"
       
  4007         using cball_def[of x e2] dist_norm[of x xy] e1 by auto
       
  4008       then have "y \<in> f ` S"
       
  4009         using xy e2 by auto
       
  4010     }
       
  4011     then have "f x \<in> rel_interior (f ` S)"
       
  4012       using mem_rel_interior_cball[of "(f x)" "(f ` S)"] * \<open>e > 0\<close> by auto
       
  4013   }
       
  4014   ultimately show ?thesis by auto
       
  4015 qed
       
  4016 
       
  4017 lemma rel_interior_injective_linear_image:
       
  4018   fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
       
  4019   assumes "bounded_linear f"
       
  4020     and "inj f"
       
  4021   shows "rel_interior (f ` S) = f ` (rel_interior S)"
       
  4022   using assms rel_interior_injective_on_span_linear_image[of f S]
       
  4023     subset_inj_on[of f "UNIV" "span S"]
       
  4024   by auto
       
  4025 
       
  4026 
       
  4027 subsection\<open>Some Properties of subset of standard basis\<close>
       
  4028 
       
  4029 lemma affine_hull_substd_basis:
       
  4030   assumes "d \<subseteq> Basis"
       
  4031   shows "affine hull (insert 0 d) = {x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
       
  4032   (is "affine hull (insert 0 ?A) = ?B")
       
  4033 proof -
       
  4034   have *: "\<And>A. op + (0::'a) ` A = A" "\<And>A. op + (- (0::'a)) ` A = A"
       
  4035     by auto
       
  4036   show ?thesis
       
  4037     unfolding affine_hull_insert_span_gen span_substd_basis[OF assms,symmetric] * ..
       
  4038 qed
       
  4039 
       
  4040 lemma affine_hull_convex_hull [simp]: "affine hull (convex hull S) = affine hull S"
       
  4041   by (metis Int_absorb1 Int_absorb2 convex_hull_subset_affine_hull hull_hull hull_mono hull_subset)
       
  4042 
       
  4043 
       
  4044 subsection \<open>Openness and compactness are preserved by convex hull operation.\<close>
       
  4045 
       
  4046 lemma open_convex_hull[intro]:
       
  4047   fixes s :: "'a::real_normed_vector set"
       
  4048   assumes "open s"
       
  4049   shows "open (convex hull s)"
       
  4050   unfolding open_contains_cball convex_hull_explicit
       
  4051   unfolding mem_Collect_eq ball_simps(8)
       
  4052 proof (rule, rule)
       
  4053   fix a
       
  4054   assume "\<exists>sa u. finite sa \<and> sa \<subseteq> s \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = a"
       
  4055   then obtain t u where obt: "finite t" "t\<subseteq>s" "\<forall>x\<in>t. 0 \<le> u x" "setsum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = a"
       
  4056     by auto
       
  4057 
       
  4058   from assms[unfolded open_contains_cball] obtain b
       
  4059     where b: "\<forall>x\<in>s. 0 < b x \<and> cball x (b x) \<subseteq> s"
       
  4060     using bchoice[of s "\<lambda>x e. e > 0 \<and> cball x e \<subseteq> s"] by auto
       
  4061   have "b ` t \<noteq> {}"
       
  4062     using obt by auto
       
  4063   define i where "i = b ` t"
       
  4064 
       
  4065   show "\<exists>e > 0.
       
  4066     cball a e \<subseteq> {y. \<exists>sa u. finite sa \<and> sa \<subseteq> s \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y}"
       
  4067     apply (rule_tac x = "Min i" in exI)
       
  4068     unfolding subset_eq
       
  4069     apply rule
       
  4070     defer
       
  4071     apply rule
       
  4072     unfolding mem_Collect_eq
       
  4073   proof -
       
  4074     show "0 < Min i"
       
  4075       unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] \<open>b ` t\<noteq>{}\<close>]
       
  4076       using b
       
  4077       apply simp
       
  4078       apply rule
       
  4079       apply (erule_tac x=x in ballE)
       
  4080       using \<open>t\<subseteq>s\<close>
       
  4081       apply auto
       
  4082       done
       
  4083   next
       
  4084     fix y
       
  4085     assume "y \<in> cball a (Min i)"
       
  4086     then have y: "norm (a - y) \<le> Min i"
       
  4087       unfolding dist_norm[symmetric] by auto
       
  4088     {
       
  4089       fix x
       
  4090       assume "x \<in> t"
       
  4091       then have "Min i \<le> b x"
       
  4092         unfolding i_def
       
  4093         apply (rule_tac Min_le)
       
  4094         using obt(1)
       
  4095         apply auto
       
  4096         done
       
  4097       then have "x + (y - a) \<in> cball x (b x)"
       
  4098         using y unfolding mem_cball dist_norm by auto
       
  4099       moreover from \<open>x\<in>t\<close> have "x \<in> s"
       
  4100         using obt(2) by auto
       
  4101       ultimately have "x + (y - a) \<in> s"
       
  4102         using y and b[THEN bspec[where x=x]] unfolding subset_eq by fast
       
  4103     }
       
  4104     moreover
       
  4105     have *: "inj_on (\<lambda>v. v + (y - a)) t"
       
  4106       unfolding inj_on_def by auto
       
  4107     have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a))) = 1"
       
  4108       unfolding setsum.reindex[OF *] o_def using obt(4) by auto
       
  4109     moreover have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a)) *\<^sub>R v) = y"
       
  4110       unfolding setsum.reindex[OF *] o_def using obt(4,5)
       
  4111       by (simp add: setsum.distrib setsum_subtractf scaleR_left.setsum[symmetric] scaleR_right_distrib)
       
  4112     ultimately
       
  4113     show "\<exists>sa u. finite sa \<and> (\<forall>x\<in>sa. x \<in> s) \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y"
       
  4114       apply (rule_tac x="(\<lambda>v. v + (y - a)) ` t" in exI)
       
  4115       apply (rule_tac x="\<lambda>v. u (v - (y - a))" in exI)
       
  4116       using obt(1, 3)
       
  4117       apply auto
       
  4118       done
       
  4119   qed
       
  4120 qed
       
  4121 
       
  4122 lemma compact_convex_combinations:
       
  4123   fixes s t :: "'a::real_normed_vector set"
       
  4124   assumes "compact s" "compact t"
       
  4125   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}"
       
  4126 proof -
       
  4127   let ?X = "{0..1} \<times> s \<times> t"
       
  4128   let ?h = "(\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))"
       
  4129   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"
       
  4130     apply (rule set_eqI)
       
  4131     unfolding image_iff mem_Collect_eq
       
  4132     apply rule
       
  4133     apply auto
       
  4134     apply (rule_tac x=u in rev_bexI)
       
  4135     apply simp
       
  4136     apply (erule rev_bexI)
       
  4137     apply (erule rev_bexI)
       
  4138     apply simp
       
  4139     apply auto
       
  4140     done
       
  4141   have "continuous_on ?X (\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))"
       
  4142     unfolding continuous_on by (rule ballI) (intro tendsto_intros)
       
  4143   then show ?thesis
       
  4144     unfolding *
       
  4145     apply (rule compact_continuous_image)
       
  4146     apply (intro compact_Times compact_Icc assms)
       
  4147     done
       
  4148 qed
       
  4149 
       
  4150 lemma finite_imp_compact_convex_hull:
       
  4151   fixes s :: "'a::real_normed_vector set"
       
  4152   assumes "finite s"
       
  4153   shows "compact (convex hull s)"
       
  4154 proof (cases "s = {}")
       
  4155   case True
       
  4156   then show ?thesis by simp
       
  4157 next
       
  4158   case False
       
  4159   with assms show ?thesis
       
  4160   proof (induct rule: finite_ne_induct)
       
  4161     case (singleton x)
       
  4162     show ?case by simp
       
  4163   next
       
  4164     case (insert x A)
       
  4165     let ?f = "\<lambda>(u, y::'a). u *\<^sub>R x + (1 - u) *\<^sub>R y"
       
  4166     let ?T = "{0..1::real} \<times> (convex hull A)"
       
  4167     have "continuous_on ?T ?f"
       
  4168       unfolding split_def continuous_on by (intro ballI tendsto_intros)
       
  4169     moreover have "compact ?T"
       
  4170       by (intro compact_Times compact_Icc insert)
       
  4171     ultimately have "compact (?f ` ?T)"
       
  4172       by (rule compact_continuous_image)
       
  4173     also have "?f ` ?T = convex hull (insert x A)"
       
  4174       unfolding convex_hull_insert [OF \<open>A \<noteq> {}\<close>]
       
  4175       apply safe
       
  4176       apply (rule_tac x=a in exI, simp)
       
  4177       apply (rule_tac x="1 - a" in exI, simp)
       
  4178       apply fast
       
  4179       apply (rule_tac x="(u, b)" in image_eqI, simp_all)
       
  4180       done
       
  4181     finally show "compact (convex hull (insert x A))" .
       
  4182   qed
       
  4183 qed
       
  4184 
       
  4185 lemma compact_convex_hull:
       
  4186   fixes s :: "'a::euclidean_space set"
       
  4187   assumes "compact s"
       
  4188   shows "compact (convex hull s)"
       
  4189 proof (cases "s = {}")
       
  4190   case True
       
  4191   then show ?thesis using compact_empty by simp
       
  4192 next
       
  4193   case False
       
  4194   then obtain w where "w \<in> s" by auto
       
  4195   show ?thesis
       
  4196     unfolding caratheodory[of s]
       
  4197   proof (induct ("DIM('a) + 1"))
       
  4198     case 0
       
  4199     have *: "{x.\<exists>sa. finite sa \<and> sa \<subseteq> s \<and> card sa \<le> 0 \<and> x \<in> convex hull sa} = {}"
       
  4200       using compact_empty by auto
       
  4201     from 0 show ?case unfolding * by simp
       
  4202   next
       
  4203     case (Suc n)
       
  4204     show ?case
       
  4205     proof (cases "n = 0")
       
  4206       case True
       
  4207       have "{x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t} = s"
       
  4208         unfolding set_eq_iff and mem_Collect_eq
       
  4209       proof (rule, rule)
       
  4210         fix x
       
  4211         assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
       
  4212         then obtain t where t: "finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t"
       
  4213           by auto
       
  4214         show "x \<in> s"
       
  4215         proof (cases "card t = 0")
       
  4216           case True
       
  4217           then show ?thesis
       
  4218             using t(4) unfolding card_0_eq[OF t(1)] by simp
       
  4219         next
       
  4220           case False
       
  4221           then have "card t = Suc 0" using t(3) \<open>n=0\<close> by auto
       
  4222           then obtain a where "t = {a}" unfolding card_Suc_eq by auto
       
  4223           then show ?thesis using t(2,4) by simp
       
  4224         qed
       
  4225       next
       
  4226         fix x assume "x\<in>s"
       
  4227         then show "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
       
  4228           apply (rule_tac x="{x}" in exI)
       
  4229           unfolding convex_hull_singleton
       
  4230           apply auto
       
  4231           done
       
  4232       qed
       
  4233       then show ?thesis using assms by simp
       
  4234     next
       
  4235       case False
       
  4236       have "{x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t} =
       
  4237         {(1 - u) *\<^sub>R x + u *\<^sub>R y | x y u.
       
  4238           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}}"
       
  4239         unfolding set_eq_iff and mem_Collect_eq
       
  4240       proof (rule, rule)
       
  4241         fix x
       
  4242         assume "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and>
       
  4243           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)"
       
  4244         then obtain u v c t where obt: "x = (1 - c) *\<^sub>R u + c *\<^sub>R v"
       
  4245           "0 \<le> c \<and> c \<le> 1" "u \<in> s" "finite t" "t \<subseteq> s" "card t \<le> n"  "v \<in> convex hull t"
       
  4246           by auto
       
  4247         moreover have "(1 - c) *\<^sub>R u + c *\<^sub>R v \<in> convex hull insert u t"
       
  4248           apply (rule convexD_alt)
       
  4249           using obt(2) and convex_convex_hull and hull_subset[of "insert u t" convex]
       
  4250           using obt(7) and hull_mono[of t "insert u t"]
       
  4251           apply auto
       
  4252           done
       
  4253         ultimately show "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
       
  4254           apply (rule_tac x="insert u t" in exI)
       
  4255           apply (auto simp add: card_insert_if)
       
  4256           done
       
  4257       next
       
  4258         fix x
       
  4259         assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
       
  4260         then obtain t where t: "finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t"
       
  4261           by auto
       
  4262         show "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and>
       
  4263           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)"
       
  4264         proof (cases "card t = Suc n")
       
  4265           case False
       
  4266           then have "card t \<le> n" using t(3) by auto
       
  4267           then show ?thesis
       
  4268             apply (rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI)
       
  4269             using \<open>w\<in>s\<close> and t
       
  4270             apply (auto intro!: exI[where x=t])
       
  4271             done
       
  4272         next
       
  4273           case True
       
  4274           then obtain a u where au: "t = insert a u" "a\<notin>u"
       
  4275             apply (drule_tac card_eq_SucD)
       
  4276             apply auto
       
  4277             done
       
  4278           show ?thesis
       
  4279           proof (cases "u = {}")
       
  4280             case True
       
  4281             then have "x = a" using t(4)[unfolded au] by auto
       
  4282             show ?thesis unfolding \<open>x = a\<close>
       
  4283               apply (rule_tac x=a in exI)
       
  4284               apply (rule_tac x=a in exI)
       
  4285               apply (rule_tac x=1 in exI)
       
  4286               using t and \<open>n \<noteq> 0\<close>
       
  4287               unfolding au
       
  4288               apply (auto intro!: exI[where x="{a}"])
       
  4289               done
       
  4290           next
       
  4291             case False
       
  4292             obtain ux vx b where obt: "ux\<ge>0" "vx\<ge>0" "ux + vx = 1"
       
  4293               "b \<in> convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b"
       
  4294               using t(4)[unfolded au convex_hull_insert[OF False]]
       
  4295               by auto
       
  4296             have *: "1 - vx = ux" using obt(3) by auto
       
  4297             show ?thesis
       
  4298               apply (rule_tac x=a in exI)
       
  4299               apply (rule_tac x=b in exI)
       
  4300               apply (rule_tac x=vx in exI)
       
  4301               using obt and t(1-3)
       
  4302               unfolding au and * using card_insert_disjoint[OF _ au(2)]
       
  4303               apply (auto intro!: exI[where x=u])
       
  4304               done
       
  4305           qed
       
  4306         qed
       
  4307       qed
       
  4308       then show ?thesis
       
  4309         using compact_convex_combinations[OF assms Suc] by simp
       
  4310     qed
       
  4311   qed
       
  4312 qed
       
  4313 
       
  4314 
       
  4315 subsection \<open>Extremal points of a simplex are some vertices.\<close>
       
  4316 
       
  4317 lemma dist_increases_online:
       
  4318   fixes a b d :: "'a::real_inner"
       
  4319   assumes "d \<noteq> 0"
       
  4320   shows "dist a (b + d) > dist a b \<or> dist a (b - d) > dist a b"
       
  4321 proof (cases "inner a d - inner b d > 0")
       
  4322   case True
       
  4323   then have "0 < inner d d + (inner a d * 2 - inner b d * 2)"
       
  4324     apply (rule_tac add_pos_pos)
       
  4325     using assms
       
  4326     apply auto
       
  4327     done
       
  4328   then show ?thesis
       
  4329     apply (rule_tac disjI2)
       
  4330     unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff
       
  4331     apply  (simp add: algebra_simps inner_commute)
       
  4332     done
       
  4333 next
       
  4334   case False
       
  4335   then have "0 < inner d d + (inner b d * 2 - inner a d * 2)"
       
  4336     apply (rule_tac add_pos_nonneg)
       
  4337     using assms
       
  4338     apply auto
       
  4339     done
       
  4340   then show ?thesis
       
  4341     apply (rule_tac disjI1)
       
  4342     unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff
       
  4343     apply (simp add: algebra_simps inner_commute)
       
  4344     done
       
  4345 qed
       
  4346 
       
  4347 lemma norm_increases_online:
       
  4348   fixes d :: "'a::real_inner"
       
  4349   shows "d \<noteq> 0 \<Longrightarrow> norm (a + d) > norm a \<or> norm(a - d) > norm a"
       
  4350   using dist_increases_online[of d a 0] unfolding dist_norm by auto
       
  4351 
       
  4352 lemma simplex_furthest_lt:
       
  4353   fixes s :: "'a::real_inner set"
       
  4354   assumes "finite s"
       
  4355   shows "\<forall>x \<in> convex hull s.  x \<notin> s \<longrightarrow> (\<exists>y \<in> convex hull s. norm (x - a) < norm(y - a))"
       
  4356   using assms
       
  4357 proof induct
       
  4358   fix x s
       
  4359   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))"
       
  4360   show "\<forall>xa\<in>convex hull insert x s. xa \<notin> insert x s \<longrightarrow>
       
  4361     (\<exists>y\<in>convex hull insert x s. norm (xa - a) < norm (y - a))"
       
  4362   proof (rule, rule, cases "s = {}")
       
  4363     case False
       
  4364     fix y
       
  4365     assume y: "y \<in> convex hull insert x s" "y \<notin> insert x s"
       
  4366     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"
       
  4367       using y(1)[unfolded convex_hull_insert[OF False]] by auto
       
  4368     show "\<exists>z\<in>convex hull insert x s. norm (y - a) < norm (z - a)"
       
  4369     proof (cases "y \<in> convex hull s")
       
  4370       case True
       
  4371       then obtain z where "z \<in> convex hull s" "norm (y - a) < norm (z - a)"
       
  4372         using as(3)[THEN bspec[where x=y]] and y(2) by auto
       
  4373       then show ?thesis
       
  4374         apply (rule_tac x=z in bexI)
       
  4375         unfolding convex_hull_insert[OF False]
       
  4376         apply auto
       
  4377         done
       
  4378     next
       
  4379       case False
       
  4380       show ?thesis
       
  4381         using obt(3)
       
  4382       proof (cases "u = 0", case_tac[!] "v = 0")
       
  4383         assume "u = 0" "v \<noteq> 0"
       
  4384         then have "y = b" using obt by auto
       
  4385         then show ?thesis using False and obt(4) by auto
       
  4386       next
       
  4387         assume "u \<noteq> 0" "v = 0"
       
  4388         then have "y = x" using obt by auto
       
  4389         then show ?thesis using y(2) by auto
       
  4390       next
       
  4391         assume "u \<noteq> 0" "v \<noteq> 0"
       
  4392         then obtain w where w: "w>0" "w<u" "w<v"
       
  4393           using real_lbound_gt_zero[of u v] and obt(1,2) by auto
       
  4394         have "x \<noteq> b"
       
  4395         proof
       
  4396           assume "x = b"
       
  4397           then have "y = b" unfolding obt(5)
       
  4398             using obt(3) by (auto simp add: scaleR_left_distrib[symmetric])
       
  4399           then show False using obt(4) and False by simp
       
  4400         qed
       
  4401         then have *: "w *\<^sub>R (x - b) \<noteq> 0" using w(1) by auto
       
  4402         show ?thesis
       
  4403           using dist_increases_online[OF *, of a y]
       
  4404         proof (elim disjE)
       
  4405           assume "dist a y < dist a (y + w *\<^sub>R (x - b))"
       
  4406           then have "norm (y - a) < norm ((u + w) *\<^sub>R x + (v - w) *\<^sub>R b - a)"
       
  4407             unfolding dist_commute[of a]
       
  4408             unfolding dist_norm obt(5)
       
  4409             by (simp add: algebra_simps)
       
  4410           moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \<in> convex hull insert x s"
       
  4411             unfolding convex_hull_insert[OF \<open>s\<noteq>{}\<close>] and mem_Collect_eq
       
  4412             apply (rule_tac x="u + w" in exI)
       
  4413             apply rule
       
  4414             defer
       
  4415             apply (rule_tac x="v - w" in exI)
       
  4416             using \<open>u \<ge> 0\<close> and w and obt(3,4)
       
  4417             apply auto
       
  4418             done
       
  4419           ultimately show ?thesis by auto
       
  4420         next
       
  4421           assume "dist a y < dist a (y - w *\<^sub>R (x - b))"
       
  4422           then have "norm (y - a) < norm ((u - w) *\<^sub>R x + (v + w) *\<^sub>R b - a)"
       
  4423             unfolding dist_commute[of a]
       
  4424             unfolding dist_norm obt(5)
       
  4425             by (simp add: algebra_simps)
       
  4426           moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \<in> convex hull insert x s"
       
  4427             unfolding convex_hull_insert[OF \<open>s\<noteq>{}\<close>] and mem_Collect_eq
       
  4428             apply (rule_tac x="u - w" in exI)
       
  4429             apply rule
       
  4430             defer
       
  4431             apply (rule_tac x="v + w" in exI)
       
  4432             using \<open>u \<ge> 0\<close> and w and obt(3,4)
       
  4433             apply auto
       
  4434             done
       
  4435           ultimately show ?thesis by auto
       
  4436         qed
       
  4437       qed auto
       
  4438     qed
       
  4439   qed auto
       
  4440 qed (auto simp add: assms)
       
  4441 
       
  4442 lemma simplex_furthest_le:
       
  4443   fixes s :: "'a::real_inner set"
       
  4444   assumes "finite s"
       
  4445     and "s \<noteq> {}"
       
  4446   shows "\<exists>y\<in>s. \<forall>x\<in> convex hull s. norm (x - a) \<le> norm (y - a)"
       
  4447 proof -
       
  4448   have "convex hull s \<noteq> {}"
       
  4449     using hull_subset[of s convex] and assms(2) by auto
       
  4450   then obtain x where x: "x \<in> convex hull s" "\<forall>y\<in>convex hull s. norm (y - a) \<le> norm (x - a)"
       
  4451     using distance_attains_sup[OF finite_imp_compact_convex_hull[OF assms(1)], of a]
       
  4452     unfolding dist_commute[of a]
       
  4453     unfolding dist_norm
       
  4454     by auto
       
  4455   show ?thesis
       
  4456   proof (cases "x \<in> s")
       
  4457     case False
       
  4458     then obtain y where "y \<in> convex hull s" "norm (x - a) < norm (y - a)"
       
  4459       using simplex_furthest_lt[OF assms(1), THEN bspec[where x=x]] and x(1)
       
  4460       by auto
       
  4461     then show ?thesis
       
  4462       using x(2)[THEN bspec[where x=y]] by auto
       
  4463   next
       
  4464     case True
       
  4465     with x show ?thesis by auto
       
  4466   qed
       
  4467 qed
       
  4468 
       
  4469 lemma simplex_furthest_le_exists:
       
  4470   fixes s :: "('a::real_inner) set"
       
  4471   shows "finite s \<Longrightarrow> \<forall>x\<in>(convex hull s). \<exists>y\<in>s. norm (x - a) \<le> norm (y - a)"
       
  4472   using simplex_furthest_le[of s] by (cases "s = {}") auto
       
  4473 
       
  4474 lemma simplex_extremal_le:
       
  4475   fixes s :: "'a::real_inner set"
       
  4476   assumes "finite s"
       
  4477     and "s \<noteq> {}"
       
  4478   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)"
       
  4479 proof -
       
  4480   have "convex hull s \<noteq> {}"
       
  4481     using hull_subset[of s convex] and assms(2) by auto
       
  4482   then obtain u v where obt: "u \<in> convex hull s" "v \<in> convex hull s"
       
  4483     "\<forall>x\<in>convex hull s. \<forall>y\<in>convex hull s. norm (x - y) \<le> norm (u - v)"
       
  4484     using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]]
       
  4485     by (auto simp: dist_norm)
       
  4486   then show ?thesis
       
  4487   proof (cases "u\<notin>s \<or> v\<notin>s", elim disjE)
       
  4488     assume "u \<notin> s"
       
  4489     then obtain y where "y \<in> convex hull s" "norm (u - v) < norm (y - v)"
       
  4490       using simplex_furthest_lt[OF assms(1), THEN bspec[where x=u]] and obt(1)
       
  4491       by auto
       
  4492     then show ?thesis
       
  4493       using obt(3)[THEN bspec[where x=y], THEN bspec[where x=v]] and obt(2)
       
  4494       by auto
       
  4495   next
       
  4496     assume "v \<notin> s"
       
  4497     then obtain y where "y \<in> convex hull s" "norm (v - u) < norm (y - u)"
       
  4498       using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2)
       
  4499       by auto
       
  4500     then show ?thesis
       
  4501       using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1)
       
  4502       by (auto simp add: norm_minus_commute)
       
  4503   qed auto
       
  4504 qed
       
  4505 
       
  4506 lemma simplex_extremal_le_exists:
       
  4507   fixes s :: "'a::real_inner set"
       
  4508   shows "finite s \<Longrightarrow> x \<in> convex hull s \<Longrightarrow> y \<in> convex hull s \<Longrightarrow>
       
  4509     \<exists>u\<in>s. \<exists>v\<in>s. norm (x - y) \<le> norm (u - v)"
       
  4510   using convex_hull_empty simplex_extremal_le[of s]
       
  4511   by(cases "s = {}") auto
       
  4512 
       
  4513 
       
  4514 subsection \<open>Closest point of a convex set is unique, with a continuous projection.\<close>
       
  4515 
       
  4516 definition closest_point :: "'a::{real_inner,heine_borel} set \<Rightarrow> 'a \<Rightarrow> 'a"
       
  4517   where "closest_point s a = (SOME x. x \<in> s \<and> (\<forall>y\<in>s. dist a x \<le> dist a y))"
       
  4518 
       
  4519 lemma closest_point_exists:
       
  4520   assumes "closed s"
       
  4521     and "s \<noteq> {}"
       
  4522   shows "closest_point s a \<in> s"
       
  4523     and "\<forall>y\<in>s. dist a (closest_point s a) \<le> dist a y"
       
  4524   unfolding closest_point_def
       
  4525   apply(rule_tac[!] someI2_ex)
       
  4526   apply (auto intro: distance_attains_inf[OF assms(1,2), of a])
       
  4527   done
       
  4528 
       
  4529 lemma closest_point_in_set: "closed s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> closest_point s a \<in> s"
       
  4530   by (meson closest_point_exists)
       
  4531 
       
  4532 lemma closest_point_le: "closed s \<Longrightarrow> x \<in> s \<Longrightarrow> dist a (closest_point s a) \<le> dist a x"
       
  4533   using closest_point_exists[of s] by auto
       
  4534 
       
  4535 lemma closest_point_self:
       
  4536   assumes "x \<in> s"
       
  4537   shows "closest_point s x = x"
       
  4538   unfolding closest_point_def
       
  4539   apply (rule some1_equality, rule ex1I[of _ x])
       
  4540   using assms
       
  4541   apply auto
       
  4542   done
       
  4543 
       
  4544 lemma closest_point_refl: "closed s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> closest_point s x = x \<longleftrightarrow> x \<in> s"
       
  4545   using closest_point_in_set[of s x] closest_point_self[of x s]
       
  4546   by auto
       
  4547 
       
  4548 lemma closer_points_lemma:
       
  4549   assumes "inner y z > 0"
       
  4550   shows "\<exists>u>0. \<forall>v>0. v \<le> u \<longrightarrow> norm(v *\<^sub>R z - y) < norm y"
       
  4551 proof -
       
  4552   have z: "inner z z > 0"
       
  4553     unfolding inner_gt_zero_iff using assms by auto
       
  4554   then show ?thesis
       
  4555     using assms
       
  4556     apply (rule_tac x = "inner y z / inner z z" in exI)
       
  4557     apply rule
       
  4558     defer
       
  4559   proof rule+
       
  4560     fix v
       
  4561     assume "0 < v" and "v \<le> inner y z / inner z z"
       
  4562     then show "norm (v *\<^sub>R z - y) < norm y"
       
  4563       unfolding norm_lt using z and assms
       
  4564       by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ \<open>0<v\<close>])
       
  4565   qed auto
       
  4566 qed
       
  4567 
       
  4568 lemma closer_point_lemma:
       
  4569   assumes "inner (y - x) (z - x) > 0"
       
  4570   shows "\<exists>u>0. u \<le> 1 \<and> dist (x + u *\<^sub>R (z - x)) y < dist x y"
       
  4571 proof -
       
  4572   obtain u where "u > 0"
       
  4573     and u: "\<forall>v>0. v \<le> u \<longrightarrow> norm (v *\<^sub>R (z - x) - (y - x)) < norm (y - x)"
       
  4574     using closer_points_lemma[OF assms] by auto
       
  4575   show ?thesis
       
  4576     apply (rule_tac x="min u 1" in exI)
       
  4577     using u[THEN spec[where x="min u 1"]] and \<open>u > 0\<close>
       
  4578     unfolding dist_norm by (auto simp add: norm_minus_commute field_simps)
       
  4579 qed
       
  4580 
       
  4581 lemma any_closest_point_dot:
       
  4582   assumes "convex s" "closed s" "x \<in> s" "y \<in> s" "\<forall>z\<in>s. dist a x \<le> dist a z"
       
  4583   shows "inner (a - x) (y - x) \<le> 0"
       
  4584 proof (rule ccontr)
       
  4585   assume "\<not> ?thesis"
       
  4586   then obtain u where u: "u>0" "u\<le>1" "dist (x + u *\<^sub>R (y - x)) a < dist x a"
       
  4587     using closer_point_lemma[of a x y] by auto
       
  4588   let ?z = "(1 - u) *\<^sub>R x + u *\<^sub>R y"
       
  4589   have "?z \<in> s"
       
  4590     using convexD_alt[OF assms(1,3,4), of u] using u by auto
       
  4591   then show False
       
  4592     using assms(5)[THEN bspec[where x="?z"]] and u(3)
       
  4593     by (auto simp add: dist_commute algebra_simps)
       
  4594 qed
       
  4595 
       
  4596 lemma any_closest_point_unique:
       
  4597   fixes x :: "'a::real_inner"
       
  4598   assumes "convex s" "closed s" "x \<in> s" "y \<in> s"
       
  4599     "\<forall>z\<in>s. dist a x \<le> dist a z" "\<forall>z\<in>s. dist a y \<le> dist a z"
       
  4600   shows "x = y"
       
  4601   using any_closest_point_dot[OF assms(1-4,5)] and any_closest_point_dot[OF assms(1-2,4,3,6)]
       
  4602   unfolding norm_pths(1) and norm_le_square
       
  4603   by (auto simp add: algebra_simps)
       
  4604 
       
  4605 lemma closest_point_unique:
       
  4606   assumes "convex s" "closed s" "x \<in> s" "\<forall>z\<in>s. dist a x \<le> dist a z"
       
  4607   shows "x = closest_point s a"
       
  4608   using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point s a"]
       
  4609   using closest_point_exists[OF assms(2)] and assms(3) by auto
       
  4610 
       
  4611 lemma closest_point_dot:
       
  4612   assumes "convex s" "closed s" "x \<in> s"
       
  4613   shows "inner (a - closest_point s a) (x - closest_point s a) \<le> 0"
       
  4614   apply (rule any_closest_point_dot[OF assms(1,2) _ assms(3)])
       
  4615   using closest_point_exists[OF assms(2)] and assms(3)
       
  4616   apply auto
       
  4617   done
       
  4618 
       
  4619 lemma closest_point_lt:
       
  4620   assumes "convex s" "closed s" "x \<in> s" "x \<noteq> closest_point s a"
       
  4621   shows "dist a (closest_point s a) < dist a x"
       
  4622   apply (rule ccontr)
       
  4623   apply (rule_tac notE[OF assms(4)])
       
  4624   apply (rule closest_point_unique[OF assms(1-3), of a])
       
  4625   using closest_point_le[OF assms(2), of _ a]
       
  4626   apply fastforce
       
  4627   done
       
  4628 
       
  4629 lemma closest_point_lipschitz:
       
  4630   assumes "convex s"
       
  4631     and "closed s" "s \<noteq> {}"
       
  4632   shows "dist (closest_point s x) (closest_point s y) \<le> dist x y"
       
  4633 proof -
       
  4634   have "inner (x - closest_point s x) (closest_point s y - closest_point s x) \<le> 0"
       
  4635     and "inner (y - closest_point s y) (closest_point s x - closest_point s y) \<le> 0"
       
  4636     apply (rule_tac[!] any_closest_point_dot[OF assms(1-2)])
       
  4637     using closest_point_exists[OF assms(2-3)]
       
  4638     apply auto
       
  4639     done
       
  4640   then show ?thesis unfolding dist_norm and norm_le
       
  4641     using inner_ge_zero[of "(x - closest_point s x) - (y - closest_point s y)"]
       
  4642     by (simp add: inner_add inner_diff inner_commute)
       
  4643 qed
       
  4644 
       
  4645 lemma continuous_at_closest_point:
       
  4646   assumes "convex s"
       
  4647     and "closed s"
       
  4648     and "s \<noteq> {}"
       
  4649   shows "continuous (at x) (closest_point s)"
       
  4650   unfolding continuous_at_eps_delta
       
  4651   using le_less_trans[OF closest_point_lipschitz[OF assms]] by auto
       
  4652 
       
  4653 lemma continuous_on_closest_point:
       
  4654   assumes "convex s"
       
  4655     and "closed s"
       
  4656     and "s \<noteq> {}"
       
  4657   shows "continuous_on t (closest_point s)"
       
  4658   by (metis continuous_at_imp_continuous_on continuous_at_closest_point[OF assms])
       
  4659 
       
  4660 
       
  4661 subsubsection \<open>Various point-to-set separating/supporting hyperplane theorems.\<close>
       
  4662 
       
  4663 lemma supporting_hyperplane_closed_point:
       
  4664   fixes z :: "'a::{real_inner,heine_borel}"
       
  4665   assumes "convex s"
       
  4666     and "closed s"
       
  4667     and "s \<noteq> {}"
       
  4668     and "z \<notin> s"
       
  4669   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)"
       
  4670 proof -
       
  4671   obtain y where "y \<in> s" and y: "\<forall>x\<in>s. dist z y \<le> dist z x"
       
  4672     by (metis distance_attains_inf[OF assms(2-3)])
       
  4673   show ?thesis
       
  4674     apply (rule_tac x="y - z" in exI)
       
  4675     apply (rule_tac x="inner (y - z) y" in exI)
       
  4676     apply (rule_tac x=y in bexI)
       
  4677     apply rule
       
  4678     defer
       
  4679     apply rule
       
  4680     defer
       
  4681     apply rule
       
  4682     apply (rule ccontr)
       
  4683     using \<open>y \<in> s\<close>
       
  4684   proof -
       
  4685     show "inner (y - z) z < inner (y - z) y"
       
  4686       apply (subst diff_gt_0_iff_gt [symmetric])
       
  4687       unfolding inner_diff_right[symmetric] and inner_gt_zero_iff
       
  4688       using \<open>y\<in>s\<close> \<open>z\<notin>s\<close>
       
  4689       apply auto
       
  4690       done
       
  4691   next
       
  4692     fix x
       
  4693     assume "x \<in> s"
       
  4694     have *: "\<forall>u. 0 \<le> u \<and> u \<le> 1 \<longrightarrow> dist z y \<le> dist z ((1 - u) *\<^sub>R y + u *\<^sub>R x)"
       
  4695       using assms(1)[unfolded convex_alt] and y and \<open>x\<in>s\<close> and \<open>y\<in>s\<close> by auto
       
  4696     assume "\<not> inner (y - z) y \<le> inner (y - z) x"
       
  4697     then obtain v where "v > 0" "v \<le> 1" "dist (y + v *\<^sub>R (x - y)) z < dist y z"
       
  4698       using closer_point_lemma[of z y x] by (auto simp add: inner_diff)
       
  4699     then show False
       
  4700       using *[THEN spec[where x=v]] by (auto simp add: dist_commute algebra_simps)
       
  4701   qed auto
       
  4702 qed
       
  4703 
       
  4704 lemma separating_hyperplane_closed_point:
       
  4705   fixes z :: "'a::{real_inner,heine_borel}"
       
  4706   assumes "convex s"
       
  4707     and "closed s"
       
  4708     and "z \<notin> s"
       
  4709   shows "\<exists>a b. inner a z < b \<and> (\<forall>x\<in>s. inner a x > b)"
       
  4710 proof (cases "s = {}")
       
  4711   case True
       
  4712   then show ?thesis
       
  4713     apply (rule_tac x="-z" in exI)
       
  4714     apply (rule_tac x=1 in exI)
       
  4715     using less_le_trans[OF _ inner_ge_zero[of z]]
       
  4716     apply auto
       
  4717     done
       
  4718 next
       
  4719   case False
       
  4720   obtain y where "y \<in> s" and y: "\<forall>x\<in>s. dist z y \<le> dist z x"
       
  4721     by (metis distance_attains_inf[OF assms(2) False])
       
  4722   show ?thesis
       
  4723     apply (rule_tac x="y - z" in exI)
       
  4724     apply (rule_tac x="inner (y - z) z + (norm (y - z))\<^sup>2 / 2" in exI)
       
  4725     apply rule
       
  4726     defer
       
  4727     apply rule
       
  4728   proof -
       
  4729     fix x
       
  4730     assume "x \<in> s"
       
  4731     have "\<not> 0 < inner (z - y) (x - y)"
       
  4732       apply (rule notI)
       
  4733       apply (drule closer_point_lemma)
       
  4734     proof -
       
  4735       assume "\<exists>u>0. u \<le> 1 \<and> dist (y + u *\<^sub>R (x - y)) z < dist y z"
       
  4736       then obtain u where "u > 0" "u \<le> 1" "dist (y + u *\<^sub>R (x - y)) z < dist y z"
       
  4737         by auto
       
  4738       then show False using y[THEN bspec[where x="y + u *\<^sub>R (x - y)"]]
       
  4739         using assms(1)[unfolded convex_alt, THEN bspec[where x=y]]
       
  4740         using \<open>x\<in>s\<close> \<open>y\<in>s\<close> by (auto simp add: dist_commute algebra_simps)
       
  4741     qed
       
  4742     moreover have "0 < (norm (y - z))\<^sup>2"
       
  4743       using \<open>y\<in>s\<close> \<open>z\<notin>s\<close> by auto
       
  4744     then have "0 < inner (y - z) (y - z)"
       
  4745       unfolding power2_norm_eq_inner by simp
       
  4746     ultimately show "inner (y - z) z + (norm (y - z))\<^sup>2 / 2 < inner (y - z) x"
       
  4747       unfolding power2_norm_eq_inner and not_less
       
  4748       by (auto simp add: field_simps inner_commute inner_diff)
       
  4749   qed (insert \<open>y\<in>s\<close> \<open>z\<notin>s\<close>, auto)
       
  4750 qed
       
  4751 
       
  4752 lemma separating_hyperplane_closed_0:
       
  4753   assumes "convex (s::('a::euclidean_space) set)"
       
  4754     and "closed s"
       
  4755     and "0 \<notin> s"
       
  4756   shows "\<exists>a b. a \<noteq> 0 \<and> 0 < b \<and> (\<forall>x\<in>s. inner a x > b)"
       
  4757 proof (cases "s = {}")
       
  4758   case True
       
  4759   have "norm ((SOME i. i\<in>Basis)::'a) = 1" "(SOME i. i\<in>Basis) \<noteq> (0::'a)"
       
  4760     defer
       
  4761     apply (subst norm_le_zero_iff[symmetric])
       
  4762     apply (auto simp: SOME_Basis)
       
  4763     done
       
  4764   then show ?thesis
       
  4765     apply (rule_tac x="SOME i. i\<in>Basis" in exI)
       
  4766     apply (rule_tac x=1 in exI)
       
  4767     using True using DIM_positive[where 'a='a]
       
  4768     apply auto
       
  4769     done
       
  4770 next
       
  4771   case False
       
  4772   then show ?thesis
       
  4773     using False using separating_hyperplane_closed_point[OF assms]
       
  4774     apply (elim exE)
       
  4775     unfolding inner_zero_right
       
  4776     apply (rule_tac x=a in exI)
       
  4777     apply (rule_tac x=b in exI)
       
  4778     apply auto
       
  4779     done
       
  4780 qed
       
  4781 
       
  4782 
       
  4783 subsubsection \<open>Now set-to-set for closed/compact sets\<close>
       
  4784 
       
  4785 lemma separating_hyperplane_closed_compact:
       
  4786   fixes s :: "'a::euclidean_space set"
       
  4787   assumes "convex s"
       
  4788     and "closed s"
       
  4789     and "convex t"
       
  4790     and "compact t"
       
  4791     and "t \<noteq> {}"
       
  4792     and "s \<inter> t = {}"
       
  4793   shows "\<exists>a b. (\<forall>x\<in>s. inner a x < b) \<and> (\<forall>x\<in>t. inner a x > b)"
       
  4794 proof (cases "s = {}")
       
  4795   case True
       
  4796   obtain b where b: "b > 0" "\<forall>x\<in>t. norm x \<le> b"
       
  4797     using compact_imp_bounded[OF assms(4)] unfolding bounded_pos by auto
       
  4798   obtain z :: 'a where z: "norm z = b + 1"
       
  4799     using vector_choose_size[of "b + 1"] and b(1) by auto
       
  4800   then have "z \<notin> t" using b(2)[THEN bspec[where x=z]] by auto
       
  4801   then obtain a b where ab: "inner a z < b" "\<forall>x\<in>t. b < inner a x"
       
  4802     using separating_hyperplane_closed_point[OF assms(3) compact_imp_closed[OF assms(4)], of z]
       
  4803     by auto
       
  4804   then show ?thesis
       
  4805     using True by auto
       
  4806 next
       
  4807   case False
       
  4808   then obtain y where "y \<in> s" by auto
       
  4809   obtain a b where "0 < b" "\<forall>x\<in>{x - y |x y. x \<in> s \<and> y \<in> t}. b < inner a x"
       
  4810     using separating_hyperplane_closed_point[OF convex_differences[OF assms(1,3)], of 0]
       
  4811     using closed_compact_differences[OF assms(2,4)]
       
  4812     using assms(6) by auto blast
       
  4813   then have ab: "\<forall>x\<in>s. \<forall>y\<in>t. b + inner a y < inner a x"
       
  4814     apply -
       
  4815     apply rule
       
  4816     apply rule
       
  4817     apply (erule_tac x="x - y" in ballE)
       
  4818     apply (auto simp add: inner_diff)
       
  4819     done
       
  4820   define k where "k = (SUP x:t. a \<bullet> x)"
       
  4821   show ?thesis
       
  4822     apply (rule_tac x="-a" in exI)
       
  4823     apply (rule_tac x="-(k + b / 2)" in exI)
       
  4824     apply (intro conjI ballI)
       
  4825     unfolding inner_minus_left and neg_less_iff_less
       
  4826   proof -
       
  4827     fix x assume "x \<in> t"
       
  4828     then have "inner a x - b / 2 < k"
       
  4829       unfolding k_def
       
  4830     proof (subst less_cSUP_iff)
       
  4831       show "t \<noteq> {}" by fact
       
  4832       show "bdd_above (op \<bullet> a ` t)"
       
  4833         using ab[rule_format, of y] \<open>y \<in> s\<close>
       
  4834         by (intro bdd_aboveI2[where M="inner a y - b"]) (auto simp: field_simps intro: less_imp_le)
       
  4835     qed (auto intro!: bexI[of _ x] \<open>0<b\<close>)
       
  4836     then show "inner a x < k + b / 2"
       
  4837       by auto
       
  4838   next
       
  4839     fix x
       
  4840     assume "x \<in> s"
       
  4841     then have "k \<le> inner a x - b"
       
  4842       unfolding k_def
       
  4843       apply (rule_tac cSUP_least)
       
  4844       using assms(5)
       
  4845       using ab[THEN bspec[where x=x]]
       
  4846       apply auto
       
  4847       done
       
  4848     then show "k + b / 2 < inner a x"
       
  4849       using \<open>0 < b\<close> by auto
       
  4850   qed
       
  4851 qed
       
  4852 
       
  4853 lemma separating_hyperplane_compact_closed:
       
  4854   fixes s :: "'a::euclidean_space set"
       
  4855   assumes "convex s"
       
  4856     and "compact s"
       
  4857     and "s \<noteq> {}"
       
  4858     and "convex t"
       
  4859     and "closed t"
       
  4860     and "s \<inter> t = {}"
       
  4861   shows "\<exists>a b. (\<forall>x\<in>s. inner a x < b) \<and> (\<forall>x\<in>t. inner a x > b)"
       
  4862 proof -
       
  4863   obtain a b where "(\<forall>x\<in>t. inner a x < b) \<and> (\<forall>x\<in>s. b < inner a x)"
       
  4864     using separating_hyperplane_closed_compact[OF assms(4-5,1-2,3)] and assms(6)
       
  4865     by auto
       
  4866   then show ?thesis
       
  4867     apply (rule_tac x="-a" in exI)
       
  4868     apply (rule_tac x="-b" in exI)
       
  4869     apply auto
       
  4870     done
       
  4871 qed
       
  4872 
       
  4873 
       
  4874 subsubsection \<open>General case without assuming closure and getting non-strict separation\<close>
       
  4875 
       
  4876 lemma separating_hyperplane_set_0:
       
  4877   assumes "convex s" "(0::'a::euclidean_space) \<notin> s"
       
  4878   shows "\<exists>a. a \<noteq> 0 \<and> (\<forall>x\<in>s. 0 \<le> inner a x)"
       
  4879 proof -
       
  4880   let ?k = "\<lambda>c. {x::'a. 0 \<le> inner c x}"
       
  4881   have *: "frontier (cball 0 1) \<inter> \<Inter>f \<noteq> {}" if as: "f \<subseteq> ?k ` s" "finite f" for f
       
  4882   proof -
       
  4883     obtain c where c: "f = ?k ` c" "c \<subseteq> s" "finite c"
       
  4884       using finite_subset_image[OF as(2,1)] by auto
       
  4885     then obtain a b where ab: "a \<noteq> 0" "0 < b" "\<forall>x\<in>convex hull c. b < inner a x"
       
  4886       using separating_hyperplane_closed_0[OF convex_convex_hull, of c]
       
  4887       using finite_imp_compact_convex_hull[OF c(3), THEN compact_imp_closed] and assms(2)
       
  4888       using subset_hull[of convex, OF assms(1), symmetric, of c]
       
  4889       by force
       
  4890     then have "\<exists>x. norm x = 1 \<and> (\<forall>y\<in>c. 0 \<le> inner y x)"
       
  4891       apply (rule_tac x = "inverse(norm a) *\<^sub>R a" in exI)
       
  4892       using hull_subset[of c convex]
       
  4893       unfolding subset_eq and inner_scaleR
       
  4894       by (auto simp add: inner_commute del: ballE elim!: ballE)
       
  4895     then show "frontier (cball 0 1) \<inter> \<Inter>f \<noteq> {}"
       
  4896       unfolding c(1) frontier_cball sphere_def dist_norm by auto
       
  4897   qed
       
  4898   have "frontier (cball 0 1) \<inter> (\<Inter>(?k ` s)) \<noteq> {}"
       
  4899     apply (rule compact_imp_fip)
       
  4900     apply (rule compact_frontier[OF compact_cball])
       
  4901     using * closed_halfspace_ge
       
  4902     by auto
       
  4903   then obtain x where "norm x = 1" "\<forall>y\<in>s. x\<in>?k y"
       
  4904     unfolding frontier_cball dist_norm sphere_def by auto
       
  4905   then show ?thesis
       
  4906     by (metis inner_commute mem_Collect_eq norm_eq_zero zero_neq_one)
       
  4907 qed
       
  4908 
       
  4909 lemma separating_hyperplane_sets:
       
  4910   fixes s t :: "'a::euclidean_space set"
       
  4911   assumes "convex s"
       
  4912     and "convex t"
       
  4913     and "s \<noteq> {}"
       
  4914     and "t \<noteq> {}"
       
  4915     and "s \<inter> t = {}"
       
  4916   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)"
       
  4917 proof -
       
  4918   from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]]
       
  4919   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"
       
  4920     using assms(3-5) by fastforce
       
  4921   then have *: "\<And>x y. x \<in> t \<Longrightarrow> y \<in> s \<Longrightarrow> inner a y \<le> inner a x"
       
  4922     by (force simp add: inner_diff)
       
  4923   then have bdd: "bdd_above ((op \<bullet> a)`s)"
       
  4924     using \<open>t \<noteq> {}\<close> by (auto intro: bdd_aboveI2[OF *])
       
  4925   show ?thesis
       
  4926     using \<open>a\<noteq>0\<close>
       
  4927     by (intro exI[of _ a] exI[of _ "SUP x:s. a \<bullet> x"])
       
  4928        (auto intro!: cSUP_upper bdd cSUP_least \<open>a \<noteq> 0\<close> \<open>s \<noteq> {}\<close> *)
       
  4929 qed
       
  4930 
       
  4931 
       
  4932 subsection \<open>More convexity generalities\<close>
       
  4933 
       
  4934 lemma convex_closure [intro,simp]:
       
  4935   fixes s :: "'a::real_normed_vector set"
       
  4936   assumes "convex s"
       
  4937   shows "convex (closure s)"
       
  4938   apply (rule convexI)
       
  4939   apply (unfold closure_sequential, elim exE)
       
  4940   apply (rule_tac x="\<lambda>n. u *\<^sub>R xa n + v *\<^sub>R xb n" in exI)
       
  4941   apply (rule,rule)
       
  4942   apply (rule convexD [OF assms])
       
  4943   apply (auto del: tendsto_const intro!: tendsto_intros)
       
  4944   done
       
  4945 
       
  4946 lemma convex_interior [intro,simp]:
       
  4947   fixes s :: "'a::real_normed_vector set"
       
  4948   assumes "convex s"
       
  4949   shows "convex (interior s)"
       
  4950   unfolding convex_alt Ball_def mem_interior
       
  4951   apply (rule,rule,rule,rule,rule,rule)
       
  4952   apply (elim exE conjE)
       
  4953 proof -
       
  4954   fix x y u
       
  4955   assume u: "0 \<le> u" "u \<le> (1::real)"
       
  4956   fix e d
       
  4957   assume ed: "ball x e \<subseteq> s" "ball y d \<subseteq> s" "0<d" "0<e"
       
  4958   show "\<exists>e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \<subseteq> s"
       
  4959     apply (rule_tac x="min d e" in exI)
       
  4960     apply rule
       
  4961     unfolding subset_eq
       
  4962     defer
       
  4963     apply rule
       
  4964   proof -
       
  4965     fix z
       
  4966     assume "z \<in> ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) (min d e)"
       
  4967     then have "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \<in> s"
       
  4968       apply (rule_tac assms[unfolded convex_alt, rule_format])
       
  4969       using ed(1,2) and u
       
  4970       unfolding subset_eq mem_ball Ball_def dist_norm
       
  4971       apply (auto simp add: algebra_simps)
       
  4972       done
       
  4973     then show "z \<in> s"
       
  4974       using u by (auto simp add: algebra_simps)
       
  4975   qed(insert u ed(3-4), auto)
       
  4976 qed
       
  4977 
       
  4978 lemma convex_hull_eq_empty[simp]: "convex hull s = {} \<longleftrightarrow> s = {}"
       
  4979   using hull_subset[of s convex] convex_hull_empty by auto
       
  4980 
       
  4981 
       
  4982 subsection \<open>Moving and scaling convex hulls.\<close>
       
  4983 
       
  4984 lemma convex_hull_set_plus:
       
  4985   "convex hull (s + t) = convex hull s + convex hull t"
       
  4986   unfolding set_plus_image
       
  4987   apply (subst convex_hull_linear_image [symmetric])
       
  4988   apply (simp add: linear_iff scaleR_right_distrib)
       
  4989   apply (simp add: convex_hull_Times)
       
  4990   done
       
  4991 
       
  4992 lemma translation_eq_singleton_plus: "(\<lambda>x. a + x) ` t = {a} + t"
       
  4993   unfolding set_plus_def by auto
       
  4994 
       
  4995 lemma convex_hull_translation:
       
  4996   "convex hull ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (convex hull s)"
       
  4997   unfolding translation_eq_singleton_plus
       
  4998   by (simp only: convex_hull_set_plus convex_hull_singleton)
       
  4999 
       
  5000 lemma convex_hull_scaling:
       
  5001   "convex hull ((\<lambda>x. c *\<^sub>R x) ` s) = (\<lambda>x. c *\<^sub>R x) ` (convex hull s)"
       
  5002   using linear_scaleR by (rule convex_hull_linear_image [symmetric])
       
  5003 
       
  5004 lemma convex_hull_affinity:
       
  5005   "convex hull ((\<lambda>x. a + c *\<^sub>R x) ` s) = (\<lambda>x. a + c *\<^sub>R x) ` (convex hull s)"
       
  5006   by(simp only: image_image[symmetric] convex_hull_scaling convex_hull_translation)
       
  5007 
       
  5008 
       
  5009 subsection \<open>Convexity of cone hulls\<close>
       
  5010 
       
  5011 lemma convex_cone_hull:
       
  5012   assumes "convex S"
       
  5013   shows "convex (cone hull S)"
       
  5014 proof (rule convexI)
       
  5015   fix x y
       
  5016   assume xy: "x \<in> cone hull S" "y \<in> cone hull S"
       
  5017   then have "S \<noteq> {}"
       
  5018     using cone_hull_empty_iff[of S] by auto
       
  5019   fix u v :: real
       
  5020   assume uv: "u \<ge> 0" "v \<ge> 0" "u + v = 1"
       
  5021   then have *: "u *\<^sub>R x \<in> cone hull S" "v *\<^sub>R y \<in> cone hull S"
       
  5022     using cone_cone_hull[of S] xy cone_def[of "cone hull S"] by auto
       
  5023   from * obtain cx :: real and xx where x: "u *\<^sub>R x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S"
       
  5024     using cone_hull_expl[of S] by auto
       
  5025   from * obtain cy :: real and yy where y: "v *\<^sub>R y = cy *\<^sub>R yy" "cy \<ge> 0" "yy \<in> S"
       
  5026     using cone_hull_expl[of S] by auto
       
  5027   {
       
  5028     assume "cx + cy \<le> 0"
       
  5029     then have "u *\<^sub>R x = 0" and "v *\<^sub>R y = 0"
       
  5030       using x y by auto
       
  5031     then have "u *\<^sub>R x + v *\<^sub>R y = 0"
       
  5032       by auto
       
  5033     then have "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S"
       
  5034       using cone_hull_contains_0[of S] \<open>S \<noteq> {}\<close> by auto
       
  5035   }
       
  5036   moreover
       
  5037   {
       
  5038     assume "cx + cy > 0"
       
  5039     then have "(cx / (cx + cy)) *\<^sub>R xx + (cy / (cx + cy)) *\<^sub>R yy \<in> S"
       
  5040       using assms mem_convex_alt[of S xx yy cx cy] x y by auto
       
  5041     then have "cx *\<^sub>R xx + cy *\<^sub>R yy \<in> cone hull S"
       
  5042       using mem_cone_hull[of "(cx/(cx+cy)) *\<^sub>R xx + (cy/(cx+cy)) *\<^sub>R yy" S "cx+cy"] \<open>cx+cy>0\<close>
       
  5043       by (auto simp add: scaleR_right_distrib)
       
  5044     then have "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S"
       
  5045       using x y by auto
       
  5046   }
       
  5047   moreover have "cx + cy \<le> 0 \<or> cx + cy > 0" by auto
       
  5048   ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S" by blast
       
  5049 qed
       
  5050 
       
  5051 lemma cone_convex_hull:
       
  5052   assumes "cone S"
       
  5053   shows "cone (convex hull S)"
       
  5054 proof (cases "S = {}")
       
  5055   case True
       
  5056   then show ?thesis by auto
       
  5057 next
       
  5058   case False
       
  5059   then have *: "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` S = S)"
       
  5060     using cone_iff[of S] assms by auto
       
  5061   {
       
  5062     fix c :: real
       
  5063     assume "c > 0"
       
  5064     then have "op *\<^sub>R c ` (convex hull S) = convex hull (op *\<^sub>R c ` S)"
       
  5065       using convex_hull_scaling[of _ S] by auto
       
  5066     also have "\<dots> = convex hull S"
       
  5067       using * \<open>c > 0\<close> by auto
       
  5068     finally have "op *\<^sub>R c ` (convex hull S) = convex hull S"
       
  5069       by auto
       
  5070   }
       
  5071   then have "0 \<in> convex hull S" "\<And>c. c > 0 \<Longrightarrow> (op *\<^sub>R c ` (convex hull S)) = (convex hull S)"
       
  5072     using * hull_subset[of S convex] by auto
       
  5073   then show ?thesis
       
  5074     using \<open>S \<noteq> {}\<close> cone_iff[of "convex hull S"] by auto
       
  5075 qed
       
  5076 
       
  5077 subsection \<open>Convex set as intersection of halfspaces\<close>
       
  5078 
       
  5079 lemma convex_halfspace_intersection:
       
  5080   fixes s :: "('a::euclidean_space) set"
       
  5081   assumes "closed s" "convex s"
       
  5082   shows "s = \<Inter>{h. s \<subseteq> h \<and> (\<exists>a b. h = {x. inner a x \<le> b})}"
       
  5083   apply (rule set_eqI)
       
  5084   apply rule
       
  5085   unfolding Inter_iff Ball_def mem_Collect_eq
       
  5086   apply (rule,rule,erule conjE)
       
  5087 proof -
       
  5088   fix x
       
  5089   assume "\<forall>xa. s \<subseteq> xa \<and> (\<exists>a b. xa = {x. inner a x \<le> b}) \<longrightarrow> x \<in> xa"
       
  5090   then have "\<forall>a b. s \<subseteq> {x. inner a x \<le> b} \<longrightarrow> x \<in> {x. inner a x \<le> b}"
       
  5091     by blast
       
  5092   then show "x \<in> s"
       
  5093     apply (rule_tac ccontr)
       
  5094     apply (drule separating_hyperplane_closed_point[OF assms(2,1)])
       
  5095     apply (erule exE)+
       
  5096     apply (erule_tac x="-a" in allE)
       
  5097     apply (erule_tac x="-b" in allE)
       
  5098     apply auto
       
  5099     done
       
  5100 qed auto
       
  5101 
       
  5102 
       
  5103 subsection \<open>Radon's theorem (from Lars Schewe)\<close>
       
  5104 
       
  5105 lemma radon_ex_lemma:
       
  5106   assumes "finite c" "affine_dependent c"
       
  5107   shows "\<exists>u. setsum u c = 0 \<and> (\<exists>v\<in>c. u v \<noteq> 0) \<and> setsum (\<lambda>v. u v *\<^sub>R v) c = 0"
       
  5108 proof -
       
  5109   from assms(2)[unfolded affine_dependent_explicit]
       
  5110   obtain s u where
       
  5111       "finite s" "s \<subseteq> c" "setsum u s = 0" "\<exists>v\<in>s. u v \<noteq> 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
       
  5112     by blast
       
  5113   then show ?thesis
       
  5114     apply (rule_tac x="\<lambda>v. if v\<in>s then u v else 0" in exI)
       
  5115     unfolding if_smult scaleR_zero_left and setsum.inter_restrict[OF assms(1), symmetric]
       
  5116     apply (auto simp add: Int_absorb1)
       
  5117     done
       
  5118 qed
       
  5119 
       
  5120 lemma radon_s_lemma:
       
  5121   assumes "finite s"
       
  5122     and "setsum f s = (0::real)"
       
  5123   shows "setsum f {x\<in>s. 0 < f x} = - setsum f {x\<in>s. f x < 0}"
       
  5124 proof -
       
  5125   have *: "\<And>x. (if f x < 0 then f x else 0) + (if 0 < f x then f x else 0) = f x"
       
  5126     by auto
       
  5127   show ?thesis
       
  5128     unfolding add_eq_0_iff[symmetric] and setsum.inter_filter[OF assms(1)]
       
  5129       and setsum.distrib[symmetric] and *
       
  5130     using assms(2)
       
  5131     by assumption
       
  5132 qed
       
  5133 
       
  5134 lemma radon_v_lemma:
       
  5135   assumes "finite s"
       
  5136     and "setsum f s = 0"
       
  5137     and "\<forall>x. g x = (0::real) \<longrightarrow> f x = (0::'a::euclidean_space)"
       
  5138   shows "(setsum f {x\<in>s. 0 < g x}) = - setsum f {x\<in>s. g x < 0}"
       
  5139 proof -
       
  5140   have *: "\<And>x. (if 0 < g x then f x else 0) + (if g x < 0 then f x else 0) = f x"
       
  5141     using assms(3) by auto
       
  5142   show ?thesis
       
  5143     unfolding eq_neg_iff_add_eq_0 and setsum.inter_filter[OF assms(1)]
       
  5144       and setsum.distrib[symmetric] and *
       
  5145     using assms(2)
       
  5146     apply assumption
       
  5147     done
       
  5148 qed
       
  5149 
       
  5150 lemma radon_partition:
       
  5151   assumes "finite c" "affine_dependent c"
       
  5152   shows "\<exists>m p. m \<inter> p = {} \<and> m \<union> p = c \<and> (convex hull m) \<inter> (convex hull p) \<noteq> {}"
       
  5153 proof -
       
  5154   obtain u v where uv: "setsum u c = 0" "v\<in>c" "u v \<noteq> 0"  "(\<Sum>v\<in>c. u v *\<^sub>R v) = 0"
       
  5155     using radon_ex_lemma[OF assms] by auto
       
  5156   have fin: "finite {x \<in> c. 0 < u x}" "finite {x \<in> c. 0 > u x}"
       
  5157     using assms(1) by auto
       
  5158   define z  where "z = inverse (setsum u {x\<in>c. u x > 0}) *\<^sub>R setsum (\<lambda>x. u x *\<^sub>R x) {x\<in>c. u x > 0}"
       
  5159   have "setsum u {x \<in> c. 0 < u x} \<noteq> 0"
       
  5160   proof (cases "u v \<ge> 0")
       
  5161     case False
       
  5162     then have "u v < 0" by auto
       
  5163     then show ?thesis
       
  5164     proof (cases "\<exists>w\<in>{x \<in> c. 0 < u x}. u w > 0")
       
  5165       case True
       
  5166       then show ?thesis
       
  5167         using setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] by auto
       
  5168     next
       
  5169       case False
       
  5170       then have "setsum u c \<le> setsum (\<lambda>x. if x=v then u v else 0) c"
       
  5171         apply (rule_tac setsum_mono)
       
  5172         apply auto
       
  5173         done
       
  5174       then show ?thesis
       
  5175         unfolding setsum.delta[OF assms(1)] using uv(2) and \<open>u v < 0\<close> and uv(1) by auto
       
  5176     qed
       
  5177   qed (insert setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] uv(2-3), auto)
       
  5178 
       
  5179   then have *: "setsum u {x\<in>c. u x > 0} > 0"
       
  5180     unfolding less_le
       
  5181     apply (rule_tac conjI)
       
  5182     apply (rule_tac setsum_nonneg)
       
  5183     apply auto
       
  5184     done
       
  5185   moreover have "setsum u ({x \<in> c. 0 < u x} \<union> {x \<in> c. u x < 0}) = setsum u c"
       
  5186     "(\<Sum>x\<in>{x \<in> c. 0 < u x} \<union> {x \<in> c. u x < 0}. u x *\<^sub>R x) = (\<Sum>x\<in>c. u x *\<^sub>R x)"
       
  5187     using assms(1)
       
  5188     apply (rule_tac[!] setsum.mono_neutral_left)
       
  5189     apply auto
       
  5190     done
       
  5191   then have "setsum u {x \<in> c. 0 < u x} = - setsum u {x \<in> c. 0 > u x}"
       
  5192     "(\<Sum>x\<in>{x \<in> c. 0 < u x}. u x *\<^sub>R x) = - (\<Sum>x\<in>{x \<in> c. 0 > u x}. u x *\<^sub>R x)"
       
  5193     unfolding eq_neg_iff_add_eq_0
       
  5194     using uv(1,4)
       
  5195     by (auto simp add: setsum.union_inter_neutral[OF fin, symmetric])
       
  5196   moreover have "\<forall>x\<in>{v \<in> c. u v < 0}. 0 \<le> inverse (setsum u {x \<in> c. 0 < u x}) * - u x"
       
  5197     apply rule
       
  5198     apply (rule mult_nonneg_nonneg)
       
  5199     using *
       
  5200     apply auto
       
  5201     done
       
  5202   ultimately have "z \<in> convex hull {v \<in> c. u v \<le> 0}"
       
  5203     unfolding convex_hull_explicit mem_Collect_eq
       
  5204     apply (rule_tac x="{v \<in> c. u v < 0}" in exI)
       
  5205     apply (rule_tac x="\<lambda>y. inverse (setsum u {x\<in>c. u x > 0}) * - u y" in exI)
       
  5206     using assms(1) unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] and z_def
       
  5207     apply (auto simp add: setsum_negf setsum_right_distrib[symmetric])
       
  5208     done
       
  5209   moreover have "\<forall>x\<in>{v \<in> c. 0 < u v}. 0 \<le> inverse (setsum u {x \<in> c. 0 < u x}) * u x"
       
  5210     apply rule
       
  5211     apply (rule mult_nonneg_nonneg)
       
  5212     using *
       
  5213     apply auto
       
  5214     done
       
  5215   then have "z \<in> convex hull {v \<in> c. u v > 0}"
       
  5216     unfolding convex_hull_explicit mem_Collect_eq
       
  5217     apply (rule_tac x="{v \<in> c. 0 < u v}" in exI)
       
  5218     apply (rule_tac x="\<lambda>y. inverse (setsum u {x\<in>c. u x > 0}) * u y" in exI)
       
  5219     using assms(1)
       
  5220     unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] and z_def
       
  5221     using *
       
  5222     apply (auto simp add: setsum_negf setsum_right_distrib[symmetric])
       
  5223     done
       
  5224   ultimately show ?thesis
       
  5225     apply (rule_tac x="{v\<in>c. u v \<le> 0}" in exI)
       
  5226     apply (rule_tac x="{v\<in>c. u v > 0}" in exI)
       
  5227     apply auto
       
  5228     done
       
  5229 qed
       
  5230 
       
  5231 lemma radon:
       
  5232   assumes "affine_dependent c"
       
  5233   obtains m p where "m \<subseteq> c" "p \<subseteq> c" "m \<inter> p = {}" "(convex hull m) \<inter> (convex hull p) \<noteq> {}"
       
  5234 proof -
       
  5235   from assms[unfolded affine_dependent_explicit]
       
  5236   obtain s u where
       
  5237       "finite s" "s \<subseteq> c" "setsum u s = 0" "\<exists>v\<in>s. u v \<noteq> 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
       
  5238     by blast
       
  5239   then have *: "finite s" "affine_dependent s" and s: "s \<subseteq> c"
       
  5240     unfolding affine_dependent_explicit by auto
       
  5241   from radon_partition[OF *]
       
  5242   obtain m p where "m \<inter> p = {}" "m \<union> p = s" "convex hull m \<inter> convex hull p \<noteq> {}"
       
  5243     by blast
       
  5244   then show ?thesis
       
  5245     apply (rule_tac that[of p m])
       
  5246     using s
       
  5247     apply auto
       
  5248     done
       
  5249 qed
       
  5250 
       
  5251 
       
  5252 subsection \<open>Helly's theorem\<close>
       
  5253 
       
  5254 lemma helly_induct:
       
  5255   fixes f :: "'a::euclidean_space set set"
       
  5256   assumes "card f = n"
       
  5257     and "n \<ge> DIM('a) + 1"
       
  5258     and "\<forall>s\<in>f. convex s" "\<forall>t\<subseteq>f. card t = DIM('a) + 1 \<longrightarrow> \<Inter>t \<noteq> {}"
       
  5259   shows "\<Inter>f \<noteq> {}"
       
  5260   using assms
       
  5261 proof (induct n arbitrary: f)
       
  5262   case 0
       
  5263   then show ?case by auto
       
  5264 next
       
  5265   case (Suc n)
       
  5266   have "finite f"
       
  5267     using \<open>card f = Suc n\<close> by (auto intro: card_ge_0_finite)
       
  5268   show "\<Inter>f \<noteq> {}"
       
  5269     apply (cases "n = DIM('a)")
       
  5270     apply (rule Suc(5)[rule_format])
       
  5271     unfolding \<open>card f = Suc n\<close>
       
  5272   proof -
       
  5273     assume ng: "n \<noteq> DIM('a)"
       
  5274     then have "\<exists>X. \<forall>s\<in>f. X s \<in> \<Inter>(f - {s})"
       
  5275       apply (rule_tac bchoice)
       
  5276       unfolding ex_in_conv
       
  5277       apply (rule, rule Suc(1)[rule_format])
       
  5278       unfolding card_Diff_singleton_if[OF \<open>finite f\<close>] \<open>card f = Suc n\<close>
       
  5279       defer
       
  5280       defer
       
  5281       apply (rule Suc(4)[rule_format])
       
  5282       defer
       
  5283       apply (rule Suc(5)[rule_format])
       
  5284       using Suc(3) \<open>finite f\<close>
       
  5285       apply auto
       
  5286       done
       
  5287     then obtain X where X: "\<forall>s\<in>f. X s \<in> \<Inter>(f - {s})" by auto
       
  5288     show ?thesis
       
  5289     proof (cases "inj_on X f")
       
  5290       case False
       
  5291       then obtain s t where st: "s\<noteq>t" "s\<in>f" "t\<in>f" "X s = X t"
       
  5292         unfolding inj_on_def by auto
       
  5293       then have *: "\<Inter>f = \<Inter>(f - {s}) \<inter> \<Inter>(f - {t})" by auto
       
  5294       show ?thesis
       
  5295         unfolding *
       
  5296         unfolding ex_in_conv[symmetric]
       
  5297         apply (rule_tac x="X s" in exI)
       
  5298         apply rule
       
  5299         apply (rule X[rule_format])
       
  5300         using X st
       
  5301         apply auto
       
  5302         done
       
  5303     next
       
  5304       case True
       
  5305       then obtain m p where mp: "m \<inter> p = {}" "m \<union> p = X ` f" "convex hull m \<inter> convex hull p \<noteq> {}"
       
  5306         using radon_partition[of "X ` f"] and affine_dependent_biggerset[of "X ` f"]
       
  5307         unfolding card_image[OF True] and \<open>card f = Suc n\<close>
       
  5308         using Suc(3) \<open>finite f\<close> and ng
       
  5309         by auto
       
  5310       have "m \<subseteq> X ` f" "p \<subseteq> X ` f"
       
  5311         using mp(2) by auto
       
  5312       then obtain g h where gh:"m = X ` g" "p = X ` h" "g \<subseteq> f" "h \<subseteq> f"
       
  5313         unfolding subset_image_iff by auto
       
  5314       then have "f \<union> (g \<union> h) = f" by auto
       
  5315       then have f: "f = g \<union> h"
       
  5316         using inj_on_Un_image_eq_iff[of X f "g \<union> h"] and True
       
  5317         unfolding mp(2)[unfolded image_Un[symmetric] gh]
       
  5318         by auto
       
  5319       have *: "g \<inter> h = {}"
       
  5320         using mp(1)
       
  5321         unfolding gh
       
  5322         using inj_on_image_Int[OF True gh(3,4)]
       
  5323         by auto
       
  5324       have "convex hull (X ` h) \<subseteq> \<Inter>g" "convex hull (X ` g) \<subseteq> \<Inter>h"
       
  5325         apply (rule_tac [!] hull_minimal)
       
  5326         using Suc gh(3-4)
       
  5327         unfolding subset_eq
       
  5328         apply (rule_tac [2] convex_Inter, rule_tac [4] convex_Inter)
       
  5329         apply rule
       
  5330         prefer 3
       
  5331         apply rule
       
  5332       proof -
       
  5333         fix x
       
  5334         assume "x \<in> X ` g"
       
  5335         then obtain y where "y \<in> g" "x = X y"
       
  5336           unfolding image_iff ..
       
  5337         then show "x \<in> \<Inter>h"
       
  5338           using X[THEN bspec[where x=y]] using * f by auto
       
  5339       next
       
  5340         fix x
       
  5341         assume "x \<in> X ` h"
       
  5342         then obtain y where "y \<in> h" "x = X y"
       
  5343           unfolding image_iff ..
       
  5344         then show "x \<in> \<Inter>g"
       
  5345           using X[THEN bspec[where x=y]] using * f by auto
       
  5346       qed auto
       
  5347       then show ?thesis
       
  5348         unfolding f using mp(3)[unfolded gh] by blast
       
  5349     qed
       
  5350   qed auto
       
  5351 qed
       
  5352 
       
  5353 lemma helly:
       
  5354   fixes f :: "'a::euclidean_space set set"
       
  5355   assumes "card f \<ge> DIM('a) + 1" "\<forall>s\<in>f. convex s"
       
  5356     and "\<forall>t\<subseteq>f. card t = DIM('a) + 1 \<longrightarrow> \<Inter>t \<noteq> {}"
       
  5357   shows "\<Inter>f \<noteq> {}"
       
  5358   apply (rule helly_induct)
       
  5359   using assms
       
  5360   apply auto
       
  5361   done
       
  5362 
       
  5363 
       
  5364 subsection \<open>Epigraphs of convex functions\<close>
       
  5365 
       
  5366 definition "epigraph s (f :: _ \<Rightarrow> real) = {xy. fst xy \<in> s \<and> f (fst xy) \<le> snd xy}"
       
  5367 
       
  5368 lemma mem_epigraph: "(x, y) \<in> epigraph s f \<longleftrightarrow> x \<in> s \<and> f x \<le> y"
       
  5369   unfolding epigraph_def by auto
       
  5370 
       
  5371 lemma convex_epigraph: "convex (epigraph s f) \<longleftrightarrow> convex_on s f \<and> convex s"
       
  5372   unfolding convex_def convex_on_def
       
  5373   unfolding Ball_def split_paired_All epigraph_def
       
  5374   unfolding mem_Collect_eq fst_conv snd_conv fst_add snd_add fst_scaleR snd_scaleR Ball_def[symmetric]
       
  5375   apply safe
       
  5376   defer
       
  5377   apply (erule_tac x=x in allE)
       
  5378   apply (erule_tac x="f x" in allE)
       
  5379   apply safe
       
  5380   apply (erule_tac x=xa in allE)
       
  5381   apply (erule_tac x="f xa" in allE)
       
  5382   prefer 3
       
  5383   apply (rule_tac y="u * f a + v * f aa" in order_trans)
       
  5384   defer
       
  5385   apply (auto intro!:mult_left_mono add_mono)
       
  5386   done
       
  5387 
       
  5388 lemma convex_epigraphI: "convex_on s f \<Longrightarrow> convex s \<Longrightarrow> convex (epigraph s f)"
       
  5389   unfolding convex_epigraph by auto
       
  5390 
       
  5391 lemma convex_epigraph_convex: "convex s \<Longrightarrow> convex_on s f \<longleftrightarrow> convex(epigraph s f)"
       
  5392   by (simp add: convex_epigraph)
       
  5393 
       
  5394 
       
  5395 subsubsection \<open>Use this to derive general bound property of convex function\<close>
       
  5396 
       
  5397 lemma convex_on:
       
  5398   assumes "convex s"
       
  5399   shows "convex_on s f \<longleftrightarrow>
       
  5400     (\<forall>k u x. (\<forall>i\<in>{1..k::nat}. 0 \<le> u i \<and> x i \<in> s) \<and> setsum u {1..k} = 1 \<longrightarrow>
       
  5401       f (setsum (\<lambda>i. u i *\<^sub>R x i) {1..k} ) \<le> setsum (\<lambda>i. u i * f(x i)) {1..k})"
       
  5402   unfolding convex_epigraph_convex[OF assms] convex epigraph_def Ball_def mem_Collect_eq
       
  5403   unfolding fst_setsum snd_setsum fst_scaleR snd_scaleR
       
  5404   apply safe
       
  5405   apply (drule_tac x=k in spec)
       
  5406   apply (drule_tac x=u in spec)
       
  5407   apply (drule_tac x="\<lambda>i. (x i, f (x i))" in spec)
       
  5408   apply simp
       
  5409   using assms[unfolded convex]
       
  5410   apply simp
       
  5411   apply (rule_tac y="\<Sum>i = 1..k. u i * f (fst (x i))" in order_trans)
       
  5412   defer
       
  5413   apply (rule setsum_mono)
       
  5414   apply (erule_tac x=i in allE)
       
  5415   unfolding real_scaleR_def
       
  5416   apply (rule mult_left_mono)
       
  5417   using assms[unfolded convex]
       
  5418   apply auto
       
  5419   done
       
  5420 
       
  5421 
       
  5422 subsection \<open>Convexity of general and special intervals\<close>
       
  5423 
       
  5424 lemma is_interval_convex:
       
  5425   fixes s :: "'a::euclidean_space set"
       
  5426   assumes "is_interval s"
       
  5427   shows "convex s"
       
  5428 proof (rule convexI)
       
  5429   fix x y and u v :: real
       
  5430   assume as: "x \<in> s" "y \<in> s" "0 \<le> u" "0 \<le> v" "u + v = 1"
       
  5431   then have *: "u = 1 - v" "1 - v \<ge> 0" and **: "v = 1 - u" "1 - u \<ge> 0"
       
  5432     by auto
       
  5433   {
       
  5434     fix a b
       
  5435     assume "\<not> b \<le> u * a + v * b"
       
  5436     then have "u * a < (1 - v) * b"
       
  5437       unfolding not_le using as(4) by (auto simp add: field_simps)
       
  5438     then have "a < b"
       
  5439       unfolding * using as(4) *(2)
       
  5440       apply (rule_tac mult_left_less_imp_less[of "1 - v"])
       
  5441       apply (auto simp add: field_simps)
       
  5442       done
       
  5443     then have "a \<le> u * a + v * b"
       
  5444       unfolding * using as(4)
       
  5445       by (auto simp add: field_simps intro!:mult_right_mono)
       
  5446   }
       
  5447   moreover
       
  5448   {
       
  5449     fix a b
       
  5450     assume "\<not> u * a + v * b \<le> a"
       
  5451     then have "v * b > (1 - u) * a"
       
  5452       unfolding not_le using as(4) by (auto simp add: field_simps)
       
  5453     then have "a < b"
       
  5454       unfolding * using as(4)
       
  5455       apply (rule_tac mult_left_less_imp_less)
       
  5456       apply (auto simp add: field_simps)
       
  5457       done
       
  5458     then have "u * a + v * b \<le> b"
       
  5459       unfolding **
       
  5460       using **(2) as(3)
       
  5461       by (auto simp add: field_simps intro!:mult_right_mono)
       
  5462   }
       
  5463   ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> s"
       
  5464     apply -
       
  5465     apply (rule assms[unfolded is_interval_def, rule_format, OF as(1,2)])
       
  5466     using as(3-) DIM_positive[where 'a='a]
       
  5467     apply (auto simp: inner_simps)
       
  5468     done
       
  5469 qed
       
  5470 
       
  5471 lemma is_interval_connected:
       
  5472   fixes s :: "'a::euclidean_space set"
       
  5473   shows "is_interval s \<Longrightarrow> connected s"
       
  5474   using is_interval_convex convex_connected by auto
       
  5475 
       
  5476 lemma convex_box [simp]: "convex (cbox a b)" "convex (box a (b::'a::euclidean_space))"
       
  5477   apply (rule_tac[!] is_interval_convex)+
       
  5478   using is_interval_box is_interval_cbox
       
  5479   apply auto
       
  5480   done
       
  5481 
       
  5482 subsection \<open>On \<open>real\<close>, \<open>is_interval\<close>, \<open>convex\<close> and \<open>connected\<close> are all equivalent.\<close>
       
  5483 
       
  5484 lemma is_interval_1:
       
  5485   "is_interval (s::real set) \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. \<forall> x. a \<le> x \<and> x \<le> b \<longrightarrow> x \<in> s)"
       
  5486   unfolding is_interval_def by auto
       
  5487 
       
  5488 lemma is_interval_connected_1:
       
  5489   fixes s :: "real set"
       
  5490   shows "is_interval s \<longleftrightarrow> connected s"
       
  5491   apply rule
       
  5492   apply (rule is_interval_connected, assumption)
       
  5493   unfolding is_interval_1
       
  5494   apply rule
       
  5495   apply rule
       
  5496   apply rule
       
  5497   apply rule
       
  5498   apply (erule conjE)
       
  5499   apply (rule ccontr)
       
  5500 proof -
       
  5501   fix a b x
       
  5502   assume as: "connected s" "a \<in> s" "b \<in> s" "a \<le> x" "x \<le> b" "x \<notin> s"
       
  5503   then have *: "a < x" "x < b"
       
  5504     unfolding not_le [symmetric] by auto
       
  5505   let ?halfl = "{..<x} "
       
  5506   let ?halfr = "{x<..}"
       
  5507   {
       
  5508     fix y
       
  5509     assume "y \<in> s"
       
  5510     with \<open>x \<notin> s\<close> have "x \<noteq> y" by auto
       
  5511     then have "y \<in> ?halfr \<union> ?halfl" by auto
       
  5512   }
       
  5513   moreover have "a \<in> ?halfl" "b \<in> ?halfr" using * by auto
       
  5514   then have "?halfl \<inter> s \<noteq> {}" "?halfr \<inter> s \<noteq> {}"
       
  5515     using as(2-3) by auto
       
  5516   ultimately show False
       
  5517     apply (rule_tac notE[OF as(1)[unfolded connected_def]])
       
  5518     apply (rule_tac x = ?halfl in exI)
       
  5519     apply (rule_tac x = ?halfr in exI)
       
  5520     apply rule
       
  5521     apply (rule open_lessThan)
       
  5522     apply rule
       
  5523     apply (rule open_greaterThan)
       
  5524     apply auto
       
  5525     done
       
  5526 qed
       
  5527 
       
  5528 lemma is_interval_convex_1:
       
  5529   fixes s :: "real set"
       
  5530   shows "is_interval s \<longleftrightarrow> convex s"
       
  5531   by (metis is_interval_convex convex_connected is_interval_connected_1)
       
  5532 
       
  5533 lemma connected_convex_1:
       
  5534   fixes s :: "real set"
       
  5535   shows "connected s \<longleftrightarrow> convex s"
       
  5536   by (metis is_interval_convex convex_connected is_interval_connected_1)
       
  5537 
       
  5538 lemma connected_convex_1_gen:
       
  5539   fixes s :: "'a :: euclidean_space set"
       
  5540   assumes "DIM('a) = 1"
       
  5541   shows "connected s \<longleftrightarrow> convex s"
       
  5542 proof -
       
  5543   obtain f:: "'a \<Rightarrow> real" where linf: "linear f" and "inj f"
       
  5544     using subspace_isomorphism [where 'a = 'a and 'b = real]
       
  5545     by (metis DIM_real dim_UNIV subspace_UNIV assms)
       
  5546   then have "f -` (f ` s) = s"
       
  5547     by (simp add: inj_vimage_image_eq)
       
  5548   then show ?thesis
       
  5549     by (metis connected_convex_1 convex_linear_vimage linf convex_connected connected_linear_image)
       
  5550 qed
       
  5551 
       
  5552 subsection \<open>Another intermediate value theorem formulation\<close>
       
  5553 
       
  5554 lemma ivt_increasing_component_on_1:
       
  5555   fixes f :: "real \<Rightarrow> 'a::euclidean_space"
       
  5556   assumes "a \<le> b"
       
  5557     and "continuous_on {a..b} f"
       
  5558     and "(f a)\<bullet>k \<le> y" "y \<le> (f b)\<bullet>k"
       
  5559   shows "\<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
       
  5560 proof -
       
  5561   have "f a \<in> f ` cbox a b" "f b \<in> f ` cbox a b"
       
  5562     apply (rule_tac[!] imageI)
       
  5563     using assms(1)
       
  5564     apply auto
       
  5565     done
       
  5566   then show ?thesis
       
  5567     using connected_ivt_component[of "f ` cbox a b" "f a" "f b" k y]
       
  5568     by (simp add: Topology_Euclidean_Space.connected_continuous_image assms)
       
  5569 qed
       
  5570 
       
  5571 lemma ivt_increasing_component_1:
       
  5572   fixes f :: "real \<Rightarrow> 'a::euclidean_space"
       
  5573   shows "a \<le> b \<Longrightarrow> \<forall>x\<in>{a..b}. continuous (at x) f \<Longrightarrow>
       
  5574     f a\<bullet>k \<le> y \<Longrightarrow> y \<le> f b\<bullet>k \<Longrightarrow> \<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
       
  5575   by (rule ivt_increasing_component_on_1) (auto simp add: continuous_at_imp_continuous_on)
       
  5576 
       
  5577 lemma ivt_decreasing_component_on_1:
       
  5578   fixes f :: "real \<Rightarrow> 'a::euclidean_space"
       
  5579   assumes "a \<le> b"
       
  5580     and "continuous_on {a..b} f"
       
  5581     and "(f b)\<bullet>k \<le> y"
       
  5582     and "y \<le> (f a)\<bullet>k"
       
  5583   shows "\<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
       
  5584   apply (subst neg_equal_iff_equal[symmetric])
       
  5585   using ivt_increasing_component_on_1[of a b "\<lambda>x. - f x" k "- y"]
       
  5586   using assms using continuous_on_minus
       
  5587   apply auto
       
  5588   done
       
  5589 
       
  5590 lemma ivt_decreasing_component_1:
       
  5591   fixes f :: "real \<Rightarrow> 'a::euclidean_space"
       
  5592   shows "a \<le> b \<Longrightarrow> \<forall>x\<in>{a..b}. continuous (at x) f \<Longrightarrow>
       
  5593     f b\<bullet>k \<le> y \<Longrightarrow> y \<le> f a\<bullet>k \<Longrightarrow> \<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
       
  5594   by (rule ivt_decreasing_component_on_1) (auto simp: continuous_at_imp_continuous_on)
       
  5595 
       
  5596 
       
  5597 subsection \<open>A bound within a convex hull, and so an interval\<close>
       
  5598 
       
  5599 lemma convex_on_convex_hull_bound:
       
  5600   assumes "convex_on (convex hull s) f"
       
  5601     and "\<forall>x\<in>s. f x \<le> b"
       
  5602   shows "\<forall>x\<in> convex hull s. f x \<le> b"
       
  5603 proof
       
  5604   fix x
       
  5605   assume "x \<in> convex hull s"
       
  5606   then obtain k u v where
       
  5607     obt: "\<forall>i\<in>{1..k::nat}. 0 \<le> u i \<and> v i \<in> s" "setsum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R v i) = x"
       
  5608     unfolding convex_hull_indexed mem_Collect_eq by auto
       
  5609   have "(\<Sum>i = 1..k. u i * f (v i)) \<le> b"
       
  5610     using setsum_mono[of "{1..k}" "\<lambda>i. u i * f (v i)" "\<lambda>i. u i * b"]
       
  5611     unfolding setsum_left_distrib[symmetric] obt(2) mult_1
       
  5612     apply (drule_tac meta_mp)
       
  5613     apply (rule mult_left_mono)
       
  5614     using assms(2) obt(1)
       
  5615     apply auto
       
  5616     done
       
  5617   then show "f x \<le> b"
       
  5618     using assms(1)[unfolded convex_on[OF convex_convex_hull], rule_format, of k u v]
       
  5619     unfolding obt(2-3)
       
  5620     using obt(1) and hull_subset[unfolded subset_eq, rule_format, of _ s]
       
  5621     by auto
       
  5622 qed
       
  5623 
       
  5624 lemma inner_setsum_Basis[simp]: "i \<in> Basis \<Longrightarrow> (\<Sum>Basis) \<bullet> i = 1"
       
  5625   by (simp add: inner_setsum_left setsum.If_cases inner_Basis)
       
  5626 
       
  5627 lemma convex_set_plus:
       
  5628   assumes "convex s" and "convex t" shows "convex (s + t)"
       
  5629 proof -
       
  5630   have "convex {x + y |x y. x \<in> s \<and> y \<in> t}"
       
  5631     using assms by (rule convex_sums)
       
  5632   moreover have "{x + y |x y. x \<in> s \<and> y \<in> t} = s + t"
       
  5633     unfolding set_plus_def by auto
       
  5634   finally show "convex (s + t)" .
       
  5635 qed
       
  5636 
       
  5637 lemma convex_set_setsum:
       
  5638   assumes "\<And>i. i \<in> A \<Longrightarrow> convex (B i)"
       
  5639   shows "convex (\<Sum>i\<in>A. B i)"
       
  5640 proof (cases "finite A")
       
  5641   case True then show ?thesis using assms
       
  5642     by induct (auto simp: convex_set_plus)
       
  5643 qed auto
       
  5644 
       
  5645 lemma finite_set_setsum:
       
  5646   assumes "finite A" and "\<forall>i\<in>A. finite (B i)" shows "finite (\<Sum>i\<in>A. B i)"
       
  5647   using assms by (induct set: finite, simp, simp add: finite_set_plus)
       
  5648 
       
  5649 lemma set_setsum_eq:
       
  5650   "finite A \<Longrightarrow> (\<Sum>i\<in>A. B i) = {\<Sum>i\<in>A. f i |f. \<forall>i\<in>A. f i \<in> B i}"
       
  5651   apply (induct set: finite)
       
  5652   apply simp
       
  5653   apply simp
       
  5654   apply (safe elim!: set_plus_elim)
       
  5655   apply (rule_tac x="fun_upd f x a" in exI)
       
  5656   apply simp
       
  5657   apply (rule_tac f="\<lambda>x. a + x" in arg_cong)
       
  5658   apply (rule setsum.cong [OF refl])
       
  5659   apply clarsimp
       
  5660   apply fast
       
  5661   done
       
  5662 
       
  5663 lemma box_eq_set_setsum_Basis:
       
  5664   shows "{x. \<forall>i\<in>Basis. x\<bullet>i \<in> B i} = (\<Sum>i\<in>Basis. image (\<lambda>x. x *\<^sub>R i) (B i))"
       
  5665   apply (subst set_setsum_eq [OF finite_Basis])
       
  5666   apply safe
       
  5667   apply (fast intro: euclidean_representation [symmetric])
       
  5668   apply (subst inner_setsum_left)
       
  5669   apply (subgoal_tac "(\<Sum>x\<in>Basis. f x \<bullet> i) = f i \<bullet> i")
       
  5670   apply (drule (1) bspec)
       
  5671   apply clarsimp
       
  5672   apply (frule setsum.remove [OF finite_Basis])
       
  5673   apply (erule trans)
       
  5674   apply simp
       
  5675   apply (rule setsum.neutral)
       
  5676   apply clarsimp
       
  5677   apply (frule_tac x=i in bspec, assumption)
       
  5678   apply (drule_tac x=x in bspec, assumption)
       
  5679   apply clarsimp
       
  5680   apply (cut_tac u=x and v=i in inner_Basis, assumption+)
       
  5681   apply (rule ccontr)
       
  5682   apply simp
       
  5683   done
       
  5684 
       
  5685 lemma convex_hull_set_setsum:
       
  5686   "convex hull (\<Sum>i\<in>A. B i) = (\<Sum>i\<in>A. convex hull (B i))"
       
  5687 proof (cases "finite A")
       
  5688   assume "finite A" then show ?thesis
       
  5689     by (induct set: finite, simp, simp add: convex_hull_set_plus)
       
  5690 qed simp
       
  5691 
       
  5692 lemma convex_hull_eq_real_cbox:
       
  5693   fixes x y :: real assumes "x \<le> y"
       
  5694   shows "convex hull {x, y} = cbox x y"
       
  5695 proof (rule hull_unique)
       
  5696   show "{x, y} \<subseteq> cbox x y" using \<open>x \<le> y\<close> by auto
       
  5697   show "convex (cbox x y)"
       
  5698     by (rule convex_box)
       
  5699 next
       
  5700   fix s assume "{x, y} \<subseteq> s" and "convex s"
       
  5701   then show "cbox x y \<subseteq> s"
       
  5702     unfolding is_interval_convex_1 [symmetric] is_interval_def Basis_real_def
       
  5703     by - (clarify, simp (no_asm_use), fast)
       
  5704 qed
       
  5705 
       
  5706 lemma unit_interval_convex_hull:
       
  5707   "cbox (0::'a::euclidean_space) One = convex hull {x. \<forall>i\<in>Basis. (x\<bullet>i = 0) \<or> (x\<bullet>i = 1)}"
       
  5708   (is "?int = convex hull ?points")
       
  5709 proof -
       
  5710   have One[simp]: "\<And>i. i \<in> Basis \<Longrightarrow> One \<bullet> i = 1"
       
  5711     by (simp add: inner_setsum_left setsum.If_cases inner_Basis)
       
  5712   have "?int = {x. \<forall>i\<in>Basis. x \<bullet> i \<in> cbox 0 1}"
       
  5713     by (auto simp: cbox_def)
       
  5714   also have "\<dots> = (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` cbox 0 1)"
       
  5715     by (simp only: box_eq_set_setsum_Basis)
       
  5716   also have "\<dots> = (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` (convex hull {0, 1}))"
       
  5717     by (simp only: convex_hull_eq_real_cbox zero_le_one)
       
  5718   also have "\<dots> = (\<Sum>i\<in>Basis. convex hull ((\<lambda>x. x *\<^sub>R i) ` {0, 1}))"
       
  5719     by (simp only: convex_hull_linear_image linear_scaleR_left)
       
  5720   also have "\<dots> = convex hull (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` {0, 1})"
       
  5721     by (simp only: convex_hull_set_setsum)
       
  5722   also have "\<dots> = convex hull {x. \<forall>i\<in>Basis. x\<bullet>i \<in> {0, 1}}"
       
  5723     by (simp only: box_eq_set_setsum_Basis)
       
  5724   also have "convex hull {x. \<forall>i\<in>Basis. x\<bullet>i \<in> {0, 1}} = convex hull ?points"
       
  5725     by simp
       
  5726   finally show ?thesis .
       
  5727 qed
       
  5728 
       
  5729 text \<open>And this is a finite set of vertices.\<close>
       
  5730 
       
  5731 lemma unit_cube_convex_hull:
       
  5732   obtains s :: "'a::euclidean_space set"
       
  5733     where "finite s" and "cbox 0 (\<Sum>Basis) = convex hull s"
       
  5734   apply (rule that[of "{x::'a. \<forall>i\<in>Basis. x\<bullet>i=0 \<or> x\<bullet>i=1}"])
       
  5735   apply (rule finite_subset[of _ "(\<lambda>s. (\<Sum>i\<in>Basis. (if i\<in>s then 1 else 0) *\<^sub>R i)::'a) ` Pow Basis"])
       
  5736   prefer 3
       
  5737   apply (rule unit_interval_convex_hull)
       
  5738   apply rule
       
  5739   unfolding mem_Collect_eq
       
  5740 proof -
       
  5741   fix x :: 'a
       
  5742   assume as: "\<forall>i\<in>Basis. x \<bullet> i = 0 \<or> x \<bullet> i = 1"
       
  5743   show "x \<in> (\<lambda>s. \<Sum>i\<in>Basis. (if i\<in>s then 1 else 0) *\<^sub>R i) ` Pow Basis"
       
  5744     apply (rule image_eqI[where x="{i. i\<in>Basis \<and> x\<bullet>i = 1}"])
       
  5745     using as
       
  5746     apply (subst euclidean_eq_iff)
       
  5747     apply auto
       
  5748     done
       
  5749 qed auto
       
  5750 
       
  5751 text \<open>Hence any cube (could do any nonempty interval).\<close>
       
  5752 
       
  5753 lemma cube_convex_hull:
       
  5754   assumes "d > 0"
       
  5755   obtains s :: "'a::euclidean_space set" where
       
  5756     "finite s" and "cbox (x - (\<Sum>i\<in>Basis. d*\<^sub>Ri)) (x + (\<Sum>i\<in>Basis. d*\<^sub>Ri)) = convex hull s"
       
  5757 proof -
       
  5758   let ?d = "(\<Sum>i\<in>Basis. d*\<^sub>Ri)::'a"
       
  5759   have *: "cbox (x - ?d) (x + ?d) = (\<lambda>y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 (\<Sum>Basis)"
       
  5760     apply (rule set_eqI, rule)
       
  5761     unfolding image_iff
       
  5762     defer
       
  5763     apply (erule bexE)
       
  5764   proof -
       
  5765     fix y
       
  5766     assume as: "y\<in>cbox (x - ?d) (x + ?d)"
       
  5767     then have "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \<in> cbox 0 (\<Sum>Basis)"
       
  5768       using assms by (simp add: mem_box field_simps inner_simps)
       
  5769     with \<open>0 < d\<close> show "\<exists>z\<in>cbox 0 (\<Sum>Basis). y = x - ?d + (2 * d) *\<^sub>R z"
       
  5770       by (intro bexI[of _ "inverse (2 * d) *\<^sub>R (y - (x - ?d))"]) auto
       
  5771   next
       
  5772     fix y z
       
  5773     assume as: "z\<in>cbox 0 (\<Sum>Basis)" "y = x - ?d + (2*d) *\<^sub>R z"
       
  5774     have "\<And>i. i\<in>Basis \<Longrightarrow> 0 \<le> d * (z \<bullet> i) \<and> d * (z \<bullet> i) \<le> d"
       
  5775       using assms as(1)[unfolded mem_box]
       
  5776       apply (erule_tac x=i in ballE)
       
  5777       apply rule
       
  5778       prefer 2
       
  5779       apply (rule mult_right_le_one_le)
       
  5780       using assms
       
  5781       apply auto
       
  5782       done
       
  5783     then show "y \<in> cbox (x - ?d) (x + ?d)"
       
  5784       unfolding as(2) mem_box
       
  5785       apply -
       
  5786       apply rule
       
  5787       using as(1)[unfolded mem_box]
       
  5788       apply (erule_tac x=i in ballE)
       
  5789       using assms
       
  5790       apply (auto simp: inner_simps)
       
  5791       done
       
  5792   qed
       
  5793   obtain s where "finite s" "cbox 0 (\<Sum>Basis::'a) = convex hull s"
       
  5794     using unit_cube_convex_hull by auto
       
  5795   then show ?thesis
       
  5796     apply (rule_tac that[of "(\<lambda>y. x - ?d + (2 * d) *\<^sub>R y)` s"])
       
  5797     unfolding * and convex_hull_affinity
       
  5798     apply auto
       
  5799     done
       
  5800 qed
       
  5801 
       
  5802 subsubsection\<open>Representation of any interval as a finite convex hull\<close>
       
  5803 
       
  5804 lemma image_stretch_interval:
       
  5805   "(\<lambda>x. \<Sum>k\<in>Basis. (m k * (x\<bullet>k)) *\<^sub>R k) ` cbox a (b::'a::euclidean_space) =
       
  5806   (if (cbox a b) = {} then {} else
       
  5807     cbox (\<Sum>k\<in>Basis. (min (m k * (a\<bullet>k)) (m k * (b\<bullet>k))) *\<^sub>R k::'a)
       
  5808      (\<Sum>k\<in>Basis. (max (m k * (a\<bullet>k)) (m k * (b\<bullet>k))) *\<^sub>R k))"
       
  5809 proof cases
       
  5810   assume *: "cbox a b \<noteq> {}"
       
  5811   show ?thesis
       
  5812     unfolding box_ne_empty if_not_P[OF *]
       
  5813     apply (simp add: cbox_def image_Collect set_eq_iff euclidean_eq_iff[where 'a='a] ball_conj_distrib[symmetric])
       
  5814     apply (subst choice_Basis_iff[symmetric])
       
  5815   proof (intro allI ball_cong refl)
       
  5816     fix x i :: 'a assume "i \<in> Basis"
       
  5817     with * have a_le_b: "a \<bullet> i \<le> b \<bullet> i"
       
  5818       unfolding box_ne_empty by auto
       
  5819     show "(\<exists>xa. x \<bullet> i = m i * xa \<and> a \<bullet> i \<le> xa \<and> xa \<le> b \<bullet> i) \<longleftrightarrow>
       
  5820         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))"
       
  5821     proof (cases "m i = 0")
       
  5822       case True
       
  5823       with a_le_b show ?thesis by auto
       
  5824     next
       
  5825       case False
       
  5826       then have *: "\<And>a b. a = m i * b \<longleftrightarrow> b = a / m i"
       
  5827         by (auto simp add: field_simps)
       
  5828       from False have
       
  5829           "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))"
       
  5830           "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))"
       
  5831         using a_le_b by (auto simp: min_def max_def mult_le_cancel_left)
       
  5832       with False show ?thesis using a_le_b
       
  5833         unfolding * by (auto simp add: le_divide_eq divide_le_eq ac_simps)
       
  5834     qed
       
  5835   qed
       
  5836 qed simp
       
  5837 
       
  5838 lemma interval_image_stretch_interval:
       
  5839   "\<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)"
       
  5840   unfolding image_stretch_interval by auto
       
  5841 
       
  5842 lemma cbox_translation: "cbox (c + a) (c + b) = image (\<lambda>x. c + x) (cbox a b)"
       
  5843   using image_affinity_cbox [of 1 c a b]
       
  5844   using box_ne_empty [of "a+c" "b+c"]  box_ne_empty [of a b]
       
  5845   by (auto simp add: inner_left_distrib add.commute)
       
  5846 
       
  5847 lemma cbox_image_unit_interval:
       
  5848   fixes a :: "'a::euclidean_space"
       
  5849   assumes "cbox a b \<noteq> {}"
       
  5850     shows "cbox a b =
       
  5851            op + a ` (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k) ` cbox 0 One"
       
  5852 using assms
       
  5853 apply (simp add: box_ne_empty image_stretch_interval cbox_translation [symmetric])
       
  5854 apply (simp add: min_def max_def algebra_simps setsum_subtractf euclidean_representation)
       
  5855 done
       
  5856 
       
  5857 lemma closed_interval_as_convex_hull:
       
  5858   fixes a :: "'a::euclidean_space"
       
  5859   obtains s where "finite s" "cbox a b = convex hull s"
       
  5860 proof (cases "cbox a b = {}")
       
  5861   case True with convex_hull_empty that show ?thesis
       
  5862     by blast
       
  5863 next
       
  5864   case False
       
  5865   obtain s::"'a set" where "finite s" and eq: "cbox 0 One = convex hull s"
       
  5866     by (blast intro: unit_cube_convex_hull)
       
  5867   have lin: "linear (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k)"
       
  5868     by (rule linear_compose_setsum) (auto simp: algebra_simps linearI)
       
  5869   have "finite (op + a ` (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k) ` s)"
       
  5870     by (rule finite_imageI \<open>finite s\<close>)+
       
  5871   then show ?thesis
       
  5872     apply (rule that)
       
  5873     apply (simp add: convex_hull_translation convex_hull_linear_image [OF lin, symmetric])
       
  5874     apply (simp add: eq [symmetric] cbox_image_unit_interval [OF False])
       
  5875     done
       
  5876 qed
       
  5877 
       
  5878 
       
  5879 subsection \<open>Bounded convex function on open set is continuous\<close>
       
  5880 
       
  5881 lemma convex_on_bounded_continuous:
       
  5882   fixes s :: "('a::real_normed_vector) set"
       
  5883   assumes "open s"
       
  5884     and "convex_on s f"
       
  5885     and "\<forall>x\<in>s. \<bar>f x\<bar> \<le> b"
       
  5886   shows "continuous_on s f"
       
  5887   apply (rule continuous_at_imp_continuous_on)
       
  5888   unfolding continuous_at_real_range
       
  5889 proof (rule,rule,rule)
       
  5890   fix x and e :: real
       
  5891   assume "x \<in> s" "e > 0"
       
  5892   define B where "B = \<bar>b\<bar> + 1"
       
  5893   have B: "0 < B" "\<And>x. x\<in>s \<Longrightarrow> \<bar>f x\<bar> \<le> B"
       
  5894     unfolding B_def
       
  5895     defer
       
  5896     apply (drule assms(3)[rule_format])
       
  5897     apply auto
       
  5898     done
       
  5899   obtain k where "k > 0" and k: "cball x k \<subseteq> s"
       
  5900     using assms(1)[unfolded open_contains_cball, THEN bspec[where x=x]]
       
  5901     using \<open>x\<in>s\<close> by auto
       
  5902   show "\<exists>d>0. \<forall>x'. norm (x' - x) < d \<longrightarrow> \<bar>f x' - f x\<bar> < e"
       
  5903     apply (rule_tac x="min (k / 2) (e / (2 * B) * k)" in exI)
       
  5904     apply rule
       
  5905     defer
       
  5906   proof (rule, rule)
       
  5907     fix y
       
  5908     assume as: "norm (y - x) < min (k / 2) (e / (2 * B) * k)"
       
  5909     show "\<bar>f y - f x\<bar> < e"
       
  5910     proof (cases "y = x")
       
  5911       case False
       
  5912       define t where "t = k / norm (y - x)"
       
  5913       have "2 < t" "0<t"
       
  5914         unfolding t_def using as False and \<open>k>0\<close>
       
  5915         by (auto simp add:field_simps)
       
  5916       have "y \<in> s"
       
  5917         apply (rule k[unfolded subset_eq,rule_format])
       
  5918         unfolding mem_cball dist_norm
       
  5919         apply (rule order_trans[of _ "2 * norm (x - y)"])
       
  5920         using as
       
  5921         by (auto simp add: field_simps norm_minus_commute)
       
  5922       {
       
  5923         define w where "w = x + t *\<^sub>R (y - x)"
       
  5924         have "w \<in> s"
       
  5925           unfolding w_def
       
  5926           apply (rule k[unfolded subset_eq,rule_format])
       
  5927           unfolding mem_cball dist_norm
       
  5928           unfolding t_def
       
  5929           using \<open>k>0\<close>
       
  5930           apply auto
       
  5931           done
       
  5932         have "(1 / t) *\<^sub>R x + - x + ((t - 1) / t) *\<^sub>R x = (1 / t - 1 + (t - 1) / t) *\<^sub>R x"
       
  5933           by (auto simp add: algebra_simps)
       
  5934         also have "\<dots> = 0"
       
  5935           using \<open>t > 0\<close> by (auto simp add:field_simps)
       
  5936         finally have w: "(1 / t) *\<^sub>R w + ((t - 1) / t) *\<^sub>R x = y"
       
  5937           unfolding w_def using False and \<open>t > 0\<close>
       
  5938           by (auto simp add: algebra_simps)
       
  5939         have  "2 * B < e * t"
       
  5940           unfolding t_def using \<open>0 < e\<close> \<open>0 < k\<close> \<open>B > 0\<close> and as and False
       
  5941           by (auto simp add:field_simps)
       
  5942         then have "(f w - f x) / t < e"
       
  5943           using B(2)[OF \<open>w\<in>s\<close>] and B(2)[OF \<open>x\<in>s\<close>]
       
  5944           using \<open>t > 0\<close> by (auto simp add:field_simps)
       
  5945         then have th1: "f y - f x < e"
       
  5946           apply -
       
  5947           apply (rule le_less_trans)
       
  5948           defer
       
  5949           apply assumption
       
  5950           using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w]
       
  5951           using \<open>0 < t\<close> \<open>2 < t\<close> and \<open>x \<in> s\<close> \<open>w \<in> s\<close>
       
  5952           by (auto simp add:field_simps)
       
  5953       }
       
  5954       moreover
       
  5955       {
       
  5956         define w where "w = x - t *\<^sub>R (y - x)"
       
  5957         have "w \<in> s"
       
  5958           unfolding w_def
       
  5959           apply (rule k[unfolded subset_eq,rule_format])
       
  5960           unfolding mem_cball dist_norm
       
  5961           unfolding t_def
       
  5962           using \<open>k > 0\<close>
       
  5963           apply auto
       
  5964           done
       
  5965         have "(1 / (1 + t)) *\<^sub>R x + (t / (1 + t)) *\<^sub>R x = (1 / (1 + t) + t / (1 + t)) *\<^sub>R x"
       
  5966           by (auto simp add: algebra_simps)
       
  5967         also have "\<dots> = x"
       
  5968           using \<open>t > 0\<close> by (auto simp add:field_simps)
       
  5969         finally have w: "(1 / (1+t)) *\<^sub>R w + (t / (1 + t)) *\<^sub>R y = x"
       
  5970           unfolding w_def using False and \<open>t > 0\<close>
       
  5971           by (auto simp add: algebra_simps)
       
  5972         have "2 * B < e * t"
       
  5973           unfolding t_def
       
  5974           using \<open>0 < e\<close> \<open>0 < k\<close> \<open>B > 0\<close> and as and False
       
  5975           by (auto simp add:field_simps)
       
  5976         then have *: "(f w - f y) / t < e"
       
  5977           using B(2)[OF \<open>w\<in>s\<close>] and B(2)[OF \<open>y\<in>s\<close>]
       
  5978           using \<open>t > 0\<close>
       
  5979           by (auto simp add:field_simps)
       
  5980         have "f x \<le> 1 / (1 + t) * f w + (t / (1 + t)) * f y"
       
  5981           using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w]
       
  5982           using \<open>0 < t\<close> \<open>2 < t\<close> and \<open>y \<in> s\<close> \<open>w \<in> s\<close>
       
  5983           by (auto simp add:field_simps)
       
  5984         also have "\<dots> = (f w + t * f y) / (1 + t)"
       
  5985           using \<open>t > 0\<close> by (auto simp add: divide_simps)
       
  5986         also have "\<dots> < e + f y"
       
  5987           using \<open>t > 0\<close> * \<open>e > 0\<close> by (auto simp add: field_simps)
       
  5988         finally have "f x - f y < e" by auto
       
  5989       }
       
  5990       ultimately show ?thesis by auto
       
  5991     qed (insert \<open>0<e\<close>, auto)
       
  5992   qed (insert \<open>0<e\<close> \<open>0<k\<close> \<open>0<B\<close>, auto simp: field_simps)
       
  5993 qed
       
  5994 
       
  5995 
       
  5996 subsection \<open>Upper bound on a ball implies upper and lower bounds\<close>
       
  5997 
       
  5998 lemma convex_bounds_lemma:
       
  5999   fixes x :: "'a::real_normed_vector"
       
  6000   assumes "convex_on (cball x e) f"
       
  6001     and "\<forall>y \<in> cball x e. f y \<le> b"
       
  6002   shows "\<forall>y \<in> cball x e. \<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>"
       
  6003   apply rule
       
  6004 proof (cases "0 \<le> e")
       
  6005   case True
       
  6006   fix y
       
  6007   assume y: "y \<in> cball x e"
       
  6008   define z where "z = 2 *\<^sub>R x - y"
       
  6009   have *: "x - (2 *\<^sub>R x - y) = y - x"
       
  6010     by (simp add: scaleR_2)
       
  6011   have z: "z \<in> cball x e"
       
  6012     using y unfolding z_def mem_cball dist_norm * by (auto simp add: norm_minus_commute)
       
  6013   have "(1 / 2) *\<^sub>R y + (1 / 2) *\<^sub>R z = x"
       
  6014     unfolding z_def by (auto simp add: algebra_simps)
       
  6015   then show "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>"
       
  6016     using assms(1)[unfolded convex_on_def,rule_format, OF y z, of "1/2" "1/2"]
       
  6017     using assms(2)[rule_format,OF y] assms(2)[rule_format,OF z]
       
  6018     by (auto simp add:field_simps)
       
  6019 next
       
  6020   case False
       
  6021   fix y
       
  6022   assume "y \<in> cball x e"
       
  6023   then have "dist x y < 0"
       
  6024     using False unfolding mem_cball not_le by (auto simp del: dist_not_less_zero)
       
  6025   then show "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>"
       
  6026     using zero_le_dist[of x y] by auto
       
  6027 qed
       
  6028 
       
  6029 
       
  6030 subsubsection \<open>Hence a convex function on an open set is continuous\<close>
       
  6031 
       
  6032 lemma real_of_nat_ge_one_iff: "1 \<le> real (n::nat) \<longleftrightarrow> 1 \<le> n"
       
  6033   by auto
       
  6034 
       
  6035 lemma convex_on_continuous:
       
  6036   assumes "open (s::('a::euclidean_space) set)" "convex_on s f"
       
  6037   shows "continuous_on s f"
       
  6038   unfolding continuous_on_eq_continuous_at[OF assms(1)]
       
  6039 proof
       
  6040   note dimge1 = DIM_positive[where 'a='a]
       
  6041   fix x
       
  6042   assume "x \<in> s"
       
  6043   then obtain e where e: "cball x e \<subseteq> s" "e > 0"
       
  6044     using assms(1) unfolding open_contains_cball by auto
       
  6045   define d where "d = e / real DIM('a)"
       
  6046   have "0 < d"
       
  6047     unfolding d_def using \<open>e > 0\<close> dimge1 by auto
       
  6048   let ?d = "(\<Sum>i\<in>Basis. d *\<^sub>R i)::'a"
       
  6049   obtain c
       
  6050     where c: "finite c"
       
  6051     and c1: "convex hull c \<subseteq> cball x e"
       
  6052     and c2: "cball x d \<subseteq> convex hull c"
       
  6053   proof
       
  6054     define c where "c = (\<Sum>i\<in>Basis. (\<lambda>a. a *\<^sub>R i) ` {x\<bullet>i - d, x\<bullet>i + d})"
       
  6055     show "finite c"
       
  6056       unfolding c_def by (simp add: finite_set_setsum)
       
  6057     have 1: "convex hull c = {a. \<forall>i\<in>Basis. a \<bullet> i \<in> cbox (x \<bullet> i - d) (x \<bullet> i + d)}"
       
  6058       unfolding box_eq_set_setsum_Basis
       
  6059       unfolding c_def convex_hull_set_setsum
       
  6060       apply (subst convex_hull_linear_image [symmetric])
       
  6061       apply (simp add: linear_iff scaleR_add_left)
       
  6062       apply (rule setsum.cong [OF refl])
       
  6063       apply (rule image_cong [OF _ refl])
       
  6064       apply (rule convex_hull_eq_real_cbox)
       
  6065       apply (cut_tac \<open>0 < d\<close>, simp)
       
  6066       done
       
  6067     then have 2: "convex hull c = {a. \<forall>i\<in>Basis. a \<bullet> i \<in> cball (x \<bullet> i) d}"
       
  6068       by (simp add: dist_norm abs_le_iff algebra_simps)
       
  6069     show "cball x d \<subseteq> convex hull c"
       
  6070       unfolding 2
       
  6071       apply clarsimp
       
  6072       apply (simp only: dist_norm)
       
  6073       apply (subst inner_diff_left [symmetric])
       
  6074       apply simp
       
  6075       apply (erule (1) order_trans [OF Basis_le_norm])
       
  6076       done
       
  6077     have e': "e = (\<Sum>(i::'a)\<in>Basis. d)"
       
  6078       by (simp add: d_def DIM_positive)
       
  6079     show "convex hull c \<subseteq> cball x e"
       
  6080       unfolding 2
       
  6081       apply clarsimp
       
  6082       apply (subst euclidean_dist_l2)
       
  6083       apply (rule order_trans [OF setL2_le_setsum])
       
  6084       apply (rule zero_le_dist)
       
  6085       unfolding e'
       
  6086       apply (rule setsum_mono)
       
  6087       apply simp
       
  6088       done
       
  6089   qed
       
  6090   define k where "k = Max (f ` c)"
       
  6091   have "convex_on (convex hull c) f"
       
  6092     apply(rule convex_on_subset[OF assms(2)])
       
  6093     apply(rule subset_trans[OF _ e(1)])
       
  6094     apply(rule c1)
       
  6095     done
       
  6096   then have k: "\<forall>y\<in>convex hull c. f y \<le> k"
       
  6097     apply (rule_tac convex_on_convex_hull_bound)
       
  6098     apply assumption
       
  6099     unfolding k_def
       
  6100     apply (rule, rule Max_ge)
       
  6101     using c(1)
       
  6102     apply auto
       
  6103     done
       
  6104   have "d \<le> e"
       
  6105     unfolding d_def
       
  6106     apply (rule mult_imp_div_pos_le)
       
  6107     using \<open>e > 0\<close>
       
  6108     unfolding mult_le_cancel_left1
       
  6109     apply (auto simp: real_of_nat_ge_one_iff Suc_le_eq DIM_positive)
       
  6110     done
       
  6111   then have dsube: "cball x d \<subseteq> cball x e"
       
  6112     by (rule subset_cball)
       
  6113   have conv: "convex_on (cball x d) f"
       
  6114     apply (rule convex_on_subset)
       
  6115     apply (rule convex_on_subset[OF assms(2)])
       
  6116     apply (rule e(1))
       
  6117     apply (rule dsube)
       
  6118     done
       
  6119   then have "\<forall>y\<in>cball x d. \<bar>f y\<bar> \<le> k + 2 * \<bar>f x\<bar>"
       
  6120     apply (rule convex_bounds_lemma)
       
  6121     apply (rule ballI)
       
  6122     apply (rule k [rule_format])
       
  6123     apply (erule rev_subsetD)
       
  6124     apply (rule c2)
       
  6125     done
       
  6126   then have "continuous_on (ball x d) f"
       
  6127     apply (rule_tac convex_on_bounded_continuous)
       
  6128     apply (rule open_ball, rule convex_on_subset[OF conv])
       
  6129     apply (rule ball_subset_cball)
       
  6130     apply force
       
  6131     done
       
  6132   then show "continuous (at x) f"
       
  6133     unfolding continuous_on_eq_continuous_at[OF open_ball]
       
  6134     using \<open>d > 0\<close> by auto
       
  6135 qed
       
  6136 
       
  6137 
       
  6138 subsection \<open>Line segments, Starlike Sets, etc.\<close>
       
  6139 
       
  6140 definition midpoint :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a"
       
  6141   where "midpoint a b = (inverse (2::real)) *\<^sub>R (a + b)"
       
  6142 
       
  6143 definition closed_segment :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a set"
       
  6144   where "closed_segment a b = {(1 - u) *\<^sub>R a + u *\<^sub>R b | u::real. 0 \<le> u \<and> u \<le> 1}"
       
  6145 
       
  6146 definition open_segment :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a set" where
       
  6147   "open_segment a b \<equiv> closed_segment a b - {a,b}"
       
  6148 
       
  6149 lemmas segment = open_segment_def closed_segment_def
       
  6150 
       
  6151 lemma in_segment:
       
  6152     "x \<in> closed_segment a b \<longleftrightarrow> (\<exists>u. 0 \<le> u \<and> u \<le> 1 \<and> x = (1 - u) *\<^sub>R a + u *\<^sub>R b)"
       
  6153     "x \<in> open_segment a b \<longleftrightarrow> a \<noteq> b \<and> (\<exists>u. 0 < u \<and> u < 1 \<and> x = (1 - u) *\<^sub>R a + u *\<^sub>R b)"
       
  6154   using less_eq_real_def by (auto simp: segment algebra_simps)
       
  6155 
       
  6156 lemma closed_segment_linear_image:
       
  6157     "linear f \<Longrightarrow> closed_segment (f a) (f b) = f ` (closed_segment a b)"
       
  6158   by (force simp add: in_segment linear_add_cmul)
       
  6159 
       
  6160 lemma open_segment_linear_image:
       
  6161     "\<lbrakk>linear f; inj f\<rbrakk> \<Longrightarrow> open_segment (f a) (f b) = f ` (open_segment a b)"
       
  6162   by (force simp: open_segment_def closed_segment_linear_image inj_on_def)
       
  6163 
       
  6164 lemma closed_segment_translation:
       
  6165     "closed_segment (c + a) (c + b) = image (\<lambda>x. c + x) (closed_segment a b)"
       
  6166 apply safe
       
  6167 apply (rule_tac x="x-c" in image_eqI)
       
  6168 apply (auto simp: in_segment algebra_simps)
       
  6169 done
       
  6170 
       
  6171 lemma open_segment_translation:
       
  6172     "open_segment (c + a) (c + b) = image (\<lambda>x. c + x) (open_segment a b)"
       
  6173 by (simp add: open_segment_def closed_segment_translation translation_diff)
       
  6174 
       
  6175 lemma open_segment_PairD:
       
  6176     "(x, x') \<in> open_segment (a, a') (b, b')
       
  6177      \<Longrightarrow> (x \<in> open_segment a b \<or> a = b) \<and> (x' \<in> open_segment a' b' \<or> a' = b')"
       
  6178   by (auto simp: in_segment)
       
  6179 
       
  6180 lemma closed_segment_PairD:
       
  6181   "(x, x') \<in> closed_segment (a, a') (b, b') \<Longrightarrow> x \<in> closed_segment a b \<and> x' \<in> closed_segment a' b'"
       
  6182   by (auto simp: closed_segment_def)
       
  6183 
       
  6184 lemma closed_segment_translation_eq [simp]:
       
  6185     "d + x \<in> closed_segment (d + a) (d + b) \<longleftrightarrow> x \<in> closed_segment a b"
       
  6186 proof -
       
  6187   have *: "\<And>d x a b. x \<in> closed_segment a b \<Longrightarrow> d + x \<in> closed_segment (d + a) (d + b)"
       
  6188     apply (simp add: closed_segment_def)
       
  6189     apply (erule ex_forward)
       
  6190     apply (simp add: algebra_simps)
       
  6191     done
       
  6192   show ?thesis
       
  6193   using * [where d = "-d"] *
       
  6194   by (fastforce simp add:)
       
  6195 qed
       
  6196 
       
  6197 lemma open_segment_translation_eq [simp]:
       
  6198     "d + x \<in> open_segment (d + a) (d + b) \<longleftrightarrow> x \<in> open_segment a b"
       
  6199   by (simp add: open_segment_def)
       
  6200 
       
  6201 definition "between = (\<lambda>(a,b) x. x \<in> closed_segment a b)"
       
  6202 
       
  6203 definition "starlike s \<longleftrightarrow> (\<exists>a\<in>s. \<forall>x\<in>s. closed_segment a x \<subseteq> s)"
       
  6204 
       
  6205 lemma starlike_UNIV [simp]: "starlike UNIV"
       
  6206   by (simp add: starlike_def)
       
  6207 
       
  6208 lemma midpoint_refl: "midpoint x x = x"
       
  6209   unfolding midpoint_def
       
  6210   unfolding scaleR_right_distrib
       
  6211   unfolding scaleR_left_distrib[symmetric]
       
  6212   by auto
       
  6213 
       
  6214 lemma midpoint_sym: "midpoint a b = midpoint b a"
       
  6215   unfolding midpoint_def by (auto simp add: scaleR_right_distrib)
       
  6216 
       
  6217 lemma midpoint_eq_iff: "midpoint a b = c \<longleftrightarrow> a + b = c + c"
       
  6218 proof -
       
  6219   have "midpoint a b = c \<longleftrightarrow> scaleR 2 (midpoint a b) = scaleR 2 c"
       
  6220     by simp
       
  6221   then show ?thesis
       
  6222     unfolding midpoint_def scaleR_2 [symmetric] by simp
       
  6223 qed
       
  6224 
       
  6225 lemma dist_midpoint:
       
  6226   fixes a b :: "'a::real_normed_vector" shows
       
  6227   "dist a (midpoint a b) = (dist a b) / 2" (is ?t1)
       
  6228   "dist b (midpoint a b) = (dist a b) / 2" (is ?t2)
       
  6229   "dist (midpoint a b) a = (dist a b) / 2" (is ?t3)
       
  6230   "dist (midpoint a b) b = (dist a b) / 2" (is ?t4)
       
  6231 proof -
       
  6232   have *: "\<And>x y::'a. 2 *\<^sub>R x = - y \<Longrightarrow> norm x = (norm y) / 2"
       
  6233     unfolding equation_minus_iff by auto
       
  6234   have **: "\<And>x y::'a. 2 *\<^sub>R x =   y \<Longrightarrow> norm x = (norm y) / 2"
       
  6235     by auto
       
  6236   note scaleR_right_distrib [simp]
       
  6237   show ?t1
       
  6238     unfolding midpoint_def dist_norm
       
  6239     apply (rule **)
       
  6240     apply (simp add: scaleR_right_diff_distrib)
       
  6241     apply (simp add: scaleR_2)
       
  6242     done
       
  6243   show ?t2
       
  6244     unfolding midpoint_def dist_norm
       
  6245     apply (rule *)
       
  6246     apply (simp add: scaleR_right_diff_distrib)
       
  6247     apply (simp add: scaleR_2)
       
  6248     done
       
  6249   show ?t3
       
  6250     unfolding midpoint_def dist_norm
       
  6251     apply (rule *)
       
  6252     apply (simp add: scaleR_right_diff_distrib)
       
  6253     apply (simp add: scaleR_2)
       
  6254     done
       
  6255   show ?t4
       
  6256     unfolding midpoint_def dist_norm
       
  6257     apply (rule **)
       
  6258     apply (simp add: scaleR_right_diff_distrib)
       
  6259     apply (simp add: scaleR_2)
       
  6260     done
       
  6261 qed
       
  6262 
       
  6263 lemma midpoint_eq_endpoint [simp]:
       
  6264   "midpoint a b = a \<longleftrightarrow> a = b"
       
  6265   "midpoint a b = b \<longleftrightarrow> a = b"
       
  6266   unfolding midpoint_eq_iff by auto
       
  6267 
       
  6268 lemma convex_contains_segment:
       
  6269   "convex s \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. closed_segment a b \<subseteq> s)"
       
  6270   unfolding convex_alt closed_segment_def by auto
       
  6271 
       
  6272 lemma closed_segment_subset: "\<lbrakk>x \<in> s; y \<in> s; convex s\<rbrakk> \<Longrightarrow> closed_segment x y \<subseteq> s"
       
  6273   by (simp add: convex_contains_segment)
       
  6274 
       
  6275 lemma closed_segment_subset_convex_hull:
       
  6276     "\<lbrakk>x \<in> convex hull s; y \<in> convex hull s\<rbrakk> \<Longrightarrow> closed_segment x y \<subseteq> convex hull s"
       
  6277   using convex_contains_segment by blast
       
  6278 
       
  6279 lemma convex_imp_starlike:
       
  6280   "convex s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> starlike s"
       
  6281   unfolding convex_contains_segment starlike_def by auto
       
  6282 
       
  6283 lemma segment_convex_hull:
       
  6284   "closed_segment a b = convex hull {a,b}"
       
  6285 proof -
       
  6286   have *: "\<And>x. {x} \<noteq> {}" by auto
       
  6287   show ?thesis
       
  6288     unfolding segment convex_hull_insert[OF *] convex_hull_singleton
       
  6289     by (safe; rule_tac x="1 - u" in exI; force)
       
  6290 qed
       
  6291 
       
  6292 lemma open_closed_segment: "u \<in> open_segment w z \<Longrightarrow> u \<in> closed_segment w z"
       
  6293   by (auto simp add: closed_segment_def open_segment_def)
       
  6294 
       
  6295 lemma segment_open_subset_closed:
       
  6296    "open_segment a b \<subseteq> closed_segment a b"
       
  6297   by (auto simp: closed_segment_def open_segment_def)
       
  6298 
       
  6299 lemma bounded_closed_segment:
       
  6300     fixes a :: "'a::euclidean_space" shows "bounded (closed_segment a b)"
       
  6301   by (simp add: segment_convex_hull compact_convex_hull compact_imp_bounded)
       
  6302 
       
  6303 lemma bounded_open_segment:
       
  6304     fixes a :: "'a::euclidean_space" shows "bounded (open_segment a b)"
       
  6305   by (rule bounded_subset [OF bounded_closed_segment segment_open_subset_closed])
       
  6306 
       
  6307 lemmas bounded_segment = bounded_closed_segment open_closed_segment
       
  6308 
       
  6309 lemma ends_in_segment [iff]: "a \<in> closed_segment a b" "b \<in> closed_segment a b"
       
  6310   unfolding segment_convex_hull
       
  6311   by (auto intro!: hull_subset[unfolded subset_eq, rule_format])
       
  6312 
       
  6313 lemma segment_furthest_le:
       
  6314   fixes a b x y :: "'a::euclidean_space"
       
  6315   assumes "x \<in> closed_segment a b"
       
  6316   shows "norm (y - x) \<le> norm (y - a) \<or>  norm (y - x) \<le> norm (y - b)"
       
  6317 proof -
       
  6318   obtain z where "z \<in> {a, b}" "norm (x - y) \<le> norm (z - y)"
       
  6319     using simplex_furthest_le[of "{a, b}" y]
       
  6320     using assms[unfolded segment_convex_hull]
       
  6321     by auto
       
  6322   then show ?thesis
       
  6323     by (auto simp add:norm_minus_commute)
       
  6324 qed
       
  6325 
       
  6326 lemma closed_segment_commute: "closed_segment a b = closed_segment b a"
       
  6327 proof -
       
  6328   have "{a, b} = {b, a}" by auto
       
  6329   thus ?thesis
       
  6330     by (simp add: segment_convex_hull)
       
  6331 qed
       
  6332 
       
  6333 lemma segment_bound1:
       
  6334   assumes "x \<in> closed_segment a b"
       
  6335   shows "norm (x - a) \<le> norm (b - a)"
       
  6336 proof -
       
  6337   obtain u where "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \<le> u" "u \<le> 1"
       
  6338     using assms by (auto simp add: closed_segment_def)
       
  6339   then show "norm (x - a) \<le> norm (b - a)"
       
  6340     apply clarify
       
  6341     apply (auto simp: algebra_simps)
       
  6342     apply (simp add: scaleR_diff_right [symmetric] mult_left_le_one_le)
       
  6343     done
       
  6344 qed
       
  6345 
       
  6346 lemma segment_bound:
       
  6347   assumes "x \<in> closed_segment a b"
       
  6348   shows "norm (x - a) \<le> norm (b - a)" "norm (x - b) \<le> norm (b - a)"
       
  6349 apply (simp add: assms segment_bound1)
       
  6350 by (metis assms closed_segment_commute dist_commute dist_norm segment_bound1)
       
  6351 
       
  6352 lemma open_segment_commute: "open_segment a b = open_segment b a"
       
  6353 proof -
       
  6354   have "{a, b} = {b, a}" by auto
       
  6355   thus ?thesis
       
  6356     by (simp add: closed_segment_commute open_segment_def)
       
  6357 qed
       
  6358 
       
  6359 lemma closed_segment_idem [simp]: "closed_segment a a = {a}"
       
  6360   unfolding segment by (auto simp add: algebra_simps)
       
  6361 
       
  6362 lemma open_segment_idem [simp]: "open_segment a a = {}"
       
  6363   by (simp add: open_segment_def)
       
  6364 
       
  6365 lemma closed_segment_eq_open: "closed_segment a b = open_segment a b \<union> {a,b}"
       
  6366   using open_segment_def by auto
       
  6367 
       
  6368 lemma convex_contains_open_segment:
       
  6369   "convex s \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. open_segment a b \<subseteq> s)"
       
  6370   by (simp add: convex_contains_segment closed_segment_eq_open)
       
  6371 
       
  6372 lemma closed_segment_eq_real_ivl:
       
  6373   fixes a b::real
       
  6374   shows "closed_segment a b = (if a \<le> b then {a .. b} else {b .. a})"
       
  6375 proof -
       
  6376   have "b \<le> a \<Longrightarrow> closed_segment b a = {b .. a}"
       
  6377     and "a \<le> b \<Longrightarrow> closed_segment a b = {a .. b}"
       
  6378     by (auto simp: convex_hull_eq_real_cbox segment_convex_hull)
       
  6379   thus ?thesis
       
  6380     by (auto simp: closed_segment_commute)
       
  6381 qed
       
  6382 
       
  6383 lemma closed_segment_real_eq:
       
  6384   fixes u::real shows "closed_segment u v = (\<lambda>x. (v - u) * x + u) ` {0..1}"
       
  6385   by (simp add: add.commute [of u] image_affinity_atLeastAtMost [where c=u] closed_segment_eq_real_ivl)
       
  6386 
       
  6387 lemma dist_in_closed_segment:
       
  6388   fixes a :: "'a :: euclidean_space"
       
  6389   assumes "x \<in> closed_segment a b"
       
  6390     shows "dist x a \<le> dist a b \<and> dist x b \<le> dist a b"
       
  6391 proof (intro conjI)
       
  6392   obtain u where u: "0 \<le> u" "u \<le> 1" and x: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b"
       
  6393     using assms by (force simp: in_segment algebra_simps)
       
  6394   have "dist x a = u * dist a b"
       
  6395     apply (simp add: dist_norm algebra_simps x)
       
  6396     by (metis \<open>0 \<le> u\<close> abs_of_nonneg norm_minus_commute norm_scaleR real_vector.scale_right_diff_distrib)
       
  6397   also have "...  \<le> dist a b"
       
  6398     by (simp add: mult_left_le_one_le u)
       
  6399   finally show "dist x a \<le> dist a b" .
       
  6400   have "dist x b = norm ((1-u) *\<^sub>R a - (1-u) *\<^sub>R b)"
       
  6401     by (simp add: dist_norm algebra_simps x)
       
  6402   also have "... = (1-u) * dist a b"
       
  6403   proof -
       
  6404     have "norm ((1 - 1 * u) *\<^sub>R (a - b)) = (1 - 1 * u) * norm (a - b)"
       
  6405       using \<open>u \<le> 1\<close> by force
       
  6406     then show ?thesis
       
  6407       by (simp add: dist_norm real_vector.scale_right_diff_distrib)
       
  6408   qed
       
  6409   also have "... \<le> dist a b"
       
  6410     by (simp add: mult_left_le_one_le u)
       
  6411   finally show "dist x b \<le> dist a b" .
       
  6412 qed
       
  6413 
       
  6414 lemma dist_in_open_segment:
       
  6415   fixes a :: "'a :: euclidean_space"
       
  6416   assumes "x \<in> open_segment a b"
       
  6417     shows "dist x a < dist a b \<and> dist x b < dist a b"
       
  6418 proof (intro conjI)
       
  6419   obtain u where u: "0 < u" "u < 1" and x: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b"
       
  6420     using assms by (force simp: in_segment algebra_simps)
       
  6421   have "dist x a = u * dist a b"
       
  6422     apply (simp add: dist_norm algebra_simps x)
       
  6423     by (metis abs_of_nonneg less_eq_real_def norm_minus_commute norm_scaleR real_vector.scale_right_diff_distrib \<open>0 < u\<close>)
       
  6424   also have *: "...  < dist a b"
       
  6425     by (metis (no_types) assms dist_eq_0_iff dist_not_less_zero in_segment(2) linorder_neqE_linordered_idom mult.left_neutral real_mult_less_iff1 \<open>u < 1\<close>)
       
  6426   finally show "dist x a < dist a b" .
       
  6427   have ab_ne0: "dist a b \<noteq> 0"
       
  6428     using * by fastforce
       
  6429   have "dist x b = norm ((1-u) *\<^sub>R a - (1-u) *\<^sub>R b)"
       
  6430     by (simp add: dist_norm algebra_simps x)
       
  6431   also have "... = (1-u) * dist a b"
       
  6432   proof -
       
  6433     have "norm ((1 - 1 * u) *\<^sub>R (a - b)) = (1 - 1 * u) * norm (a - b)"
       
  6434       using \<open>u < 1\<close> by force
       
  6435     then show ?thesis
       
  6436       by (simp add: dist_norm real_vector.scale_right_diff_distrib)
       
  6437   qed
       
  6438   also have "... < dist a b"
       
  6439     using ab_ne0 \<open>0 < u\<close> by simp
       
  6440   finally show "dist x b < dist a b" .
       
  6441 qed
       
  6442 
       
  6443 lemma dist_decreases_open_segment_0:
       
  6444   fixes x :: "'a :: euclidean_space"
       
  6445   assumes "x \<in> open_segment 0 b"
       
  6446     shows "dist c x < dist c 0 \<or> dist c x < dist c b"
       
  6447 proof (rule ccontr, clarsimp simp: not_less)
       
  6448   obtain u where u: "0 \<noteq> b" "0 < u" "u < 1" and x: "x = u *\<^sub>R b"
       
  6449     using assms by (auto simp: in_segment)
       
  6450   have xb: "x \<bullet> b < b \<bullet> b"
       
  6451     using u x by auto
       
  6452   assume "norm c \<le> dist c x"
       
  6453   then have "c \<bullet> c \<le> (c - x) \<bullet> (c - x)"
       
  6454     by (simp add: dist_norm norm_le)
       
  6455   moreover have "0 < x \<bullet> b"
       
  6456     using u x by auto
       
  6457   ultimately have less: "c \<bullet> b < x \<bullet> b"
       
  6458     by (simp add: x algebra_simps inner_commute u)
       
  6459   assume "dist c b \<le> dist c x"
       
  6460   then have "(c - b) \<bullet> (c - b) \<le> (c - x) \<bullet> (c - x)"
       
  6461     by (simp add: dist_norm norm_le)
       
  6462   then have "(b \<bullet> b) * (1 - u*u) \<le> 2 * (b \<bullet> c) * (1-u)"
       
  6463     by (simp add: x algebra_simps inner_commute)
       
  6464   then have "(1+u) * (b \<bullet> b) * (1-u) \<le> 2 * (b \<bullet> c) * (1-u)"
       
  6465     by (simp add: algebra_simps)
       
  6466   then have "(1+u) * (b \<bullet> b) \<le> 2 * (b \<bullet> c)"
       
  6467     using \<open>u < 1\<close> by auto
       
  6468   with xb have "c \<bullet> b \<ge> x \<bullet> b"
       
  6469     by (auto simp: x algebra_simps inner_commute)
       
  6470   with less show False by auto
       
  6471 qed
       
  6472 
       
  6473 proposition dist_decreases_open_segment:
       
  6474   fixes a :: "'a :: euclidean_space"
       
  6475   assumes "x \<in> open_segment a b"
       
  6476     shows "dist c x < dist c a \<or> dist c x < dist c b"
       
  6477 proof -
       
  6478   have *: "x - a \<in> open_segment 0 (b - a)" using assms
       
  6479     by (metis diff_self open_segment_translation_eq uminus_add_conv_diff)
       
  6480   show ?thesis
       
  6481     using dist_decreases_open_segment_0 [OF *, of "c-a"] assms
       
  6482     by (simp add: dist_norm)
       
  6483 qed
       
  6484 
       
  6485 lemma dist_decreases_closed_segment:
       
  6486   fixes a :: "'a :: euclidean_space"
       
  6487   assumes "x \<in> closed_segment a b"
       
  6488     shows "dist c x \<le> dist c a \<or> dist c x \<le> dist c b"
       
  6489 apply (cases "x \<in> open_segment a b")
       
  6490  using dist_decreases_open_segment less_eq_real_def apply blast
       
  6491 by (metis DiffI assms empty_iff insertE open_segment_def order_refl)
       
  6492 
       
  6493 lemma convex_intermediate_ball:
       
  6494   fixes a :: "'a :: euclidean_space"
       
  6495   shows "\<lbrakk>ball a r \<subseteq> T; T \<subseteq> cball a r\<rbrakk> \<Longrightarrow> convex T"
       
  6496 apply (simp add: convex_contains_open_segment, clarify)
       
  6497 by (metis (no_types, hide_lams) less_le_trans mem_ball mem_cball subsetCE dist_decreases_open_segment)
       
  6498 
       
  6499 lemma csegment_midpoint_subset: "closed_segment (midpoint a b) b \<subseteq> closed_segment a b"
       
  6500   apply (clarsimp simp: midpoint_def in_segment)
       
  6501   apply (rule_tac x="(1 + u) / 2" in exI)
       
  6502   apply (auto simp: algebra_simps add_divide_distrib diff_divide_distrib)
       
  6503   by (metis real_sum_of_halves scaleR_left.add)
       
  6504 
       
  6505 lemma notin_segment_midpoint:
       
  6506   fixes a :: "'a :: euclidean_space"
       
  6507   shows "a \<noteq> b \<Longrightarrow> a \<notin> closed_segment (midpoint a b) b"
       
  6508 by (auto simp: dist_midpoint dest!: dist_in_closed_segment)
       
  6509 
       
  6510 subsubsection\<open>More lemmas, especially for working with the underlying formula\<close>
       
  6511 
       
  6512 lemma segment_eq_compose:
       
  6513   fixes a :: "'a :: real_vector"
       
  6514   shows "(\<lambda>u. (1 - u) *\<^sub>R a + u *\<^sub>R b) = (\<lambda>x. a + x) o (\<lambda>u. u *\<^sub>R (b - a))"
       
  6515     by (simp add: o_def algebra_simps)
       
  6516 
       
  6517 lemma segment_degen_1:
       
  6518   fixes a :: "'a :: real_vector"
       
  6519   shows "(1 - u) *\<^sub>R a + u *\<^sub>R b = b \<longleftrightarrow> a=b \<or> u=1"
       
  6520 proof -
       
  6521   { assume "(1 - u) *\<^sub>R a + u *\<^sub>R b = b"
       
  6522     then have "(1 - u) *\<^sub>R a = (1 - u) *\<^sub>R b"
       
  6523       by (simp add: algebra_simps)
       
  6524     then have "a=b \<or> u=1"
       
  6525       by simp
       
  6526   } then show ?thesis
       
  6527       by (auto simp: algebra_simps)
       
  6528 qed
       
  6529 
       
  6530 lemma segment_degen_0:
       
  6531     fixes a :: "'a :: real_vector"
       
  6532     shows "(1 - u) *\<^sub>R a + u *\<^sub>R b = a \<longleftrightarrow> a=b \<or> u=0"
       
  6533   using segment_degen_1 [of "1-u" b a]
       
  6534   by (auto simp: algebra_simps)
       
  6535 
       
  6536 lemma closed_segment_image_interval:
       
  6537      "closed_segment a b = (\<lambda>u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0..1}"
       
  6538   by (auto simp: set_eq_iff image_iff closed_segment_def)
       
  6539 
       
  6540 lemma open_segment_image_interval:
       
  6541      "open_segment a b = (if a=b then {} else (\<lambda>u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0<..<1})"
       
  6542   by (auto simp:  open_segment_def closed_segment_def segment_degen_0 segment_degen_1)
       
  6543 
       
  6544 lemmas segment_image_interval = closed_segment_image_interval open_segment_image_interval
       
  6545 
       
  6546 lemma open_segment_bound1:
       
  6547   assumes "x \<in> open_segment a b"
       
  6548   shows "norm (x - a) < norm (b - a)"
       
  6549 proof -
       
  6550   obtain u where "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 < u" "u < 1" "a \<noteq> b"
       
  6551     using assms by (auto simp add: open_segment_image_interval split: if_split_asm)
       
  6552   then show "norm (x - a) < norm (b - a)"
       
  6553     apply clarify
       
  6554     apply (auto simp: algebra_simps)
       
  6555     apply (simp add: scaleR_diff_right [symmetric])
       
  6556     done
       
  6557 qed
       
  6558 
       
  6559 lemma compact_segment [simp]:
       
  6560   fixes a :: "'a::real_normed_vector"
       
  6561   shows "compact (closed_segment a b)"
       
  6562   by (auto simp: segment_image_interval intro!: compact_continuous_image continuous_intros)
       
  6563 
       
  6564 lemma closed_segment [simp]:
       
  6565   fixes a :: "'a::real_normed_vector"
       
  6566   shows "closed (closed_segment a b)"
       
  6567   by (simp add: compact_imp_closed)
       
  6568 
       
  6569 lemma closure_closed_segment [simp]:
       
  6570   fixes a :: "'a::real_normed_vector"
       
  6571   shows "closure(closed_segment a b) = closed_segment a b"
       
  6572   by simp
       
  6573 
       
  6574 lemma open_segment_bound:
       
  6575   assumes "x \<in> open_segment a b"
       
  6576   shows "norm (x - a) < norm (b - a)" "norm (x - b) < norm (b - a)"
       
  6577 apply (simp add: assms open_segment_bound1)
       
  6578 by (metis assms norm_minus_commute open_segment_bound1 open_segment_commute)
       
  6579 
       
  6580 lemma closure_open_segment [simp]:
       
  6581     fixes a :: "'a::euclidean_space"
       
  6582     shows "closure(open_segment a b) = (if a = b then {} else closed_segment a b)"
       
  6583 proof -
       
  6584   have "closure ((\<lambda>u. u *\<^sub>R (b - a)) ` {0<..<1}) = (\<lambda>u. u *\<^sub>R (b - a)) ` closure {0<..<1}" if "a \<noteq> b"
       
  6585     apply (rule closure_injective_linear_image [symmetric])
       
  6586     apply (simp add:)
       
  6587     using that by (simp add: inj_on_def)
       
  6588   then show ?thesis
       
  6589     by (simp add: segment_image_interval segment_eq_compose closure_greaterThanLessThan [symmetric]
       
  6590          closure_translation image_comp [symmetric] del: closure_greaterThanLessThan)
       
  6591 qed
       
  6592 
       
  6593 lemma closed_open_segment_iff [simp]:
       
  6594     fixes a :: "'a::euclidean_space"  shows "closed(open_segment a b) \<longleftrightarrow> a = b"
       
  6595   by (metis open_segment_def DiffE closure_eq closure_open_segment ends_in_segment(1) insert_iff segment_image_interval(2))
       
  6596 
       
  6597 lemma compact_open_segment_iff [simp]:
       
  6598     fixes a :: "'a::euclidean_space"  shows "compact(open_segment a b) \<longleftrightarrow> a = b"
       
  6599   by (simp add: bounded_open_segment compact_eq_bounded_closed)
       
  6600 
       
  6601 lemma convex_closed_segment [iff]: "convex (closed_segment a b)"
       
  6602   unfolding segment_convex_hull by(rule convex_convex_hull)
       
  6603 
       
  6604 lemma convex_open_segment [iff]: "convex(open_segment a b)"
       
  6605 proof -
       
  6606   have "convex ((\<lambda>u. u *\<^sub>R (b-a)) ` {0<..<1})"
       
  6607     by (rule convex_linear_image) auto
       
  6608   then show ?thesis
       
  6609     apply (simp add: open_segment_image_interval segment_eq_compose)
       
  6610     by (metis image_comp convex_translation)
       
  6611 qed
       
  6612 
       
  6613 lemmas convex_segment = convex_closed_segment convex_open_segment
       
  6614 
       
  6615 lemma connected_segment [iff]:
       
  6616   fixes x :: "'a :: real_normed_vector"
       
  6617   shows "connected (closed_segment x y)"
       
  6618   by (simp add: convex_connected)
       
  6619 
       
  6620 lemma affine_hull_closed_segment [simp]:
       
  6621      "affine hull (closed_segment a b) = affine hull {a,b}"
       
  6622   by (simp add: segment_convex_hull)
       
  6623 
       
  6624 lemma affine_hull_open_segment [simp]:
       
  6625     fixes a :: "'a::euclidean_space"
       
  6626     shows "affine hull (open_segment a b) = (if a = b then {} else affine hull {a,b})"
       
  6627 by (metis affine_hull_convex_hull affine_hull_empty closure_open_segment closure_same_affine_hull segment_convex_hull)
       
  6628 
       
  6629 lemma rel_interior_closure_convex_segment:
       
  6630   fixes S :: "_::euclidean_space set"
       
  6631   assumes "convex S" "a \<in> rel_interior S" "b \<in> closure S"
       
  6632     shows "open_segment a b \<subseteq> rel_interior S"
       
  6633 proof
       
  6634   fix x
       
  6635   have [simp]: "(1 - u) *\<^sub>R a + u *\<^sub>R b = b - (1 - u) *\<^sub>R (b - a)" for u
       
  6636     by (simp add: algebra_simps)
       
  6637   assume "x \<in> open_segment a b"
       
  6638   then show "x \<in> rel_interior S"
       
  6639     unfolding closed_segment_def open_segment_def  using assms
       
  6640     by (auto intro: rel_interior_closure_convex_shrink)
       
  6641 qed
       
  6642 
       
  6643 subsection\<open>More results about segments\<close>
       
  6644 
       
  6645 lemma dist_half_times2:
       
  6646   fixes a :: "'a :: real_normed_vector"
       
  6647   shows "dist ((1 / 2) *\<^sub>R (a + b)) x * 2 = dist (a+b) (2 *\<^sub>R x)"
       
  6648 proof -
       
  6649   have "norm ((1 / 2) *\<^sub>R (a + b) - x) * 2 = norm (2 *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x))"
       
  6650     by simp
       
  6651   also have "... = norm ((a + b) - 2 *\<^sub>R x)"
       
  6652     by (simp add: real_vector.scale_right_diff_distrib)
       
  6653   finally show ?thesis
       
  6654     by (simp only: dist_norm)
       
  6655 qed
       
  6656 
       
  6657 lemma closed_segment_as_ball:
       
  6658     "closed_segment a b = affine hull {a,b} \<inter> cball(inverse 2 *\<^sub>R (a + b))(norm(b - a) / 2)"
       
  6659 proof (cases "b = a")
       
  6660   case True then show ?thesis by (auto simp: hull_inc)
       
  6661 next
       
  6662   case False
       
  6663   then have *: "((\<exists>u v. x = u *\<^sub>R a + v *\<^sub>R b \<and> u + v = 1) \<and>
       
  6664                   dist ((1 / 2) *\<^sub>R (a + b)) x * 2 \<le> norm (b - a)) =
       
  6665                  (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1)" for x
       
  6666   proof -
       
  6667     have "((\<exists>u v. x = u *\<^sub>R a + v *\<^sub>R b \<and> u + v = 1) \<and>
       
  6668                   dist ((1 / 2) *\<^sub>R (a + b)) x * 2 \<le> norm (b - a)) =
       
  6669           ((\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b) \<and>
       
  6670                   dist ((1 / 2) *\<^sub>R (a + b)) x * 2 \<le> norm (b - a))"
       
  6671       unfolding eq_diff_eq [symmetric] by simp
       
  6672     also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and>
       
  6673                           norm ((a+b) - (2 *\<^sub>R x)) \<le> norm (b - a))"
       
  6674       by (simp add: dist_half_times2) (simp add: dist_norm)
       
  6675     also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and>
       
  6676             norm ((a+b) - (2 *\<^sub>R ((1 - u) *\<^sub>R a + u *\<^sub>R b))) \<le> norm (b - a))"
       
  6677       by auto
       
  6678     also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and>
       
  6679                 norm ((1 - u * 2) *\<^sub>R (b - a)) \<le> norm (b - a))"
       
  6680       by (simp add: algebra_simps scaleR_2)
       
  6681     also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and>
       
  6682                           \<bar>1 - u * 2\<bar> * norm (b - a) \<le> norm (b - a))"
       
  6683       by simp
       
  6684     also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> \<bar>1 - u * 2\<bar> \<le> 1)"
       
  6685       by (simp add: mult_le_cancel_right2 False)
       
  6686     also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1)"
       
  6687       by auto
       
  6688     finally show ?thesis .
       
  6689   qed
       
  6690   show ?thesis
       
  6691     by (simp add: affine_hull_2 Set.set_eq_iff closed_segment_def *)
       
  6692 qed
       
  6693 
       
  6694 lemma open_segment_as_ball:
       
  6695     "open_segment a b =
       
  6696      affine hull {a,b} \<inter> ball(inverse 2 *\<^sub>R (a + b))(norm(b - a) / 2)"
       
  6697 proof (cases "b = a")
       
  6698   case True then show ?thesis by (auto simp: hull_inc)
       
  6699 next
       
  6700   case False
       
  6701   then have *: "((\<exists>u v. x = u *\<^sub>R a + v *\<^sub>R b \<and> u + v = 1) \<and>
       
  6702                   dist ((1 / 2) *\<^sub>R (a + b)) x * 2 < norm (b - a)) =
       
  6703                  (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 < u \<and> u < 1)" for x
       
  6704   proof -
       
  6705     have "((\<exists>u v. x = u *\<^sub>R a + v *\<^sub>R b \<and> u + v = 1) \<and>
       
  6706                   dist ((1 / 2) *\<^sub>R (a + b)) x * 2 < norm (b - a)) =
       
  6707           ((\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b) \<and>
       
  6708                   dist ((1 / 2) *\<^sub>R (a + b)) x * 2 < norm (b - a))"
       
  6709       unfolding eq_diff_eq [symmetric] by simp
       
  6710     also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and>
       
  6711                           norm ((a+b) - (2 *\<^sub>R x)) < norm (b - a))"
       
  6712       by (simp add: dist_half_times2) (simp add: dist_norm)
       
  6713     also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and>
       
  6714             norm ((a+b) - (2 *\<^sub>R ((1 - u) *\<^sub>R a + u *\<^sub>R b))) < norm (b - a))"
       
  6715       by auto
       
  6716     also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and>
       
  6717                 norm ((1 - u * 2) *\<^sub>R (b - a)) < norm (b - a))"
       
  6718       by (simp add: algebra_simps scaleR_2)
       
  6719     also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and>
       
  6720                           \<bar>1 - u * 2\<bar> * norm (b - a) < norm (b - a))"
       
  6721       by simp
       
  6722     also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> \<bar>1 - u * 2\<bar> < 1)"
       
  6723       by (simp add: mult_le_cancel_right2 False)
       
  6724     also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 < u \<and> u < 1)"
       
  6725       by auto
       
  6726     finally show ?thesis .
       
  6727   qed
       
  6728   show ?thesis
       
  6729     using False by (force simp: affine_hull_2 Set.set_eq_iff open_segment_image_interval *)
       
  6730 qed
       
  6731 
       
  6732 lemmas segment_as_ball = closed_segment_as_ball open_segment_as_ball
       
  6733 
       
  6734 lemma closed_segment_neq_empty [simp]: "closed_segment a b \<noteq> {}"
       
  6735   by auto
       
  6736 
       
  6737 lemma open_segment_eq_empty [simp]: "open_segment a b = {} \<longleftrightarrow> a = b"
       
  6738 proof -
       
  6739   { assume a1: "open_segment a b = {}"
       
  6740     have "{} \<noteq> {0::real<..<1}"
       
  6741       by simp
       
  6742     then have "a = b"
       
  6743       using a1 open_segment_image_interval by fastforce
       
  6744   } then show ?thesis by auto
       
  6745 qed
       
  6746 
       
  6747 lemma open_segment_eq_empty' [simp]: "{} = open_segment a b \<longleftrightarrow> a = b"
       
  6748   using open_segment_eq_empty by blast
       
  6749 
       
  6750 lemmas segment_eq_empty = closed_segment_neq_empty open_segment_eq_empty
       
  6751 
       
  6752 lemma inj_segment:
       
  6753   fixes a :: "'a :: real_vector"
       
  6754   assumes "a \<noteq> b"
       
  6755     shows "inj_on (\<lambda>u. (1 - u) *\<^sub>R a + u *\<^sub>R b) I"
       
  6756 proof
       
  6757   fix x y
       
  6758   assume "(1 - x) *\<^sub>R a + x *\<^sub>R b = (1 - y) *\<^sub>R a + y *\<^sub>R b"
       
  6759   then have "x *\<^sub>R (b - a) = y *\<^sub>R (b - a)"
       
  6760     by (simp add: algebra_simps)
       
  6761   with assms show "x = y"
       
  6762     by (simp add: real_vector.scale_right_imp_eq)
       
  6763 qed
       
  6764 
       
  6765 lemma finite_closed_segment [simp]: "finite(closed_segment a b) \<longleftrightarrow> a = b"
       
  6766   apply auto
       
  6767   apply (rule ccontr)
       
  6768   apply (simp add: segment_image_interval)
       
  6769   using infinite_Icc [OF zero_less_one] finite_imageD [OF _ inj_segment] apply blast
       
  6770   done
       
  6771 
       
  6772 lemma finite_open_segment [simp]: "finite(open_segment a b) \<longleftrightarrow> a = b"
       
  6773   by (auto simp: open_segment_def)
       
  6774 
       
  6775 lemmas finite_segment = finite_closed_segment finite_open_segment
       
  6776 
       
  6777 lemma closed_segment_eq_sing: "closed_segment a b = {c} \<longleftrightarrow> a = c \<and> b = c"
       
  6778   by auto
       
  6779 
       
  6780 lemma open_segment_eq_sing: "open_segment a b \<noteq> {c}"
       
  6781   by (metis finite_insert finite_open_segment insert_not_empty open_segment_image_interval)
       
  6782 
       
  6783 lemmas segment_eq_sing = closed_segment_eq_sing open_segment_eq_sing
       
  6784 
       
  6785 lemma subset_closed_segment:
       
  6786     "closed_segment a b \<subseteq> closed_segment c d \<longleftrightarrow>
       
  6787      a \<in> closed_segment c d \<and> b \<in> closed_segment c d"
       
  6788   by auto (meson contra_subsetD convex_closed_segment convex_contains_segment)
       
  6789 
       
  6790 lemma subset_co_segment:
       
  6791     "closed_segment a b \<subseteq> open_segment c d \<longleftrightarrow>
       
  6792      a \<in> open_segment c d \<and> b \<in> open_segment c d"
       
  6793 using closed_segment_subset by blast
       
  6794 
       
  6795 lemma subset_open_segment:
       
  6796   fixes a :: "'a::euclidean_space"
       
  6797   shows "open_segment a b \<subseteq> open_segment c d \<longleftrightarrow>
       
  6798          a = b \<or> a \<in> closed_segment c d \<and> b \<in> closed_segment c d"
       
  6799         (is "?lhs = ?rhs")
       
  6800 proof (cases "a = b")
       
  6801   case True then show ?thesis by simp
       
  6802 next
       
  6803   case False show ?thesis
       
  6804   proof
       
  6805     assume rhs: ?rhs
       
  6806     with \<open>a \<noteq> b\<close> have "c \<noteq> d"
       
  6807       using closed_segment_idem singleton_iff by auto
       
  6808     have "\<exists>uc. (1 - u) *\<^sub>R ((1 - ua) *\<^sub>R c + ua *\<^sub>R d) + u *\<^sub>R ((1 - ub) *\<^sub>R c + ub *\<^sub>R d) =
       
  6809                (1 - uc) *\<^sub>R c + uc *\<^sub>R d \<and> 0 < uc \<and> uc < 1"
       
  6810         if neq: "(1 - ua) *\<^sub>R c + ua *\<^sub>R d \<noteq> (1 - ub) *\<^sub>R c + ub *\<^sub>R d" "c \<noteq> d"
       
  6811            and "a = (1 - ua) *\<^sub>R c + ua *\<^sub>R d" "b = (1 - ub) *\<^sub>R c + ub *\<^sub>R d"
       
  6812            and u: "0 < u" "u < 1" and uab: "0 \<le> ua" "ua \<le> 1" "0 \<le> ub" "ub \<le> 1"
       
  6813         for u ua ub
       
  6814     proof -
       
  6815       have "ua \<noteq> ub"
       
  6816         using neq by auto
       
  6817       moreover have "(u - 1) * ua \<le> 0" using u uab
       
  6818         by (simp add: mult_nonpos_nonneg)
       
  6819       ultimately have lt: "(u - 1) * ua < u * ub" using u uab
       
  6820         by (metis antisym_conv diff_ge_0_iff_ge le_less_trans mult_eq_0_iff mult_le_0_iff not_less)
       
  6821       have "p * ua + q * ub < p+q" if p: "0 < p" and  q: "0 < q" for p q
       
  6822       proof -
       
  6823         have "\<not> p \<le> 0" "\<not> q \<le> 0"
       
  6824           using p q not_less by blast+
       
  6825         then show ?thesis
       
  6826           by (metis \<open>ua \<noteq> ub\<close> add_less_cancel_left add_less_cancel_right add_mono_thms_linordered_field(5)
       
  6827                     less_eq_real_def mult_cancel_left1 mult_less_cancel_left2 uab(2) uab(4))
       
  6828       qed
       
  6829       then have "(1 - u) * ua + u * ub < 1" using u \<open>ua \<noteq> ub\<close>
       
  6830         by (metis diff_add_cancel diff_gt_0_iff_gt)
       
  6831       with lt show ?thesis
       
  6832         by (rule_tac x="ua + u*(ub-ua)" in exI) (simp add: algebra_simps)
       
  6833     qed
       
  6834     with rhs \<open>a \<noteq> b\<close> \<open>c \<noteq> d\<close> show ?lhs
       
  6835       unfolding open_segment_image_interval closed_segment_def
       
  6836       by (fastforce simp add:)
       
  6837   next
       
  6838     assume lhs: ?lhs
       
  6839     with \<open>a \<noteq> b\<close> have "c \<noteq> d"
       
  6840       by (meson finite_open_segment rev_finite_subset)
       
  6841     have "closure (open_segment a b) \<subseteq> closure (open_segment c d)"
       
  6842       using lhs closure_mono by blast
       
  6843     then have "closed_segment a b \<subseteq> closed_segment c d"
       
  6844       by (simp add: \<open>a \<noteq> b\<close> \<open>c \<noteq> d\<close>)
       
  6845     then show ?rhs
       
  6846       by (force simp: \<open>a \<noteq> b\<close>)
       
  6847   qed
       
  6848 qed
       
  6849 
       
  6850 lemma subset_oc_segment:
       
  6851   fixes a :: "'a::euclidean_space"
       
  6852   shows "open_segment a b \<subseteq> closed_segment c d \<longleftrightarrow>
       
  6853          a = b \<or> a \<in> closed_segment c d \<and> b \<in> closed_segment c d"
       
  6854 apply (simp add: subset_open_segment [symmetric])
       
  6855 apply (rule iffI)
       
  6856  apply (metis closure_closed_segment closure_mono closure_open_segment subset_closed_segment subset_open_segment)
       
  6857 apply (meson dual_order.trans segment_open_subset_closed)
       
  6858 done
       
  6859 
       
  6860 lemmas subset_segment = subset_closed_segment subset_co_segment subset_oc_segment subset_open_segment
       
  6861 
       
  6862 
       
  6863 subsection\<open>Betweenness\<close>
       
  6864 
       
  6865 lemma between_mem_segment: "between (a,b) x \<longleftrightarrow> x \<in> closed_segment a b"
       
  6866   unfolding between_def by auto
       
  6867 
       
  6868 lemma between: "between (a, b) (x::'a::euclidean_space) \<longleftrightarrow> dist a b = (dist a x) + (dist x b)"
       
  6869 proof (cases "a = b")
       
  6870   case True
       
  6871   then show ?thesis
       
  6872     unfolding between_def split_conv
       
  6873     by (auto simp add: dist_commute)
       
  6874 next
       
  6875   case False
       
  6876   then have Fal: "norm (a - b) \<noteq> 0" and Fal2: "norm (a - b) > 0"
       
  6877     by auto
       
  6878   have *: "\<And>u. a - ((1 - u) *\<^sub>R a + u *\<^sub>R b) = u *\<^sub>R (a - b)"
       
  6879     by (auto simp add: algebra_simps)
       
  6880   show ?thesis
       
  6881     unfolding between_def split_conv closed_segment_def mem_Collect_eq
       
  6882     apply rule
       
  6883     apply (elim exE conjE)
       
  6884     apply (subst dist_triangle_eq)
       
  6885   proof -
       
  6886     fix u
       
  6887     assume as: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \<le> u" "u \<le> 1"
       
  6888     then have *: "a - x = u *\<^sub>R (a - b)" "x - b = (1 - u) *\<^sub>R (a - b)"
       
  6889       unfolding as(1) by (auto simp add:algebra_simps)
       
  6890     show "norm (a - x) *\<^sub>R (x - b) = norm (x - b) *\<^sub>R (a - x)"
       
  6891       unfolding norm_minus_commute[of x a] * using as(2,3)
       
  6892       by (auto simp add: field_simps)
       
  6893   next
       
  6894     assume as: "dist a b = dist a x + dist x b"
       
  6895     have "norm (a - x) / norm (a - b) \<le> 1"
       
  6896       using Fal2 unfolding as[unfolded dist_norm] norm_ge_zero by auto
       
  6897     then show "\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1"
       
  6898       apply (rule_tac x="dist a x / dist a b" in exI)
       
  6899       unfolding dist_norm
       
  6900       apply (subst euclidean_eq_iff)
       
  6901       apply rule
       
  6902       defer
       
  6903       apply rule
       
  6904       prefer 3
       
  6905       apply rule
       
  6906     proof -
       
  6907       fix i :: 'a
       
  6908       assume i: "i \<in> Basis"
       
  6909       have "((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) \<bullet> i =
       
  6910         ((norm (a - b) - norm (a - x)) * (a \<bullet> i) + norm (a - x) * (b \<bullet> i)) / norm (a - b)"
       
  6911         using Fal by (auto simp add: field_simps inner_simps)
       
  6912       also have "\<dots> = x\<bullet>i"
       
  6913         apply (rule divide_eq_imp[OF Fal])
       
  6914         unfolding as[unfolded dist_norm]
       
  6915         using as[unfolded dist_triangle_eq]
       
  6916         apply -
       
  6917         apply (subst (asm) euclidean_eq_iff)
       
  6918         using i
       
  6919         apply (erule_tac x=i in ballE)
       
  6920         apply (auto simp add: field_simps inner_simps)
       
  6921         done
       
  6922       finally show "x \<bullet> i =
       
  6923         ((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) \<bullet> i"
       
  6924         by auto
       
  6925     qed (insert Fal2, auto)
       
  6926   qed
       
  6927 qed
       
  6928 
       
  6929 lemma between_midpoint:
       
  6930   fixes a :: "'a::euclidean_space"
       
  6931   shows "between (a,b) (midpoint a b)" (is ?t1)
       
  6932     and "between (b,a) (midpoint a b)" (is ?t2)
       
  6933 proof -
       
  6934   have *: "\<And>x y z. x = (1/2::real) *\<^sub>R z \<Longrightarrow> y = (1/2) *\<^sub>R z \<Longrightarrow> norm z = norm x + norm y"
       
  6935     by auto
       
  6936   show ?t1 ?t2
       
  6937     unfolding between midpoint_def dist_norm
       
  6938     apply(rule_tac[!] *)
       
  6939     unfolding euclidean_eq_iff[where 'a='a]
       
  6940     apply (auto simp add: field_simps inner_simps)
       
  6941     done
       
  6942 qed
       
  6943 
       
  6944 lemma between_mem_convex_hull:
       
  6945   "between (a,b) x \<longleftrightarrow> x \<in> convex hull {a,b}"
       
  6946   unfolding between_mem_segment segment_convex_hull ..
       
  6947 
       
  6948 
       
  6949 subsection \<open>Shrinking towards the interior of a convex set\<close>
       
  6950 
       
  6951 lemma mem_interior_convex_shrink:
       
  6952   fixes s :: "'a::euclidean_space set"
       
  6953   assumes "convex s"
       
  6954     and "c \<in> interior s"
       
  6955     and "x \<in> s"
       
  6956     and "0 < e"
       
  6957     and "e \<le> 1"
       
  6958   shows "x - e *\<^sub>R (x - c) \<in> interior s"
       
  6959 proof -
       
  6960   obtain d where "d > 0" and d: "ball c d \<subseteq> s"
       
  6961     using assms(2) unfolding mem_interior by auto
       
  6962   show ?thesis
       
  6963     unfolding mem_interior
       
  6964     apply (rule_tac x="e*d" in exI)
       
  6965     apply rule
       
  6966     defer
       
  6967     unfolding subset_eq Ball_def mem_ball
       
  6968   proof (rule, rule)
       
  6969     fix y
       
  6970     assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d"
       
  6971     have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x"
       
  6972       using \<open>e > 0\<close> by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib)
       
  6973     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)"
       
  6974       unfolding dist_norm
       
  6975       unfolding norm_scaleR[symmetric]
       
  6976       apply (rule arg_cong[where f=norm])
       
  6977       using \<open>e > 0\<close>
       
  6978       by (auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps)
       
  6979     also have "\<dots> = \<bar>1/e\<bar> * norm (x - e *\<^sub>R (x - c) - y)"
       
  6980       by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps)
       
  6981     also have "\<dots> < d"
       
  6982       using as[unfolded dist_norm] and \<open>e > 0\<close>
       
  6983       by (auto simp add:pos_divide_less_eq[OF \<open>e > 0\<close>] mult.commute)
       
  6984     finally show "y \<in> s"
       
  6985       apply (subst *)
       
  6986       apply (rule assms(1)[unfolded convex_alt,rule_format])
       
  6987       apply (rule d[unfolded subset_eq,rule_format])
       
  6988       unfolding mem_ball
       
  6989       using assms(3-5)
       
  6990       apply auto
       
  6991       done
       
  6992   qed (insert \<open>e>0\<close> \<open>d>0\<close>, auto)
       
  6993 qed
       
  6994 
       
  6995 lemma mem_interior_closure_convex_shrink:
       
  6996   fixes s :: "'a::euclidean_space set"
       
  6997   assumes "convex s"
       
  6998     and "c \<in> interior s"
       
  6999     and "x \<in> closure s"
       
  7000     and "0 < e"
       
  7001     and "e \<le> 1"
       
  7002   shows "x - e *\<^sub>R (x - c) \<in> interior s"
       
  7003 proof -
       
  7004   obtain d where "d > 0" and d: "ball c d \<subseteq> s"
       
  7005     using assms(2) unfolding mem_interior by auto
       
  7006   have "\<exists>y\<in>s. norm (y - x) * (1 - e) < e * d"
       
  7007   proof (cases "x \<in> s")
       
  7008     case True
       
  7009     then show ?thesis
       
  7010       using \<open>e > 0\<close> \<open>d > 0\<close>
       
  7011       apply (rule_tac bexI[where x=x])
       
  7012       apply (auto)
       
  7013       done
       
  7014   next
       
  7015     case False
       
  7016     then have x: "x islimpt s"
       
  7017       using assms(3)[unfolded closure_def] by auto
       
  7018     show ?thesis
       
  7019     proof (cases "e = 1")
       
  7020       case True
       
  7021       obtain y where "y \<in> s" "y \<noteq> x" "dist y x < 1"
       
  7022         using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto
       
  7023       then show ?thesis
       
  7024         apply (rule_tac x=y in bexI)
       
  7025         unfolding True
       
  7026         using \<open>d > 0\<close>
       
  7027         apply auto
       
  7028         done
       
  7029     next
       
  7030       case False
       
  7031       then have "0 < e * d / (1 - e)" and *: "1 - e > 0"
       
  7032         using \<open>e \<le> 1\<close> \<open>e > 0\<close> \<open>d > 0\<close> by auto
       
  7033       then obtain y where "y \<in> s" "y \<noteq> x" "dist y x < e * d / (1 - e)"
       
  7034         using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto
       
  7035       then show ?thesis
       
  7036         apply (rule_tac x=y in bexI)
       
  7037         unfolding dist_norm
       
  7038         using pos_less_divide_eq[OF *]
       
  7039         apply auto
       
  7040         done
       
  7041     qed
       
  7042   qed
       
  7043   then obtain y where "y \<in> s" and y: "norm (y - x) * (1 - e) < e * d"
       
  7044     by auto
       
  7045   define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)"
       
  7046   have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)"
       
  7047     unfolding z_def using \<open>e > 0\<close>
       
  7048     by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib)
       
  7049   have "z \<in> interior s"
       
  7050     apply (rule interior_mono[OF d,unfolded subset_eq,rule_format])
       
  7051     unfolding interior_open[OF open_ball] mem_ball z_def dist_norm using y and assms(4,5)
       
  7052     apply (auto simp add:field_simps norm_minus_commute)
       
  7053     done
       
  7054   then show ?thesis
       
  7055     unfolding *
       
  7056     apply -
       
  7057     apply (rule mem_interior_convex_shrink)
       
  7058     using assms(1,4-5) \<open>y\<in>s\<close>
       
  7059     apply auto
       
  7060     done
       
  7061 qed
       
  7062 
       
  7063 lemma in_interior_closure_convex_segment:
       
  7064   fixes S :: "'a::euclidean_space set"
       
  7065   assumes "convex S" and a: "a \<in> interior S" and b: "b \<in> closure S"
       
  7066     shows "open_segment a b \<subseteq> interior S"
       
  7067 proof (clarsimp simp: in_segment)
       
  7068   fix u::real
       
  7069   assume u: "0 < u" "u < 1"
       
  7070   have "(1 - u) *\<^sub>R a + u *\<^sub>R b = b - (1 - u) *\<^sub>R (b - a)"
       
  7071     by (simp add: algebra_simps)
       
  7072   also have "... \<in> interior S" using mem_interior_closure_convex_shrink [OF assms] u
       
  7073     by simp
       
  7074   finally show "(1 - u) *\<^sub>R a + u *\<^sub>R b \<in> interior S" .
       
  7075 qed
       
  7076 
       
  7077 
       
  7078 subsection \<open>Some obvious but surprisingly hard simplex lemmas\<close>
       
  7079 
       
  7080 lemma simplex:
       
  7081   assumes "finite s"
       
  7082     and "0 \<notin> s"
       
  7083   shows "convex hull (insert 0 s) =
       
  7084     {y. (\<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s \<le> 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y)}"
       
  7085   unfolding convex_hull_finite[OF finite.insertI[OF assms(1)]]
       
  7086   apply (rule set_eqI, rule)
       
  7087   unfolding mem_Collect_eq
       
  7088   apply (erule_tac[!] exE)
       
  7089   apply (erule_tac[!] conjE)+
       
  7090   unfolding setsum_clauses(2)[OF \<open>finite s\<close>]
       
  7091   apply (rule_tac x=u in exI)
       
  7092   defer
       
  7093   apply (rule_tac x="\<lambda>x. if x = 0 then 1 - setsum u s else u x" in exI)
       
  7094   using assms(2)
       
  7095   unfolding if_smult and setsum_delta_notmem[OF assms(2)]
       
  7096   apply auto
       
  7097   done
       
  7098 
       
  7099 lemma substd_simplex:
       
  7100   assumes d: "d \<subseteq> Basis"
       
  7101   shows "convex hull (insert 0 d) =
       
  7102     {x. (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> (\<Sum>i\<in>d. x\<bullet>i) \<le> 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)}"
       
  7103   (is "convex hull (insert 0 ?p) = ?s")
       
  7104 proof -
       
  7105   let ?D = d
       
  7106   have "0 \<notin> ?p"
       
  7107     using assms by (auto simp: image_def)
       
  7108   from d have "finite d"
       
  7109     by (blast intro: finite_subset finite_Basis)
       
  7110   show ?thesis
       
  7111     unfolding simplex[OF \<open>finite d\<close> \<open>0 \<notin> ?p\<close>]
       
  7112     apply (rule set_eqI)
       
  7113     unfolding mem_Collect_eq
       
  7114     apply rule
       
  7115     apply (elim exE conjE)
       
  7116     apply (erule_tac[2] conjE)+
       
  7117   proof -
       
  7118     fix x :: "'a::euclidean_space"
       
  7119     fix u
       
  7120     assume as: "\<forall>x\<in>?D. 0 \<le> u x" "setsum u ?D \<le> 1" "(\<Sum>x\<in>?D. u x *\<^sub>R x) = x"
       
  7121     have *: "\<forall>i\<in>Basis. i:d \<longrightarrow> u i = x\<bullet>i"
       
  7122       and "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)"
       
  7123       using as(3)
       
  7124       unfolding substdbasis_expansion_unique[OF assms]
       
  7125       by auto
       
  7126     then have **: "setsum u ?D = setsum (op \<bullet> x) ?D"
       
  7127       apply -
       
  7128       apply (rule setsum.cong)
       
  7129       using assms
       
  7130       apply auto
       
  7131       done
       
  7132     have "(\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> setsum (op \<bullet> x) ?D \<le> 1"
       
  7133     proof (rule,rule)
       
  7134       fix i :: 'a
       
  7135       assume i: "i \<in> Basis"
       
  7136       have "i \<in> d \<Longrightarrow> 0 \<le> x\<bullet>i"
       
  7137         unfolding *[rule_format,OF i,symmetric]
       
  7138          apply (rule_tac as(1)[rule_format])
       
  7139          apply auto
       
  7140          done
       
  7141       moreover have "i \<notin> d \<Longrightarrow> 0 \<le> x\<bullet>i"
       
  7142         using \<open>(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)\<close>[rule_format, OF i] by auto
       
  7143       ultimately show "0 \<le> x\<bullet>i" by auto
       
  7144     qed (insert as(2)[unfolded **], auto)
       
  7145     then show "(\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> setsum (op \<bullet> x) ?D \<le> 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)"
       
  7146       using \<open>(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)\<close> by auto
       
  7147   next
       
  7148     fix x :: "'a::euclidean_space"
       
  7149     assume as: "\<forall>i\<in>Basis. 0 \<le> x \<bullet> i" "setsum (op \<bullet> x) ?D \<le> 1" "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)"
       
  7150     show "\<exists>u. (\<forall>x\<in>?D. 0 \<le> u x) \<and> setsum u ?D \<le> 1 \<and> (\<Sum>x\<in>?D. u x *\<^sub>R x) = x"
       
  7151       using as d
       
  7152       unfolding substdbasis_expansion_unique[OF assms]
       
  7153       apply (rule_tac x="inner x" in exI)
       
  7154       apply auto
       
  7155       done
       
  7156   qed
       
  7157 qed
       
  7158 
       
  7159 lemma std_simplex:
       
  7160   "convex hull (insert 0 Basis) =
       
  7161     {x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> setsum (\<lambda>i. x\<bullet>i) Basis \<le> 1}"
       
  7162   using substd_simplex[of Basis] by auto
       
  7163 
       
  7164 lemma interior_std_simplex:
       
  7165   "interior (convex hull (insert 0 Basis)) =
       
  7166     {x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 < x\<bullet>i) \<and> setsum (\<lambda>i. x\<bullet>i) Basis < 1}"
       
  7167   apply (rule set_eqI)
       
  7168   unfolding mem_interior std_simplex
       
  7169   unfolding subset_eq mem_Collect_eq Ball_def mem_ball
       
  7170   unfolding Ball_def[symmetric]
       
  7171   apply rule
       
  7172   apply (elim exE conjE)
       
  7173   defer
       
  7174   apply (erule conjE)
       
  7175 proof -
       
  7176   fix x :: 'a
       
  7177   fix e
       
  7178   assume "e > 0" and as: "\<forall>xa. dist x xa < e \<longrightarrow> (\<forall>x\<in>Basis. 0 \<le> xa \<bullet> x) \<and> setsum (op \<bullet> xa) Basis \<le> 1"
       
  7179   show "(\<forall>xa\<in>Basis. 0 < x \<bullet> xa) \<and> setsum (op \<bullet> x) Basis < 1"
       
  7180     apply safe
       
  7181   proof -
       
  7182     fix i :: 'a
       
  7183     assume i: "i \<in> Basis"
       
  7184     then show "0 < x \<bullet> i"
       
  7185       using as[THEN spec[where x="x - (e / 2) *\<^sub>R i"]] and \<open>e > 0\<close>
       
  7186       unfolding dist_norm
       
  7187       by (auto elim!: ballE[where x=i] simp: inner_simps)
       
  7188   next
       
  7189     have **: "dist x (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis)) < e" using \<open>e > 0\<close>
       
  7190       unfolding dist_norm
       
  7191       by (auto intro!: mult_strict_left_mono simp: SOME_Basis)
       
  7192     have "\<And>i. i \<in> Basis \<Longrightarrow> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis)) \<bullet> i =
       
  7193       x\<bullet>i + (if i = (SOME i. i\<in>Basis) then e/2 else 0)"
       
  7194       by (auto simp: SOME_Basis inner_Basis inner_simps)
       
  7195     then have *: "setsum (op \<bullet> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis))) Basis =
       
  7196       setsum (\<lambda>i. x\<bullet>i + (if (SOME i. i\<in>Basis) = i then e/2 else 0)) Basis"
       
  7197       apply (rule_tac setsum.cong)
       
  7198       apply auto
       
  7199       done
       
  7200     have "setsum (op \<bullet> x) Basis < setsum (op \<bullet> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis))) Basis"
       
  7201       unfolding * setsum.distrib
       
  7202       using \<open>e > 0\<close> DIM_positive[where 'a='a]
       
  7203       apply (subst setsum.delta')
       
  7204       apply (auto simp: SOME_Basis)
       
  7205       done
       
  7206     also have "\<dots> \<le> 1"
       
  7207       using **
       
  7208       apply (drule_tac as[rule_format])
       
  7209       apply auto
       
  7210       done
       
  7211     finally show "setsum (op \<bullet> x) Basis < 1" by auto
       
  7212   qed
       
  7213 next
       
  7214   fix x :: 'a
       
  7215   assume as: "\<forall>i\<in>Basis. 0 < x \<bullet> i" "setsum (op \<bullet> x) Basis < 1"
       
  7216   obtain a :: 'b where "a \<in> UNIV" using UNIV_witness ..
       
  7217   let ?d = "(1 - setsum (op \<bullet> x) Basis) / real (DIM('a))"
       
  7218   have "Min ((op \<bullet> x) ` Basis) > 0"
       
  7219     apply (rule Min_grI)
       
  7220     using as(1)
       
  7221     apply auto
       
  7222     done
       
  7223   moreover have "?d > 0"
       
  7224     using as(2) by (auto simp: Suc_le_eq DIM_positive)
       
  7225   ultimately show "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> (\<forall>i\<in>Basis. 0 \<le> y \<bullet> i) \<and> setsum (op \<bullet> y) Basis \<le> 1"
       
  7226     apply (rule_tac x="min (Min ((op \<bullet> x) ` Basis)) D" for D in exI)
       
  7227     apply rule
       
  7228     defer
       
  7229     apply (rule, rule)
       
  7230   proof -
       
  7231     fix y
       
  7232     assume y: "dist x y < min (Min (op \<bullet> x ` Basis)) ?d"
       
  7233     have "setsum (op \<bullet> y) Basis \<le> setsum (\<lambda>i. x\<bullet>i + ?d) Basis"
       
  7234     proof (rule setsum_mono)
       
  7235       fix i :: 'a
       
  7236       assume i: "i \<in> Basis"
       
  7237       then have "\<bar>y\<bullet>i - x\<bullet>i\<bar> < ?d"
       
  7238         apply -
       
  7239         apply (rule le_less_trans)
       
  7240         using Basis_le_norm[OF i, of "y - x"]
       
  7241         using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2]
       
  7242         apply (auto simp add: norm_minus_commute inner_diff_left)
       
  7243         done
       
  7244       then show "y \<bullet> i \<le> x \<bullet> i + ?d" by auto
       
  7245     qed
       
  7246     also have "\<dots> \<le> 1"
       
  7247       unfolding setsum.distrib setsum_constant
       
  7248       by (auto simp add: Suc_le_eq)
       
  7249     finally show "(\<forall>i\<in>Basis. 0 \<le> y \<bullet> i) \<and> setsum (op \<bullet> y) Basis \<le> 1"
       
  7250     proof safe
       
  7251       fix i :: 'a
       
  7252       assume i: "i \<in> Basis"
       
  7253       have "norm (x - y) < x\<bullet>i"
       
  7254         apply (rule less_le_trans)
       
  7255         apply (rule y[unfolded min_less_iff_conj dist_norm, THEN conjunct1])
       
  7256         using i
       
  7257         apply auto
       
  7258         done
       
  7259       then show "0 \<le> y\<bullet>i"
       
  7260         using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format, OF i]
       
  7261         by (auto simp: inner_simps)
       
  7262     qed
       
  7263   qed auto
       
  7264 qed
       
  7265 
       
  7266 lemma interior_std_simplex_nonempty:
       
  7267   obtains a :: "'a::euclidean_space" where
       
  7268     "a \<in> interior(convex hull (insert 0 Basis))"
       
  7269 proof -
       
  7270   let ?D = "Basis :: 'a set"
       
  7271   let ?a = "setsum (\<lambda>b::'a. inverse (2 * real DIM('a)) *\<^sub>R b) Basis"
       
  7272   {
       
  7273     fix i :: 'a
       
  7274     assume i: "i \<in> Basis"
       
  7275     have "?a \<bullet> i = inverse (2 * real DIM('a))"
       
  7276       by (rule trans[of _ "setsum (\<lambda>j. if i = j then inverse (2 * real DIM('a)) else 0) ?D"])
       
  7277          (simp_all add: setsum.If_cases i) }
       
  7278   note ** = this
       
  7279   show ?thesis
       
  7280     apply (rule that[of ?a])
       
  7281     unfolding interior_std_simplex mem_Collect_eq
       
  7282   proof safe
       
  7283     fix i :: 'a
       
  7284     assume i: "i \<in> Basis"
       
  7285     show "0 < ?a \<bullet> i"
       
  7286       unfolding **[OF i] by (auto simp add: Suc_le_eq DIM_positive)
       
  7287   next
       
  7288     have "setsum (op \<bullet> ?a) ?D = setsum (\<lambda>i. inverse (2 * real DIM('a))) ?D"
       
  7289       apply (rule setsum.cong)
       
  7290       apply rule
       
  7291       apply auto
       
  7292       done
       
  7293     also have "\<dots> < 1"
       
  7294       unfolding setsum_constant divide_inverse[symmetric]
       
  7295       by (auto simp add: field_simps)
       
  7296     finally show "setsum (op \<bullet> ?a) ?D < 1" by auto
       
  7297   qed
       
  7298 qed
       
  7299 
       
  7300 lemma rel_interior_substd_simplex:
       
  7301   assumes d: "d \<subseteq> Basis"
       
  7302   shows "rel_interior (convex hull (insert 0 d)) =
       
  7303     {x::'a::euclidean_space. (\<forall>i\<in>d. 0 < x\<bullet>i) \<and> (\<Sum>i\<in>d. x\<bullet>i) < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)}"
       
  7304   (is "rel_interior (convex hull (insert 0 ?p)) = ?s")
       
  7305 proof -
       
  7306   have "finite d"
       
  7307     apply (rule finite_subset)
       
  7308     using assms
       
  7309     apply auto
       
  7310     done
       
  7311   show ?thesis
       
  7312   proof (cases "d = {}")
       
  7313     case True
       
  7314     then show ?thesis
       
  7315       using rel_interior_sing using euclidean_eq_iff[of _ 0] by auto
       
  7316   next
       
  7317     case False
       
  7318     have h0: "affine hull (convex hull (insert 0 ?p)) =
       
  7319       {x::'a::euclidean_space. (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)}"
       
  7320       using affine_hull_convex_hull affine_hull_substd_basis assms by auto
       
  7321     have aux: "\<And>x::'a. \<forall>i\<in>Basis. (\<forall>i\<in>d. 0 \<le> x\<bullet>i) \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0) \<longrightarrow> 0 \<le> x\<bullet>i"
       
  7322       by auto
       
  7323     {
       
  7324       fix x :: "'a::euclidean_space"
       
  7325       assume x: "x \<in> rel_interior (convex hull (insert 0 ?p))"
       
  7326       then obtain e where e0: "e > 0" and
       
  7327         "ball x e \<inter> {xa. (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> xa\<bullet>i = 0)} \<subseteq> convex hull (insert 0 ?p)"
       
  7328         using mem_rel_interior_ball[of x "convex hull (insert 0 ?p)"] h0 by auto
       
  7329       then have as: "\<forall>xa. dist x xa < e \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> xa\<bullet>i = 0) \<longrightarrow>
       
  7330         (\<forall>i\<in>d. 0 \<le> xa \<bullet> i) \<and> setsum (op \<bullet> xa) d \<le> 1"
       
  7331         unfolding ball_def unfolding substd_simplex[OF assms] using assms by auto
       
  7332       have x0: "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)"
       
  7333         using x rel_interior_subset  substd_simplex[OF assms] by auto
       
  7334       have "(\<forall>i\<in>d. 0 < x \<bullet> i) \<and> setsum (op \<bullet> x) d < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)"
       
  7335         apply rule
       
  7336         apply rule
       
  7337       proof -
       
  7338         fix i :: 'a
       
  7339         assume "i \<in> d"
       
  7340         then have "\<forall>ia\<in>d. 0 \<le> (x - (e / 2) *\<^sub>R i) \<bullet> ia"
       
  7341           apply -
       
  7342           apply (rule as[rule_format,THEN conjunct1])
       
  7343           unfolding dist_norm
       
  7344           using d \<open>e > 0\<close> x0
       
  7345           apply (auto simp: inner_simps inner_Basis)
       
  7346           done
       
  7347         then show "0 < x \<bullet> i"
       
  7348           apply (erule_tac x=i in ballE)
       
  7349           using \<open>e > 0\<close> \<open>i \<in> d\<close> d
       
  7350           apply (auto simp: inner_simps inner_Basis)
       
  7351           done
       
  7352       next
       
  7353         obtain a where a: "a \<in> d"
       
  7354           using \<open>d \<noteq> {}\<close> by auto
       
  7355         then have **: "dist x (x + (e / 2) *\<^sub>R a) < e"
       
  7356           using \<open>e > 0\<close> norm_Basis[of a] d
       
  7357           unfolding dist_norm
       
  7358           by auto
       
  7359         have "\<And>i. i \<in> Basis \<Longrightarrow> (x + (e / 2) *\<^sub>R a) \<bullet> i = x\<bullet>i + (if i = a then e/2 else 0)"
       
  7360           using a d by (auto simp: inner_simps inner_Basis)
       
  7361         then have *: "setsum (op \<bullet> (x + (e / 2) *\<^sub>R a)) d =
       
  7362           setsum (\<lambda>i. x\<bullet>i + (if a = i then e/2 else 0)) d"
       
  7363           using d by (intro setsum.cong) auto
       
  7364         have "a \<in> Basis"
       
  7365           using \<open>a \<in> d\<close> d by auto
       
  7366         then have h1: "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> (x + (e / 2) *\<^sub>R a) \<bullet> i = 0)"
       
  7367           using x0 d \<open>a\<in>d\<close> by (auto simp add: inner_add_left inner_Basis)
       
  7368         have "setsum (op \<bullet> x) d < setsum (op \<bullet> (x + (e / 2) *\<^sub>R a)) d"
       
  7369           unfolding * setsum.distrib
       
  7370           using \<open>e > 0\<close> \<open>a \<in> d\<close>
       
  7371           using \<open>finite d\<close>
       
  7372           by (auto simp add: setsum.delta')
       
  7373         also have "\<dots> \<le> 1"
       
  7374           using ** h1 as[rule_format, of "x + (e / 2) *\<^sub>R a"]
       
  7375           by auto
       
  7376         finally show "setsum (op \<bullet> x) d < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)"
       
  7377           using x0 by auto
       
  7378       qed
       
  7379     }
       
  7380     moreover
       
  7381     {
       
  7382       fix x :: "'a::euclidean_space"
       
  7383       assume as: "x \<in> ?s"
       
  7384       have "\<forall>i. 0 < x\<bullet>i \<or> 0 = x\<bullet>i \<longrightarrow> 0 \<le> x\<bullet>i"
       
  7385         by auto
       
  7386       moreover have "\<forall>i. i \<in> d \<or> i \<notin> d" by auto
       
  7387       ultimately
       
  7388       have "\<forall>i. (\<forall>i\<in>d. 0 < x\<bullet>i) \<and> (\<forall>i. i \<notin> d \<longrightarrow> x\<bullet>i = 0) \<longrightarrow> 0 \<le> x\<bullet>i"
       
  7389         by metis
       
  7390       then have h2: "x \<in> convex hull (insert 0 ?p)"
       
  7391         using as assms
       
  7392         unfolding substd_simplex[OF assms] by fastforce
       
  7393       obtain a where a: "a \<in> d"
       
  7394         using \<open>d \<noteq> {}\<close> by auto
       
  7395       let ?d = "(1 - setsum (op \<bullet> x) d) / real (card d)"
       
  7396       have "0 < card d" using \<open>d \<noteq> {}\<close> \<open>finite d\<close>
       
  7397         by (simp add: card_gt_0_iff)
       
  7398       have "Min ((op \<bullet> x) ` d) > 0"
       
  7399         using as \<open>d \<noteq> {}\<close> \<open>finite d\<close> by (simp add: Min_grI)
       
  7400       moreover have "?d > 0" using as using \<open>0 < card d\<close> by auto
       
  7401       ultimately have h3: "min (Min ((op \<bullet> x) ` d)) ?d > 0"
       
  7402         by auto
       
  7403 
       
  7404       have "x \<in> rel_interior (convex hull (insert 0 ?p))"
       
  7405         unfolding rel_interior_ball mem_Collect_eq h0
       
  7406         apply (rule,rule h2)
       
  7407         unfolding substd_simplex[OF assms]
       
  7408         apply (rule_tac x="min (Min ((op \<bullet> x) ` d)) ?d" in exI)
       
  7409         apply (rule, rule h3)
       
  7410         apply safe
       
  7411         unfolding mem_ball
       
  7412       proof -
       
  7413         fix y :: 'a
       
  7414         assume y: "dist x y < min (Min (op \<bullet> x ` d)) ?d"
       
  7415         assume y2: "\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> y\<bullet>i = 0"
       
  7416         have "setsum (op \<bullet> y) d \<le> setsum (\<lambda>i. x\<bullet>i + ?d) d"
       
  7417         proof (rule setsum_mono)
       
  7418           fix i
       
  7419           assume "i \<in> d"
       
  7420           with d have i: "i \<in> Basis"
       
  7421             by auto
       
  7422           have "\<bar>y\<bullet>i - x\<bullet>i\<bar> < ?d"
       
  7423             apply (rule le_less_trans)
       
  7424             using Basis_le_norm[OF i, of "y - x"]
       
  7425             using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2]
       
  7426             apply (auto simp add: norm_minus_commute inner_simps)
       
  7427             done
       
  7428           then show "y \<bullet> i \<le> x \<bullet> i + ?d" by auto
       
  7429         qed
       
  7430         also have "\<dots> \<le> 1"
       
  7431           unfolding setsum.distrib setsum_constant  using \<open>0 < card d\<close>
       
  7432           by auto
       
  7433         finally show "setsum (op \<bullet> y) d \<le> 1" .
       
  7434 
       
  7435         fix i :: 'a
       
  7436         assume i: "i \<in> Basis"
       
  7437         then show "0 \<le> y\<bullet>i"
       
  7438         proof (cases "i\<in>d")
       
  7439           case True
       
  7440           have "norm (x - y) < x\<bullet>i"
       
  7441             using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1]
       
  7442             using Min_gr_iff[of "op \<bullet> x ` d" "norm (x - y)"] \<open>0 < card d\<close> \<open>i:d\<close>
       
  7443             by (simp add: card_gt_0_iff)
       
  7444           then show "0 \<le> y\<bullet>i"
       
  7445             using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format]
       
  7446             by (auto simp: inner_simps)
       
  7447         qed (insert y2, auto)
       
  7448       qed
       
  7449     }
       
  7450     ultimately have
       
  7451       "\<And>x. x \<in> rel_interior (convex hull insert 0 d) \<longleftrightarrow>
       
  7452         x \<in> {x. (\<forall>i\<in>d. 0 < x \<bullet> i) \<and> setsum (op \<bullet> x) d < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)}"
       
  7453       by blast
       
  7454     then show ?thesis by (rule set_eqI)
       
  7455   qed
       
  7456 qed
       
  7457 
       
  7458 lemma rel_interior_substd_simplex_nonempty:
       
  7459   assumes "d \<noteq> {}"
       
  7460     and "d \<subseteq> Basis"
       
  7461   obtains a :: "'a::euclidean_space"
       
  7462     where "a \<in> rel_interior (convex hull (insert 0 d))"
       
  7463 proof -
       
  7464   let ?D = d
       
  7465   let ?a = "setsum (\<lambda>b::'a::euclidean_space. inverse (2 * real (card d)) *\<^sub>R b) ?D"
       
  7466   have "finite d"
       
  7467     apply (rule finite_subset)
       
  7468     using assms(2)
       
  7469     apply auto
       
  7470     done
       
  7471   then have d1: "0 < real (card d)"
       
  7472     using \<open>d \<noteq> {}\<close> by auto
       
  7473   {
       
  7474     fix i
       
  7475     assume "i \<in> d"
       
  7476     have "?a \<bullet> i = inverse (2 * real (card d))"
       
  7477       apply (rule trans[of _ "setsum (\<lambda>j. if i = j then inverse (2 * real (card d)) else 0) ?D"])
       
  7478       unfolding inner_setsum_left
       
  7479       apply (rule setsum.cong)
       
  7480       using \<open>i \<in> d\<close> \<open>finite d\<close> setsum.delta'[of d i "(\<lambda>k. inverse (2 * real (card d)))"]
       
  7481         d1 assms(2)
       
  7482       by (auto simp: inner_Basis set_rev_mp[OF _ assms(2)])
       
  7483   }
       
  7484   note ** = this
       
  7485   show ?thesis
       
  7486     apply (rule that[of ?a])
       
  7487     unfolding rel_interior_substd_simplex[OF assms(2)] mem_Collect_eq
       
  7488   proof safe
       
  7489     fix i
       
  7490     assume "i \<in> d"
       
  7491     have "0 < inverse (2 * real (card d))"
       
  7492       using d1 by auto
       
  7493     also have "\<dots> = ?a \<bullet> i" using **[of i] \<open>i \<in> d\<close>
       
  7494       by auto
       
  7495     finally show "0 < ?a \<bullet> i" by auto
       
  7496   next
       
  7497     have "setsum (op \<bullet> ?a) ?D = setsum (\<lambda>i. inverse (2 * real (card d))) ?D"
       
  7498       by (rule setsum.cong) (rule refl, rule **)
       
  7499     also have "\<dots> < 1"
       
  7500       unfolding setsum_constant divide_real_def[symmetric]
       
  7501       by (auto simp add: field_simps)
       
  7502     finally show "setsum (op \<bullet> ?a) ?D < 1" by auto
       
  7503   next
       
  7504     fix i
       
  7505     assume "i \<in> Basis" and "i \<notin> d"
       
  7506     have "?a \<in> span d"
       
  7507     proof (rule span_setsum[of d "(\<lambda>b. b /\<^sub>R (2 * real (card d)))" d])
       
  7508       {
       
  7509         fix x :: "'a::euclidean_space"
       
  7510         assume "x \<in> d"
       
  7511         then have "x \<in> span d"
       
  7512           using span_superset[of _ "d"] by auto
       
  7513         then have "x /\<^sub>R (2 * real (card d)) \<in> span d"
       
  7514           using span_mul[of x "d" "(inverse (real (card d)) / 2)"] by auto
       
  7515       }
       
  7516       then show "\<And>x. x\<in>d \<Longrightarrow> x /\<^sub>R (2 * real (card d)) \<in> span d"
       
  7517         by auto
       
  7518     qed
       
  7519     then show "?a \<bullet> i = 0 "
       
  7520       using \<open>i \<notin> d\<close> unfolding span_substd_basis[OF assms(2)] using \<open>i \<in> Basis\<close> by auto
       
  7521   qed
       
  7522 qed
       
  7523 
       
  7524 
       
  7525 subsection \<open>Relative interior of convex set\<close>
       
  7526 
       
  7527 lemma rel_interior_convex_nonempty_aux:
       
  7528   fixes S :: "'n::euclidean_space set"
       
  7529   assumes "convex S"
       
  7530     and "0 \<in> S"
       
  7531   shows "rel_interior S \<noteq> {}"
       
  7532 proof (cases "S = {0}")
       
  7533   case True
       
  7534   then show ?thesis using rel_interior_sing by auto
       
  7535 next
       
  7536   case False
       
  7537   obtain B where B: "independent B \<and> B \<le> S \<and> S \<le> span B \<and> card B = dim S"
       
  7538     using basis_exists[of S] by auto
       
  7539   then have "B \<noteq> {}"
       
  7540     using B assms \<open>S \<noteq> {0}\<close> span_empty by auto
       
  7541   have "insert 0 B \<le> span B"
       
  7542     using subspace_span[of B] subspace_0[of "span B"] span_inc by auto
       
  7543   then have "span (insert 0 B) \<le> span B"
       
  7544     using span_span[of B] span_mono[of "insert 0 B" "span B"] by blast
       
  7545   then have "convex hull insert 0 B \<le> span B"
       
  7546     using convex_hull_subset_span[of "insert 0 B"] by auto
       
  7547   then have "span (convex hull insert 0 B) \<le> span B"
       
  7548     using span_span[of B] span_mono[of "convex hull insert 0 B" "span B"] by blast
       
  7549   then have *: "span (convex hull insert 0 B) = span B"
       
  7550     using span_mono[of B "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto
       
  7551   then have "span (convex hull insert 0 B) = span S"
       
  7552     using B span_mono[of B S] span_mono[of S "span B"] span_span[of B] by auto
       
  7553   moreover have "0 \<in> affine hull (convex hull insert 0 B)"
       
  7554     using hull_subset[of "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto
       
  7555   ultimately have **: "affine hull (convex hull insert 0 B) = affine hull S"
       
  7556     using affine_hull_span_0[of "convex hull insert 0 B"] affine_hull_span_0[of "S"]
       
  7557       assms hull_subset[of S]
       
  7558     by auto
       
  7559   obtain d and f :: "'n \<Rightarrow> 'n" where
       
  7560     fd: "card d = card B" "linear f" "f ` B = d"
       
  7561       "f ` span B = {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = (0::real)} \<and> inj_on f (span B)"
       
  7562     and d: "d \<subseteq> Basis"
       
  7563     using basis_to_substdbasis_subspace_isomorphism[of B,OF _ ] B by auto
       
  7564   then have "bounded_linear f"
       
  7565     using linear_conv_bounded_linear by auto
       
  7566   have "d \<noteq> {}"
       
  7567     using fd B \<open>B \<noteq> {}\<close> by auto
       
  7568   have "insert 0 d = f ` (insert 0 B)"
       
  7569     using fd linear_0 by auto
       
  7570   then have "(convex hull (insert 0 d)) = f ` (convex hull (insert 0 B))"
       
  7571     using convex_hull_linear_image[of f "(insert 0 d)"]
       
  7572       convex_hull_linear_image[of f "(insert 0 B)"] \<open>linear f\<close>
       
  7573     by auto
       
  7574   moreover have "rel_interior (f ` (convex hull insert 0 B)) =
       
  7575     f ` rel_interior (convex hull insert 0 B)"
       
  7576     apply (rule  rel_interior_injective_on_span_linear_image[of f "(convex hull insert 0 B)"])
       
  7577     using \<open>bounded_linear f\<close> fd *
       
  7578     apply auto
       
  7579     done
       
  7580   ultimately have "rel_interior (convex hull insert 0 B) \<noteq> {}"
       
  7581     using rel_interior_substd_simplex_nonempty[OF \<open>d \<noteq> {}\<close> d]
       
  7582     apply auto
       
  7583     apply blast
       
  7584     done
       
  7585   moreover have "convex hull (insert 0 B) \<subseteq> S"
       
  7586     using B assms hull_mono[of "insert 0 B" "S" "convex"] convex_hull_eq
       
  7587     by auto
       
  7588   ultimately show ?thesis
       
  7589     using subset_rel_interior[of "convex hull insert 0 B" S] ** by auto
       
  7590 qed
       
  7591 
       
  7592 lemma rel_interior_eq_empty:
       
  7593   fixes S :: "'n::euclidean_space set"
       
  7594   assumes "convex S"
       
  7595   shows "rel_interior S = {} \<longleftrightarrow> S = {}"
       
  7596 proof -
       
  7597   {
       
  7598     assume "S \<noteq> {}"
       
  7599     then obtain a where "a \<in> S" by auto
       
  7600     then have "0 \<in> op + (-a) ` S"
       
  7601       using assms exI[of "(\<lambda>x. x \<in> S \<and> - a + x = 0)" a] by auto
       
  7602     then have "rel_interior (op + (-a) ` S) \<noteq> {}"
       
  7603       using rel_interior_convex_nonempty_aux[of "op + (-a) ` S"]
       
  7604         convex_translation[of S "-a"] assms
       
  7605       by auto
       
  7606     then have "rel_interior S \<noteq> {}"
       
  7607       using rel_interior_translation by auto
       
  7608   }
       
  7609   then show ?thesis
       
  7610     using rel_interior_empty by auto
       
  7611 qed
       
  7612 
       
  7613 lemma convex_rel_interior:
       
  7614   fixes S :: "'n::euclidean_space set"
       
  7615   assumes "convex S"
       
  7616   shows "convex (rel_interior S)"
       
  7617 proof -
       
  7618   {
       
  7619     fix x y and u :: real
       
  7620     assume assm: "x \<in> rel_interior S" "y \<in> rel_interior S" "0 \<le> u" "u \<le> 1"
       
  7621     then have "x \<in> S"
       
  7622       using rel_interior_subset by auto
       
  7623     have "x - u *\<^sub>R (x-y) \<in> rel_interior S"
       
  7624     proof (cases "0 = u")
       
  7625       case False
       
  7626       then have "0 < u" using assm by auto
       
  7627       then show ?thesis
       
  7628         using assm rel_interior_convex_shrink[of S y x u] assms \<open>x \<in> S\<close> by auto
       
  7629     next
       
  7630       case True
       
  7631       then show ?thesis using assm by auto
       
  7632     qed
       
  7633     then have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> rel_interior S"
       
  7634       by (simp add: algebra_simps)
       
  7635   }
       
  7636   then show ?thesis
       
  7637     unfolding convex_alt by auto
       
  7638 qed
       
  7639 
       
  7640 lemma convex_closure_rel_interior:
       
  7641   fixes S :: "'n::euclidean_space set"
       
  7642   assumes "convex S"
       
  7643   shows "closure (rel_interior S) = closure S"
       
  7644 proof -
       
  7645   have h1: "closure (rel_interior S) \<le> closure S"
       
  7646     using closure_mono[of "rel_interior S" S] rel_interior_subset[of S] by auto
       
  7647   show ?thesis
       
  7648   proof (cases "S = {}")
       
  7649     case False
       
  7650     then obtain a where a: "a \<in> rel_interior S"
       
  7651       using rel_interior_eq_empty assms by auto
       
  7652     { fix x
       
  7653       assume x: "x \<in> closure S"
       
  7654       {
       
  7655         assume "x = a"
       
  7656         then have "x \<in> closure (rel_interior S)"
       
  7657           using a unfolding closure_def by auto
       
  7658       }
       
  7659       moreover
       
  7660       {
       
  7661         assume "x \<noteq> a"
       
  7662          {
       
  7663            fix e :: real
       
  7664            assume "e > 0"
       
  7665            define e1 where "e1 = min 1 (e/norm (x - a))"
       
  7666            then have e1: "e1 > 0" "e1 \<le> 1" "e1 * norm (x - a) \<le> e"
       
  7667              using \<open>x \<noteq> a\<close> \<open>e > 0\<close> le_divide_eq[of e1 e "norm (x - a)"]
       
  7668              by simp_all
       
  7669            then have *: "x - e1 *\<^sub>R (x - a) : rel_interior S"
       
  7670              using rel_interior_closure_convex_shrink[of S a x e1] assms x a e1_def
       
  7671              by auto
       
  7672            have "\<exists>y. y \<in> rel_interior S \<and> y \<noteq> x \<and> dist y x \<le> e"
       
  7673               apply (rule_tac x="x - e1 *\<^sub>R (x - a)" in exI)
       
  7674               using * e1 dist_norm[of "x - e1 *\<^sub>R (x - a)" x] \<open>x \<noteq> a\<close>
       
  7675               apply simp
       
  7676               done
       
  7677         }
       
  7678         then have "x islimpt rel_interior S"
       
  7679           unfolding islimpt_approachable_le by auto
       
  7680         then have "x \<in> closure(rel_interior S)"
       
  7681           unfolding closure_def by auto
       
  7682       }
       
  7683       ultimately have "x \<in> closure(rel_interior S)" by auto
       
  7684     }
       
  7685     then show ?thesis using h1 by auto
       
  7686   next
       
  7687     case True
       
  7688     then have "rel_interior S = {}"
       
  7689       using rel_interior_empty by auto
       
  7690     then have "closure (rel_interior S) = {}"
       
  7691       using closure_empty by auto
       
  7692     with True show ?thesis by auto
       
  7693   qed
       
  7694 qed
       
  7695 
       
  7696 lemma rel_interior_same_affine_hull:
       
  7697   fixes S :: "'n::euclidean_space set"
       
  7698   assumes "convex S"
       
  7699   shows "affine hull (rel_interior S) = affine hull S"
       
  7700   by (metis assms closure_same_affine_hull convex_closure_rel_interior)
       
  7701 
       
  7702 lemma rel_interior_aff_dim:
       
  7703   fixes S :: "'n::euclidean_space set"
       
  7704   assumes "convex S"
       
  7705   shows "aff_dim (rel_interior S) = aff_dim S"
       
  7706   by (metis aff_dim_affine_hull2 assms rel_interior_same_affine_hull)
       
  7707 
       
  7708 lemma rel_interior_rel_interior:
       
  7709   fixes S :: "'n::euclidean_space set"
       
  7710   assumes "convex S"
       
  7711   shows "rel_interior (rel_interior S) = rel_interior S"
       
  7712 proof -
       
  7713   have "openin (subtopology euclidean (affine hull (rel_interior S))) (rel_interior S)"
       
  7714     using openin_rel_interior[of S] rel_interior_same_affine_hull[of S] assms by auto
       
  7715   then show ?thesis
       
  7716     using rel_interior_def by auto
       
  7717 qed
       
  7718 
       
  7719 lemma rel_interior_rel_open:
       
  7720   fixes S :: "'n::euclidean_space set"
       
  7721   assumes "convex S"
       
  7722   shows "rel_open (rel_interior S)"
       
  7723   unfolding rel_open_def using rel_interior_rel_interior assms by auto
       
  7724 
       
  7725 lemma convex_rel_interior_closure_aux:
       
  7726   fixes x y z :: "'n::euclidean_space"
       
  7727   assumes "0 < a" "0 < b" "(a + b) *\<^sub>R z = a *\<^sub>R x + b *\<^sub>R y"
       
  7728   obtains e where "0 < e" "e \<le> 1" "z = y - e *\<^sub>R (y - x)"
       
  7729 proof -
       
  7730   define e where "e = a / (a + b)"
       
  7731   have "z = (1 / (a + b)) *\<^sub>R ((a + b) *\<^sub>R z)"
       
  7732     apply auto
       
  7733     using assms
       
  7734     apply simp
       
  7735     done
       
  7736   also have "\<dots> = (1 / (a + b)) *\<^sub>R (a *\<^sub>R x + b *\<^sub>R y)"
       
  7737     using assms scaleR_cancel_left[of "1/(a+b)" "(a + b) *\<^sub>R z" "a *\<^sub>R x + b *\<^sub>R y"]
       
  7738     by auto
       
  7739   also have "\<dots> = y - e *\<^sub>R (y-x)"
       
  7740     using e_def
       
  7741     apply (simp add: algebra_simps)
       
  7742     using scaleR_left_distrib[of "a/(a+b)" "b/(a+b)" y] assms add_divide_distrib[of a b "a+b"]
       
  7743     apply auto
       
  7744     done
       
  7745   finally have "z = y - e *\<^sub>R (y-x)"
       
  7746     by auto
       
  7747   moreover have "e > 0" using e_def assms by auto
       
  7748   moreover have "e \<le> 1" using e_def assms by auto
       
  7749   ultimately show ?thesis using that[of e] by auto
       
  7750 qed
       
  7751 
       
  7752 lemma convex_rel_interior_closure:
       
  7753   fixes S :: "'n::euclidean_space set"
       
  7754   assumes "convex S"
       
  7755   shows "rel_interior (closure S) = rel_interior S"
       
  7756 proof (cases "S = {}")
       
  7757   case True
       
  7758   then show ?thesis
       
  7759     using assms rel_interior_eq_empty by auto
       
  7760 next
       
  7761   case False
       
  7762   have "rel_interior (closure S) \<supseteq> rel_interior S"
       
  7763     using subset_rel_interior[of S "closure S"] closure_same_affine_hull closure_subset
       
  7764     by auto
       
  7765   moreover
       
  7766   {
       
  7767     fix z
       
  7768     assume z: "z \<in> rel_interior (closure S)"
       
  7769     obtain x where x: "x \<in> rel_interior S"
       
  7770       using \<open>S \<noteq> {}\<close> assms rel_interior_eq_empty by auto
       
  7771     have "z \<in> rel_interior S"
       
  7772     proof (cases "x = z")
       
  7773       case True
       
  7774       then show ?thesis using x by auto
       
  7775     next
       
  7776       case False
       
  7777       obtain e where e: "e > 0" "cball z e \<inter> affine hull closure S \<le> closure S"
       
  7778         using z rel_interior_cball[of "closure S"] by auto
       
  7779       hence *: "0 < e/norm(z-x)" using e False by auto
       
  7780       define y where "y = z + (e/norm(z-x)) *\<^sub>R (z-x)"
       
  7781       have yball: "y \<in> cball z e"
       
  7782         using mem_cball y_def dist_norm[of z y] e by auto
       
  7783       have "x \<in> affine hull closure S"
       
  7784         using x rel_interior_subset_closure hull_inc[of x "closure S"] by blast
       
  7785       moreover have "z \<in> affine hull closure S"
       
  7786         using z rel_interior_subset hull_subset[of "closure S"] by blast
       
  7787       ultimately have "y \<in> affine hull closure S"
       
  7788         using y_def affine_affine_hull[of "closure S"]
       
  7789           mem_affine_3_minus [of "affine hull closure S" z z x "e/norm(z-x)"] by auto
       
  7790       then have "y \<in> closure S" using e yball by auto
       
  7791       have "(1 + (e/norm(z-x))) *\<^sub>R z = (e/norm(z-x)) *\<^sub>R x + y"
       
  7792         using y_def by (simp add: algebra_simps)
       
  7793       then obtain e1 where "0 < e1" "e1 \<le> 1" "z = y - e1 *\<^sub>R (y - x)"
       
  7794         using * convex_rel_interior_closure_aux[of "e / norm (z - x)" 1 z x y]
       
  7795         by (auto simp add: algebra_simps)
       
  7796       then show ?thesis
       
  7797         using rel_interior_closure_convex_shrink assms x \<open>y \<in> closure S\<close>
       
  7798         by auto
       
  7799     qed
       
  7800   }
       
  7801   ultimately show ?thesis by auto
       
  7802 qed
       
  7803 
       
  7804 lemma convex_interior_closure:
       
  7805   fixes S :: "'n::euclidean_space set"
       
  7806   assumes "convex S"
       
  7807   shows "interior (closure S) = interior S"
       
  7808   using closure_aff_dim[of S] interior_rel_interior_gen[of S]
       
  7809     interior_rel_interior_gen[of "closure S"]
       
  7810     convex_rel_interior_closure[of S] assms
       
  7811   by auto
       
  7812 
       
  7813 lemma closure_eq_rel_interior_eq:
       
  7814   fixes S1 S2 :: "'n::euclidean_space set"
       
  7815   assumes "convex S1"
       
  7816     and "convex S2"
       
  7817   shows "closure S1 = closure S2 \<longleftrightarrow> rel_interior S1 = rel_interior S2"
       
  7818   by (metis convex_rel_interior_closure convex_closure_rel_interior assms)
       
  7819 
       
  7820 lemma closure_eq_between:
       
  7821   fixes S1 S2 :: "'n::euclidean_space set"
       
  7822   assumes "convex S1"
       
  7823     and "convex S2"
       
  7824   shows "closure S1 = closure S2 \<longleftrightarrow> rel_interior S1 \<le> S2 \<and> S2 \<subseteq> closure S1"
       
  7825   (is "?A \<longleftrightarrow> ?B")
       
  7826 proof
       
  7827   assume ?A
       
  7828   then show ?B
       
  7829     by (metis assms closure_subset convex_rel_interior_closure rel_interior_subset)
       
  7830 next
       
  7831   assume ?B
       
  7832   then have "closure S1 \<subseteq> closure S2"
       
  7833     by (metis assms(1) convex_closure_rel_interior closure_mono)
       
  7834   moreover from \<open>?B\<close> have "closure S1 \<supseteq> closure S2"
       
  7835     by (metis closed_closure closure_minimal)
       
  7836   ultimately show ?A ..
       
  7837 qed
       
  7838 
       
  7839 lemma open_inter_closure_rel_interior:
       
  7840   fixes S A :: "'n::euclidean_space set"
       
  7841   assumes "convex S"
       
  7842     and "open A"
       
  7843   shows "A \<inter> closure S = {} \<longleftrightarrow> A \<inter> rel_interior S = {}"
       
  7844   by (metis assms convex_closure_rel_interior open_Int_closure_eq_empty)
       
  7845 
       
  7846 lemma rel_interior_open_segment:
       
  7847   fixes a :: "'a :: euclidean_space"
       
  7848   shows "rel_interior(open_segment a b) = open_segment a b"
       
  7849 proof (cases "a = b")
       
  7850   case True then show ?thesis by auto
       
  7851 next
       
  7852   case False then show ?thesis
       
  7853     apply (simp add: rel_interior_eq openin_open)
       
  7854     apply (rule_tac x="ball (inverse 2 *\<^sub>R (a + b)) (norm(b - a) / 2)" in exI)
       
  7855     apply (simp add: open_segment_as_ball)
       
  7856     done
       
  7857 qed
       
  7858 
       
  7859 lemma rel_interior_closed_segment:
       
  7860   fixes a :: "'a :: euclidean_space"
       
  7861   shows "rel_interior(closed_segment a b) =
       
  7862          (if a = b then {a} else open_segment a b)"
       
  7863 proof (cases "a = b")
       
  7864   case True then show ?thesis by auto
       
  7865 next
       
  7866   case False then show ?thesis
       
  7867     by simp
       
  7868        (metis closure_open_segment convex_open_segment convex_rel_interior_closure
       
  7869               rel_interior_open_segment)
       
  7870 qed
       
  7871 
       
  7872 lemmas rel_interior_segment = rel_interior_closed_segment rel_interior_open_segment
       
  7873 
       
  7874 lemma starlike_convex_tweak_boundary_points:
       
  7875   fixes S :: "'a::euclidean_space set"
       
  7876   assumes "convex S" "S \<noteq> {}" and ST: "rel_interior S \<subseteq> T" and TS: "T \<subseteq> closure S"
       
  7877   shows "starlike T"
       
  7878 proof -
       
  7879   have "rel_interior S \<noteq> {}"
       
  7880     by (simp add: assms rel_interior_eq_empty)
       
  7881   then obtain a where a: "a \<in> rel_interior S"  by blast
       
  7882   with ST have "a \<in> T"  by blast
       
  7883   have *: "\<And>x. x \<in> T \<Longrightarrow> open_segment a x \<subseteq> rel_interior S"
       
  7884     apply (rule rel_interior_closure_convex_segment [OF \<open>convex S\<close> a])
       
  7885     using assms by blast
       
  7886   show ?thesis
       
  7887     unfolding starlike_def
       
  7888     apply (rule bexI [OF _ \<open>a \<in> T\<close>])
       
  7889     apply (simp add: closed_segment_eq_open)
       
  7890     apply (intro conjI ballI a \<open>a \<in> T\<close> rel_interior_closure_convex_segment [OF \<open>convex S\<close> a])
       
  7891     apply (simp add: order_trans [OF * ST])
       
  7892     done
       
  7893 qed
       
  7894 
       
  7895 subsection\<open>The relative frontier of a set\<close>
       
  7896 
       
  7897 definition "rel_frontier S = closure S - rel_interior S"
       
  7898 
       
  7899 lemma rel_frontier_empty [simp]: "rel_frontier {} = {}"
       
  7900   by (simp add: rel_frontier_def)
       
  7901 
       
  7902 lemma rel_frontier_eq_empty:
       
  7903     fixes S :: "'n::euclidean_space set"
       
  7904     shows "rel_frontier S = {} \<longleftrightarrow> affine S"
       
  7905   apply (simp add: rel_frontier_def)
       
  7906   apply (simp add: rel_interior_eq_closure [symmetric])
       
  7907   using rel_interior_subset_closure by blast
       
  7908 
       
  7909 lemma rel_frontier_sing [simp]:
       
  7910     fixes a :: "'n::euclidean_space"
       
  7911     shows "rel_frontier {a} = {}"
       
  7912   by (simp add: rel_frontier_def)
       
  7913 
       
  7914 lemma rel_frontier_affine_hull:
       
  7915   fixes S :: "'a::euclidean_space set"
       
  7916   shows "rel_frontier S \<subseteq> affine hull S"
       
  7917 using closure_affine_hull rel_frontier_def by fastforce
       
  7918 
       
  7919 lemma rel_frontier_cball [simp]:
       
  7920     fixes a :: "'n::euclidean_space"
       
  7921     shows "rel_frontier(cball a r) = (if r = 0 then {} else sphere a r)"
       
  7922 proof (cases rule: linorder_cases [of r 0])
       
  7923   case less then show ?thesis
       
  7924     by (force simp: sphere_def)
       
  7925 next
       
  7926   case equal then show ?thesis by simp
       
  7927 next
       
  7928   case greater then show ?thesis
       
  7929     apply simp
       
  7930     by (metis centre_in_ball empty_iff frontier_cball frontier_def interior_cball interior_rel_interior_gen rel_frontier_def)
       
  7931 qed
       
  7932 
       
  7933 lemma rel_frontier_translation:
       
  7934   fixes a :: "'a::euclidean_space"
       
  7935   shows "rel_frontier((\<lambda>x. a + x) ` S) = (\<lambda>x. a + x) ` (rel_frontier S)"
       
  7936 by (simp add: rel_frontier_def translation_diff rel_interior_translation closure_translation)
       
  7937 
       
  7938 lemma closed_affine_hull [iff]:
       
  7939   fixes S :: "'n::euclidean_space set"
       
  7940   shows "closed (affine hull S)"
       
  7941   by (metis affine_affine_hull affine_closed)
       
  7942 
       
  7943 lemma rel_frontier_nonempty_interior:
       
  7944   fixes S :: "'n::euclidean_space set"
       
  7945   shows "interior S \<noteq> {} \<Longrightarrow> rel_frontier S = frontier S"
       
  7946 by (metis frontier_def interior_rel_interior_gen rel_frontier_def)
       
  7947 
       
  7948 lemma rel_frontier_frontier:
       
  7949   fixes S :: "'n::euclidean_space set"
       
  7950   shows "affine hull S = UNIV \<Longrightarrow> rel_frontier S = frontier S"
       
  7951 by (simp add: frontier_def rel_frontier_def rel_interior_interior)
       
  7952 
       
  7953 lemma closed_rel_frontier [iff]:
       
  7954   fixes S :: "'n::euclidean_space set"
       
  7955   shows "closed (rel_frontier S)"
       
  7956 proof -
       
  7957   have *: "closedin (subtopology euclidean (affine hull S)) (closure S - rel_interior S)"
       
  7958     apply (rule closedin_diff[of "subtopology euclidean (affine hull S)""closure S" "rel_interior S"])
       
  7959     using closed_closedin_trans[of "affine hull S" "closure S"] closed_affine_hull[of S]
       
  7960       closure_affine_hull[of S] openin_rel_interior[of S]
       
  7961     apply auto
       
  7962     done
       
  7963   show ?thesis
       
  7964     apply (rule closedin_closed_trans[of "affine hull S" "rel_frontier S"])
       
  7965     unfolding rel_frontier_def
       
  7966     using * closed_affine_hull
       
  7967     apply auto
       
  7968     done
       
  7969 qed
       
  7970 
       
  7971 lemma closed_rel_boundary:
       
  7972   fixes S :: "'n::euclidean_space set"
       
  7973   shows "closed S \<Longrightarrow> closed(S - rel_interior S)"
       
  7974 by (metis closed_rel_frontier closure_closed rel_frontier_def)
       
  7975 
       
  7976 lemma compact_rel_boundary:
       
  7977   fixes S :: "'n::euclidean_space set"
       
  7978   shows "compact S \<Longrightarrow> compact(S - rel_interior S)"
       
  7979 by (metis bounded_diff closed_rel_boundary closure_eq compact_closure compact_imp_closed)
       
  7980 
       
  7981 lemma bounded_rel_frontier:
       
  7982   fixes S :: "'n::euclidean_space set"
       
  7983   shows "bounded S \<Longrightarrow> bounded(rel_frontier S)"
       
  7984 by (simp add: bounded_closure bounded_diff rel_frontier_def)
       
  7985 
       
  7986 lemma compact_rel_frontier_bounded:
       
  7987   fixes S :: "'n::euclidean_space set"
       
  7988   shows "bounded S \<Longrightarrow> compact(rel_frontier S)"
       
  7989 using bounded_rel_frontier closed_rel_frontier compact_eq_bounded_closed by blast
       
  7990 
       
  7991 lemma compact_rel_frontier:
       
  7992   fixes S :: "'n::euclidean_space set"
       
  7993   shows "compact S \<Longrightarrow> compact(rel_frontier S)"
       
  7994 by (meson compact_eq_bounded_closed compact_rel_frontier_bounded)
       
  7995 
       
  7996 lemma convex_same_rel_interior_closure:
       
  7997   fixes S :: "'n::euclidean_space set"
       
  7998   shows "\<lbrakk>convex S; convex T\<rbrakk>
       
  7999          \<Longrightarrow> rel_interior S = rel_interior T \<longleftrightarrow> closure S = closure T"
       
  8000 by (simp add: closure_eq_rel_interior_eq)
       
  8001 
       
  8002 lemma convex_same_rel_interior_closure_straddle:
       
  8003   fixes S :: "'n::euclidean_space set"
       
  8004   shows "\<lbrakk>convex S; convex T\<rbrakk>
       
  8005          \<Longrightarrow> rel_interior S = rel_interior T \<longleftrightarrow>
       
  8006              rel_interior S \<subseteq> T \<and> T \<subseteq> closure S"
       
  8007 by (simp add: closure_eq_between convex_same_rel_interior_closure)
       
  8008 
       
  8009 lemma convex_rel_frontier_aff_dim:
       
  8010   fixes S1 S2 :: "'n::euclidean_space set"
       
  8011   assumes "convex S1"
       
  8012     and "convex S2"
       
  8013     and "S2 \<noteq> {}"
       
  8014     and "S1 \<le> rel_frontier S2"
       
  8015   shows "aff_dim S1 < aff_dim S2"
       
  8016 proof -
       
  8017   have "S1 \<subseteq> closure S2"
       
  8018     using assms unfolding rel_frontier_def by auto
       
  8019   then have *: "affine hull S1 \<subseteq> affine hull S2"
       
  8020     using hull_mono[of "S1" "closure S2"] closure_same_affine_hull[of S2] by blast
       
  8021   then have "aff_dim S1 \<le> aff_dim S2"
       
  8022     using * aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2]
       
  8023       aff_dim_subset[of "affine hull S1" "affine hull S2"]
       
  8024     by auto
       
  8025   moreover
       
  8026   {
       
  8027     assume eq: "aff_dim S1 = aff_dim S2"
       
  8028     then have "S1 \<noteq> {}"
       
  8029       using aff_dim_empty[of S1] aff_dim_empty[of S2] \<open>S2 \<noteq> {}\<close> by auto
       
  8030     have **: "affine hull S1 = affine hull S2"
       
  8031        apply (rule affine_dim_equal)
       
  8032        using * affine_affine_hull
       
  8033        apply auto
       
  8034        using \<open>S1 \<noteq> {}\<close> hull_subset[of S1]
       
  8035        apply auto
       
  8036        using eq aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2]
       
  8037        apply auto
       
  8038        done
       
  8039     obtain a where a: "a \<in> rel_interior S1"
       
  8040       using \<open>S1 \<noteq> {}\<close> rel_interior_eq_empty assms by auto
       
  8041     obtain T where T: "open T" "a \<in> T \<inter> S1" "T \<inter> affine hull S1 \<subseteq> S1"
       
  8042        using mem_rel_interior[of a S1] a by auto
       
  8043     then have "a \<in> T \<inter> closure S2"
       
  8044       using a assms unfolding rel_frontier_def by auto
       
  8045     then obtain b where b: "b \<in> T \<inter> rel_interior S2"
       
  8046       using open_inter_closure_rel_interior[of S2 T] assms T by auto
       
  8047     then have "b \<in> affine hull S1"
       
  8048       using rel_interior_subset hull_subset[of S2] ** by auto
       
  8049     then have "b \<in> S1"
       
  8050       using T b by auto
       
  8051     then have False
       
  8052       using b assms unfolding rel_frontier_def by auto
       
  8053   }
       
  8054   ultimately show ?thesis
       
  8055     using less_le by auto
       
  8056 qed
       
  8057 
       
  8058 
       
  8059 lemma convex_rel_interior_if:
       
  8060   fixes S ::  "'n::euclidean_space set"
       
  8061   assumes "convex S"
       
  8062     and "z \<in> rel_interior S"
       
  8063   shows "\<forall>x\<in>affine hull S. \<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S)"
       
  8064 proof -
       
  8065   obtain e1 where e1: "e1 > 0 \<and> cball z e1 \<inter> affine hull S \<subseteq> S"
       
  8066     using mem_rel_interior_cball[of z S] assms by auto
       
  8067   {
       
  8068     fix x
       
  8069     assume x: "x \<in> affine hull S"
       
  8070     {
       
  8071       assume "x \<noteq> z"
       
  8072       define m where "m = 1 + e1/norm(x-z)"
       
  8073       hence "m > 1" using e1 \<open>x \<noteq> z\<close> by auto
       
  8074       {
       
  8075         fix e
       
  8076         assume e: "e > 1 \<and> e \<le> m"
       
  8077         have "z \<in> affine hull S"
       
  8078           using assms rel_interior_subset hull_subset[of S] by auto
       
  8079         then have *: "(1 - e)*\<^sub>R x + e *\<^sub>R z \<in> affine hull S"
       
  8080           using mem_affine[of "affine hull S" x z "(1-e)" e] affine_affine_hull[of S] x
       
  8081           by auto
       
  8082         have "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) = norm ((e - 1) *\<^sub>R (x - z))"
       
  8083           by (simp add: algebra_simps)
       
  8084         also have "\<dots> = (e - 1) * norm (x-z)"
       
  8085           using norm_scaleR e by auto
       
  8086         also have "\<dots> \<le> (m - 1) * norm (x - z)"
       
  8087           using e mult_right_mono[of _ _ "norm(x-z)"] by auto
       
  8088         also have "\<dots> = (e1 / norm (x - z)) * norm (x - z)"
       
  8089           using m_def by auto
       
  8090         also have "\<dots> = e1"
       
  8091           using \<open>x \<noteq> z\<close> e1 by simp
       
  8092         finally have **: "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) \<le> e1"
       
  8093           by auto
       
  8094         have "(1 - e)*\<^sub>R x+ e *\<^sub>R z \<in> cball z e1"
       
  8095           using m_def **
       
  8096           unfolding cball_def dist_norm
       
  8097           by (auto simp add: algebra_simps)
       
  8098         then have "(1 - e) *\<^sub>R x+ e *\<^sub>R z \<in> S"
       
  8099           using e * e1 by auto
       
  8100       }
       
  8101       then have "\<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S )"
       
  8102         using \<open>m> 1 \<close> by auto
       
  8103     }
       
  8104     moreover
       
  8105     {
       
  8106       assume "x = z"
       
  8107       define m where "m = 1 + e1"
       
  8108       then have "m > 1"
       
  8109         using e1 by auto
       
  8110       {
       
  8111         fix e
       
  8112         assume e: "e > 1 \<and> e \<le> m"
       
  8113         then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S"
       
  8114           using e1 x \<open>x = z\<close> by (auto simp add: algebra_simps)
       
  8115         then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S"
       
  8116           using e by auto
       
  8117       }
       
  8118       then have "\<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S)"
       
  8119         using \<open>m > 1\<close> by auto
       
  8120     }
       
  8121     ultimately have "\<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S )"
       
  8122       by blast
       
  8123   }
       
  8124   then show ?thesis by auto
       
  8125 qed
       
  8126 
       
  8127 lemma convex_rel_interior_if2:
       
  8128   fixes S :: "'n::euclidean_space set"
       
  8129   assumes "convex S"
       
  8130   assumes "z \<in> rel_interior S"
       
  8131   shows "\<forall>x\<in>affine hull S. \<exists>e. e > 1 \<and> (1 - e)*\<^sub>R x + e *\<^sub>R z \<in> S"
       
  8132   using convex_rel_interior_if[of S z] assms by auto
       
  8133 
       
  8134 lemma convex_rel_interior_only_if:
       
  8135   fixes S :: "'n::euclidean_space set"
       
  8136   assumes "convex S"
       
  8137     and "S \<noteq> {}"
       
  8138   assumes "\<forall>x\<in>S. \<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S"
       
  8139   shows "z \<in> rel_interior S"
       
  8140 proof -
       
  8141   obtain x where x: "x \<in> rel_interior S"
       
  8142     using rel_interior_eq_empty assms by auto
       
  8143   then have "x \<in> S"
       
  8144     using rel_interior_subset by auto
       
  8145   then obtain e where e: "e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S"
       
  8146     using assms by auto
       
  8147   define y where [abs_def]: "y = (1 - e) *\<^sub>R x + e *\<^sub>R z"
       
  8148   then have "y \<in> S" using e by auto
       
  8149   define e1 where "e1 = 1/e"
       
  8150   then have "0 < e1 \<and> e1 < 1" using e by auto
       
  8151   then have "z  =y - (1 - e1) *\<^sub>R (y - x)"
       
  8152     using e1_def y_def by (auto simp add: algebra_simps)
       
  8153   then show ?thesis
       
  8154     using rel_interior_convex_shrink[of S x y "1-e1"] \<open>0 < e1 \<and> e1 < 1\<close> \<open>y \<in> S\<close> x assms
       
  8155     by auto
       
  8156 qed
       
  8157 
       
  8158 lemma convex_rel_interior_iff:
       
  8159   fixes S :: "'n::euclidean_space set"
       
  8160   assumes "convex S"
       
  8161     and "S \<noteq> {}"
       
  8162   shows "z \<in> rel_interior S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S)"
       
  8163   using assms hull_subset[of S "affine"]
       
  8164     convex_rel_interior_if[of S z] convex_rel_interior_only_if[of S z]
       
  8165   by auto
       
  8166 
       
  8167 lemma convex_rel_interior_iff2:
       
  8168   fixes S :: "'n::euclidean_space set"
       
  8169   assumes "convex S"
       
  8170     and "S \<noteq> {}"
       
  8171   shows "z \<in> rel_interior S \<longleftrightarrow> (\<forall>x\<in>affine hull S. \<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S)"
       
  8172   using assms hull_subset[of S] convex_rel_interior_if2[of S z] convex_rel_interior_only_if[of S z]
       
  8173   by auto
       
  8174 
       
  8175 lemma convex_interior_iff:
       
  8176   fixes S :: "'n::euclidean_space set"
       
  8177   assumes "convex S"
       
  8178   shows "z \<in> interior S \<longleftrightarrow> (\<forall>x. \<exists>e. e > 0 \<and> z + e *\<^sub>R x \<in> S)"
       
  8179 proof (cases "aff_dim S = int DIM('n)")
       
  8180   case False
       
  8181   {
       
  8182     assume "z \<in> interior S"
       
  8183     then have False
       
  8184       using False interior_rel_interior_gen[of S] by auto
       
  8185   }
       
  8186   moreover
       
  8187   {
       
  8188     assume r: "\<forall>x. \<exists>e. e > 0 \<and> z + e *\<^sub>R x \<in> S"
       
  8189     {
       
  8190       fix x
       
  8191       obtain e1 where e1: "e1 > 0 \<and> z + e1 *\<^sub>R (x - z) \<in> S"
       
  8192         using r by auto
       
  8193       obtain e2 where e2: "e2 > 0 \<and> z + e2 *\<^sub>R (z - x) \<in> S"
       
  8194         using r by auto
       
  8195       define x1 where [abs_def]: "x1 = z + e1 *\<^sub>R (x - z)"
       
  8196       then have x1: "x1 \<in> affine hull S"
       
  8197         using e1 hull_subset[of S] by auto
       
  8198       define x2 where [abs_def]: "x2 = z + e2 *\<^sub>R (z - x)"
       
  8199       then have x2: "x2 \<in> affine hull S"
       
  8200         using e2 hull_subset[of S] by auto
       
  8201       have *: "e1/(e1+e2) + e2/(e1+e2) = 1"
       
  8202         using add_divide_distrib[of e1 e2 "e1+e2"] e1 e2 by simp
       
  8203       then have "z = (e2/(e1+e2)) *\<^sub>R x1 + (e1/(e1+e2)) *\<^sub>R x2"
       
  8204         using x1_def x2_def
       
  8205         apply (auto simp add: algebra_simps)
       
  8206         using scaleR_left_distrib[of "e1/(e1+e2)" "e2/(e1+e2)" z]
       
  8207         apply auto
       
  8208         done
       
  8209       then have z: "z \<in> affine hull S"
       
  8210         using mem_affine[of "affine hull S" x1 x2 "e2/(e1+e2)" "e1/(e1+e2)"]
       
  8211           x1 x2 affine_affine_hull[of S] *
       
  8212         by auto
       
  8213       have "x1 - x2 = (e1 + e2) *\<^sub>R (x - z)"
       
  8214         using x1_def x2_def by (auto simp add: algebra_simps)
       
  8215       then have "x = z+(1/(e1+e2)) *\<^sub>R (x1-x2)"
       
  8216         using e1 e2 by simp
       
  8217       then have "x \<in> affine hull S"
       
  8218         using mem_affine_3_minus[of "affine hull S" z x1 x2 "1/(e1+e2)"]
       
  8219           x1 x2 z affine_affine_hull[of S]
       
  8220         by auto
       
  8221     }
       
  8222     then have "affine hull S = UNIV"
       
  8223       by auto
       
  8224     then have "aff_dim S = int DIM('n)"
       
  8225       using aff_dim_affine_hull[of S] by (simp add: aff_dim_UNIV)
       
  8226     then have False
       
  8227       using False by auto
       
  8228   }
       
  8229   ultimately show ?thesis by auto
       
  8230 next
       
  8231   case True
       
  8232   then have "S \<noteq> {}"
       
  8233     using aff_dim_empty[of S] by auto
       
  8234   have *: "affine hull S = UNIV"
       
  8235     using True affine_hull_UNIV by auto
       
  8236   {
       
  8237     assume "z \<in> interior S"
       
  8238     then have "z \<in> rel_interior S"
       
  8239       using True interior_rel_interior_gen[of S] by auto
       
  8240     then have **: "\<forall>x. \<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S"
       
  8241       using convex_rel_interior_iff2[of S z] assms \<open>S \<noteq> {}\<close> * by auto
       
  8242     fix x
       
  8243     obtain e1 where e1: "e1 > 1" "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z \<in> S"
       
  8244       using **[rule_format, of "z-x"] by auto
       
  8245     define e where [abs_def]: "e = e1 - 1"
       
  8246     then have "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z = z + e *\<^sub>R x"
       
  8247       by (simp add: algebra_simps)
       
  8248     then have "e > 0" "z + e *\<^sub>R x \<in> S"
       
  8249       using e1 e_def by auto
       
  8250     then have "\<exists>e. e > 0 \<and> z + e *\<^sub>R x \<in> S"
       
  8251       by auto
       
  8252   }
       
  8253   moreover
       
  8254   {
       
  8255     assume r: "\<forall>x. \<exists>e. e > 0 \<and> z + e *\<^sub>R x \<in> S"
       
  8256     {
       
  8257       fix x
       
  8258       obtain e1 where e1: "e1 > 0" "z + e1 *\<^sub>R (z - x) \<in> S"
       
  8259         using r[rule_format, of "z-x"] by auto
       
  8260       define e where "e = e1 + 1"
       
  8261       then have "z + e1 *\<^sub>R (z - x) = (1 - e) *\<^sub>R x + e *\<^sub>R z"
       
  8262         by (simp add: algebra_simps)
       
  8263       then have "e > 1" "(1 - e)*\<^sub>R x + e *\<^sub>R z \<in> S"
       
  8264         using e1 e_def by auto
       
  8265       then have "\<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S" by auto
       
  8266     }
       
  8267     then have "z \<in> rel_interior S"
       
  8268       using convex_rel_interior_iff2[of S z] assms \<open>S \<noteq> {}\<close> by auto
       
  8269     then have "z \<in> interior S"
       
  8270       using True interior_rel_interior_gen[of S] by auto
       
  8271   }
       
  8272   ultimately show ?thesis by auto
       
  8273 qed
       
  8274 
       
  8275 
       
  8276 subsubsection \<open>Relative interior and closure under common operations\<close>
       
  8277 
       
  8278 lemma rel_interior_inter_aux: "\<Inter>{rel_interior S |S. S : I} \<subseteq> \<Inter>I"
       
  8279 proof -
       
  8280   {
       
  8281     fix y
       
  8282     assume "y \<in> \<Inter>{rel_interior S |S. S : I}"
       
  8283     then have y: "\<forall>S \<in> I. y \<in> rel_interior S"
       
  8284       by auto
       
  8285     {
       
  8286       fix S
       
  8287       assume "S \<in> I"
       
  8288       then have "y \<in> S"
       
  8289         using rel_interior_subset y by auto
       
  8290     }
       
  8291     then have "y \<in> \<Inter>I" by auto
       
  8292   }
       
  8293   then show ?thesis by auto
       
  8294 qed
       
  8295 
       
  8296 lemma closure_Int: "closure (\<Inter>I) \<le> \<Inter>{closure S |S. S \<in> I}"
       
  8297 proof -
       
  8298   {
       
  8299     fix y
       
  8300     assume "y \<in> \<Inter>I"
       
  8301     then have y: "\<forall>S \<in> I. y \<in> S" by auto
       
  8302     {
       
  8303       fix S
       
  8304       assume "S \<in> I"
       
  8305       then have "y \<in> closure S"
       
  8306         using closure_subset y by auto
       
  8307     }
       
  8308     then have "y \<in> \<Inter>{closure S |S. S \<in> I}"
       
  8309       by auto
       
  8310   }
       
  8311   then have "\<Inter>I \<subseteq> \<Inter>{closure S |S. S \<in> I}"
       
  8312     by auto
       
  8313   moreover have "closed (\<Inter>{closure S |S. S \<in> I})"
       
  8314     unfolding closed_Inter closed_closure by auto
       
  8315   ultimately show ?thesis using closure_hull[of "\<Inter>I"]
       
  8316     hull_minimal[of "\<Inter>I" "\<Inter>{closure S |S. S \<in> I}" "closed"] by auto
       
  8317 qed
       
  8318 
       
  8319 lemma convex_closure_rel_interior_inter:
       
  8320   assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set)"
       
  8321     and "\<Inter>{rel_interior S |S. S \<in> I} \<noteq> {}"
       
  8322   shows "\<Inter>{closure S |S. S \<in> I} \<le> closure (\<Inter>{rel_interior S |S. S \<in> I})"
       
  8323 proof -
       
  8324   obtain x where x: "\<forall>S\<in>I. x \<in> rel_interior S"
       
  8325     using assms by auto
       
  8326   {
       
  8327     fix y
       
  8328     assume "y \<in> \<Inter>{closure S |S. S \<in> I}"
       
  8329     then have y: "\<forall>S \<in> I. y \<in> closure S"
       
  8330       by auto
       
  8331     {
       
  8332       assume "y = x"
       
  8333       then have "y \<in> closure (\<Inter>{rel_interior S |S. S \<in> I})"
       
  8334         using x closure_subset[of "\<Inter>{rel_interior S |S. S \<in> I}"] by auto
       
  8335     }
       
  8336     moreover
       
  8337     {
       
  8338       assume "y \<noteq> x"
       
  8339       { fix e :: real
       
  8340         assume e: "e > 0"
       
  8341         define e1 where "e1 = min 1 (e/norm (y - x))"
       
  8342         then have e1: "e1 > 0" "e1 \<le> 1" "e1 * norm (y - x) \<le> e"
       
  8343           using \<open>y \<noteq> x\<close> \<open>e > 0\<close> le_divide_eq[of e1 e "norm (y - x)"]
       
  8344           by simp_all
       
  8345         define z where "z = y - e1 *\<^sub>R (y - x)"
       
  8346         {
       
  8347           fix S
       
  8348           assume "S \<in> I"
       
  8349           then have "z \<in> rel_interior S"
       
  8350             using rel_interior_closure_convex_shrink[of S x y e1] assms x y e1 z_def
       
  8351             by auto
       
  8352         }
       
  8353         then have *: "z \<in> \<Inter>{rel_interior S |S. S \<in> I}"
       
  8354           by auto
       
  8355         have "\<exists>z. z \<in> \<Inter>{rel_interior S |S. S \<in> I} \<and> z \<noteq> y \<and> dist z y \<le> e"
       
  8356           apply (rule_tac x="z" in exI)
       
  8357           using \<open>y \<noteq> x\<close> z_def * e1 e dist_norm[of z y]
       
  8358           apply simp
       
  8359           done
       
  8360       }
       
  8361       then have "y islimpt \<Inter>{rel_interior S |S. S \<in> I}"
       
  8362         unfolding islimpt_approachable_le by blast
       
  8363       then have "y \<in> closure (\<Inter>{rel_interior S |S. S \<in> I})"
       
  8364         unfolding closure_def by auto
       
  8365     }
       
  8366     ultimately have "y \<in> closure (\<Inter>{rel_interior S |S. S \<in> I})"
       
  8367       by auto
       
  8368   }
       
  8369   then show ?thesis by auto
       
  8370 qed
       
  8371 
       
  8372 lemma convex_closure_inter:
       
  8373   assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set)"
       
  8374     and "\<Inter>{rel_interior S |S. S \<in> I} \<noteq> {}"
       
  8375   shows "closure (\<Inter>I) = \<Inter>{closure S |S. S \<in> I}"
       
  8376 proof -
       
  8377   have "\<Inter>{closure S |S. S \<in> I} \<le> closure (\<Inter>{rel_interior S |S. S \<in> I})"
       
  8378     using convex_closure_rel_interior_inter assms by auto
       
  8379   moreover
       
  8380   have "closure (\<Inter>{rel_interior S |S. S \<in> I}) \<le> closure (\<Inter>I)"
       
  8381     using rel_interior_inter_aux closure_mono[of "\<Inter>{rel_interior S |S. S \<in> I}" "\<Inter>I"]
       
  8382     by auto
       
  8383   ultimately show ?thesis
       
  8384     using closure_Int[of I] by auto
       
  8385 qed
       
  8386 
       
  8387 lemma convex_inter_rel_interior_same_closure:
       
  8388   assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set)"
       
  8389     and "\<Inter>{rel_interior S |S. S \<in> I} \<noteq> {}"
       
  8390   shows "closure (\<Inter>{rel_interior S |S. S \<in> I}) = closure (\<Inter>I)"
       
  8391 proof -
       
  8392   have "\<Inter>{closure S |S. S \<in> I} \<le> closure (\<Inter>{rel_interior S |S. S \<in> I})"
       
  8393     using convex_closure_rel_interior_inter assms by auto
       
  8394   moreover
       
  8395   have "closure (\<Inter>{rel_interior S |S. S \<in> I}) \<le> closure (\<Inter>I)"
       
  8396     using rel_interior_inter_aux closure_mono[of "\<Inter>{rel_interior S |S. S \<in> I}" "\<Inter>I"]
       
  8397     by auto
       
  8398   ultimately show ?thesis
       
  8399     using closure_Int[of I] by auto
       
  8400 qed
       
  8401 
       
  8402 lemma convex_rel_interior_inter:
       
  8403   assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set)"
       
  8404     and "\<Inter>{rel_interior S |S. S \<in> I} \<noteq> {}"
       
  8405   shows "rel_interior (\<Inter>I) \<subseteq> \<Inter>{rel_interior S |S. S \<in> I}"
       
  8406 proof -
       
  8407   have "convex (\<Inter>I)"
       
  8408     using assms convex_Inter by auto
       
  8409   moreover
       
  8410   have "convex (\<Inter>{rel_interior S |S. S \<in> I})"
       
  8411     apply (rule convex_Inter)
       
  8412     using assms convex_rel_interior
       
  8413     apply auto
       
  8414     done
       
  8415   ultimately
       
  8416   have "rel_interior (\<Inter>{rel_interior S |S. S \<in> I}) = rel_interior (\<Inter>I)"
       
  8417     using convex_inter_rel_interior_same_closure assms
       
  8418       closure_eq_rel_interior_eq[of "\<Inter>{rel_interior S |S. S \<in> I}" "\<Inter>I"]
       
  8419     by blast
       
  8420   then show ?thesis
       
  8421     using rel_interior_subset[of "\<Inter>{rel_interior S |S. S \<in> I}"] by auto
       
  8422 qed
       
  8423 
       
  8424 lemma convex_rel_interior_finite_inter:
       
  8425   assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set)"
       
  8426     and "\<Inter>{rel_interior S |S. S \<in> I} \<noteq> {}"
       
  8427     and "finite I"
       
  8428   shows "rel_interior (\<Inter>I) = \<Inter>{rel_interior S |S. S \<in> I}"
       
  8429 proof -
       
  8430   have "\<Inter>I \<noteq> {}"
       
  8431     using assms rel_interior_inter_aux[of I] by auto
       
  8432   have "convex (\<Inter>I)"
       
  8433     using convex_Inter assms by auto
       
  8434   show ?thesis
       
  8435   proof (cases "I = {}")
       
  8436     case True
       
  8437     then show ?thesis
       
  8438       using Inter_empty rel_interior_UNIV by auto
       
  8439   next
       
  8440     case False
       
  8441     {
       
  8442       fix z
       
  8443       assume z: "z \<in> \<Inter>{rel_interior S |S. S \<in> I}"
       
  8444       {
       
  8445         fix x
       
  8446         assume x: "x \<in> \<Inter>I"
       
  8447         {
       
  8448           fix S
       
  8449           assume S: "S \<in> I"
       
  8450           then have "z \<in> rel_interior S" "x \<in> S"
       
  8451             using z x by auto
       
  8452           then have "\<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e)*\<^sub>R x + e *\<^sub>R z \<in> S)"
       
  8453             using convex_rel_interior_if[of S z] S assms hull_subset[of S] by auto
       
  8454         }
       
  8455         then obtain mS where
       
  8456           mS: "\<forall>S\<in>I. mS S > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> mS S \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S)" by metis
       
  8457         define e where "e = Min (mS ` I)"
       
  8458         then have "e \<in> mS ` I" using assms \<open>I \<noteq> {}\<close> by simp
       
  8459         then have "e > 1" using mS by auto
       
  8460         moreover have "\<forall>S\<in>I. e \<le> mS S"
       
  8461           using e_def assms by auto
       
  8462         ultimately have "\<exists>e > 1. (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> \<Inter>I"
       
  8463           using mS by auto
       
  8464       }
       
  8465       then have "z \<in> rel_interior (\<Inter>I)"
       
  8466         using convex_rel_interior_iff[of "\<Inter>I" z] \<open>\<Inter>I \<noteq> {}\<close> \<open>convex (\<Inter>I)\<close> by auto
       
  8467     }
       
  8468     then show ?thesis
       
  8469       using convex_rel_interior_inter[of I] assms by auto
       
  8470   qed
       
  8471 qed
       
  8472 
       
  8473 lemma convex_closure_inter_two:
       
  8474   fixes S T :: "'n::euclidean_space set"
       
  8475   assumes "convex S"
       
  8476     and "convex T"
       
  8477   assumes "rel_interior S \<inter> rel_interior T \<noteq> {}"
       
  8478   shows "closure (S \<inter> T) = closure S \<inter> closure T"
       
  8479   using convex_closure_inter[of "{S,T}"] assms by auto
       
  8480 
       
  8481 lemma convex_rel_interior_inter_two:
       
  8482   fixes S T :: "'n::euclidean_space set"
       
  8483   assumes "convex S"
       
  8484     and "convex T"
       
  8485     and "rel_interior S \<inter> rel_interior T \<noteq> {}"
       
  8486   shows "rel_interior (S \<inter> T) = rel_interior S \<inter> rel_interior T"
       
  8487   using convex_rel_interior_finite_inter[of "{S,T}"] assms by auto
       
  8488 
       
  8489 lemma convex_affine_closure_Int:
       
  8490   fixes S T :: "'n::euclidean_space set"
       
  8491   assumes "convex S"
       
  8492     and "affine T"
       
  8493     and "rel_interior S \<inter> T \<noteq> {}"
       
  8494   shows "closure (S \<inter> T) = closure S \<inter> T"
       
  8495 proof -
       
  8496   have "affine hull T = T"
       
  8497     using assms by auto
       
  8498   then have "rel_interior T = T"
       
  8499     using rel_interior_affine_hull[of T] by metis
       
  8500   moreover have "closure T = T"
       
  8501     using assms affine_closed[of T] by auto
       
  8502   ultimately show ?thesis
       
  8503     using convex_closure_inter_two[of S T] assms affine_imp_convex by auto
       
  8504 qed
       
  8505 
       
  8506 lemma connected_component_1_gen:
       
  8507   fixes S :: "'a :: euclidean_space set"
       
  8508   assumes "DIM('a) = 1"
       
  8509   shows "connected_component S a b \<longleftrightarrow> closed_segment a b \<subseteq> S"
       
  8510 unfolding connected_component_def
       
  8511 by (metis (no_types, lifting) assms subsetD subsetI convex_contains_segment convex_segment(1)
       
  8512             ends_in_segment connected_convex_1_gen)
       
  8513 
       
  8514 lemma connected_component_1:
       
  8515   fixes S :: "real set"
       
  8516   shows "connected_component S a b \<longleftrightarrow> closed_segment a b \<subseteq> S"
       
  8517 by (simp add: connected_component_1_gen)
       
  8518 
       
  8519 lemma convex_affine_rel_interior_Int:
       
  8520   fixes S T :: "'n::euclidean_space set"
       
  8521   assumes "convex S"
       
  8522     and "affine T"
       
  8523     and "rel_interior S \<inter> T \<noteq> {}"
       
  8524   shows "rel_interior (S \<inter> T) = rel_interior S \<inter> T"
       
  8525 proof -
       
  8526   have "affine hull T = T"
       
  8527     using assms by auto
       
  8528   then have "rel_interior T = T"
       
  8529     using rel_interior_affine_hull[of T] by metis
       
  8530   moreover have "closure T = T"
       
  8531     using assms affine_closed[of T] by auto
       
  8532   ultimately show ?thesis
       
  8533     using convex_rel_interior_inter_two[of S T] assms affine_imp_convex by auto
       
  8534 qed
       
  8535 
       
  8536 lemma convex_affine_rel_frontier_Int:
       
  8537    fixes S T :: "'n::euclidean_space set"
       
  8538   assumes "convex S"
       
  8539     and "affine T"
       
  8540     and "interior S \<inter> T \<noteq> {}"
       
  8541   shows "rel_frontier(S \<inter> T) = frontier S \<inter> T"
       
  8542 using assms
       
  8543 apply (simp add: rel_frontier_def convex_affine_closure_Int frontier_def)
       
  8544 by (metis Diff_Int_distrib2 Int_emptyI convex_affine_closure_Int convex_affine_rel_interior_Int empty_iff interior_rel_interior_gen)
       
  8545 
       
  8546 lemma subset_rel_interior_convex:
       
  8547   fixes S T :: "'n::euclidean_space set"
       
  8548   assumes "convex S"
       
  8549     and "convex T"
       
  8550     and "S \<le> closure T"
       
  8551     and "\<not> S \<subseteq> rel_frontier T"
       
  8552   shows "rel_interior S \<subseteq> rel_interior T"
       
  8553 proof -
       
  8554   have *: "S \<inter> closure T = S"
       
  8555     using assms by auto
       
  8556   have "\<not> rel_interior S \<subseteq> rel_frontier T"
       
  8557     using closure_mono[of "rel_interior S" "rel_frontier T"] closed_rel_frontier[of T]
       
  8558       closure_closed[of S] convex_closure_rel_interior[of S] closure_subset[of S] assms
       
  8559     by auto
       
  8560   then have "rel_interior S \<inter> rel_interior (closure T) \<noteq> {}"
       
  8561     using assms rel_frontier_def[of T] rel_interior_subset convex_rel_interior_closure[of T]
       
  8562     by auto
       
  8563   then have "rel_interior S \<inter> rel_interior T = rel_interior (S \<inter> closure T)"
       
  8564     using assms convex_closure convex_rel_interior_inter_two[of S "closure T"]
       
  8565       convex_rel_interior_closure[of T]
       
  8566     by auto
       
  8567   also have "\<dots> = rel_interior S"
       
  8568     using * by auto
       
  8569   finally show ?thesis
       
  8570     by auto
       
  8571 qed
       
  8572 
       
  8573 lemma rel_interior_convex_linear_image:
       
  8574   fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
       
  8575   assumes "linear f"
       
  8576     and "convex S"
       
  8577   shows "f ` (rel_interior S) = rel_interior (f ` S)"
       
  8578 proof (cases "S = {}")
       
  8579   case True
       
  8580   then show ?thesis
       
  8581     using assms rel_interior_empty rel_interior_eq_empty by auto
       
  8582 next
       
  8583   case False
       
  8584   have *: "f ` (rel_interior S) \<subseteq> f ` S"
       
  8585     unfolding image_mono using rel_interior_subset by auto
       
  8586   have "f ` S \<subseteq> f ` (closure S)"
       
  8587     unfolding image_mono using closure_subset by auto
       
  8588   also have "\<dots> = f ` (closure (rel_interior S))"
       
  8589     using convex_closure_rel_interior assms by auto
       
  8590   also have "\<dots> \<subseteq> closure (f ` (rel_interior S))"
       
  8591     using closure_linear_image_subset assms by auto
       
  8592   finally have "closure (f ` S) = closure (f ` rel_interior S)"
       
  8593     using closure_mono[of "f ` S" "closure (f ` rel_interior S)"] closure_closure
       
  8594       closure_mono[of "f ` rel_interior S" "f ` S"] *
       
  8595     by auto
       
  8596   then have "rel_interior (f ` S) = rel_interior (f ` rel_interior S)"
       
  8597     using assms convex_rel_interior
       
  8598       linear_conv_bounded_linear[of f] convex_linear_image[of _ S]
       
  8599       convex_linear_image[of _ "rel_interior S"]
       
  8600       closure_eq_rel_interior_eq[of "f ` S" "f ` rel_interior S"]
       
  8601     by auto
       
  8602   then have "rel_interior (f ` S) \<subseteq> f ` rel_interior S"
       
  8603     using rel_interior_subset by auto
       
  8604   moreover
       
  8605   {
       
  8606     fix z
       
  8607     assume "z \<in> f ` rel_interior S"
       
  8608     then obtain z1 where z1: "z1 \<in> rel_interior S" "f z1 = z" by auto
       
  8609     {
       
  8610       fix x
       
  8611       assume "x \<in> f ` S"
       
  8612       then obtain x1 where x1: "x1 \<in> S" "f x1 = x" by auto
       
  8613       then obtain e where e: "e > 1" "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1 : S"
       
  8614         using convex_rel_interior_iff[of S z1] \<open>convex S\<close> x1 z1 by auto
       
  8615       moreover have "f ((1 - e) *\<^sub>R x1 + e *\<^sub>R z1) = (1 - e) *\<^sub>R x + e *\<^sub>R z"
       
  8616         using x1 z1 \<open>linear f\<close> by (simp add: linear_add_cmul)
       
  8617       ultimately have "(1 - e) *\<^sub>R x + e *\<^sub>R z : f ` S"
       
  8618         using imageI[of "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1" S f] by auto
       
  8619       then have "\<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z : f ` S"
       
  8620         using e by auto
       
  8621     }
       
  8622     then have "z \<in> rel_interior (f ` S)"
       
  8623       using convex_rel_interior_iff[of "f ` S" z] \<open>convex S\<close>
       
  8624         \<open>linear f\<close> \<open>S \<noteq> {}\<close> convex_linear_image[of f S]  linear_conv_bounded_linear[of f]
       
  8625       by auto
       
  8626   }
       
  8627   ultimately show ?thesis by auto
       
  8628 qed
       
  8629 
       
  8630 lemma rel_interior_convex_linear_preimage:
       
  8631   fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
       
  8632   assumes "linear f"
       
  8633     and "convex S"
       
  8634     and "f -` (rel_interior S) \<noteq> {}"
       
  8635   shows "rel_interior (f -` S) = f -` (rel_interior S)"
       
  8636 proof -
       
  8637   have "S \<noteq> {}"
       
  8638     using assms rel_interior_empty by auto
       
  8639   have nonemp: "f -` S \<noteq> {}"
       
  8640     by (metis assms(3) rel_interior_subset subset_empty vimage_mono)
       
  8641   then have "S \<inter> (range f) \<noteq> {}"
       
  8642     by auto
       
  8643   have conv: "convex (f -` S)"
       
  8644     using convex_linear_vimage assms by auto
       
  8645   then have "convex (S \<inter> range f)"
       
  8646     by (metis assms(1) assms(2) convex_Int subspace_UNIV subspace_imp_convex subspace_linear_image)
       
  8647   {
       
  8648     fix z
       
  8649     assume "z \<in> f -` (rel_interior S)"
       
  8650     then have z: "f z : rel_interior S"
       
  8651       by auto
       
  8652     {
       
  8653       fix x
       
  8654       assume "x \<in> f -` S"
       
  8655       then have "f x \<in> S" by auto
       
  8656       then obtain e where e: "e > 1" "(1 - e) *\<^sub>R f x + e *\<^sub>R f z \<in> S"
       
  8657         using convex_rel_interior_iff[of S "f z"] z assms \<open>S \<noteq> {}\<close> by auto
       
  8658       moreover have "(1 - e) *\<^sub>R f x + e *\<^sub>R f z = f ((1 - e) *\<^sub>R x + e *\<^sub>R z)"
       
  8659         using \<open>linear f\<close> by (simp add: linear_iff)
       
  8660       ultimately have "\<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> f -` S"
       
  8661         using e by auto
       
  8662     }
       
  8663     then have "z \<in> rel_interior (f -` S)"
       
  8664       using convex_rel_interior_iff[of "f -` S" z] conv nonemp by auto
       
  8665   }
       
  8666   moreover
       
  8667   {
       
  8668     fix z
       
  8669     assume z: "z \<in> rel_interior (f -` S)"
       
  8670     {
       
  8671       fix x
       
  8672       assume "x \<in> S \<inter> range f"
       
  8673       then obtain y where y: "f y = x" "y \<in> f -` S" by auto
       
  8674       then obtain e where e: "e > 1" "(1 - e) *\<^sub>R y + e *\<^sub>R z \<in> f -` S"
       
  8675         using convex_rel_interior_iff[of "f -` S" z] z conv by auto
       
  8676       moreover have "(1 - e) *\<^sub>R x + e *\<^sub>R f z = f ((1 - e) *\<^sub>R y + e *\<^sub>R z)"
       
  8677         using \<open>linear f\<close> y by (simp add: linear_iff)
       
  8678       ultimately have "\<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R f z \<in> S \<inter> range f"
       
  8679         using e by auto
       
  8680     }
       
  8681     then have "f z \<in> rel_interior (S \<inter> range f)"
       
  8682       using \<open>convex (S \<inter> (range f))\<close> \<open>S \<inter> range f \<noteq> {}\<close>
       
  8683         convex_rel_interior_iff[of "S \<inter> (range f)" "f z"]
       
  8684       by auto
       
  8685     moreover have "affine (range f)"
       
  8686       by (metis assms(1) subspace_UNIV subspace_imp_affine subspace_linear_image)
       
  8687     ultimately have "f z \<in> rel_interior S"
       
  8688       using convex_affine_rel_interior_Int[of S "range f"] assms by auto
       
  8689     then have "z \<in> f -` (rel_interior S)"
       
  8690       by auto
       
  8691   }
       
  8692   ultimately show ?thesis by auto
       
  8693 qed
       
  8694 
       
  8695 lemma rel_interior_Times:
       
  8696   fixes S :: "'n::euclidean_space set"
       
  8697     and T :: "'m::euclidean_space set"
       
  8698   assumes "convex S"
       
  8699     and "convex T"
       
  8700   shows "rel_interior (S \<times> T) = rel_interior S \<times> rel_interior T"
       
  8701 proof -
       
  8702   { assume "S = {}"
       
  8703     then have ?thesis
       
  8704       by auto
       
  8705   }
       
  8706   moreover
       
  8707   { assume "T = {}"
       
  8708     then have ?thesis
       
  8709        by auto
       
  8710   }
       
  8711   moreover
       
  8712   {
       
  8713     assume "S \<noteq> {}" "T \<noteq> {}"
       
  8714     then have ri: "rel_interior S \<noteq> {}" "rel_interior T \<noteq> {}"
       
  8715       using rel_interior_eq_empty assms by auto
       
  8716     then have "fst -` rel_interior S \<noteq> {}"
       
  8717       using fst_vimage_eq_Times[of "rel_interior S"] by auto
       
  8718     then have "rel_interior ((fst :: 'n * 'm \<Rightarrow> 'n) -` S) = fst -` rel_interior S"
       
  8719       using fst_linear \<open>convex S\<close> rel_interior_convex_linear_preimage[of fst S] by auto
       
  8720     then have s: "rel_interior (S \<times> (UNIV :: 'm set)) = rel_interior S \<times> UNIV"
       
  8721       by (simp add: fst_vimage_eq_Times)
       
  8722     from ri have "snd -` rel_interior T \<noteq> {}"
       
  8723       using snd_vimage_eq_Times[of "rel_interior T"] by auto
       
  8724     then have "rel_interior ((snd :: 'n * 'm \<Rightarrow> 'm) -` T) = snd -` rel_interior T"
       
  8725       using snd_linear \<open>convex T\<close> rel_interior_convex_linear_preimage[of snd T] by auto
       
  8726     then have t: "rel_interior ((UNIV :: 'n set) \<times> T) = UNIV \<times> rel_interior T"
       
  8727       by (simp add: snd_vimage_eq_Times)
       
  8728     from s t have *: "rel_interior (S \<times> (UNIV :: 'm set)) \<inter> rel_interior ((UNIV :: 'n set) \<times> T) =
       
  8729       rel_interior S \<times> rel_interior T" by auto
       
  8730     have "S \<times> T = S \<times> (UNIV :: 'm set) \<inter> (UNIV :: 'n set) \<times> T"
       
  8731       by auto
       
  8732     then have "rel_interior (S \<times> T) = rel_interior ((S \<times> (UNIV :: 'm set)) \<inter> ((UNIV :: 'n set) \<times> T))"
       
  8733       by auto
       
  8734     also have "\<dots> = rel_interior (S \<times> (UNIV :: 'm set)) \<inter> rel_interior ((UNIV :: 'n set) \<times> T)"
       
  8735        apply (subst convex_rel_interior_inter_two[of "S \<times> (UNIV :: 'm set)" "(UNIV :: 'n set) \<times> T"])
       
  8736        using * ri assms convex_Times
       
  8737        apply auto
       
  8738        done
       
  8739     finally have ?thesis using * by auto
       
  8740   }
       
  8741   ultimately show ?thesis by blast
       
  8742 qed
       
  8743 
       
  8744 lemma rel_interior_scaleR:
       
  8745   fixes S :: "'n::euclidean_space set"
       
  8746   assumes "c \<noteq> 0"
       
  8747   shows "(op *\<^sub>R c) ` (rel_interior S) = rel_interior ((op *\<^sub>R c) ` S)"
       
  8748   using rel_interior_injective_linear_image[of "(op *\<^sub>R c)" S]
       
  8749     linear_conv_bounded_linear[of "op *\<^sub>R c"] linear_scaleR injective_scaleR[of c] assms
       
  8750   by auto
       
  8751 
       
  8752 lemma rel_interior_convex_scaleR:
       
  8753   fixes S :: "'n::euclidean_space set"
       
  8754   assumes "convex S"
       
  8755   shows "(op *\<^sub>R c) ` (rel_interior S) = rel_interior ((op *\<^sub>R c) ` S)"
       
  8756   by (metis assms linear_scaleR rel_interior_convex_linear_image)
       
  8757 
       
  8758 lemma convex_rel_open_scaleR:
       
  8759   fixes S :: "'n::euclidean_space set"
       
  8760   assumes "convex S"
       
  8761     and "rel_open S"
       
  8762   shows "convex ((op *\<^sub>R c) ` S) \<and> rel_open ((op *\<^sub>R c) ` S)"
       
  8763   by (metis assms convex_scaling rel_interior_convex_scaleR rel_open_def)
       
  8764 
       
  8765 lemma convex_rel_open_finite_inter:
       
  8766   assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set) \<and> rel_open S"
       
  8767     and "finite I"
       
  8768   shows "convex (\<Inter>I) \<and> rel_open (\<Inter>I)"
       
  8769 proof (cases "\<Inter>{rel_interior S |S. S \<in> I} = {}")
       
  8770   case True
       
  8771   then have "\<Inter>I = {}"
       
  8772     using assms unfolding rel_open_def by auto
       
  8773   then show ?thesis
       
  8774     unfolding rel_open_def using rel_interior_empty by auto
       
  8775 next
       
  8776   case False
       
  8777   then have "rel_open (\<Inter>I)"
       
  8778     using assms unfolding rel_open_def
       
  8779     using convex_rel_interior_finite_inter[of I]
       
  8780     by auto
       
  8781   then show ?thesis
       
  8782     using convex_Inter assms by auto
       
  8783 qed
       
  8784 
       
  8785 lemma convex_rel_open_linear_image:
       
  8786   fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
       
  8787   assumes "linear f"
       
  8788     and "convex S"
       
  8789     and "rel_open S"
       
  8790   shows "convex (f ` S) \<and> rel_open (f ` S)"
       
  8791   by (metis assms convex_linear_image rel_interior_convex_linear_image rel_open_def)
       
  8792 
       
  8793 lemma convex_rel_open_linear_preimage:
       
  8794   fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
       
  8795   assumes "linear f"
       
  8796     and "convex S"
       
  8797     and "rel_open S"
       
  8798   shows "convex (f -` S) \<and> rel_open (f -` S)"
       
  8799 proof (cases "f -` (rel_interior S) = {}")
       
  8800   case True
       
  8801   then have "f -` S = {}"
       
  8802     using assms unfolding rel_open_def by auto
       
  8803   then show ?thesis
       
  8804     unfolding rel_open_def using rel_interior_empty by auto
       
  8805 next
       
  8806   case False
       
  8807   then have "rel_open (f -` S)"
       
  8808     using assms unfolding rel_open_def
       
  8809     using rel_interior_convex_linear_preimage[of f S]
       
  8810     by auto
       
  8811   then show ?thesis
       
  8812     using convex_linear_vimage assms
       
  8813     by auto
       
  8814 qed
       
  8815 
       
  8816 lemma rel_interior_projection:
       
  8817   fixes S :: "('m::euclidean_space \<times> 'n::euclidean_space) set"
       
  8818     and f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space set"
       
  8819   assumes "convex S"
       
  8820     and "f = (\<lambda>y. {z. (y, z) \<in> S})"
       
  8821   shows "(y, z) \<in> rel_interior S \<longleftrightarrow> (y \<in> rel_interior {y. (f y \<noteq> {})} \<and> z \<in> rel_interior (f y))"
       
  8822 proof -
       
  8823   {
       
  8824     fix y
       
  8825     assume "y \<in> {y. f y \<noteq> {}}"
       
  8826     then obtain z where "(y, z) \<in> S"
       
  8827       using assms by auto
       
  8828     then have "\<exists>x. x \<in> S \<and> y = fst x"
       
  8829       apply (rule_tac x="(y, z)" in exI)
       
  8830       apply auto
       
  8831       done
       
  8832     then obtain x where "x \<in> S" "y = fst x"
       
  8833       by blast
       
  8834     then have "y \<in> fst ` S"
       
  8835       unfolding image_def by auto
       
  8836   }
       
  8837   then have "fst ` S = {y. f y \<noteq> {}}"
       
  8838     unfolding fst_def using assms by auto
       
  8839   then have h1: "fst ` rel_interior S = rel_interior {y. f y \<noteq> {}}"
       
  8840     using rel_interior_convex_linear_image[of fst S] assms fst_linear by auto
       
  8841   {
       
  8842     fix y
       
  8843     assume "y \<in> rel_interior {y. f y \<noteq> {}}"
       
  8844     then have "y \<in> fst ` rel_interior S"
       
  8845       using h1 by auto
       
  8846     then have *: "rel_interior S \<inter> fst -` {y} \<noteq> {}"
       
  8847       by auto
       
  8848     moreover have aff: "affine (fst -` {y})"
       
  8849       unfolding affine_alt by (simp add: algebra_simps)
       
  8850     ultimately have **: "rel_interior (S \<inter> fst -` {y}) = rel_interior S \<inter> fst -` {y}"
       
  8851       using convex_affine_rel_interior_Int[of S "fst -` {y}"] assms by auto
       
  8852     have conv: "convex (S \<inter> fst -` {y})"
       
  8853       using convex_Int assms aff affine_imp_convex by auto
       
  8854     {
       
  8855       fix x
       
  8856       assume "x \<in> f y"
       
  8857       then have "(y, x) \<in> S \<inter> (fst -` {y})"
       
  8858         using assms by auto
       
  8859       moreover have "x = snd (y, x)" by auto
       
  8860       ultimately have "x \<in> snd ` (S \<inter> fst -` {y})"
       
  8861         by blast
       
  8862     }
       
  8863     then have "snd ` (S \<inter> fst -` {y}) = f y"
       
  8864       using assms by auto
       
  8865     then have ***: "rel_interior (f y) = snd ` rel_interior (S \<inter> fst -` {y})"
       
  8866       using rel_interior_convex_linear_image[of snd "S \<inter> fst -` {y}"] snd_linear conv
       
  8867       by auto
       
  8868     {
       
  8869       fix z
       
  8870       assume "z \<in> rel_interior (f y)"
       
  8871       then have "z \<in> snd ` rel_interior (S \<inter> fst -` {y})"
       
  8872         using *** by auto
       
  8873       moreover have "{y} = fst ` rel_interior (S \<inter> fst -` {y})"
       
  8874         using * ** rel_interior_subset by auto
       
  8875       ultimately have "(y, z) \<in> rel_interior (S \<inter> fst -` {y})"
       
  8876         by force
       
  8877       then have "(y,z) \<in> rel_interior S"
       
  8878         using ** by auto
       
  8879     }
       
  8880     moreover
       
  8881     {
       
  8882       fix z
       
  8883       assume "(y, z) \<in> rel_interior S"
       
  8884       then have "(y, z) \<in> rel_interior (S \<inter> fst -` {y})"
       
  8885         using ** by auto
       
  8886       then have "z \<in> snd ` rel_interior (S \<inter> fst -` {y})"
       
  8887         by (metis Range_iff snd_eq_Range)
       
  8888       then have "z \<in> rel_interior (f y)"
       
  8889         using *** by auto
       
  8890     }
       
  8891     ultimately have "\<And>z. (y, z) \<in> rel_interior S \<longleftrightarrow> z \<in> rel_interior (f y)"
       
  8892       by auto
       
  8893   }
       
  8894   then have h2: "\<And>y z. y \<in> rel_interior {t. f t \<noteq> {}} \<Longrightarrow>
       
  8895     (y, z) \<in> rel_interior S \<longleftrightarrow> z \<in> rel_interior (f y)"
       
  8896     by auto
       
  8897   {
       
  8898     fix y z
       
  8899     assume asm: "(y, z) \<in> rel_interior S"
       
  8900     then have "y \<in> fst ` rel_interior S"
       
  8901       by (metis Domain_iff fst_eq_Domain)
       
  8902     then have "y \<in> rel_interior {t. f t \<noteq> {}}"
       
  8903       using h1 by auto
       
  8904     then have "y \<in> rel_interior {t. f t \<noteq> {}}" and "(z : rel_interior (f y))"
       
  8905       using h2 asm by auto
       
  8906   }
       
  8907   then show ?thesis using h2 by blast
       
  8908 qed
       
  8909 
       
  8910 lemma rel_frontier_Times:
       
  8911   fixes S :: "'n::euclidean_space set"
       
  8912     and T :: "'m::euclidean_space set"
       
  8913   assumes "convex S"
       
  8914     and "convex T"
       
  8915   shows "rel_frontier S \<times> rel_frontier T \<subseteq> rel_frontier (S \<times> T)"
       
  8916     by (force simp: rel_frontier_def rel_interior_Times assms closure_Times)
       
  8917 
       
  8918 
       
  8919 subsubsection \<open>Relative interior of convex cone\<close>
       
  8920 
       
  8921 lemma cone_rel_interior:
       
  8922   fixes S :: "'m::euclidean_space set"
       
  8923   assumes "cone S"
       
  8924   shows "cone ({0} \<union> rel_interior S)"
       
  8925 proof (cases "S = {}")
       
  8926   case True
       
  8927   then show ?thesis
       
  8928     by (simp add: rel_interior_empty cone_0)
       
  8929 next
       
  8930   case False
       
  8931   then have *: "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` S = S)"
       
  8932     using cone_iff[of S] assms by auto
       
  8933   then have *: "0 \<in> ({0} \<union> rel_interior S)"
       
  8934     and "\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` ({0} \<union> rel_interior S) = ({0} \<union> rel_interior S)"
       
  8935     by (auto simp add: rel_interior_scaleR)
       
  8936   then show ?thesis
       
  8937     using cone_iff[of "{0} \<union> rel_interior S"] by auto
       
  8938 qed
       
  8939 
       
  8940 lemma rel_interior_convex_cone_aux:
       
  8941   fixes S :: "'m::euclidean_space set"
       
  8942   assumes "convex S"
       
  8943   shows "(c, x) \<in> rel_interior (cone hull ({(1 :: real)} \<times> S)) \<longleftrightarrow>
       
  8944     c > 0 \<and> x \<in> ((op *\<^sub>R c) ` (rel_interior S))"
       
  8945 proof (cases "S = {}")
       
  8946   case True
       
  8947   then show ?thesis
       
  8948     by (simp add: rel_interior_empty cone_hull_empty)
       
  8949 next
       
  8950   case False
       
  8951   then obtain s where "s \<in> S" by auto
       
  8952   have conv: "convex ({(1 :: real)} \<times> S)"
       
  8953     using convex_Times[of "{(1 :: real)}" S] assms convex_singleton[of "1 :: real"]
       
  8954     by auto
       
  8955   define f where "f y = {z. (y, z) \<in> cone hull ({1 :: real} \<times> S)}" for y
       
  8956   then have *: "(c, x) \<in> rel_interior (cone hull ({(1 :: real)} \<times> S)) =
       
  8957     (c \<in> rel_interior {y. f y \<noteq> {}} \<and> x \<in> rel_interior (f c))"
       
  8958     apply (subst rel_interior_projection[of "cone hull ({(1 :: real)} \<times> S)" f c x])
       
  8959     using convex_cone_hull[of "{(1 :: real)} \<times> S"] conv
       
  8960     apply auto
       
  8961     done
       
  8962   {
       
  8963     fix y :: real
       
  8964     assume "y \<ge> 0"
       
  8965     then have "y *\<^sub>R (1,s) \<in> cone hull ({1 :: real} \<times> S)"
       
  8966       using cone_hull_expl[of "{(1 :: real)} \<times> S"] \<open>s \<in> S\<close> by auto
       
  8967     then have "f y \<noteq> {}"
       
  8968       using f_def by auto
       
  8969   }
       
  8970   then have "{y. f y \<noteq> {}} = {0..}"
       
  8971     using f_def cone_hull_expl[of "{1 :: real} \<times> S"] by auto
       
  8972   then have **: "rel_interior {y. f y \<noteq> {}} = {0<..}"
       
  8973     using rel_interior_real_semiline by auto
       
  8974   {
       
  8975     fix c :: real
       
  8976     assume "c > 0"
       
  8977     then have "f c = (op *\<^sub>R c ` S)"
       
  8978       using f_def cone_hull_expl[of "{1 :: real} \<times> S"] by auto
       
  8979     then have "rel_interior (f c) = op *\<^sub>R c ` rel_interior S"
       
  8980       using rel_interior_convex_scaleR[of S c] assms by auto
       
  8981   }
       
  8982   then show ?thesis using * ** by auto
       
  8983 qed
       
  8984 
       
  8985 lemma rel_interior_convex_cone:
       
  8986   fixes S :: "'m::euclidean_space set"
       
  8987   assumes "convex S"
       
  8988   shows "rel_interior (cone hull ({1 :: real} \<times> S)) =
       
  8989     {(c, c *\<^sub>R x) | c x. c > 0 \<and> x \<in> rel_interior S}"
       
  8990   (is "?lhs = ?rhs")
       
  8991 proof -
       
  8992   {
       
  8993     fix z
       
  8994     assume "z \<in> ?lhs"
       
  8995     have *: "z = (fst z, snd z)"
       
  8996       by auto
       
  8997     have "z \<in> ?rhs"
       
  8998       using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms \<open>z \<in> ?lhs\<close>
       
  8999       apply auto
       
  9000       apply (rule_tac x = "fst z" in exI)
       
  9001       apply (rule_tac x = x in exI)
       
  9002       using *
       
  9003       apply auto
       
  9004       done
       
  9005   }
       
  9006   moreover
       
  9007   {
       
  9008     fix z
       
  9009     assume "z \<in> ?rhs"
       
  9010     then have "z \<in> ?lhs"
       
  9011       using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms
       
  9012       by auto
       
  9013   }
       
  9014   ultimately show ?thesis by blast
       
  9015 qed
       
  9016 
       
  9017 lemma convex_hull_finite_union:
       
  9018   assumes "finite I"
       
  9019   assumes "\<forall>i\<in>I. convex (S i) \<and> (S i) \<noteq> {}"
       
  9020   shows "convex hull (\<Union>(S ` I)) =
       
  9021     {setsum (\<lambda>i. c i *\<^sub>R s i) I | c s. (\<forall>i\<in>I. c i \<ge> 0) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. s i \<in> S i)}"
       
  9022   (is "?lhs = ?rhs")
       
  9023 proof -
       
  9024   have "?lhs \<supseteq> ?rhs"
       
  9025   proof
       
  9026     fix x
       
  9027     assume "x : ?rhs"
       
  9028     then obtain c s where *: "setsum (\<lambda>i. c i *\<^sub>R s i) I = x" "setsum c I = 1"
       
  9029       "(\<forall>i\<in>I. c i \<ge> 0) \<and> (\<forall>i\<in>I. s i \<in> S i)" by auto
       
  9030     then have "\<forall>i\<in>I. s i \<in> convex hull (\<Union>(S ` I))"
       
  9031       using hull_subset[of "\<Union>(S ` I)" convex] by auto
       
  9032     then show "x \<in> ?lhs"
       
  9033       unfolding *(1)[symmetric]
       
  9034       apply (subst convex_setsum[of I "convex hull \<Union>(S ` I)" c s])
       
  9035       using * assms convex_convex_hull
       
  9036       apply auto
       
  9037       done
       
  9038   qed
       
  9039 
       
  9040   {
       
  9041     fix i
       
  9042     assume "i \<in> I"
       
  9043     with assms have "\<exists>p. p \<in> S i" by auto
       
  9044   }
       
  9045   then obtain p where p: "\<forall>i\<in>I. p i \<in> S i" by metis
       
  9046 
       
  9047   {
       
  9048     fix i
       
  9049     assume "i \<in> I"
       
  9050     {
       
  9051       fix x
       
  9052       assume "x \<in> S i"
       
  9053       define c where "c j = (if j = i then 1::real else 0)" for j
       
  9054       then have *: "setsum c I = 1"
       
  9055         using \<open>finite I\<close> \<open>i \<in> I\<close> setsum.delta[of I i "\<lambda>j::'a. 1::real"]
       
  9056         by auto
       
  9057       define s where "s j = (if j = i then x else p j)" for j
       
  9058       then have "\<forall>j. c j *\<^sub>R s j = (if j = i then x else 0)"
       
  9059         using c_def by (auto simp add: algebra_simps)
       
  9060       then have "x = setsum (\<lambda>i. c i *\<^sub>R s i) I"
       
  9061         using s_def c_def \<open>finite I\<close> \<open>i \<in> I\<close> setsum.delta[of I i "\<lambda>j::'a. x"]
       
  9062         by auto
       
  9063       then have "x \<in> ?rhs"
       
  9064         apply auto
       
  9065         apply (rule_tac x = c in exI)
       
  9066         apply (rule_tac x = s in exI)
       
  9067         using * c_def s_def p \<open>x \<in> S i\<close>
       
  9068         apply auto
       
  9069         done
       
  9070     }
       
  9071     then have "?rhs \<supseteq> S i" by auto
       
  9072   }
       
  9073   then have *: "?rhs \<supseteq> \<Union>(S ` I)" by auto
       
  9074 
       
  9075   {
       
  9076     fix u v :: real
       
  9077     assume uv: "u \<ge> 0 \<and> v \<ge> 0 \<and> u + v = 1"
       
  9078     fix x y
       
  9079     assume xy: "x \<in> ?rhs \<and> y \<in> ?rhs"
       
  9080     from xy obtain c s where
       
  9081       xc: "x = setsum (\<lambda>i. c i *\<^sub>R s i) I \<and> (\<forall>i\<in>I. c i \<ge> 0) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. s i \<in> S i)"
       
  9082       by auto
       
  9083     from xy obtain d t where
       
  9084       yc: "y = setsum (\<lambda>i. d i *\<^sub>R t i) I \<and> (\<forall>i\<in>I. d i \<ge> 0) \<and> setsum d I = 1 \<and> (\<forall>i\<in>I. t i \<in> S i)"
       
  9085       by auto
       
  9086     define e where "e i = u * c i + v * d i" for i
       
  9087     have ge0: "\<forall>i\<in>I. e i \<ge> 0"
       
  9088       using e_def xc yc uv by simp
       
  9089     have "setsum (\<lambda>i. u * c i) I = u * setsum c I"
       
  9090       by (simp add: setsum_right_distrib)
       
  9091     moreover have "setsum (\<lambda>i. v * d i) I = v * setsum d I"
       
  9092       by (simp add: setsum_right_distrib)
       
  9093     ultimately have sum1: "setsum e I = 1"
       
  9094       using e_def xc yc uv by (simp add: setsum.distrib)
       
  9095     define q where "q i = (if e i = 0 then p i else (u * c i / e i) *\<^sub>R s i + (v * d i / e i) *\<^sub>R t i)"
       
  9096       for i
       
  9097     {
       
  9098       fix i
       
  9099       assume i: "i \<in> I"
       
  9100       have "q i \<in> S i"
       
  9101       proof (cases "e i = 0")
       
  9102         case True
       
  9103         then show ?thesis using i p q_def by auto
       
  9104       next
       
  9105         case False
       
  9106         then show ?thesis
       
  9107           using mem_convex_alt[of "S i" "s i" "t i" "u * (c i)" "v * (d i)"]
       
  9108             mult_nonneg_nonneg[of u "c i"] mult_nonneg_nonneg[of v "d i"]
       
  9109             assms q_def e_def i False xc yc uv
       
  9110           by (auto simp del: mult_nonneg_nonneg)
       
  9111       qed
       
  9112     }
       
  9113     then have qs: "\<forall>i\<in>I. q i \<in> S i" by auto
       
  9114     {
       
  9115       fix i
       
  9116       assume i: "i \<in> I"
       
  9117       have "(u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i = e i *\<^sub>R q i"
       
  9118       proof (cases "e i = 0")
       
  9119         case True
       
  9120         have ge: "u * (c i) \<ge> 0 \<and> v * d i \<ge> 0"
       
  9121           using xc yc uv i by simp
       
  9122         moreover from ge have "u * c i \<le> 0 \<and> v * d i \<le> 0"
       
  9123           using True e_def i by simp
       
  9124         ultimately have "u * c i = 0 \<and> v * d i = 0" by auto
       
  9125         with True show ?thesis by auto
       
  9126       next
       
  9127         case False
       
  9128         then have "(u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i) = q i"
       
  9129           using q_def by auto
       
  9130         then have "e i *\<^sub>R ((u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i))
       
  9131                = (e i) *\<^sub>R (q i)" by auto
       
  9132         with False show ?thesis by (simp add: algebra_simps)
       
  9133       qed
       
  9134     }
       
  9135     then have *: "\<forall>i\<in>I. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i = e i *\<^sub>R q i"
       
  9136       by auto
       
  9137     have "u *\<^sub>R x + v *\<^sub>R y = setsum (\<lambda>i. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i) I"
       
  9138       using xc yc by (simp add: algebra_simps scaleR_right.setsum setsum.distrib)
       
  9139     also have "\<dots> = setsum (\<lambda>i. e i *\<^sub>R q i) I"
       
  9140       using * by auto
       
  9141     finally have "u *\<^sub>R x + v *\<^sub>R y = setsum (\<lambda>i. (e i) *\<^sub>R (q i)) I"
       
  9142       by auto
       
  9143     then have "u *\<^sub>R x + v *\<^sub>R y \<in> ?rhs"
       
  9144       using ge0 sum1 qs by auto
       
  9145   }
       
  9146   then have "convex ?rhs" unfolding convex_def by auto
       
  9147   then show ?thesis
       
  9148     using \<open>?lhs \<supseteq> ?rhs\<close> * hull_minimal[of "\<Union>(S ` I)" ?rhs convex]
       
  9149     by blast
       
  9150 qed
       
  9151 
       
  9152 lemma convex_hull_union_two:
       
  9153   fixes S T :: "'m::euclidean_space set"
       
  9154   assumes "convex S"
       
  9155     and "S \<noteq> {}"
       
  9156     and "convex T"
       
  9157     and "T \<noteq> {}"
       
  9158   shows "convex hull (S \<union> T) =
       
  9159     {u *\<^sub>R s + v *\<^sub>R t | u v s t. u \<ge> 0 \<and> v \<ge> 0 \<and> u + v = 1 \<and> s \<in> S \<and> t \<in> T}"
       
  9160   (is "?lhs = ?rhs")
       
  9161 proof
       
  9162   define I :: "nat set" where "I = {1, 2}"
       
  9163   define s where "s i = (if i = (1::nat) then S else T)" for i
       
  9164   have "\<Union>(s ` I) = S \<union> T"
       
  9165     using s_def I_def by auto
       
  9166   then have "convex hull (\<Union>(s ` I)) = convex hull (S \<union> T)"
       
  9167     by auto
       
  9168   moreover have "convex hull \<Union>(s ` I) =
       
  9169     {\<Sum> i\<in>I. c i *\<^sub>R sa i | c sa. (\<forall>i\<in>I. 0 \<le> c i) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. sa i \<in> s i)}"
       
  9170       apply (subst convex_hull_finite_union[of I s])
       
  9171       using assms s_def I_def
       
  9172       apply auto
       
  9173       done
       
  9174   moreover have
       
  9175     "{\<Sum>i\<in>I. c i *\<^sub>R sa i | c sa. (\<forall>i\<in>I. 0 \<le> c i) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. sa i \<in> s i)} \<le> ?rhs"
       
  9176     using s_def I_def by auto
       
  9177   ultimately show "?lhs \<subseteq> ?rhs" by auto
       
  9178   {
       
  9179     fix x
       
  9180     assume "x \<in> ?rhs"
       
  9181     then obtain u v s t where *: "x = u *\<^sub>R s + v *\<^sub>R t \<and> u \<ge> 0 \<and> v \<ge> 0 \<and> u + v = 1 \<and> s \<in> S \<and> t \<in> T"
       
  9182       by auto
       
  9183     then have "x \<in> convex hull {s, t}"
       
  9184       using convex_hull_2[of s t] by auto
       
  9185     then have "x \<in> convex hull (S \<union> T)"
       
  9186       using * hull_mono[of "{s, t}" "S \<union> T"] by auto
       
  9187   }
       
  9188   then show "?lhs \<supseteq> ?rhs" by blast
       
  9189 qed
       
  9190 
       
  9191 
       
  9192 subsection \<open>Convexity on direct sums\<close>
       
  9193 
       
  9194 lemma closure_sum:
       
  9195   fixes S T :: "'a::real_normed_vector set"
       
  9196   shows "closure S + closure T \<subseteq> closure (S + T)"
       
  9197   unfolding set_plus_image closure_Times [symmetric] split_def
       
  9198   by (intro closure_bounded_linear_image_subset bounded_linear_add
       
  9199     bounded_linear_fst bounded_linear_snd)
       
  9200 
       
  9201 lemma rel_interior_sum:
       
  9202   fixes S T :: "'n::euclidean_space set"
       
  9203   assumes "convex S"
       
  9204     and "convex T"
       
  9205   shows "rel_interior (S + T) = rel_interior S + rel_interior T"
       
  9206 proof -
       
  9207   have "rel_interior S + rel_interior T = (\<lambda>(x,y). x + y) ` (rel_interior S \<times> rel_interior T)"
       
  9208     by (simp add: set_plus_image)
       
  9209   also have "\<dots> = (\<lambda>(x,y). x + y) ` rel_interior (S \<times> T)"
       
  9210     using rel_interior_Times assms by auto
       
  9211   also have "\<dots> = rel_interior (S + T)"
       
  9212     using fst_snd_linear convex_Times assms
       
  9213       rel_interior_convex_linear_image[of "(\<lambda>(x,y). x + y)" "S \<times> T"]
       
  9214     by (auto simp add: set_plus_image)
       
  9215   finally show ?thesis ..
       
  9216 qed
       
  9217 
       
  9218 lemma rel_interior_sum_gen:
       
  9219   fixes S :: "'a \<Rightarrow> 'n::euclidean_space set"
       
  9220   assumes "\<forall>i\<in>I. convex (S i)"
       
  9221   shows "rel_interior (setsum S I) = setsum (\<lambda>i. rel_interior (S i)) I"
       
  9222   apply (subst setsum_set_cond_linear[of convex])
       
  9223   using rel_interior_sum rel_interior_sing[of "0"] assms
       
  9224   apply (auto simp add: convex_set_plus)
       
  9225   done
       
  9226 
       
  9227 lemma convex_rel_open_direct_sum:
       
  9228   fixes S T :: "'n::euclidean_space set"
       
  9229   assumes "convex S"
       
  9230     and "rel_open S"
       
  9231     and "convex T"
       
  9232     and "rel_open T"
       
  9233   shows "convex (S \<times> T) \<and> rel_open (S \<times> T)"
       
  9234   by (metis assms convex_Times rel_interior_Times rel_open_def)
       
  9235 
       
  9236 lemma convex_rel_open_sum:
       
  9237   fixes S T :: "'n::euclidean_space set"
       
  9238   assumes "convex S"
       
  9239     and "rel_open S"
       
  9240     and "convex T"
       
  9241     and "rel_open T"
       
  9242   shows "convex (S + T) \<and> rel_open (S + T)"
       
  9243   by (metis assms convex_set_plus rel_interior_sum rel_open_def)
       
  9244 
       
  9245 lemma convex_hull_finite_union_cones:
       
  9246   assumes "finite I"
       
  9247     and "I \<noteq> {}"
       
  9248   assumes "\<forall>i\<in>I. convex (S i) \<and> cone (S i) \<and> S i \<noteq> {}"
       
  9249   shows "convex hull (\<Union>(S ` I)) = setsum S I"
       
  9250   (is "?lhs = ?rhs")
       
  9251 proof -
       
  9252   {
       
  9253     fix x
       
  9254     assume "x \<in> ?lhs"
       
  9255     then obtain c xs where
       
  9256       x: "x = setsum (\<lambda>i. c i *\<^sub>R xs i) I \<and> (\<forall>i\<in>I. c i \<ge> 0) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. xs i \<in> S i)"
       
  9257       using convex_hull_finite_union[of I S] assms by auto
       
  9258     define s where "s i = c i *\<^sub>R xs i" for i
       
  9259     {
       
  9260       fix i
       
  9261       assume "i \<in> I"
       
  9262       then have "s i \<in> S i"
       
  9263         using s_def x assms mem_cone[of "S i" "xs i" "c i"] by auto
       
  9264     }
       
  9265     then have "\<forall>i\<in>I. s i \<in> S i" by auto
       
  9266     moreover have "x = setsum s I" using x s_def by auto
       
  9267     ultimately have "x \<in> ?rhs"
       
  9268       using set_setsum_alt[of I S] assms by auto
       
  9269   }
       
  9270   moreover
       
  9271   {
       
  9272     fix x
       
  9273     assume "x \<in> ?rhs"
       
  9274     then obtain s where x: "x = setsum s I \<and> (\<forall>i\<in>I. s i \<in> S i)"
       
  9275       using set_setsum_alt[of I S] assms by auto
       
  9276     define xs where "xs i = of_nat(card I) *\<^sub>R s i" for i
       
  9277     then have "x = setsum (\<lambda>i. ((1 :: real) / of_nat(card I)) *\<^sub>R xs i) I"
       
  9278       using x assms by auto
       
  9279     moreover have "\<forall>i\<in>I. xs i \<in> S i"
       
  9280       using x xs_def assms by (simp add: cone_def)
       
  9281     moreover have "\<forall>i\<in>I. (1 :: real) / of_nat (card I) \<ge> 0"
       
  9282       by auto
       
  9283     moreover have "setsum (\<lambda>i. (1 :: real) / of_nat (card I)) I = 1"
       
  9284       using assms by auto
       
  9285     ultimately have "x \<in> ?lhs"
       
  9286       apply (subst convex_hull_finite_union[of I S])
       
  9287       using assms
       
  9288       apply blast
       
  9289       using assms
       
  9290       apply blast
       
  9291       apply rule
       
  9292       apply (rule_tac x = "(\<lambda>i. (1 :: real) / of_nat (card I))" in exI)
       
  9293       apply auto
       
  9294       done
       
  9295   }
       
  9296   ultimately show ?thesis by auto
       
  9297 qed
       
  9298 
       
  9299 lemma convex_hull_union_cones_two:
       
  9300   fixes S T :: "'m::euclidean_space set"
       
  9301   assumes "convex S"
       
  9302     and "cone S"
       
  9303     and "S \<noteq> {}"
       
  9304   assumes "convex T"
       
  9305     and "cone T"
       
  9306     and "T \<noteq> {}"
       
  9307   shows "convex hull (S \<union> T) = S + T"
       
  9308 proof -
       
  9309   define I :: "nat set" where "I = {1, 2}"
       
  9310   define A where "A i = (if i = (1::nat) then S else T)" for i
       
  9311   have "\<Union>(A ` I) = S \<union> T"
       
  9312     using A_def I_def by auto
       
  9313   then have "convex hull (\<Union>(A ` I)) = convex hull (S \<union> T)"
       
  9314     by auto
       
  9315   moreover have "convex hull \<Union>(A ` I) = setsum A I"
       
  9316     apply (subst convex_hull_finite_union_cones[of I A])
       
  9317     using assms A_def I_def
       
  9318     apply auto
       
  9319     done
       
  9320   moreover have "setsum A I = S + T"
       
  9321     using A_def I_def
       
  9322     unfolding set_plus_def
       
  9323     apply auto
       
  9324     unfolding set_plus_def
       
  9325     apply auto
       
  9326     done
       
  9327   ultimately show ?thesis by auto
       
  9328 qed
       
  9329 
       
  9330 lemma rel_interior_convex_hull_union:
       
  9331   fixes S :: "'a \<Rightarrow> 'n::euclidean_space set"
       
  9332   assumes "finite I"
       
  9333     and "\<forall>i\<in>I. convex (S i) \<and> S i \<noteq> {}"
       
  9334   shows "rel_interior (convex hull (\<Union>(S ` I))) =
       
  9335     {setsum (\<lambda>i. c i *\<^sub>R s i) I | c s. (\<forall>i\<in>I. c i > 0) \<and> setsum c I = 1 \<and>
       
  9336       (\<forall>i\<in>I. s i \<in> rel_interior(S i))}"
       
  9337   (is "?lhs = ?rhs")
       
  9338 proof (cases "I = {}")
       
  9339   case True
       
  9340   then show ?thesis
       
  9341     using convex_hull_empty rel_interior_empty by auto
       
  9342 next
       
  9343   case False
       
  9344   define C0 where "C0 = convex hull (\<Union>(S ` I))"
       
  9345   have "\<forall>i\<in>I. C0 \<ge> S i"
       
  9346     unfolding C0_def using hull_subset[of "\<Union>(S ` I)"] by auto
       
  9347   define K0 where "K0 = cone hull ({1 :: real} \<times> C0)"
       
  9348   define K where "K i = cone hull ({1 :: real} \<times> S i)" for i
       
  9349   have "\<forall>i\<in>I. K i \<noteq> {}"
       
  9350     unfolding K_def using assms
       
  9351     by (simp add: cone_hull_empty_iff[symmetric])
       
  9352   {
       
  9353     fix i
       
  9354     assume "i \<in> I"
       
  9355     then have "convex (K i)"
       
  9356       unfolding K_def
       
  9357       apply (subst convex_cone_hull)
       
  9358       apply (subst convex_Times)
       
  9359       using assms
       
  9360       apply auto
       
  9361       done
       
  9362   }
       
  9363   then have convK: "\<forall>i\<in>I. convex (K i)"
       
  9364     by auto
       
  9365   {
       
  9366     fix i
       
  9367     assume "i \<in> I"
       
  9368     then have "K0 \<supseteq> K i"
       
  9369       unfolding K0_def K_def
       
  9370       apply (subst hull_mono)
       
  9371       using \<open>\<forall>i\<in>I. C0 \<ge> S i\<close>
       
  9372       apply auto
       
  9373       done
       
  9374   }
       
  9375   then have "K0 \<supseteq> \<Union>(K ` I)" by auto
       
  9376   moreover have "convex K0"
       
  9377     unfolding K0_def
       
  9378     apply (subst convex_cone_hull)
       
  9379     apply (subst convex_Times)
       
  9380     unfolding C0_def
       
  9381     using convex_convex_hull
       
  9382     apply auto
       
  9383     done
       
  9384   ultimately have geq: "K0 \<supseteq> convex hull (\<Union>(K ` I))"
       
  9385     using hull_minimal[of _ "K0" "convex"] by blast
       
  9386   have "\<forall>i\<in>I. K i \<supseteq> {1 :: real} \<times> S i"
       
  9387     using K_def by (simp add: hull_subset)
       
  9388   then have "\<Union>(K ` I) \<supseteq> {1 :: real} \<times> \<Union>(S ` I)"
       
  9389     by auto
       
  9390   then have "convex hull \<Union>(K ` I) \<supseteq> convex hull ({1 :: real} \<times> \<Union>(S ` I))"
       
  9391     by (simp add: hull_mono)
       
  9392   then have "convex hull \<Union>(K ` I) \<supseteq> {1 :: real} \<times> C0"
       
  9393     unfolding C0_def
       
  9394     using convex_hull_Times[of "{(1 :: real)}" "\<Union>(S ` I)"] convex_hull_singleton
       
  9395     by auto
       
  9396   moreover have "cone (convex hull (\<Union>(K ` I)))"
       
  9397     apply (subst cone_convex_hull)
       
  9398     using cone_Union[of "K ` I"]
       
  9399     apply auto
       
  9400     unfolding K_def
       
  9401     using cone_cone_hull
       
  9402     apply auto
       
  9403     done
       
  9404   ultimately have "convex hull (\<Union>(K ` I)) \<supseteq> K0"
       
  9405     unfolding K0_def
       
  9406     using hull_minimal[of _ "convex hull (\<Union>(K ` I))" "cone"]
       
  9407     by blast
       
  9408   then have "K0 = convex hull (\<Union>(K ` I))"
       
  9409     using geq by auto
       
  9410   also have "\<dots> = setsum K I"
       
  9411     apply (subst convex_hull_finite_union_cones[of I K])
       
  9412     using assms
       
  9413     apply blast
       
  9414     using False
       
  9415     apply blast
       
  9416     unfolding K_def
       
  9417     apply rule
       
  9418     apply (subst convex_cone_hull)
       
  9419     apply (subst convex_Times)
       
  9420     using assms cone_cone_hull \<open>\<forall>i\<in>I. K i \<noteq> {}\<close> K_def
       
  9421     apply auto
       
  9422     done
       
  9423   finally have "K0 = setsum K I" by auto
       
  9424   then have *: "rel_interior K0 = setsum (\<lambda>i. (rel_interior (K i))) I"
       
  9425     using rel_interior_sum_gen[of I K] convK by auto
       
  9426   {
       
  9427     fix x
       
  9428     assume "x \<in> ?lhs"
       
  9429     then have "(1::real, x) \<in> rel_interior K0"
       
  9430       using K0_def C0_def rel_interior_convex_cone_aux[of C0 "1::real" x] convex_convex_hull
       
  9431       by auto
       
  9432     then obtain k where k: "(1::real, x) = setsum k I \<and> (\<forall>i\<in>I. k i \<in> rel_interior (K i))"
       
  9433       using \<open>finite I\<close> * set_setsum_alt[of I "\<lambda>i. rel_interior (K i)"] by auto
       
  9434     {
       
  9435       fix i
       
  9436       assume "i \<in> I"
       
  9437       then have "convex (S i) \<and> k i \<in> rel_interior (cone hull {1} \<times> S i)"
       
  9438         using k K_def assms by auto
       
  9439       then have "\<exists>ci si. k i = (ci, ci *\<^sub>R si) \<and> 0 < ci \<and> si \<in> rel_interior (S i)"
       
  9440         using rel_interior_convex_cone[of "S i"] by auto
       
  9441     }
       
  9442     then obtain c s where
       
  9443       cs: "\<forall>i\<in>I. k i = (c i, c i *\<^sub>R s i) \<and> 0 < c i \<and> s i \<in> rel_interior (S i)"
       
  9444       by metis
       
  9445     then have "x = (\<Sum>i\<in>I. c i *\<^sub>R s i) \<and> setsum c I = 1"
       
  9446       using k by (simp add: setsum_prod)
       
  9447     then have "x \<in> ?rhs"
       
  9448       using k
       
  9449       apply auto
       
  9450       apply (rule_tac x = c in exI)
       
  9451       apply (rule_tac x = s in exI)
       
  9452       using cs
       
  9453       apply auto
       
  9454       done
       
  9455   }
       
  9456   moreover
       
  9457   {
       
  9458     fix x
       
  9459     assume "x \<in> ?rhs"
       
  9460     then obtain c s where cs: "x = setsum (\<lambda>i. c i *\<^sub>R s i) I \<and>
       
  9461         (\<forall>i\<in>I. c i > 0) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. s i \<in> rel_interior (S i))"
       
  9462       by auto
       
  9463     define k where "k i = (c i, c i *\<^sub>R s i)" for i
       
  9464     {
       
  9465       fix i assume "i:I"
       
  9466       then have "k i \<in> rel_interior (K i)"
       
  9467         using k_def K_def assms cs rel_interior_convex_cone[of "S i"]
       
  9468         by auto
       
  9469     }
       
  9470     then have "(1::real, x) \<in> rel_interior K0"
       
  9471       using K0_def * set_setsum_alt[of I "(\<lambda>i. rel_interior (K i))"] assms k_def cs
       
  9472       apply auto
       
  9473       apply (rule_tac x = k in exI)
       
  9474       apply (simp add: setsum_prod)
       
  9475       done
       
  9476     then have "x \<in> ?lhs"
       
  9477       using K0_def C0_def rel_interior_convex_cone_aux[of C0 1 x]
       
  9478       by (auto simp add: convex_convex_hull)
       
  9479   }
       
  9480   ultimately show ?thesis by blast
       
  9481 qed
       
  9482 
       
  9483 
       
  9484 lemma convex_le_Inf_differential:
       
  9485   fixes f :: "real \<Rightarrow> real"
       
  9486   assumes "convex_on I f"
       
  9487     and "x \<in> interior I"
       
  9488     and "y \<in> I"
       
  9489   shows "f y \<ge> f x + Inf ((\<lambda>t. (f x - f t) / (x - t)) ` ({x<..} \<inter> I)) * (y - x)"
       
  9490   (is "_ \<ge> _ + Inf (?F x) * (y - x)")
       
  9491 proof (cases rule: linorder_cases)
       
  9492   assume "x < y"
       
  9493   moreover
       
  9494   have "open (interior I)" by auto
       
  9495   from openE[OF this \<open>x \<in> interior I\<close>]
       
  9496   obtain e where e: "0 < e" "ball x e \<subseteq> interior I" .
       
  9497   moreover define t where "t = min (x + e / 2) ((x + y) / 2)"
       
  9498   ultimately have "x < t" "t < y" "t \<in> ball x e"
       
  9499     by (auto simp: dist_real_def field_simps split: split_min)
       
  9500   with \<open>x \<in> interior I\<close> e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto
       
  9501 
       
  9502   have "open (interior I)" by auto
       
  9503   from openE[OF this \<open>x \<in> interior I\<close>]
       
  9504   obtain e where "0 < e" "ball x e \<subseteq> interior I" .
       
  9505   moreover define K where "K = x - e / 2"
       
  9506   with \<open>0 < e\<close> have "K \<in> ball x e" "K < x"
       
  9507     by (auto simp: dist_real_def)
       
  9508   ultimately have "K \<in> I" "K < x" "x \<in> I"
       
  9509     using interior_subset[of I] \<open>x \<in> interior I\<close> by auto
       
  9510 
       
  9511   have "Inf (?F x) \<le> (f x - f y) / (x - y)"
       
  9512   proof (intro bdd_belowI cInf_lower2)
       
  9513     show "(f x - f t) / (x - t) \<in> ?F x"
       
  9514       using \<open>t \<in> I\<close> \<open>x < t\<close> by auto
       
  9515     show "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)"
       
  9516       using \<open>convex_on I f\<close> \<open>x \<in> I\<close> \<open>y \<in> I\<close> \<open>x < t\<close> \<open>t < y\<close>
       
  9517       by (rule convex_on_diff)
       
  9518   next
       
  9519     fix y
       
  9520     assume "y \<in> ?F x"
       
  9521     with order_trans[OF convex_on_diff[OF \<open>convex_on I f\<close> \<open>K \<in> I\<close> _ \<open>K < x\<close> _]]
       
  9522     show "(f K - f x) / (K - x) \<le> y" by auto
       
  9523   qed
       
  9524   then show ?thesis
       
  9525     using \<open>x < y\<close> by (simp add: field_simps)
       
  9526 next
       
  9527   assume "y < x"
       
  9528   moreover
       
  9529   have "open (interior I)" by auto
       
  9530   from openE[OF this \<open>x \<in> interior I\<close>]
       
  9531   obtain e where e: "0 < e" "ball x e \<subseteq> interior I" .
       
  9532   moreover define t where "t = x + e / 2"
       
  9533   ultimately have "x < t" "t \<in> ball x e"
       
  9534     by (auto simp: dist_real_def field_simps)
       
  9535   with \<open>x \<in> interior I\<close> e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto
       
  9536 
       
  9537   have "(f x - f y) / (x - y) \<le> Inf (?F x)"
       
  9538   proof (rule cInf_greatest)
       
  9539     have "(f x - f y) / (x - y) = (f y - f x) / (y - x)"
       
  9540       using \<open>y < x\<close> by (auto simp: field_simps)
       
  9541     also
       
  9542     fix z
       
  9543     assume "z \<in> ?F x"
       
  9544     with order_trans[OF convex_on_diff[OF \<open>convex_on I f\<close> \<open>y \<in> I\<close> _ \<open>y < x\<close>]]
       
  9545     have "(f y - f x) / (y - x) \<le> z"
       
  9546       by auto
       
  9547     finally show "(f x - f y) / (x - y) \<le> z" .
       
  9548   next
       
  9549     have "open (interior I)" by auto
       
  9550     from openE[OF this \<open>x \<in> interior I\<close>]
       
  9551     obtain e where e: "0 < e" "ball x e \<subseteq> interior I" .
       
  9552     then have "x + e / 2 \<in> ball x e"
       
  9553       by (auto simp: dist_real_def)
       
  9554     with e interior_subset[of I] have "x + e / 2 \<in> {x<..} \<inter> I"
       
  9555       by auto
       
  9556     then show "?F x \<noteq> {}"
       
  9557       by blast
       
  9558   qed
       
  9559   then show ?thesis
       
  9560     using \<open>y < x\<close> by (simp add: field_simps)
       
  9561 qed simp
       
  9562 
       
  9563 subsection\<open>Explicit formulas for interior and relative interior of convex hull\<close>
       
  9564 
       
  9565 lemma interior_atLeastAtMost [simp]:
       
  9566   fixes a::real shows "interior {a..b} = {a<..<b}"
       
  9567   using interior_cbox [of a b] by auto
       
  9568 
       
  9569 lemma interior_atLeastLessThan [simp]:
       
  9570   fixes a::real shows "interior {a..<b} = {a<..<b}"
       
  9571   by (metis atLeastLessThan_def greaterThanLessThan_def interior_atLeastAtMost interior_Int interior_interior interior_real_semiline)
       
  9572 
       
  9573 lemma interior_lessThanAtMost [simp]:
       
  9574   fixes a::real shows "interior {a<..b} = {a<..<b}"
       
  9575   by (metis atLeastAtMost_def greaterThanAtMost_def interior_atLeastAtMost interior_Int
       
  9576             interior_interior interior_real_semiline)
       
  9577 
       
  9578 lemma at_within_closed_interval:
       
  9579     fixes x::real
       
  9580     shows "a < x \<Longrightarrow> x < b \<Longrightarrow> (at x within {a..b}) = at x"
       
  9581   by (metis at_within_interior greaterThanLessThan_iff interior_atLeastAtMost)
       
  9582 
       
  9583 lemma affine_independent_convex_affine_hull:
       
  9584   fixes s :: "'a::euclidean_space set"
       
  9585   assumes "~affine_dependent s" "t \<subseteq> s"
       
  9586     shows "convex hull t = affine hull t \<inter> convex hull s"
       
  9587 proof -
       
  9588   have fin: "finite s" "finite t" using assms aff_independent_finite finite_subset by auto
       
  9589     { fix u v x
       
  9590       assume uv: "setsum u t = 1" "\<forall>x\<in>s. 0 \<le> v x" "setsum v s = 1"
       
  9591                  "(\<Sum>x\<in>s. v x *\<^sub>R x) = (\<Sum>v\<in>t. u v *\<^sub>R v)" "x \<in> t"
       
  9592       then have s: "s = (s - t) \<union> t" \<comment>\<open>split into separate cases\<close>
       
  9593         using assms by auto
       
  9594       have [simp]: "(\<Sum>x\<in>t. v x *\<^sub>R x) + (\<Sum>x\<in>s - t. v x *\<^sub>R x) = (\<Sum>x\<in>t. u x *\<^sub>R x)"
       
  9595                    "setsum v t + setsum v (s - t) = 1"
       
  9596         using uv fin s
       
  9597         by (auto simp: setsum.union_disjoint [symmetric] Un_commute)
       
  9598       have "(\<Sum>x\<in>s. if x \<in> t then v x - u x else v x) = 0"
       
  9599            "(\<Sum>x\<in>s. (if x \<in> t then v x - u x else v x) *\<^sub>R x) = 0"
       
  9600         using uv fin
       
  9601         by (subst s, subst setsum.union_disjoint, auto simp: algebra_simps setsum_subtractf)+
       
  9602     } note [simp] = this
       
  9603   have "convex hull t \<subseteq> affine hull t"
       
  9604     using convex_hull_subset_affine_hull by blast
       
  9605   moreover have "convex hull t \<subseteq> convex hull s"
       
  9606     using assms hull_mono by blast
       
  9607   moreover have "affine hull t \<inter> convex hull s \<subseteq> convex hull t"
       
  9608     using assms
       
  9609     apply (simp add: convex_hull_finite affine_hull_finite fin affine_dependent_explicit)
       
  9610     apply (drule_tac x=s in spec)
       
  9611     apply (auto simp: fin)
       
  9612     apply (rule_tac x=u in exI)
       
  9613     apply (rename_tac v)
       
  9614     apply (drule_tac x="\<lambda>x. if x \<in> t then v x - u x else v x" in spec)
       
  9615     apply (force)+
       
  9616     done
       
  9617   ultimately show ?thesis
       
  9618     by blast
       
  9619 qed
       
  9620 
       
  9621 lemma affine_independent_span_eq:
       
  9622   fixes s :: "'a::euclidean_space set"
       
  9623   assumes "~affine_dependent s" "card s = Suc (DIM ('a))"
       
  9624     shows "affine hull s = UNIV"
       
  9625 proof (cases "s = {}")
       
  9626   case True then show ?thesis
       
  9627     using assms by simp
       
  9628 next
       
  9629   case False
       
  9630     then obtain a t where t: "a \<notin> t" "s = insert a t"
       
  9631       by blast
       
  9632     then have fin: "finite t" using assms
       
  9633       by (metis finite_insert aff_independent_finite)
       
  9634     show ?thesis
       
  9635     using assms t fin
       
  9636       apply (simp add: affine_dependent_iff_dependent affine_hull_insert_span_gen)
       
  9637       apply (rule subset_antisym)
       
  9638       apply force
       
  9639       apply (rule Fun.vimage_subsetD)
       
  9640       apply (metis add.commute diff_add_cancel surj_def)
       
  9641       apply (rule card_ge_dim_independent)
       
  9642       apply (auto simp: card_image inj_on_def dim_subset_UNIV)
       
  9643       done
       
  9644 qed
       
  9645 
       
  9646 lemma affine_independent_span_gt:
       
  9647   fixes s :: "'a::euclidean_space set"
       
  9648   assumes ind: "~ affine_dependent s" and dim: "DIM ('a) < card s"
       
  9649     shows "affine hull s = UNIV"
       
  9650   apply (rule affine_independent_span_eq [OF ind])
       
  9651   apply (rule antisym)
       
  9652   using assms
       
  9653   apply auto
       
  9654   apply (metis add_2_eq_Suc' not_less_eq_eq affine_dependent_biggerset aff_independent_finite)
       
  9655   done
       
  9656 
       
  9657 lemma empty_interior_affine_hull:
       
  9658   fixes s :: "'a::euclidean_space set"
       
  9659   assumes "finite s" and dim: "card s \<le> DIM ('a)"
       
  9660     shows "interior(affine hull s) = {}"
       
  9661   using assms
       
  9662   apply (induct s rule: finite_induct)
       
  9663   apply (simp_all add:  affine_dependent_iff_dependent affine_hull_insert_span_gen interior_translation)
       
  9664   apply (rule empty_interior_lowdim)
       
  9665   apply (simp add: affine_dependent_iff_dependent affine_hull_insert_span_gen)
       
  9666   apply (metis Suc_le_lessD not_less order_trans card_image_le finite_imageI dim_le_card)
       
  9667   done
       
  9668 
       
  9669 lemma empty_interior_convex_hull:
       
  9670   fixes s :: "'a::euclidean_space set"
       
  9671   assumes "finite s" and dim: "card s \<le> DIM ('a)"
       
  9672     shows "interior(convex hull s) = {}"
       
  9673   by (metis Diff_empty Diff_eq_empty_iff convex_hull_subset_affine_hull
       
  9674             interior_mono empty_interior_affine_hull [OF assms])
       
  9675 
       
  9676 lemma explicit_subset_rel_interior_convex_hull:
       
  9677   fixes s :: "'a::euclidean_space set"
       
  9678   shows "finite s
       
  9679          \<Longrightarrow> {y. \<exists>u. (\<forall>x \<in> s. 0 < u x \<and> u x < 1) \<and> setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}
       
  9680              \<subseteq> rel_interior (convex hull s)"
       
  9681   by (force simp add:  rel_interior_convex_hull_union [where S="\<lambda>x. {x}" and I=s, simplified])
       
  9682 
       
  9683 lemma explicit_subset_rel_interior_convex_hull_minimal:
       
  9684   fixes s :: "'a::euclidean_space set"
       
  9685   shows "finite s
       
  9686          \<Longrightarrow> {y. \<exists>u. (\<forall>x \<in> s. 0 < u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}
       
  9687              \<subseteq> rel_interior (convex hull s)"
       
  9688   by (force simp add:  rel_interior_convex_hull_union [where S="\<lambda>x. {x}" and I=s, simplified])
       
  9689 
       
  9690 lemma rel_interior_convex_hull_explicit:
       
  9691   fixes s :: "'a::euclidean_space set"
       
  9692   assumes "~ affine_dependent s"
       
  9693   shows "rel_interior(convex hull s) =
       
  9694          {y. \<exists>u. (\<forall>x \<in> s. 0 < u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}"
       
  9695          (is "?lhs = ?rhs")
       
  9696 proof
       
  9697   show "?rhs \<le> ?lhs"
       
  9698     by (simp add: aff_independent_finite explicit_subset_rel_interior_convex_hull_minimal assms)
       
  9699 next
       
  9700   show "?lhs \<le> ?rhs"
       
  9701   proof (cases "\<exists>a. s = {a}")
       
  9702     case True then show "?lhs \<le> ?rhs"
       
  9703       by force
       
  9704   next
       
  9705     case False
       
  9706     have fs: "finite s"
       
  9707       using assms by (simp add: aff_independent_finite)
       
  9708     { fix a b and d::real
       
  9709       assume ab: "a \<in> s" "b \<in> s" "a \<noteq> b"
       
  9710       then have s: "s = (s - {a,b}) \<union> {a,b}" \<comment>\<open>split into separate cases\<close>
       
  9711         by auto
       
  9712       have "(\<Sum>x\<in>s. if x = a then - d else if x = b then d else 0) = 0"
       
  9713            "(\<Sum>x\<in>s. (if x = a then - d else if x = b then d else 0) *\<^sub>R x) = d *\<^sub>R b - d *\<^sub>R a"
       
  9714         using ab fs
       
  9715         by (subst s, subst setsum.union_disjoint, auto)+
       
  9716     } note [simp] = this
       
  9717     { fix y
       
  9718       assume y: "y \<in> convex hull s" "y \<notin> ?rhs"
       
  9719       { fix u T a
       
  9720         assume ua: "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = 1" "\<not> 0 < u a" "a \<in> s"
       
  9721            and yT: "y = (\<Sum>x\<in>s. u x *\<^sub>R x)" "y \<in> T" "open T"
       
  9722            and sb: "T \<inter> affine hull s \<subseteq> {w. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = w}"
       
  9723         have ua0: "u a = 0"
       
  9724           using ua by auto
       
  9725         obtain b where b: "b\<in>s" "a \<noteq> b"
       
  9726           using ua False by auto
       
  9727         obtain e where e: "0 < e" "ball (\<Sum>x\<in>s. u x *\<^sub>R x) e \<subseteq> T"
       
  9728           using yT by (auto elim: openE)
       
  9729         with b obtain d where d: "0 < d" "norm(d *\<^sub>R (a-b)) < e"
       
  9730           by (auto intro: that [of "e / 2 / norm(a-b)"])
       
  9731         have "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> affine hull s"
       
  9732           using yT y by (metis affine_hull_convex_hull hull_redundant_eq)
       
  9733         then have "(\<Sum>x\<in>s. u x *\<^sub>R x) - d *\<^sub>R (a - b) \<in> affine hull s"
       
  9734           using ua b by (auto simp: hull_inc intro: mem_affine_3_minus2)
       
  9735         then have "y - d *\<^sub>R (a - b) \<in> T \<inter> affine hull s"
       
  9736           using d e yT by auto
       
  9737         then obtain v where "\<forall>x\<in>s. 0 \<le> v x"
       
  9738                             "setsum v s = 1"
       
  9739                             "(\<Sum>x\<in>s. v x *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x) - d *\<^sub>R (a - b)"
       
  9740           using subsetD [OF sb] yT
       
  9741           by auto
       
  9742         then have False
       
  9743           using assms
       
  9744           apply (simp add: affine_dependent_explicit_finite fs)
       
  9745           apply (drule_tac x="\<lambda>x. (v x - u x) - (if x = a then -d else if x = b then d else 0)" in spec)
       
  9746           using ua b d
       
  9747           apply (auto simp: algebra_simps setsum_subtractf setsum.distrib)
       
  9748           done
       
  9749       } note * = this
       
  9750       have "y \<notin> rel_interior (convex hull s)"
       
  9751         using y
       
  9752         apply (simp add: mem_rel_interior affine_hull_convex_hull)
       
  9753         apply (auto simp: convex_hull_finite [OF fs])
       
  9754         apply (drule_tac x=u in spec)
       
  9755         apply (auto intro: *)
       
  9756         done
       
  9757     } with rel_interior_subset show "?lhs \<le> ?rhs"
       
  9758       by blast
       
  9759   qed
       
  9760 qed
       
  9761 
       
  9762 lemma interior_convex_hull_explicit_minimal:
       
  9763   fixes s :: "'a::euclidean_space set"
       
  9764   shows
       
  9765    "~ affine_dependent s
       
  9766         ==> interior(convex hull s) =
       
  9767              (if card(s) \<le> DIM('a) then {}
       
  9768               else {y. \<exists>u. (\<forall>x \<in> s. 0 < u x) \<and> setsum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = y})"
       
  9769   apply (simp add: aff_independent_finite empty_interior_convex_hull, clarify)
       
  9770   apply (rule trans [of _ "rel_interior(convex hull s)"])
       
  9771   apply (simp add: affine_hull_convex_hull affine_independent_span_gt rel_interior_interior)
       
  9772   by (simp add: rel_interior_convex_hull_explicit)
       
  9773 
       
  9774 lemma interior_convex_hull_explicit:
       
  9775   fixes s :: "'a::euclidean_space set"
       
  9776   assumes "~ affine_dependent s"
       
  9777   shows
       
  9778    "interior(convex hull s) =
       
  9779              (if card(s) \<le> DIM('a) then {}
       
  9780               else {y. \<exists>u. (\<forall>x \<in> s. 0 < u x \<and> u x < 1) \<and> setsum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = y})"
       
  9781 proof -
       
  9782   { fix u :: "'a \<Rightarrow> real" and a
       
  9783     assume "card Basis < card s" and u: "\<And>x. x\<in>s \<Longrightarrow> 0 < u x" "setsum u s = 1" and a: "a \<in> s"
       
  9784     then have cs: "Suc 0 < card s"
       
  9785       by (metis DIM_positive less_trans_Suc)
       
  9786     obtain b where b: "b \<in> s" "a \<noteq> b"
       
  9787     proof (cases "s \<le> {a}")
       
  9788       case True
       
  9789       then show thesis
       
  9790         using cs subset_singletonD by fastforce
       
  9791     next
       
  9792       case False
       
  9793       then show thesis
       
  9794       by (blast intro: that)
       
  9795     qed
       
  9796     have "u a + u b \<le> setsum u {a,b}"
       
  9797       using a b by simp
       
  9798     also have "... \<le> setsum u s"
       
  9799       apply (rule Groups_Big.setsum_mono2)
       
  9800       using a b u
       
  9801       apply (auto simp: less_imp_le aff_independent_finite assms)
       
  9802       done
       
  9803     finally have "u a < 1"
       
  9804       using \<open>b \<in> s\<close> u by fastforce
       
  9805   } note [simp] = this
       
  9806   show ?thesis
       
  9807     using assms
       
  9808     apply (auto simp: interior_convex_hull_explicit_minimal)
       
  9809     apply (rule_tac x=u in exI)
       
  9810     apply (auto simp: not_le)
       
  9811     done
       
  9812 qed
       
  9813 
       
  9814 lemma interior_closed_segment_ge2:
       
  9815   fixes a :: "'a::euclidean_space"
       
  9816   assumes "2 \<le> DIM('a)"
       
  9817     shows  "interior(closed_segment a b) = {}"
       
  9818 using assms unfolding segment_convex_hull
       
  9819 proof -
       
  9820   have "card {a, b} \<le> DIM('a)"
       
  9821     using assms
       
  9822     by (simp add: card_insert_if linear not_less_eq_eq numeral_2_eq_2)
       
  9823   then show "interior (convex hull {a, b}) = {}"
       
  9824     by (metis empty_interior_convex_hull finite.insertI finite.emptyI)
       
  9825 qed
       
  9826 
       
  9827 lemma interior_open_segment:
       
  9828   fixes a :: "'a::euclidean_space"
       
  9829   shows  "interior(open_segment a b) =
       
  9830                  (if 2 \<le> DIM('a) then {} else open_segment a b)"
       
  9831 proof (simp add: not_le, intro conjI impI)
       
  9832   assume "2 \<le> DIM('a)"
       
  9833   then show "interior (open_segment a b) = {}"
       
  9834     apply (simp add: segment_convex_hull open_segment_def)
       
  9835     apply (metis Diff_subset interior_mono segment_convex_hull subset_empty interior_closed_segment_ge2)
       
  9836     done
       
  9837 next
       
  9838   assume le2: "DIM('a) < 2"
       
  9839   show "interior (open_segment a b) = open_segment a b"
       
  9840   proof (cases "a = b")
       
  9841     case True then show ?thesis by auto
       
  9842   next
       
  9843     case False
       
  9844     with le2 have "affine hull (open_segment a b) = UNIV"
       
  9845       apply simp
       
  9846       apply (rule affine_independent_span_gt)
       
  9847       apply (simp_all add: affine_dependent_def insert_Diff_if)
       
  9848       done
       
  9849     then show "interior (open_segment a b) = open_segment a b"
       
  9850       using rel_interior_interior rel_interior_open_segment by blast
       
  9851   qed
       
  9852 qed
       
  9853 
       
  9854 lemma interior_closed_segment:
       
  9855   fixes a :: "'a::euclidean_space"
       
  9856   shows "interior(closed_segment a b) =
       
  9857                  (if 2 \<le> DIM('a) then {} else open_segment a b)"
       
  9858 proof (cases "a = b")
       
  9859   case True then show ?thesis by simp
       
  9860 next
       
  9861   case False
       
  9862   then have "closure (open_segment a b) = closed_segment a b"
       
  9863     by simp
       
  9864   then show ?thesis
       
  9865     by (metis (no_types) convex_interior_closure convex_open_segment interior_open_segment)
       
  9866 qed
       
  9867 
       
  9868 lemmas interior_segment = interior_closed_segment interior_open_segment
       
  9869 
       
  9870 lemma closed_segment_eq [simp]:
       
  9871   fixes a :: "'a::euclidean_space"
       
  9872   shows "closed_segment a b = closed_segment c d \<longleftrightarrow> {a,b} = {c,d}"
       
  9873 proof
       
  9874   assume abcd: "closed_segment a b = closed_segment c d"
       
  9875   show "{a,b} = {c,d}"
       
  9876   proof (cases "a=b \<or> c=d")
       
  9877     case True with abcd show ?thesis by force
       
  9878   next
       
  9879     case False
       
  9880     then have neq: "a \<noteq> b \<and> c \<noteq> d" by force
       
  9881     have *: "closed_segment c d - {a, b} = rel_interior (closed_segment c d)"
       
  9882       using neq abcd by (metis (no_types) open_segment_def rel_interior_closed_segment)
       
  9883     have "b \<in> {c, d}"
       
  9884     proof -
       
  9885       have "insert b (closed_segment c d) = closed_segment c d"
       
  9886         using abcd by blast
       
  9887       then show ?thesis
       
  9888         by (metis DiffD2 Diff_insert2 False * insertI1 insert_Diff_if open_segment_def rel_interior_closed_segment)
       
  9889     qed
       
  9890     moreover have "a \<in> {c, d}"
       
  9891       by (metis Diff_iff False * abcd ends_in_segment(1) insertI1 open_segment_def rel_interior_closed_segment)
       
  9892     ultimately show "{a, b} = {c, d}"
       
  9893       using neq by fastforce
       
  9894   qed
       
  9895 next
       
  9896   assume "{a,b} = {c,d}"
       
  9897   then show "closed_segment a b = closed_segment c d"
       
  9898     by (simp add: segment_convex_hull)
       
  9899 qed
       
  9900 
       
  9901 lemma closed_open_segment_eq [simp]:
       
  9902   fixes a :: "'a::euclidean_space"
       
  9903   shows "closed_segment a b \<noteq> open_segment c d"
       
  9904 by (metis DiffE closed_segment_neq_empty closure_closed_segment closure_open_segment ends_in_segment(1) insertI1 open_segment_def)
       
  9905 
       
  9906 lemma open_closed_segment_eq [simp]:
       
  9907   fixes a :: "'a::euclidean_space"
       
  9908   shows "open_segment a b \<noteq> closed_segment c d"
       
  9909 using closed_open_segment_eq by blast
       
  9910 
       
  9911 lemma open_segment_eq [simp]:
       
  9912   fixes a :: "'a::euclidean_space"
       
  9913   shows "open_segment a b = open_segment c d \<longleftrightarrow> a = b \<and> c = d \<or> {a,b} = {c,d}"
       
  9914         (is "?lhs = ?rhs")
       
  9915 proof
       
  9916   assume abcd: ?lhs
       
  9917   show ?rhs
       
  9918   proof (cases "a=b \<or> c=d")
       
  9919     case True with abcd show ?thesis
       
  9920       using finite_open_segment by fastforce
       
  9921   next
       
  9922     case False
       
  9923     then have a2: "a \<noteq> b \<and> c \<noteq> d" by force
       
  9924     with abcd show ?rhs
       
  9925       unfolding open_segment_def
       
  9926       by (metis (no_types) abcd closed_segment_eq closure_open_segment)
       
  9927   qed
       
  9928 next
       
  9929   assume ?rhs
       
  9930   then show ?lhs
       
  9931     by (metis Diff_cancel convex_hull_singleton insert_absorb2 open_segment_def segment_convex_hull)
       
  9932 qed
       
  9933 
       
  9934 subsection\<open>Similar results for closure and (relative or absolute) frontier.\<close>
       
  9935 
       
  9936 lemma closure_convex_hull [simp]:
       
  9937   fixes s :: "'a::euclidean_space set"
       
  9938   shows "compact s ==> closure(convex hull s) = convex hull s"
       
  9939   by (simp add: compact_imp_closed compact_convex_hull)
       
  9940 
       
  9941 lemma rel_frontier_convex_hull_explicit:
       
  9942   fixes s :: "'a::euclidean_space set"
       
  9943   assumes "~ affine_dependent s"
       
  9944   shows "rel_frontier(convex hull s) =
       
  9945          {y. \<exists>u. (\<forall>x \<in> s. 0 \<le> u x) \<and> (\<exists>x \<in> s. u x = 0) \<and> setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}"
       
  9946 proof -
       
  9947   have fs: "finite s"
       
  9948     using assms by (simp add: aff_independent_finite)
       
  9949   show ?thesis
       
  9950     apply (simp add: rel_frontier_def finite_imp_compact rel_interior_convex_hull_explicit assms fs)
       
  9951     apply (auto simp: convex_hull_finite fs)
       
  9952     apply (drule_tac x=u in spec)
       
  9953     apply (rule_tac x=u in exI)
       
  9954     apply force
       
  9955     apply (rename_tac v)
       
  9956     apply (rule notE [OF assms])
       
  9957     apply (simp add: affine_dependent_explicit)
       
  9958     apply (rule_tac x=s in exI)
       
  9959     apply (auto simp: fs)
       
  9960     apply (rule_tac x = "\<lambda>x. u x - v x" in exI)
       
  9961     apply (force simp: setsum_subtractf scaleR_diff_left)
       
  9962     done
       
  9963 qed
       
  9964 
       
  9965 lemma frontier_convex_hull_explicit:
       
  9966   fixes s :: "'a::euclidean_space set"
       
  9967   assumes "~ affine_dependent s"
       
  9968   shows "frontier(convex hull s) =
       
  9969          {y. \<exists>u. (\<forall>x \<in> s. 0 \<le> u x) \<and> (DIM ('a) < card s \<longrightarrow> (\<exists>x \<in> s. u x = 0)) \<and>
       
  9970              setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}"
       
  9971 proof -
       
  9972   have fs: "finite s"
       
  9973     using assms by (simp add: aff_independent_finite)
       
  9974   show ?thesis
       
  9975   proof (cases "DIM ('a) < card s")
       
  9976     case True
       
  9977     with assms fs show ?thesis
       
  9978       by (simp add: rel_frontier_def frontier_def rel_frontier_convex_hull_explicit [symmetric]
       
  9979                     interior_convex_hull_explicit_minimal rel_interior_convex_hull_explicit)
       
  9980   next
       
  9981     case False
       
  9982     then have "card s \<le> DIM ('a)"
       
  9983       by linarith
       
  9984     then show ?thesis
       
  9985       using assms fs
       
  9986       apply (simp add: frontier_def interior_convex_hull_explicit finite_imp_compact)
       
  9987       apply (simp add: convex_hull_finite)
       
  9988       done
       
  9989   qed
       
  9990 qed
       
  9991 
       
  9992 lemma rel_frontier_convex_hull_cases:
       
  9993   fixes s :: "'a::euclidean_space set"
       
  9994   assumes "~ affine_dependent s"
       
  9995   shows "rel_frontier(convex hull s) = \<Union>{convex hull (s - {x}) |x. x \<in> s}"
       
  9996 proof -
       
  9997   have fs: "finite s"
       
  9998     using assms by (simp add: aff_independent_finite)
       
  9999   { fix u a
       
 10000   have "\<forall>x\<in>s. 0 \<le> u x \<Longrightarrow> a \<in> s \<Longrightarrow> u a = 0 \<Longrightarrow> setsum u s = 1 \<Longrightarrow>
       
 10001             \<exists>x v. x \<in> s \<and>
       
 10002                   (\<forall>x\<in>s - {x}. 0 \<le> v x) \<and>
       
 10003                       setsum v (s - {x}) = 1 \<and> (\<Sum>x\<in>s - {x}. v x *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x)"
       
 10004     apply (rule_tac x=a in exI)
       
 10005     apply (rule_tac x=u in exI)
       
 10006     apply (simp add: Groups_Big.setsum_diff1 fs)
       
 10007     done }
       
 10008   moreover
       
 10009   { fix a u
       
 10010     have "a \<in> s \<Longrightarrow> \<forall>x\<in>s - {a}. 0 \<le> u x \<Longrightarrow> setsum u (s - {a}) = 1 \<Longrightarrow>
       
 10011             \<exists>v. (\<forall>x\<in>s. 0 \<le> v x) \<and>
       
 10012                  (\<exists>x\<in>s. v x = 0) \<and> setsum v s = 1 \<and> (\<Sum>x\<in>s. v x *\<^sub>R x) = (\<Sum>x\<in>s - {a}. u x *\<^sub>R x)"
       
 10013     apply (rule_tac x="\<lambda>x. if x = a then 0 else u x" in exI)
       
 10014     apply (auto simp: setsum.If_cases Diff_eq if_smult fs)
       
 10015     done }
       
 10016   ultimately show ?thesis
       
 10017     using assms
       
 10018     apply (simp add: rel_frontier_convex_hull_explicit)
       
 10019     apply (simp add: convex_hull_finite fs Union_SetCompr_eq, auto)
       
 10020     done
       
 10021 qed
       
 10022 
       
 10023 lemma frontier_convex_hull_eq_rel_frontier:
       
 10024   fixes s :: "'a::euclidean_space set"
       
 10025   assumes "~ affine_dependent s"
       
 10026   shows "frontier(convex hull s) =
       
 10027            (if card s \<le> DIM ('a) then convex hull s else rel_frontier(convex hull s))"
       
 10028   using assms
       
 10029   unfolding rel_frontier_def frontier_def
       
 10030   by (simp add: affine_independent_span_gt rel_interior_interior
       
 10031                 finite_imp_compact empty_interior_convex_hull aff_independent_finite)
       
 10032 
       
 10033 lemma frontier_convex_hull_cases:
       
 10034   fixes s :: "'a::euclidean_space set"
       
 10035   assumes "~ affine_dependent s"
       
 10036   shows "frontier(convex hull s) =
       
 10037            (if card s \<le> DIM ('a) then convex hull s else \<Union>{convex hull (s - {x}) |x. x \<in> s})"
       
 10038 by (simp add: assms frontier_convex_hull_eq_rel_frontier rel_frontier_convex_hull_cases)
       
 10039 
       
 10040 lemma in_frontier_convex_hull:
       
 10041   fixes s :: "'a::euclidean_space set"
       
 10042   assumes "finite s" "card s \<le> Suc (DIM ('a))" "x \<in> s"
       
 10043   shows   "x \<in> frontier(convex hull s)"
       
 10044 proof (cases "affine_dependent s")
       
 10045   case True
       
 10046   with assms show ?thesis
       
 10047     apply (auto simp: affine_dependent_def frontier_def finite_imp_compact hull_inc)
       
 10048     by (metis card.insert_remove convex_hull_subset_affine_hull empty_interior_affine_hull finite_Diff hull_redundant insert_Diff insert_Diff_single insert_not_empty interior_mono not_less_eq_eq subset_empty)
       
 10049 next
       
 10050   case False
       
 10051   { assume "card s = Suc (card Basis)"
       
 10052     then have cs: "Suc 0 < card s"
       
 10053       by (simp add: DIM_positive)
       
 10054     with subset_singletonD have "\<exists>y \<in> s. y \<noteq> x"
       
 10055       by (cases "s \<le> {x}") fastforce+
       
 10056   } note [dest!] = this
       
 10057   show ?thesis using assms
       
 10058     unfolding frontier_convex_hull_cases [OF False] Union_SetCompr_eq
       
 10059     by (auto simp: le_Suc_eq hull_inc)
       
 10060 qed
       
 10061 
       
 10062 lemma not_in_interior_convex_hull:
       
 10063   fixes s :: "'a::euclidean_space set"
       
 10064   assumes "finite s" "card s \<le> Suc (DIM ('a))" "x \<in> s"
       
 10065   shows   "x \<notin> interior(convex hull s)"
       
 10066 using in_frontier_convex_hull [OF assms]
       
 10067 by (metis Diff_iff frontier_def)
       
 10068 
       
 10069 lemma interior_convex_hull_eq_empty:
       
 10070   fixes s :: "'a::euclidean_space set"
       
 10071   assumes "card s = Suc (DIM ('a))"
       
 10072   shows   "interior(convex hull s) = {} \<longleftrightarrow> affine_dependent s"
       
 10073 proof -
       
 10074   { fix a b
       
 10075     assume ab: "a \<in> interior (convex hull s)" "b \<in> s" "b \<in> affine hull (s - {b})"
       
 10076     then have "interior(affine hull s) = {}" using assms
       
 10077       by (metis DIM_positive One_nat_def Suc_mono card.remove card_infinite empty_interior_affine_hull eq_iff hull_redundant insert_Diff not_less zero_le_one)
       
 10078     then have False using ab
       
 10079       by (metis convex_hull_subset_affine_hull equals0D interior_mono subset_eq)
       
 10080   } then
       
 10081   show ?thesis
       
 10082     using assms
       
 10083     apply auto
       
 10084     apply (metis UNIV_I affine_hull_convex_hull affine_hull_empty affine_independent_span_eq convex_convex_hull empty_iff rel_interior_interior rel_interior_same_affine_hull)
       
 10085     apply (auto simp: affine_dependent_def)
       
 10086     done
       
 10087 qed
       
 10088 
       
 10089 
       
 10090 subsection \<open>Coplanarity, and collinearity in terms of affine hull\<close>
       
 10091 
       
 10092 definition coplanar  where
       
 10093    "coplanar s \<equiv> \<exists>u v w. s \<subseteq> affine hull {u,v,w}"
       
 10094 
       
 10095 lemma collinear_affine_hull:
       
 10096   "collinear s \<longleftrightarrow> (\<exists>u v. s \<subseteq> affine hull {u,v})"
       
 10097 proof (cases "s={}")
       
 10098   case True then show ?thesis
       
 10099     by simp
       
 10100 next
       
 10101   case False
       
 10102   then obtain x where x: "x \<in> s" by auto
       
 10103   { fix u
       
 10104     assume *: "\<And>x y. \<lbrakk>x\<in>s; y\<in>s\<rbrakk> \<Longrightarrow> \<exists>c. x - y = c *\<^sub>R u"
       
 10105     have "\<exists>u v. s \<subseteq> {a *\<^sub>R u + b *\<^sub>R v |a b. a + b = 1}"
       
 10106       apply (rule_tac x=x in exI)
       
 10107       apply (rule_tac x="x+u" in exI, clarify)
       
 10108       apply (erule exE [OF * [OF x]])
       
 10109       apply (rename_tac c)
       
 10110       apply (rule_tac x="1+c" in exI)
       
 10111       apply (rule_tac x="-c" in exI)
       
 10112       apply (simp add: algebra_simps)
       
 10113       done
       
 10114   } moreover
       
 10115   { fix u v x y
       
 10116     assume *: "s \<subseteq> {a *\<^sub>R u + b *\<^sub>R v |a b. a + b = 1}"
       
 10117     have "x\<in>s \<Longrightarrow> y\<in>s \<Longrightarrow> \<exists>c. x - y = c *\<^sub>R (v-u)"
       
 10118       apply (drule subsetD [OF *])+
       
 10119       apply simp
       
 10120       apply clarify
       
 10121       apply (rename_tac r1 r2)
       
 10122       apply (rule_tac x="r1-r2" in exI)
       
 10123       apply (simp add: algebra_simps)
       
 10124       apply (metis scaleR_left.add)
       
 10125       done
       
 10126   } ultimately
       
 10127   show ?thesis
       
 10128   unfolding collinear_def affine_hull_2
       
 10129     by blast
       
 10130 qed
       
 10131 
       
 10132 lemma collinear_closed_segment [simp]: "collinear (closed_segment a b)"
       
 10133 by (metis affine_hull_convex_hull collinear_affine_hull hull_subset segment_convex_hull)
       
 10134 
       
 10135 lemma collinear_open_segment [simp]: "collinear (open_segment a b)"
       
 10136   unfolding open_segment_def
       
 10137   by (metis convex_hull_subset_affine_hull segment_convex_hull dual_order.trans
       
 10138     convex_hull_subset_affine_hull Diff_subset collinear_affine_hull)
       
 10139 
       
 10140 lemma subset_continuous_image_segment_1:
       
 10141   fixes f :: "'a::euclidean_space \<Rightarrow> real"
       
 10142   assumes "continuous_on (closed_segment a b) f"
       
 10143   shows "closed_segment (f a) (f b) \<subseteq> image f (closed_segment a b)"
       
 10144 by (metis connected_segment convex_contains_segment ends_in_segment imageI
       
 10145            is_interval_connected_1 is_interval_convex connected_continuous_image [OF assms])
       
 10146 
       
 10147 lemma collinear_imp_coplanar:
       
 10148   "collinear s ==> coplanar s"
       
 10149 by (metis collinear_affine_hull coplanar_def insert_absorb2)
       
 10150 
       
 10151 lemma collinear_small:
       
 10152   assumes "finite s" "card s \<le> 2"
       
 10153     shows "collinear s"
       
 10154 proof -
       
 10155   have "card s = 0 \<or> card s = 1 \<or> card s = 2"
       
 10156     using assms by linarith
       
 10157   then show ?thesis using assms
       
 10158     using card_eq_SucD
       
 10159     by auto (metis collinear_2 numeral_2_eq_2)
       
 10160 qed
       
 10161 
       
 10162 lemma coplanar_small:
       
 10163   assumes "finite s" "card s \<le> 3"
       
 10164     shows "coplanar s"
       
 10165 proof -
       
 10166   have "card s \<le> 2 \<or> card s = Suc (Suc (Suc 0))"
       
 10167     using assms by linarith
       
 10168   then show ?thesis using assms
       
 10169     apply safe
       
 10170     apply (simp add: collinear_small collinear_imp_coplanar)
       
 10171     apply (safe dest!: card_eq_SucD)
       
 10172     apply (auto simp: coplanar_def)
       
 10173     apply (metis hull_subset insert_subset)
       
 10174     done
       
 10175 qed
       
 10176 
       
 10177 lemma coplanar_empty: "coplanar {}"
       
 10178   by (simp add: coplanar_small)
       
 10179 
       
 10180 lemma coplanar_sing: "coplanar {a}"
       
 10181   by (simp add: coplanar_small)
       
 10182 
       
 10183 lemma coplanar_2: "coplanar {a,b}"
       
 10184   by (auto simp: card_insert_if coplanar_small)
       
 10185 
       
 10186 lemma coplanar_3: "coplanar {a,b,c}"
       
 10187   by (auto simp: card_insert_if coplanar_small)
       
 10188 
       
 10189 lemma collinear_affine_hull_collinear: "collinear(affine hull s) \<longleftrightarrow> collinear s"
       
 10190   unfolding collinear_affine_hull
       
 10191   by (metis affine_affine_hull subset_hull hull_hull hull_mono)
       
 10192 
       
 10193 lemma coplanar_affine_hull_coplanar: "coplanar(affine hull s) \<longleftrightarrow> coplanar s"
       
 10194   unfolding coplanar_def
       
 10195   by (metis affine_affine_hull subset_hull hull_hull hull_mono)
       
 10196 
       
 10197 lemma coplanar_linear_image:
       
 10198   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
       
 10199   assumes "coplanar s" "linear f" shows "coplanar(f ` s)"
       
 10200 proof -
       
 10201   { fix u v w
       
 10202     assume "s \<subseteq> affine hull {u, v, w}"
       
 10203     then have "f ` s \<subseteq> f ` (affine hull {u, v, w})"
       
 10204       by (simp add: image_mono)
       
 10205     then have "f ` s \<subseteq> affine hull (f ` {u, v, w})"
       
 10206       by (metis assms(2) linear_conv_bounded_linear affine_hull_linear_image)
       
 10207   } then
       
 10208   show ?thesis
       
 10209     by auto (meson assms(1) coplanar_def)
       
 10210 qed
       
 10211 
       
 10212 lemma coplanar_translation_imp: "coplanar s \<Longrightarrow> coplanar ((\<lambda>x. a + x) ` s)"
       
 10213   unfolding coplanar_def
       
 10214   apply clarify
       
 10215   apply (rule_tac x="u+a" in exI)
       
 10216   apply (rule_tac x="v+a" in exI)
       
 10217   apply (rule_tac x="w+a" in exI)
       
 10218   using affine_hull_translation [of a "{u,v,w}" for u v w]
       
 10219   apply (force simp: add.commute)
       
 10220   done
       
 10221 
       
 10222 lemma coplanar_translation_eq: "coplanar((\<lambda>x. a + x) ` s) \<longleftrightarrow> coplanar s"
       
 10223     by (metis (no_types) coplanar_translation_imp translation_galois)
       
 10224 
       
 10225 lemma coplanar_linear_image_eq:
       
 10226   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
 10227   assumes "linear f" "inj f" shows "coplanar(f ` s) = coplanar s"
       
 10228 proof
       
 10229   assume "coplanar s"
       
 10230   then show "coplanar (f ` s)"
       
 10231     unfolding coplanar_def
       
 10232     using affine_hull_linear_image [of f "{u,v,w}" for u v w]  assms
       
 10233     by (meson coplanar_def coplanar_linear_image)
       
 10234 next
       
 10235   obtain g where g: "linear g" "g \<circ> f = id"
       
 10236     using linear_injective_left_inverse [OF assms]
       
 10237     by blast
       
 10238   assume "coplanar (f ` s)"
       
 10239   then obtain u v w where "f ` s \<subseteq> affine hull {u, v, w}"
       
 10240     by (auto simp: coplanar_def)
       
 10241   then have "g ` f ` s \<subseteq> g ` (affine hull {u, v, w})"
       
 10242     by blast
       
 10243   then have "s \<subseteq> g ` (affine hull {u, v, w})"
       
 10244     using g by (simp add: Fun.image_comp)
       
 10245   then show "coplanar s"
       
 10246     unfolding coplanar_def
       
 10247     using affine_hull_linear_image [of g "{u,v,w}" for u v w]  \<open>linear g\<close> linear_conv_bounded_linear
       
 10248     by fastforce
       
 10249 qed
       
 10250 (*The HOL Light proof is simply
       
 10251     MATCH_ACCEPT_TAC(LINEAR_INVARIANT_RULE COPLANAR_LINEAR_IMAGE));;
       
 10252 *)
       
 10253 
       
 10254 lemma coplanar_subset: "\<lbrakk>coplanar t; s \<subseteq> t\<rbrakk> \<Longrightarrow> coplanar s"
       
 10255   by (meson coplanar_def order_trans)
       
 10256 
       
 10257 lemma affine_hull_3_imp_collinear: "c \<in> affine hull {a,b} \<Longrightarrow> collinear {a,b,c}"
       
 10258   by (metis collinear_2 collinear_affine_hull_collinear hull_redundant insert_commute)
       
 10259 
       
 10260 lemma collinear_3_imp_in_affine_hull: "\<lbrakk>collinear {a,b,c}; a \<noteq> b\<rbrakk> \<Longrightarrow> c \<in> affine hull {a,b}"
       
 10261   unfolding collinear_def
       
 10262   apply clarify
       
 10263   apply (frule_tac x=b in bspec, blast, drule_tac x=a in bspec, blast, erule exE)
       
 10264   apply (drule_tac x=c in bspec, blast, drule_tac x=a in bspec, blast, erule exE)
       
 10265   apply (rename_tac y x)
       
 10266   apply (simp add: affine_hull_2)
       
 10267   apply (rule_tac x="1 - x/y" in exI)
       
 10268   apply (simp add: algebra_simps)
       
 10269   done
       
 10270 
       
 10271 lemma collinear_3_affine_hull:
       
 10272   assumes "a \<noteq> b"
       
 10273     shows "collinear {a,b,c} \<longleftrightarrow> c \<in> affine hull {a,b}"
       
 10274 using affine_hull_3_imp_collinear assms collinear_3_imp_in_affine_hull by blast
       
 10275 
       
 10276 lemma collinear_3_eq_affine_dependent:
       
 10277   "collinear{a,b,c} \<longleftrightarrow> a = b \<or> a = c \<or> b = c \<or> affine_dependent {a,b,c}"
       
 10278 apply (case_tac "a=b", simp)
       
 10279 apply (case_tac "a=c")
       
 10280 apply (simp add: insert_commute)
       
 10281 apply (case_tac "b=c")
       
 10282 apply (simp add: insert_commute)
       
 10283 apply (auto simp: affine_dependent_def collinear_3_affine_hull insert_Diff_if)
       
 10284 apply (metis collinear_3_affine_hull insert_commute)+
       
 10285 done
       
 10286 
       
 10287 lemma affine_dependent_imp_collinear_3:
       
 10288   "affine_dependent {a,b,c} \<Longrightarrow> collinear{a,b,c}"
       
 10289 by (simp add: collinear_3_eq_affine_dependent)
       
 10290 
       
 10291 lemma collinear_3: "NO_MATCH 0 x \<Longrightarrow> collinear {x,y,z} \<longleftrightarrow> collinear {0, x-y, z-y}"
       
 10292   by (auto simp add: collinear_def)
       
 10293 
       
 10294 
       
 10295 thm affine_hull_nonempty
       
 10296 corollary affine_hull_eq_empty [simp]: "affine hull S = {} \<longleftrightarrow> S = {}"
       
 10297   using affine_hull_nonempty by blast
       
 10298 
       
 10299 lemma affine_hull_2_alt:
       
 10300   fixes a b :: "'a::real_vector"
       
 10301   shows "affine hull {a,b} = range (\<lambda>u. a + u *\<^sub>R (b - a))"
       
 10302 apply (simp add: affine_hull_2, safe)
       
 10303 apply (rule_tac x=v in image_eqI)
       
 10304 apply (simp add: algebra_simps)
       
 10305 apply (metis scaleR_add_left scaleR_one, simp)
       
 10306 apply (rule_tac x="1-u" in exI)
       
 10307 apply (simp add: algebra_simps)
       
 10308 done
       
 10309 
       
 10310 lemma interior_convex_hull_3_minimal:
       
 10311   fixes a :: "'a::euclidean_space"
       
 10312   shows "\<lbrakk>~ collinear{a,b,c}; DIM('a) = 2\<rbrakk>
       
 10313          \<Longrightarrow> interior(convex hull {a,b,c}) =
       
 10314                 {v. \<exists>x y z. 0 < x \<and> 0 < y \<and> 0 < z \<and> x + y + z = 1 \<and>
       
 10315                             x *\<^sub>R a + y *\<^sub>R b + z *\<^sub>R c = v}"
       
 10316 apply (simp add: collinear_3_eq_affine_dependent interior_convex_hull_explicit_minimal, safe)
       
 10317 apply (rule_tac x="u a" in exI, simp)
       
 10318 apply (rule_tac x="u b" in exI, simp)
       
 10319 apply (rule_tac x="u c" in exI, simp)
       
 10320 apply (rename_tac uu x y z)
       
 10321 apply (rule_tac x="\<lambda>r. (if r=a then x else if r=b then y else if r=c then z else 0)" in exI)
       
 10322 apply simp
       
 10323 done
       
 10324 
       
 10325 subsection\<open>The infimum of the distance between two sets\<close>
       
 10326 
       
 10327 definition setdist :: "'a::metric_space set \<Rightarrow> 'a set \<Rightarrow> real" where
       
 10328   "setdist s t \<equiv>
       
 10329        (if s = {} \<or> t = {} then 0
       
 10330         else Inf {dist x y| x y. x \<in> s \<and> y \<in> t})"
       
 10331 
       
 10332 lemma setdist_empty1 [simp]: "setdist {} t = 0"
       
 10333   by (simp add: setdist_def)
       
 10334 
       
 10335 lemma setdist_empty2 [simp]: "setdist t {} = 0"
       
 10336   by (simp add: setdist_def)
       
 10337 
       
 10338 lemma setdist_pos_le [simp]: "0 \<le> setdist s t"
       
 10339   by (auto simp: setdist_def ex_in_conv [symmetric] intro: cInf_greatest)
       
 10340 
       
 10341 lemma le_setdistI:
       
 10342   assumes "s \<noteq> {}" "t \<noteq> {}" "\<And>x y. \<lbrakk>x \<in> s; y \<in> t\<rbrakk> \<Longrightarrow> d \<le> dist x y"
       
 10343     shows "d \<le> setdist s t"
       
 10344   using assms
       
 10345   by (auto simp: setdist_def Set.ex_in_conv [symmetric] intro: cInf_greatest)
       
 10346 
       
 10347 lemma setdist_le_dist: "\<lbrakk>x \<in> s; y \<in> t\<rbrakk> \<Longrightarrow> setdist s t \<le> dist x y"
       
 10348   unfolding setdist_def
       
 10349   by (auto intro!: bdd_belowI [where m=0] cInf_lower)
       
 10350 
       
 10351 lemma le_setdist_iff:
       
 10352         "d \<le> setdist s t \<longleftrightarrow>
       
 10353         (\<forall>x \<in> s. \<forall>y \<in> t. d \<le> dist x y) \<and> (s = {} \<or> t = {} \<longrightarrow> d \<le> 0)"
       
 10354   apply (cases "s = {} \<or> t = {}")
       
 10355   apply (force simp add: setdist_def)
       
 10356   apply (intro iffI conjI)
       
 10357   using setdist_le_dist apply fastforce
       
 10358   apply (auto simp: intro: le_setdistI)
       
 10359   done
       
 10360 
       
 10361 lemma setdist_ltE:
       
 10362   assumes "setdist s t < b" "s \<noteq> {}" "t \<noteq> {}"
       
 10363     obtains x y where "x \<in> s" "y \<in> t" "dist x y < b"
       
 10364 using assms
       
 10365 by (auto simp: not_le [symmetric] le_setdist_iff)
       
 10366 
       
 10367 lemma setdist_refl: "setdist s s = 0"
       
 10368   apply (cases "s = {}")
       
 10369   apply (force simp add: setdist_def)
       
 10370   apply (rule antisym [OF _ setdist_pos_le])
       
 10371   apply (metis all_not_in_conv dist_self setdist_le_dist)
       
 10372   done
       
 10373 
       
 10374 lemma setdist_sym: "setdist s t = setdist t s"
       
 10375   by (force simp: setdist_def dist_commute intro!: arg_cong [where f=Inf])
       
 10376 
       
 10377 lemma setdist_triangle: "setdist s t \<le> setdist s {a} + setdist {a} t"
       
 10378 proof (cases "s = {} \<or> t = {}")
       
 10379   case True then show ?thesis
       
 10380     using setdist_pos_le by fastforce
       
 10381 next
       
 10382   case False
       
 10383   have "\<And>x. x \<in> s \<Longrightarrow> setdist s t - dist x a \<le> setdist {a} t"
       
 10384     apply (rule le_setdistI, blast)
       
 10385     using False apply (fastforce intro: le_setdistI)
       
 10386     apply (simp add: algebra_simps)
       
 10387     apply (metis dist_commute dist_triangle3 order_trans [OF setdist_le_dist])
       
 10388     done
       
 10389   then have "setdist s t - setdist {a} t \<le> setdist s {a}"
       
 10390     using False by (fastforce intro: le_setdistI)
       
 10391   then show ?thesis
       
 10392     by (simp add: algebra_simps)
       
 10393 qed
       
 10394 
       
 10395 lemma setdist_singletons [simp]: "setdist {x} {y} = dist x y"
       
 10396   by (simp add: setdist_def)
       
 10397 
       
 10398 lemma setdist_Lipschitz: "\<bar>setdist {x} s - setdist {y} s\<bar> \<le> dist x y"
       
 10399   apply (subst setdist_singletons [symmetric])
       
 10400   by (metis abs_diff_le_iff diff_le_eq setdist_triangle setdist_sym)
       
 10401 
       
 10402 lemma continuous_at_setdist [continuous_intros]: "continuous (at x) (\<lambda>y. (setdist {y} s))"
       
 10403   by (force simp: continuous_at_eps_delta dist_real_def intro: le_less_trans [OF setdist_Lipschitz])
       
 10404 
       
 10405 lemma continuous_on_setdist [continuous_intros]: "continuous_on t (\<lambda>y. (setdist {y} s))"
       
 10406   by (metis continuous_at_setdist continuous_at_imp_continuous_on)
       
 10407 
       
 10408 lemma uniformly_continuous_on_setdist: "uniformly_continuous_on t (\<lambda>y. (setdist {y} s))"
       
 10409   by (force simp: uniformly_continuous_on_def dist_real_def intro: le_less_trans [OF setdist_Lipschitz])
       
 10410 
       
 10411 lemma setdist_subset_right: "\<lbrakk>t \<noteq> {}; t \<subseteq> u\<rbrakk> \<Longrightarrow> setdist s u \<le> setdist s t"
       
 10412   apply (cases "s = {} \<or> u = {}", force)
       
 10413   apply (auto simp: setdist_def intro!: bdd_belowI [where m=0] cInf_superset_mono)
       
 10414   done
       
 10415 
       
 10416 lemma setdist_subset_left: "\<lbrakk>s \<noteq> {}; s \<subseteq> t\<rbrakk> \<Longrightarrow> setdist t u \<le> setdist s u"
       
 10417   by (metis setdist_subset_right setdist_sym)
       
 10418 
       
 10419 lemma setdist_closure_1 [simp]: "setdist (closure s) t = setdist s t"
       
 10420 proof (cases "s = {} \<or> t = {}")
       
 10421   case True then show ?thesis by force
       
 10422 next
       
 10423   case False
       
 10424   { fix y
       
 10425     assume "y \<in> t"
       
 10426     have "continuous_on (closure s) (\<lambda>a. dist a y)"
       
 10427       by (auto simp: continuous_intros dist_norm)
       
 10428     then have *: "\<And>x. x \<in> closure s \<Longrightarrow> setdist s t \<le> dist x y"
       
 10429       apply (rule continuous_ge_on_closure)
       
 10430       apply assumption
       
 10431       apply (blast intro: setdist_le_dist \<open>y \<in> t\<close> )
       
 10432       done
       
 10433   } note * = this
       
 10434   show ?thesis
       
 10435     apply (rule antisym)
       
 10436      using False closure_subset apply (blast intro: setdist_subset_left)
       
 10437     using False *
       
 10438     apply (force simp add: closure_eq_empty intro!: le_setdistI)
       
 10439     done
       
 10440 qed
       
 10441 
       
 10442 lemma setdist_closure_2 [simp]: "setdist t (closure s) = setdist t s"
       
 10443 by (metis setdist_closure_1 setdist_sym)
       
 10444 
       
 10445 lemma setdist_compact_closed:
       
 10446   fixes s :: "'a::euclidean_space set"
       
 10447   assumes s: "compact s" and t: "closed t"
       
 10448       and "s \<noteq> {}" "t \<noteq> {}"
       
 10449     shows "\<exists>x \<in> s. \<exists>y \<in> t. dist x y = setdist s t"
       
 10450 proof -
       
 10451   have "{x - y |x y. x \<in> s \<and> y \<in> t} \<noteq> {}"
       
 10452     using assms by blast
       
 10453   then
       
 10454   have "\<exists>x \<in> s. \<exists>y \<in> t. dist x y \<le> setdist s t"
       
 10455     apply (rule distance_attains_inf [where a=0, OF compact_closed_differences [OF s t]])
       
 10456     apply (simp add: dist_norm le_setdist_iff)
       
 10457     apply blast
       
 10458     done
       
 10459   then show ?thesis
       
 10460     by (blast intro!: antisym [OF _ setdist_le_dist] )
       
 10461 qed
       
 10462 
       
 10463 lemma setdist_closed_compact:
       
 10464   fixes s :: "'a::euclidean_space set"
       
 10465   assumes s: "closed s" and t: "compact t"
       
 10466       and "s \<noteq> {}" "t \<noteq> {}"
       
 10467     shows "\<exists>x \<in> s. \<exists>y \<in> t. dist x y = setdist s t"
       
 10468   using setdist_compact_closed [OF t s \<open>t \<noteq> {}\<close> \<open>s \<noteq> {}\<close>]
       
 10469   by (metis dist_commute setdist_sym)
       
 10470 
       
 10471 lemma setdist_eq_0I: "\<lbrakk>x \<in> s; x \<in> t\<rbrakk> \<Longrightarrow> setdist s t = 0"
       
 10472   by (metis antisym dist_self setdist_le_dist setdist_pos_le)
       
 10473 
       
 10474 lemma setdist_eq_0_compact_closed:
       
 10475   fixes s :: "'a::euclidean_space set"
       
 10476   assumes s: "compact s" and t: "closed t"
       
 10477     shows "setdist s t = 0 \<longleftrightarrow> s = {} \<or> t = {} \<or> s \<inter> t \<noteq> {}"
       
 10478   apply (cases "s = {} \<or> t = {}", force)
       
 10479   using setdist_compact_closed [OF s t]
       
 10480   apply (force intro: setdist_eq_0I )
       
 10481   done
       
 10482 
       
 10483 corollary setdist_gt_0_compact_closed:
       
 10484   fixes s :: "'a::euclidean_space set"
       
 10485   assumes s: "compact s" and t: "closed t"
       
 10486     shows "setdist s t > 0 \<longleftrightarrow> (s \<noteq> {} \<and> t \<noteq> {} \<and> s \<inter> t = {})"
       
 10487   using setdist_pos_le [of s t] setdist_eq_0_compact_closed [OF assms]
       
 10488   by linarith
       
 10489 
       
 10490 lemma setdist_eq_0_closed_compact:
       
 10491   fixes s :: "'a::euclidean_space set"
       
 10492   assumes s: "closed s" and t: "compact t"
       
 10493     shows "setdist s t = 0 \<longleftrightarrow> s = {} \<or> t = {} \<or> s \<inter> t \<noteq> {}"
       
 10494   using setdist_eq_0_compact_closed [OF t s]
       
 10495   by (metis Int_commute setdist_sym)
       
 10496 
       
 10497 lemma setdist_eq_0_bounded:
       
 10498   fixes s :: "'a::euclidean_space set"
       
 10499   assumes "bounded s \<or> bounded t"
       
 10500     shows "setdist s t = 0 \<longleftrightarrow> s = {} \<or> t = {} \<or> closure s \<inter> closure t \<noteq> {}"
       
 10501   apply (cases "s = {} \<or> t = {}", force)
       
 10502   using setdist_eq_0_compact_closed [of "closure s" "closure t"]
       
 10503         setdist_eq_0_closed_compact [of "closure s" "closure t"] assms
       
 10504   apply (force simp add:  bounded_closure compact_eq_bounded_closed)
       
 10505   done
       
 10506 
       
 10507 lemma setdist_unique:
       
 10508   "\<lbrakk>a \<in> s; b \<in> t; \<And>x y. x \<in> s \<and> y \<in> t ==> dist a b \<le> dist x y\<rbrakk>
       
 10509    \<Longrightarrow> setdist s t = dist a b"
       
 10510   by (force simp add: setdist_le_dist le_setdist_iff intro: antisym)
       
 10511 
       
 10512 lemma setdist_closest_point:
       
 10513     "\<lbrakk>closed s; s \<noteq> {}\<rbrakk> \<Longrightarrow> setdist {a} s = dist a (closest_point s a)"
       
 10514   apply (rule setdist_unique)
       
 10515   using closest_point_le
       
 10516   apply (auto simp: closest_point_in_set)
       
 10517   done
       
 10518 
       
 10519 lemma setdist_eq_0_sing_1:
       
 10520     fixes s :: "'a::euclidean_space set"
       
 10521     shows "setdist {x} s = 0 \<longleftrightarrow> s = {} \<or> x \<in> closure s"
       
 10522   by (auto simp: setdist_eq_0_bounded)
       
 10523 
       
 10524 lemma setdist_eq_0_sing_2:
       
 10525     fixes s :: "'a::euclidean_space set"
       
 10526     shows "setdist s {x} = 0 \<longleftrightarrow> s = {} \<or> x \<in> closure s"
       
 10527   by (auto simp: setdist_eq_0_bounded)
       
 10528 
       
 10529 lemma setdist_neq_0_sing_1:
       
 10530     fixes s :: "'a::euclidean_space set"
       
 10531     shows "\<lbrakk>setdist {x} s = a; a \<noteq> 0\<rbrakk> \<Longrightarrow> s \<noteq> {} \<and> x \<notin> closure s"
       
 10532   by (auto simp: setdist_eq_0_sing_1)
       
 10533 
       
 10534 lemma setdist_neq_0_sing_2:
       
 10535     fixes s :: "'a::euclidean_space set"
       
 10536     shows "\<lbrakk>setdist s {x} = a; a \<noteq> 0\<rbrakk> \<Longrightarrow> s \<noteq> {} \<and> x \<notin> closure s"
       
 10537   by (auto simp: setdist_eq_0_sing_2)
       
 10538 
       
 10539 lemma setdist_sing_in_set:
       
 10540     fixes s :: "'a::euclidean_space set"
       
 10541     shows "x \<in> s \<Longrightarrow> setdist {x} s = 0"
       
 10542   using closure_subset by (auto simp: setdist_eq_0_sing_1)
       
 10543 
       
 10544 lemma setdist_le_sing: "x \<in> s ==> setdist s t \<le> setdist {x} t"
       
 10545   using setdist_subset_left by auto
       
 10546 
       
 10547 lemma setdist_eq_0_closed:
       
 10548   fixes s :: "'a::euclidean_space set"
       
 10549   shows  "closed s \<Longrightarrow> (setdist {x} s = 0 \<longleftrightarrow> s = {} \<or> x \<in> s)"
       
 10550 by (simp add: setdist_eq_0_sing_1)
       
 10551 
       
 10552 lemma setdist_eq_0_closedin:
       
 10553   fixes s :: "'a::euclidean_space set"
       
 10554   shows "\<lbrakk>closedin (subtopology euclidean u) s; x \<in> u\<rbrakk>
       
 10555          \<Longrightarrow> (setdist {x} s = 0 \<longleftrightarrow> s = {} \<or> x \<in> s)"
       
 10556   by (auto simp: closedin_limpt setdist_eq_0_sing_1 closure_def)
       
 10557 
       
 10558 subsection\<open>Basic lemmas about hyperplanes and halfspaces\<close>
       
 10559 
       
 10560 lemma hyperplane_eq_Ex:
       
 10561   assumes "a \<noteq> 0" obtains x where "a \<bullet> x = b"
       
 10562   by (rule_tac x = "(b / (a \<bullet> a)) *\<^sub>R a" in that) (simp add: assms)
       
 10563 
       
 10564 lemma hyperplane_eq_empty:
       
 10565      "{x. a \<bullet> x = b} = {} \<longleftrightarrow> a = 0 \<and> b \<noteq> 0"
       
 10566   using hyperplane_eq_Ex apply auto[1]
       
 10567   using inner_zero_right by blast
       
 10568 
       
 10569 lemma hyperplane_eq_UNIV:
       
 10570    "{x. a \<bullet> x = b} = UNIV \<longleftrightarrow> a = 0 \<and> b = 0"
       
 10571 proof -
       
 10572   have "UNIV \<subseteq> {x. a \<bullet> x = b} \<Longrightarrow> a = 0 \<and> b = 0"
       
 10573     apply (drule_tac c = "((b+1) / (a \<bullet> a)) *\<^sub>R a" in subsetD)
       
 10574     apply simp_all
       
 10575     by (metis add_cancel_right_right divide_1 zero_neq_one)
       
 10576   then show ?thesis by force
       
 10577 qed
       
 10578 
       
 10579 lemma halfspace_eq_empty_lt:
       
 10580    "{x. a \<bullet> x < b} = {} \<longleftrightarrow> a = 0 \<and> b \<le> 0"
       
 10581 proof -
       
 10582   have "{x. a \<bullet> x < b} \<subseteq> {} \<Longrightarrow> a = 0 \<and> b \<le> 0"
       
 10583     apply (rule ccontr)
       
 10584     apply (drule_tac c = "((b-1) / (a \<bullet> a)) *\<^sub>R a" in subsetD)
       
 10585     apply force+
       
 10586     done
       
 10587   then show ?thesis by force
       
 10588 qed
       
 10589 
       
 10590 lemma halfspace_eq_empty_gt:
       
 10591    "{x. a \<bullet> x > b} = {} \<longleftrightarrow> a = 0 \<and> b \<ge> 0"
       
 10592 using halfspace_eq_empty_lt [of "-a" "-b"]
       
 10593 by simp
       
 10594 
       
 10595 lemma halfspace_eq_empty_le:
       
 10596    "{x. a \<bullet> x \<le> b} = {} \<longleftrightarrow> a = 0 \<and> b < 0"
       
 10597 proof -
       
 10598   have "{x. a \<bullet> x \<le> b} \<subseteq> {} \<Longrightarrow> a = 0 \<and> b < 0"
       
 10599     apply (rule ccontr)
       
 10600     apply (drule_tac c = "((b-1) / (a \<bullet> a)) *\<^sub>R a" in subsetD)
       
 10601     apply force+
       
 10602     done
       
 10603   then show ?thesis by force
       
 10604 qed
       
 10605 
       
 10606 lemma halfspace_eq_empty_ge:
       
 10607    "{x. a \<bullet> x \<ge> b} = {} \<longleftrightarrow> a = 0 \<and> b > 0"
       
 10608 using halfspace_eq_empty_le [of "-a" "-b"]
       
 10609 by simp
       
 10610 
       
 10611 subsection\<open>Use set distance for an easy proof of separation properties\<close>
       
 10612 
       
 10613 proposition separation_closures:
       
 10614   fixes S :: "'a::euclidean_space set"
       
 10615   assumes "S \<inter> closure T = {}" "T \<inter> closure S = {}"
       
 10616   obtains U V where "U \<inter> V = {}" "open U" "open V" "S \<subseteq> U" "T \<subseteq> V"
       
 10617 proof (cases "S = {} \<or> T = {}")
       
 10618   case True with that show ?thesis by auto
       
 10619 next
       
 10620   case False
       
 10621   define f where "f \<equiv> \<lambda>x. setdist {x} T - setdist {x} S"
       
 10622   have contf: "continuous_on UNIV f"
       
 10623     unfolding f_def by (intro continuous_intros continuous_on_setdist)
       
 10624   show ?thesis
       
 10625   proof (rule_tac U = "{x. f x > 0}" and V = "{x. f x < 0}" in that)
       
 10626     show "{x. 0 < f x} \<inter> {x. f x < 0} = {}"
       
 10627       by auto
       
 10628     show "open {x. 0 < f x}"
       
 10629       by (simp add: open_Collect_less contf continuous_on_const)
       
 10630     show "open {x. f x < 0}"
       
 10631       by (simp add: open_Collect_less contf continuous_on_const)
       
 10632     show "S \<subseteq> {x. 0 < f x}"
       
 10633       apply (clarsimp simp add: f_def setdist_sing_in_set)
       
 10634       using assms
       
 10635       by (metis False IntI empty_iff le_less setdist_eq_0_sing_2 setdist_pos_le setdist_sym)
       
 10636     show "T \<subseteq> {x. f x < 0}"
       
 10637       apply (clarsimp simp add: f_def setdist_sing_in_set)
       
 10638       using assms
       
 10639       by (metis False IntI empty_iff le_less setdist_eq_0_sing_2 setdist_pos_le setdist_sym)
       
 10640   qed
       
 10641 qed
       
 10642 
       
 10643 lemma separation_normal:
       
 10644   fixes S :: "'a::euclidean_space set"
       
 10645   assumes "closed S" "closed T" "S \<inter> T = {}"
       
 10646   obtains U V where "open U" "open V" "S \<subseteq> U" "T \<subseteq> V" "U \<inter> V = {}"
       
 10647 using separation_closures [of S T]
       
 10648 by (metis assms closure_closed disjnt_def inf_commute)
       
 10649 
       
 10650 
       
 10651 lemma separation_normal_local:
       
 10652   fixes S :: "'a::euclidean_space set"
       
 10653   assumes US: "closedin (subtopology euclidean U) S"
       
 10654       and UT: "closedin (subtopology euclidean U) T"
       
 10655       and "S \<inter> T = {}"
       
 10656   obtains S' T' where "openin (subtopology euclidean U) S'"
       
 10657                       "openin (subtopology euclidean U) T'"
       
 10658                       "S \<subseteq> S'"  "T \<subseteq> T'"  "S' \<inter> T' = {}"
       
 10659 proof (cases "S = {} \<or> T = {}")
       
 10660   case True with that show ?thesis
       
 10661     apply safe
       
 10662     using UT closedin_subset apply blast
       
 10663     using US closedin_subset apply blast
       
 10664     done
       
 10665 next
       
 10666   case False
       
 10667   define f where "f \<equiv> \<lambda>x. setdist {x} T - setdist {x} S"
       
 10668   have contf: "continuous_on U f"
       
 10669     unfolding f_def by (intro continuous_intros)
       
 10670   show ?thesis
       
 10671   proof (rule_tac S' = "{x\<in>U. f x > 0}" and T' = "{x\<in>U. f x < 0}" in that)
       
 10672     show "{x \<in> U. 0 < f x} \<inter> {x \<in> U. f x < 0} = {}"
       
 10673       by auto
       
 10674     have "openin (subtopology euclidean U) {x \<in> U. f x \<in> {0<..}}"
       
 10675       apply (rule continuous_openin_preimage [where T=UNIV])
       
 10676         apply (simp_all add: contf)
       
 10677       using open_greaterThan open_openin by blast
       
 10678     then show "openin (subtopology euclidean U) {x \<in> U. 0 < f x}" by simp
       
 10679   next
       
 10680     have "openin (subtopology euclidean U) {x \<in> U. f x \<in> {..<0}}"
       
 10681       apply (rule continuous_openin_preimage [where T=UNIV])
       
 10682         apply (simp_all add: contf)
       
 10683       using open_lessThan open_openin by blast
       
 10684     then show "openin (subtopology euclidean U) {x \<in> U. f x < 0}" by simp
       
 10685   next
       
 10686     show "S \<subseteq> {x \<in> U. 0 < f x}"
       
 10687       apply (clarsimp simp add: f_def setdist_sing_in_set)
       
 10688       using assms
       
 10689       by (metis False Int_insert_right closedin_imp_subset empty_not_insert insert_absorb insert_subset linorder_neqE_linordered_idom not_le setdist_eq_0_closedin setdist_pos_le)
       
 10690     show "T \<subseteq> {x \<in> U. f x < 0}"
       
 10691       apply (clarsimp simp add: f_def setdist_sing_in_set)
       
 10692       using assms
       
 10693       by (metis False closedin_subset disjoint_iff_not_equal insert_absorb insert_subset linorder_neqE_linordered_idom not_less setdist_eq_0_closedin setdist_pos_le topspace_euclidean_subtopology)
       
 10694   qed
       
 10695 qed
       
 10696 
       
 10697 lemma separation_normal_compact:
       
 10698   fixes S :: "'a::euclidean_space set"
       
 10699   assumes "compact S" "closed T" "S \<inter> T = {}"
       
 10700   obtains U V where "open U" "compact(closure U)" "open V" "S \<subseteq> U" "T \<subseteq> V" "U \<inter> V = {}"
       
 10701 proof -
       
 10702   have "closed S" "bounded S"
       
 10703     using assms by (auto simp: compact_eq_bounded_closed)
       
 10704   then obtain r where "r>0" and r: "S \<subseteq> ball 0 r"
       
 10705     by (auto dest!: bounded_subset_ballD)
       
 10706   have **: "closed (T \<union> - ball 0 r)" "S \<inter> (T \<union> - ball 0 r) = {}"
       
 10707     using assms r by blast+
       
 10708   then show ?thesis
       
 10709     apply (rule separation_normal [OF \<open>closed S\<close>])
       
 10710     apply (rule_tac U=U and V=V in that)
       
 10711     by auto (meson bounded_ball bounded_subset compl_le_swap2 disjoint_eq_subset_Compl)
       
 10712 qed
       
 10713 
       
 10714 subsection\<open>Proper maps, including projections out of compact sets\<close>
       
 10715 
       
 10716 lemma finite_indexed_bound:
       
 10717   assumes A: "finite A" "\<And>x. x \<in> A \<Longrightarrow> \<exists>n::'a::linorder. P x n"
       
 10718     shows "\<exists>m. \<forall>x \<in> A. \<exists>k\<le>m. P x k"
       
 10719 using A
       
 10720 proof (induction A)
       
 10721   case empty then show ?case by force
       
 10722 next
       
 10723   case (insert a A)
       
 10724     then obtain m n where "\<forall>x \<in> A. \<exists>k\<le>m. P x k" "P a n"
       
 10725       by force
       
 10726     then show ?case
       
 10727       apply (rule_tac x="max m n" in exI, safe)
       
 10728       using max.cobounded2 apply blast
       
 10729       by (meson le_max_iff_disj)
       
 10730 qed
       
 10731 
       
 10732 proposition proper_map:
       
 10733   fixes f :: "'a::heine_borel \<Rightarrow> 'b::heine_borel"
       
 10734   assumes "closedin (subtopology euclidean S) K"
       
 10735       and com: "\<And>U. \<lbrakk>U \<subseteq> T; compact U\<rbrakk> \<Longrightarrow> compact {x \<in> S. f x \<in> U}"
       
 10736       and "f ` S \<subseteq> T"
       
 10737     shows "closedin (subtopology euclidean T) (f ` K)"
       
 10738 proof -
       
 10739   have "K \<subseteq> S"
       
 10740     using assms closedin_imp_subset by metis
       
 10741   obtain C where "closed C" and Keq: "K = S \<inter> C"
       
 10742     using assms by (auto simp: closedin_closed)
       
 10743   have *: "y \<in> f ` K" if "y \<in> T" and y: "y islimpt f ` K" for y
       
 10744   proof -
       
 10745     obtain h where "\<forall>n. (\<exists>x\<in>K. h n = f x) \<and> h n \<noteq> y" "inj h" and hlim: "(h \<longlongrightarrow> y) sequentially"
       
 10746       using \<open>y \<in> T\<close> y by (force simp: limpt_sequential_inj)
       
 10747     then obtain X where X: "\<And>n. X n \<in> K \<and> h n = f (X n) \<and> h n \<noteq> y"
       
 10748       by metis
       
 10749     then have fX: "\<And>n. f (X n) = h n"
       
 10750       by metis
       
 10751     have "compact (C \<inter> {a \<in> S. f a \<in> insert y (range (\<lambda>i. f(X(n + i))))})" for n
       
 10752       apply (rule closed_Int_compact [OF \<open>closed C\<close>])
       
 10753       apply (rule com)
       
 10754        using X \<open>K \<subseteq> S\<close> \<open>f ` S \<subseteq> T\<close> \<open>y \<in> T\<close> apply blast
       
 10755       apply (rule compact_sequence_with_limit)
       
 10756       apply (simp add: fX add.commute [of n] LIMSEQ_ignore_initial_segment [OF hlim])
       
 10757       done
       
 10758     then have comf: "compact {a \<in> K. f a \<in> insert y (range (\<lambda>i. f(X(n + i))))}" for n
       
 10759       by (simp add: Keq Int_def conj_commute)
       
 10760     have ne: "\<Inter>\<F> \<noteq> {}"
       
 10761              if "finite \<F>"
       
 10762                 and \<F>: "\<And>t. t \<in> \<F> \<Longrightarrow>
       
 10763                            (\<exists>n. t = {a \<in> K. f a \<in> insert y (range (\<lambda>i. f (X (n + i))))})"
       
 10764              for \<F>
       
 10765     proof -
       
 10766       obtain m where m: "\<And>t. t \<in> \<F> \<Longrightarrow> \<exists>k\<le>m. t = {a \<in> K. f a \<in> insert y (range (\<lambda>i. f (X (k + i))))}"
       
 10767         apply (rule exE)
       
 10768         apply (rule finite_indexed_bound [OF \<open>finite \<F>\<close> \<F>], assumption, force)
       
 10769         done
       
 10770       have "X m \<in> \<Inter>\<F>"
       
 10771         using X le_Suc_ex by (fastforce dest: m)
       
 10772       then show ?thesis by blast
       
 10773     qed
       
 10774     have "\<Inter>{{a. a \<in> K \<and> f a \<in> insert y (range (\<lambda>i. f(X(n + i))))} |n. n \<in> UNIV}
       
 10775                \<noteq> {}"
       
 10776       apply (rule compact_fip_heine_borel)
       
 10777        using comf apply force
       
 10778       using ne  apply (simp add: subset_iff del: insert_iff)
       
 10779       done
       
 10780     then have "\<exists>x. x \<in> (\<Inter>n. {a \<in> K. f a \<in> insert y (range (\<lambda>i. f (X (n + i))))})"
       
 10781       by blast
       
 10782     then show ?thesis
       
 10783       apply (simp add: image_iff fX)
       
 10784       by (metis \<open>inj h\<close> le_add1 not_less_eq_eq rangeI range_ex1_eq)
       
 10785   qed
       
 10786   with assms closedin_subset show ?thesis
       
 10787     by (force simp: closedin_limpt)
       
 10788 qed
       
 10789 
       
 10790 
       
 10791 lemma compact_continuous_image_eq:
       
 10792   fixes f :: "'a::heine_borel \<Rightarrow> 'b::heine_borel"
       
 10793   assumes f: "inj_on f S"
       
 10794   shows "continuous_on S f \<longleftrightarrow> (\<forall>T. compact T \<and> T \<subseteq> S \<longrightarrow> compact(f ` T))"
       
 10795            (is "?lhs = ?rhs")
       
 10796 proof
       
 10797   assume ?lhs then show ?rhs
       
 10798     by (metis continuous_on_subset compact_continuous_image)
       
 10799 next
       
 10800   assume RHS: ?rhs
       
 10801   obtain g where gf: "\<And>x. x \<in> S \<Longrightarrow> g (f x) = x"
       
 10802     by (metis inv_into_f_f f)
       
 10803   then have *: "{x \<in> S. f x \<in> U} = g ` U" if "U \<subseteq> f ` S" for U
       
 10804     using that by fastforce
       
 10805   have gfim: "g ` f ` S \<subseteq> S" using gf by auto
       
 10806   have **: "compact {x \<in> f ` S. g x \<in> C}" if C: "C \<subseteq> S" "compact C" for C
       
 10807   proof -
       
 10808     obtain h :: "'a set \<Rightarrow> 'a" where "h C \<in> C \<and> h C \<notin> S \<or> compact (f ` C)"
       
 10809       by (force simp: C RHS)
       
 10810     moreover have "f ` C = {b \<in> f ` S. g b \<in> C}"
       
 10811       using C gf by auto
       
 10812     ultimately show "compact {b \<in> f ` S. g b \<in> C}"
       
 10813       using C by auto
       
 10814   qed
       
 10815   show ?lhs
       
 10816     using proper_map [OF _ _ gfim] **
       
 10817     by (simp add: continuous_on_closed * closedin_imp_subset)
       
 10818 qed
       
 10819 
       
 10820 lemma proper_map_from_compact:
       
 10821   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
 10822   assumes contf: "continuous_on S f" and imf: "f ` S \<subseteq> T" and "compact S"
       
 10823           "closedin (subtopology euclidean T) K"
       
 10824   shows "compact {x. x \<in> S \<and> f x \<in> K}"
       
 10825 by (rule closedin_compact [OF \<open>compact S\<close>] continuous_closedin_preimage_gen assms)+
       
 10826 
       
 10827 lemma proper_map_fst:
       
 10828   assumes "compact T" "K \<subseteq> S" "compact K"
       
 10829     shows "compact {z \<in> S \<times> T. fst z \<in> K}"
       
 10830 proof -
       
 10831   have "{z \<in> S \<times> T. fst z \<in> K} = K \<times> T"
       
 10832     using assms by auto
       
 10833   then show ?thesis by (simp add: assms compact_Times)
       
 10834 qed
       
 10835 
       
 10836 lemma closed_map_fst:
       
 10837   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
 10838   assumes "compact T" "closedin (subtopology euclidean (S \<times> T)) c"
       
 10839    shows "closedin (subtopology euclidean S) (fst ` c)"
       
 10840 proof -
       
 10841   have *: "fst ` (S \<times> T) \<subseteq> S"
       
 10842     by auto
       
 10843   show ?thesis
       
 10844     using proper_map [OF _ _ *] by (simp add: proper_map_fst assms)
       
 10845 qed
       
 10846 
       
 10847 lemma proper_map_snd:
       
 10848   assumes "compact S" "K \<subseteq> T" "compact K"
       
 10849     shows "compact {z \<in> S \<times> T. snd z \<in> K}"
       
 10850 proof -
       
 10851   have "{z \<in> S \<times> T. snd z \<in> K} = S \<times> K"
       
 10852     using assms by auto
       
 10853   then show ?thesis by (simp add: assms compact_Times)
       
 10854 qed
       
 10855 
       
 10856 lemma closed_map_snd:
       
 10857   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
 10858   assumes "compact S" "closedin (subtopology euclidean (S \<times> T)) c"
       
 10859    shows "closedin (subtopology euclidean T) (snd ` c)"
       
 10860 proof -
       
 10861   have *: "snd ` (S \<times> T) \<subseteq> T"
       
 10862     by auto
       
 10863   show ?thesis
       
 10864     using proper_map [OF _ _ *] by (simp add: proper_map_snd assms)
       
 10865 qed
       
 10866 
       
 10867 lemma closedin_compact_projection:
       
 10868   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
 10869   assumes "compact S" and clo: "closedin (subtopology euclidean (S \<times> T)) U"
       
 10870     shows "closedin (subtopology euclidean T) {y. \<exists>x. x \<in> S \<and> (x, y) \<in> U}"
       
 10871 proof -
       
 10872   have "U \<subseteq> S \<times> T"
       
 10873     by (metis clo closedin_imp_subset)
       
 10874   then have "{y. \<exists>x. x \<in> S \<and> (x, y) \<in> U} = snd ` U"
       
 10875     by force
       
 10876   moreover have "closedin (subtopology euclidean T) (snd ` U)"
       
 10877     by (rule closed_map_snd [OF assms])
       
 10878   ultimately show ?thesis
       
 10879     by simp
       
 10880 qed
       
 10881 
       
 10882 
       
 10883 lemma closed_compact_projection:
       
 10884   fixes S :: "'a::euclidean_space set"
       
 10885     and T :: "('a * 'b::euclidean_space) set"
       
 10886   assumes "compact S" and clo: "closed T"
       
 10887     shows "closed {y. \<exists>x. x \<in> S \<and> (x, y) \<in> T}"
       
 10888 proof -
       
 10889   have *: "{y. \<exists>x. x \<in> S \<and> Pair x y \<in> T} =
       
 10890         {y. \<exists>x. x \<in> S \<and> Pair x y \<in> ((S \<times> UNIV) \<inter> T)}"
       
 10891     by auto
       
 10892   show ?thesis
       
 10893     apply (subst *)
       
 10894     apply (rule closedin_closed_trans [OF _ closed_UNIV])
       
 10895     apply (rule closedin_compact_projection [OF \<open>compact S\<close>])
       
 10896     by (simp add: clo closedin_closed_Int)
       
 10897 qed
       
 10898 
       
 10899 subsubsection\<open>Representing affine hull as a finite intersection of hyperplanes\<close>
       
 10900 
       
 10901 proposition affine_hull_convex_Int_nonempty_interior:
       
 10902   fixes S :: "'a::real_normed_vector set"
       
 10903   assumes "convex S" "S \<inter> interior T \<noteq> {}"
       
 10904     shows "affine hull (S \<inter> T) = affine hull S"
       
 10905 proof
       
 10906   show "affine hull (S \<inter> T) \<subseteq> affine hull S"
       
 10907     by (simp add: hull_mono)
       
 10908 next
       
 10909   obtain a where "a \<in> S" "a \<in> T" and at: "a \<in> interior T"
       
 10910     using assms interior_subset by blast
       
 10911   then obtain e where "e > 0" and e: "cball a e \<subseteq> T"
       
 10912     using mem_interior_cball by blast
       
 10913   have *: "x \<in> op + a ` span ((\<lambda>x. x - a) ` (S \<inter> T))" if "x \<in> S" for x
       
 10914   proof (cases "x = a")
       
 10915     case True with that span_0 eq_add_iff image_def mem_Collect_eq show ?thesis
       
 10916       by blast
       
 10917   next
       
 10918     case False
       
 10919     define k where "k = min (1/2) (e / norm (x-a))"
       
 10920     have k: "0 < k" "k < 1"
       
 10921       using \<open>e > 0\<close> False by (auto simp: k_def)
       
 10922     then have xa: "(x-a) = inverse k *\<^sub>R k *\<^sub>R (x-a)"
       
 10923       by simp
       
 10924     have "e / norm (x - a) \<ge> k"
       
 10925       using k_def by linarith
       
 10926     then have "a + k *\<^sub>R (x - a) \<in> cball a e"
       
 10927       using \<open>0 < k\<close> False by (simp add: dist_norm field_simps)
       
 10928     then have T: "a + k *\<^sub>R (x - a) \<in> T"
       
 10929       using e by blast
       
 10930     have S: "a + k *\<^sub>R (x - a) \<in> S"
       
 10931       using k \<open>a \<in> S\<close> convexD [OF \<open>convex S\<close> \<open>a \<in> S\<close> \<open>x \<in> S\<close>, of "1-k" k]
       
 10932       by (simp add: algebra_simps)
       
 10933     have "inverse k *\<^sub>R k *\<^sub>R (x-a) \<in> span ((\<lambda>x. x - a) ` (S \<inter> T))"
       
 10934       apply (rule span_mul)
       
 10935       apply (rule span_superset)
       
 10936       apply (rule image_eqI [where x = "a + k *\<^sub>R (x - a)"])
       
 10937       apply (auto simp: S T)
       
 10938       done
       
 10939     with xa image_iff show ?thesis  by fastforce
       
 10940   qed
       
 10941   show "affine hull S \<subseteq> affine hull (S \<inter> T)"
       
 10942     apply (simp add: subset_hull)
       
 10943     apply (simp add: \<open>a \<in> S\<close> \<open>a \<in> T\<close> hull_inc affine_hull_span_gen [of a])
       
 10944     apply (force simp: *)
       
 10945     done
       
 10946 qed
       
 10947 
       
 10948 corollary affine_hull_convex_Int_open:
       
 10949   fixes S :: "'a::real_normed_vector set"
       
 10950   assumes "convex S" "open T" "S \<inter> T \<noteq> {}"
       
 10951     shows "affine hull (S \<inter> T) = affine hull S"
       
 10952 using affine_hull_convex_Int_nonempty_interior assms interior_eq by blast
       
 10953 
       
 10954 corollary affine_hull_affine_Int_nonempty_interior:
       
 10955   fixes S :: "'a::real_normed_vector set"
       
 10956   assumes "affine S" "S \<inter> interior T \<noteq> {}"
       
 10957     shows "affine hull (S \<inter> T) = affine hull S"
       
 10958 by (simp add: affine_hull_convex_Int_nonempty_interior affine_imp_convex assms)
       
 10959 
       
 10960 corollary affine_hull_affine_Int_open:
       
 10961   fixes S :: "'a::real_normed_vector set"
       
 10962   assumes "affine S" "open T" "S \<inter> T \<noteq> {}"
       
 10963     shows "affine hull (S \<inter> T) = affine hull S"
       
 10964 by (simp add: affine_hull_convex_Int_open affine_imp_convex assms)
       
 10965 
       
 10966 corollary affine_hull_convex_Int_openin:
       
 10967   fixes S :: "'a::real_normed_vector set"
       
 10968   assumes "convex S" "openin (subtopology euclidean (affine hull S)) T" "S \<inter> T \<noteq> {}"
       
 10969     shows "affine hull (S \<inter> T) = affine hull S"
       
 10970 using assms unfolding openin_open
       
 10971 by (metis affine_hull_convex_Int_open hull_subset inf.orderE inf_assoc)
       
 10972 
       
 10973 corollary affine_hull_open_in:
       
 10974   fixes S :: "'a::real_normed_vector set"
       
 10975   assumes "openin (subtopology euclidean (affine hull T)) S" "S \<noteq> {}"
       
 10976     shows "affine hull S = affine hull T"
       
 10977 using assms unfolding openin_open
       
 10978 by (metis affine_affine_hull affine_hull_affine_Int_open hull_hull)
       
 10979 
       
 10980 corollary affine_hull_open:
       
 10981   fixes S :: "'a::real_normed_vector set"
       
 10982   assumes "open S" "S \<noteq> {}"
       
 10983     shows "affine hull S = UNIV"
       
 10984 by (metis affine_hull_convex_Int_nonempty_interior assms convex_UNIV hull_UNIV inf_top.left_neutral interior_open)
       
 10985 
       
 10986 lemma aff_dim_convex_Int_nonempty_interior:
       
 10987   fixes S :: "'a::euclidean_space set"
       
 10988   shows "\<lbrakk>convex S; S \<inter> interior T \<noteq> {}\<rbrakk> \<Longrightarrow> aff_dim(S \<inter> T) = aff_dim S"
       
 10989 using aff_dim_affine_hull2 affine_hull_convex_Int_nonempty_interior by blast
       
 10990 
       
 10991 lemma aff_dim_convex_Int_open:
       
 10992   fixes S :: "'a::euclidean_space set"
       
 10993   shows "\<lbrakk>convex S; open T; S \<inter> T \<noteq> {}\<rbrakk> \<Longrightarrow>  aff_dim(S \<inter> T) = aff_dim S"
       
 10994 using aff_dim_convex_Int_nonempty_interior interior_eq by blast
       
 10995 
       
 10996 lemma affine_hull_halfspace_lt:
       
 10997   fixes a :: "'a::euclidean_space"
       
 10998   shows "affine hull {x. a \<bullet> x < r} = (if a = 0 \<and> r \<le> 0 then {} else UNIV)"
       
 10999 using halfspace_eq_empty_lt [of a r]
       
 11000 by (simp add: open_halfspace_lt affine_hull_open)
       
 11001 
       
 11002 lemma affine_hull_halfspace_le:
       
 11003   fixes a :: "'a::euclidean_space"
       
 11004   shows "affine hull {x. a \<bullet> x \<le> r} = (if a = 0 \<and> r < 0 then {} else UNIV)"
       
 11005 proof (cases "a = 0")
       
 11006   case True then show ?thesis by simp
       
 11007 next
       
 11008   case False
       
 11009   then have "affine hull closure {x. a \<bullet> x < r} = UNIV"
       
 11010     using affine_hull_halfspace_lt closure_same_affine_hull by fastforce
       
 11011   moreover have "{x. a \<bullet> x < r} \<subseteq> {x. a \<bullet> x \<le> r}"
       
 11012     by (simp add: Collect_mono)
       
 11013   ultimately show ?thesis using False antisym_conv hull_mono top_greatest
       
 11014     by (metis affine_hull_halfspace_lt)
       
 11015 qed
       
 11016 
       
 11017 lemma affine_hull_halfspace_gt:
       
 11018   fixes a :: "'a::euclidean_space"
       
 11019   shows "affine hull {x. a \<bullet> x > r} = (if a = 0 \<and> r \<ge> 0 then {} else UNIV)"
       
 11020 using halfspace_eq_empty_gt [of r a]
       
 11021 by (simp add: open_halfspace_gt affine_hull_open)
       
 11022 
       
 11023 lemma affine_hull_halfspace_ge:
       
 11024   fixes a :: "'a::euclidean_space"
       
 11025   shows "affine hull {x. a \<bullet> x \<ge> r} = (if a = 0 \<and> r > 0 then {} else UNIV)"
       
 11026 using affine_hull_halfspace_le [of "-a" "-r"] by simp
       
 11027 
       
 11028 lemma aff_dim_halfspace_lt:
       
 11029   fixes a :: "'a::euclidean_space"
       
 11030   shows "aff_dim {x. a \<bullet> x < r} =
       
 11031         (if a = 0 \<and> r \<le> 0 then -1 else DIM('a))"
       
 11032 by simp (metis aff_dim_open halfspace_eq_empty_lt open_halfspace_lt)
       
 11033 
       
 11034 lemma aff_dim_halfspace_le:
       
 11035   fixes a :: "'a::euclidean_space"
       
 11036   shows "aff_dim {x. a \<bullet> x \<le> r} =
       
 11037         (if a = 0 \<and> r < 0 then -1 else DIM('a))"
       
 11038 proof -
       
 11039   have "int (DIM('a)) = aff_dim (UNIV::'a set)"
       
 11040     by (simp add: aff_dim_UNIV)
       
 11041   then have "aff_dim (affine hull {x. a \<bullet> x \<le> r}) = DIM('a)" if "(a = 0 \<longrightarrow> r \<ge> 0)"
       
 11042     using that by (simp add: affine_hull_halfspace_le not_less)
       
 11043   then show ?thesis
       
 11044     by (force simp: aff_dim_affine_hull)
       
 11045 qed
       
 11046 
       
 11047 lemma aff_dim_halfspace_gt:
       
 11048   fixes a :: "'a::euclidean_space"
       
 11049   shows "aff_dim {x. a \<bullet> x > r} =
       
 11050         (if a = 0 \<and> r \<ge> 0 then -1 else DIM('a))"
       
 11051 by simp (metis aff_dim_open halfspace_eq_empty_gt open_halfspace_gt)
       
 11052 
       
 11053 lemma aff_dim_halfspace_ge:
       
 11054   fixes a :: "'a::euclidean_space"
       
 11055   shows "aff_dim {x. a \<bullet> x \<ge> r} =
       
 11056         (if a = 0 \<and> r > 0 then -1 else DIM('a))"
       
 11057 using aff_dim_halfspace_le [of "-a" "-r"] by simp
       
 11058 
       
 11059 subsection\<open>Properties of special hyperplanes\<close>
       
 11060 
       
 11061 lemma subspace_hyperplane: "subspace {x. a \<bullet> x = 0}"
       
 11062   by (simp add: subspace_def inner_right_distrib)
       
 11063 
       
 11064 lemma subspace_hyperplane2: "subspace {x. x \<bullet> a = 0}"
       
 11065   by (simp add: inner_commute inner_right_distrib subspace_def)
       
 11066 
       
 11067 lemma special_hyperplane_span:
       
 11068   fixes S :: "'n::euclidean_space set"
       
 11069   assumes "k \<in> Basis"
       
 11070   shows "{x. k \<bullet> x = 0} = span (Basis - {k})"
       
 11071 proof -
       
 11072   have *: "x \<in> span (Basis - {k})" if "k \<bullet> x = 0" for x
       
 11073   proof -
       
 11074     have "x = (\<Sum>b\<in>Basis. (x \<bullet> b) *\<^sub>R b)"
       
 11075       by (simp add: euclidean_representation)
       
 11076     also have "... = (\<Sum>b \<in> Basis - {k}. (x \<bullet> b) *\<^sub>R b)"
       
 11077       by (auto simp: setsum.remove [of _ k] inner_commute assms that)
       
 11078     finally have "x = (\<Sum>b\<in>Basis - {k}. (x \<bullet> b) *\<^sub>R b)" .
       
 11079     then show ?thesis
       
 11080       by (simp add: Linear_Algebra.span_finite) metis
       
 11081   qed
       
 11082   show ?thesis
       
 11083     apply (rule span_subspace [symmetric])
       
 11084     using assms
       
 11085     apply (auto simp: inner_not_same_Basis intro: * subspace_hyperplane)
       
 11086     done
       
 11087 qed
       
 11088 
       
 11089 lemma dim_special_hyperplane:
       
 11090   fixes k :: "'n::euclidean_space"
       
 11091   shows "k \<in> Basis \<Longrightarrow> dim {x. k \<bullet> x = 0} = DIM('n) - 1"
       
 11092 apply (simp add: special_hyperplane_span)
       
 11093 apply (rule Linear_Algebra.dim_unique [OF subset_refl])
       
 11094 apply (auto simp: Diff_subset independent_substdbasis)
       
 11095 apply (metis member_remove remove_def span_clauses(1))
       
 11096 done
       
 11097 
       
 11098 proposition dim_hyperplane:
       
 11099   fixes a :: "'a::euclidean_space"
       
 11100   assumes "a \<noteq> 0"
       
 11101     shows "dim {x. a \<bullet> x = 0} = DIM('a) - 1"
       
 11102 proof -
       
 11103   have span0: "span {x. a \<bullet> x = 0} = {x. a \<bullet> x = 0}"
       
 11104     by (rule span_unique) (auto simp: subspace_hyperplane)
       
 11105   then obtain B where "independent B"
       
 11106               and Bsub: "B \<subseteq> {x. a \<bullet> x = 0}"
       
 11107               and subspB: "{x. a \<bullet> x = 0} \<subseteq> span B"
       
 11108               and card0: "(card B = dim {x. a \<bullet> x = 0})"
       
 11109               and ortho: "pairwise orthogonal B"
       
 11110     using orthogonal_basis_exists by metis
       
 11111   with assms have "a \<notin> span B"
       
 11112     by (metis (mono_tags, lifting) span_eq inner_eq_zero_iff mem_Collect_eq span0 span_subspace)
       
 11113   then have ind: "independent (insert a B)"
       
 11114     by (simp add: \<open>independent B\<close> independent_insert)
       
 11115   have "finite B"
       
 11116     using \<open>independent B\<close> independent_bound by blast
       
 11117   have "UNIV \<subseteq> span (insert a B)"
       
 11118   proof fix y::'a
       
 11119     obtain r z where z: "y = r *\<^sub>R a + z" "a \<bullet> z = 0"
       
 11120       apply (rule_tac r="(a \<bullet> y) / (a \<bullet> a)" and z = "y - ((a \<bullet> y) / (a \<bullet> a)) *\<^sub>R a" in that)
       
 11121       using assms
       
 11122       by (auto simp: algebra_simps)
       
 11123     show "y \<in> span (insert a B)"
       
 11124       by (metis (mono_tags, lifting) z Bsub Convex_Euclidean_Space.span_eq
       
 11125          add_diff_cancel_left' mem_Collect_eq span0 span_breakdown_eq span_subspace subspB)
       
 11126   qed
       
 11127   then have dima: "DIM('a) = dim(insert a B)"
       
 11128     by (metis antisym dim_UNIV dim_subset_UNIV subset_le_dim)
       
 11129   then show ?thesis
       
 11130     by (metis (mono_tags, lifting) Bsub Diff_insert_absorb \<open>a \<notin> span B\<close> ind card0 card_Diff_singleton dim_span indep_card_eq_dim_span insertI1 subsetCE subspB)
       
 11131 qed
       
 11132 
       
 11133 lemma lowdim_eq_hyperplane:
       
 11134   fixes S :: "'a::euclidean_space set"
       
 11135   assumes "dim S = DIM('a) - 1"
       
 11136   obtains a where "a \<noteq> 0" and "span S = {x. a \<bullet> x = 0}"
       
 11137 proof -
       
 11138   have [simp]: "dim S < DIM('a)"
       
 11139     by (simp add: DIM_positive assms)
       
 11140   then obtain b where b: "b \<noteq> 0" "span S \<subseteq> {a. b \<bullet> a = 0}"
       
 11141     using lowdim_subset_hyperplane [of S] by fastforce
       
 11142   show ?thesis
       
 11143     using b that real_vector_class.subspace_span [of S]
       
 11144     by (simp add: assms dim_hyperplane subspace_dim_equal subspace_hyperplane)
       
 11145 qed
       
 11146 
       
 11147 lemma dim_eq_hyperplane:
       
 11148   fixes S :: "'n::euclidean_space set"
       
 11149   shows "dim S = DIM('n) - 1 \<longleftrightarrow> (\<exists>a. a \<noteq> 0 \<and> span S = {x. a \<bullet> x = 0})"
       
 11150 by (metis One_nat_def dim_hyperplane dim_span lowdim_eq_hyperplane)
       
 11151 
       
 11152 proposition aff_dim_eq_hyperplane:
       
 11153   fixes S :: "'a::euclidean_space set"
       
 11154   shows "aff_dim S = DIM('a) - 1 \<longleftrightarrow> (\<exists>a b. a \<noteq> 0 \<and> affine hull S = {x. a \<bullet> x = b})"
       
 11155 proof (cases "S = {}")
       
 11156   case True then show ?thesis
       
 11157     by (auto simp: dest: hyperplane_eq_Ex)
       
 11158 next
       
 11159   case False
       
 11160   then obtain c where "c \<in> S" by blast
       
 11161   show ?thesis
       
 11162   proof (cases "c = 0")
       
 11163     case True show ?thesis
       
 11164     apply (simp add: aff_dim_eq_dim [of c] affine_hull_span_gen [of c] \<open>c \<in> S\<close> hull_inc dim_eq_hyperplane
       
 11165                 del: One_nat_def)
       
 11166     apply (rule ex_cong)
       
 11167     apply (metis (mono_tags) span_0 \<open>c = 0\<close> image_add_0 inner_zero_right mem_Collect_eq)
       
 11168     done
       
 11169   next
       
 11170     case False
       
 11171     have xc_im: "x \<in> op + c ` {y. a \<bullet> y = 0}" if "a \<bullet> x = a \<bullet> c" for a x
       
 11172     proof -
       
 11173       have "\<exists>y. a \<bullet> y = 0 \<and> c + y = x"
       
 11174         by (metis that add.commute diff_add_cancel inner_commute inner_diff_left right_minus_eq)
       
 11175       then show "x \<in> op + c ` {y. a \<bullet> y = 0}"
       
 11176         by blast
       
 11177     qed
       
 11178     have 2: "span ((\<lambda>x. x - c) ` S) = {x. a \<bullet> x = 0}"
       
 11179          if "op + c ` span ((\<lambda>x. x - c) ` S) = {x. a \<bullet> x = b}" for a b
       
 11180     proof -
       
 11181       have "b = a \<bullet> c"
       
 11182         using span_0 that by fastforce
       
 11183       with that have "op + c ` span ((\<lambda>x. x - c) ` S) = {x. a \<bullet> x = a \<bullet> c}"
       
 11184         by simp
       
 11185       then have "span ((\<lambda>x. x - c) ` S) = (\<lambda>x. x - c) ` {x. a \<bullet> x = a \<bullet> c}"
       
 11186         by (metis (no_types) image_cong translation_galois uminus_add_conv_diff)
       
 11187       also have "... = {x. a \<bullet> x = 0}"
       
 11188         by (force simp: inner_distrib inner_diff_right
       
 11189              intro: image_eqI [where x="x+c" for x])
       
 11190       finally show ?thesis .
       
 11191     qed
       
 11192     show ?thesis
       
 11193       apply (simp add: aff_dim_eq_dim [of c] affine_hull_span_gen [of c] \<open>c \<in> S\<close> hull_inc dim_eq_hyperplane
       
 11194                   del: One_nat_def, safe)
       
 11195       apply (fastforce simp add: inner_distrib intro: xc_im)
       
 11196       apply (force simp: intro!: 2)
       
 11197       done
       
 11198   qed
       
 11199 qed
       
 11200 
       
 11201 corollary aff_dim_hyperplane [simp]:
       
 11202   fixes a :: "'a::euclidean_space"
       
 11203   shows "a \<noteq> 0 \<Longrightarrow> aff_dim {x. a \<bullet> x = r} = DIM('a) - 1"
       
 11204 by (metis aff_dim_eq_hyperplane affine_hull_eq affine_hyperplane)
       
 11205 
       
 11206 subsection\<open>Some stepping theorems\<close>
       
 11207 
       
 11208 lemma dim_empty [simp]: "dim ({} :: 'a::euclidean_space set) = 0"
       
 11209   by (force intro!: dim_unique)
       
 11210 
       
 11211 lemma dim_insert:
       
 11212   fixes x :: "'a::euclidean_space"
       
 11213   shows "dim (insert x S) = (if x \<in> span S then dim S else dim S + 1)"
       
 11214 proof -
       
 11215   show ?thesis
       
 11216   proof (cases "x \<in> span S")
       
 11217     case True then show ?thesis
       
 11218       by (metis dim_span span_redundant)
       
 11219   next
       
 11220     case False
       
 11221     obtain B where B: "B \<subseteq> span S" "independent B" "span S \<subseteq> span B" "card B = dim (span S)"
       
 11222       using basis_exists [of "span S"] by blast
       
 11223     have 1: "insert x B \<subseteq> span (insert x S)"
       
 11224       by (meson \<open>B \<subseteq> span S\<close> dual_order.trans insertI1 insert_subsetI span_mono span_superset subset_insertI)
       
 11225     have 2: "span (insert x S) \<subseteq> span (insert x B)"
       
 11226       by (metis \<open>B \<subseteq> span S\<close> \<open>span S \<subseteq> span B\<close> span_breakdown_eq span_subspace subsetI subspace_span)
       
 11227     have 3: "independent (insert x B)"
       
 11228       by (metis B independent_insert span_subspace subspace_span False)
       
 11229     have "dim (span (insert x S)) = Suc (dim S)"
       
 11230       apply (rule dim_unique [OF 1 2 3])
       
 11231       by (metis B False card_insert_disjoint dim_span independent_imp_finite subsetCE)
       
 11232     then show ?thesis
       
 11233       by (simp add: False)
       
 11234   qed
       
 11235 qed
       
 11236 
       
 11237 lemma dim_singleton [simp]:
       
 11238   fixes x :: "'a::euclidean_space"
       
 11239   shows "dim{x} = (if x = 0 then 0 else 1)"
       
 11240 by (simp add: dim_insert)
       
 11241 
       
 11242 lemma dim_eq_0 [simp]:
       
 11243   fixes S :: "'a::euclidean_space set"
       
 11244   shows "dim S = 0 \<longleftrightarrow> S \<subseteq> {0}"
       
 11245 apply safe
       
 11246 apply (metis DIM_positive DIM_real card_ge_dim_independent contra_subsetD dim_empty dim_insert dim_singleton empty_subsetI independent_empty less_not_refl zero_le)
       
 11247 by (metis dim_singleton dim_subset le_0_eq)
       
 11248 
       
 11249 lemma aff_dim_insert:
       
 11250   fixes a :: "'a::euclidean_space"
       
 11251   shows "aff_dim (insert a S) = (if a \<in> affine hull S then aff_dim S else aff_dim S + 1)"
       
 11252 proof (cases "S = {}")
       
 11253   case True then show ?thesis
       
 11254     by simp
       
 11255 next
       
 11256   case False
       
 11257   then obtain x s' where S: "S = insert x s'" "x \<notin> s'"
       
 11258     by (meson Set.set_insert all_not_in_conv)
       
 11259   show ?thesis using S
       
 11260     apply (simp add: hull_redundant cong: aff_dim_affine_hull2)
       
 11261     apply (simp add: affine_hull_insert_span_gen hull_inc)
       
 11262     apply (simp add: insert_commute [of a] hull_inc aff_dim_eq_dim [of x] dim_insert span_0)
       
 11263     apply (metis (no_types, lifting) add_minus_cancel image_iff uminus_add_conv_diff)
       
 11264     done
       
 11265 qed
       
 11266 
       
 11267 lemma subspace_bounded_eq_trivial:
       
 11268   fixes S :: "'a::real_normed_vector set"
       
 11269   assumes "subspace S"
       
 11270     shows "bounded S \<longleftrightarrow> S = {0}"
       
 11271 proof -
       
 11272   have "False" if "bounded S" "x \<in> S" "x \<noteq> 0" for x
       
 11273   proof -
       
 11274     obtain B where B: "\<And>y. y \<in> S \<Longrightarrow> norm y < B" "B > 0"
       
 11275       using \<open>bounded S\<close> by (force simp: bounded_pos_less)
       
 11276     have "(B / norm x) *\<^sub>R x \<in> S"
       
 11277       using assms subspace_mul \<open>x \<in> S\<close> by auto
       
 11278     moreover have "norm ((B / norm x) *\<^sub>R x) = B"
       
 11279       using that B by (simp add: algebra_simps)
       
 11280     ultimately show False using B by force
       
 11281   qed
       
 11282   then have "bounded S \<Longrightarrow> S = {0}"
       
 11283     using assms subspace_0 by fastforce
       
 11284   then show ?thesis
       
 11285     by blast
       
 11286 qed
       
 11287 
       
 11288 lemma affine_bounded_eq_trivial:
       
 11289   fixes S :: "'a::real_normed_vector set"
       
 11290   assumes "affine S"
       
 11291     shows "bounded S \<longleftrightarrow> S = {} \<or> (\<exists>a. S = {a})"
       
 11292 proof (cases "S = {}")
       
 11293   case True then show ?thesis
       
 11294     by simp
       
 11295 next
       
 11296   case False
       
 11297   then obtain b where "b \<in> S" by blast
       
 11298   with False assms show ?thesis
       
 11299     apply safe
       
 11300     using affine_diffs_subspace [OF assms \<open>b \<in> S\<close>]
       
 11301     apply (metis (no_types, lifting) subspace_bounded_eq_trivial ab_left_minus bounded_translation
       
 11302                 image_empty image_insert translation_invert)
       
 11303     apply force
       
 11304     done
       
 11305 qed
       
 11306 
       
 11307 lemma affine_bounded_eq_lowdim:
       
 11308   fixes S :: "'a::euclidean_space set"
       
 11309   assumes "affine S"
       
 11310     shows "bounded S \<longleftrightarrow> aff_dim S \<le> 0"
       
 11311 apply safe
       
 11312 using affine_bounded_eq_trivial assms apply fastforce
       
 11313 by (metis aff_dim_sing aff_dim_subset affine_dim_equal affine_sing all_not_in_conv assms bounded_empty bounded_insert dual_order.antisym empty_subsetI insert_subset)
       
 11314 
       
 11315 
       
 11316 lemma bounded_hyperplane_eq_trivial_0:
       
 11317   fixes a :: "'a::euclidean_space"
       
 11318   assumes "a \<noteq> 0"
       
 11319   shows "bounded {x. a \<bullet> x = 0} \<longleftrightarrow> DIM('a) = 1"
       
 11320 proof
       
 11321   assume "bounded {x. a \<bullet> x = 0}"
       
 11322   then have "aff_dim {x. a \<bullet> x = 0} \<le> 0"
       
 11323     by (simp add: affine_bounded_eq_lowdim affine_hyperplane)
       
 11324   with assms show "DIM('a) = 1"
       
 11325     by (simp add: le_Suc_eq aff_dim_hyperplane)
       
 11326 next
       
 11327   assume "DIM('a) = 1"
       
 11328   then show "bounded {x. a \<bullet> x = 0}"
       
 11329     by (simp add: aff_dim_hyperplane affine_bounded_eq_lowdim affine_hyperplane assms)
       
 11330 qed
       
 11331 
       
 11332 lemma bounded_hyperplane_eq_trivial:
       
 11333   fixes a :: "'a::euclidean_space"
       
 11334   shows "bounded {x. a \<bullet> x = r} \<longleftrightarrow> (if a = 0 then r \<noteq> 0 else DIM('a) = 1)"
       
 11335 proof (simp add: bounded_hyperplane_eq_trivial_0, clarify)
       
 11336   assume "r \<noteq> 0" "a \<noteq> 0"
       
 11337   have "aff_dim {x. y \<bullet> x = 0} = aff_dim {x. a \<bullet> x = r}" if "y \<noteq> 0" for y::'a
       
 11338     by (metis that \<open>a \<noteq> 0\<close> aff_dim_hyperplane)
       
 11339   then show "bounded {x. a \<bullet> x = r} = (DIM('a) = Suc 0)"
       
 11340     by (metis One_nat_def \<open>a \<noteq> 0\<close> affine_bounded_eq_lowdim affine_hyperplane bounded_hyperplane_eq_trivial_0)
       
 11341 qed
       
 11342 
       
 11343 subsection\<open>General case without assuming closure and getting non-strict separation\<close>
       
 11344 
       
 11345 proposition separating_hyperplane_closed_point_inset:
       
 11346   fixes S :: "'a::euclidean_space set"
       
 11347   assumes "convex S" "closed S" "S \<noteq> {}" "z \<notin> S"
       
 11348   obtains a b where "a \<in> S" "(a - z) \<bullet> z < b" "\<And>x. x \<in> S \<Longrightarrow> b < (a - z) \<bullet> x"
       
 11349 proof -
       
 11350   obtain y where "y \<in> S" and y: "\<And>u. u \<in> S \<Longrightarrow> dist z y \<le> dist z u"
       
 11351     using distance_attains_inf [of S z] assms by auto
       
 11352   then have *: "(y - z) \<bullet> z < (y - z) \<bullet> z + (norm (y - z))\<^sup>2 / 2"
       
 11353     using \<open>y \<in> S\<close> \<open>z \<notin> S\<close> by auto
       
 11354   show ?thesis
       
 11355   proof (rule that [OF \<open>y \<in> S\<close> *])
       
 11356     fix x
       
 11357     assume "x \<in> S"
       
 11358     have yz: "0 < (y - z) \<bullet> (y - z)"
       
 11359       using \<open>y \<in> S\<close> \<open>z \<notin> S\<close> by auto
       
 11360     { assume 0: "0 < ((z - y) \<bullet> (x - y))"
       
 11361       with any_closest_point_dot [OF \<open>convex S\<close> \<open>closed S\<close>]
       
 11362       have False
       
 11363         using y \<open>x \<in> S\<close> \<open>y \<in> S\<close> not_less by blast
       
 11364     }
       
 11365     then have "0 \<le> ((y - z) \<bullet> (x - y))"
       
 11366       by (force simp: not_less inner_diff_left)
       
 11367     with yz have "0 < 2 * ((y - z) \<bullet> (x - y)) + (y - z) \<bullet> (y - z)"
       
 11368       by (simp add: algebra_simps)
       
 11369     then show "(y - z) \<bullet> z + (norm (y - z))\<^sup>2 / 2 < (y - z) \<bullet> x"
       
 11370       by (simp add: field_simps inner_diff_left inner_diff_right dot_square_norm [symmetric])
       
 11371   qed
       
 11372 qed
       
 11373 
       
 11374 lemma separating_hyperplane_closed_0_inset:
       
 11375   fixes S :: "'a::euclidean_space set"
       
 11376   assumes "convex S" "closed S" "S \<noteq> {}" "0 \<notin> S"
       
 11377   obtains a b where "a \<in> S" "a \<noteq> 0" "0 < b" "\<And>x. x \<in> S \<Longrightarrow> a \<bullet> x > b"
       
 11378 using separating_hyperplane_closed_point_inset [OF assms]
       
 11379 by simp (metis \<open>0 \<notin> S\<close>)
       
 11380 
       
 11381 
       
 11382 proposition separating_hyperplane_set_0_inspan:
       
 11383   fixes S :: "'a::euclidean_space set"
       
 11384   assumes "convex S" "S \<noteq> {}" "0 \<notin> S"
       
 11385   obtains a where "a \<in> span S" "a \<noteq> 0" "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> a \<bullet> x"
       
 11386 proof -
       
 11387   define k where [abs_def]: "k c = {x. 0 \<le> c \<bullet> x}" for c :: 'a
       
 11388   have *: "span S \<inter> frontier (cball 0 1) \<inter> \<Inter>f' \<noteq> {}"
       
 11389           if f': "finite f'" "f' \<subseteq> k ` S" for f'
       
 11390   proof -
       
 11391     obtain C where "C \<subseteq> S" "finite C" and C: "f' = k ` C"
       
 11392       using finite_subset_image [OF f'] by blast
       
 11393     obtain a where "a \<in> S" "a \<noteq> 0"
       
 11394       using \<open>S \<noteq> {}\<close> \<open>0 \<notin> S\<close> ex_in_conv by blast
       
 11395     then have "norm (a /\<^sub>R (norm a)) = 1"
       
 11396       by simp
       
 11397     moreover have "a /\<^sub>R (norm a) \<in> span S"
       
 11398       by (simp add: \<open>a \<in> S\<close> span_mul span_superset)
       
 11399     ultimately have ass: "a /\<^sub>R (norm a) \<in> span S \<inter> sphere 0 1"
       
 11400       by simp
       
 11401     show ?thesis
       
 11402     proof (cases "C = {}")
       
 11403       case True with C ass show ?thesis
       
 11404         by auto
       
 11405     next
       
 11406       case False
       
 11407       have "closed (convex hull C)"
       
 11408         using \<open>finite C\<close> compact_eq_bounded_closed finite_imp_compact_convex_hull by auto
       
 11409       moreover have "convex hull C \<noteq> {}"
       
 11410         by (simp add: False)
       
 11411       moreover have "0 \<notin> convex hull C"
       
 11412         by (metis \<open>C \<subseteq> S\<close> \<open>convex S\<close> \<open>0 \<notin> S\<close> convex_hull_subset hull_same insert_absorb insert_subset)
       
 11413       ultimately obtain a b
       
 11414             where "a \<in> convex hull C" "a \<noteq> 0" "0 < b"
       
 11415                   and ab: "\<And>x. x \<in> convex hull C \<Longrightarrow> a \<bullet> x > b"
       
 11416         using separating_hyperplane_closed_0_inset by blast
       
 11417       then have "a \<in> S"
       
 11418         by (metis \<open>C \<subseteq> S\<close> assms(1) subsetCE subset_hull)
       
 11419       moreover have "norm (a /\<^sub>R (norm a)) = 1"
       
 11420         using \<open>a \<noteq> 0\<close> by simp
       
 11421       moreover have "a /\<^sub>R (norm a) \<in> span S"
       
 11422         by (simp add: \<open>a \<in> S\<close> span_mul span_superset)
       
 11423       ultimately have ass: "a /\<^sub>R (norm a) \<in> span S \<inter> sphere 0 1"
       
 11424         by simp
       
 11425       have aa: "a /\<^sub>R (norm a) \<in> (\<Inter>c\<in>C. {x. 0 \<le> c \<bullet> x})"
       
 11426         apply (clarsimp simp add: divide_simps)
       
 11427         using ab \<open>0 < b\<close>
       
 11428         by (metis hull_inc inner_commute less_eq_real_def less_trans)
       
 11429       show ?thesis
       
 11430         apply (simp add: C k_def)
       
 11431         using ass aa Int_iff empty_iff by blast
       
 11432     qed
       
 11433   qed
       
 11434   have "(span S \<inter> frontier(cball 0 1)) \<inter> (\<Inter> (k ` S)) \<noteq> {}"
       
 11435     apply (rule compact_imp_fip)
       
 11436     apply (blast intro: compact_cball)
       
 11437     using closed_halfspace_ge k_def apply blast
       
 11438     apply (metis *)
       
 11439     done
       
 11440   then show ?thesis
       
 11441     unfolding set_eq_iff k_def
       
 11442     by simp (metis inner_commute norm_eq_zero that zero_neq_one)
       
 11443 qed
       
 11444 
       
 11445 
       
 11446 lemma separating_hyperplane_set_point_inaff:
       
 11447   fixes S :: "'a::euclidean_space set"
       
 11448   assumes "convex S" "S \<noteq> {}" and zno: "z \<notin> S"
       
 11449   obtains a b where "(z + a) \<in> affine hull (insert z S)"
       
 11450                 and "a \<noteq> 0" and "a \<bullet> z \<le> b"
       
 11451                 and "\<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<ge> b"
       
 11452 proof -
       
 11453 from separating_hyperplane_set_0_inspan [of "image (\<lambda>x. -z + x) S"]
       
 11454   have "convex (op + (- z) ` S)"
       
 11455     by (simp add: \<open>convex S\<close>)
       
 11456   moreover have "op + (- z) ` S \<noteq> {}"
       
 11457     by (simp add: \<open>S \<noteq> {}\<close>)
       
 11458   moreover have "0 \<notin> op + (- z) ` S"
       
 11459     using zno by auto
       
 11460   ultimately obtain a where "a \<in> span (op + (- z) ` S)" "a \<noteq> 0"
       
 11461                   and a:  "\<And>x. x \<in> (op + (- z) ` S) \<Longrightarrow> 0 \<le> a \<bullet> x"
       
 11462     using separating_hyperplane_set_0_inspan [of "image (\<lambda>x. -z + x) S"]
       
 11463     by blast
       
 11464   then have szx: "\<And>x. x \<in> S \<Longrightarrow> a \<bullet> z \<le> a \<bullet> x"
       
 11465     by (metis (no_types, lifting) imageI inner_minus_right inner_right_distrib minus_add neg_le_0_iff_le neg_le_iff_le real_add_le_0_iff)
       
 11466   show ?thesis
       
 11467     apply (rule_tac a=a and b = "a  \<bullet> z" in that, simp_all)
       
 11468     using \<open>a \<in> span (op + (- z) ` S)\<close> affine_hull_insert_span_gen apply blast
       
 11469     apply (simp_all add: \<open>a \<noteq> 0\<close> szx)
       
 11470     done
       
 11471 qed
       
 11472 
       
 11473 proposition supporting_hyperplane_rel_boundary:
       
 11474   fixes S :: "'a::euclidean_space set"
       
 11475   assumes "convex S" "x \<in> S" and xno: "x \<notin> rel_interior S"
       
 11476   obtains a where "a \<noteq> 0"
       
 11477               and "\<And>y. y \<in> S \<Longrightarrow> a \<bullet> x \<le> a \<bullet> y"
       
 11478               and "\<And>y. y \<in> rel_interior S \<Longrightarrow> a \<bullet> x < a \<bullet> y"
       
 11479 proof -
       
 11480   obtain a b where aff: "(x + a) \<in> affine hull (insert x (rel_interior S))"
       
 11481                   and "a \<noteq> 0" and "a \<bullet> x \<le> b"
       
 11482                   and ageb: "\<And>u. u \<in> (rel_interior S) \<Longrightarrow> a \<bullet> u \<ge> b"
       
 11483     using separating_hyperplane_set_point_inaff [of "rel_interior S" x] assms
       
 11484     by (auto simp: rel_interior_eq_empty convex_rel_interior)
       
 11485   have le_ay: "a \<bullet> x \<le> a \<bullet> y" if "y \<in> S" for y
       
 11486   proof -
       
 11487     have con: "continuous_on (closure (rel_interior S)) (op \<bullet> a)"
       
 11488       by (rule continuous_intros continuous_on_subset | blast)+
       
 11489     have y: "y \<in> closure (rel_interior S)"
       
 11490       using \<open>convex S\<close> closure_def convex_closure_rel_interior \<open>y \<in> S\<close>
       
 11491       by fastforce
       
 11492     show ?thesis
       
 11493       using continuous_ge_on_closure [OF con y] ageb \<open>a \<bullet> x \<le> b\<close>
       
 11494       by fastforce
       
 11495   qed
       
 11496   have 3: "a \<bullet> x < a \<bullet> y" if "y \<in> rel_interior S" for y
       
 11497   proof -
       
 11498     obtain e where "0 < e" "y \<in> S" and e: "cball y e \<inter> affine hull S \<subseteq> S"
       
 11499       using \<open>y \<in> rel_interior S\<close> by (force simp: rel_interior_cball)
       
 11500     define y' where "y' = y - (e / norm a) *\<^sub>R ((x + a) - x)"
       
 11501     have "y' \<in> cball y e"
       
 11502       unfolding y'_def using \<open>0 < e\<close> by force
       
 11503     moreover have "y' \<in> affine hull S"
       
 11504       unfolding y'_def
       
 11505       by (metis \<open>x \<in> S\<close> \<open>y \<in> S\<close> \<open>convex S\<close> aff affine_affine_hull hull_redundant
       
 11506                 rel_interior_same_affine_hull hull_inc mem_affine_3_minus2)
       
 11507     ultimately have "y' \<in> S"
       
 11508       using e by auto
       
 11509     have "a \<bullet> x \<le> a \<bullet> y"
       
 11510       using le_ay \<open>a \<noteq> 0\<close> \<open>y \<in> S\<close> by blast
       
 11511     moreover have "a \<bullet> x \<noteq> a \<bullet> y"
       
 11512       using le_ay [OF \<open>y' \<in> S\<close>] \<open>a \<noteq> 0\<close>
       
 11513       apply (simp add: y'_def inner_diff dot_square_norm power2_eq_square)
       
 11514       by (metis \<open>0 < e\<close> add_le_same_cancel1 inner_commute inner_real_def inner_zero_left le_diff_eq norm_le_zero_iff real_mult_le_cancel_iff2)
       
 11515     ultimately show ?thesis by force
       
 11516   qed
       
 11517   show ?thesis
       
 11518     by (rule that [OF \<open>a \<noteq> 0\<close> le_ay 3])
       
 11519 qed
       
 11520 
       
 11521 lemma supporting_hyperplane_relative_frontier:
       
 11522   fixes S :: "'a::euclidean_space set"
       
 11523   assumes "convex S" "x \<in> closure S" "x \<notin> rel_interior S"
       
 11524   obtains a where "a \<noteq> 0"
       
 11525               and "\<And>y. y \<in> closure S \<Longrightarrow> a \<bullet> x \<le> a \<bullet> y"
       
 11526               and "\<And>y. y \<in> rel_interior S \<Longrightarrow> a \<bullet> x < a \<bullet> y"
       
 11527 using supporting_hyperplane_rel_boundary [of "closure S" x]
       
 11528 by (metis assms convex_closure convex_rel_interior_closure)
       
 11529 
       
 11530 subsection\<open> Some results on decomposing convex hulls, e.g. simplicial subdivision\<close>
       
 11531 
       
 11532 lemma
       
 11533   fixes s :: "'a::euclidean_space set"
       
 11534   assumes "~ (affine_dependent(s \<union> t))"
       
 11535     shows convex_hull_Int_subset: "convex hull s \<inter> convex hull t \<subseteq> convex hull (s \<inter> t)" (is ?C)
       
 11536       and affine_hull_Int_subset: "affine hull s \<inter> affine hull t \<subseteq> affine hull (s \<inter> t)" (is ?A)
       
 11537 proof -
       
 11538   have [simp]: "finite s" "finite t"
       
 11539     using aff_independent_finite assms by blast+
       
 11540     have "setsum u (s \<inter> t) = 1 \<and>
       
 11541           (\<Sum>v\<in>s \<inter> t. u v *\<^sub>R v) = (\<Sum>v\<in>s. u v *\<^sub>R v)"
       
 11542       if [simp]:  "setsum u s = 1"
       
 11543                  "setsum v t = 1"
       
 11544          and eq: "(\<Sum>x\<in>t. v x *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x)" for u v
       
 11545     proof -
       
 11546     define f where "f x = (if x \<in> s then u x else 0) - (if x \<in> t then v x else 0)" for x
       
 11547     have "setsum f (s \<union> t) = 0"
       
 11548       apply (simp add: f_def setsum_Un setsum_subtractf)
       
 11549       apply (simp add: setsum.inter_restrict [symmetric] Int_commute)
       
 11550       done
       
 11551     moreover have "(\<Sum>x\<in>(s \<union> t). f x *\<^sub>R x) = 0"
       
 11552       apply (simp add: f_def setsum_Un scaleR_left_diff_distrib setsum_subtractf)
       
 11553       apply (simp add: if_smult setsum.inter_restrict [symmetric] Int_commute eq
       
 11554           cong del: if_weak_cong)
       
 11555       done
       
 11556     ultimately have "\<And>v. v \<in> s \<union> t \<Longrightarrow> f v = 0"
       
 11557       using aff_independent_finite assms unfolding affine_dependent_explicit
       
 11558       by blast
       
 11559     then have u [simp]: "\<And>x. x \<in> s \<Longrightarrow> u x = (if x \<in> t then v x else 0)"
       
 11560       by (simp add: f_def) presburger
       
 11561     have "setsum u (s \<inter> t) = setsum u s"
       
 11562       by (simp add: setsum.inter_restrict)
       
 11563     then have "setsum u (s \<inter> t) = 1"
       
 11564       using that by linarith
       
 11565     moreover have "(\<Sum>v\<in>s \<inter> t. u v *\<^sub>R v) = (\<Sum>v\<in>s. u v *\<^sub>R v)"
       
 11566       by (auto simp: if_smult setsum.inter_restrict intro: setsum.cong)
       
 11567     ultimately show ?thesis
       
 11568       by force
       
 11569     qed
       
 11570     then show ?A ?C
       
 11571       by (auto simp: convex_hull_finite affine_hull_finite)
       
 11572 qed
       
 11573 
       
 11574 
       
 11575 proposition affine_hull_Int:
       
 11576   fixes s :: "'a::euclidean_space set"
       
 11577   assumes "~ (affine_dependent(s \<union> t))"
       
 11578     shows "affine hull (s \<inter> t) = affine hull s \<inter> affine hull t"
       
 11579 apply (rule subset_antisym)
       
 11580 apply (simp add: hull_mono)
       
 11581 by (simp add: affine_hull_Int_subset assms)
       
 11582 
       
 11583 proposition convex_hull_Int:
       
 11584   fixes s :: "'a::euclidean_space set"
       
 11585   assumes "~ (affine_dependent(s \<union> t))"
       
 11586     shows "convex hull (s \<inter> t) = convex hull s \<inter> convex hull t"
       
 11587 apply (rule subset_antisym)
       
 11588 apply (simp add: hull_mono)
       
 11589 by (simp add: convex_hull_Int_subset assms)
       
 11590 
       
 11591 proposition
       
 11592   fixes s :: "'a::euclidean_space set set"
       
 11593   assumes "~ (affine_dependent (\<Union>s))"
       
 11594     shows affine_hull_Inter: "affine hull (\<Inter>s) = (\<Inter>t\<in>s. affine hull t)" (is "?A")
       
 11595       and convex_hull_Inter: "convex hull (\<Inter>s) = (\<Inter>t\<in>s. convex hull t)" (is "?C")
       
 11596 proof -
       
 11597   have "finite s"
       
 11598     using aff_independent_finite assms finite_UnionD by blast
       
 11599   then have "?A \<and> ?C" using assms
       
 11600   proof (induction s rule: finite_induct)
       
 11601     case empty then show ?case by auto
       
 11602   next
       
 11603     case (insert t F)
       
 11604     then show ?case
       
 11605     proof (cases "F={}")
       
 11606       case True then show ?thesis by simp
       
 11607     next
       
 11608       case False
       
 11609       with "insert.prems" have [simp]: "\<not> affine_dependent (t \<union> \<Inter>F)"
       
 11610         by (auto intro: affine_dependent_subset)
       
 11611       have [simp]: "\<not> affine_dependent (\<Union>F)"
       
 11612         using affine_independent_subset insert.prems by fastforce
       
 11613       show ?thesis
       
 11614         by (simp add: affine_hull_Int convex_hull_Int insert.IH)
       
 11615     qed
       
 11616   qed
       
 11617   then show "?A" "?C"
       
 11618     by auto
       
 11619 qed
       
 11620 
       
 11621 lemma affine_hull_finite_intersection_hyperplanes:
       
 11622   fixes s :: "'a::euclidean_space set"
       
 11623   obtains f where
       
 11624      "finite f"
       
 11625      "of_nat (card f) + aff_dim s = DIM('a)"
       
 11626      "affine hull s = \<Inter>f"
       
 11627      "\<And>h. h \<in> f \<Longrightarrow> \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x = b}"
       
 11628 proof -
       
 11629   obtain b where "b \<subseteq> s"
       
 11630              and indb: "\<not> affine_dependent b"
       
 11631              and eq: "affine hull s = affine hull b"
       
 11632     using affine_basis_exists by blast
       
 11633   obtain c where indc: "\<not> affine_dependent c" and "b \<subseteq> c"
       
 11634              and affc: "affine hull c = UNIV"
       
 11635     by (metis extend_to_affine_basis affine_UNIV hull_same indb subset_UNIV)
       
 11636   then have "finite c"
       
 11637     by (simp add: aff_independent_finite)
       
 11638   then have fbc: "finite b" "card b \<le> card c"
       
 11639     using \<open>b \<subseteq> c\<close> infinite_super by (auto simp: card_mono)
       
 11640   have imeq: "(\<lambda>x. affine hull x) ` ((\<lambda>a. c - {a}) ` (c - b)) = ((\<lambda>a. affine hull (c - {a})) ` (c - b))"
       
 11641     by blast
       
 11642   have card1: "card ((\<lambda>a. affine hull (c - {a})) ` (c - b)) = card (c - b)"
       
 11643     apply (rule card_image [OF inj_onI])
       
 11644     by (metis Diff_eq_empty_iff Diff_iff indc affine_dependent_def hull_subset insert_iff)
       
 11645   have card2: "(card (c - b)) + aff_dim s = DIM('a)"
       
 11646   proof -
       
 11647     have aff: "aff_dim (UNIV::'a set) = aff_dim c"
       
 11648       by (metis aff_dim_affine_hull affc)
       
 11649     have "aff_dim b = aff_dim s"
       
 11650       by (metis (no_types) aff_dim_affine_hull eq)
       
 11651     then have "int (card b) = 1 + aff_dim s"
       
 11652       by (simp add: aff_dim_affine_independent indb)
       
 11653     then show ?thesis
       
 11654       using fbc aff
       
 11655       by (simp add: \<open>\<not> affine_dependent c\<close> \<open>b \<subseteq> c\<close> aff_dim_affine_independent aff_dim_UNIV card_Diff_subset of_nat_diff)
       
 11656   qed
       
 11657   show ?thesis
       
 11658   proof (cases "c = b")
       
 11659     case True show ?thesis
       
 11660       apply (rule_tac f="{}" in that)
       
 11661       using True affc
       
 11662       apply (simp_all add: eq [symmetric])
       
 11663       by (metis aff_dim_UNIV aff_dim_affine_hull)
       
 11664   next
       
 11665     case False
       
 11666     have ind: "\<not> affine_dependent (\<Union>a\<in>c - b. c - {a})"
       
 11667       by (rule affine_independent_subset [OF indc]) auto
       
 11668     have affeq: "affine hull s = (\<Inter>x\<in>(\<lambda>a. c - {a}) ` (c - b). affine hull x)"
       
 11669       using \<open>b \<subseteq> c\<close> False
       
 11670       apply (subst affine_hull_Inter [OF ind, symmetric])
       
 11671       apply (simp add: eq double_diff)
       
 11672       done
       
 11673     have *: "1 + aff_dim (c - {t}) = int (DIM('a))"
       
 11674             if t: "t \<in> c" for t
       
 11675     proof -
       
 11676       have "insert t c = c"
       
 11677         using t by blast
       
 11678       then show ?thesis
       
 11679         by (metis (full_types) add.commute aff_dim_affine_hull aff_dim_insert aff_dim_UNIV affc affine_dependent_def indc insert_Diff_single t)
       
 11680     qed
       
 11681     show ?thesis
       
 11682       apply (rule_tac f = "(\<lambda>x. affine hull x) ` ((\<lambda>a. c - {a}) ` (c - b))" in that)
       
 11683          using \<open>finite c\<close> apply blast
       
 11684         apply (simp add: imeq card1 card2)
       
 11685       apply (simp add: affeq, clarify)
       
 11686       apply (metis DIM_positive One_nat_def Suc_leI add_diff_cancel_left' of_nat_1 aff_dim_eq_hyperplane of_nat_diff *)
       
 11687       done
       
 11688   qed
       
 11689 qed
       
 11690 
       
 11691 subsection\<open>Misc results about span\<close>
       
 11692 
       
 11693 lemma eq_span_insert_eq:
       
 11694   assumes "(x - y) \<in> span S"
       
 11695     shows "span(insert x S) = span(insert y S)"
       
 11696 proof -
       
 11697   have *: "span(insert x S) \<subseteq> span(insert y S)" if "(x - y) \<in> span S" for x y
       
 11698   proof -
       
 11699     have 1: "(r *\<^sub>R x - r *\<^sub>R y) \<in> span S" for r
       
 11700       by (metis real_vector.scale_right_diff_distrib span_mul that)
       
 11701     have 2: "(z - k *\<^sub>R y) - k *\<^sub>R (x - y) = z - k *\<^sub>R x" for  z k
       
 11702       by (simp add: real_vector.scale_right_diff_distrib)
       
 11703   show ?thesis
       
 11704     apply (clarsimp simp add: span_breakdown_eq)
       
 11705     by (metis 1 2 diff_add_cancel real_vector.scale_right_diff_distrib span_add_eq)
       
 11706   qed
       
 11707   show ?thesis
       
 11708     apply (intro subset_antisym * assms)
       
 11709     using assms subspace_neg subspace_span minus_diff_eq by force
       
 11710 qed
       
 11711 
       
 11712 lemma dim_psubset:
       
 11713     fixes S :: "'a :: euclidean_space set"
       
 11714     shows "span S \<subset> span T \<Longrightarrow> dim S < dim T"
       
 11715 by (metis (no_types, hide_lams) dim_span less_le not_le subspace_dim_equal subspace_span)
       
 11716 
       
 11717 
       
 11718 lemma basis_subspace_exists:
       
 11719   fixes S :: "'a::euclidean_space set"
       
 11720   shows
       
 11721    "subspace S
       
 11722         \<Longrightarrow> \<exists>b. finite b \<and> b \<subseteq> S \<and>
       
 11723                 independent b \<and> span b = S \<and> card b = dim S"
       
 11724 by (metis span_subspace basis_exists independent_imp_finite)
       
 11725 
       
 11726 lemma affine_hyperplane_sums_eq_UNIV_0:
       
 11727   fixes S :: "'a :: euclidean_space set"
       
 11728   assumes "affine S"
       
 11729      and "0 \<in> S" and "w \<in> S"
       
 11730      and "a \<bullet> w \<noteq> 0"
       
 11731    shows "{x + y| x y. x \<in> S \<and> a \<bullet> y = 0} = UNIV"
       
 11732 proof -
       
 11733   have "subspace S"
       
 11734     by (simp add: assms subspace_affine)
       
 11735   have span1: "span {y. a \<bullet> y = 0} \<subseteq> span {x + y |x y. x \<in> S \<and> a \<bullet> y = 0}"
       
 11736     apply (rule span_mono)
       
 11737     using \<open>0 \<in> S\<close> add.left_neutral by force
       
 11738   have "w \<notin> span {y. a \<bullet> y = 0}"
       
 11739     using \<open>a \<bullet> w \<noteq> 0\<close> span_induct subspace_hyperplane by auto
       
 11740   moreover have "w \<in> span {x + y |x y. x \<in> S \<and> a \<bullet> y = 0}"
       
 11741     using \<open>w \<in> S\<close>
       
 11742     by (metis (mono_tags, lifting) inner_zero_right mem_Collect_eq pth_d span_superset)
       
 11743   ultimately have span2: "span {y. a \<bullet> y = 0} \<noteq> span {x + y |x y. x \<in> S \<and> a \<bullet> y = 0}"
       
 11744     by blast
       
 11745   have "a \<noteq> 0" using assms inner_zero_left by blast
       
 11746   then have "DIM('a) - 1 = dim {y. a \<bullet> y = 0}"
       
 11747     by (simp add: dim_hyperplane)
       
 11748   also have "... < dim {x + y |x y. x \<in> S \<and> a \<bullet> y = 0}"
       
 11749     using span1 span2 by (blast intro: dim_psubset)
       
 11750   finally have DIM_lt: "DIM('a) - 1 < dim {x + y |x y. x \<in> S \<and> a \<bullet> y = 0}" .
       
 11751   have subs: "subspace {x + y| x y. x \<in> S \<and> a \<bullet> y = 0}"
       
 11752     using subspace_sums [OF \<open>subspace S\<close> subspace_hyperplane] by simp
       
 11753   moreover have "span {x + y| x y. x \<in> S \<and> a \<bullet> y = 0} = UNIV"
       
 11754     apply (rule dim_eq_full [THEN iffD1])
       
 11755     apply (rule antisym [OF dim_subset_UNIV])
       
 11756     using DIM_lt apply simp
       
 11757     done
       
 11758   ultimately show ?thesis
       
 11759     by (simp add: subs) (metis (lifting) span_eq subs)
       
 11760 qed
       
 11761 
       
 11762 proposition affine_hyperplane_sums_eq_UNIV:
       
 11763   fixes S :: "'a :: euclidean_space set"
       
 11764   assumes "affine S"
       
 11765       and "S \<inter> {v. a \<bullet> v = b} \<noteq> {}"
       
 11766       and "S - {v. a \<bullet> v = b} \<noteq> {}"
       
 11767     shows "{x + y| x y. x \<in> S \<and> a \<bullet> y = b} = UNIV"
       
 11768 proof (cases "a = 0")
       
 11769   case True with assms show ?thesis
       
 11770     by (auto simp: if_splits)
       
 11771 next
       
 11772   case False
       
 11773   obtain c where "c \<in> S" and c: "a \<bullet> c = b"
       
 11774     using assms by force
       
 11775   with affine_diffs_subspace [OF \<open>affine S\<close>]
       
 11776   have "subspace (op + (- c) ` S)" by blast
       
 11777   then have aff: "affine (op + (- c) ` S)"
       
 11778     by (simp add: subspace_imp_affine)
       
 11779   have 0: "0 \<in> op + (- c) ` S"
       
 11780     by (simp add: \<open>c \<in> S\<close>)
       
 11781   obtain d where "d \<in> S" and "a \<bullet> d \<noteq> b" and dc: "d-c \<in> op + (- c) ` S"
       
 11782     using assms by auto
       
 11783   then have adc: "a \<bullet> (d - c) \<noteq> 0"
       
 11784     by (simp add: c inner_diff_right)
       
 11785   let ?U = "op + (c+c) ` {x + y |x y. x \<in> op + (- c) ` S \<and> a \<bullet> y = 0}"
       
 11786   have "u + v \<in> op + (c + c) ` {x + v |x v. x \<in> op + (- c) ` S \<and> a \<bullet> v = 0}"
       
 11787               if "u \<in> S" "b = a \<bullet> v" for u v
       
 11788     apply (rule_tac x="u+v-c-c" in image_eqI)
       
 11789     apply (simp_all add: algebra_simps)
       
 11790     apply (rule_tac x="u-c" in exI)
       
 11791     apply (rule_tac x="v-c" in exI)
       
 11792     apply (simp add: algebra_simps that c)
       
 11793     done
       
 11794   moreover have "\<lbrakk>a \<bullet> v = 0; u \<in> S\<rbrakk>
       
 11795        \<Longrightarrow> \<exists>x ya. v + (u + c) = x + ya \<and> x \<in> S \<and> a \<bullet> ya = b" for v u
       
 11796     by (metis add.left_commute c inner_right_distrib pth_d)
       
 11797   ultimately have "{x + y |x y. x \<in> S \<and> a \<bullet> y = b} = ?U"
       
 11798     by (fastforce simp: algebra_simps)
       
 11799   also have "... = op + (c+c) ` UNIV"
       
 11800     by (simp add: affine_hyperplane_sums_eq_UNIV_0 [OF aff 0 dc adc])
       
 11801   also have "... = UNIV"
       
 11802     by (simp add: translation_UNIV)
       
 11803   finally show ?thesis .
       
 11804 qed
       
 11805 
       
 11806 proposition dim_sums_Int:
       
 11807     fixes S :: "'a :: euclidean_space set"
       
 11808   assumes "subspace S" "subspace T"
       
 11809   shows "dim {x + y |x y. x \<in> S \<and> y \<in> T} + dim(S \<inter> T) = dim S + dim T"
       
 11810 proof -
       
 11811   obtain B where B: "B \<subseteq> S \<inter> T" "S \<inter> T \<subseteq> span B"
       
 11812              and indB: "independent B"
       
 11813              and cardB: "card B = dim (S \<inter> T)"
       
 11814     using basis_exists by blast
       
 11815   then obtain C D where "B \<subseteq> C" "C \<subseteq> S" "independent C" "S \<subseteq> span C"
       
 11816                     and "B \<subseteq> D" "D \<subseteq> T" "independent D" "T \<subseteq> span D"
       
 11817     using maximal_independent_subset_extend
       
 11818     by (metis Int_subset_iff \<open>B \<subseteq> S \<inter> T\<close> indB)
       
 11819   then have "finite B" "finite C" "finite D"
       
 11820     by (simp_all add: independent_imp_finite indB independent_bound)
       
 11821   have Beq: "B = C \<inter> D"
       
 11822     apply (rule sym)
       
 11823     apply (rule spanning_subset_independent)
       
 11824     using \<open>B \<subseteq> C\<close> \<open>B \<subseteq> D\<close> apply blast
       
 11825     apply (meson \<open>independent C\<close> independent_mono inf.cobounded1)
       
 11826     using B \<open>C \<subseteq> S\<close> \<open>D \<subseteq> T\<close> apply auto
       
 11827     done
       
 11828   then have Deq: "D = B \<union> (D - C)"
       
 11829     by blast
       
 11830   have CUD: "C \<union> D \<subseteq> {x + y |x y. x \<in> S \<and> y \<in> T}"
       
 11831     apply safe
       
 11832     apply (metis add.right_neutral subsetCE \<open>C \<subseteq> S\<close> \<open>subspace T\<close> set_eq_subset span_0 span_minimal)
       
 11833     apply (metis add.left_neutral subsetCE \<open>D \<subseteq> T\<close> \<open>subspace S\<close> set_eq_subset span_0 span_minimal)
       
 11834     done
       
 11835   have "a v = 0" if 0: "(\<Sum>v\<in>C. a v *\<^sub>R v) + (\<Sum>v\<in>D - C. a v *\<^sub>R v) = 0"
       
 11836                  and v: "v \<in> C \<union> (D-C)" for a v
       
 11837   proof -
       
 11838     have eq: "(\<Sum>v\<in>D - C. a v *\<^sub>R v) = - (\<Sum>v\<in>C. a v *\<^sub>R v)"
       
 11839       using that add_eq_0_iff by blast
       
 11840     have "(\<Sum>v\<in>D - C. a v *\<^sub>R v) \<in> S"
       
 11841       apply (subst eq)
       
 11842       apply (rule subspace_neg [OF \<open>subspace S\<close>])
       
 11843       apply (rule subspace_setsum [OF \<open>subspace S\<close>])
       
 11844       by (meson subsetCE subspace_mul \<open>C \<subseteq> S\<close> \<open>subspace S\<close>)
       
 11845     moreover have "(\<Sum>v\<in>D - C. a v *\<^sub>R v) \<in> T"
       
 11846       apply (rule subspace_setsum [OF \<open>subspace T\<close>])
       
 11847       by (meson DiffD1 \<open>D \<subseteq> T\<close> \<open>subspace T\<close> subset_eq subspace_def)
       
 11848     ultimately have "(\<Sum>v \<in> D-C. a v *\<^sub>R v) \<in> span B"
       
 11849       using B by blast
       
 11850     then obtain e where e: "(\<Sum>v\<in>B. e v *\<^sub>R v) = (\<Sum>v \<in> D-C. a v *\<^sub>R v)"
       
 11851       using span_finite [OF \<open>finite B\<close>] by blast
       
 11852     have "\<And>c v. \<lbrakk>(\<Sum>v\<in>C. c v *\<^sub>R v) = 0; v \<in> C\<rbrakk> \<Longrightarrow> c v = 0"
       
 11853       using independent_explicit \<open>independent C\<close> by blast
       
 11854     define cc where "cc x = (if x \<in> B then a x + e x else a x)" for x
       
 11855     have [simp]: "C \<inter> B = B" "D \<inter> B = B" "C \<inter> - B = C-D" "B \<inter> (D - C) = {}"
       
 11856       using \<open>B \<subseteq> C\<close> \<open>B \<subseteq> D\<close> Beq by blast+
       
 11857     have f2: "(\<Sum>v\<in>C \<inter> D. e v *\<^sub>R v) = (\<Sum>v\<in>D - C. a v *\<^sub>R v)"
       
 11858       using Beq e by presburger
       
 11859     have f3: "(\<Sum>v\<in>C \<union> D. a v *\<^sub>R v) = (\<Sum>v\<in>C - D. a v *\<^sub>R v) + (\<Sum>v\<in>D - C. a v *\<^sub>R v) + (\<Sum>v\<in>C \<inter> D. a v *\<^sub>R v)"
       
 11860       using \<open>finite C\<close> \<open>finite D\<close> setsum.union_diff2 by blast
       
 11861     have f4: "(\<Sum>v\<in>C \<union> (D - C). a v *\<^sub>R v) = (\<Sum>v\<in>C. a v *\<^sub>R v) + (\<Sum>v\<in>D - C. a v *\<^sub>R v)"
       
 11862       by (meson Diff_disjoint \<open>finite C\<close> \<open>finite D\<close> finite_Diff setsum.union_disjoint)
       
 11863     have "(\<Sum>v\<in>C. cc v *\<^sub>R v) = 0"
       
 11864       using 0 f2 f3 f4
       
 11865       apply (simp add: cc_def Beq if_smult \<open>finite C\<close> setsum.If_cases algebra_simps setsum.distrib)
       
 11866       apply (simp add: add.commute add.left_commute diff_eq)
       
 11867       done
       
 11868     then have "\<And>v. v \<in> C \<Longrightarrow> cc v = 0"
       
 11869       using independent_explicit \<open>independent C\<close> by blast
       
 11870     then have C0: "\<And>v. v \<in> C - B \<Longrightarrow> a v = 0"
       
 11871       by (simp add: cc_def Beq) meson
       
 11872     then have [simp]: "(\<Sum>x\<in>C - B. a x *\<^sub>R x) = 0"
       
 11873       by simp
       
 11874     have "(\<Sum>x\<in>C. a x *\<^sub>R x) = (\<Sum>x\<in>B. a x *\<^sub>R x)"
       
 11875     proof -
       
 11876       have "C - D = C - B"
       
 11877         using Beq by blast
       
 11878       then show ?thesis
       
 11879         using Beq \<open>(\<Sum>x\<in>C - B. a x *\<^sub>R x) = 0\<close> f3 f4 by auto
       
 11880     qed
       
 11881     with 0 have Dcc0: "(\<Sum>v\<in>D. a v *\<^sub>R v) = 0"
       
 11882       apply (subst Deq)
       
 11883       by (simp add: \<open>finite B\<close> \<open>finite D\<close> setsum_Un)
       
 11884     then have D0: "\<And>v. v \<in> D \<Longrightarrow> a v = 0"
       
 11885       using independent_explicit \<open>independent D\<close> by blast
       
 11886     show ?thesis
       
 11887       using v C0 D0 Beq by blast
       
 11888   qed
       
 11889   then have "independent (C \<union> (D - C))"
       
 11890     by (simp add: independent_explicit \<open>finite C\<close> \<open>finite D\<close> setsum_Un del: Un_Diff_cancel)
       
 11891   then have indCUD: "independent (C \<union> D)" by simp
       
 11892   have "dim (S \<inter> T) = card B"
       
 11893     by (rule dim_unique [OF B indB refl])
       
 11894   moreover have "dim S = card C"
       
 11895     by (metis \<open>C \<subseteq> S\<close> \<open>independent C\<close> \<open>S \<subseteq> span C\<close> basis_card_eq_dim)
       
 11896   moreover have "dim T = card D"
       
 11897     by (metis \<open>D \<subseteq> T\<close> \<open>independent D\<close> \<open>T \<subseteq> span D\<close> basis_card_eq_dim)
       
 11898   moreover have "dim {x + y |x y. x \<in> S \<and> y \<in> T} = card(C \<union> D)"
       
 11899     apply (rule dim_unique [OF CUD _ indCUD refl], clarify)
       
 11900     apply (meson \<open>S \<subseteq> span C\<close> \<open>T \<subseteq> span D\<close> span_add span_inc span_minimal subsetCE subspace_span sup.bounded_iff)
       
 11901     done
       
 11902   ultimately show ?thesis
       
 11903     using \<open>B = C \<inter> D\<close> [symmetric]
       
 11904     by (simp add:  \<open>independent C\<close> \<open>independent D\<close> card_Un_Int independent_finite)
       
 11905 qed
       
 11906 
       
 11907 
       
 11908 lemma aff_dim_sums_Int_0:
       
 11909   assumes "affine S"
       
 11910       and "affine T"
       
 11911       and "0 \<in> S" "0 \<in> T"
       
 11912     shows "aff_dim {x + y| x y. x \<in> S \<and> y \<in> T} = (aff_dim S + aff_dim T) - aff_dim(S \<inter> T)"
       
 11913 proof -
       
 11914   have "0 \<in> {x + y |x y. x \<in> S \<and> y \<in> T}"
       
 11915     using assms by force
       
 11916   then have 0: "0 \<in> affine hull {x + y |x y. x \<in> S \<and> y \<in> T}"
       
 11917     by (metis (lifting) hull_inc)
       
 11918   have sub: "subspace S"  "subspace T"
       
 11919     using assms by (auto simp: subspace_affine)
       
 11920   show ?thesis
       
 11921     using dim_sums_Int [OF sub] by (simp add: aff_dim_zero assms 0 hull_inc)
       
 11922 qed
       
 11923 
       
 11924 proposition aff_dim_sums_Int:
       
 11925   assumes "affine S"
       
 11926       and "affine T"
       
 11927       and "S \<inter> T \<noteq> {}"
       
 11928     shows "aff_dim {x + y| x y. x \<in> S \<and> y \<in> T} = (aff_dim S + aff_dim T) - aff_dim(S \<inter> T)"
       
 11929 proof -
       
 11930   obtain a where a: "a \<in> S" "a \<in> T" using assms by force
       
 11931   have aff: "affine (op+ (-a) ` S)"  "affine (op+ (-a) ` T)"
       
 11932     using assms by (auto simp: affine_translation [symmetric])
       
 11933   have zero: "0 \<in> (op+ (-a) ` S)"  "0 \<in> (op+ (-a) ` T)"
       
 11934     using a assms by auto
       
 11935   have [simp]: "{x + y |x y. x \<in> op + (- a) ` S \<and> y \<in> op + (- a) ` T} =
       
 11936         op + (- 2 *\<^sub>R a) ` {x + y| x y. x \<in> S \<and> y \<in> T}"
       
 11937     by (force simp: algebra_simps scaleR_2)
       
 11938   have [simp]: "op + (- a) ` S \<inter> op + (- a) ` T = op + (- a) ` (S \<inter> T)"
       
 11939     by auto
       
 11940   show ?thesis
       
 11941     using aff_dim_sums_Int_0 [OF aff zero]
       
 11942     by (auto simp: aff_dim_translation_eq)
       
 11943 qed
       
 11944 
       
 11945 lemma aff_dim_affine_Int_hyperplane:
       
 11946   fixes a :: "'a::euclidean_space"
       
 11947   assumes "affine S"
       
 11948     shows "aff_dim(S \<inter> {x. a \<bullet> x = b}) =
       
 11949              (if S \<inter> {v. a \<bullet> v = b} = {} then - 1
       
 11950               else if S \<subseteq> {v. a \<bullet> v = b} then aff_dim S
       
 11951               else aff_dim S - 1)"
       
 11952 proof (cases "a = 0")
       
 11953   case True with assms show ?thesis
       
 11954     by auto
       
 11955 next
       
 11956   case False
       
 11957   then have "aff_dim (S \<inter> {x. a \<bullet> x = b}) = aff_dim S - 1"
       
 11958             if "x \<in> S" "a \<bullet> x \<noteq> b" and non: "S \<inter> {v. a \<bullet> v = b} \<noteq> {}" for x
       
 11959   proof -
       
 11960     have [simp]: "{x + y| x y. x \<in> S \<and> a \<bullet> y = b} = UNIV"
       
 11961       using affine_hyperplane_sums_eq_UNIV [OF assms non] that  by blast
       
 11962     show ?thesis
       
 11963       using aff_dim_sums_Int [OF assms affine_hyperplane non]
       
 11964       by (simp add: of_nat_diff False)
       
 11965   qed
       
 11966   then show ?thesis
       
 11967     by (metis (mono_tags, lifting) inf.orderE aff_dim_empty_eq mem_Collect_eq subsetI)
       
 11968 qed
       
 11969 
       
 11970 lemma aff_dim_lt_full:
       
 11971   fixes S :: "'a::euclidean_space set"
       
 11972   shows "aff_dim S < DIM('a) \<longleftrightarrow> (affine hull S \<noteq> UNIV)"
       
 11973 by (metis (no_types) aff_dim_affine_hull aff_dim_le_DIM aff_dim_UNIV affine_hull_UNIV less_le)
       
 11974 
       
 11975 subsection\<open> Orthogonal bases, Gram-Schmidt process, and related theorems\<close>
       
 11976 
       
 11977 lemma span_delete_0 [simp]: "span(S - {0}) = span S"
       
 11978 proof
       
 11979   show "span (S - {0}) \<subseteq> span S"
       
 11980     by (blast intro!: span_mono)
       
 11981 next
       
 11982   have "span S \<subseteq> span(insert 0 (S - {0}))"
       
 11983     by (blast intro!: span_mono)
       
 11984   also have "... \<subseteq> span(S - {0})"
       
 11985     using span_insert_0 by blast
       
 11986   finally show "span S \<subseteq> span (S - {0})" .
       
 11987 qed
       
 11988 
       
 11989 lemma span_image_scale:
       
 11990   assumes "finite S" and nz: "\<And>x. x \<in> S \<Longrightarrow> c x \<noteq> 0"
       
 11991     shows "span ((\<lambda>x. c x *\<^sub>R x) ` S) = span S"
       
 11992 using assms
       
 11993 proof (induction S arbitrary: c)
       
 11994   case (empty c) show ?case by simp
       
 11995 next
       
 11996   case (insert x F c)
       
 11997   show ?case
       
 11998   proof (intro set_eqI iffI)
       
 11999     fix y
       
 12000       assume "y \<in> span ((\<lambda>x. c x *\<^sub>R x) ` insert x F)"
       
 12001       then show "y \<in> span (insert x F)"
       
 12002         using insert by (force simp: span_breakdown_eq)
       
 12003   next
       
 12004     fix y
       
 12005       assume "y \<in> span (insert x F)"
       
 12006       then show "y \<in> span ((\<lambda>x. c x *\<^sub>R x) ` insert x F)"
       
 12007         using insert
       
 12008         apply (clarsimp simp: span_breakdown_eq)
       
 12009         apply (rule_tac x="k / c x" in exI)
       
 12010         by simp
       
 12011   qed
       
 12012 qed
       
 12013 
       
 12014 lemma pairwise_orthogonal_independent:
       
 12015   assumes "pairwise orthogonal S" and "0 \<notin> S"
       
 12016     shows "independent S"
       
 12017 proof -
       
 12018   have 0: "\<And>x y. \<lbrakk>x \<noteq> y; x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> x \<bullet> y = 0"
       
 12019     using assms by (simp add: pairwise_def orthogonal_def)
       
 12020   have "False" if "a \<in> S" and a: "a \<in> span (S - {a})" for a
       
 12021   proof -
       
 12022     obtain T U where "T \<subseteq> S - {a}" "a = (\<Sum>v\<in>T. U v *\<^sub>R v)"
       
 12023       using a by (force simp: span_explicit)
       
 12024     then have "a \<bullet> a = a \<bullet> (\<Sum>v\<in>T. U v *\<^sub>R v)"
       
 12025       by simp
       
 12026     also have "... = 0"
       
 12027       apply (simp add: inner_setsum_right)
       
 12028       apply (rule comm_monoid_add_class.setsum.neutral)
       
 12029       by (metis "0" DiffE \<open>T \<subseteq> S - {a}\<close> mult_not_zero singletonI subsetCE \<open>a \<in> S\<close>)
       
 12030     finally show ?thesis
       
 12031       using \<open>0 \<notin> S\<close> \<open>a \<in> S\<close> by auto
       
 12032   qed
       
 12033   then show ?thesis
       
 12034     by (force simp: dependent_def)
       
 12035 qed
       
 12036 
       
 12037 lemma pairwise_orthogonal_imp_finite:
       
 12038   fixes S :: "'a::euclidean_space set"
       
 12039   assumes "pairwise orthogonal S"
       
 12040     shows "finite S"
       
 12041 proof -
       
 12042   have "independent (S - {0})"
       
 12043     apply (rule pairwise_orthogonal_independent)
       
 12044      apply (metis Diff_iff assms pairwise_def)
       
 12045     by blast
       
 12046   then show ?thesis
       
 12047     by (meson independent_imp_finite infinite_remove)
       
 12048 qed
       
 12049 
       
 12050 lemma subspace_orthogonal_to_vector: "subspace {y. orthogonal x y}"
       
 12051   by (simp add: subspace_def orthogonal_clauses)
       
 12052 
       
 12053 lemma subspace_orthogonal_to_vectors: "subspace {y. \<forall>x \<in> S. orthogonal x y}"
       
 12054   by (simp add: subspace_def orthogonal_clauses)
       
 12055 
       
 12056 lemma orthogonal_to_span:
       
 12057   assumes a: "a \<in> span S" and x: "\<And>y. y \<in> S \<Longrightarrow> orthogonal x y"
       
 12058     shows "orthogonal x a"
       
 12059 apply (rule span_induct [OF a subspace_orthogonal_to_vector])
       
 12060 apply (simp add: x)
       
 12061 done
       
 12062 
       
 12063 proposition Gram_Schmidt_step:
       
 12064   fixes S :: "'a::euclidean_space set"
       
 12065   assumes S: "pairwise orthogonal S" and x: "x \<in> span S"
       
 12066     shows "orthogonal x (a - (\<Sum>b\<in>S. (b \<bullet> a / (b \<bullet> b)) *\<^sub>R b))"
       
 12067 proof -
       
 12068   have "finite S"
       
 12069     by (simp add: S pairwise_orthogonal_imp_finite)
       
 12070   have "orthogonal (a - (\<Sum>b\<in>S. (b \<bullet> a / (b \<bullet> b)) *\<^sub>R b)) x"
       
 12071        if "x \<in> S" for x
       
 12072   proof -
       
 12073     have "a \<bullet> x = (\<Sum>y\<in>S. if y = x then y \<bullet> a else 0)"
       
 12074       by (simp add: \<open>finite S\<close> inner_commute setsum.delta that)
       
 12075     also have "... =  (\<Sum>b\<in>S. b \<bullet> a * (b \<bullet> x) / (b \<bullet> b))"
       
 12076       apply (rule setsum.cong [OF refl], simp)
       
 12077       by (meson S orthogonal_def pairwise_def that)
       
 12078    finally show ?thesis
       
 12079      by (simp add: orthogonal_def algebra_simps inner_setsum_left)
       
 12080   qed
       
 12081   then show ?thesis
       
 12082     using orthogonal_to_span orthogonal_commute x by blast
       
 12083 qed
       
 12084 
       
 12085 
       
 12086 lemma orthogonal_extension_aux:
       
 12087   fixes S :: "'a::euclidean_space set"
       
 12088   assumes "finite T" "finite S" "pairwise orthogonal S"
       
 12089     shows "\<exists>U. pairwise orthogonal (S \<union> U) \<and> span (S \<union> U) = span (S \<union> T)"
       
 12090 using assms
       
 12091 proof (induction arbitrary: S)
       
 12092   case empty then show ?case
       
 12093     by simp (metis sup_bot_right)
       
 12094 next
       
 12095   case (insert a T)
       
 12096   have 0: "\<And>x y. \<lbrakk>x \<noteq> y; x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> x \<bullet> y = 0"
       
 12097     using insert by (simp add: pairwise_def orthogonal_def)
       
 12098   define a' where "a' = a - (\<Sum>b\<in>S. (b \<bullet> a / (b \<bullet> b)) *\<^sub>R b)"
       
 12099   obtain U where orthU: "pairwise orthogonal (S \<union> insert a' U)"
       
 12100              and spanU: "span (insert a' S \<union> U) = span (insert a' S \<union> T)"
       
 12101     apply (rule exE [OF insert.IH [of "insert a' S"]])
       
 12102     apply (auto simp: Gram_Schmidt_step a'_def insert.prems orthogonal_commute pairwise_orthogonal_insert span_clauses)
       
 12103     done
       
 12104   have orthS: "\<And>x. x \<in> S \<Longrightarrow> a' \<bullet> x = 0"
       
 12105     apply (simp add: a'_def)
       
 12106     using Gram_Schmidt_step [OF \<open>pairwise orthogonal S\<close>]
       
 12107     apply (force simp: orthogonal_def inner_commute span_inc [THEN subsetD])
       
 12108     done
       
 12109   have "span (S \<union> insert a' U) = span (insert a' (S \<union> T))"
       
 12110     using spanU by simp
       
 12111   also have "... = span (insert a (S \<union> T))"
       
 12112     apply (rule eq_span_insert_eq)
       
 12113     apply (simp add: a'_def span_neg span_setsum span_clauses(1) span_mul)
       
 12114     done
       
 12115   also have "... = span (S \<union> insert a T)"
       
 12116     by simp
       
 12117   finally show ?case
       
 12118     apply (rule_tac x="insert a' U" in exI)
       
 12119     using orthU apply auto
       
 12120     done
       
 12121 qed
       
 12122 
       
 12123 
       
 12124 proposition orthogonal_extension:
       
 12125   fixes S :: "'a::euclidean_space set"
       
 12126   assumes S: "pairwise orthogonal S"
       
 12127   obtains U where "pairwise orthogonal (S \<union> U)" "span (S \<union> U) = span (S \<union> T)"
       
 12128 proof -
       
 12129   obtain B where "finite B" "span B = span T"
       
 12130     using basis_subspace_exists [of "span T"] subspace_span by auto
       
 12131   with orthogonal_extension_aux [of B S]
       
 12132   obtain U where "pairwise orthogonal (S \<union> U)" "span (S \<union> U) = span (S \<union> B)"
       
 12133     using assms pairwise_orthogonal_imp_finite by auto
       
 12134   show ?thesis
       
 12135     apply (rule_tac U=U in that)
       
 12136      apply (simp add: \<open>pairwise orthogonal (S \<union> U)\<close>)
       
 12137     by (metis \<open>span (S \<union> U) = span (S \<union> B)\<close> \<open>span B = span T\<close> span_Un)
       
 12138 qed
       
 12139 
       
 12140 corollary orthogonal_extension_strong:
       
 12141   fixes S :: "'a::euclidean_space set"
       
 12142   assumes S: "pairwise orthogonal S"
       
 12143   obtains U where "U \<inter> (insert 0 S) = {}" "pairwise orthogonal (S \<union> U)"
       
 12144                    "span (S \<union> U) = span (S \<union> T)"
       
 12145 proof -
       
 12146   obtain U where "pairwise orthogonal (S \<union> U)" "span (S \<union> U) = span (S \<union> T)"
       
 12147     using orthogonal_extension assms by blast
       
 12148   then show ?thesis
       
 12149     apply (rule_tac U = "U - (insert 0 S)" in that)
       
 12150       apply blast
       
 12151      apply (force simp: pairwise_def)
       
 12152     apply (metis (no_types, lifting) Un_Diff_cancel span_insert_0 span_Un)
       
 12153   done
       
 12154 qed
       
 12155 
       
 12156 subsection\<open>Decomposing a vector into parts in orthogonal subspaces.\<close>
       
 12157 
       
 12158 text\<open>existence of orthonormal basis for a subspace.\<close>
       
 12159 
       
 12160 lemma orthogonal_spanningset_subspace:
       
 12161   fixes S :: "'a :: euclidean_space set"
       
 12162   assumes "subspace S"
       
 12163   obtains B where "B \<subseteq> S" "pairwise orthogonal B" "span B = S"
       
 12164 proof -
       
 12165   obtain B where "B \<subseteq> S" "independent B" "S \<subseteq> span B" "card B = dim S"
       
 12166     using basis_exists by blast
       
 12167   with orthogonal_extension [of "{}" B]
       
 12168   show ?thesis
       
 12169     by (metis Un_empty_left assms pairwise_empty span_inc span_subspace that)
       
 12170 qed
       
 12171 
       
 12172 lemma orthogonal_basis_subspace:
       
 12173   fixes S :: "'a :: euclidean_space set"
       
 12174   assumes "subspace S"
       
 12175   obtains B where "0 \<notin> B" "B \<subseteq> S" "pairwise orthogonal B" "independent B"
       
 12176                   "card B = dim S" "span B = S"
       
 12177 proof -
       
 12178   obtain B where "B \<subseteq> S" "pairwise orthogonal B" "span B = S"
       
 12179     using assms orthogonal_spanningset_subspace by blast
       
 12180   then show ?thesis
       
 12181     apply (rule_tac B = "B - {0}" in that)
       
 12182     apply (auto simp: indep_card_eq_dim_span pairwise_subset Diff_subset pairwise_orthogonal_independent elim: pairwise_subset)
       
 12183     done
       
 12184 qed
       
 12185 
       
 12186 proposition orthonormal_basis_subspace:
       
 12187   fixes S :: "'a :: euclidean_space set"
       
 12188   assumes "subspace S"
       
 12189   obtains B where "B \<subseteq> S" "pairwise orthogonal B"
       
 12190               and "\<And>x. x \<in> B \<Longrightarrow> norm x = 1"
       
 12191               and "independent B" "card B = dim S" "span B = S"
       
 12192 proof -
       
 12193   obtain B where "0 \<notin> B" "B \<subseteq> S"
       
 12194              and orth: "pairwise orthogonal B"
       
 12195              and "independent B" "card B = dim S" "span B = S"
       
 12196     by (blast intro: orthogonal_basis_subspace [OF assms])
       
 12197   have 1: "(\<lambda>x. x /\<^sub>R norm x) ` B \<subseteq> S"
       
 12198     using \<open>span B = S\<close> span_clauses(1) span_mul by fastforce
       
 12199   have 2: "pairwise orthogonal ((\<lambda>x. x /\<^sub>R norm x) ` B)"
       
 12200     using orth by (force simp: pairwise_def orthogonal_clauses)
       
 12201   have 3: "\<And>x. x \<in> (\<lambda>x. x /\<^sub>R norm x) ` B \<Longrightarrow> norm x = 1"
       
 12202     by (metis (no_types, lifting) \<open>0 \<notin> B\<close> image_iff norm_sgn sgn_div_norm)
       
 12203   have 4: "independent ((\<lambda>x. x /\<^sub>R norm x) ` B)"
       
 12204     by (metis "2" "3" norm_zero pairwise_orthogonal_independent zero_neq_one)
       
 12205   have "inj_on (\<lambda>x. x /\<^sub>R norm x) B"
       
 12206   proof
       
 12207     fix x y
       
 12208     assume "x \<in> B" "y \<in> B" "x /\<^sub>R norm x = y /\<^sub>R norm y"
       
 12209     moreover have "\<And>i. i \<in> B \<Longrightarrow> norm (i /\<^sub>R norm i) = 1"
       
 12210       using 3 by blast
       
 12211     ultimately show "x = y"
       
 12212       by (metis norm_eq_1 orth orthogonal_clauses(7) orthogonal_commute orthogonal_def pairwise_def zero_neq_one)
       
 12213   qed
       
 12214   then have 5: "card ((\<lambda>x. x /\<^sub>R norm x) ` B) = dim S"
       
 12215     by (metis \<open>card B = dim S\<close> card_image)
       
 12216   have 6: "span ((\<lambda>x. x /\<^sub>R norm x) ` B) = S"
       
 12217     by (metis "1" "4" "5" assms card_eq_dim independent_finite span_subspace)
       
 12218   show ?thesis
       
 12219     by (rule that [OF 1 2 3 4 5 6])
       
 12220 qed
       
 12221 
       
 12222 proposition orthogonal_subspace_decomp_exists:
       
 12223   fixes S :: "'a :: euclidean_space set"
       
 12224   obtains y z where "y \<in> span S" "\<And>w. w \<in> span S \<Longrightarrow> orthogonal z w" "x = y + z"
       
 12225 proof -
       
 12226   obtain T where "0 \<notin> T" "T \<subseteq> span S" "pairwise orthogonal T" "independent T" "card T = dim (span S)" "span T = span S"
       
 12227     using orthogonal_basis_subspace subspace_span by blast
       
 12228   let ?a = "\<Sum>b\<in>T. (b \<bullet> x / (b \<bullet> b)) *\<^sub>R b"
       
 12229   have orth: "orthogonal (x - ?a) w" if "w \<in> span S" for w
       
 12230     by (simp add: Gram_Schmidt_step \<open>pairwise orthogonal T\<close> \<open>span T = span S\<close> orthogonal_commute that)
       
 12231   show ?thesis
       
 12232     apply (rule_tac y = "?a" and z = "x - ?a" in that)
       
 12233       apply (meson \<open>T \<subseteq> span S\<close> span_mul span_setsum subsetCE)
       
 12234      apply (fact orth, simp)
       
 12235     done
       
 12236 qed
       
 12237 
       
 12238 lemma orthogonal_subspace_decomp_unique:
       
 12239   fixes S :: "'a :: euclidean_space set"
       
 12240   assumes "x + y = x' + y'"
       
 12241       and ST: "x \<in> span S" "x' \<in> span S" "y \<in> span T" "y' \<in> span T"
       
 12242       and orth: "\<And>a b. \<lbrakk>a \<in> S; b \<in> T\<rbrakk> \<Longrightarrow> orthogonal a b"
       
 12243   shows "x = x' \<and> y = y'"
       
 12244 proof -
       
 12245   have "x + y - y' = x'"
       
 12246     by (simp add: assms)
       
 12247   moreover have "\<And>a b. \<lbrakk>a \<in> span S; b \<in> span T\<rbrakk> \<Longrightarrow> orthogonal a b"
       
 12248     by (meson orth orthogonal_commute orthogonal_to_span)
       
 12249   ultimately have "0 = x' - x"
       
 12250     by (metis (full_types) add_diff_cancel_left' ST diff_right_commute orthogonal_clauses(10) orthogonal_clauses(5) orthogonal_self)
       
 12251   with assms show ?thesis by auto
       
 12252 qed
       
 12253 
       
 12254 proposition dim_orthogonal_sum:
       
 12255   fixes A :: "'a::euclidean_space set"
       
 12256   assumes "\<And>x y. \<lbrakk>x \<in> A; y \<in> B\<rbrakk> \<Longrightarrow> x \<bullet> y = 0"
       
 12257     shows "dim(A \<union> B) = dim A + dim B"
       
 12258 proof -
       
 12259   have 1: "\<And>x y. \<lbrakk>x \<in> span A; y \<in> B\<rbrakk> \<Longrightarrow> x \<bullet> y = 0"
       
 12260     by (erule span_induct [OF _ subspace_hyperplane2]; simp add: assms)
       
 12261   have "\<And>x y. \<lbrakk>x \<in> span A; y \<in> span B\<rbrakk> \<Longrightarrow> x \<bullet> y = 0"
       
 12262     apply (erule span_induct [OF _ subspace_hyperplane])
       
 12263     using 1 by (simp add: )
       
 12264   then have 0: "\<And>x y. \<lbrakk>x \<in> span A; y \<in> span B\<rbrakk> \<Longrightarrow> x \<bullet> y = 0"
       
 12265     by simp
       
 12266   have "dim(A \<union> B) = dim (span (A \<union> B))"
       
 12267     by simp
       
 12268   also have "... = dim ((\<lambda>(a, b). a + b) ` (span A \<times> span B))"
       
 12269     by (simp add: span_Un)
       
 12270   also have "... = dim {x + y |x y. x \<in> span A \<and> y \<in> span B}"
       
 12271     by (auto intro!: arg_cong [where f=dim])
       
 12272   also have "... = dim {x + y |x y. x \<in> span A \<and> y \<in> span B} + dim(span A \<inter> span B)"
       
 12273     by (auto simp: dest: 0)
       
 12274   also have "... = dim (span A) + dim (span B)"
       
 12275     by (rule dim_sums_Int) auto
       
 12276   also have "... = dim A + dim B"
       
 12277     by simp
       
 12278   finally show ?thesis .
       
 12279 qed
       
 12280 
       
 12281 lemma dim_subspace_orthogonal_to_vectors:
       
 12282   fixes A :: "'a::euclidean_space set"
       
 12283   assumes "subspace A" "subspace B" "A \<subseteq> B"
       
 12284     shows "dim {y \<in> B. \<forall>x \<in> A. orthogonal x y} + dim A = dim B"
       
 12285 proof -
       
 12286   have "dim (span ({y \<in> B. \<forall>x\<in>A. orthogonal x y} \<union> A)) = dim (span B)"
       
 12287   proof (rule arg_cong [where f=dim, OF subset_antisym])
       
 12288     show "span ({y \<in> B. \<forall>x\<in>A. orthogonal x y} \<union> A) \<subseteq> span B"
       
 12289       by (simp add: \<open>A \<subseteq> B\<close> Collect_restrict span_mono)
       
 12290   next
       
 12291     have *: "x \<in> span ({y \<in> B. \<forall>x\<in>A. orthogonal x y} \<union> A)"
       
 12292          if "x \<in> B" for x
       
 12293     proof -
       
 12294       obtain y z where "x = y + z" "y \<in> span A" and orth: "\<And>w. w \<in> span A \<Longrightarrow> orthogonal z w"
       
 12295         using orthogonal_subspace_decomp_exists [of A x] that by auto
       
 12296       have "y \<in> span B"
       
 12297         by (metis span_eq \<open>y \<in> span A\<close> assms subset_iff)
       
 12298       then have "z \<in> {a \<in> B. \<forall>x. x \<in> A \<longrightarrow> orthogonal x a}"
       
 12299         by simp (metis (no_types) span_eq \<open>x = y + z\<close> \<open>subspace A\<close> \<open>subspace B\<close> orth orthogonal_commute span_add_eq that)
       
 12300       then have z: "z \<in> span {y \<in> B. \<forall>x\<in>A. orthogonal x y}"
       
 12301         by (meson span_inc subset_iff)
       
 12302       then show ?thesis
       
 12303         apply (simp add: span_Un image_def)
       
 12304         apply (rule bexI [OF _ z])
       
 12305         apply (simp add: \<open>x = y + z\<close> \<open>y \<in> span A\<close>)
       
 12306         done
       
 12307     qed
       
 12308     show "span B \<subseteq> span ({y \<in> B. \<forall>x\<in>A. orthogonal x y} \<union> A)"
       
 12309       by (rule span_minimal) (auto intro: * span_minimal elim: )
       
 12310   qed
       
 12311   then show ?thesis
       
 12312     by (metis (no_types, lifting) dim_orthogonal_sum dim_span mem_Collect_eq orthogonal_commute orthogonal_def)
       
 12313 qed
       
 12314 
       
 12315 subsection\<open>Parallel slices, etc.\<close>
       
 12316 
       
 12317 text\<open> If we take a slice out of a set, we can do it perpendicularly,
       
 12318   with the normal vector to the slice parallel to the affine hull.\<close>
       
 12319 
       
 12320 proposition affine_parallel_slice:
       
 12321   fixes S :: "'a :: euclidean_space set"
       
 12322   assumes "affine S"
       
 12323       and "S \<inter> {x. a \<bullet> x \<le> b} \<noteq> {}"
       
 12324       and "~ (S \<subseteq> {x. a \<bullet> x \<le> b})"
       
 12325   obtains a' b' where "a' \<noteq> 0"
       
 12326                    "S \<inter> {x. a' \<bullet> x \<le> b'} = S \<inter> {x. a \<bullet> x \<le> b}"
       
 12327                    "S \<inter> {x. a' \<bullet> x = b'} = S \<inter> {x. a \<bullet> x = b}"
       
 12328                    "\<And>w. w \<in> S \<Longrightarrow> (w + a') \<in> S"
       
 12329 proof (cases "S \<inter> {x. a \<bullet> x = b} = {}")
       
 12330   case True
       
 12331   then obtain u v where "u \<in> S" "v \<in> S" "a \<bullet> u \<le> b" "a \<bullet> v > b"
       
 12332     using assms by (auto simp: not_le)
       
 12333   define \<eta> where "\<eta> = u + ((b - a \<bullet> u) / (a \<bullet> v - a \<bullet> u)) *\<^sub>R (v - u)"
       
 12334   have "\<eta> \<in> S"
       
 12335     by (simp add: \<eta>_def \<open>u \<in> S\<close> \<open>v \<in> S\<close> \<open>affine S\<close> mem_affine_3_minus)
       
 12336   moreover have "a \<bullet> \<eta> = b"
       
 12337     using \<open>a \<bullet> u \<le> b\<close> \<open>b < a \<bullet> v\<close>
       
 12338     by (simp add: \<eta>_def algebra_simps) (simp add: field_simps)
       
 12339   ultimately have False
       
 12340     using True by force
       
 12341   then show ?thesis ..
       
 12342 next
       
 12343   case False
       
 12344   then obtain z where "z \<in> S" and z: "a \<bullet> z = b"
       
 12345     using assms by auto
       
 12346   with affine_diffs_subspace [OF \<open>affine S\<close>]
       
 12347   have sub: "subspace (op + (- z) ` S)" by blast
       
 12348   then have aff: "affine (op + (- z) ` S)" and span: "span (op + (- z) ` S) = (op + (- z) ` S)"
       
 12349     by (auto simp: subspace_imp_affine)
       
 12350   obtain a' a'' where a': "a' \<in> span (op + (- z) ` S)" and a: "a = a' + a''"
       
 12351                   and "\<And>w. w \<in> span (op + (- z) ` S) \<Longrightarrow> orthogonal a'' w"
       
 12352       using orthogonal_subspace_decomp_exists [of "op + (- z) ` S" "a"] by metis
       
 12353   then have "\<And>w. w \<in> S \<Longrightarrow> a'' \<bullet> (w-z) = 0"
       
 12354     by (simp add: imageI orthogonal_def span)
       
 12355   then have a'': "\<And>w. w \<in> S \<Longrightarrow> a'' \<bullet> w = (a - a') \<bullet> z"
       
 12356     by (simp add: a inner_diff_right)
       
 12357   then have ba'': "\<And>w. w \<in> S \<Longrightarrow> a'' \<bullet> w = b - a' \<bullet> z"
       
 12358     by (simp add: inner_diff_left z)
       
 12359   have "\<And>w. w \<in> op + (- z) ` S \<Longrightarrow> (w + a') \<in> op + (- z) ` S"
       
 12360     by (metis subspace_add a' span_eq sub)
       
 12361   then have Sclo: "\<And>w. w \<in> S \<Longrightarrow> (w + a') \<in> S"
       
 12362     by fastforce
       
 12363   show ?thesis
       
 12364   proof (cases "a' = 0")
       
 12365     case True
       
 12366     with a assms True a'' diff_zero less_irrefl show ?thesis
       
 12367       by auto
       
 12368   next
       
 12369     case False
       
 12370     show ?thesis
       
 12371       apply (rule_tac a' = "a'" and b' = "a' \<bullet> z" in that)
       
 12372       apply (auto simp: a ba'' inner_left_distrib False Sclo)
       
 12373       done
       
 12374   qed
       
 12375 qed
       
 12376 
       
 12377 lemma diffs_affine_hull_span:
       
 12378   assumes "a \<in> S"
       
 12379     shows "{x - a |x. x \<in> affine hull S} = span {x - a |x. x \<in> S}"
       
 12380 proof -
       
 12381   have *: "((\<lambda>x. x - a) ` (S - {a})) = {x. x + a \<in> S} - {0}"
       
 12382     by (auto simp: algebra_simps)
       
 12383   show ?thesis
       
 12384     apply (simp add: affine_hull_span2 [OF assms] *)
       
 12385     apply (auto simp: algebra_simps)
       
 12386     done
       
 12387 qed
       
 12388 
       
 12389 lemma aff_dim_dim_affine_diffs:
       
 12390   fixes S :: "'a :: euclidean_space set"
       
 12391   assumes "affine S" "a \<in> S"
       
 12392     shows "aff_dim S = dim {x - a |x. x \<in> S}"
       
 12393 proof -
       
 12394   obtain B where aff: "affine hull B = affine hull S"
       
 12395              and ind: "\<not> affine_dependent B"
       
 12396              and card: "of_nat (card B) = aff_dim S + 1"
       
 12397     using aff_dim_basis_exists by blast
       
 12398   then have "B \<noteq> {}" using assms
       
 12399     by (metis affine_hull_eq_empty ex_in_conv)
       
 12400   then obtain c where "c \<in> B" by auto
       
 12401   then have "c \<in> S"
       
 12402     by (metis aff affine_hull_eq \<open>affine S\<close> hull_inc)
       
 12403   have xy: "x - c = y - a \<longleftrightarrow> y = x + 1 *\<^sub>R (a - c)" for x y c and a::'a
       
 12404     by (auto simp: algebra_simps)
       
 12405   have *: "{x - c |x. x \<in> S} = {x - a |x. x \<in> S}"
       
 12406     apply safe
       
 12407     apply (simp_all only: xy)
       
 12408     using mem_affine_3_minus [OF \<open>affine S\<close>] \<open>a \<in> S\<close> \<open>c \<in> S\<close> apply blast+
       
 12409     done
       
 12410   have affS: "affine hull S = S"
       
 12411     by (simp add: \<open>affine S\<close>)
       
 12412   have "aff_dim S = of_nat (card B) - 1"
       
 12413     using card by simp
       
 12414   also have "... = dim {x - c |x. x \<in> B}"
       
 12415     by (simp add: affine_independent_card_dim_diffs [OF ind \<open>c \<in> B\<close>])
       
 12416   also have "... = dim {x - c | x. x \<in> affine hull B}"
       
 12417      by (simp add: diffs_affine_hull_span \<open>c \<in> B\<close>)
       
 12418   also have "... = dim {x - a |x. x \<in> S}"
       
 12419      by (simp add: affS aff *)
       
 12420    finally show ?thesis .
       
 12421 qed
       
 12422 
       
 12423 lemma aff_dim_linear_image_le:
       
 12424   assumes "linear f"
       
 12425     shows "aff_dim(f ` S) \<le> aff_dim S"
       
 12426 proof -
       
 12427   have "aff_dim (f ` T) \<le> aff_dim T" if "affine T" for T
       
 12428   proof (cases "T = {}")
       
 12429     case True then show ?thesis by (simp add: aff_dim_geq)
       
 12430   next
       
 12431     case False
       
 12432     then obtain a where "a \<in> T" by auto
       
 12433     have 1: "((\<lambda>x. x - f a) ` f ` T) = {x - f a |x. x \<in> f ` T}"
       
 12434       by auto
       
 12435     have 2: "{x - f a| x. x \<in> f ` T} = f ` {x - a| x. x \<in> T}"
       
 12436       by (force simp: linear_diff [OF assms])
       
 12437     have "aff_dim (f ` T) = int (dim {x - f a |x. x \<in> f ` T})"
       
 12438       by (simp add: \<open>a \<in> T\<close> hull_inc aff_dim_eq_dim [of "f a"] 1)
       
 12439     also have "... = int (dim (f ` {x - a| x. x \<in> T}))"
       
 12440       by (force simp: linear_diff [OF assms] 2)
       
 12441     also have "... \<le> int (dim {x - a| x. x \<in> T})"
       
 12442       by (simp add: dim_image_le [OF assms])
       
 12443     also have "... \<le> aff_dim T"
       
 12444       by (simp add: aff_dim_dim_affine_diffs [symmetric] \<open>a \<in> T\<close> \<open>affine T\<close>)
       
 12445     finally show ?thesis .
       
 12446   qed
       
 12447   then
       
 12448   have "aff_dim (f ` (affine hull S)) \<le> aff_dim (affine hull S)"
       
 12449     using affine_affine_hull [of S] by blast
       
 12450   then show ?thesis
       
 12451     using affine_hull_linear_image assms linear_conv_bounded_linear by fastforce
       
 12452 qed
       
 12453 
       
 12454 lemma aff_dim_injective_linear_image [simp]:
       
 12455   assumes "linear f" "inj f"
       
 12456     shows "aff_dim (f ` S) = aff_dim S"
       
 12457 proof (rule antisym)
       
 12458   show "aff_dim (f ` S) \<le> aff_dim S"
       
 12459     by (simp add: aff_dim_linear_image_le assms(1))
       
 12460 next
       
 12461   obtain g where "linear g" "g \<circ> f = id"
       
 12462     using linear_injective_left_inverse assms by blast
       
 12463   then have "aff_dim S \<le> aff_dim(g ` f ` S)"
       
 12464     by (simp add: image_comp)
       
 12465   also have "... \<le> aff_dim (f ` S)"
       
 12466     by (simp add: \<open>linear g\<close> aff_dim_linear_image_le)
       
 12467   finally show "aff_dim S \<le> aff_dim (f ` S)" .
       
 12468 qed
       
 12469 
       
 12470 text\<open>Choosing a subspace of a given dimension\<close>
       
 12471 proposition choose_subspace_of_subspace:
       
 12472   fixes S :: "'n::euclidean_space set"
       
 12473   assumes "n \<le> dim S"
       
 12474   obtains T where "subspace T" "T \<subseteq> span S" "dim T = n"
       
 12475 proof -
       
 12476   have "\<exists>T. subspace T \<and> T \<subseteq> span S \<and> dim T = n"
       
 12477   using assms
       
 12478   proof (induction n)
       
 12479     case 0 then show ?case by force
       
 12480   next
       
 12481     case (Suc n)
       
 12482     then obtain T where "subspace T" "T \<subseteq> span S" "dim T = n"
       
 12483       by force
       
 12484     then show ?case
       
 12485     proof (cases "span S \<subseteq> span T")
       
 12486       case True
       
 12487       have "dim S = dim T"
       
 12488         apply (rule span_eq_dim [OF subset_antisym [OF True]])
       
 12489         by (simp add: \<open>T \<subseteq> span S\<close> span_minimal subspace_span)
       
 12490       then show ?thesis
       
 12491         using Suc.prems \<open>dim T = n\<close> by linarith
       
 12492     next
       
 12493       case False
       
 12494       then obtain y where y: "y \<in> S" "y \<notin> T"
       
 12495         by (meson span_mono subsetI)
       
 12496       then have "span (insert y T) \<subseteq> span S"
       
 12497         by (metis (no_types) \<open>T \<subseteq> span S\<close> subsetD insert_subset span_inc span_mono span_span)
       
 12498       with \<open>dim T = n\<close>  \<open>subspace T\<close> y show ?thesis
       
 12499         apply (rule_tac x="span(insert y T)" in exI)
       
 12500         apply (auto simp: dim_insert)
       
 12501         using span_eq by blast
       
 12502     qed
       
 12503   qed
       
 12504   with that show ?thesis by blast
       
 12505 qed
       
 12506 
       
 12507 lemma choose_affine_subset:
       
 12508   assumes "affine S" "-1 \<le> d" and dle: "d \<le> aff_dim S"
       
 12509   obtains T where "affine T" "T \<subseteq> S" "aff_dim T = d"
       
 12510 proof (cases "d = -1 \<or> S={}")
       
 12511   case True with assms show ?thesis
       
 12512     by (metis aff_dim_empty affine_empty bot.extremum that eq_iff)
       
 12513 next
       
 12514   case False
       
 12515   with assms obtain a where "a \<in> S" "0 \<le> d" by auto
       
 12516   with assms have ss: "subspace (op + (- a) ` S)"
       
 12517     by (simp add: affine_diffs_subspace)
       
 12518   have "nat d \<le> dim (op + (- a) ` S)"
       
 12519     by (metis aff_dim_subspace aff_dim_translation_eq dle nat_int nat_mono ss)
       
 12520   then obtain T where "subspace T" and Tsb: "T \<subseteq> span (op + (- a) ` S)"
       
 12521                   and Tdim: "dim T = nat d"
       
 12522     using choose_subspace_of_subspace [of "nat d" "op + (- a) ` S"] by blast
       
 12523   then have "affine T"
       
 12524     using subspace_affine by blast
       
 12525   then have "affine (op + a ` T)"
       
 12526     by (metis affine_hull_eq affine_hull_translation)
       
 12527   moreover have "op + a ` T \<subseteq> S"
       
 12528   proof -
       
 12529     have "T \<subseteq> op + (- a) ` S"
       
 12530       by (metis (no_types) span_eq Tsb ss)
       
 12531     then show "op + a ` T \<subseteq> S"
       
 12532       using add_ac by auto
       
 12533   qed
       
 12534   moreover have "aff_dim (op + a ` T) = d"
       
 12535     by (simp add: aff_dim_subspace Tdim \<open>0 \<le> d\<close> \<open>subspace T\<close> aff_dim_translation_eq)
       
 12536   ultimately show ?thesis
       
 12537     by (rule that)
       
 12538 qed
       
 12539 
       
 12540 subsection\<open>Several Variants of Paracompactness\<close>
       
 12541 
       
 12542 proposition paracompact:
       
 12543   fixes S :: "'a :: euclidean_space set"
       
 12544   assumes "S \<subseteq> \<Union>\<C>" and opC: "\<And>T. T \<in> \<C> \<Longrightarrow> open T"
       
 12545   obtains \<C>' where "S \<subseteq> \<Union> \<C>'"
       
 12546                and "\<And>U. U \<in> \<C>' \<Longrightarrow> open U \<and> (\<exists>T. T \<in> \<C> \<and> U \<subseteq> T)"
       
 12547                and "\<And>x. x \<in> S
       
 12548                        \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and>
       
 12549                                finite {U. U \<in> \<C>' \<and> (U \<inter> V \<noteq> {})}"
       
 12550 proof (cases "S = {}")
       
 12551   case True with that show ?thesis by blast
       
 12552 next
       
 12553   case False
       
 12554   have "\<exists>T U. x \<in> U \<and> open U \<and> closure U \<subseteq> T \<and> T \<in> \<C>" if "x \<in> S" for x
       
 12555   proof -
       
 12556     obtain T where "x \<in> T" "T \<in> \<C>" "open T"
       
 12557       using assms \<open>x \<in> S\<close> by blast
       
 12558     then obtain e where "e > 0" "cball x e \<subseteq> T"
       
 12559       by (force simp: open_contains_cball)
       
 12560     then show ?thesis
       
 12561       apply (rule_tac x = T in exI)
       
 12562       apply (rule_tac x = "ball x e" in exI)
       
 12563       using  \<open>T \<in> \<C>\<close>
       
 12564       apply (simp add: closure_minimal)
       
 12565       done
       
 12566   qed
       
 12567   then obtain F G where Gin: "x \<in> G x" and oG: "open (G x)"
       
 12568                     and clos: "closure (G x) \<subseteq> F x" and Fin: "F x \<in> \<C>"
       
 12569          if "x \<in> S" for x
       
 12570     by metis
       
 12571   then obtain \<F> where "\<F> \<subseteq> G ` S" "countable \<F>" "\<Union>\<F> = UNION S G"
       
 12572     using Lindelof [of "G ` S"] by (metis image_iff)
       
 12573   then obtain K where K: "K \<subseteq> S" "countable K" and eq: "UNION K G = UNION S G"
       
 12574     by (metis countable_subset_image)
       
 12575   with False Gin have "K \<noteq> {}" by force
       
 12576   then obtain a :: "nat \<Rightarrow> 'a" where "range a = K"
       
 12577     by (metis range_from_nat_into \<open>countable K\<close>)
       
 12578   then have odif: "\<And>n. open (F (a n) - \<Union>{closure (G (a m)) |m. m < n})"
       
 12579     using \<open>K \<subseteq> S\<close> Fin opC by (fastforce simp add:)
       
 12580   let ?C = "range (\<lambda>n. F(a n) - \<Union>{closure(G(a m)) |m. m < n})"
       
 12581   have enum_S: "\<exists>n. x \<in> F(a n) \<and> x \<in> G(a n)" if "x \<in> S" for x
       
 12582   proof -
       
 12583     have "\<exists>y \<in> K. x \<in> G y" using eq that Gin by fastforce
       
 12584     then show ?thesis
       
 12585       using clos K \<open>range a = K\<close> closure_subset by blast
       
 12586   qed
       
 12587   have 1: "S \<subseteq> Union ?C"
       
 12588   proof
       
 12589     fix x assume "x \<in> S"
       
 12590     define n where "n \<equiv> LEAST n. x \<in> F(a n)"
       
 12591     have n: "x \<in> F(a n)"
       
 12592       using enum_S [OF \<open>x \<in> S\<close>] by (force simp: n_def intro: LeastI)
       
 12593     have notn: "x \<notin> F(a m)" if "m < n" for m
       
 12594       using that not_less_Least by (force simp: n_def)
       
 12595     then have "x \<notin> \<Union>{closure (G (a m)) |m. m < n}"
       
 12596       using n \<open>K \<subseteq> S\<close> \<open>range a = K\<close> clos notn by fastforce
       
 12597     with n show "x \<in> Union ?C"
       
 12598       by blast
       
 12599   qed
       
 12600   have 3: "\<exists>V. open V \<and> x \<in> V \<and> finite {U. U \<in> ?C \<and> (U \<inter> V \<noteq> {})}" if "x \<in> S" for x
       
 12601   proof -
       
 12602     obtain n where n: "x \<in> F(a n)" "x \<in> G(a n)"
       
 12603       using \<open>x \<in> S\<close> enum_S by auto
       
 12604     have "{U \<in> ?C. U \<inter> G (a n) \<noteq> {}} \<subseteq> (\<lambda>n. F(a n) - \<Union>{closure(G(a m)) |m. m < n}) ` atMost n"
       
 12605     proof clarsimp
       
 12606       fix k  assume "(F (a k) - \<Union>{closure (G (a m)) |m. m < k}) \<inter> G (a n) \<noteq> {}"
       
 12607       then have "k \<le> n"
       
 12608         by auto (metis closure_subset not_le subsetCE)
       
 12609       then show "F (a k) - \<Union>{closure (G (a m)) |m. m < k}
       
 12610                  \<in> (\<lambda>n. F (a n) - \<Union>{closure (G (a m)) |m. m < n}) ` {..n}"
       
 12611         by force
       
 12612     qed
       
 12613     moreover have "finite ((\<lambda>n. F(a n) - \<Union>{closure(G(a m)) |m. m < n}) ` atMost n)"
       
 12614       by force
       
 12615     ultimately have *: "finite {U \<in> ?C. U \<inter> G (a n) \<noteq> {}}"
       
 12616       using finite_subset by blast
       
 12617     show ?thesis
       
 12618       apply (rule_tac x="G (a n)" in exI)
       
 12619       apply (intro conjI oG n *)
       
 12620       using \<open>K \<subseteq> S\<close> \<open>range a = K\<close> apply blast
       
 12621       done
       
 12622   qed
       
 12623   show ?thesis
       
 12624     apply (rule that [OF 1 _ 3])
       
 12625     using Fin \<open>K \<subseteq> S\<close> \<open>range a = K\<close>  apply (auto simp: odif)
       
 12626     done
       
 12627 qed
       
 12628 
       
 12629 
       
 12630 corollary paracompact_closedin:
       
 12631   fixes S :: "'a :: euclidean_space set"
       
 12632   assumes cin: "closedin (subtopology euclidean U) S"
       
 12633       and oin: "\<And>T. T \<in> \<C> \<Longrightarrow> openin (subtopology euclidean U) T"
       
 12634       and "S \<subseteq> \<Union>\<C>"
       
 12635   obtains \<C>' where "S \<subseteq> \<Union> \<C>'"
       
 12636                and "\<And>V. V \<in> \<C>' \<Longrightarrow> openin (subtopology euclidean U) V \<and> (\<exists>T. T \<in> \<C> \<and> V \<subseteq> T)"
       
 12637                and "\<And>x. x \<in> U
       
 12638                        \<Longrightarrow> \<exists>V. openin (subtopology euclidean U) V \<and> x \<in> V \<and>
       
 12639                                finite {X. X \<in> \<C>' \<and> (X \<inter> V \<noteq> {})}"
       
 12640 proof -
       
 12641   have "\<exists>Z. open Z \<and> (T = U \<inter> Z)" if "T \<in> \<C>" for T
       
 12642     using oin [OF that] by (auto simp: openin_open)
       
 12643   then obtain F where opF: "open (F T)" and intF: "U \<inter> F T = T" if "T \<in> \<C>" for T
       
 12644     by metis
       
 12645   obtain K where K: "closed K" "U \<inter> K = S"
       
 12646     using cin by (auto simp: closedin_closed)
       
 12647   have 1: "U \<subseteq> \<Union>insert (- K) (F ` \<C>)"
       
 12648     by clarsimp (metis Int_iff Union_iff \<open>U \<inter> K = S\<close> \<open>S \<subseteq> \<Union>\<C>\<close> subsetD intF)
       
 12649   have 2: "\<And>T. T \<in> insert (- K) (F ` \<C>) \<Longrightarrow> open T"
       
 12650     using \<open>closed K\<close> by (auto simp: opF)
       
 12651   obtain \<D> where "U \<subseteq> \<Union>\<D>"
       
 12652              and D1: "\<And>U. U \<in> \<D> \<Longrightarrow> open U \<and> (\<exists>T. T \<in> insert (- K) (F ` \<C>) \<and> U \<subseteq> T)"
       
 12653              and D2: "\<And>x. x \<in> U \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> finite {U \<in> \<D>. U \<inter> V \<noteq> {}}"
       
 12654     using paracompact [OF 1 2] by auto
       
 12655   let ?C = "{U \<inter> V |V. V \<in> \<D> \<and> (V \<inter> K \<noteq> {})}"
       
 12656   show ?thesis
       
 12657   proof (rule_tac \<C>' = "{U \<inter> V |V. V \<in> \<D> \<and> (V \<inter> K \<noteq> {})}" in that)
       
 12658     show "S \<subseteq> \<Union>?C"
       
 12659       using \<open>U \<inter> K = S\<close> \<open>U \<subseteq> \<Union>\<D>\<close> K by (blast dest!: subsetD)
       
 12660     show "\<And>V. V \<in> ?C \<Longrightarrow> openin (subtopology euclidean U) V \<and> (\<exists>T. T \<in> \<C> \<and> V \<subseteq> T)"
       
 12661       using D1 intF by fastforce
       
 12662     have *: "{X. (\<exists>V. X = U \<inter> V \<and> V \<in> \<D> \<and> V \<inter> K \<noteq> {}) \<and> X \<inter> (U \<inter> V) \<noteq> {}} \<subseteq>
       
 12663              (\<lambda>x. U \<inter> x) ` {U \<in> \<D>. U \<inter> V \<noteq> {}}" for V
       
 12664       by blast
       
 12665     show "\<exists>V. openin (subtopology euclidean U) V \<and> x \<in> V \<and> finite {X \<in> ?C. X \<inter> V \<noteq> {}}"
       
 12666          if "x \<in> U" for x
       
 12667       using D2 [OF that]
       
 12668       apply clarify
       
 12669       apply (rule_tac x="U \<inter> V" in exI)
       
 12670       apply (auto intro: that finite_subset [OF *])
       
 12671       done
       
 12672     qed
       
 12673 qed
       
 12674 
       
 12675 corollary paracompact_closed:
       
 12676   fixes S :: "'a :: euclidean_space set"
       
 12677   assumes "closed S"
       
 12678       and opC: "\<And>T. T \<in> \<C> \<Longrightarrow> open T"
       
 12679       and "S \<subseteq> \<Union>\<C>"
       
 12680   obtains \<C>' where "S \<subseteq> \<Union>\<C>'"
       
 12681                and "\<And>U. U \<in> \<C>' \<Longrightarrow> open U \<and> (\<exists>T. T \<in> \<C> \<and> U \<subseteq> T)"
       
 12682                and "\<And>x. \<exists>V. open V \<and> x \<in> V \<and>
       
 12683                                finite {U. U \<in> \<C>' \<and> (U \<inter> V \<noteq> {})}"
       
 12684 using paracompact_closedin [of UNIV S \<C>] assms
       
 12685 by (auto simp: open_openin [symmetric] closed_closedin [symmetric])
       
 12686 
       
 12687 end