src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy
changeset 49531 8d68162b7826
parent 49530 7faf67b411b4
child 49962 a8cc904a6820
equal deleted inserted replaced
49530:7faf67b411b4 49531:8d68162b7826
    34   using assms convex_def[of S] by auto
    34   using assms convex_def[of S] by auto
    35 
    35 
    36 lemma mem_convex_alt:
    36 lemma mem_convex_alt:
    37   assumes "convex S" "x : S" "y : S" "u>=0" "v>=0" "u+v>0"
    37   assumes "convex S" "x : S" "y : S" "u>=0" "v>=0" "u+v>0"
    38   shows "((u/(u+v)) *\<^sub>R x + (v/(u+v)) *\<^sub>R y) : S"
    38   shows "((u/(u+v)) *\<^sub>R x + (v/(u+v)) *\<^sub>R y) : S"
    39   apply (subst mem_convex_2) 
    39   apply (subst mem_convex_2)
    40   using assms apply (auto simp add: algebra_simps zero_le_divide_iff)
    40   using assms apply (auto simp add: algebra_simps zero_le_divide_iff)
    41   using add_divide_distrib[of u v "u+v"] apply auto
    41   using add_divide_distrib[of u v "u+v"] apply auto
    42   done
    42   done
    43 
    43 
    44 lemma inj_on_image_mem_iff: "inj_on f B ==> (A <= B) ==> (f a : f`A) ==> (a : B) ==> (a : A)" 
    44 lemma inj_on_image_mem_iff: "inj_on f B ==> (A <= B) ==> (f a : f`A) ==> (a : B) ==> (a : A)"
    45   by (blast dest: inj_onD)
    45   by (blast dest: inj_onD)
    46 
    46 
    47 lemma independent_injective_on_span_image:
    47 lemma independent_injective_on_span_image:
    48   assumes iS: "independent S" 
    48   assumes iS: "independent S"
    49     and lf: "linear f" and fi: "inj_on f (span S)"
    49     and lf: "linear f" and fi: "inj_on f (span S)"
    50   shows "independent (f ` S)"
    50   shows "independent (f ` S)"
    51 proof -
    51 proof -
    52   {
    52   {
    53     fix a
    53     fix a
    63   then show ?thesis unfolding dependent_def by blast
    63   then show ?thesis unfolding dependent_def by blast
    64 qed
    64 qed
    65 
    65 
    66 lemma dim_image_eq:
    66 lemma dim_image_eq:
    67   fixes f :: "'n::euclidean_space => 'm::euclidean_space"
    67   fixes f :: "'n::euclidean_space => 'm::euclidean_space"
    68   assumes lf: "linear f" and fi: "inj_on f (span S)" 
    68   assumes lf: "linear f" and fi: "inj_on f (span S)"
    69   shows "dim (f ` S) = dim (S:: ('n::euclidean_space) set)"
    69   shows "dim (f ` S) = dim (S:: ('n::euclidean_space) set)"
    70 proof -
    70 proof -
    71   obtain B where B_def: "B<=S & independent B & S <= span B & card B = dim S" 
    71   obtain B where B_def: "B<=S & independent B & S <= span B & card B = dim S"
    72     using basis_exists[of S] by auto
    72     using basis_exists[of S] by auto
    73   then have "span S = span B"
    73   then have "span S = span B"
    74     using span_mono[of B S] span_mono[of S "span B"] span_span[of B] by auto
    74     using span_mono[of B S] span_mono[of S "span B"] span_span[of B] by auto
    75   then have "independent (f ` B)"
    75   then have "independent (f ` B)"
    76     using independent_injective_on_span_image[of B f] B_def assms by auto
    76     using independent_injective_on_span_image[of B f] B_def assms by auto
    88 proof -
    88 proof -
    89   have "inj_on f S <-> (!x : S. !y : S. f x = f y --> x = y)" by (simp add: inj_on_def)
    89   have "inj_on f S <-> (!x : S. !y : S. f x = f y --> x = y)" by (simp add: inj_on_def)
    90   also have "... <-> (!x : S. !y : S. f x - f y = 0 --> x - y = 0)" by simp
    90   also have "... <-> (!x : S. !y : S. f x - f y = 0 --> x - y = 0)" by simp
    91   also have "... <-> (!x : S. !y : S. f (x - y) = 0 --> x - y = 0)"
    91   also have "... <-> (!x : S. !y : S. f (x - y) = 0 --> x - y = 0)"
    92     by (simp add: linear_sub[OF lf])
    92     by (simp add: linear_sub[OF lf])
    93   also have "... <-> (! x : S. f x = 0 --> x = 0)" 
    93   also have "... <-> (! x : S. f x = 0 --> x = 0)"
    94     using `subspace S` subspace_def[of S] subspace_sub[of S] by auto
    94     using `subspace S` subspace_def[of S] subspace_sub[of S] by auto
    95   finally show ?thesis .
    95   finally show ?thesis .
    96 qed
    96 qed
    97 
    97 
    98 lemma subspace_Inter: "(!s : f. subspace s) ==> subspace (Inter f)"
    98 lemma subspace_Inter: "(!s : f. subspace s) ==> subspace (Inter f)"
    99   unfolding subspace_def by auto 
    99   unfolding subspace_def by auto
   100 
   100 
   101 lemma span_eq[simp]: "(span s = s) <-> subspace s"
   101 lemma span_eq[simp]: "(span s = s) <-> subspace s"
   102   unfolding span_def by (rule hull_eq, rule subspace_Inter)
   102   unfolding span_def by (rule hull_eq, rule subspace_Inter)
   103 
   103 
   104 lemma basis_inj_on: "d \<subseteq> {..<DIM('n)} \<Longrightarrow> inj_on (basis :: nat => 'n::euclidean_space) d"
   104 lemma basis_inj_on: "d \<subseteq> {..<DIM('n)} \<Longrightarrow> inj_on (basis :: nat => 'n::euclidean_space) d"
   105   by (auto simp add: inj_on_def euclidean_eq[where 'a='n])
   105   by (auto simp add: inj_on_def euclidean_eq[where 'a='n])
   106   
   106 
   107 lemma finite_substdbasis: "finite {basis i ::'n::euclidean_space |i. i : (d:: nat set)}" (is "finite ?S")
   107 lemma finite_substdbasis: "finite {basis i ::'n::euclidean_space |i. i : (d:: nat set)}" (is "finite ?S")
   108 proof -
   108 proof -
   109   have eq: "?S = basis ` d" by blast
   109   have eq: "?S = basis ` d" by blast
   110   show ?thesis
   110   show ?thesis
   111     unfolding eq
   111     unfolding eq
   150     apply(intro independent_mono[of "{basis i ::'a |i. i : {..<DIM('a::euclidean_space)}}" "?A"] )
   150     apply(intro independent_mono[of "{basis i ::'a |i. i : {..<DIM('a::euclidean_space)}}" "?A"] )
   151     using independent_basis[where 'a='a] assms apply (auto simp: *)
   151     using independent_basis[where 'a='a] assms apply (auto simp: *)
   152     done
   152     done
   153 qed
   153 qed
   154 
   154 
   155 lemma dim_cball: 
   155 lemma dim_cball:
   156   assumes "0<e"
   156   assumes "0<e"
   157   shows "dim (cball (0 :: 'n::euclidean_space) e) = DIM('n)"
   157   shows "dim (cball (0 :: 'n::euclidean_space) e) = DIM('n)"
   158 proof -
   158 proof -
   159   { fix x :: "'n::euclidean_space"
   159   { fix x :: "'n::euclidean_space"
   160     def y == "(e/norm x) *\<^sub>R x"
   160     def y == "(e/norm x) *\<^sub>R x"
   161     then have "y : cball 0 e" using cball_def dist_norm[of 0 y] assms by auto
   161     then have "y : cball 0 e" using cball_def dist_norm[of 0 y] assms by auto
   162     moreover have "x = (norm x/e) *\<^sub>R y" using y_def assms by simp
   162     moreover have "x = (norm x/e) *\<^sub>R y" using y_def assms by simp
   163     moreover hence "x = (norm x/e) *\<^sub>R y" by auto
   163     moreover hence "x = (norm x/e) *\<^sub>R y" by auto
   164     ultimately have "x : span (cball 0 e)"
   164     ultimately have "x : span (cball 0 e)"
   165       using span_mul[of y "cball 0 e" "norm x/e"] span_inc[of "cball 0 e"] by auto
   165       using span_mul[of y "cball 0 e" "norm x/e"] span_inc[of "cball 0 e"] by auto
   166   } then have "span (cball 0 e) = (UNIV :: ('n::euclidean_space) set)" by auto 
   166   } then have "span (cball 0 e) = (UNIV :: ('n::euclidean_space) set)" by auto
   167   then show ?thesis
   167   then show ?thesis
   168     using dim_span[of "cball (0 :: 'n::euclidean_space) e"] by (auto simp add: dim_UNIV)
   168     using dim_span[of "cball (0 :: 'n::euclidean_space) e"] by (auto simp add: dim_UNIV)
   169 qed
   169 qed
   170 
   170 
   171 lemma indep_card_eq_dim_span:
   171 lemma indep_card_eq_dim_span:
   172   fixes B :: "('n::euclidean_space) set"
   172   fixes B :: "('n::euclidean_space) set"
   173   assumes "independent B"
   173   assumes "independent B"
   174   shows "finite B & card B = dim (span B)" 
   174   shows "finite B & card B = dim (span B)"
   175   using assms basis_card_eq_dim[of B "span B"] span_inc by auto
   175   using assms basis_card_eq_dim[of B "span B"] span_inc by auto
   176 
   176 
   177 lemma setsum_not_0: "setsum f A ~= 0 ==> EX a:A. f a ~= 0"
   177 lemma setsum_not_0: "setsum f A ~= 0 ==> EX a:A. f a ~= 0"
   178   by (rule ccontr) auto
   178   by (rule ccontr) auto
   179 
   179 
   180 lemma translate_inj_on: 
   180 lemma translate_inj_on:
   181   fixes A :: "('a::ab_group_add) set"
   181   fixes A :: "('a::ab_group_add) set"
   182   shows "inj_on (%x. a+x) A"
   182   shows "inj_on (%x. a+x) A"
   183   unfolding inj_on_def by auto
   183   unfolding inj_on_def by auto
   184 
   184 
   185 lemma translation_assoc:
   185 lemma translation_assoc:
   204   using translation_assoc[of "-a" a S] apply auto
   204   using translation_assoc[of "-a" a S] apply auto
   205   using translation_assoc[of a "-a" T] apply auto
   205   using translation_assoc[of a "-a" T] apply auto
   206   done
   206   done
   207 
   207 
   208 lemma translation_inverse_subset:
   208 lemma translation_inverse_subset:
   209   assumes "((%x. -a+x) ` V) <= (S :: 'n::ab_group_add set)" 
   209   assumes "((%x. -a+x) ` V) <= (S :: 'n::ab_group_add set)"
   210   shows "V <= ((%x. a+x) ` S)"
   210   shows "V <= ((%x. a+x) ` S)"
   211 proof -
   211 proof -
   212   { fix x
   212   { fix x
   213     assume "x:V"
   213     assume "x:V"
   214     then have "x-a : S" using assms by auto
   214     then have "x-a : S" using assms by auto
   392 
   392 
   393 lemma affine_UNIV[intro]: "affine UNIV"
   393 lemma affine_UNIV[intro]: "affine UNIV"
   394   unfolding affine_def by auto
   394   unfolding affine_def by auto
   395 
   395 
   396 lemma affine_Inter: "(\<forall>s\<in>f. affine s) \<Longrightarrow> affine (\<Inter> f)"
   396 lemma affine_Inter: "(\<forall>s\<in>f. affine s) \<Longrightarrow> affine (\<Inter> f)"
   397   unfolding affine_def by auto 
   397   unfolding affine_def by auto
   398 
   398 
   399 lemma affine_Int: "affine s \<Longrightarrow> affine t \<Longrightarrow> affine (s \<inter> t)"
   399 lemma affine_Int: "affine s \<Longrightarrow> affine t \<Longrightarrow> affine (s \<inter> t)"
   400   unfolding affine_def by auto
   400   unfolding affine_def by auto
   401 
   401 
   402 lemma affine_affine_hull: "affine(affine hull s)"
   402 lemma affine_affine_hull: "affine(affine hull s)"
   414   shows "affine V \<longleftrightarrow>
   414   shows "affine V \<longleftrightarrow>
   415     (\<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)"
   415     (\<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)"
   416   unfolding affine_def
   416   unfolding affine_def
   417   apply rule
   417   apply rule
   418   apply(rule, rule, rule)
   418   apply(rule, rule, rule)
   419   apply(erule conjE)+ 
   419   apply(erule conjE)+
   420   defer
   420   defer
   421   apply (rule, rule, rule, rule, rule)
   421   apply (rule, rule, rule, rule, rule)
   422 proof -
   422 proof -
   423   fix x y u v
   423   fix x y u v
   424   assume as: "x \<in> V" "y \<in> V" "u + v = (1::real)"
   424   assume as: "x \<in> V" "y \<in> V" "u + v = (1::real)"
   478       have "(\<Sum>xa\<in>s - {x}. (inverse (1 - u x) * u xa) *\<^sub>R xa) \<in> V"
   478       have "(\<Sum>xa\<in>s - {x}. (inverse (1 - u x) * u xa) *\<^sub>R xa) \<in> V"
   479       proof (cases "card (s - {x}) > 2")
   479       proof (cases "card (s - {x}) > 2")
   480         case True
   480         case True
   481         then have "s - {x} \<noteq> {}" "card (s - {x}) = n"
   481         then have "s - {x} \<noteq> {}" "card (s - {x}) = n"
   482           unfolding c and as(1)[symmetric]
   482           unfolding c and as(1)[symmetric]
   483         proof (rule_tac ccontr) 
   483         proof (rule_tac ccontr)
   484           assume "\<not> s - {x} \<noteq> {}"
   484           assume "\<not> s - {x} \<noteq> {}"
   485           then have "card (s - {x}) = 0" unfolding card_0_eq[OF *(2)] by simp 
   485           then have "card (s - {x}) = 0" unfolding card_0_eq[OF *(2)] by simp
   486           then show False using True by auto
   486           then show False using True by auto
   487         qed auto
   487         qed auto
   488         then show ?thesis
   488         then show ?thesis
   489           apply (rule_tac IA[of "s - {x}" "\<lambda>y. (inverse (1 - u x) * u y)"])
   489           apply (rule_tac IA[of "s - {x}" "\<lambda>y. (inverse (1 - u x) * u y)"])
   490           unfolding setsum_right_distrib[symmetric] using as and *** and True
   490           unfolding setsum_right_distrib[symmetric] using as and *** and True
   638     proof (cases "a \<in> s")
   638     proof (cases "a \<in> s")
   639       case True
   639       case True
   640       then show ?thesis
   640       then show ?thesis
   641         apply (rule_tac x="\<lambda>x. (if x=a then v else 0) + u x" in exI)
   641         apply (rule_tac x="\<lambda>x. (if x=a then v else 0) + u x" in exI)
   642         unfolding setsum_clauses(2)[OF `?as`] apply simp
   642         unfolding setsum_clauses(2)[OF `?as`] apply simp
   643         unfolding scaleR_left_distrib and setsum_addf 
   643         unfolding scaleR_left_distrib and setsum_addf
   644         unfolding vu and * and scaleR_zero_left
   644         unfolding vu and * and scaleR_zero_left
   645         apply (auto simp add: setsum_delta[OF `?as`])
   645         apply (auto simp add: setsum_delta[OF `?as`])
   646         done
   646         done
   647     next
   647     next
   648       case False 
   648       case False
   649       then have **:
   649       then have **:
   650         "\<And>x. x \<in> s \<Longrightarrow> u x = (if x = a then v else u x)"
   650         "\<And>x. x \<in> s \<Longrightarrow> u x = (if x = a then v else u x)"
   651         "\<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
   651         "\<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
   652       from False show ?thesis
   652       from False show ?thesis
   653         apply (rule_tac x="\<lambda>x. if x=a then v else u x" in exI)
   653         apply (rule_tac x="\<lambda>x. if x=a then v else u x" in exI)
   664 lemma affine_hull_2:
   664 lemma affine_hull_2:
   665   fixes a b :: "'a::real_vector"
   665   fixes a b :: "'a::real_vector"
   666   shows "affine hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b| u v. (u + v = 1)}" (is "?lhs = ?rhs")
   666   shows "affine hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b| u v. (u + v = 1)}" (is "?lhs = ?rhs")
   667 proof -
   667 proof -
   668   have *:
   668   have *:
   669     "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)" 
   669     "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)"
   670     "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto
   670     "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto
   671   have "?lhs = {y. \<exists>u. setsum u {a, b} = 1 \<and> (\<Sum>v\<in>{a, b}. u v *\<^sub>R v) = y}"
   671   have "?lhs = {y. \<exists>u. setsum u {a, b} = 1 \<and> (\<Sum>v\<in>{a, b}. u v *\<^sub>R v) = y}"
   672     using affine_hull_finite[of "{a,b}"] by auto
   672     using affine_hull_finite[of "{a,b}"] by auto
   673   also have "\<dots> = {y. \<exists>v u. u b = 1 - v \<and> u b *\<^sub>R b = y - v *\<^sub>R a}"
   673   also have "\<dots> = {y. \<exists>v u. u b = 1 - v \<and> u b *\<^sub>R b = y - v *\<^sub>R a}"
   674     by (simp add: affine_hull_finite_step(2)[of "{b}" a])
   674     by (simp add: affine_hull_finite_step(2)[of "{b}" a])
   679 lemma affine_hull_3:
   679 lemma affine_hull_3:
   680   fixes a b c :: "'a::real_vector"
   680   fixes a b c :: "'a::real_vector"
   681   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}" (is "?lhs = ?rhs")
   681   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}" (is "?lhs = ?rhs")
   682 proof -
   682 proof -
   683   have *:
   683   have *:
   684     "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)" 
   684     "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)"
   685     "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto
   685     "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto
   686   show ?thesis
   686   show ?thesis
   687     apply (simp add: affine_hull_finite affine_hull_finite_step)
   687     apply (simp add: affine_hull_finite affine_hull_finite_step)
   688     unfolding *
   688     unfolding *
   689     apply auto
   689     apply auto
   706   moreover
   706   moreover
   707   have "affine hull {x, y, z} <= affine hull S"
   707   have "affine hull {x, y, z} <= affine hull S"
   708     using hull_mono[of "{x, y, z}" "S"] assms by auto
   708     using hull_mono[of "{x, y, z}" "S"] assms by auto
   709   moreover
   709   moreover
   710   have "affine hull S = S" using assms affine_hull_eq[of S] by auto
   710   have "affine hull S = S" using assms affine_hull_eq[of S] by auto
   711   ultimately show ?thesis by auto 
   711   ultimately show ?thesis by auto
   712 qed
   712 qed
   713 
   713 
   714 lemma mem_affine_3_minus:
   714 lemma mem_affine_3_minus:
   715   assumes "affine S" "x : S" "y : S" "z : S"
   715   assumes "affine S" "x : S" "y : S" "z : S"
   716   shows "x + v *\<^sub>R (y-z) : S"
   716   shows "x + v *\<^sub>R (y-z) : S"
   748   shows "affine hull (insert a s) = {a + v | v . v \<in> span {x - a | x.  x \<in> s}}"
   748   shows "affine hull (insert a s) = {a + v | v . v \<in> span {x - a | x.  x \<in> s}}"
   749   apply (rule, rule affine_hull_insert_subset_span)
   749   apply (rule, rule affine_hull_insert_subset_span)
   750   unfolding subset_eq Ball_def
   750   unfolding subset_eq Ball_def
   751   unfolding affine_hull_explicit and mem_Collect_eq
   751   unfolding affine_hull_explicit and mem_Collect_eq
   752 proof (rule, rule, erule exE, erule conjE)
   752 proof (rule, rule, erule exE, erule conjE)
   753   fix y v 
   753   fix y v
   754   assume "y = a + v" "v \<in> span {x - a |x. x \<in> s}"
   754   assume "y = a + v" "v \<in> span {x - a |x. x \<in> s}"
   755   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"
   755   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"
   756     unfolding span_explicit by auto
   756     unfolding span_explicit by auto
   757   def f \<equiv> "(\<lambda>x. x + a) ` t"
   757   def f \<equiv> "(\<lambda>x. x + a) ` t"
   758   have f:"finite f" "f \<subseteq> s" "(\<Sum>v\<in>f. u (v - a) *\<^sub>R (v - a)) = y - a"
   758   have f:"finite f" "f \<subseteq> s" "(\<Sum>v\<in>f. u (v - a) *\<^sub>R (v - a)) = y - a"
   778 definition affine_parallel :: "'a::real_vector set => 'a::real_vector set => bool"
   778 definition affine_parallel :: "'a::real_vector set => 'a::real_vector set => bool"
   779   where "affine_parallel S T = (? a. T = ((%x. a + x) ` S))"
   779   where "affine_parallel S T = (? a. T = ((%x. a + x) ` S))"
   780 
   780 
   781 lemma affine_parallel_expl_aux:
   781 lemma affine_parallel_expl_aux:
   782   fixes S T :: "'a::real_vector set"
   782   fixes S T :: "'a::real_vector set"
   783   assumes "!x. (x : S <-> (a+x) : T)" 
   783   assumes "!x. (x : S <-> (a+x) : T)"
   784   shows "T = ((%x. a + x) ` S)"
   784   shows "T = ((%x. a + x) ` S)"
   785 proof -
   785 proof -
   786   { fix x
   786   { fix x
   787     assume "x : T"
   787     assume "x : T"
   788     then have "(-a)+x : S" using assms by auto
   788     then have "(-a)+x : S" using assms by auto
   789     then have "x : ((%x. a + x) ` S)"
   789     then have "x : ((%x. a + x) ` S)"
   790       using imageI[of "-a+x" S "(%x. a+x)"] by auto }
   790       using imageI[of "-a+x" S "(%x. a+x)"] by auto }
   791   moreover have "T >= ((%x. a + x) ` S)" using assms by auto 
   791   moreover have "T >= ((%x. a + x) ` S)" using assms by auto
   792   ultimately show ?thesis by auto
   792   ultimately show ?thesis by auto
   793 qed
   793 qed
   794 
   794 
   795 lemma affine_parallel_expl: "affine_parallel S T = (? a. !x. (x : S <-> (a+x) : T))"
   795 lemma affine_parallel_expl: "affine_parallel S T = (? a. !x. (x : S <-> (a+x) : T))"
   796   unfolding affine_parallel_def
   796   unfolding affine_parallel_def
   809     using translation_galois[of B a A] unfolding affine_parallel_def by auto
   809     using translation_galois[of B a A] unfolding affine_parallel_def by auto
   810 qed
   810 qed
   811 
   811 
   812 lemma affine_parallel_assoc:
   812 lemma affine_parallel_assoc:
   813   assumes "affine_parallel A B" "affine_parallel B C"
   813   assumes "affine_parallel A B" "affine_parallel B C"
   814   shows "affine_parallel A C" 
   814   shows "affine_parallel A C"
   815 proof -
   815 proof -
   816   from assms obtain ab where "B=((%x. ab + x) ` A)"
   816   from assms obtain ab where "B=((%x. ab + x) ` A)"
   817     unfolding affine_parallel_def by auto 
   817     unfolding affine_parallel_def by auto
   818   moreover 
   818   moreover
   819   from assms obtain bc where "C=((%x. bc + x) ` B)"
   819   from assms obtain bc where "C=((%x. bc + x) ` B)"
   820     unfolding affine_parallel_def by auto
   820     unfolding affine_parallel_def by auto
   821   ultimately show ?thesis
   821   ultimately show ?thesis
   822     using translation_assoc[of bc ab A] unfolding affine_parallel_def by auto
   822     using translation_assoc[of bc ab A] unfolding affine_parallel_def by auto
   823 qed
   823 qed
   854   fixes S T :: "'a::real_vector set"
   854   fixes S T :: "'a::real_vector set"
   855   assumes "affine S" "affine_parallel S T"
   855   assumes "affine S" "affine_parallel S T"
   856   shows "affine T"
   856   shows "affine T"
   857 proof -
   857 proof -
   858   from assms obtain a where "T=((%x. a + x) ` S)"
   858   from assms obtain a where "T=((%x. a + x) ` S)"
   859     unfolding affine_parallel_def by auto 
   859     unfolding affine_parallel_def by auto
   860   then show ?thesis using affine_translation assms by auto
   860   then show ?thesis using affine_translation assms by auto
   861 qed
   861 qed
   862 
   862 
   863 lemma subspace_imp_affine: "subspace s \<Longrightarrow> affine s"
   863 lemma subspace_imp_affine: "subspace s \<Longrightarrow> affine s"
   864   unfolding subspace_def affine_def by auto
   864   unfolding subspace_def affine_def by auto
   869 lemma subspace_affine: "subspace S <-> (affine S & 0 : S)"
   869 lemma subspace_affine: "subspace S <-> (affine S & 0 : S)"
   870 proof -
   870 proof -
   871   have h0: "subspace S ==> (affine S & 0 : S)"
   871   have h0: "subspace S ==> (affine S & 0 : S)"
   872     using subspace_imp_affine[of S] subspace_0 by auto
   872     using subspace_imp_affine[of S] subspace_0 by auto
   873   { assume assm: "affine S & 0 : S"
   873   { assume assm: "affine S & 0 : S"
   874     { fix c :: real 
   874     { fix c :: real
   875       fix x assume x_def: "x : S"
   875       fix x assume x_def: "x : S"
   876       have "c *\<^sub>R x = (1-c) *\<^sub>R 0 + c *\<^sub>R x" by auto
   876       have "c *\<^sub>R x = (1-c) *\<^sub>R 0 + c *\<^sub>R x" by auto
   877       moreover
   877       moreover
   878       have "(1-c) *\<^sub>R 0 + c *\<^sub>R x : S" using affine_alt[of S] assm x_def by auto
   878       have "(1-c) *\<^sub>R 0 + c *\<^sub>R x : S" using affine_alt[of S] assm x_def by auto
   879       ultimately have "c *\<^sub>R x : S" by auto
   879       ultimately have "c *\<^sub>R x : S" by auto
   892       moreover
   892       moreover
   893       have "(x+y) = 2 *\<^sub>R ((1/2) *\<^sub>R (x+y))" by auto
   893       have "(x+y) = 2 *\<^sub>R ((1/2) *\<^sub>R (x+y))" by auto
   894       ultimately
   894       ultimately
   895       have "(x+y) : S" using h1[rule_format, of "(1/2) *\<^sub>R (x+y)" "2"] by auto
   895       have "(x+y) : S" using h1[rule_format, of "(1/2) *\<^sub>R (x+y)" "2"] by auto
   896     }
   896     }
   897     then have "!x : S. !y : S. (x+y) : S" by auto 
   897     then have "!x : S. !y : S. (x+y) : S" by auto
   898     then have "subspace S" using h1 assm unfolding subspace_def by auto
   898     then have "subspace S" using h1 assm unfolding subspace_def by auto
   899   }
   899   }
   900   then show ?thesis using h0 by metis
   900   then show ?thesis using h0 by metis
   901 qed
   901 qed
   902 
   902 
   903 lemma affine_diffs_subspace:
   903 lemma affine_diffs_subspace:
   904   assumes "affine S" "a : S"
   904   assumes "affine S" "a : S"
   905   shows "subspace ((%x. (-a)+x) ` S)"
   905   shows "subspace ((%x. (-a)+x) ` S)"
   906 proof -
   906 proof -
   907   have "affine ((%x. (-a)+x) ` S)"
   907   have "affine ((%x. (-a)+x) ` S)"
   908     using  affine_translation assms by auto  
   908     using  affine_translation assms by auto
   909   moreover have "0 : ((%x. (-a)+x) ` S)"
   909   moreover have "0 : ((%x. (-a)+x) ` S)"
   910     using assms exI[of "(%x. x:S & -a+x=0)" a] by auto
   910     using assms exI[of "(%x. x:S & -a+x=0)" a] by auto
   911   ultimately show ?thesis using subspace_affine by auto 
   911   ultimately show ?thesis using subspace_affine by auto
   912 qed
   912 qed
   913 
   913 
   914 lemma parallel_subspace_explicit:
   914 lemma parallel_subspace_explicit:
   915   assumes "affine S" "a : S"
   915   assumes "affine S" "a : S"
   916   assumes "L == {y. ? x : S. (-a)+x=y}" 
   916   assumes "L == {y. ? x : S. (-a)+x=y}"
   917   shows "subspace L & affine_parallel S L" 
   917   shows "subspace L & affine_parallel S L"
   918 proof -
   918 proof -
   919   have par: "affine_parallel S L"
   919   have par: "affine_parallel S L"
   920     unfolding affine_parallel_def using assms by auto
   920     unfolding affine_parallel_def using assms by auto
   921   then have "affine L" using assms parallel_is_affine by auto  
   921   then have "affine L" using assms parallel_is_affine by auto
   922   moreover have "0 : L"
   922   moreover have "0 : L"
   923     using assms apply auto
   923     using assms apply auto
   924     using exI[of "(%x. x:S & -a+x=0)" a] apply auto
   924     using exI[of "(%x. x:S & -a+x=0)" a] apply auto
   925     done
   925     done
   926   ultimately show ?thesis using subspace_affine par by auto 
   926   ultimately show ?thesis using subspace_affine par by auto
   927 qed
   927 qed
   928 
   928 
   929 lemma parallel_subspace_aux:
   929 lemma parallel_subspace_aux:
   930   assumes "subspace A" "subspace B" "affine_parallel A B"
   930   assumes "subspace A" "subspace B" "affine_parallel A B"
   931   shows "A>=B"
   931   shows "A>=B"
   947     using assms parallel_subspace_aux[of B A] affine_parallel_commut by auto
   947     using assms parallel_subspace_aux[of B A] affine_parallel_commut by auto
   948 qed
   948 qed
   949 
   949 
   950 lemma affine_parallel_subspace:
   950 lemma affine_parallel_subspace:
   951   assumes "affine S" "S ~= {}"
   951   assumes "affine S" "S ~= {}"
   952   shows "?!L. subspace L & affine_parallel S L" 
   952   shows "?!L. subspace L & affine_parallel S L"
   953 proof -
   953 proof -
   954   have ex: "? L. subspace L & affine_parallel S L"
   954   have ex: "? L. subspace L & affine_parallel S L"
   955     using assms parallel_subspace_explicit by auto 
   955     using assms parallel_subspace_explicit by auto
   956   { fix L1 L2
   956   { fix L1 L2
   957     assume ass: "subspace L1 & affine_parallel S L1" "subspace L2 & affine_parallel S L2"
   957     assume ass: "subspace L1 & affine_parallel S L1" "subspace L2 & affine_parallel S L2"
   958     then have "affine_parallel L1 L2"
   958     then have "affine_parallel L1 L2"
   959       using affine_parallel_commut[of S L1] affine_parallel_assoc[of L1 S L2] by auto
   959       using affine_parallel_commut[of S L1] affine_parallel_assoc[of L1 S L2] by auto
   960     then have "L1 = L2"
   960     then have "L1 = L2"
  1082   } then have "cone ?rhs" unfolding cone_def by auto
  1082   } then have "cone ?rhs" unfolding cone_def by auto
  1083   then have "?rhs : Collect cone" unfolding mem_Collect_eq by auto
  1083   then have "?rhs : Collect cone" unfolding mem_Collect_eq by auto
  1084   { fix x
  1084   { fix x
  1085     assume "x : S"
  1085     assume "x : S"
  1086     then have "1 *\<^sub>R x : ?rhs"
  1086     then have "1 *\<^sub>R x : ?rhs"
  1087       apply auto  
  1087       apply auto
  1088       apply (rule_tac x="1" in exI)
  1088       apply (rule_tac x="1" in exI)
  1089       apply auto
  1089       apply auto
  1090       done
  1090       done
  1091     then have "x : ?rhs" by auto
  1091     then have "x : ?rhs" by auto
  1092   } then have "S <= ?rhs" by auto
  1092   } then have "S <= ?rhs" by auto
  1251 
  1251 
  1252 lemma convex_connected:
  1252 lemma convex_connected:
  1253   fixes s :: "'a::real_normed_vector set"
  1253   fixes s :: "'a::real_normed_vector set"
  1254   assumes "convex s" shows "connected s"
  1254   assumes "convex s" shows "connected s"
  1255 proof-
  1255 proof-
  1256   { fix e1 e2 assume as:"open e1" "open e2" "e1 \<inter> e2 \<inter> s = {}" "s \<subseteq> e1 \<union> e2" 
  1256   { fix e1 e2 assume as:"open e1" "open e2" "e1 \<inter> e2 \<inter> s = {}" "s \<subseteq> e1 \<union> e2"
  1257     assume "e1 \<inter> s \<noteq> {}" "e2 \<inter> s \<noteq> {}"
  1257     assume "e1 \<inter> s \<noteq> {}" "e2 \<inter> s \<noteq> {}"
  1258     then obtain x1 x2 where x1:"x1\<in>e1" "x1\<in>s" and x2:"x2\<in>e2" "x2\<in>s" by auto
  1258     then obtain x1 x2 where x1:"x1\<in>e1" "x1\<in>s" and x2:"x2\<in>e2" "x2\<in>s" by auto
  1259     hence n:"norm (x1 - x2) > 0" unfolding zero_less_norm_iff using as(3) by auto
  1259     hence n:"norm (x1 - x2) > 0" unfolding zero_less_norm_iff using as(3) by auto
  1260 
  1260 
  1261     { fix x e::real assume as:"0 \<le> x" "x \<le> 1" "0 < e"
  1261     { fix x e::real assume as:"0 \<le> x" "x \<le> 1" "0 < e"
  1275       using as(1,2)[unfolded open_dist] apply simp
  1275       using as(1,2)[unfolded open_dist] apply simp
  1276       using as(1,2)[unfolded open_dist] apply simp
  1276       using as(1,2)[unfolded open_dist] apply simp
  1277       using assms[unfolded convex_alt, THEN bspec[where x=x1], THEN bspec[where x=x2]] using x1 x2
  1277       using assms[unfolded convex_alt, THEN bspec[where x=x1], THEN bspec[where x=x2]] using x1 x2
  1278       using as(3) by auto
  1278       using as(3) by auto
  1279     then obtain x where "x\<ge>0" "x\<le>1" "(1 - x) *\<^sub>R x1 + x *\<^sub>R x2 \<notin> e1"  "(1 - x) *\<^sub>R x1 + x *\<^sub>R x2 \<notin> e2" by auto
  1279     then obtain x where "x\<ge>0" "x\<le>1" "(1 - x) *\<^sub>R x1 + x *\<^sub>R x2 \<notin> e1"  "(1 - x) *\<^sub>R x1 + x *\<^sub>R x2 \<notin> e2" by auto
  1280     hence False using as(4) 
  1280     hence False using as(4)
  1281       using assms[unfolded convex_alt, THEN bspec[where x=x1], THEN bspec[where x=x2]]
  1281       using assms[unfolded convex_alt, THEN bspec[where x=x1], THEN bspec[where x=x2]]
  1282       using x1(2) x2(2) by auto  }
  1282       using x1(2) x2(2) by auto  }
  1283   thus ?thesis unfolding connected_def by auto
  1283   thus ?thesis unfolding connected_def by auto
  1284 qed
  1284 qed
  1285 
  1285 
  1320   ultimately show False using mult_strict_left_mono[OF y `u>0`] unfolding left_diff_distrib by auto
  1320   ultimately show False using mult_strict_left_mono[OF y `u>0`] unfolding left_diff_distrib by auto
  1321 qed
  1321 qed
  1322 
  1322 
  1323 lemma convex_ball:
  1323 lemma convex_ball:
  1324   fixes x :: "'a::real_normed_vector"
  1324   fixes x :: "'a::real_normed_vector"
  1325   shows "convex (ball x e)" 
  1325   shows "convex (ball x e)"
  1326 proof(auto simp add: convex_def)
  1326 proof(auto simp add: convex_def)
  1327   fix y z assume yz:"dist x y < e" "dist x z < e"
  1327   fix y z assume yz:"dist x y < e" "dist x z < e"
  1328   fix u v ::real assume uv:"0 \<le> u" "0 \<le> v" "u + v = 1"
  1328   fix u v ::real assume uv:"0 \<le> u" "0 \<le> v" "u + v = 1"
  1329   have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z" using uv yz
  1329   have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z" using uv yz
  1330     using convex_distance[of "ball x e" x, unfolded convex_on_def, THEN bspec[where x=y], THEN bspec[where x=z]] by auto
  1330     using convex_distance[of "ball x e" x, unfolded convex_on_def, THEN bspec[where x=y], THEN bspec[where x=z]] by auto
  1337 proof(auto simp add: convex_def Ball_def)
  1337 proof(auto simp add: convex_def Ball_def)
  1338   fix y z assume yz:"dist x y \<le> e" "dist x z \<le> e"
  1338   fix y z assume yz:"dist x y \<le> e" "dist x z \<le> e"
  1339   fix u v ::real assume uv:" 0 \<le> u" "0 \<le> v" "u + v = 1"
  1339   fix u v ::real assume uv:" 0 \<le> u" "0 \<le> v" "u + v = 1"
  1340   have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z" using uv yz
  1340   have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z" using uv yz
  1341     using convex_distance[of "cball x e" x, unfolded convex_on_def, THEN bspec[where x=y], THEN bspec[where x=z]] by auto
  1341     using convex_distance[of "cball x e" x, unfolded convex_on_def, THEN bspec[where x=y], THEN bspec[where x=z]] by auto
  1342   thus "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> e" using convex_bound_le[OF yz uv] by auto 
  1342   thus "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> e" using convex_bound_le[OF yz uv] by auto
  1343 qed
  1343 qed
  1344 
  1344 
  1345 lemma connected_ball:
  1345 lemma connected_ball:
  1346   fixes x :: "'a::real_normed_vector"
  1346   fixes x :: "'a::real_normed_vector"
  1347   shows "connected (ball x e)"
  1347   shows "connected (ball x e)"
  1377 subsubsection {* Convex hull is "preserved" by a linear function *}
  1377 subsubsection {* Convex hull is "preserved" by a linear function *}
  1378 
  1378 
  1379 lemma convex_hull_linear_image:
  1379 lemma convex_hull_linear_image:
  1380   assumes "bounded_linear f"
  1380   assumes "bounded_linear f"
  1381   shows "f ` (convex hull s) = convex hull (f ` s)"
  1381   shows "f ` (convex hull s) = convex hull (f ` s)"
  1382   apply rule unfolding subset_eq ball_simps apply(rule_tac[!] hull_induct, rule hull_inc) prefer 3  
  1382   apply rule unfolding subset_eq ball_simps apply(rule_tac[!] hull_induct, rule hull_inc) prefer 3
  1383   apply(erule imageE)apply(rule_tac x=xa in image_eqI) apply assumption
  1383   apply(erule imageE)apply(rule_tac x=xa in image_eqI) apply assumption
  1384   apply(rule hull_subset[unfolded subset_eq, rule_format]) apply assumption
  1384   apply(rule hull_subset[unfolded subset_eq, rule_format]) apply assumption
  1385 proof-
  1385 proof-
  1386   interpret f: bounded_linear f by fact
  1386   interpret f: bounded_linear f by fact
  1387   show "convex {x. f x \<in> convex hull f ` s}" 
  1387   show "convex {x. f x \<in> convex hull f ` s}"
  1388   unfolding convex_def by(auto simp add: f.scaleR f.add convex_convex_hull[unfolded convex_def, rule_format]) next
  1388   unfolding convex_def by(auto simp add: f.scaleR f.add convex_convex_hull[unfolded convex_def, rule_format]) next
  1389   interpret f: bounded_linear f by fact
  1389   interpret f: bounded_linear f by fact
  1390   show "convex {x. x \<in> f ` (convex hull s)}" using  convex_convex_hull[unfolded convex_def, of s] 
  1390   show "convex {x. x \<in> f ` (convex hull s)}" using  convex_convex_hull[unfolded convex_def, of s]
  1391     unfolding convex_def by (auto simp add: f.scaleR [symmetric] f.add [symmetric])
  1391     unfolding convex_def by (auto simp add: f.scaleR [symmetric] f.add [symmetric])
  1392 qed auto
  1392 qed auto
  1393 
  1393 
  1394 lemma in_convex_hull_linear_image:
  1394 lemma in_convex_hull_linear_image:
  1395   assumes "bounded_linear f" "x \<in> convex hull s"
  1395   assumes "bounded_linear f" "x \<in> convex hull s"
  1409   assumes "s \<noteq> {}"
  1409   assumes "s \<noteq> {}"
  1410   shows "convex hull (insert a s) = {x. \<exists>u\<ge>0. \<exists>v\<ge>0. \<exists>b. (u + v = 1) \<and>
  1410   shows "convex hull (insert a s) = {x. \<exists>u\<ge>0. \<exists>v\<ge>0. \<exists>b. (u + v = 1) \<and>
  1411                                     b \<in> (convex hull s) \<and> (x = u *\<^sub>R a + v *\<^sub>R b)}" (is "?xyz = ?hull")
  1411                                     b \<in> (convex hull s) \<and> (x = u *\<^sub>R a + v *\<^sub>R b)}" (is "?xyz = ?hull")
  1412  apply(rule,rule hull_minimal,rule) unfolding insert_iff prefer 3 apply rule proof-
  1412  apply(rule,rule hull_minimal,rule) unfolding insert_iff prefer 3 apply rule proof-
  1413  fix x assume x:"x = a \<or> x \<in> s"
  1413  fix x assume x:"x = a \<or> x \<in> s"
  1414  thus "x\<in>?hull" apply rule unfolding mem_Collect_eq apply(rule_tac x=1 in exI) defer 
  1414  thus "x\<in>?hull" apply rule unfolding mem_Collect_eq apply(rule_tac x=1 in exI) defer
  1415    apply(rule_tac x=0 in exI) using assms hull_subset[of s convex] by auto
  1415    apply(rule_tac x=0 in exI) using assms hull_subset[of s convex] by auto
  1416 next
  1416 next
  1417   fix x assume "x\<in>?hull"
  1417   fix x assume "x\<in>?hull"
  1418   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" by auto
  1418   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" by auto
  1419   have "a\<in>convex hull insert a s" "b\<in>convex hull insert a s"
  1419   have "a\<in>convex hull insert a s" "b\<in>convex hull insert a s"
  1433         using mult_nonneg_nonneg[OF `u\<ge>0` `v1\<ge>0`] mult_nonneg_nonneg[OF `v\<ge>0` `v2\<ge>0`] by arith+
  1433         using mult_nonneg_nonneg[OF `u\<ge>0` `v1\<ge>0`] mult_nonneg_nonneg[OF `v\<ge>0` `v2\<ge>0`] by arith+
  1434       hence "u * u1 + v * u2 = 1" using as(3) obt1(3) obt2(3) by auto
  1434       hence "u * u1 + v * u2 = 1" using as(3) obt1(3) obt2(3) by auto
  1435       thus ?thesis unfolding obt1(5) obt2(5) * using assms hull_subset[of s convex] by(auto simp add: ** scaleR_right_distrib)
  1435       thus ?thesis unfolding obt1(5) obt2(5) * using assms hull_subset[of s convex] by(auto simp add: ** scaleR_right_distrib)
  1436     next
  1436     next
  1437       have "1 - (u * u1 + v * u2) = (u + v) - (u * u1 + v * u2)" using as(3) obt1(3) obt2(3) by (auto simp add: field_simps)
  1437       have "1 - (u * u1 + v * u2) = (u + v) - (u * u1 + v * u2)" using as(3) obt1(3) obt2(3) by (auto simp add: field_simps)
  1438       also have "\<dots> = u * (v1 + u1 - u1) + v * (v2 + u2 - u2)" using as(3) obt1(3) obt2(3) by (auto simp add: field_simps) 
  1438       also have "\<dots> = u * (v1 + u1 - u1) + v * (v2 + u2 - u2)" using as(3) obt1(3) obt2(3) by (auto simp add: field_simps)
  1439       also have "\<dots> = u * v1 + v * v2" by simp finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" by auto
  1439       also have "\<dots> = u * v1 + v * v2" by simp finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" by auto
  1440       case False have "0 \<le> u * v1 + v * v2" "0 \<le> u * v1" "0 \<le> u * v1 + v * v2" "0 \<le> v * v2" apply -
  1440       case False have "0 \<le> u * v1 + v * v2" "0 \<le> u * v1" "0 \<le> u * v1 + v * v2" "0 \<le> v * v2" apply -
  1441         apply(rule add_nonneg_nonneg) prefer 4 apply(rule add_nonneg_nonneg) apply(rule_tac [!] mult_nonneg_nonneg)
  1441         apply(rule add_nonneg_nonneg) prefer 4 apply(rule add_nonneg_nonneg) apply(rule_tac [!] mult_nonneg_nonneg)
  1442         using as(1,2) obt1(1,2) obt2(1,2) by auto 
  1442         using as(1,2) obt1(1,2) obt2(1,2) by auto
  1443       thus ?thesis unfolding obt1(5) obt2(5) unfolding * and ** using False
  1443       thus ?thesis unfolding obt1(5) obt2(5) unfolding * and ** using False
  1444         apply(rule_tac x="((u * v1) / (u * v1 + v * v2)) *\<^sub>R b1 + ((v * v2) / (u * v1 + v * v2)) *\<^sub>R b2" in bexI) defer
  1444         apply(rule_tac x="((u * v1) / (u * v1 + v * v2)) *\<^sub>R b1 + ((v * v2) / (u * v1 + v * v2)) *\<^sub>R b2" in bexI) defer
  1445         apply(rule convex_convex_hull[of s, unfolded convex_def, rule_format]) using obt1(4) obt2(4)
  1445         apply(rule convex_convex_hull[of s, unfolded convex_def, rule_format]) using obt1(4) obt2(4)
  1446         unfolding add_divide_distrib[symmetric] and zero_le_divide_iff
  1446         unfolding add_divide_distrib[symmetric] and zero_le_divide_iff
  1447         by (auto simp add: scaleR_left_distrib scaleR_right_distrib)
  1447         by (auto simp add: scaleR_left_distrib scaleR_right_distrib)
  1449     have u1:"u1 \<le> 1" unfolding obt1(3)[symmetric] and not_le using obt1(2) by auto
  1449     have u1:"u1 \<le> 1" unfolding obt1(3)[symmetric] and not_le using obt1(2) by auto
  1450     have u2:"u2 \<le> 1" unfolding obt2(3)[symmetric] and not_le using obt2(2) by auto
  1450     have u2:"u2 \<le> 1" unfolding obt2(3)[symmetric] and not_le using obt2(2) by auto
  1451     have "u1 * u + u2 * v \<le> (max u1 u2) * u + (max u1 u2) * v" apply(rule add_mono)
  1451     have "u1 * u + u2 * v \<le> (max u1 u2) * u + (max u1 u2) * v" apply(rule add_mono)
  1452       apply(rule_tac [!] mult_right_mono) using as(1,2) obt1(1,2) obt2(1,2) by auto
  1452       apply(rule_tac [!] mult_right_mono) using as(1,2) obt1(1,2) obt2(1,2) by auto
  1453     also have "\<dots> \<le> 1" unfolding right_distrib[symmetric] and as(3) using u1 u2 by auto
  1453     also have "\<dots> \<le> 1" unfolding right_distrib[symmetric] and as(3) using u1 u2 by auto
  1454     finally 
  1454     finally
  1455     show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull" unfolding mem_Collect_eq apply(rule_tac x="u * u1 + v * u2" in exI)
  1455     show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull" unfolding mem_Collect_eq apply(rule_tac x="u * u1 + v * u2" in exI)
  1456       apply(rule conjI) defer apply(rule_tac x="1 - u * u1 - v * u2" in exI) unfolding Bex_def
  1456       apply(rule conjI) defer apply(rule_tac x="1 - u * u1 - v * u2" in exI) unfolding Bex_def
  1457       using as(1,2) obt1(1,2) obt2(1,2) * by(auto intro!: mult_nonneg_nonneg add_nonneg_nonneg simp add: algebra_simps)
  1457       using as(1,2) obt1(1,2) obt2(1,2) * by(auto intro!: mult_nonneg_nonneg add_nonneg_nonneg simp add: algebra_simps)
  1458   qed
  1458   qed
  1459 qed
  1459 qed
  1482   from xy obtain k1 u1 x1 where 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" by auto
  1482   from xy obtain k1 u1 x1 where 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" by auto
  1483   from xy obtain k2 u2 x2 where 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" by auto
  1483   from xy obtain k2 u2 x2 where 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" by auto
  1484   have *:"\<And>P (x1::'a) x2 s1 s2 i.(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)"
  1484   have *:"\<And>P (x1::'a) x2 s1 s2 i.(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)"
  1485     "{1..k1 + k2} \<inter> {1..k1} = {1..k1}" "{1..k1 + k2} \<inter> - {1..k1} = (\<lambda>i. i + k1) ` {1..k2}"
  1485     "{1..k1 + k2} \<inter> {1..k1} = {1..k1}" "{1..k1 + k2} \<inter> - {1..k1} = (\<lambda>i. i + k1) ` {1..k2}"
  1486     prefer 3 apply(rule,rule) unfolding image_iff apply(rule_tac x="x - k1" in bexI) by(auto simp add: not_le)
  1486     prefer 3 apply(rule,rule) unfolding image_iff apply(rule_tac x="x - k1" in bexI) by(auto simp add: not_le)
  1487   have inj:"inj_on (\<lambda>i. i + k1) {1..k2}" unfolding inj_on_def by auto  
  1487   have inj:"inj_on (\<lambda>i. i + k1) {1..k2}" unfolding inj_on_def by auto
  1488   show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull" apply(rule)
  1488   show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull" apply(rule)
  1489     apply(rule_tac x="k1 + k2" in exI, rule_tac x="\<lambda>i. if i \<in> {1..k1} then u * u1 i else v * u2 (i - k1)" in exI)
  1489     apply(rule_tac x="k1 + k2" in exI, rule_tac x="\<lambda>i. if i \<in> {1..k1} then u * u1 i else v * u2 (i - k1)" in exI)
  1490     apply(rule_tac x="\<lambda>i. if i \<in> {1..k1} then x1 i else x2 (i - k1)" in exI) apply(rule,rule) defer apply(rule)
  1490     apply(rule_tac x="\<lambda>i. if i \<in> {1..k1} then x1 i else x2 (i - k1)" in exI) apply(rule,rule) defer apply(rule)
  1491     unfolding * and setsum_cases[OF finite_atLeastAtMost[of 1 "k1 + k2"]] and setsum_reindex[OF inj] and o_def Collect_mem_eq
  1491     unfolding * and setsum_cases[OF finite_atLeastAtMost[of 1 "k1 + k2"]] and setsum_reindex[OF inj] and o_def Collect_mem_eq
  1492     unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] setsum_right_distrib[symmetric] proof-
  1492     unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] setsum_right_distrib[symmetric] proof-
  1505   fixes s :: "'a::real_vector set"
  1505   fixes s :: "'a::real_vector set"
  1506   assumes "finite s"
  1506   assumes "finite s"
  1507   shows "convex hull s = {y. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and>
  1507   shows "convex hull s = {y. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and>
  1508          setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}" (is "?HULL = ?set")
  1508          setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}" (is "?HULL = ?set")
  1509 proof(rule hull_unique, auto simp add: convex_def[of ?set])
  1509 proof(rule hull_unique, auto simp add: convex_def[of ?set])
  1510   fix x assume "x\<in>s" thus " \<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" 
  1510   fix x assume "x\<in>s" thus " \<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"
  1511     apply(rule_tac x="\<lambda>y. if x=y then 1 else 0" in exI) apply auto
  1511     apply(rule_tac x="\<lambda>y. if x=y then 1 else 0" in exI) apply auto
  1512     unfolding setsum_delta'[OF assms] and setsum_delta''[OF assms] by auto 
  1512     unfolding setsum_delta'[OF assms] and setsum_delta''[OF assms] by auto
  1513 next
  1513 next
  1514   fix u v ::real assume uv:"0 \<le> u" "0 \<le> v" "u + v = 1"
  1514   fix u v ::real assume uv:"0 \<le> u" "0 \<le> v" "u + v = 1"
  1515   fix ux assume ux:"\<forall>x\<in>s. 0 \<le> ux x" "setsum ux s = (1::real)"
  1515   fix ux assume ux:"\<forall>x\<in>s. 0 \<le> ux x" "setsum ux s = (1::real)"
  1516   fix uy assume uy:"\<forall>x\<in>s. 0 \<le> uy x" "setsum uy s = (1::real)"
  1516   fix uy assume uy:"\<forall>x\<in>s. 0 \<le> uy x" "setsum uy s = (1::real)"
  1517   { fix x assume "x\<in>s"
  1517   { fix x assume "x\<in>s"
  1520   moreover have "(\<Sum>x\<in>s. u * ux x + v * uy x) = 1"
  1520   moreover have "(\<Sum>x\<in>s. u * ux x + v * uy x) = 1"
  1521     unfolding setsum_addf and setsum_right_distrib[symmetric] and ux(2) uy(2) using uv(3) by auto
  1521     unfolding setsum_addf and setsum_right_distrib[symmetric] and ux(2) uy(2) using uv(3) by auto
  1522   moreover 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)"
  1522   moreover 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)"
  1523     unfolding scaleR_left_distrib and setsum_addf and scaleR_scaleR[symmetric] and scaleR_right.setsum [symmetric] by auto
  1523     unfolding scaleR_left_distrib and setsum_addf and scaleR_scaleR[symmetric] and scaleR_right.setsum [symmetric] by auto
  1524   ultimately show "\<exists>uc. (\<forall>x\<in>s. 0 \<le> uc x) \<and> setsum uc s = 1 \<and> (\<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)"
  1524   ultimately show "\<exists>uc. (\<forall>x\<in>s. 0 \<le> uc x) \<and> setsum uc s = 1 \<and> (\<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)"
  1525     apply(rule_tac x="\<lambda>x. u * ux x + v * uy x" in exI) by auto 
  1525     apply(rule_tac x="\<lambda>x. u * ux x + v * uy x" in exI) by auto
  1526 next
  1526 next
  1527   fix t assume t:"s \<subseteq> t" "convex t" 
  1527   fix t assume t:"s \<subseteq> t" "convex t"
  1528   fix u assume u:"\<forall>x\<in>s. 0 \<le> u x" "setsum u s = (1::real)"
  1528   fix u assume u:"\<forall>x\<in>s. 0 \<le> u x" "setsum u s = (1::real)"
  1529   thus "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> t" using t(2)[unfolded convex_explicit, THEN spec[where x=s], THEN spec[where x=u]]
  1529   thus "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> t" using t(2)[unfolded convex_explicit, THEN spec[where x=s], THEN spec[where x=u]]
  1530     using assms and t(1) by auto
  1530     using assms and t(1) by auto
  1531 qed
  1531 qed
  1532 
  1532 
  1552     have fin:"finite {1..k}" by auto
  1552     have fin:"finite {1..k}" by auto
  1553     have fin':"\<And>v. finite {i \<in> {1..k}. y i = v}" by auto
  1553     have fin':"\<And>v. finite {i \<in> {1..k}. y i = v}" by auto
  1554     { fix j assume "j\<in>{1..k}"
  1554     { fix j assume "j\<in>{1..k}"
  1555       hence "y j \<in> p" "0 \<le> setsum u {i. Suc 0 \<le> i \<and> i \<le> k \<and> y i = y j}"
  1555       hence "y j \<in> p" "0 \<le> setsum u {i. Suc 0 \<le> i \<and> i \<le> k \<and> y i = y j}"
  1556         using obt(1)[THEN bspec[where x=j]] and obt(2) apply simp
  1556         using obt(1)[THEN bspec[where x=j]] and obt(2) apply simp
  1557         apply(rule setsum_nonneg) using obt(1) by auto } 
  1557         apply(rule setsum_nonneg) using obt(1) by auto }
  1558     moreover
  1558     moreover
  1559     have "(\<Sum>v\<in>y ` {1..k}. setsum u {i \<in> {1..k}. y i = v}) = 1"  
  1559     have "(\<Sum>v\<in>y ` {1..k}. setsum u {i \<in> {1..k}. y i = v}) = 1"
  1560       unfolding setsum_image_gen[OF fin, symmetric] using obt(2) by auto
  1560       unfolding setsum_image_gen[OF fin, symmetric] using obt(2) by auto
  1561     moreover have "(\<Sum>v\<in>y ` {1..k}. setsum u {i \<in> {1..k}. y i = v} *\<^sub>R v) = x"
  1561     moreover have "(\<Sum>v\<in>y ` {1..k}. setsum u {i \<in> {1..k}. y i = v} *\<^sub>R v) = x"
  1562       using setsum_image_gen[OF fin, of "\<lambda>i. u i *\<^sub>R y i" y, symmetric]
  1562       using setsum_image_gen[OF fin, of "\<lambda>i. u i *\<^sub>R y i" y, symmetric]
  1563       unfolding scaleR_left.setsum using obt(3) by auto
  1563       unfolding scaleR_left.setsum using obt(3) by auto
  1564     ultimately 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"
  1564     ultimately 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"
  1568   moreover
  1568   moreover
  1569   { fix y assume "y\<in>?rhs"
  1569   { fix y assume "y\<in>?rhs"
  1570     then obtain s u where 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" by auto
  1570     then obtain s u where 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" by auto
  1571 
  1571 
  1572     obtain f where f:"inj_on f {1..card s}" "f ` {1..card s} = s" using ex_bij_betw_nat_finite_1[OF obt(1)] unfolding bij_betw_def by auto
  1572     obtain f where f:"inj_on f {1..card s}" "f ` {1..card s} = s" using ex_bij_betw_nat_finite_1[OF obt(1)] unfolding bij_betw_def by auto
  1573     
  1573 
  1574     { fix i::nat assume "i\<in>{1..card s}"
  1574     { fix i::nat assume "i\<in>{1..card s}"
  1575       hence "f i \<in> s"  apply(subst f(2)[symmetric]) by auto
  1575       hence "f i \<in> s"  apply(subst f(2)[symmetric]) by auto
  1576       hence "0 \<le> u (f i)" "f i \<in> p" using obt(2,3) by auto  }
  1576       hence "0 \<le> u (f i)" "f i \<in> p" using obt(2,3) by auto  }
  1577     moreover have *:"finite {1..card s}" by auto
  1577     moreover have *:"finite {1..card s}" by auto
  1578     { fix y assume "y\<in>s"
  1578     { fix y assume "y\<in>s"
  1582       hence "(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x)) = u y"
  1582       hence "(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x)) = u y"
  1583             "(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y"
  1583             "(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y"
  1584         by (auto simp add: setsum_constant_scaleR)   }
  1584         by (auto simp add: setsum_constant_scaleR)   }
  1585 
  1585 
  1586     hence "(\<Sum>x = 1..card s. u (f x)) = 1" "(\<Sum>i = 1..card s. u (f i) *\<^sub>R f i) = y"
  1586     hence "(\<Sum>x = 1..card s. u (f x)) = 1" "(\<Sum>i = 1..card s. u (f i) *\<^sub>R f i) = y"
  1587       unfolding setsum_image_gen[OF *(1), of "\<lambda>x. u (f x) *\<^sub>R f x" f] and setsum_image_gen[OF *(1), of "\<lambda>x. u (f x)" f] 
  1587       unfolding setsum_image_gen[OF *(1), of "\<lambda>x. u (f x) *\<^sub>R f x" f] and setsum_image_gen[OF *(1), of "\<lambda>x. u (f x)" f]
  1588       unfolding f using setsum_cong2[of 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"]
  1588       unfolding f using setsum_cong2[of 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"]
  1589       using setsum_cong2 [of s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x))" u] unfolding obt(4,5) by auto
  1589       using setsum_cong2 [of s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x))" u] unfolding obt(4,5) by auto
  1590     
  1590 
  1591     ultimately 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> (\<Sum>i::nat = 1..k. u i *\<^sub>R x i) = y"
  1591     ultimately 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> (\<Sum>i::nat = 1..k. u i *\<^sub>R x i) = y"
  1592       apply(rule_tac x="card s" in exI) apply(rule_tac x="u \<circ> f" in exI) apply(rule_tac x=f in exI) by fastforce
  1592       apply(rule_tac x="card s" in exI) apply(rule_tac x="u \<circ> f" in exI) apply(rule_tac x=f in exI) by fastforce
  1593     hence "y \<in> ?lhs" unfolding convex_hull_indexed by auto  }
  1593     hence "y \<in> ?lhs" unfolding convex_hull_indexed by auto  }
  1594   ultimately show ?thesis unfolding set_eq_iff by blast
  1594   ultimately show ?thesis unfolding set_eq_iff by blast
  1595 qed
  1595 qed
  1670 by (metis affine_affine_hull affine_imp_convex hull_minimal hull_subset)
  1670 by (metis affine_affine_hull affine_imp_convex hull_minimal hull_subset)
  1671 
  1671 
  1672 
  1672 
  1673 lemma affine_dependent_imp_dependent:
  1673 lemma affine_dependent_imp_dependent:
  1674   shows "affine_dependent s \<Longrightarrow> dependent s"
  1674   shows "affine_dependent s \<Longrightarrow> dependent s"
  1675   unfolding affine_dependent_def dependent_def 
  1675   unfolding affine_dependent_def dependent_def
  1676   using affine_hull_subset_span by auto
  1676   using affine_hull_subset_span by auto
  1677 
  1677 
  1678 lemma dependent_imp_affine_dependent:
  1678 lemma dependent_imp_affine_dependent:
  1679   assumes "dependent {x - a| x . x \<in> s}" "a \<notin> s"
  1679   assumes "dependent {x - a| x . x \<in> s}" "a \<notin> s"
  1680   shows "affine_dependent (insert a s)"
  1680   shows "affine_dependent (insert a s)"
  1681 proof-
  1681 proof-
  1682   from assms(1)[unfolded dependent_explicit] obtain S u v 
  1682   from assms(1)[unfolded dependent_explicit] obtain S u v
  1683     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" by auto
  1683     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" by auto
  1684   def t \<equiv> "(\<lambda>x. x + a) ` S"
  1684   def t \<equiv> "(\<lambda>x. x + a) ` S"
  1685 
  1685 
  1686   have inj:"inj_on (\<lambda>x. x + a) S" unfolding inj_on_def by auto
  1686   have inj:"inj_on (\<lambda>x. x + a) S" unfolding inj_on_def by auto
  1687   have "0\<notin>S" using obt(2) assms(2) unfolding subset_eq by auto
  1687   have "0\<notin>S" using obt(2) assms(2) unfolding subset_eq by auto
  1688   have fin:"finite t" and  "t\<subseteq>s" unfolding t_def using obt(1,2) by auto 
  1688   have fin:"finite t" and  "t\<subseteq>s" unfolding t_def using obt(1,2) by auto
  1689 
  1689 
  1690   hence "finite (insert a t)" and "insert a t \<subseteq> insert a s" by auto 
  1690   hence "finite (insert a t)" and "insert a t \<subseteq> insert a s" by auto
  1691   moreover have *:"\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x)) = (\<Sum>x\<in>t. Q x)"
  1691   moreover have *:"\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x)) = (\<Sum>x\<in>t. Q x)"
  1692     apply(rule setsum_cong2) using `a\<notin>s` `t\<subseteq>s` by auto
  1692     apply(rule setsum_cong2) using `a\<notin>s` `t\<subseteq>s` by auto
  1693   have "(\<Sum>x\<in>insert a t. if x = a then - (\<Sum>x\<in>t. u (x - a)) else u (x - a)) = 0"
  1693   have "(\<Sum>x\<in>insert a t. if x = a then - (\<Sum>x\<in>t. u (x - a)) else u (x - a)) = 0"
  1694     unfolding setsum_clauses(2)[OF fin] using `a\<notin>s` `t\<subseteq>s` apply auto unfolding * by auto
  1694     unfolding setsum_clauses(2)[OF fin] using `a\<notin>s` `t\<subseteq>s` apply auto unfolding * by auto
  1695   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"
  1695   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"
  1696     apply(rule_tac x="v + a" in bexI) using obt(3,4) and `0\<notin>S` unfolding t_def by auto
  1696     apply(rule_tac x="v + a" in bexI) using obt(3,4) and `0\<notin>S` unfolding t_def by auto
  1697   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)"
  1697   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)"
  1698     apply(rule setsum_cong2) using `a\<notin>s` `t\<subseteq>s` by auto
  1698     apply(rule setsum_cong2) using `a\<notin>s` `t\<subseteq>s` by auto
  1699   have "(\<Sum>x\<in>t. u (x - a)) *\<^sub>R a = (\<Sum>v\<in>t. u (v - a) *\<^sub>R v)" 
  1699   have "(\<Sum>x\<in>t. u (x - a)) *\<^sub>R a = (\<Sum>v\<in>t. u (v - a) *\<^sub>R v)"
  1700     unfolding scaleR_left.setsum unfolding t_def and setsum_reindex[OF inj] and o_def
  1700     unfolding scaleR_left.setsum unfolding t_def and setsum_reindex[OF inj] and o_def
  1701     using obt(5) by (auto simp add: setsum_addf scaleR_right_distrib)
  1701     using obt(5) by (auto simp add: setsum_addf scaleR_right_distrib)
  1702   hence "(\<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"
  1702   hence "(\<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"
  1703     unfolding setsum_clauses(2)[OF fin] using `a\<notin>s` `t\<subseteq>s` by (auto simp add: *)
  1703     unfolding setsum_clauses(2)[OF fin] using `a\<notin>s` `t\<subseteq>s` by (auto simp add: *)
  1704   ultimately show ?thesis unfolding affine_dependent_explicit
  1704   ultimately show ?thesis unfolding affine_dependent_explicit
  1705     apply(rule_tac x="insert a t" in exI) by auto 
  1705     apply(rule_tac x="insert a t" in exI) by auto
  1706 qed
  1706 qed
  1707 
  1707 
  1708 lemma convex_cone:
  1708 lemma convex_cone:
  1709   "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)" (is "?lhs = ?rhs")
  1709   "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)" (is "?lhs = ?rhs")
  1710 proof-
  1710 proof-
  1720   assumes "finite s" "card s \<ge> DIM('a) + 2"
  1720   assumes "finite s" "card s \<ge> DIM('a) + 2"
  1721   shows "affine_dependent s"
  1721   shows "affine_dependent s"
  1722 proof-
  1722 proof-
  1723   have "s\<noteq>{}" using assms by auto then obtain a where "a\<in>s" by auto
  1723   have "s\<noteq>{}" using assms by auto then obtain a where "a\<in>s" by auto
  1724   have *:"{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})" by auto
  1724   have *:"{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})" by auto
  1725   have "card {x - a |x. x \<in> s - {a}} = card (s - {a})" unfolding * 
  1725   have "card {x - a |x. x \<in> s - {a}} = card (s - {a})" unfolding *
  1726     apply(rule card_image) unfolding inj_on_def by auto
  1726     apply(rule card_image) unfolding inj_on_def by auto
  1727   also have "\<dots> > DIM('a)" using assms(2)
  1727   also have "\<dots> > DIM('a)" using assms(2)
  1728     unfolding card_Diff_singleton[OF assms(1) `a\<in>s`] by auto
  1728     unfolding card_Diff_singleton[OF assms(1) `a\<in>s`] by auto
  1729   finally show ?thesis apply(subst insert_Diff[OF `a\<in>s`, symmetric])
  1729   finally show ?thesis apply(subst insert_Diff[OF `a\<in>s`, symmetric])
  1730     apply(rule dependent_imp_affine_dependent)
  1730     apply(rule dependent_imp_affine_dependent)
  1735   shows "affine_dependent s"
  1735   shows "affine_dependent s"
  1736 proof-
  1736 proof-
  1737   from assms(2) have "s \<noteq> {}" by auto
  1737   from assms(2) have "s \<noteq> {}" by auto
  1738   then obtain a where "a\<in>s" by auto
  1738   then obtain a where "a\<in>s" by auto
  1739   have *:"{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})" by auto
  1739   have *:"{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})" by auto
  1740   have **:"card {x - a |x. x \<in> s - {a}} = card (s - {a})" unfolding * 
  1740   have **:"card {x - a |x. x \<in> s - {a}} = card (s - {a})" unfolding *
  1741     apply(rule card_image) unfolding inj_on_def by auto
  1741     apply(rule card_image) unfolding inj_on_def by auto
  1742   have "dim {x - a |x. x \<in> s - {a}} \<le> dim s"
  1742   have "dim {x - a |x. x \<in> s - {a}} \<le> dim s"
  1743     apply(rule subset_le_dim) unfolding subset_eq
  1743     apply(rule subset_le_dim) unfolding subset_eq
  1744     using `a\<in>s` by (auto simp add:span_superset span_sub)
  1744     using `a\<in>s` by (auto simp add:span_superset span_sub)
  1745   also have "\<dots> < dim s + 1" by auto
  1745   also have "\<dots> < dim s + 1" by auto
  1775         using as[THEN bspec[where x=v]] and `v\<in>s` using `w v \<noteq> 0` by auto
  1775         using as[THEN bspec[where x=v]] and `v\<in>s` using `w v \<noteq> 0` by auto
  1776       thus False using wv(1) by auto
  1776       thus False using wv(1) by auto
  1777     qed hence "i\<noteq>{}" unfolding i_def by auto
  1777     qed hence "i\<noteq>{}" unfolding i_def by auto
  1778 
  1778 
  1779     hence "t \<ge> 0" using Min_ge_iff[of i 0 ] and obt(1) unfolding t_def i_def
  1779     hence "t \<ge> 0" using Min_ge_iff[of i 0 ] and obt(1) unfolding t_def i_def
  1780       using obt(4)[unfolded le_less] apply auto unfolding divide_le_0_iff by auto 
  1780       using obt(4)[unfolded le_less] apply auto unfolding divide_le_0_iff by auto
  1781     have t:"\<forall>v\<in>s. u v + t * w v \<ge> 0" proof
  1781     have t:"\<forall>v\<in>s. u v + t * w v \<ge> 0" proof
  1782       fix v assume "v\<in>s" hence v:"0\<le>u v" using obt(4)[THEN bspec[where x=v]] by auto
  1782       fix v assume "v\<in>s" hence v:"0\<le>u v" using obt(4)[THEN bspec[where x=v]] by auto
  1783       show"0 \<le> u v + t * w v" proof(cases "w v < 0")
  1783       show"0 \<le> u v + t * w v" proof(cases "w v < 0")
  1784         case False thus ?thesis apply(rule_tac add_nonneg_nonneg) 
  1784         case False thus ?thesis apply(rule_tac add_nonneg_nonneg)
  1785           using v apply simp apply(rule mult_nonneg_nonneg) using `t\<ge>0` by auto next
  1785           using v apply simp apply(rule mult_nonneg_nonneg) using `t\<ge>0` by auto next
  1786         case True hence "t \<le> u v / (- w v)" using `v\<in>s`
  1786         case True hence "t \<le> u v / (- w v)" using `v\<in>s`
  1787           unfolding t_def i_def apply(rule_tac Min_le) using obt(1) by auto 
  1787           unfolding t_def i_def apply(rule_tac Min_le) using obt(1) by auto
  1788         thus ?thesis unfolding real_0_le_add_iff
  1788         thus ?thesis unfolding real_0_le_add_iff
  1789           using pos_le_divide_eq[OF True[unfolded neg_0_less_iff_less[symmetric]]] by auto
  1789           using pos_le_divide_eq[OF True[unfolded neg_0_less_iff_less[symmetric]]] by auto
  1790       qed qed
  1790       qed qed
  1791 
  1791 
  1792     obtain a where "a\<in>s" and "t = (\<lambda>v. (u v) / (- w v)) a" and "w a < 0"
  1792     obtain a where "a\<in>s" and "t = (\<lambda>v. (u v) / (- w v)) a" and "w a < 0"
  1793       using Min_in[OF _ `i\<noteq>{}`] and obt(1) unfolding i_def t_def by auto
  1793       using Min_in[OF _ `i\<noteq>{}`] and obt(1) unfolding i_def t_def by auto
  1794     hence a:"a\<in>s" "u a + t * w a = 0" by auto
  1794     hence a:"a\<in>s" "u a + t * w a = 0" by auto
  1795     have *:"\<And>f. setsum f (s - {a}) = setsum f s - ((f a)::'b::ab_group_add)"
  1795     have *:"\<And>f. setsum f (s - {a}) = setsum f s - ((f a)::'b::ab_group_add)"
  1796       unfolding setsum_diff1'[OF obt(1) `a\<in>s`] by auto 
  1796       unfolding setsum_diff1'[OF obt(1) `a\<in>s`] by auto
  1797     have "(\<Sum>v\<in>s. u v + t * w v) = 1"
  1797     have "(\<Sum>v\<in>s. u v + t * w v) = 1"
  1798       unfolding setsum_addf wv(1) setsum_right_distrib[symmetric] obt(5) by auto
  1798       unfolding setsum_addf wv(1) setsum_right_distrib[symmetric] obt(5) by auto
  1799     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" 
  1799     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"
  1800       unfolding setsum_addf obt(6) scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] wv(4)
  1800       unfolding setsum_addf obt(6) scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] wv(4)
  1801       using a(2) [THEN eq_neg_iff_add_eq_0 [THEN iffD2]] by simp
  1801       using a(2) [THEN eq_neg_iff_add_eq_0 [THEN iffD2]] by simp
  1802     ultimately have "?P (n - 1)" apply(rule_tac x="(s - {a})" in exI)
  1802     ultimately have "?P (n - 1)" apply(rule_tac x="(s - {a})" in exI)
  1803       apply(rule_tac x="\<lambda>v. u v + t * w v" in exI) using obt(1-3) and t and a
  1803       apply(rule_tac x="\<lambda>v. u v + t * w v" in exI) using obt(1-3) and t and a
  1804       by (auto simp add: * scaleR_left_distrib)
  1804       by (auto simp add: * scaleR_left_distrib)
  1838 proof-
  1838 proof-
  1839 have "affine ((%x. a + x) ` (affine hull S))" using affine_translation affine_affine_hull by auto
  1839 have "affine ((%x. a + x) ` (affine hull S))" using affine_translation affine_affine_hull by auto
  1840 moreover have "(%x. a + x) `  S <= (%x. a + x) ` (affine hull S)" using hull_subset[of S] by auto
  1840 moreover have "(%x. a + x) `  S <= (%x. a + x) ` (affine hull S)" using hull_subset[of S] by auto
  1841 ultimately have h1: "affine hull ((%x. a + x) `  S) <= (%x. a + x) ` (affine hull S)" by (metis hull_minimal)
  1841 ultimately have h1: "affine hull ((%x. a + x) `  S) <= (%x. a + x) ` (affine hull S)" by (metis hull_minimal)
  1842 have "affine((%x. -a + x) ` (affine hull ((%x. a + x) `  S)))"  using affine_translation affine_affine_hull by auto
  1842 have "affine((%x. -a + x) ` (affine hull ((%x. a + x) `  S)))"  using affine_translation affine_affine_hull by auto
  1843 moreover have "(%x. -a + x) ` (%x. a + x) `  S <= (%x. -a + x) ` (affine hull ((%x. a + x) `  S))" using hull_subset[of "(%x. a + x) `  S"] by auto 
  1843 moreover have "(%x. -a + x) ` (%x. a + x) `  S <= (%x. -a + x) ` (affine hull ((%x. a + x) `  S))" using hull_subset[of "(%x. a + x) `  S"] by auto
  1844 moreover have "S=(%x. -a + x) ` (%x. a + x) `  S" using  translation_assoc[of "-a" a] by auto
  1844 moreover have "S=(%x. -a + x) ` (%x. a + x) `  S" using  translation_assoc[of "-a" a] by auto
  1845 ultimately have "(%x. -a + x) ` (affine hull ((%x. a + x) `  S)) >= (affine hull S)" by (metis hull_minimal)
  1845 ultimately have "(%x. -a + x) ` (affine hull ((%x. a + x) `  S)) >= (affine hull S)" by (metis hull_minimal)
  1846 hence "affine hull ((%x. a + x) `  S) >= (%x. a + x) ` (affine hull S)" by auto
  1846 hence "affine hull ((%x. a + x) `  S) >= (%x. a + x) ` (affine hull S)" by auto
  1847 from this show ?thesis using h1 by auto
  1847 from this show ?thesis using h1 by auto
  1848 qed
  1848 qed
  1852   shows "affine_dependent ((%x. a + x) ` S)"
  1852   shows "affine_dependent ((%x. a + x) ` S)"
  1853 proof-
  1853 proof-
  1854 obtain x where x_def: "x : S & x : affine hull (S - {x})" using assms affine_dependent_def by auto
  1854 obtain x where x_def: "x : S & x : affine hull (S - {x})" using assms affine_dependent_def by auto
  1855 have "op + a ` (S - {x}) = op + a ` S - {a + x}" by auto
  1855 have "op + a ` (S - {x}) = op + a ` S - {a + x}" by auto
  1856 hence "a+x : affine hull ((%x. a + x) ` S - {a+x})" using  affine_hull_translation[of a "S-{x}"] x_def by auto
  1856 hence "a+x : affine hull ((%x. a + x) ` S - {a+x})" using  affine_hull_translation[of a "S-{x}"] x_def by auto
  1857 moreover have "a+x : (%x. a + x) ` S" using x_def by auto  
  1857 moreover have "a+x : (%x. a + x) ` S" using x_def by auto
  1858 ultimately show ?thesis unfolding affine_dependent_def by auto 
  1858 ultimately show ?thesis unfolding affine_dependent_def by auto
  1859 qed
  1859 qed
  1860 
  1860 
  1861 lemma affine_dependent_translation_eq:
  1861 lemma affine_dependent_translation_eq:
  1862   "affine_dependent S <-> affine_dependent ((%x. a + x) ` S)"
  1862   "affine_dependent S <-> affine_dependent ((%x. a + x) ` S)"
  1863 proof-
  1863 proof-
  1864 { assume "affine_dependent ((%x. a + x) ` S)" 
  1864 { assume "affine_dependent ((%x. a + x) ` S)"
  1865   hence "affine_dependent S" using affine_dependent_translation[of "((%x. a + x) ` S)" "-a"] translation_assoc[of "-a" a] by auto  
  1865   hence "affine_dependent S" using affine_dependent_translation[of "((%x. a + x) ` S)" "-a"] translation_assoc[of "-a" a] by auto
  1866 } from this show ?thesis using affine_dependent_translation by auto
  1866 } from this show ?thesis using affine_dependent_translation by auto
  1867 qed
  1867 qed
  1868 
  1868 
  1869 lemma affine_hull_0_dependent:
  1869 lemma affine_hull_0_dependent:
  1870   assumes "0 : affine hull S"
  1870   assumes "0 : affine hull S"
  1871   shows "dependent S"
  1871   shows "dependent S"
  1872 proof-
  1872 proof-
  1873 obtain s u where s_u_def: "finite s & s ~= {} & s <= S & setsum u s = 1 & (SUM v:s. u v *\<^sub>R v) = 0" using assms affine_hull_explicit[of S] by auto
  1873 obtain s u where s_u_def: "finite s & s ~= {} & s <= S & setsum u s = 1 & (SUM v:s. u v *\<^sub>R v) = 0" using assms affine_hull_explicit[of S] by auto
  1874 hence "EX v:s. u v ~= 0" using setsum_not_0[of "u" "s"] by auto 
  1874 hence "EX v:s. u v ~= 0" using setsum_not_0[of "u" "s"] by auto
  1875 hence "finite s & s <= S & (EX v:s. u v ~= 0 & (SUM v:s. u v *\<^sub>R v) = 0)" using s_u_def by auto
  1875 hence "finite s & s <= S & (EX v:s. u v ~= 0 & (SUM v:s. u v *\<^sub>R v) = 0)" using s_u_def by auto
  1876 from this show ?thesis unfolding dependent_explicit[of S] by auto
  1876 from this show ?thesis unfolding dependent_explicit[of S] by auto
  1877 qed
  1877 qed
  1878 
  1878 
  1879 lemma affine_dependent_imp_dependent2:
  1879 lemma affine_dependent_imp_dependent2:
  1885 moreover have "span (insert 0 S - {x}) = span (S - {x})" using insert_Diff_if[of "0" S "{x}"] span_insert_0[of "S-{x}"] by auto
  1885 moreover have "span (insert 0 S - {x}) = span (S - {x})" using insert_Diff_if[of "0" S "{x}"] span_insert_0[of "S-{x}"] by auto
  1886 ultimately have "x : span (S - {x})" by auto
  1886 ultimately have "x : span (S - {x})" by auto
  1887 hence "(x~=0) ==> dependent S" using x_def dependent_def by auto
  1887 hence "(x~=0) ==> dependent S" using x_def dependent_def by auto
  1888 moreover
  1888 moreover
  1889 { assume "x=0" hence "0 : affine hull S" using x_def hull_mono[of "S - {0}" S] by auto
  1889 { assume "x=0" hence "0 : affine hull S" using x_def hull_mono[of "S - {0}" S] by auto
  1890                hence "dependent S" using affine_hull_0_dependent by auto  
  1890                hence "dependent S" using affine_hull_0_dependent by auto
  1891 } ultimately show ?thesis by auto
  1891 } ultimately show ?thesis by auto
  1892 qed
  1892 qed
  1893 
  1893 
  1894 lemma affine_dependent_iff_dependent:
  1894 lemma affine_dependent_iff_dependent:
  1895   assumes "a ~: S"
  1895   assumes "a ~: S"
  1896   shows "affine_dependent (insert a S) <-> dependent ((%x. -a + x) ` S)" 
  1896   shows "affine_dependent (insert a S) <-> dependent ((%x. -a + x) ` S)"
  1897 proof-
  1897 proof-
  1898 have "(op + (- a) ` S)={x - a| x . x : S}" by auto
  1898 have "(op + (- a) ` S)={x - a| x . x : S}" by auto
  1899 from this show ?thesis using affine_dependent_translation_eq[of "(insert a S)" "-a"] 
  1899 from this show ?thesis using affine_dependent_translation_eq[of "(insert a S)" "-a"]
  1900       affine_dependent_imp_dependent2 assms 
  1900       affine_dependent_imp_dependent2 assms
  1901       dependent_imp_affine_dependent[of a S] by auto
  1901       dependent_imp_affine_dependent[of a S] by auto
  1902 qed
  1902 qed
  1903 
  1903 
  1904 lemma affine_dependent_iff_dependent2:
  1904 lemma affine_dependent_iff_dependent2:
  1905   assumes "a : S"
  1905   assumes "a : S"
  1906   shows "affine_dependent S <-> dependent ((%x. -a + x) ` (S-{a}))"
  1906   shows "affine_dependent S <-> dependent ((%x. -a + x) ` (S-{a}))"
  1907 proof-
  1907 proof-
  1908 have "insert a (S - {a})=S" using assms by auto
  1908 have "insert a (S - {a})=S" using assms by auto
  1909 from this show ?thesis using assms affine_dependent_iff_dependent[of a "S-{a}"] by auto 
  1909 from this show ?thesis using assms affine_dependent_iff_dependent[of a "S-{a}"] by auto
  1910 qed
  1910 qed
  1911 
  1911 
  1912 lemma affine_hull_insert_span_gen:
  1912 lemma affine_hull_insert_span_gen:
  1913   shows "affine hull (insert a s) = (%x. a+x) ` span ((%x. -a+x) ` s)" 
  1913   shows "affine hull (insert a s) = (%x. a+x) ` span ((%x. -a+x) ` s)"
  1914 proof-
  1914 proof-
  1915 have h1: "{x - a |x. x : s}=((%x. -a+x) ` s)" by auto
  1915 have h1: "{x - a |x. x : s}=((%x. -a+x) ` s)" by auto
  1916 { assume "a ~: s" hence ?thesis using affine_hull_insert_span[of a s] h1 by auto}  
  1916 { assume "a ~: s" hence ?thesis using affine_hull_insert_span[of a s] h1 by auto}
  1917 moreover
  1917 moreover
  1918 { assume a1: "a : s"
  1918 { assume a1: "a : s"
  1919   have "EX x. x:s & -a+x=0" apply (rule exI[of _ a]) using a1 by auto
  1919   have "EX x. x:s & -a+x=0" apply (rule exI[of _ a]) using a1 by auto
  1920   hence "insert 0 ((%x. -a+x) ` (s - {a}))=(%x. -a+x) ` s" by auto
  1920   hence "insert 0 ((%x. -a+x) ` (s - {a}))=(%x. -a+x) ` s" by auto
  1921   hence "span ((%x. -a+x) ` (s - {a}))=span ((%x. -a+x) ` s)" 
  1921   hence "span ((%x. -a+x) ` (s - {a}))=span ((%x. -a+x) ` s)"
  1922     using span_insert_0[of "op + (- a) ` (s - {a})"] by auto
  1922     using span_insert_0[of "op + (- a) ` (s - {a})"] by auto
  1923   moreover have "{x - a |x. x : (s - {a})}=((%x. -a+x) ` (s - {a}))" by auto 
  1923   moreover have "{x - a |x. x : (s - {a})}=((%x. -a+x) ` (s - {a}))" by auto
  1924   moreover have "insert a (s - {a})=(insert a s)" using assms by auto
  1924   moreover have "insert a (s - {a})=(insert a s)" using assms by auto
  1925   ultimately have ?thesis using assms affine_hull_insert_span[of "a" "s-{a}"] by auto
  1925   ultimately have ?thesis using assms affine_hull_insert_span[of "a" "s-{a}"] by auto
  1926 } 
  1926 }
  1927 ultimately show ?thesis by auto  
  1927 ultimately show ?thesis by auto
  1928 qed
  1928 qed
  1929 
  1929 
  1930 lemma affine_hull_span2:
  1930 lemma affine_hull_span2:
  1931   assumes "a : s"
  1931   assumes "a : s"
  1932   shows "affine hull s = (%x. a+x) ` span ((%x. -a+x) ` (s-{a}))"
  1932   shows "affine hull s = (%x. a+x) ` span ((%x. -a+x) ` (s-{a}))"
  1951 assumes "~(affine_dependent S)" "S <= V" "S~={}"
  1951 assumes "~(affine_dependent S)" "S <= V" "S~={}"
  1952 shows "? T. ~(affine_dependent T) & S<=T & T<=V & affine hull T = affine hull V"
  1952 shows "? T. ~(affine_dependent T) & S<=T & T<=V & affine hull T = affine hull V"
  1953 proof-
  1953 proof-
  1954 obtain a where a_def: "a : S" using assms by auto
  1954 obtain a where a_def: "a : S" using assms by auto
  1955 hence h0: "independent  ((%x. -a + x) ` (S-{a}))" using affine_dependent_iff_dependent2 assms by auto
  1955 hence h0: "independent  ((%x. -a + x) ` (S-{a}))" using affine_dependent_iff_dependent2 assms by auto
  1956 from this obtain B 
  1956 from this obtain B
  1957    where B_def: "(%x. -a+x) ` (S - {a}) <= B & B <= (%x. -a+x) ` V & independent B & (%x. -a+x) ` V <= span B" 
  1957    where B_def: "(%x. -a+x) ` (S - {a}) <= B & B <= (%x. -a+x) ` V & independent B & (%x. -a+x) ` V <= span B"
  1958    using maximal_independent_subset_extend[of "(%x. -a+x) ` (S-{a})" "(%x. -a + x) ` V"] assms by blast
  1958    using maximal_independent_subset_extend[of "(%x. -a+x) ` (S-{a})" "(%x. -a + x) ` V"] assms by blast
  1959 def T == "(%x. a+x) ` (insert 0 B)" hence "T=insert a ((%x. a+x) ` B)" by auto
  1959 def T == "(%x. a+x) ` (insert 0 B)" hence "T=insert a ((%x. a+x) ` B)" by auto
  1960 hence "affine hull T = (%x. a+x) ` span B" using affine_hull_insert_span_gen[of a "((%x. a+x) ` B)"] translation_assoc[of "-a" a B] by auto
  1960 hence "affine hull T = (%x. a+x) ` span B" using affine_hull_insert_span_gen[of a "((%x. a+x) ` B)"] translation_assoc[of "-a" a B] by auto
  1961 hence "V <= affine hull T" using B_def assms translation_inverse_subset[of a V "span B"] by auto
  1961 hence "V <= affine hull T" using B_def assms translation_inverse_subset[of a V "span B"] by auto
  1962 moreover have "T<=V" using T_def B_def a_def assms by auto
  1962 moreover have "T<=V" using T_def B_def a_def assms by auto
  1963 ultimately have "affine hull T = affine hull V" 
  1963 ultimately have "affine hull T = affine hull V"
  1964     by (metis Int_absorb1 Int_absorb2 hull_hull hull_mono)
  1964     by (metis Int_absorb1 Int_absorb2 hull_hull hull_mono)
  1965 moreover have "S<=T" using T_def B_def translation_inverse_subset[of a "S-{a}" B] by auto
  1965 moreover have "S<=T" using T_def B_def translation_inverse_subset[of a "S-{a}" B] by auto
  1966 moreover have "~(affine_dependent T)" using T_def affine_dependent_translation_eq[of "insert 0 B"] affine_dependent_imp_dependent2 B_def by auto
  1966 moreover have "~(affine_dependent T)" using T_def affine_dependent_translation_eq[of "insert 0 B"] affine_dependent_imp_dependent2 B_def by auto
  1967 ultimately show ?thesis using `T<=V` by auto
  1967 ultimately show ?thesis using `T<=V` by auto
  1968 qed
  1968 qed
  1969 
  1969 
  1970 lemma affine_basis_exists: 
  1970 lemma affine_basis_exists:
  1971 fixes V :: "('n::euclidean_space) set"
  1971 fixes V :: "('n::euclidean_space) set"
  1972 shows "? B. B <= V & ~(affine_dependent B) & affine hull V = affine hull B"
  1972 shows "? B. B <= V & ~(affine_dependent B) & affine hull V = affine hull B"
  1973 proof-
  1973 proof-
  1974 { assume empt: "V={}" have "? B. B <= V & ~(affine_dependent B) & (affine hull V=affine hull B)" using empt affine_independent_empty by auto
  1974 { assume empt: "V={}" have "? B. B <= V & ~(affine_dependent B) & (affine hull V=affine hull B)" using empt affine_independent_empty by auto
  1975 }
  1975 }
  1984 subsection {* Affine Dimension of a Set *}
  1984 subsection {* Affine Dimension of a Set *}
  1985 
  1985 
  1986 definition "aff_dim V = (SOME d :: int. ? B. (affine hull B=affine hull V) & ~(affine_dependent B) & (of_nat(card B) = d+1))"
  1986 definition "aff_dim V = (SOME d :: int. ? B. (affine hull B=affine hull V) & ~(affine_dependent B) & (of_nat(card B) = d+1))"
  1987 
  1987 
  1988 lemma aff_dim_basis_exists:
  1988 lemma aff_dim_basis_exists:
  1989   fixes V :: "('n::euclidean_space) set" 
  1989   fixes V :: "('n::euclidean_space) set"
  1990   shows "? B. (affine hull B=affine hull V) & ~(affine_dependent B) & (of_nat(card B) = aff_dim V+1)"
  1990   shows "? B. (affine hull B=affine hull V) & ~(affine_dependent B) & (of_nat(card B) = aff_dim V+1)"
  1991 proof-
  1991 proof-
  1992 obtain B where B_def: "~(affine_dependent B) & (affine hull B=affine hull V)" using affine_basis_exists[of V] by auto
  1992 obtain B where B_def: "~(affine_dependent B) & (affine hull B=affine hull V)" using affine_basis_exists[of V] by auto
  1993 from this show ?thesis unfolding aff_dim_def some_eq_ex[of "%d. ? (B :: ('n::euclidean_space) set). (affine hull B=affine hull V) & ~(affine_dependent B) & (of_nat(card B) = d+1)"] apply auto apply (rule exI[of _ "int (card B)-(1 :: int)"]) apply (rule exI[of _ "B"]) by auto
  1993 from this show ?thesis unfolding aff_dim_def some_eq_ex[of "%d. ? (B :: ('n::euclidean_space) set). (affine hull B=affine hull V) & ~(affine_dependent B) & (of_nat(card B) = d+1)"] apply auto apply (rule exI[of _ "int (card B)-(1 :: int)"]) apply (rule exI[of _ "B"]) by auto
  1994 qed
  1994 qed
  1995 
  1995 
  1996 lemma affine_hull_nonempty: "(S ~= {}) <-> affine hull S ~= {}"
  1996 lemma affine_hull_nonempty: "(S ~= {}) <-> affine hull S ~= {}"
  1997 proof-
  1997 proof-
  1998 have "(S = {}) ==> affine hull S = {}"using affine_hull_empty by auto 
  1998 have "(S = {}) ==> affine hull S = {}"using affine_hull_empty by auto
  1999 moreover have "affine hull S = {} ==> S = {}" unfolding hull_def by auto
  1999 moreover have "affine hull S = {} ==> S = {}" unfolding hull_def by auto
  2000 ultimately show "(S ~= {}) <-> affine hull S ~= {}" by blast
  2000 ultimately show "(S ~= {}) <-> affine hull S ~= {}" by blast
  2001 qed
  2001 qed
  2002 
  2002 
  2003 lemma aff_dim_parallel_subspace_aux:
  2003 lemma aff_dim_parallel_subspace_aux:
  2004 fixes B :: "('n::euclidean_space) set"
  2004 fixes B :: "('n::euclidean_space) set"
  2005 assumes "~(affine_dependent B)" "a:B"
  2005 assumes "~(affine_dependent B)" "a:B"
  2006 shows "finite B & ((card B) - 1 = dim (span ((%x. -a+x) ` (B-{a}))))" 
  2006 shows "finite B & ((card B) - 1 = dim (span ((%x. -a+x) ` (B-{a}))))"
  2007 proof-
  2007 proof-
  2008 have "independent ((%x. -a + x) ` (B-{a}))" using affine_dependent_iff_dependent2 assms by auto
  2008 have "independent ((%x. -a + x) ` (B-{a}))" using affine_dependent_iff_dependent2 assms by auto
  2009 hence fin: "dim (span ((%x. -a+x) ` (B-{a}))) = card ((%x. -a + x) ` (B-{a}))" "finite ((%x. -a + x) ` (B - {a}))"  using indep_card_eq_dim_span[of "(%x. -a+x) ` (B-{a})"] by auto
  2009 hence fin: "dim (span ((%x. -a+x) ` (B-{a}))) = card ((%x. -a + x) ` (B-{a}))" "finite ((%x. -a + x) ` (B - {a}))"  using indep_card_eq_dim_span[of "(%x. -a+x) ` (B-{a})"] by auto
  2010 { assume emp: "(%x. -a + x) ` (B - {a}) = {}" 
  2010 { assume emp: "(%x. -a + x) ` (B - {a}) = {}"
  2011   have "B=insert a ((%x. a + x) ` (%x. -a + x) ` (B - {a}))" using translation_assoc[of "a" "-a" "(B - {a})"] assms by auto
  2011   have "B=insert a ((%x. a + x) ` (%x. -a + x) ` (B - {a}))" using translation_assoc[of "a" "-a" "(B - {a})"] assms by auto
  2012   hence "B={a}" using emp by auto
  2012   hence "B={a}" using emp by auto
  2013   hence ?thesis using assms fin by auto  
  2013   hence ?thesis using assms fin by auto
  2014 }
  2014 }
  2015 moreover
  2015 moreover
  2016 { assume "(%x. -a + x) ` (B - {a}) ~= {}"
  2016 { assume "(%x. -a + x) ` (B - {a}) ~= {}"
  2017   hence "card ((%x. -a + x) ` (B - {a}))>0" using fin by auto
  2017   hence "card ((%x. -a + x) ` (B - {a}))>0" using fin by auto
  2018   moreover have h1: "card ((%x. -a + x) ` (B-{a})) = card (B-{a})"  
  2018   moreover have h1: "card ((%x. -a + x) ` (B-{a})) = card (B-{a})"
  2019      apply (rule card_image) using translate_inj_on by auto
  2019      apply (rule card_image) using translate_inj_on by auto
  2020   ultimately have "card (B-{a})>0" by auto
  2020   ultimately have "card (B-{a})>0" by auto
  2021   hence "finite(B-{a})" using card_gt_0_iff[of "(B - {a})"] by auto
  2021   hence "finite(B-{a})" using card_gt_0_iff[of "(B - {a})"] by auto
  2022   moreover hence "(card (B-{a})= (card B) - 1)" using card_Diff_singleton assms by auto
  2022   moreover hence "(card (B-{a})= (card B) - 1)" using card_Diff_singleton assms by auto
  2023   ultimately have ?thesis using fin h1 by auto
  2023   ultimately have ?thesis using fin h1 by auto
  2029 assumes "V ~= {}"
  2029 assumes "V ~= {}"
  2030 assumes "subspace L" "affine_parallel (affine hull V) L"
  2030 assumes "subspace L" "affine_parallel (affine hull V) L"
  2031 shows "aff_dim V=int(dim L)"
  2031 shows "aff_dim V=int(dim L)"
  2032 proof-
  2032 proof-
  2033 obtain B where B_def: "affine hull B = affine hull V & ~ affine_dependent B & int (card B) = aff_dim V + 1" using aff_dim_basis_exists by auto
  2033 obtain B where B_def: "affine hull B = affine hull V & ~ affine_dependent B & int (card B) = aff_dim V + 1" using aff_dim_basis_exists by auto
  2034 hence "B~={}" using assms B_def  affine_hull_nonempty[of V] affine_hull_nonempty[of B] by auto 
  2034 hence "B~={}" using assms B_def  affine_hull_nonempty[of V] affine_hull_nonempty[of B] by auto
  2035 from this obtain a where a_def: "a : B" by auto
  2035 from this obtain a where a_def: "a : B" by auto
  2036 def Lb == "span ((%x. -a+x) ` (B-{a}))"
  2036 def Lb == "span ((%x. -a+x) ` (B-{a}))"
  2037   moreover have "affine_parallel (affine hull B) Lb"
  2037   moreover have "affine_parallel (affine hull B) Lb"
  2038      using Lb_def B_def assms affine_hull_span2[of a B] a_def  affine_parallel_commut[of "Lb" "(affine hull B)"] unfolding affine_parallel_def by auto
  2038      using Lb_def B_def assms affine_hull_span2[of a B] a_def  affine_parallel_commut[of "Lb" "(affine hull B)"] unfolding affine_parallel_def by auto
  2039   moreover have "subspace Lb" using Lb_def subspace_span by auto
  2039   moreover have "subspace Lb" using Lb_def subspace_span by auto
  2040   moreover have "affine hull B ~= {}" using assms B_def affine_hull_nonempty[of V] by auto
  2040   moreover have "affine hull B ~= {}" using assms B_def affine_hull_nonempty[of V] by auto
  2041   ultimately have "L=Lb" using assms affine_parallel_subspace[of "affine hull B"] affine_affine_hull[of B] B_def by auto 
  2041   ultimately have "L=Lb" using assms affine_parallel_subspace[of "affine hull B"] affine_affine_hull[of B] B_def by auto
  2042   hence "dim L=dim Lb" by auto 
  2042   hence "dim L=dim Lb" by auto
  2043   moreover have "(card B) - 1 = dim Lb" "finite B" using Lb_def aff_dim_parallel_subspace_aux a_def B_def by auto
  2043   moreover have "(card B) - 1 = dim Lb" "finite B" using Lb_def aff_dim_parallel_subspace_aux a_def B_def by auto
  2044 (*  hence "card B=dim Lb+1" using `B~={}` card_gt_0_iff[of B] by auto *)
  2044 (*  hence "card B=dim Lb+1" using `B~={}` card_gt_0_iff[of B] by auto *)
  2045   ultimately show ?thesis using B_def `B~={}` card_gt_0_iff[of B] by auto
  2045   ultimately show ?thesis using B_def `B~={}` card_gt_0_iff[of B] by auto
  2046 qed
  2046 qed
  2047 
  2047 
  2048 lemma aff_independent_finite:
  2048 lemma aff_independent_finite:
  2049 fixes B :: "('n::euclidean_space) set"
  2049 fixes B :: "('n::euclidean_space) set"
  2050 assumes "~(affine_dependent B)"
  2050 assumes "~(affine_dependent B)"
  2051 shows "finite B"
  2051 shows "finite B"
  2052 proof-
  2052 proof-
  2053 { assume "B~={}" from this obtain a where "a:B" by auto 
  2053 { assume "B~={}" from this obtain a where "a:B" by auto
  2054   hence ?thesis using aff_dim_parallel_subspace_aux assms by auto 
  2054   hence ?thesis using aff_dim_parallel_subspace_aux assms by auto
  2055 } from this show ?thesis by auto
  2055 } from this show ?thesis by auto
  2056 qed
  2056 qed
  2057 
  2057 
  2058 lemma independent_finite:
  2058 lemma independent_finite:
  2059 fixes B :: "('n::euclidean_space) set"
  2059 fixes B :: "('n::euclidean_space) set"
  2060 assumes "independent B" 
  2060 assumes "independent B"
  2061 shows "finite B"
  2061 shows "finite B"
  2062 using affine_dependent_imp_dependent[of B] aff_independent_finite[of B] assms by auto
  2062 using affine_dependent_imp_dependent[of B] aff_independent_finite[of B] assms by auto
  2063 
  2063 
  2064 lemma subspace_dim_equal:
  2064 lemma subspace_dim_equal:
  2065 assumes "subspace (S :: ('n::euclidean_space) set)" "subspace T" "S <= T" "dim S >= dim T"
  2065 assumes "subspace (S :: ('n::euclidean_space) set)" "subspace T" "S <= T" "dim S >= dim T"
  2066 shows "S=T"
  2066 shows "S=T"
  2067 proof- 
  2067 proof-
  2068 obtain B where B_def: "B <= S & independent B & S <= span B & (card B = dim S)" using basis_exists[of S] by auto
  2068 obtain B where B_def: "B <= S & independent B & S <= span B & (card B = dim S)" using basis_exists[of S] by auto
  2069 hence "span B <= S" using span_mono[of B S] span_eq[of S] assms by metis 
  2069 hence "span B <= S" using span_mono[of B S] span_eq[of S] assms by metis
  2070 hence "span B = S" using B_def by auto
  2070 hence "span B = S" using B_def by auto
  2071 have "dim S = dim T" using assms dim_subset[of S T] by auto
  2071 have "dim S = dim T" using assms dim_subset[of S T] by auto
  2072 hence "T <= span B" using card_eq_dim[of B T] B_def independent_finite assms by auto
  2072 hence "T <= span B" using card_eq_dim[of B T] B_def independent_finite assms by auto
  2073 from this show ?thesis using assms `span B=S` by auto
  2073 from this show ?thesis using assms `span B=S` by auto
  2074 qed
  2074 qed
  2078   (is "span ?A = ?B")
  2078   (is "span ?A = ?B")
  2079 proof-
  2079 proof-
  2080 have "?A <= ?B" by auto
  2080 have "?A <= ?B" by auto
  2081 moreover have s: "subspace ?B" using subspace_substandard[of "%i. i ~: d"] .
  2081 moreover have s: "subspace ?B" using subspace_substandard[of "%i. i ~: d"] .
  2082 ultimately have "span ?A <= ?B" using span_mono[of "?A" "?B"] span_eq[of "?B"] by blast
  2082 ultimately have "span ?A <= ?B" using span_mono[of "?A" "?B"] span_eq[of "?B"] by blast
  2083 moreover have "card d <= dim (span ?A)" using independent_card_le_dim[of "?A" "span ?A"] 
  2083 moreover have "card d <= dim (span ?A)" using independent_card_le_dim[of "?A" "span ?A"]
  2084    independent_substdbasis[OF assms] card_substdbasis[OF assms] span_inc[of "?A"] by auto
  2084    independent_substdbasis[OF assms] card_substdbasis[OF assms] span_inc[of "?A"] by auto
  2085 moreover hence "dim ?B <= dim (span ?A)" using dim_substandard[OF assms] by auto
  2085 moreover hence "dim ?B <= dim (span ?A)" using dim_substandard[OF assms] by auto
  2086 ultimately show ?thesis using s subspace_dim_equal[of "span ?A" "?B"] 
  2086 ultimately show ?thesis using s subspace_dim_equal[of "span ?A" "?B"]
  2087   subspace_span[of "?A"] by auto
  2087   subspace_span[of "?A"] by auto
  2088 qed
  2088 qed
  2089 
  2089 
  2090 lemma basis_to_substdbasis_subspace_isomorphism:
  2090 lemma basis_to_substdbasis_subspace_isomorphism:
  2091 fixes B :: "('a::euclidean_space) set" 
  2091 fixes B :: "('a::euclidean_space) set"
  2092 assumes "independent B"
  2092 assumes "independent B"
  2093 shows "EX f d. card d = card B & linear f & f ` B = {basis i::'a |i. i : (d :: nat set)} & 
  2093 shows "EX f d. card d = card B & linear f & f ` B = {basis i::'a |i. i : (d :: nat set)} &
  2094        f ` span B = {x. ALL i<DIM('a). i ~: d --> x $$ i = (0::real)} &  inj_on f (span B) \<and> d\<subseteq>{..<DIM('a)}" 
  2094        f ` span B = {x. ALL i<DIM('a). i ~: d --> x $$ i = (0::real)} &  inj_on f (span B) \<and> d\<subseteq>{..<DIM('a)}"
  2095 proof-
  2095 proof-
  2096   have B:"card B=dim B" using dim_unique[of B B "card B"] assms span_inc[of B] by auto
  2096   have B:"card B=dim B" using dim_unique[of B B "card B"] assms span_inc[of B] by auto
  2097   def d \<equiv> "{..<dim B}" have t:"card d = dim B" unfolding d_def by auto
  2097   def d \<equiv> "{..<dim B}" have t:"card d = dim B" unfolding d_def by auto
  2098   have "dim B <= DIM('a)" using dim_subset_UNIV[of B] by auto
  2098   have "dim B <= DIM('a)" using dim_subset_UNIV[of B] by auto
  2099   hence d:"d\<subseteq>{..<DIM('a)}" unfolding d_def by auto
  2099   hence d:"d\<subseteq>{..<DIM('a)}" unfolding d_def by auto
  2104     apply(rule subspace_span) apply(rule subspace_substandard) defer
  2104     apply(rule subspace_span) apply(rule subspace_substandard) defer
  2105     apply(rule span_inc) apply(rule assms) defer unfolding dim_span[of B] apply(rule B)
  2105     apply(rule span_inc) apply(rule assms) defer unfolding dim_span[of B] apply(rule B)
  2106     unfolding span_substd_basis[OF d,symmetric] card_substdbasis[OF d] apply(rule span_inc)
  2106     unfolding span_substd_basis[OF d,symmetric] card_substdbasis[OF d] apply(rule span_inc)
  2107     apply(rule independent_substdbasis[OF d]) apply(rule,assumption)
  2107     apply(rule independent_substdbasis[OF d]) apply(rule,assumption)
  2108     unfolding t[symmetric] span_substd_basis[OF d] dim_substandard[OF d] by auto
  2108     unfolding t[symmetric] span_substd_basis[OF d] dim_substandard[OF d] by auto
  2109   from this t `card B=dim B` show ?thesis using d by auto 
  2109   from this t `card B=dim B` show ?thesis using d by auto
  2110 qed
  2110 qed
  2111 
  2111 
  2112 lemma aff_dim_empty:
  2112 lemma aff_dim_empty:
  2113 fixes S :: "('n::euclidean_space) set" 
  2113 fixes S :: "('n::euclidean_space) set"
  2114 shows "S = {} <-> aff_dim S = -1"
  2114 shows "S = {} <-> aff_dim S = -1"
  2115 proof-
  2115 proof-
  2116 obtain B where "affine hull B = affine hull S & ~ affine_dependent B & int (card B) = aff_dim S + 1" using aff_dim_basis_exists by auto
  2116 obtain B where "affine hull B = affine hull S & ~ affine_dependent B & int (card B) = aff_dim S + 1" using aff_dim_basis_exists by auto
  2117 moreover hence "S={} <-> B={}" using affine_hull_nonempty[of B] affine_hull_nonempty[of S] by auto
  2117 moreover hence "S={} <-> B={}" using affine_hull_nonempty[of B] affine_hull_nonempty[of S] by auto
  2118 ultimately show ?thesis using aff_independent_finite[of B] card_gt_0_iff[of B] by auto
  2118 ultimately show ?thesis using aff_independent_finite[of B] card_gt_0_iff[of B] by auto
  2119 qed
  2119 qed
  2120 
  2120 
  2121 lemma aff_dim_affine_hull:
  2121 lemma aff_dim_affine_hull:
  2122 shows "aff_dim (affine hull S)=aff_dim S" 
  2122 shows "aff_dim (affine hull S)=aff_dim S"
  2123 unfolding aff_dim_def using hull_hull[of _ S] by auto 
  2123 unfolding aff_dim_def using hull_hull[of _ S] by auto
  2124 
  2124 
  2125 lemma aff_dim_affine_hull2:
  2125 lemma aff_dim_affine_hull2:
  2126 assumes "affine hull S=affine hull T"
  2126 assumes "affine hull S=affine hull T"
  2127 shows "aff_dim S=aff_dim T" unfolding aff_dim_def using assms by auto
  2127 shows "aff_dim S=aff_dim T" unfolding aff_dim_def using assms by auto
  2128 
  2128 
  2129 lemma aff_dim_unique: 
  2129 lemma aff_dim_unique:
  2130 fixes B V :: "('n::euclidean_space) set" 
  2130 fixes B V :: "('n::euclidean_space) set"
  2131 assumes "(affine hull B=affine hull V) & ~(affine_dependent B)"
  2131 assumes "(affine hull B=affine hull V) & ~(affine_dependent B)"
  2132 shows "of_nat(card B) = aff_dim V+1"
  2132 shows "of_nat(card B) = aff_dim V+1"
  2133 proof-
  2133 proof-
  2134 { assume "B={}" hence "V={}" using affine_hull_nonempty[of V] affine_hull_nonempty[of B] assms by auto
  2134 { assume "B={}" hence "V={}" using affine_hull_nonempty[of V] affine_hull_nonempty[of B] assms by auto
  2135   hence "aff_dim V = (-1::int)"  using aff_dim_empty by auto  
  2135   hence "aff_dim V = (-1::int)"  using aff_dim_empty by auto
  2136   hence ?thesis using `B={}` by auto
  2136   hence ?thesis using `B={}` by auto
  2137 }
  2137 }
  2138 moreover
  2138 moreover
  2139 { assume "B~={}" from this obtain a where a_def: "a:B" by auto 
  2139 { assume "B~={}" from this obtain a where a_def: "a:B" by auto
  2140   def Lb == "span ((%x. -a+x) ` (B-{a}))"
  2140   def Lb == "span ((%x. -a+x) ` (B-{a}))"
  2141   have "affine_parallel (affine hull B) Lb"
  2141   have "affine_parallel (affine hull B) Lb"
  2142      using Lb_def affine_hull_span2[of a B] a_def  affine_parallel_commut[of "Lb" "(affine hull B)"] 
  2142      using Lb_def affine_hull_span2[of a B] a_def  affine_parallel_commut[of "Lb" "(affine hull B)"]
  2143      unfolding affine_parallel_def by auto
  2143      unfolding affine_parallel_def by auto
  2144   moreover have "subspace Lb" using Lb_def subspace_span by auto
  2144   moreover have "subspace Lb" using Lb_def subspace_span by auto
  2145   ultimately have "aff_dim B=int(dim Lb)" using aff_dim_parallel_subspace[of B Lb] `B~={}` by auto 
  2145   ultimately have "aff_dim B=int(dim Lb)" using aff_dim_parallel_subspace[of B Lb] `B~={}` by auto
  2146   moreover have "(card B) - 1 = dim Lb" "finite B" using Lb_def aff_dim_parallel_subspace_aux a_def assms by auto
  2146   moreover have "(card B) - 1 = dim Lb" "finite B" using Lb_def aff_dim_parallel_subspace_aux a_def assms by auto
  2147   ultimately have "(of_nat(card B) = aff_dim B+1)" using  `B~={}` card_gt_0_iff[of B] by auto
  2147   ultimately have "(of_nat(card B) = aff_dim B+1)" using  `B~={}` card_gt_0_iff[of B] by auto
  2148   hence ?thesis using aff_dim_affine_hull2 assms by auto
  2148   hence ?thesis using aff_dim_affine_hull2 assms by auto
  2149 } ultimately show ?thesis by blast
  2149 } ultimately show ?thesis by blast
  2150 qed
  2150 qed
  2151 
  2151 
  2152 lemma aff_dim_affine_independent: 
  2152 lemma aff_dim_affine_independent:
  2153 fixes B :: "('n::euclidean_space) set" 
  2153 fixes B :: "('n::euclidean_space) set"
  2154 assumes "~(affine_dependent B)"
  2154 assumes "~(affine_dependent B)"
  2155 shows "of_nat(card B) = aff_dim B+1"
  2155 shows "of_nat(card B) = aff_dim B+1"
  2156   using aff_dim_unique[of B B] assms by auto
  2156   using aff_dim_unique[of B B] assms by auto
  2157 
  2157 
  2158 lemma aff_dim_sing: 
  2158 lemma aff_dim_sing:
  2159 fixes a :: "'n::euclidean_space" 
  2159 fixes a :: "'n::euclidean_space"
  2160 shows "aff_dim {a}=0"
  2160 shows "aff_dim {a}=0"
  2161   using aff_dim_affine_independent[of "{a}"] affine_independent_sing by auto
  2161   using aff_dim_affine_independent[of "{a}"] affine_independent_sing by auto
  2162 
  2162 
  2163 lemma aff_dim_inner_basis_exists:
  2163 lemma aff_dim_inner_basis_exists:
  2164   fixes V :: "('n::euclidean_space) set" 
  2164   fixes V :: "('n::euclidean_space) set"
  2165   shows "? B. B<=V & (affine hull B=affine hull V) & ~(affine_dependent B) & (of_nat(card B) = aff_dim V+1)"
  2165   shows "? B. B<=V & (affine hull B=affine hull V) & ~(affine_dependent B) & (of_nat(card B) = aff_dim V+1)"
  2166 proof-
  2166 proof-
  2167 obtain B where B_def: "~(affine_dependent B) & B<=V & (affine hull B=affine hull V)" using affine_basis_exists[of V] by auto
  2167 obtain B where B_def: "~(affine_dependent B) & B<=V & (affine hull B=affine hull V)" using affine_basis_exists[of V] by auto
  2168 moreover hence "of_nat(card B) = aff_dim V+1" using aff_dim_unique by auto
  2168 moreover hence "of_nat(card B) = aff_dim V+1" using aff_dim_unique by auto
  2169 ultimately show ?thesis by auto
  2169 ultimately show ?thesis by auto
  2170 qed
  2170 qed
  2171 
  2171 
  2172 lemma aff_dim_le_card:
  2172 lemma aff_dim_le_card:
  2173 fixes V :: "('n::euclidean_space) set" 
  2173 fixes V :: "('n::euclidean_space) set"
  2174 assumes "finite V"
  2174 assumes "finite V"
  2175 shows "aff_dim V <= of_nat(card V) - 1"
  2175 shows "aff_dim V <= of_nat(card V) - 1"
  2176  proof-
  2176  proof-
  2177  obtain B where B_def: "B<=V & (of_nat(card B) = aff_dim V+1)" using aff_dim_inner_basis_exists[of V] by auto 
  2177  obtain B where B_def: "B<=V & (of_nat(card B) = aff_dim V+1)" using aff_dim_inner_basis_exists[of V] by auto
  2178  moreover hence "card B <= card V" using assms card_mono by auto
  2178  moreover hence "card B <= card V" using assms card_mono by auto
  2179  ultimately show ?thesis by auto
  2179  ultimately show ?thesis by auto
  2180 qed
  2180 qed
  2181 
  2181 
  2182 lemma aff_dim_parallel_eq:
  2182 lemma aff_dim_parallel_eq:
  2183 fixes S T :: "('n::euclidean_space) set"
  2183 fixes S T :: "('n::euclidean_space) set"
  2184 assumes "affine_parallel (affine hull S) (affine hull T)"
  2184 assumes "affine_parallel (affine hull S) (affine hull T)"
  2185 shows "aff_dim S=aff_dim T"
  2185 shows "aff_dim S=aff_dim T"
  2186 proof-
  2186 proof-
  2187 { assume "T~={}" "S~={}" 
  2187 { assume "T~={}" "S~={}"
  2188   from this obtain L where L_def: "subspace L & affine_parallel (affine hull T) L" 
  2188   from this obtain L where L_def: "subspace L & affine_parallel (affine hull T) L"
  2189        using affine_parallel_subspace[of "affine hull T"] affine_affine_hull[of T] affine_hull_nonempty by auto
  2189        using affine_parallel_subspace[of "affine hull T"] affine_affine_hull[of T] affine_hull_nonempty by auto
  2190   hence "aff_dim T = int(dim L)" using aff_dim_parallel_subspace `T~={}` by auto
  2190   hence "aff_dim T = int(dim L)" using aff_dim_parallel_subspace `T~={}` by auto
  2191   moreover have "subspace L & affine_parallel (affine hull S) L" 
  2191   moreover have "subspace L & affine_parallel (affine hull S) L"
  2192      using L_def affine_parallel_assoc[of "affine hull S" "affine hull T" L] assms by auto
  2192      using L_def affine_parallel_assoc[of "affine hull S" "affine hull T" L] assms by auto
  2193   moreover hence "aff_dim S = int(dim L)" using aff_dim_parallel_subspace `S~={}` by auto 
  2193   moreover hence "aff_dim S = int(dim L)" using aff_dim_parallel_subspace `S~={}` by auto
  2194   ultimately have ?thesis by auto
  2194   ultimately have ?thesis by auto
  2195 }
  2195 }
  2196 moreover
  2196 moreover
  2197 { assume "S={}" hence "S={} & T={}" using assms affine_hull_nonempty unfolding affine_parallel_def by auto
  2197 { assume "S={}" hence "S={} & T={}" using assms affine_hull_nonempty unfolding affine_parallel_def by auto
  2198   hence ?thesis using aff_dim_empty by auto
  2198   hence ?thesis using aff_dim_empty by auto
  2204 ultimately show ?thesis by blast
  2204 ultimately show ?thesis by blast
  2205 qed
  2205 qed
  2206 
  2206 
  2207 lemma aff_dim_translation_eq:
  2207 lemma aff_dim_translation_eq:
  2208 fixes a :: "'n::euclidean_space"
  2208 fixes a :: "'n::euclidean_space"
  2209 shows "aff_dim ((%x. a + x) ` S)=aff_dim S" 
  2209 shows "aff_dim ((%x. a + x) ` S)=aff_dim S"
  2210 proof-
  2210 proof-
  2211 have "affine_parallel (affine hull S) (affine hull ((%x. a + x) ` S))" unfolding affine_parallel_def apply (rule exI[of _ "a"]) using affine_hull_translation[of a S] by auto
  2211 have "affine_parallel (affine hull S) (affine hull ((%x. a + x) ` S))" unfolding affine_parallel_def apply (rule exI[of _ "a"]) using affine_hull_translation[of a S] by auto
  2212 from this show ?thesis using  aff_dim_parallel_eq[of S "(%x. a + x) ` S"] by auto 
  2212 from this show ?thesis using  aff_dim_parallel_eq[of S "(%x. a + x) ` S"] by auto
  2213 qed
  2213 qed
  2214 
  2214 
  2215 lemma aff_dim_affine:
  2215 lemma aff_dim_affine:
  2216 fixes S L :: "('n::euclidean_space) set"
  2216 fixes S L :: "('n::euclidean_space) set"
  2217 assumes "S ~= {}" "affine S"
  2217 assumes "S ~= {}" "affine S"
  2218 assumes "subspace L" "affine_parallel S L"
  2218 assumes "subspace L" "affine_parallel S L"
  2219 shows "aff_dim S=int(dim L)" 
  2219 shows "aff_dim S=int(dim L)"
  2220 proof-
  2220 proof-
  2221 have 1: "(affine hull S) = S" using assms affine_hull_eq[of S] by auto 
  2221 have 1: "(affine hull S) = S" using assms affine_hull_eq[of S] by auto
  2222 hence "affine_parallel (affine hull S) L" using assms by (simp add:1)
  2222 hence "affine_parallel (affine hull S) L" using assms by (simp add:1)
  2223 from this show ?thesis using assms aff_dim_parallel_subspace[of S L] by blast 
  2223 from this show ?thesis using assms aff_dim_parallel_subspace[of S L] by blast
  2224 qed
  2224 qed
  2225 
  2225 
  2226 lemma dim_affine_hull:
  2226 lemma dim_affine_hull:
  2227 fixes S :: "('n::euclidean_space) set"
  2227 fixes S :: "('n::euclidean_space) set"
  2228 shows "dim (affine hull S)=dim S"
  2228 shows "dim (affine hull S)=dim S"
  2234 qed
  2234 qed
  2235 
  2235 
  2236 lemma aff_dim_subspace:
  2236 lemma aff_dim_subspace:
  2237 fixes S :: "('n::euclidean_space) set"
  2237 fixes S :: "('n::euclidean_space) set"
  2238 assumes "S ~= {}" "subspace S"
  2238 assumes "S ~= {}" "subspace S"
  2239 shows "aff_dim S=int(dim S)" using aff_dim_affine[of S S] assms subspace_imp_affine[of S] affine_parallel_reflex[of S] by auto 
  2239 shows "aff_dim S=int(dim S)" using aff_dim_affine[of S S] assms subspace_imp_affine[of S] affine_parallel_reflex[of S] by auto
  2240 
  2240 
  2241 lemma aff_dim_zero:
  2241 lemma aff_dim_zero:
  2242 fixes S :: "('n::euclidean_space) set"
  2242 fixes S :: "('n::euclidean_space) set"
  2243 assumes "0 : affine hull S"
  2243 assumes "0 : affine hull S"
  2244 shows "aff_dim S=int(dim S)"
  2244 shows "aff_dim S=int(dim S)"
  2245 proof-
  2245 proof-
  2246 have "subspace(affine hull S)" using subspace_affine[of "affine hull S"] affine_affine_hull assms by auto
  2246 have "subspace(affine hull S)" using subspace_affine[of "affine hull S"] affine_affine_hull assms by auto
  2247 hence "aff_dim (affine hull S) =int(dim (affine hull S))" using assms aff_dim_subspace[of "affine hull S"] by auto  
  2247 hence "aff_dim (affine hull S) =int(dim (affine hull S))" using assms aff_dim_subspace[of "affine hull S"] by auto
  2248 from this show ?thesis using aff_dim_affine_hull[of S] dim_affine_hull[of S] by auto
  2248 from this show ?thesis using aff_dim_affine_hull[of S] dim_affine_hull[of S] by auto
  2249 qed
  2249 qed
  2250 
  2250 
  2251 lemma aff_dim_univ: "aff_dim (UNIV :: ('n::euclidean_space) set) = int(DIM('n))"
  2251 lemma aff_dim_univ: "aff_dim (UNIV :: ('n::euclidean_space) set) = int(DIM('n))"
  2252   using aff_dim_subspace[of "(UNIV :: ('n::euclidean_space) set)"]
  2252   using aff_dim_subspace[of "(UNIV :: ('n::euclidean_space) set)"]
  2258 proof-
  2258 proof-
  2259 obtain B where B_def: "affine hull B = affine hull V & ~ affine_dependent B & int (card B) = aff_dim V + 1" using aff_dim_basis_exists by auto
  2259 obtain B where B_def: "affine hull B = affine hull V & ~ affine_dependent B & int (card B) = aff_dim V + 1" using aff_dim_basis_exists by auto
  2260 from this show ?thesis by auto
  2260 from this show ?thesis by auto
  2261 qed
  2261 qed
  2262 
  2262 
  2263 lemma independent_card_le_aff_dim: 
  2263 lemma independent_card_le_aff_dim:
  2264   assumes "(B::('n::euclidean_space) set) <= V"
  2264   assumes "(B::('n::euclidean_space) set) <= V"
  2265   assumes "~(affine_dependent B)" 
  2265   assumes "~(affine_dependent B)"
  2266   shows "int(card B) <= aff_dim V+1"
  2266   shows "int(card B) <= aff_dim V+1"
  2267 proof-
  2267 proof-
  2268 { assume "B~={}" 
  2268 { assume "B~={}"
  2269   from this obtain T where T_def: "~(affine_dependent T) & B<=T & T<=V & affine hull T = affine hull V" 
  2269   from this obtain T where T_def: "~(affine_dependent T) & B<=T & T<=V & affine hull T = affine hull V"
  2270   using assms extend_to_affine_basis[of B V] by auto
  2270   using assms extend_to_affine_basis[of B V] by auto
  2271   hence "of_nat(card T) = aff_dim V+1" using aff_dim_unique by auto
  2271   hence "of_nat(card T) = aff_dim V+1" using aff_dim_unique by auto
  2272   hence ?thesis using T_def card_mono[of T B] aff_independent_finite[of T] by auto
  2272   hence ?thesis using T_def card_mono[of T B] aff_independent_finite[of T] by auto
  2273 }
  2273 }
  2274 moreover
  2274 moreover
  2289 qed
  2289 qed
  2290 
  2290 
  2291 lemma aff_dim_subset_univ:
  2291 lemma aff_dim_subset_univ:
  2292 fixes S :: "('n::euclidean_space) set"
  2292 fixes S :: "('n::euclidean_space) set"
  2293 shows "aff_dim S <= int(DIM('n))"
  2293 shows "aff_dim S <= int(DIM('n))"
  2294 proof - 
  2294 proof -
  2295   have "aff_dim (UNIV :: ('n::euclidean_space) set) = int(DIM('n))" using aff_dim_univ by auto
  2295   have "aff_dim (UNIV :: ('n::euclidean_space) set) = int(DIM('n))" using aff_dim_univ by auto
  2296   from this show "aff_dim (S:: ('n::euclidean_space) set) <= int(DIM('n))" using assms aff_dim_subset[of S "(UNIV :: ('n::euclidean_space) set)"] subset_UNIV by auto
  2296   from this show "aff_dim (S:: ('n::euclidean_space) set) <= int(DIM('n))" using assms aff_dim_subset[of S "(UNIV :: ('n::euclidean_space) set)"] subset_UNIV by auto
  2297 qed
  2297 qed
  2298 
  2298 
  2299 lemma affine_dim_equal:
  2299 lemma affine_dim_equal:
  2300 assumes "affine (S :: ('n::euclidean_space) set)" "affine T" "S ~= {}" "S <= T" "aff_dim S = aff_dim T"
  2300 assumes "affine (S :: ('n::euclidean_space) set)" "affine T" "S ~= {}" "S <= T" "aff_dim S = aff_dim T"
  2301 shows "S=T"
  2301 shows "S=T"
  2302 proof-
  2302 proof-
  2303 obtain a where "a : S" using assms by auto 
  2303 obtain a where "a : S" using assms by auto
  2304 hence "a : T" using assms by auto
  2304 hence "a : T" using assms by auto
  2305 def LS == "{y. ? x : S. (-a)+x=y}"
  2305 def LS == "{y. ? x : S. (-a)+x=y}"
  2306 hence ls: "subspace LS & affine_parallel S LS" using assms parallel_subspace_explicit[of S a LS] `a : S` by auto 
  2306 hence ls: "subspace LS & affine_parallel S LS" using assms parallel_subspace_explicit[of S a LS] `a : S` by auto
  2307 hence h1: "int(dim LS) = aff_dim S" using assms aff_dim_affine[of S LS] by auto
  2307 hence h1: "int(dim LS) = aff_dim S" using assms aff_dim_affine[of S LS] by auto
  2308 have "T ~= {}" using assms by auto
  2308 have "T ~= {}" using assms by auto
  2309 def LT == "{y. ? x : T. (-a)+x=y}" 
  2309 def LT == "{y. ? x : T. (-a)+x=y}"
  2310 hence lt: "subspace LT & affine_parallel T LT" using assms parallel_subspace_explicit[of T a LT] `a : T` by auto
  2310 hence lt: "subspace LT & affine_parallel T LT" using assms parallel_subspace_explicit[of T a LT] `a : T` by auto
  2311 hence "int(dim LT) = aff_dim T" using assms aff_dim_affine[of T LT] `T ~= {}` by auto 
  2311 hence "int(dim LT) = aff_dim T" using assms aff_dim_affine[of T LT] `T ~= {}` by auto
  2312 hence "dim LS = dim LT" using h1 assms by auto
  2312 hence "dim LS = dim LT" using h1 assms by auto
  2313 moreover have "LS <= LT" using LS_def LT_def assms by auto
  2313 moreover have "LS <= LT" using LS_def LT_def assms by auto
  2314 ultimately have "LS=LT" using subspace_dim_equal[of LS LT] ls lt by auto
  2314 ultimately have "LS=LT" using subspace_dim_equal[of LS LT] ls lt by auto
  2315 moreover have "S = {x. ? y : LS. a+y=x}" using LS_def by auto 
  2315 moreover have "S = {x. ? y : LS. a+y=x}" using LS_def by auto
  2316 moreover have "T = {x. ? y : LT. a+y=x}" using LT_def by auto
  2316 moreover have "T = {x. ? y : LT. a+y=x}" using LT_def by auto
  2317 ultimately show ?thesis by auto 
  2317 ultimately show ?thesis by auto
  2318 qed
  2318 qed
  2319 
  2319 
  2320 lemma affine_hull_univ:
  2320 lemma affine_hull_univ:
  2321 fixes S :: "('n::euclidean_space) set"
  2321 fixes S :: "('n::euclidean_space) set"
  2322 assumes "aff_dim S = int(DIM('n))"
  2322 assumes "aff_dim S = int(DIM('n))"
  2323 shows "affine hull S = (UNIV :: ('n::euclidean_space) set)"
  2323 shows "affine hull S = (UNIV :: ('n::euclidean_space) set)"
  2324 proof-
  2324 proof-
  2325 have "S ~= {}" using assms aff_dim_empty[of S] by auto
  2325 have "S ~= {}" using assms aff_dim_empty[of S] by auto
  2326 have h0: "S <= affine hull S" using hull_subset[of S _] by auto
  2326 have h0: "S <= affine hull S" using hull_subset[of S _] by auto
  2327 have h1: "aff_dim (UNIV :: ('n::euclidean_space) set) = aff_dim S" using aff_dim_univ assms by auto
  2327 have h1: "aff_dim (UNIV :: ('n::euclidean_space) set) = aff_dim S" using aff_dim_univ assms by auto
  2328 hence h2: "aff_dim (affine hull S) <= aff_dim (UNIV :: ('n::euclidean_space) set)" using aff_dim_subset_univ[of "affine hull S"] assms h0 by auto  
  2328 hence h2: "aff_dim (affine hull S) <= aff_dim (UNIV :: ('n::euclidean_space) set)" using aff_dim_subset_univ[of "affine hull S"] assms h0 by auto
  2329 have h3: "aff_dim S <= aff_dim (affine hull S)" using h0 aff_dim_subset[of S "affine hull S"] assms by auto
  2329 have h3: "aff_dim S <= aff_dim (affine hull S)" using h0 aff_dim_subset[of S "affine hull S"] assms by auto
  2330 hence h4: "aff_dim (affine hull S) = aff_dim (UNIV :: ('n::euclidean_space) set)" using h0 h1 h2 by auto
  2330 hence h4: "aff_dim (affine hull S) = aff_dim (UNIV :: ('n::euclidean_space) set)" using h0 h1 h2 by auto
  2331 from this show ?thesis using affine_dim_equal[of "affine hull S" "(UNIV :: ('n::euclidean_space) set)"] affine_affine_hull[of S] affine_UNIV assms h4 h0 `S ~= {}` by auto
  2331 from this show ?thesis using affine_dim_equal[of "affine hull S" "(UNIV :: ('n::euclidean_space) set)"] affine_affine_hull[of S] affine_UNIV assms h4 h0 `S ~= {}` by auto
  2332 qed
  2332 qed
  2333 
  2333 
  2334 lemma aff_dim_convex_hull:
  2334 lemma aff_dim_convex_hull:
  2335 fixes S :: "('n::euclidean_space) set"
  2335 fixes S :: "('n::euclidean_space) set"
  2336 shows "aff_dim (convex hull S)=aff_dim S"
  2336 shows "aff_dim (convex hull S)=aff_dim S"
  2337   using aff_dim_affine_hull[of S] convex_hull_subset_affine_hull[of S] 
  2337   using aff_dim_affine_hull[of S] convex_hull_subset_affine_hull[of S]
  2338   hull_subset[of S "convex"] aff_dim_subset[of S "convex hull S"] 
  2338   hull_subset[of S "convex"] aff_dim_subset[of S "convex hull S"]
  2339   aff_dim_subset[of "convex hull S" "affine hull S"] by auto
  2339   aff_dim_subset[of "convex hull S" "affine hull S"] by auto
  2340 
  2340 
  2341 lemma aff_dim_cball:
  2341 lemma aff_dim_cball:
  2342 fixes a :: "'n::euclidean_space" 
  2342 fixes a :: "'n::euclidean_space"
  2343 assumes "0<e"
  2343 assumes "0<e"
  2344 shows "aff_dim (cball a e) = int (DIM('n))"
  2344 shows "aff_dim (cball a e) = int (DIM('n))"
  2345 proof-
  2345 proof-
  2346 have "(%x. a + x) ` (cball 0 e)<=cball a e" unfolding cball_def dist_norm by auto
  2346 have "(%x. a + x) ` (cball 0 e)<=cball a e" unfolding cball_def dist_norm by auto
  2347 hence "aff_dim (cball (0 :: 'n::euclidean_space) e) <= aff_dim (cball a e)"
  2347 hence "aff_dim (cball (0 :: 'n::euclidean_space) e) <= aff_dim (cball a e)"
  2348   using aff_dim_translation_eq[of a "cball 0 e"] 
  2348   using aff_dim_translation_eq[of a "cball 0 e"]
  2349         aff_dim_subset[of "op + a ` cball 0 e" "cball a e"] by auto
  2349         aff_dim_subset[of "op + a ` cball 0 e" "cball a e"] by auto
  2350 moreover have "aff_dim (cball (0 :: 'n::euclidean_space) e) = int (DIM('n))" 
  2350 moreover have "aff_dim (cball (0 :: 'n::euclidean_space) e) = int (DIM('n))"
  2351    using hull_inc[of "(0 :: 'n::euclidean_space)" "cball 0 e"] centre_in_cball[of "(0 :: 'n::euclidean_space)"] assms 
  2351    using hull_inc[of "(0 :: 'n::euclidean_space)" "cball 0 e"] centre_in_cball[of "(0 :: 'n::euclidean_space)"] assms
  2352    by (simp add: dim_cball[of e] aff_dim_zero[of "cball 0 e"])
  2352    by (simp add: dim_cball[of e] aff_dim_zero[of "cball 0 e"])
  2353 ultimately show ?thesis using aff_dim_subset_univ[of "cball a e"] by auto 
  2353 ultimately show ?thesis using aff_dim_subset_univ[of "cball a e"] by auto
  2354 qed
  2354 qed
  2355 
  2355 
  2356 lemma aff_dim_open:
  2356 lemma aff_dim_open:
  2357 fixes S :: "('n::euclidean_space) set"
  2357 fixes S :: "('n::euclidean_space) set"
  2358 assumes "open S" "S ~= {}"
  2358 assumes "open S" "S ~= {}"
  2359 shows "aff_dim S = int (DIM('n))"
  2359 shows "aff_dim S = int (DIM('n))"
  2360 proof-
  2360 proof-
  2361 obtain x where "x:S" using assms by auto
  2361 obtain x where "x:S" using assms by auto
  2362 from this obtain e where e_def: "e>0 & cball x e <= S" using open_contains_cball[of S] assms by auto
  2362 from this obtain e where e_def: "e>0 & cball x e <= S" using open_contains_cball[of S] assms by auto
  2363 from this have "aff_dim (cball x e) <= aff_dim S" using aff_dim_subset by auto
  2363 from this have "aff_dim (cball x e) <= aff_dim S" using aff_dim_subset by auto
  2364 from this show ?thesis using aff_dim_cball[of e x] aff_dim_subset_univ[of S] e_def by auto     
  2364 from this show ?thesis using aff_dim_cball[of e x] aff_dim_subset_univ[of S] e_def by auto
  2365 qed
  2365 qed
  2366 
  2366 
  2367 lemma low_dim_interior:
  2367 lemma low_dim_interior:
  2368 fixes S :: "('n::euclidean_space) set"
  2368 fixes S :: "('n::euclidean_space) set"
  2369 assumes "~(aff_dim S = int (DIM('n)))"
  2369 assumes "~(aff_dim S = int (DIM('n)))"
  2370 shows "interior S = {}"
  2370 shows "interior S = {}"
  2371 proof-
  2371 proof-
  2372 have "aff_dim(interior S) <= aff_dim S" 
  2372 have "aff_dim(interior S) <= aff_dim S"
  2373    using interior_subset aff_dim_subset[of "interior S" S] by auto 
  2373    using interior_subset aff_dim_subset[of "interior S" S] by auto
  2374 from this show ?thesis using aff_dim_open[of "interior S"] aff_dim_subset_univ[of S] assms by auto   
  2374 from this show ?thesis using aff_dim_open[of "interior S"] aff_dim_subset_univ[of S] assms by auto
  2375 qed
  2375 qed
  2376 
  2376 
  2377 subsection {* Relative interior of a set *}
  2377 subsection {* Relative interior of a set *}
  2378 
  2378 
  2379 definition "rel_interior S = {x. ? T. openin (subtopology euclidean (affine hull S)) T & x : T & T <= S}"
  2379 definition "rel_interior S = {x. ? T. openin (subtopology euclidean (affine hull S)) T & x : T & T <= S}"
  2386 show "EX Tb. (EX Ta. open Ta & Tb = affine hull S Int Ta) & x : Tb & Tb <= S"
  2386 show "EX Tb. (EX Ta. open Ta & Tb = affine hull S Int Ta) & x : Tb & Tb <= S"
  2387 apply (rule_tac x="T Int (affine hull S)" in exI)
  2387 apply (rule_tac x="T Int (affine hull S)" in exI)
  2388 using a h1 by auto
  2388 using a h1 by auto
  2389 qed
  2389 qed
  2390 
  2390 
  2391 lemma mem_rel_interior: 
  2391 lemma mem_rel_interior:
  2392      "x : rel_interior S <-> (? T. open T & x : (T Int S) & (T Int (affine hull S)) <= S)" 
  2392      "x : rel_interior S <-> (? T. open T & x : (T Int S) & (T Int (affine hull S)) <= S)"
  2393      by (auto simp add: rel_interior)
  2393      by (auto simp add: rel_interior)
  2394 
  2394 
  2395 lemma mem_rel_interior_ball: "x : rel_interior S <-> x : S & (? e. 0 < e & ((ball x e) Int (affine hull S)) <= S)"
  2395 lemma mem_rel_interior_ball: "x : rel_interior S <-> x : S & (? e. 0 < e & ((ball x e) Int (affine hull S)) <= S)"
  2396   apply (simp add: rel_interior, safe)
  2396   apply (simp add: rel_interior, safe)
  2397   apply (force simp add: open_contains_ball)
  2397   apply (force simp add: open_contains_ball)
  2398   apply (rule_tac x="ball x e" in exI)
  2398   apply (rule_tac x="ball x e" in exI)
  2399   apply simp
  2399   apply simp
  2400   done
  2400   done
  2401 
  2401 
  2402 lemma rel_interior_ball: 
  2402 lemma rel_interior_ball:
  2403       "rel_interior S = {x : S. ? e. e>0 & ((ball x e) Int (affine hull S)) <= S}" 
  2403       "rel_interior S = {x : S. ? e. e>0 & ((ball x e) Int (affine hull S)) <= S}"
  2404       using mem_rel_interior_ball [of _ S] by auto 
  2404       using mem_rel_interior_ball [of _ S] by auto
  2405 
  2405 
  2406 lemma mem_rel_interior_cball: "x : rel_interior S <-> x : S & (? e. 0 < e & ((cball x e) Int (affine hull S)) <= S)"
  2406 lemma mem_rel_interior_cball: "x : rel_interior S <-> x : S & (? e. 0 < e & ((cball x e) Int (affine hull S)) <= S)"
  2407   apply (simp add: rel_interior, safe) 
  2407   apply (simp add: rel_interior, safe)
  2408   apply (force simp add: open_contains_cball)
  2408   apply (force simp add: open_contains_cball)
  2409   apply (rule_tac x="ball x e" in exI)
  2409   apply (rule_tac x="ball x e" in exI)
  2410   apply (simp add: subset_trans [OF ball_subset_cball])
  2410   apply (simp add: subset_trans [OF ball_subset_cball])
  2411   apply auto
  2411   apply auto
  2412   done
  2412   done
  2413 
  2413 
  2414 lemma rel_interior_cball: "rel_interior S = {x : S. ? e. e>0 & ((cball x e) Int (affine hull S)) <= S}"       using mem_rel_interior_cball [of _ S] by auto
  2414 lemma rel_interior_cball: "rel_interior S = {x : S. ? e. e>0 & ((cball x e) Int (affine hull S)) <= S}"       using mem_rel_interior_cball [of _ S] by auto
  2415 
  2415 
  2416 lemma rel_interior_empty: "rel_interior {} = {}" 
  2416 lemma rel_interior_empty: "rel_interior {} = {}"
  2417    by (auto simp add: rel_interior_def) 
  2417    by (auto simp add: rel_interior_def)
  2418 
  2418 
  2419 lemma affine_hull_sing: "affine hull {a :: 'n::euclidean_space} = {a}"
  2419 lemma affine_hull_sing: "affine hull {a :: 'n::euclidean_space} = {a}"
  2420 by (metis affine_hull_eq affine_sing)
  2420 by (metis affine_hull_eq affine_sing)
  2421 
  2421 
  2422 lemma rel_interior_sing: "rel_interior {a :: 'n::euclidean_space} = {a}"
  2422 lemma rel_interior_sing: "rel_interior {a :: 'n::euclidean_space} = {a}"
  2426 
  2426 
  2427 lemma subset_rel_interior:
  2427 lemma subset_rel_interior:
  2428 fixes S T :: "('n::euclidean_space) set"
  2428 fixes S T :: "('n::euclidean_space) set"
  2429 assumes "S<=T" "affine hull S=affine hull T"
  2429 assumes "S<=T" "affine hull S=affine hull T"
  2430 shows "rel_interior S <= rel_interior T"
  2430 shows "rel_interior S <= rel_interior T"
  2431   using assms by (auto simp add: rel_interior_def)  
  2431   using assms by (auto simp add: rel_interior_def)
  2432 
  2432 
  2433 lemma rel_interior_subset: "rel_interior S <= S" 
  2433 lemma rel_interior_subset: "rel_interior S <= S"
  2434    by (auto simp add: rel_interior_def)
  2434    by (auto simp add: rel_interior_def)
  2435 
  2435 
  2436 lemma rel_interior_subset_closure: "rel_interior S <= closure S" 
  2436 lemma rel_interior_subset_closure: "rel_interior S <= closure S"
  2437    using rel_interior_subset by (auto simp add: closure_def) 
  2437    using rel_interior_subset by (auto simp add: closure_def)
  2438 
  2438 
  2439 lemma interior_subset_rel_interior: "interior S <= rel_interior S" 
  2439 lemma interior_subset_rel_interior: "interior S <= rel_interior S"
  2440    by (auto simp add: rel_interior interior_def)
  2440    by (auto simp add: rel_interior interior_def)
  2441 
  2441 
  2442 lemma interior_rel_interior:
  2442 lemma interior_rel_interior:
  2443 fixes S :: "('n::euclidean_space) set"
  2443 fixes S :: "('n::euclidean_space) set"
  2444 assumes "aff_dim S = int(DIM('n))"
  2444 assumes "aff_dim S = int(DIM('n))"
  2445 shows "rel_interior S = interior S"
  2445 shows "rel_interior S = interior S"
  2446 proof -
  2446 proof -
  2447 have "affine hull S = UNIV" using assms affine_hull_univ[of S] by auto 
  2447 have "affine hull S = UNIV" using assms affine_hull_univ[of S] by auto
  2448 from this show ?thesis unfolding rel_interior interior_def by auto
  2448 from this show ?thesis unfolding rel_interior interior_def by auto
  2449 qed
  2449 qed
  2450 
  2450 
  2451 lemma rel_interior_open:
  2451 lemma rel_interior_open:
  2452 fixes S :: "('n::euclidean_space) set"
  2452 fixes S :: "('n::euclidean_space) set"
  2457 lemma interior_rel_interior_gen:
  2457 lemma interior_rel_interior_gen:
  2458 fixes S :: "('n::euclidean_space) set"
  2458 fixes S :: "('n::euclidean_space) set"
  2459 shows "interior S = (if aff_dim S = int(DIM('n)) then rel_interior S else {})"
  2459 shows "interior S = (if aff_dim S = int(DIM('n)) then rel_interior S else {})"
  2460 by (metis interior_rel_interior low_dim_interior)
  2460 by (metis interior_rel_interior low_dim_interior)
  2461 
  2461 
  2462 lemma rel_interior_univ: 
  2462 lemma rel_interior_univ:
  2463 fixes S :: "('n::euclidean_space) set"
  2463 fixes S :: "('n::euclidean_space) set"
  2464 shows "rel_interior (affine hull S) = affine hull S"
  2464 shows "rel_interior (affine hull S) = affine hull S"
  2465 proof-
  2465 proof-
  2466 have h1: "rel_interior (affine hull S) <= affine hull S" using rel_interior_subset by auto 
  2466 have h1: "rel_interior (affine hull S) <= affine hull S" using rel_interior_subset by auto
  2467 { fix x assume x_def: "x : affine hull S"
  2467 { fix x assume x_def: "x : affine hull S"
  2468   obtain e :: real where "e=1" by auto
  2468   obtain e :: real where "e=1" by auto
  2469   hence "e>0 & ball x e Int affine hull (affine hull S) <= affine hull S" using hull_hull[of _ S] by auto
  2469   hence "e>0 & ball x e Int affine hull (affine hull S) <= affine hull S" using hull_hull[of _ S] by auto
  2470   hence "x : rel_interior (affine hull S)" using x_def rel_interior_ball[of "affine hull S"] by auto
  2470   hence "x : rel_interior (affine hull S)" using x_def rel_interior_ball[of "affine hull S"] by auto
  2471 } from this show ?thesis using h1 by auto 
  2471 } from this show ?thesis using h1 by auto
  2472 qed
  2472 qed
  2473 
  2473 
  2474 lemma rel_interior_univ2: "rel_interior (UNIV :: ('n::euclidean_space) set) = UNIV"
  2474 lemma rel_interior_univ2: "rel_interior (UNIV :: ('n::euclidean_space) set) = UNIV"
  2475 by (metis open_UNIV rel_interior_open)
  2475 by (metis open_UNIV rel_interior_open)
  2476 
  2476 
  2477 lemma rel_interior_convex_shrink:
  2477 lemma rel_interior_convex_shrink:
  2478   fixes S :: "('a::euclidean_space) set"
  2478   fixes S :: "('a::euclidean_space) set"
  2479   assumes "convex S" "c : rel_interior S" "x : S" "0 < e" "e <= 1"
  2479   assumes "convex S" "c : rel_interior S" "x : S" "0 < e" "e <= 1"
  2480   shows "x - e *\<^sub>R (x - c) : rel_interior S"
  2480   shows "x - e *\<^sub>R (x - c) : rel_interior S"
  2481 proof- 
  2481 proof-
  2482 (* Proof is a modified copy of the proof of similar lemma mem_interior_convex_shrink 
  2482 (* Proof is a modified copy of the proof of similar lemma mem_interior_convex_shrink
  2483 *)
  2483 *)
  2484 obtain d where "d>0" and d:"ball c d Int affine hull S <= S" 
  2484 obtain d where "d>0" and d:"ball c d Int affine hull S <= S"
  2485   using assms(2) unfolding  mem_rel_interior_ball by auto
  2485   using assms(2) unfolding  mem_rel_interior_ball by auto
  2486 {   fix y assume as:"dist (x - e *\<^sub>R (x - c)) y < e * d & y : affine hull S"
  2486 {   fix y assume as:"dist (x - e *\<^sub>R (x - c)) y < e * d & y : affine hull S"
  2487     have *:"y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" using `e>0` by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib)
  2487     have *:"y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" using `e>0` by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib)
  2488     have "x : affine hull S" using assms hull_subset[of S] by auto
  2488     have "x : affine hull S" using assms hull_subset[of S] by auto
  2489     moreover have "1 / e + - ((1 - e) / e) = 1" 
  2489     moreover have "1 / e + - ((1 - e) / e) = 1"
  2490        using `e>0` left_diff_distrib[of "1" "(1-e)" "1/e"] by auto
  2490        using `e>0` left_diff_distrib[of "1" "(1-e)" "1/e"] by auto
  2491     ultimately have **: "(1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x : affine hull S"
  2491     ultimately have **: "(1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x : affine hull S"
  2492         using as affine_affine_hull[of S] mem_affine[of "affine hull S" y x "(1 / e)" "-((1 - e) / e)"] by (simp add: algebra_simps)     
  2492         using as affine_affine_hull[of S] mem_affine[of "affine hull S" y x "(1 / e)" "-((1 - e) / e)"] by (simp add: algebra_simps)
  2493     have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = abs(1/e) * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)"
  2493     have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = abs(1/e) * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)"
  2494       unfolding dist_norm unfolding norm_scaleR[symmetric] apply(rule arg_cong[where f=norm]) using `e>0`
  2494       unfolding dist_norm unfolding norm_scaleR[symmetric] apply(rule arg_cong[where f=norm]) using `e>0`
  2495       by(auto simp add:euclidean_eq[where 'a='a] field_simps) 
  2495       by(auto simp add:euclidean_eq[where 'a='a] field_simps)
  2496     also have "... = abs(1/e) * norm (x - e *\<^sub>R (x - c) - y)" by(auto intro!:arg_cong[where f=norm] simp add: algebra_simps)
  2496     also have "... = abs(1/e) * norm (x - e *\<^sub>R (x - c) - y)" by(auto intro!:arg_cong[where f=norm] simp add: algebra_simps)
  2497     also have "... < d" using as[unfolded dist_norm] and `e>0`
  2497     also have "... < d" using as[unfolded dist_norm] and `e>0`
  2498       by(auto simp add:pos_divide_less_eq[OF `e>0`] mult_commute)
  2498       by(auto simp add:pos_divide_less_eq[OF `e>0`] mult_commute)
  2499     finally have "y : S" apply(subst *) 
  2499     finally have "y : S" apply(subst *)
  2500 apply(rule assms(1)[unfolded convex_alt,rule_format])
  2500 apply(rule assms(1)[unfolded convex_alt,rule_format])
  2501       apply(rule d[unfolded subset_eq,rule_format]) unfolding mem_ball using assms(3-5) ** by auto
  2501       apply(rule d[unfolded subset_eq,rule_format]) unfolding mem_ball using assms(3-5) ** by auto
  2502 } hence "ball (x - e *\<^sub>R (x - c)) (e*d) Int affine hull S <= S" by auto
  2502 } hence "ball (x - e *\<^sub>R (x - c)) (e*d) Int affine hull S <= S" by auto
  2503 moreover have "0 < e*d" using `0<e` `0<d` by (rule mult_pos_pos)
  2503 moreover have "0 < e*d" using `0<e` `0<d` by (rule mult_pos_pos)
  2504 moreover have "c : S" using assms rel_interior_subset by auto
  2504 moreover have "c : S" using assms rel_interior_subset by auto
  2505 moreover hence "x - e *\<^sub>R (x - c) : S"
  2505 moreover hence "x - e *\<^sub>R (x - c) : S"
  2506    using mem_convex[of S x c e] apply (simp add: algebra_simps) using assms by auto
  2506    using mem_convex[of S x c e] apply (simp add: algebra_simps) using assms by auto
  2507 ultimately show ?thesis 
  2507 ultimately show ?thesis
  2508   using mem_rel_interior_ball[of "x - e *\<^sub>R (x - c)" S] `e>0` by auto
  2508   using mem_rel_interior_ball[of "x - e *\<^sub>R (x - c)" S] `e>0` by auto
  2509 qed
  2509 qed
  2510 
  2510 
  2511 lemma interior_real_semiline:
  2511 lemma interior_real_semiline:
  2512 fixes a :: real
  2512 fixes a :: real
  2513 shows "interior {a..} = {a<..}"
  2513 shows "interior {a..} = {a<..}"
  2514 proof-
  2514 proof-
  2515 { fix y assume "a<y" hence "y : interior {a..}"
  2515 { fix y assume "a<y" hence "y : interior {a..}"
  2516   apply (simp add: mem_interior) apply (rule_tac x="(y-a)" in exI) apply (auto simp add: dist_norm) 
  2516   apply (simp add: mem_interior) apply (rule_tac x="(y-a)" in exI) apply (auto simp add: dist_norm)
  2517   done }
  2517   done }
  2518 moreover
  2518 moreover
  2519 { fix y assume "y : interior {a..}" (*hence "a<=y" using interior_subset by auto*)
  2519 { fix y assume "y : interior {a..}" (*hence "a<=y" using interior_subset by auto*)
  2520   from this obtain e where e_def: "e>0 & cball y e \<subseteq> {a..}" 
  2520   from this obtain e where e_def: "e>0 & cball y e \<subseteq> {a..}"
  2521      using mem_interior_cball[of y "{a..}"] by auto
  2521      using mem_interior_cball[of y "{a..}"] by auto
  2522   moreover hence "y-e : cball y e" by (auto simp add: cball_def dist_norm) 
  2522   moreover hence "y-e : cball y e" by (auto simp add: cball_def dist_norm)
  2523   ultimately have "a<=y-e" by auto
  2523   ultimately have "a<=y-e" by auto
  2524   hence "a<y" using e_def by auto
  2524   hence "a<y" using e_def by auto
  2525 } ultimately show ?thesis by auto
  2525 } ultimately show ?thesis by auto
  2526 qed
  2526 qed
  2527 
  2527 
  2549 
  2549 
  2550 lemma rel_open: "rel_open S <-> openin (subtopology euclidean (affine hull S)) S"
  2550 lemma rel_open: "rel_open S <-> openin (subtopology euclidean (affine hull S)) S"
  2551  unfolding rel_open_def rel_interior_def apply auto
  2551  unfolding rel_open_def rel_interior_def apply auto
  2552  using openin_subopen[of "subtopology euclidean (affine hull S)" S] by auto
  2552  using openin_subopen[of "subtopology euclidean (affine hull S)" S] by auto
  2553 
  2553 
  2554 lemma opein_rel_interior: 
  2554 lemma opein_rel_interior:
  2555   "openin (subtopology euclidean (affine hull S)) (rel_interior S)"
  2555   "openin (subtopology euclidean (affine hull S)) (rel_interior S)"
  2556   apply (simp add: rel_interior_def)
  2556   apply (simp add: rel_interior_def)
  2557   apply (subst openin_subopen) by blast
  2557   apply (subst openin_subopen) by blast
  2558 
  2558 
  2559 lemma affine_rel_open: 
  2559 lemma affine_rel_open:
  2560   fixes S :: "('n::euclidean_space) set"
  2560   fixes S :: "('n::euclidean_space) set"
  2561   assumes "affine S" shows "rel_open S" 
  2561   assumes "affine S" shows "rel_open S"
  2562   unfolding rel_open_def using assms rel_interior_univ[of S] affine_hull_eq[of S] by metis
  2562   unfolding rel_open_def using assms rel_interior_univ[of S] affine_hull_eq[of S] by metis
  2563 
  2563 
  2564 lemma affine_closed: 
  2564 lemma affine_closed:
  2565   fixes S :: "('n::euclidean_space) set"
  2565   fixes S :: "('n::euclidean_space) set"
  2566   assumes "affine S" shows "closed S"
  2566   assumes "affine S" shows "closed S"
  2567 proof-
  2567 proof-
  2568 { assume "S ~= {}"
  2568 { assume "S ~= {}"
  2569   from this obtain L where L_def: "subspace L & affine_parallel S L"
  2569   from this obtain L where L_def: "subspace L & affine_parallel S L"
  2570      using assms affine_parallel_subspace[of S] by auto
  2570      using assms affine_parallel_subspace[of S] by auto
  2571   from this obtain "a" where a_def: "S=(op + a ` L)" 
  2571   from this obtain "a" where a_def: "S=(op + a ` L)"
  2572      using affine_parallel_def[of L S] affine_parallel_commut by auto 
  2572      using affine_parallel_def[of L S] affine_parallel_commut by auto
  2573   have "closed L" using L_def closed_subspace by auto
  2573   have "closed L" using L_def closed_subspace by auto
  2574   hence "closed S" using closed_translation a_def by auto
  2574   hence "closed S" using closed_translation a_def by auto
  2575 } from this show ?thesis by auto
  2575 } from this show ?thesis by auto
  2576 qed
  2576 qed
  2577 
  2577 
  2584   fixes S :: "('n::euclidean_space) set"
  2584   fixes S :: "('n::euclidean_space) set"
  2585   shows "affine hull (closure S) = affine hull S"
  2585   shows "affine hull (closure S) = affine hull S"
  2586 proof-
  2586 proof-
  2587 have "affine hull (closure S) <= affine hull S"
  2587 have "affine hull (closure S) <= affine hull S"
  2588    using hull_mono[of "closure S" "affine hull S" "affine"] closure_affine_hull[of S] hull_hull[of "affine" S] by auto
  2588    using hull_mono[of "closure S" "affine hull S" "affine"] closure_affine_hull[of S] hull_hull[of "affine" S] by auto
  2589 moreover have "affine hull (closure S) >= affine hull S"  
  2589 moreover have "affine hull (closure S) >= affine hull S"
  2590    using hull_mono[of "S" "closure S" "affine"] closure_subset by auto
  2590    using hull_mono[of "S" "closure S" "affine"] closure_subset by auto
  2591 ultimately show ?thesis by auto  
  2591 ultimately show ?thesis by auto
  2592 qed
  2592 qed
  2593 
  2593 
  2594 lemma closure_aff_dim: 
  2594 lemma closure_aff_dim:
  2595   fixes S :: "('n::euclidean_space) set"
  2595   fixes S :: "('n::euclidean_space) set"
  2596   shows "aff_dim (closure S) = aff_dim S"
  2596   shows "aff_dim (closure S) = aff_dim S"
  2597 proof-
  2597 proof-
  2598 have "aff_dim S <= aff_dim (closure S)" using aff_dim_subset closure_subset by auto
  2598 have "aff_dim S <= aff_dim (closure S)" using aff_dim_subset closure_subset by auto
  2599 moreover have "aff_dim (closure S) <= aff_dim (affine hull S)" 
  2599 moreover have "aff_dim (closure S) <= aff_dim (affine hull S)"
  2600   using aff_dim_subset closure_affine_hull by auto
  2600   using aff_dim_subset closure_affine_hull by auto
  2601 moreover have "aff_dim (affine hull S) = aff_dim S" using aff_dim_affine_hull by auto
  2601 moreover have "aff_dim (affine hull S) = aff_dim S" using aff_dim_affine_hull by auto
  2602 ultimately show ?thesis by auto
  2602 ultimately show ?thesis by auto
  2603 qed
  2603 qed
  2604 
  2604 
  2605 lemma rel_interior_closure_convex_shrink:
  2605 lemma rel_interior_closure_convex_shrink:
  2606   fixes S :: "(_::euclidean_space) set"
  2606   fixes S :: "(_::euclidean_space) set"
  2607   assumes "convex S" "c : rel_interior S" "x : closure S" "0 < e" "e <= 1"
  2607   assumes "convex S" "c : rel_interior S" "x : closure S" "0 < e" "e <= 1"
  2608   shows "x - e *\<^sub>R (x - c) : rel_interior S"
  2608   shows "x - e *\<^sub>R (x - c) : rel_interior S"
  2609 proof- 
  2609 proof-
  2610 (* Proof is a modified copy of the proof of similar lemma mem_interior_closure_convex_shrink
  2610 (* Proof is a modified copy of the proof of similar lemma mem_interior_closure_convex_shrink
  2611 *)
  2611 *)
  2612 obtain d where "d>0" and d:"ball c d Int affine hull S <= S" 
  2612 obtain d where "d>0" and d:"ball c d Int affine hull S <= S"
  2613   using assms(2) unfolding mem_rel_interior_ball by auto
  2613   using assms(2) unfolding mem_rel_interior_ball by auto
  2614 have "EX y : S. norm (y - x) * (1 - e) < e * d" proof(cases "x : S")
  2614 have "EX y : S. norm (y - x) * (1 - e) < e * d" proof(cases "x : S")
  2615     case True thus ?thesis using `e>0` `d>0` by(rule_tac bexI[where x=x], auto intro!: mult_pos_pos) next
  2615     case True thus ?thesis using `e>0` `d>0` by(rule_tac bexI[where x=x], auto intro!: mult_pos_pos) next
  2616     case False hence x:"x islimpt S" using assms(3)[unfolded closure_def] by auto
  2616     case False hence x:"x islimpt S" using assms(3)[unfolded closure_def] by auto
  2617     show ?thesis proof(cases "e=1")
  2617     show ?thesis proof(cases "e=1")
  2629   have zball: "z\<in>ball c d"
  2629   have zball: "z\<in>ball c d"
  2630     using mem_ball z_def dist_norm[of c] using y and assms(4,5) by (auto simp add:field_simps norm_minus_commute)
  2630     using mem_ball z_def dist_norm[of c] using y and assms(4,5) by (auto simp add:field_simps norm_minus_commute)
  2631   have "x : affine hull S" using closure_affine_hull assms by auto
  2631   have "x : affine hull S" using closure_affine_hull assms by auto
  2632   moreover have "y : affine hull S" using `y : S` hull_subset[of S] by auto
  2632   moreover have "y : affine hull S" using `y : S` hull_subset[of S] by auto
  2633   moreover have "c : affine hull S" using assms rel_interior_subset hull_subset[of S] by auto
  2633   moreover have "c : affine hull S" using assms rel_interior_subset hull_subset[of S] by auto
  2634   ultimately have "z : affine hull S" 
  2634   ultimately have "z : affine hull S"
  2635     using z_def affine_affine_hull[of S] 
  2635     using z_def affine_affine_hull[of S]
  2636           mem_affine_3_minus [of "affine hull S" c x y "(1 - e) / e"] 
  2636           mem_affine_3_minus [of "affine hull S" c x y "(1 - e) / e"]
  2637           assms by (auto simp add: field_simps)
  2637           assms by (auto simp add: field_simps)
  2638   hence "z : S" using d zball by auto
  2638   hence "z : S" using d zball by auto
  2639   obtain d1 where "d1>0" and d1:"ball z d1 <= ball c d"
  2639   obtain d1 where "d1>0" and d1:"ball z d1 <= ball c d"
  2640     using zball open_ball[of c d] openE[of "ball c d" z] by auto
  2640     using zball open_ball[of c d] openE[of "ball c d" z] by auto
  2641   hence "(ball z d1) Int (affine hull S) <= (ball c d) Int (affine hull S)" by auto
  2641   hence "(ball z d1) Int (affine hull S) <= (ball c d) Int (affine hull S)" by auto
  2642   hence "(ball z d1) Int (affine hull S) <= S" using d by auto 
  2642   hence "(ball z d1) Int (affine hull S) <= S" using d by auto
  2643   hence "z : rel_interior S" using mem_rel_interior_ball using `d1>0` `z : S` by auto
  2643   hence "z : rel_interior S" using mem_rel_interior_ball using `d1>0` `z : S` by auto
  2644   hence "y - e *\<^sub>R (y - z) : rel_interior S" using rel_interior_convex_shrink[of S z y e] assms`y : S` by auto
  2644   hence "y - e *\<^sub>R (y - z) : rel_interior S" using rel_interior_convex_shrink[of S z y e] assms`y : S` by auto
  2645   thus ?thesis using * by auto 
  2645   thus ?thesis using * by auto
  2646 qed
  2646 qed
  2647 
  2647 
  2648 subsubsection{* Relative interior preserves under linear transformations *}
  2648 subsubsection{* Relative interior preserves under linear transformations *}
  2649 
  2649 
  2650 lemma rel_interior_translation_aux:
  2650 lemma rel_interior_translation_aux:
  2651 fixes a :: "'n::euclidean_space"
  2651 fixes a :: "'n::euclidean_space"
  2652 shows "((%x. a + x) ` rel_interior S) <= rel_interior ((%x. a + x) ` S)"
  2652 shows "((%x. a + x) ` rel_interior S) <= rel_interior ((%x. a + x) ` S)"
  2653 proof-
  2653 proof-
  2654 { fix x assume x_def: "x : rel_interior S"
  2654 { fix x assume x_def: "x : rel_interior S"
  2655   from this obtain T where T_def: "open T & x : (T Int S) & (T Int (affine hull S)) <= S" using mem_rel_interior[of x S] by auto 
  2655   from this obtain T where T_def: "open T & x : (T Int S) & (T Int (affine hull S)) <= S" using mem_rel_interior[of x S] by auto
  2656   from this have "open ((%x. a + x) ` T)" and 
  2656   from this have "open ((%x. a + x) ` T)" and
  2657     "(a + x) : (((%x. a + x) ` T) Int ((%x. a + x) ` S))" and 
  2657     "(a + x) : (((%x. a + x) ` T) Int ((%x. a + x) ` S))" and
  2658     "(((%x. a + x) ` T) Int (affine hull ((%x. a + x) ` S))) <= ((%x. a + x) ` S)" 
  2658     "(((%x. a + x) ` T) Int (affine hull ((%x. a + x) ` S))) <= ((%x. a + x) ` S)"
  2659     using affine_hull_translation[of a S] open_translation[of T a] x_def by auto 
  2659     using affine_hull_translation[of a S] open_translation[of T a] x_def by auto
  2660   from this have "(a+x) : rel_interior ((%x. a + x) ` S)" 
  2660   from this have "(a+x) : rel_interior ((%x. a + x) ` S)"
  2661     using mem_rel_interior[of "a+x" "((%x. a + x) ` S)"] by auto 
  2661     using mem_rel_interior[of "a+x" "((%x. a + x) ` S)"] by auto
  2662 } from this show ?thesis by auto 
  2662 } from this show ?thesis by auto
  2663 qed
  2663 qed
  2664 
  2664 
  2665 lemma rel_interior_translation:
  2665 lemma rel_interior_translation:
  2666 fixes a :: "'n::euclidean_space"
  2666 fixes a :: "'n::euclidean_space"
  2667 shows "rel_interior ((%x. a + x) ` S) = ((%x. a + x) ` rel_interior S)"
  2667 shows "rel_interior ((%x. a + x) ` S) = ((%x. a + x) ` rel_interior S)"
  2668 proof-
  2668 proof-
  2669 have "(%x. (-a) + x) ` rel_interior ((%x. a + x) ` S) <= rel_interior S" 
  2669 have "(%x. (-a) + x) ` rel_interior ((%x. a + x) ` S) <= rel_interior S"
  2670    using rel_interior_translation_aux[of "-a" "(%x. a + x) ` S"] 
  2670    using rel_interior_translation_aux[of "-a" "(%x. a + x) ` S"]
  2671          translation_assoc[of "-a" "a"] by auto
  2671          translation_assoc[of "-a" "a"] by auto
  2672 hence "((%x. a + x) ` rel_interior S) >= rel_interior ((%x. a + x) ` S)" 
  2672 hence "((%x. a + x) ` rel_interior S) >= rel_interior ((%x. a + x) ` S)"
  2673    using translation_inverse_subset[of a "rel_interior (op + a ` S)" "rel_interior S"] 
  2673    using translation_inverse_subset[of a "rel_interior (op + a ` S)" "rel_interior S"]
  2674    by auto
  2674    by auto
  2675 from this show ?thesis using  rel_interior_translation_aux[of a S] by auto 
  2675 from this show ?thesis using  rel_interior_translation_aux[of a S] by auto
  2676 qed
  2676 qed
  2677 
  2677 
  2678 
  2678 
  2679 lemma affine_hull_linear_image:
  2679 lemma affine_hull_linear_image:
  2680 assumes "bounded_linear f"
  2680 assumes "bounded_linear f"
  2681 shows "f ` (affine hull s) = affine hull f ` s"
  2681 shows "f ` (affine hull s) = affine hull f ` s"
  2682 (* Proof is a modified copy of the proof of similar lemma convex_hull_linear_image
  2682 (* Proof is a modified copy of the proof of similar lemma convex_hull_linear_image
  2683 *)
  2683 *)
  2684   apply rule unfolding subset_eq ball_simps apply(rule_tac[!] hull_induct, rule hull_inc) prefer 3  
  2684   apply rule unfolding subset_eq ball_simps apply(rule_tac[!] hull_induct, rule hull_inc) prefer 3
  2685   apply(erule imageE)apply(rule_tac x=xa in image_eqI) apply assumption
  2685   apply(erule imageE)apply(rule_tac x=xa in image_eqI) apply assumption
  2686   apply(rule hull_subset[unfolded subset_eq, rule_format]) apply assumption
  2686   apply(rule hull_subset[unfolded subset_eq, rule_format]) apply assumption
  2687 proof-
  2687 proof-
  2688   interpret f: bounded_linear f by fact
  2688   interpret f: bounded_linear f by fact
  2689   show "affine {x. f x : affine hull f ` s}" 
  2689   show "affine {x. f x : affine hull f ` s}"
  2690   unfolding affine_def by(auto simp add: f.scaleR f.add affine_affine_hull[unfolded affine_def, rule_format]) next
  2690   unfolding affine_def by(auto simp add: f.scaleR f.add affine_affine_hull[unfolded affine_def, rule_format]) next
  2691   interpret f: bounded_linear f by fact
  2691   interpret f: bounded_linear f by fact
  2692   show "affine {x. x : f ` (affine hull s)}" using affine_affine_hull[unfolded affine_def, of s] 
  2692   show "affine {x. x : f ` (affine hull s)}" using affine_affine_hull[unfolded affine_def, of s]
  2693     unfolding affine_def by (auto simp add: f.scaleR [symmetric] f.add [symmetric])
  2693     unfolding affine_def by (auto simp add: f.scaleR [symmetric] f.add [symmetric])
  2694 qed auto
  2694 qed auto
  2695 
  2695 
  2696 
  2696 
  2697 lemma rel_interior_injective_on_span_linear_image:
  2697 lemma rel_interior_injective_on_span_linear_image:
  2701 shows "rel_interior (f ` S) = f ` (rel_interior S)"
  2701 shows "rel_interior (f ` S) = f ` (rel_interior S)"
  2702 proof-
  2702 proof-
  2703 { fix z assume z_def: "z : rel_interior (f ` S)"
  2703 { fix z assume z_def: "z : rel_interior (f ` S)"
  2704   have "z : f ` S" using z_def rel_interior_subset[of "f ` S"] by auto
  2704   have "z : f ` S" using z_def rel_interior_subset[of "f ` S"] by auto
  2705   from this obtain x where x_def: "x : S & (f x = z)" by auto
  2705   from this obtain x where x_def: "x : S & (f x = z)" by auto
  2706   obtain e2 where e2_def: "e2>0 & cball z e2 Int affine hull (f ` S) <= (f ` S)" 
  2706   obtain e2 where e2_def: "e2>0 & cball z e2 Int affine hull (f ` S) <= (f ` S)"
  2707     using z_def rel_interior_cball[of "f ` S"] by auto
  2707     using z_def rel_interior_cball[of "f ` S"] by auto
  2708   obtain K where K_def: "K>0 & (! x. norm (f x) <= norm x * K)" 
  2708   obtain K where K_def: "K>0 & (! x. norm (f x) <= norm x * K)"
  2709    using assms RealVector.bounded_linear.pos_bounded[of f] by auto
  2709    using assms RealVector.bounded_linear.pos_bounded[of f] by auto
  2710   def e1 == "1/K" hence e1_def: "e1>0 & (! x. e1 * norm (f x) <= norm x)" 
  2710   def e1 == "1/K" hence e1_def: "e1>0 & (! x. e1 * norm (f x) <= norm x)"
  2711    using K_def pos_le_divide_eq[of e1] by auto
  2711    using K_def pos_le_divide_eq[of e1] by auto
  2712   def e == "e1 * e2" hence "e>0" using e1_def e2_def mult_pos_pos by auto 
  2712   def e == "e1 * e2" hence "e>0" using e1_def e2_def mult_pos_pos by auto
  2713   { fix y assume y_def: "y : cball x e Int affine hull S"
  2713   { fix y assume y_def: "y : cball x e Int affine hull S"
  2714     from this have h1: "f y : affine hull (f ` S)" 
  2714     from this have h1: "f y : affine hull (f ` S)"
  2715       using affine_hull_linear_image[of f S] assms by auto 
  2715       using affine_hull_linear_image[of f S] assms by auto
  2716     from y_def have "norm (x-y)<=e1 * e2" 
  2716     from y_def have "norm (x-y)<=e1 * e2"
  2717       using cball_def[of x e] dist_norm[of x y] e_def by auto
  2717       using cball_def[of x e] dist_norm[of x y] e_def by auto
  2718     moreover have "(f x)-(f y)=f (x-y)"
  2718     moreover have "(f x)-(f y)=f (x-y)"
  2719        using assms linear_sub[of f x y] linear_conv_bounded_linear[of f] by auto
  2719        using assms linear_sub[of f x y] linear_conv_bounded_linear[of f] by auto
  2720     moreover have "e1 * norm (f (x-y)) <= norm (x-y)" using e1_def by auto
  2720     moreover have "e1 * norm (f (x-y)) <= norm (x-y)" using e1_def by auto
  2721     ultimately have "e1 * norm ((f x)-(f y)) <= e1 * e2" by auto
  2721     ultimately have "e1 * norm ((f x)-(f y)) <= e1 * e2" by auto
  2722     hence "(f y) : (cball z e2)" 
  2722     hence "(f y) : (cball z e2)"
  2723       using cball_def[of "f x" e2] dist_norm[of "f x" "f y"] e1_def x_def by auto
  2723       using cball_def[of "f x" e2] dist_norm[of "f x" "f y"] e1_def x_def by auto
  2724     hence "f y : (f ` S)" using y_def e2_def h1 by auto
  2724     hence "f y : (f ` S)" using y_def e2_def h1 by auto
  2725     hence "y : S" using assms y_def hull_subset[of S] affine_hull_subset_span 
  2725     hence "y : S" using assms y_def hull_subset[of S] affine_hull_subset_span
  2726          inj_on_image_mem_iff[of f "span S" S y] by auto
  2726          inj_on_image_mem_iff[of f "span S" S y] by auto
  2727   } 
  2727   }
  2728   hence "z : f ` (rel_interior S)" using mem_rel_interior_cball[of x S] `e>0` x_def by auto
  2728   hence "z : f ` (rel_interior S)" using mem_rel_interior_cball[of x S] `e>0` x_def by auto
  2729 } 
  2729 }
  2730 moreover
  2730 moreover
  2731 { fix x assume x_def: "x : rel_interior S"
  2731 { fix x assume x_def: "x : rel_interior S"
  2732   from this obtain e2 where e2_def: "e2>0 & cball x e2 Int affine hull S <= S" 
  2732   from this obtain e2 where e2_def: "e2>0 & cball x e2 Int affine hull S <= S"
  2733     using rel_interior_cball[of S] by auto
  2733     using rel_interior_cball[of S] by auto
  2734   have "x : S" using x_def rel_interior_subset by auto
  2734   have "x : S" using x_def rel_interior_subset by auto
  2735   hence *: "f x : f ` S" by auto
  2735   hence *: "f x : f ` S" by auto
  2736   have "! x:span S. f x = 0 --> x = 0" 
  2736   have "! x:span S. f x = 0 --> x = 0"
  2737     using assms subspace_span linear_conv_bounded_linear[of f] 
  2737     using assms subspace_span linear_conv_bounded_linear[of f]
  2738           linear_injective_on_subspace_0[of f "span S"] by auto
  2738           linear_injective_on_subspace_0[of f "span S"] by auto
  2739   from this obtain e1 where e1_def: "e1>0 & (! x : span S. e1 * norm x <= norm (f x))" 
  2739   from this obtain e1 where e1_def: "e1>0 & (! x : span S. e1 * norm x <= norm (f x))"
  2740    using assms injective_imp_isometric[of "span S" f] 
  2740    using assms injective_imp_isometric[of "span S" f]
  2741          subspace_span[of S] closed_subspace[of "span S"] by auto
  2741          subspace_span[of S] closed_subspace[of "span S"] by auto
  2742   def e == "e1 * e2" hence "e>0" using e1_def e2_def mult_pos_pos by auto 
  2742   def e == "e1 * e2" hence "e>0" using e1_def e2_def mult_pos_pos by auto
  2743   { fix y assume y_def: "y : cball (f x) e Int affine hull (f ` S)"
  2743   { fix y assume y_def: "y : cball (f x) e Int affine hull (f ` S)"
  2744     from this have "y : f ` (affine hull S)" using affine_hull_linear_image[of f S] assms by auto 
  2744     from this have "y : f ` (affine hull S)" using affine_hull_linear_image[of f S] assms by auto
  2745     from this obtain xy where xy_def: "xy : affine hull S & (f xy = y)" by auto
  2745     from this obtain xy where xy_def: "xy : affine hull S & (f xy = y)" by auto
  2746     from this y_def have "norm ((f x)-(f xy))<=e1 * e2" 
  2746     from this y_def have "norm ((f x)-(f xy))<=e1 * e2"
  2747       using cball_def[of "f x" e] dist_norm[of "f x" y] e_def by auto
  2747       using cball_def[of "f x" e] dist_norm[of "f x" y] e_def by auto
  2748     moreover have "(f x)-(f xy)=f (x-xy)"
  2748     moreover have "(f x)-(f xy)=f (x-xy)"
  2749        using assms linear_sub[of f x xy] linear_conv_bounded_linear[of f] by auto
  2749        using assms linear_sub[of f x xy] linear_conv_bounded_linear[of f] by auto
  2750     moreover have "x-xy : span S" 
  2750     moreover have "x-xy : span S"
  2751        using subspace_sub[of "span S" x xy] subspace_span `x : S` xy_def 
  2751        using subspace_sub[of "span S" x xy] subspace_span `x : S` xy_def
  2752              affine_hull_subset_span[of S] span_inc by auto
  2752              affine_hull_subset_span[of S] span_inc by auto
  2753     moreover hence "e1 * norm (x-xy) <= norm (f (x-xy))" using e1_def by auto
  2753     moreover hence "e1 * norm (x-xy) <= norm (f (x-xy))" using e1_def by auto
  2754     ultimately have "e1 * norm (x-xy) <= e1 * e2" by auto
  2754     ultimately have "e1 * norm (x-xy) <= e1 * e2" by auto
  2755     hence "xy : (cball x e2)"  using cball_def[of x e2] dist_norm[of x xy] e1_def by auto
  2755     hence "xy : (cball x e2)"  using cball_def[of x e2] dist_norm[of x xy] e1_def by auto
  2756     hence "y : (f ` S)" using xy_def e2_def by auto
  2756     hence "y : (f ` S)" using xy_def e2_def by auto
  2757   } 
  2757   }
  2758   hence "(f x) : rel_interior (f ` S)" 
  2758   hence "(f x) : rel_interior (f ` S)"
  2759      using mem_rel_interior_cball[of "(f x)" "(f ` S)"] * `e>0` by auto
  2759      using mem_rel_interior_cball[of "(f x)" "(f ` S)"] * `e>0` by auto
  2760 } 
  2760 }
  2761 ultimately show ?thesis by auto
  2761 ultimately show ?thesis by auto
  2762 qed
  2762 qed
  2763 
  2763 
  2764 lemma rel_interior_injective_linear_image:
  2764 lemma rel_interior_injective_linear_image:
  2765 fixes f :: "('m::euclidean_space) => ('n::euclidean_space)"
  2765 fixes f :: "('m::euclidean_space) => ('n::euclidean_space)"
  2766 assumes "bounded_linear f" and "inj f"
  2766 assumes "bounded_linear f" and "inj f"
  2767 shows "rel_interior (f ` S) = f ` (rel_interior S)"
  2767 shows "rel_interior (f ` S) = f ` (rel_interior S)"
  2768 using assms rel_interior_injective_on_span_linear_image[of f S] 
  2768 using assms rel_interior_injective_on_span_linear_image[of f S]
  2769       subset_inj_on[of f "UNIV" "span S"] by auto
  2769       subset_inj_on[of f "UNIV" "span S"] by auto
  2770 
  2770 
  2771 subsection{* Some Properties of subset of standard basis *}
  2771 subsection{* Some Properties of subset of standard basis *}
  2772 
  2772 
  2773 lemma affine_hull_substd_basis: assumes "d\<subseteq>{..<DIM('a::euclidean_space)}"
  2773 lemma affine_hull_substd_basis: assumes "d\<subseteq>{..<DIM('a::euclidean_space)}"
  2880   case True thus ?thesis using compact_empty by simp
  2880   case True thus ?thesis using compact_empty by simp
  2881 next
  2881 next
  2882   case False then obtain w where "w\<in>s" by auto
  2882   case False then obtain w where "w\<in>s" by auto
  2883   show ?thesis unfolding caratheodory[of s]
  2883   show ?thesis unfolding caratheodory[of s]
  2884   proof(induct ("DIM('a) + 1"))
  2884   proof(induct ("DIM('a) + 1"))
  2885     have *:"{x.\<exists>sa. finite sa \<and> sa \<subseteq> s \<and> card sa \<le> 0 \<and> x \<in> convex hull sa} = {}" 
  2885     have *:"{x.\<exists>sa. finite sa \<and> sa \<subseteq> s \<and> card sa \<le> 0 \<and> x \<in> convex hull sa} = {}"
  2886       using compact_empty by auto
  2886       using compact_empty by auto
  2887     case 0 thus ?case unfolding * by simp
  2887     case 0 thus ?case unfolding * by simp
  2888   next
  2888   next
  2889     case (Suc n)
  2889     case (Suc n)
  2890     show ?case proof(cases "n=0")
  2890     show ?case proof(cases "n=0")
  2900           thus ?thesis using t(2,4) by simp
  2900           thus ?thesis using t(2,4) by simp
  2901         qed
  2901         qed
  2902       next
  2902       next
  2903         fix x assume "x\<in>s"
  2903         fix x assume "x\<in>s"
  2904         thus "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
  2904         thus "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
  2905           apply(rule_tac x="{x}" in exI) unfolding convex_hull_singleton by auto 
  2905           apply(rule_tac x="{x}" in exI) unfolding convex_hull_singleton by auto
  2906       qed thus ?thesis using assms by simp
  2906       qed thus ?thesis using assms by simp
  2907     next
  2907     next
  2908       case False have "{x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t} =
  2908       case False have "{x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t} =
  2909         { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 
  2909         { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u.
  2910         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}}"
  2910         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}}"
  2911         unfolding set_eq_iff and mem_Collect_eq proof(rule,rule)
  2911         unfolding set_eq_iff and mem_Collect_eq proof(rule,rule)
  2912         fix x assume "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and>
  2912         fix x assume "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and>
  2913           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)"
  2913           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)"
  2914         then obtain u v c t where obt:"x = (1 - c) *\<^sub>R u + c *\<^sub>R v"
  2914         then obtain u v c t where obt:"x = (1 - c) *\<^sub>R u + c *\<^sub>R v"
  2941               using obt and t(1-3) unfolding au and * using card_insert_disjoint[OF _ au(2)]
  2941               using obt and t(1-3) unfolding au and * using card_insert_disjoint[OF _ au(2)]
  2942               by(auto intro!: exI[where x=u])
  2942               by(auto intro!: exI[where x=u])
  2943           qed
  2943           qed
  2944         qed
  2944         qed
  2945       qed
  2945       qed
  2946       thus ?thesis using compact_convex_combinations[OF assms Suc] by simp 
  2946       thus ?thesis using compact_convex_combinations[OF assms Suc] by simp
  2947     qed
  2947     qed
  2948   qed
  2948   qed
  2949 qed
  2949 qed
  2950 
  2950 
  2951 subsection {* Extremal points of a simplex are some vertices. *}
  2951 subsection {* Extremal points of a simplex are some vertices. *}
  2953 lemma dist_increases_online:
  2953 lemma dist_increases_online:
  2954   fixes a b d :: "'a::real_inner"
  2954   fixes a b d :: "'a::real_inner"
  2955   assumes "d \<noteq> 0"
  2955   assumes "d \<noteq> 0"
  2956   shows "dist a (b + d) > dist a b \<or> dist a (b - d) > dist a b"
  2956   shows "dist a (b + d) > dist a b \<or> dist a (b - d) > dist a b"
  2957 proof(cases "inner a d - inner b d > 0")
  2957 proof(cases "inner a d - inner b d > 0")
  2958   case True hence "0 < inner d d + (inner a d * 2 - inner b d * 2)" 
  2958   case True hence "0 < inner d d + (inner a d * 2 - inner b d * 2)"
  2959     apply(rule_tac add_pos_pos) using assms by auto
  2959     apply(rule_tac add_pos_pos) using assms by auto
  2960   thus ?thesis apply(rule_tac disjI2) unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff
  2960   thus ?thesis apply(rule_tac disjI2) unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff
  2961     by (simp add: algebra_simps inner_commute)
  2961     by (simp add: algebra_simps inner_commute)
  2962 next
  2962 next
  2963   case False hence "0 < inner d d + (inner b d * 2 - inner a d * 2)" 
  2963   case False hence "0 < inner d d + (inner b d * 2 - inner a d * 2)"
  2964     apply(rule_tac add_pos_nonneg) using assms by auto
  2964     apply(rule_tac add_pos_nonneg) using assms by auto
  2965   thus ?thesis apply(rule_tac disjI1) unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff
  2965   thus ?thesis apply(rule_tac disjI1) unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff
  2966     by (simp add: algebra_simps inner_commute)
  2966     by (simp add: algebra_simps inner_commute)
  2967 qed
  2967 qed
  2968 
  2968 
  2994         assume "u\<noteq>0" "v=0" hence "y = x" using obt by auto
  2994         assume "u\<noteq>0" "v=0" hence "y = x" using obt by auto
  2995         thus ?thesis using y(2) by auto
  2995         thus ?thesis using y(2) by auto
  2996       next
  2996       next
  2997         assume "u\<noteq>0" "v\<noteq>0"
  2997         assume "u\<noteq>0" "v\<noteq>0"
  2998         then obtain w where w:"w>0" "w<u" "w<v" using real_lbound_gt_zero[of u v] and obt(1,2) by auto
  2998         then obtain w where w:"w>0" "w<u" "w<v" using real_lbound_gt_zero[of u v] and obt(1,2) by auto
  2999         have "x\<noteq>b" proof(rule ccontr) 
  2999         have "x\<noteq>b" proof(rule ccontr)
  3000           assume "\<not> x\<noteq>b" hence "y=b" unfolding obt(5)
  3000           assume "\<not> x\<noteq>b" hence "y=b" unfolding obt(5)
  3001             using obt(3) by(auto simp add: scaleR_left_distrib[symmetric])
  3001             using obt(3) by(auto simp add: scaleR_left_distrib[symmetric])
  3002           thus False using obt(4) and False by simp qed
  3002           thus False using obt(4) and False by simp qed
  3003         hence *:"w *\<^sub>R (x - b) \<noteq> 0" using w(1) by auto
  3003         hence *:"w *\<^sub>R (x - b) \<noteq> 0" using w(1) by auto
  3004         show ?thesis using dist_increases_online[OF *, of a y]
  3004         show ?thesis using dist_increases_online[OF *, of a y]
  3006           assume "dist a y < dist a (y + w *\<^sub>R (x - b))"
  3006           assume "dist a y < dist a (y + w *\<^sub>R (x - b))"
  3007           hence "norm (y - a) < norm ((u + w) *\<^sub>R x + (v - w) *\<^sub>R b - a)"
  3007           hence "norm (y - a) < norm ((u + w) *\<^sub>R x + (v - w) *\<^sub>R b - a)"
  3008             unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps)
  3008             unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps)
  3009           moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \<in> convex hull insert x s"
  3009           moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \<in> convex hull insert x s"
  3010             unfolding convex_hull_insert[OF `s\<noteq>{}`] and mem_Collect_eq
  3010             unfolding convex_hull_insert[OF `s\<noteq>{}`] and mem_Collect_eq
  3011             apply(rule_tac x="u + w" in exI) apply rule defer 
  3011             apply(rule_tac x="u + w" in exI) apply rule defer
  3012             apply(rule_tac x="v - w" in exI) using `u\<ge>0` and w and obt(3,4) by auto
  3012             apply(rule_tac x="v - w" in exI) using `u\<ge>0` and w and obt(3,4) by auto
  3013           ultimately show ?thesis by auto
  3013           ultimately show ?thesis by auto
  3014         next
  3014         next
  3015           assume "dist a y < dist a (y - w *\<^sub>R (x - b))"
  3015           assume "dist a y < dist a (y - w *\<^sub>R (x - b))"
  3016           hence "norm (y - a) < norm ((u - w) *\<^sub>R x + (v + w) *\<^sub>R b - a)"
  3016           hence "norm (y - a) < norm ((u - w) *\<^sub>R x + (v + w) *\<^sub>R b - a)"
  3017             unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps)
  3017             unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps)
  3018           moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \<in> convex hull insert x s"
  3018           moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \<in> convex hull insert x s"
  3019             unfolding convex_hull_insert[OF `s\<noteq>{}`] and mem_Collect_eq
  3019             unfolding convex_hull_insert[OF `s\<noteq>{}`] and mem_Collect_eq
  3020             apply(rule_tac x="u - w" in exI) apply rule defer 
  3020             apply(rule_tac x="u - w" in exI) apply rule defer
  3021             apply(rule_tac x="v + w" in exI) using `u\<ge>0` and w and obt(3,4) by auto
  3021             apply(rule_tac x="v + w" in exI) using `u\<ge>0` and w and obt(3,4) by auto
  3022           ultimately show ?thesis by auto
  3022           ultimately show ?thesis by auto
  3023         qed
  3023         qed
  3024       qed auto
  3024       qed auto
  3025     qed
  3025     qed
  3064     assume "v\<notin>s" then obtain y where "y\<in>convex hull s" "norm (v - u) < norm (y - u)"
  3064     assume "v\<notin>s" then obtain y where "y\<in>convex hull s" "norm (v - u) < norm (y - u)"
  3065       using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2) by auto
  3065       using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2) by auto
  3066     thus ?thesis using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1)
  3066     thus ?thesis using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1)
  3067       by (auto simp add: norm_minus_commute)
  3067       by (auto simp add: norm_minus_commute)
  3068   qed auto
  3068   qed auto
  3069 qed 
  3069 qed
  3070 
  3070 
  3071 lemma simplex_extremal_le_exists:
  3071 lemma simplex_extremal_le_exists:
  3072   fixes s :: "('a::real_inner) set"
  3072   fixes s :: "('a::real_inner) set"
  3073   shows "finite s \<Longrightarrow> x \<in> convex hull s \<Longrightarrow> y \<in> convex hull s
  3073   shows "finite s \<Longrightarrow> x \<in> convex hull s \<Longrightarrow> y \<in> convex hull s
  3074   \<Longrightarrow> (\<exists>u\<in>s. \<exists>v\<in>s. norm(x - y) \<le> norm(u - v))"
  3074   \<Longrightarrow> (\<exists>u\<in>s. \<exists>v\<in>s. norm(x - y) \<le> norm(u - v))"
  3081  "closest_point s a = (SOME x. x \<in> s \<and> (\<forall>y\<in>s. dist a x \<le> dist a y))"
  3081  "closest_point s a = (SOME x. x \<in> s \<and> (\<forall>y\<in>s. dist a x \<le> dist a y))"
  3082 
  3082 
  3083 lemma closest_point_exists:
  3083 lemma closest_point_exists:
  3084   assumes "closed s" "s \<noteq> {}"
  3084   assumes "closed s" "s \<noteq> {}"
  3085   shows  "closest_point s a \<in> s" "\<forall>y\<in>s. dist a (closest_point s a) \<le> dist a y"
  3085   shows  "closest_point s a \<in> s" "\<forall>y\<in>s. dist a (closest_point s a) \<le> dist a y"
  3086   unfolding closest_point_def apply(rule_tac[!] someI2_ex) 
  3086   unfolding closest_point_def apply(rule_tac[!] someI2_ex)
  3087   using distance_attains_inf[OF assms(1,2), of a] by auto
  3087   using distance_attains_inf[OF assms(1,2), of a] by auto
  3088 
  3088 
  3089 lemma closest_point_in_set:
  3089 lemma closest_point_in_set:
  3090   "closed s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> (closest_point s a) \<in> s"
  3090   "closed s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> (closest_point s a) \<in> s"
  3091   by(meson closest_point_exists)
  3091   by(meson closest_point_exists)
  3094   "closed s \<Longrightarrow> x \<in> s \<Longrightarrow> dist a (closest_point s a) \<le> dist a x"
  3094   "closed s \<Longrightarrow> x \<in> s \<Longrightarrow> dist a (closest_point s a) \<le> dist a x"
  3095   using closest_point_exists[of s] by auto
  3095   using closest_point_exists[of s] by auto
  3096 
  3096 
  3097 lemma closest_point_self:
  3097 lemma closest_point_self:
  3098   assumes "x \<in> s"  shows "closest_point s x = x"
  3098   assumes "x \<in> s"  shows "closest_point s x = x"
  3099   unfolding closest_point_def apply(rule some1_equality, rule ex1I[of _ x]) 
  3099   unfolding closest_point_def apply(rule some1_equality, rule ex1I[of _ x])
  3100   using assms by auto
  3100   using assms by auto
  3101 
  3101 
  3102 lemma closest_point_refl:
  3102 lemma closest_point_refl:
  3103  "closed s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> (closest_point s x = x \<longleftrightarrow> x \<in> s)"
  3103  "closed s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> (closest_point s x = x \<longleftrightarrow> x \<in> s)"
  3104   using closest_point_in_set[of s x] closest_point_self[of x s] by auto
  3104   using closest_point_in_set[of s x] closest_point_self[of x s] by auto
  3138   by (auto simp add: algebra_simps)
  3138   by (auto simp add: algebra_simps)
  3139 
  3139 
  3140 lemma closest_point_unique:
  3140 lemma closest_point_unique:
  3141   assumes "convex s" "closed s" "x \<in> s" "\<forall>z\<in>s. dist a x \<le> dist a z"
  3141   assumes "convex s" "closed s" "x \<in> s" "\<forall>z\<in>s. dist a x \<le> dist a z"
  3142   shows "x = closest_point s a"
  3142   shows "x = closest_point s a"
  3143   using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point s a"] 
  3143   using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point s a"]
  3144   using closest_point_exists[OF assms(2)] and assms(3) by auto
  3144   using closest_point_exists[OF assms(2)] and assms(3) by auto
  3145 
  3145 
  3146 lemma closest_point_dot:
  3146 lemma closest_point_dot:
  3147   assumes "convex s" "closed s" "x \<in> s"
  3147   assumes "convex s" "closed s" "x \<in> s"
  3148   shows "inner (a - closest_point s a) (x - closest_point s a) \<le> 0"
  3148   shows "inner (a - closest_point s a) (x - closest_point s a) \<le> 0"
  3169     by (simp add: inner_add inner_diff inner_commute) qed
  3169     by (simp add: inner_add inner_diff inner_commute) qed
  3170 
  3170 
  3171 lemma continuous_at_closest_point:
  3171 lemma continuous_at_closest_point:
  3172   assumes "convex s" "closed s" "s \<noteq> {}"
  3172   assumes "convex s" "closed s" "s \<noteq> {}"
  3173   shows "continuous (at x) (closest_point s)"
  3173   shows "continuous (at x) (closest_point s)"
  3174   unfolding continuous_at_eps_delta 
  3174   unfolding continuous_at_eps_delta
  3175   using le_less_trans[OF closest_point_lipschitz[OF assms]] by auto
  3175   using le_less_trans[OF closest_point_lipschitz[OF assms]] by auto
  3176 
  3176 
  3177 lemma continuous_on_closest_point:
  3177 lemma continuous_on_closest_point:
  3178   assumes "convex s" "closed s" "s \<noteq> {}"
  3178   assumes "convex s" "closed s" "s \<noteq> {}"
  3179   shows "continuous_on t (closest_point s)"
  3179   shows "continuous_on t (closest_point s)"
  3263     from ab have "((\<lambda>x. inner a x) ` t) *<= (inner a y - b)"
  3263     from ab have "((\<lambda>x. inner a x) ` t) *<= (inner a y - b)"
  3264       apply(erule_tac x=y in ballE) apply(rule setleI) using `y\<in>s` by auto
  3264       apply(erule_tac x=y in ballE) apply(rule setleI) using `y\<in>s` by auto
  3265     hence k:"isLub UNIV ((\<lambda>x. inner a x) ` t) k" unfolding k_def apply(rule_tac Sup) using assms(5) by auto
  3265     hence k:"isLub UNIV ((\<lambda>x. inner a x) ` t) k" unfolding k_def apply(rule_tac Sup) using assms(5) by auto
  3266     fix x assume "x\<in>t" thus "inner a x < (k + b / 2)" using `0<b` and isLubD2[OF k, of "inner a x"] by auto
  3266     fix x assume "x\<in>t" thus "inner a x < (k + b / 2)" using `0<b` and isLubD2[OF k, of "inner a x"] by auto
  3267   next
  3267   next
  3268     fix x assume "x\<in>s" 
  3268     fix x assume "x\<in>s"
  3269     hence "k \<le> inner a x - b" unfolding k_def apply(rule_tac Sup_least) using assms(5)
  3269     hence "k \<le> inner a x - b" unfolding k_def apply(rule_tac Sup_least) using assms(5)
  3270       using ab[THEN bspec[where x=x]] by auto
  3270       using ab[THEN bspec[where x=x]] by auto
  3271     thus "k + b / 2 < inner a x" using `0 < b` by auto
  3271     thus "k + b / 2 < inner a x" using `0 < b` by auto
  3272   qed
  3272   qed
  3273 qed
  3273 qed
  3306 
  3306 
  3307 lemma separating_hyperplane_sets:
  3307 lemma separating_hyperplane_sets:
  3308   assumes "convex s" "convex (t::('a::euclidean_space) set)" "s \<noteq> {}" "t \<noteq> {}" "s \<inter> t = {}"
  3308   assumes "convex s" "convex (t::('a::euclidean_space) set)" "s \<noteq> {}" "t \<noteq> {}" "s \<inter> t = {}"
  3309   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)"
  3309   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)"
  3310 proof- from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]]
  3310 proof- from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]]
  3311   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" 
  3311   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"
  3312     using assms(3-5) by auto 
  3312     using assms(3-5) by auto
  3313   hence "\<forall>x\<in>t. \<forall>y\<in>s. inner a y \<le> inner a x"
  3313   hence "\<forall>x\<in>t. \<forall>y\<in>s. inner a y \<le> inner a x"
  3314     by (force simp add: inner_diff)
  3314     by (force simp add: inner_diff)
  3315   thus ?thesis
  3315   thus ?thesis
  3316     apply(rule_tac x=a in exI, rule_tac x="Sup ((\<lambda>x. inner a x) ` s)" in exI) using `a\<noteq>0`
  3316     apply(rule_tac x=a in exI, rule_tac x="Sup ((\<lambda>x. inner a x) ` s)" in exI) using `a\<noteq>0`
  3317     apply auto
  3317     apply auto
  3318     apply (rule Sup[THEN isLubD2]) 
  3318     apply (rule Sup[THEN isLubD2])
  3319     prefer 4
  3319     prefer 4
  3320     apply (rule Sup_least) 
  3320     apply (rule Sup_least)
  3321      using assms(3-5) apply (auto simp add: setle_def)
  3321      using assms(3-5) apply (auto simp add: setle_def)
  3322     apply metis
  3322     apply metis
  3323     done
  3323     done
  3324 qed
  3324 qed
  3325 
  3325 
  3337 lemma convex_interior:
  3337 lemma convex_interior:
  3338   fixes s :: "'a::real_normed_vector set"
  3338   fixes s :: "'a::real_normed_vector set"
  3339   assumes "convex s" shows "convex(interior s)"
  3339   assumes "convex s" shows "convex(interior s)"
  3340   unfolding convex_alt Ball_def mem_interior apply(rule,rule,rule,rule,rule,rule) apply(erule exE | erule conjE)+ proof-
  3340   unfolding convex_alt Ball_def mem_interior apply(rule,rule,rule,rule,rule,rule) apply(erule exE | erule conjE)+ proof-
  3341   fix x y u assume u:"0 \<le> u" "u \<le> (1::real)"
  3341   fix x y u assume u:"0 \<le> u" "u \<le> (1::real)"
  3342   fix e d assume ed:"ball x e \<subseteq> s" "ball y d \<subseteq> s" "0<d" "0<e" 
  3342   fix e d assume ed:"ball x e \<subseteq> s" "ball y d \<subseteq> s" "0<d" "0<e"
  3343   show "\<exists>e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \<subseteq> s" apply(rule_tac x="min d e" in exI)
  3343   show "\<exists>e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \<subseteq> s" apply(rule_tac x="min d e" in exI)
  3344     apply rule unfolding subset_eq defer apply rule proof-
  3344     apply rule unfolding subset_eq defer apply rule proof-
  3345     fix z assume "z \<in> ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) (min d e)"
  3345     fix z assume "z \<in> ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) (min d e)"
  3346     hence "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \<in> s"
  3346     hence "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \<in> s"
  3347       apply(rule_tac assms[unfolded convex_alt, rule_format])
  3347       apply(rule_tac assms[unfolded convex_alt, rule_format])
  3359 
  3359 
  3360 lemma convex_hull_bilemma: fixes neg
  3360 lemma convex_hull_bilemma: fixes neg
  3361   assumes "(\<forall>s a. (convex hull (up a s)) \<subseteq> up a (convex hull s))"
  3361   assumes "(\<forall>s a. (convex hull (up a s)) \<subseteq> up a (convex hull s))"
  3362   shows "(\<forall>s. up a (up (neg a) s) = s) \<and> (\<forall>s. up (neg a) (up a s) = s) \<and> (\<forall>s t a. s \<subseteq> t \<longrightarrow> up a s \<subseteq> up a t)
  3362   shows "(\<forall>s. up a (up (neg a) s) = s) \<and> (\<forall>s. up (neg a) (up a s) = s) \<and> (\<forall>s t a. s \<subseteq> t \<longrightarrow> up a s \<subseteq> up a t)
  3363   \<Longrightarrow> \<forall>s. (convex hull (up a s)) = up a (convex hull s)"
  3363   \<Longrightarrow> \<forall>s. (convex hull (up a s)) = up a (convex hull s)"
  3364   using assms by(metis subset_antisym) 
  3364   using assms by(metis subset_antisym)
  3365 
  3365 
  3366 lemma convex_hull_translation:
  3366 lemma convex_hull_translation:
  3367   "convex hull ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (convex hull s)"
  3367   "convex hull ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (convex hull s)"
  3368   apply(rule convex_hull_bilemma[rule_format, of _ _ "\<lambda>a. -a"], rule convex_hull_translation_lemma) unfolding image_image by auto
  3368   apply(rule convex_hull_bilemma[rule_format, of _ _ "\<lambda>a. -a"], rule convex_hull_translation_lemma) unfolding image_image by auto
  3369 
  3369 
  3437 
  3437 
  3438 lemma convex_halfspace_intersection:
  3438 lemma convex_halfspace_intersection:
  3439   fixes s :: "('a::euclidean_space) set"
  3439   fixes s :: "('a::euclidean_space) set"
  3440   assumes "closed s" "convex s"
  3440   assumes "closed s" "convex s"
  3441   shows "s = \<Inter> {h. s \<subseteq> h \<and> (\<exists>a b. h = {x. inner a x \<le> b})}"
  3441   shows "s = \<Inter> {h. s \<subseteq> h \<and> (\<exists>a b. h = {x. inner a x \<le> b})}"
  3442   apply(rule set_eqI, rule) unfolding Inter_iff Ball_def mem_Collect_eq apply(rule,rule,erule conjE) proof- 
  3442   apply(rule set_eqI, rule) unfolding Inter_iff Ball_def mem_Collect_eq apply(rule,rule,erule conjE) proof-
  3443   fix x  assume "\<forall>xa. s \<subseteq> xa \<and> (\<exists>a b. xa = {x. inner a x \<le> b}) \<longrightarrow> x \<in> xa"
  3443   fix x  assume "\<forall>xa. s \<subseteq> xa \<and> (\<exists>a b. xa = {x. inner a x \<le> b}) \<longrightarrow> x \<in> xa"
  3444   hence "\<forall>a b. s \<subseteq> {x. inner a x \<le> b} \<longrightarrow> x \<in> {x. inner a x \<le> b}" by blast
  3444   hence "\<forall>a b. s \<subseteq> {x. inner a x \<le> b} \<longrightarrow> x \<in> {x. inner a x \<le> b}" by blast
  3445   thus "x\<in>s" apply(rule_tac ccontr) apply(drule separating_hyperplane_closed_point[OF assms(2,1)])
  3445   thus "x\<in>s" apply(rule_tac ccontr) apply(drule separating_hyperplane_closed_point[OF assms(2,1)])
  3446     apply(erule exE)+ apply(erule_tac x="-a" in allE, erule_tac x="-b" in allE) by auto
  3446     apply(erule exE)+ apply(erule_tac x="-a" in allE, erule_tac x="-b" in allE) by auto
  3447 qed auto
  3447 qed auto
  3464 
  3464 
  3465 lemma radon_v_lemma:
  3465 lemma radon_v_lemma:
  3466   assumes "finite s" "setsum f s = 0" "\<forall>x. g x = (0::real) \<longrightarrow> f x = (0::'a::euclidean_space)"
  3466   assumes "finite s" "setsum f s = 0" "\<forall>x. g x = (0::real) \<longrightarrow> f x = (0::'a::euclidean_space)"
  3467   shows "(setsum f {x\<in>s. 0 < g x}) = - setsum f {x\<in>s. g x < 0}"
  3467   shows "(setsum f {x\<in>s. 0 < g x}) = - setsum f {x\<in>s. g x < 0}"
  3468 proof-
  3468 proof-
  3469   have *:"\<And>x. (if 0 < g x then f x else 0) + (if g x < 0 then f x else 0) = f x" using assms(3) by auto 
  3469   have *:"\<And>x. (if 0 < g x then f x else 0) + (if g x < 0 then f x else 0) = f x" using assms(3) by auto
  3470   show ?thesis unfolding eq_neg_iff_add_eq_0 and setsum_restrict_set''[OF assms(1)] and setsum_addf[symmetric] and *
  3470   show ?thesis unfolding eq_neg_iff_add_eq_0 and setsum_restrict_set''[OF assms(1)] and setsum_addf[symmetric] and *
  3471     using assms(2) by assumption qed
  3471     using assms(2) by assumption qed
  3472 
  3472 
  3473 lemma radon_partition:
  3473 lemma radon_partition:
  3474   assumes "finite c" "affine_dependent c"
  3474   assumes "finite c" "affine_dependent c"
  3476   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" using radon_ex_lemma[OF assms] by auto
  3476   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" using radon_ex_lemma[OF assms] by auto
  3477   have fin:"finite {x \<in> c. 0 < u x}" "finite {x \<in> c. 0 > u x}" using assms(1) by auto
  3477   have fin:"finite {x \<in> c. 0 < u x}" "finite {x \<in> c. 0 > u x}" using assms(1) by auto
  3478   def z \<equiv> "(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}"
  3478   def z \<equiv> "(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}"
  3479   have "setsum u {x \<in> c. 0 < u x} \<noteq> 0" proof(cases "u v \<ge> 0")
  3479   have "setsum u {x \<in> c. 0 < u x} \<noteq> 0" proof(cases "u v \<ge> 0")
  3480     case False hence "u v < 0" by auto
  3480     case False hence "u v < 0" by auto
  3481     thus ?thesis proof(cases "\<exists>w\<in>{x \<in> c. 0 < u x}. u w > 0") 
  3481     thus ?thesis proof(cases "\<exists>w\<in>{x \<in> c. 0 < u x}. u w > 0")
  3482       case True thus ?thesis using setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] by auto
  3482       case True thus ?thesis using setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] by auto
  3483     next
  3483     next
  3484       case False hence "setsum u c \<le> setsum (\<lambda>x. if x=v then u v else 0) c" apply(rule_tac setsum_mono) by auto
  3484       case False hence "setsum u c \<le> setsum (\<lambda>x. if x=v then u v else 0) c" apply(rule_tac setsum_mono) by auto
  3485       thus ?thesis unfolding setsum_delta[OF assms(1)] using uv(2) and `u v < 0` and uv(1) by auto qed
  3485       thus ?thesis unfolding setsum_delta[OF assms(1)] using uv(2) and `u v < 0` and uv(1) by auto qed
  3486   qed (insert setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] uv(2-3), auto)
  3486   qed (insert setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] uv(2-3), auto)
  3488   hence *:"setsum u {x\<in>c. u x > 0} > 0" unfolding less_le apply(rule_tac conjI, rule_tac setsum_nonneg) by auto
  3488   hence *:"setsum u {x\<in>c. u x > 0} > 0" unfolding less_le apply(rule_tac conjI, rule_tac setsum_nonneg) by auto
  3489   moreover have "setsum u ({x \<in> c. 0 < u x} \<union> {x \<in> c. u x < 0}) = setsum u c"
  3489   moreover have "setsum u ({x \<in> c. 0 < u x} \<union> {x \<in> c. u x < 0}) = setsum u c"
  3490     "(\<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)"
  3490     "(\<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)"
  3491     using assms(1) apply(rule_tac[!] setsum_mono_zero_left) by auto
  3491     using assms(1) apply(rule_tac[!] setsum_mono_zero_left) by auto
  3492   hence "setsum u {x \<in> c. 0 < u x} = - setsum u {x \<in> c. 0 > u x}"
  3492   hence "setsum u {x \<in> c. 0 < u x} = - setsum u {x \<in> c. 0 > u x}"
  3493    "(\<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)" 
  3493    "(\<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)"
  3494     unfolding eq_neg_iff_add_eq_0 using uv(1,4) by (auto simp add:  setsum_Un_zero[OF fin, symmetric]) 
  3494     unfolding eq_neg_iff_add_eq_0 using uv(1,4) by (auto simp add:  setsum_Un_zero[OF fin, symmetric])
  3495   moreover have "\<forall>x\<in>{v \<in> c. u v < 0}. 0 \<le> inverse (setsum u {x \<in> c. 0 < u x}) * - u x" 
  3495   moreover have "\<forall>x\<in>{v \<in> c. u v < 0}. 0 \<le> inverse (setsum u {x \<in> c. 0 < u x}) * - u x"
  3496     apply (rule) apply (rule mult_nonneg_nonneg) using * by auto
  3496     apply (rule) apply (rule mult_nonneg_nonneg) using * by auto
  3497 
  3497 
  3498   ultimately have "z \<in> convex hull {v \<in> c. u v \<le> 0}" unfolding convex_hull_explicit mem_Collect_eq
  3498   ultimately have "z \<in> convex hull {v \<in> c. u v \<le> 0}" unfolding convex_hull_explicit mem_Collect_eq
  3499     apply(rule_tac x="{v \<in> c. u v < 0}" in exI, rule_tac x="\<lambda>y. inverse (setsum u {x\<in>c. u x > 0}) * - u y" in exI)
  3499     apply(rule_tac x="{v \<in> c. u v < 0}" in exI, rule_tac x="\<lambda>y. inverse (setsum u {x\<in>c. u x > 0}) * - u y" in exI)
  3500     using assms(1) unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] and z_def
  3500     using assms(1) unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] and z_def
  3501     by(auto simp add: setsum_negf setsum_right_distrib[symmetric])
  3501     by(auto simp add: setsum_negf setsum_right_distrib[symmetric])
  3502   moreover have "\<forall>x\<in>{v \<in> c. 0 < u v}. 0 \<le> inverse (setsum u {x \<in> c. 0 < u x}) * u x" 
  3502   moreover have "\<forall>x\<in>{v \<in> c. 0 < u v}. 0 \<le> inverse (setsum u {x \<in> c. 0 < u x}) * u x"
  3503     apply (rule) apply (rule mult_nonneg_nonneg) using * by auto 
  3503     apply (rule) apply (rule mult_nonneg_nonneg) using * by auto
  3504   hence "z \<in> convex hull {v \<in> c. u v > 0}" unfolding convex_hull_explicit mem_Collect_eq
  3504   hence "z \<in> convex hull {v \<in> c. u v > 0}" unfolding convex_hull_explicit mem_Collect_eq
  3505     apply(rule_tac x="{v \<in> c. 0 < u v}" in exI, rule_tac x="\<lambda>y. inverse (setsum u {x\<in>c. u x > 0}) * u y" in exI)
  3505     apply(rule_tac x="{v \<in> c. 0 < u v}" in exI, rule_tac x="\<lambda>y. inverse (setsum u {x\<in>c. u x > 0}) * u y" in exI)
  3506     using assms(1) unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] and z_def using *
  3506     using assms(1) unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] and z_def using *
  3507     by(auto simp add: setsum_negf setsum_right_distrib[symmetric])
  3507     by(auto simp add: setsum_negf setsum_right_distrib[symmetric])
  3508   ultimately show ?thesis apply(rule_tac x="{v\<in>c. u v \<le> 0}" in exI, rule_tac x="{v\<in>c. u v > 0}" in exI) by auto
  3508   ultimately show ?thesis apply(rule_tac x="{v\<in>c. u v \<le> 0}" in exI, rule_tac x="{v\<in>c. u v > 0}" in exI) by auto
  3537       apply(rule, rule X[rule_format]) using X st by auto
  3537       apply(rule, rule X[rule_format]) using X st by auto
  3538   next case True then obtain m p where mp:"m \<inter> p = {}" "m \<union> p = X ` f" "convex hull m \<inter> convex hull p \<noteq> {}"
  3538   next case True then obtain m p where mp:"m \<inter> p = {}" "m \<union> p = X ` f" "convex hull m \<inter> convex hull p \<noteq> {}"
  3539       using radon_partition[of "X ` f"] and affine_dependent_biggerset[of "X ` f"]
  3539       using radon_partition[of "X ` f"] and affine_dependent_biggerset[of "X ` f"]
  3540       unfolding card_image[OF True] and `card f = Suc n` using Suc(3) `finite f` and ng by auto
  3540       unfolding card_image[OF True] and `card f = Suc n` using Suc(3) `finite f` and ng by auto
  3541     have "m \<subseteq> X ` f" "p \<subseteq> X ` f" using mp(2) by auto
  3541     have "m \<subseteq> X ` f" "p \<subseteq> X ` f" using mp(2) by auto
  3542     then obtain g h where gh:"m = X ` g" "p = X ` h" "g \<subseteq> f" "h \<subseteq> f" unfolding subset_image_iff by auto 
  3542     then obtain g h where gh:"m = X ` g" "p = X ` h" "g \<subseteq> f" "h \<subseteq> f" unfolding subset_image_iff by auto
  3543     hence "f \<union> (g \<union> h) = f" by auto
  3543     hence "f \<union> (g \<union> h) = f" by auto
  3544     hence f:"f = g \<union> h" using inj_on_Un_image_eq_iff[of X f "g \<union> h"] and True
  3544     hence f:"f = g \<union> h" using inj_on_Un_image_eq_iff[of X f "g \<union> h"] and True
  3545       unfolding mp(2)[unfolded image_Un[symmetric] gh] by auto
  3545       unfolding mp(2)[unfolded image_Un[symmetric] gh] by auto
  3546     have *:"g \<inter> h = {}" using mp(1) unfolding gh using inj_on_image_Int[OF True gh(3,4)] by auto
  3546     have *:"g \<inter> h = {}" using mp(1) unfolding gh using inj_on_image_Int[OF True gh(3,4)] by auto
  3547     have "convex hull (X ` h) \<subseteq> \<Inter> g" "convex hull (X ` g) \<subseteq> \<Inter> h"
  3547     have "convex hull (X ` h) \<subseteq> \<Inter> g" "convex hull (X ` g) \<subseteq> \<Inter> h"
  3563 
  3563 
  3564 subsection {* Homeomorphism of all convex compact sets with nonempty interior *}
  3564 subsection {* Homeomorphism of all convex compact sets with nonempty interior *}
  3565 
  3565 
  3566 lemma compact_frontier_line_lemma:
  3566 lemma compact_frontier_line_lemma:
  3567   fixes s :: "('a::euclidean_space) set"
  3567   fixes s :: "('a::euclidean_space) set"
  3568   assumes "compact s" "0 \<in> s" "x \<noteq> 0" 
  3568   assumes "compact s" "0 \<in> s" "x \<noteq> 0"
  3569   obtains u where "0 \<le> u" "(u *\<^sub>R x) \<in> frontier s" "\<forall>v>u. (v *\<^sub>R x) \<notin> s"
  3569   obtains u where "0 \<le> u" "(u *\<^sub>R x) \<in> frontier s" "\<forall>v>u. (v *\<^sub>R x) \<notin> s"
  3570 proof-
  3570 proof-
  3571   obtain b where b:"b>0" "\<forall>x\<in>s. norm x \<le> b" using compact_imp_bounded[OF assms(1), unfolded bounded_pos] by auto
  3571   obtain b where b:"b>0" "\<forall>x\<in>s. norm x \<le> b" using compact_imp_bounded[OF assms(1), unfolded bounded_pos] by auto
  3572   let ?A = "{y. \<exists>u. 0 \<le> u \<and> u \<le> b / norm(x) \<and> (y = u *\<^sub>R x)}"
  3572   let ?A = "{y. \<exists>u. 0 \<le> u \<and> u \<le> b / norm(x) \<and> (y = u *\<^sub>R x)}"
  3573   have A:"?A = (\<lambda>u. u *\<^sub>R x) ` {0 .. b / norm x}"
  3573   have A:"?A = (\<lambda>u. u *\<^sub>R x) ` {0 .. b / norm x}"
  3581   ultimately obtain u y where obt: "u\<ge>0" "u \<le> b / norm x" "y = u *\<^sub>R x"
  3581   ultimately obtain u y where obt: "u\<ge>0" "u \<le> b / norm x" "y = u *\<^sub>R x"
  3582     "y\<in>?A" "y\<in>s" "\<forall>z\<in>?A \<inter> s. dist 0 z \<le> dist 0 y" using distance_attains_sup[OF compact_inter[OF _ assms(1), of ?A], of 0] by auto
  3582     "y\<in>?A" "y\<in>s" "\<forall>z\<in>?A \<inter> s. dist 0 z \<le> dist 0 y" using distance_attains_sup[OF compact_inter[OF _ assms(1), of ?A], of 0] by auto
  3583 
  3583 
  3584   have "norm x > 0" using assms(3)[unfolded zero_less_norm_iff[symmetric]] by auto
  3584   have "norm x > 0" using assms(3)[unfolded zero_less_norm_iff[symmetric]] by auto
  3585   { fix v assume as:"v > u" "v *\<^sub>R x \<in> s"
  3585   { fix v assume as:"v > u" "v *\<^sub>R x \<in> s"
  3586     hence "v \<le> b / norm x" using b(2)[rule_format, OF as(2)] 
  3586     hence "v \<le> b / norm x" using b(2)[rule_format, OF as(2)]
  3587       using `u\<ge>0` unfolding pos_le_divide_eq[OF `norm x > 0`] by auto
  3587       using `u\<ge>0` unfolding pos_le_divide_eq[OF `norm x > 0`] by auto
  3588     hence "norm (v *\<^sub>R x) \<le> norm y" apply(rule_tac obt(6)[rule_format, unfolded dist_0_norm]) apply(rule IntI) defer 
  3588     hence "norm (v *\<^sub>R x) \<le> norm y" apply(rule_tac obt(6)[rule_format, unfolded dist_0_norm]) apply(rule IntI) defer
  3589       apply(rule as(2)) unfolding mem_Collect_eq apply(rule_tac x=v in exI) 
  3589       apply(rule as(2)) unfolding mem_Collect_eq apply(rule_tac x=v in exI)
  3590       using as(1) `u\<ge>0` by(auto simp add:field_simps) 
  3590       using as(1) `u\<ge>0` by(auto simp add:field_simps)
  3591     hence False unfolding obt(3) using `u\<ge>0` `norm x > 0` `v>u` by(auto simp add:field_simps)
  3591     hence False unfolding obt(3) using `u\<ge>0` `norm x > 0` `v>u` by(auto simp add:field_simps)
  3592   } note u_max = this
  3592   } note u_max = this
  3593 
  3593 
  3594   have "u *\<^sub>R x \<in> frontier s" unfolding frontier_straddle apply(rule,rule,rule) apply(rule_tac x="u *\<^sub>R x" in bexI) unfolding obt(3)[symmetric]
  3594   have "u *\<^sub>R x \<in> frontier s" unfolding frontier_straddle apply(rule,rule,rule) apply(rule_tac x="u *\<^sub>R x" in bexI) unfolding obt(3)[symmetric]
  3595     prefer 3 apply(rule_tac x="(u + (e / 2) / norm x) *\<^sub>R x" in exI) apply(rule, rule) proof-
  3595     prefer 3 apply(rule_tac x="(u + (e / 2) / norm x) *\<^sub>R x" in exI) apply(rule, rule) proof-
  3633       assume "u\<le>1" thus "u *\<^sub>R x \<in> s" apply(cases "u=1")
  3633       assume "u\<le>1" thus "u *\<^sub>R x \<in> s" apply(cases "u=1")
  3634         using assms(3)[THEN bspec[where x=x], THEN spec[where x=u]] using `0\<le>u` and x and fs by auto qed auto qed
  3634         using assms(3)[THEN bspec[where x=x], THEN spec[where x=u]] using `0\<le>u` and x and fs by auto qed auto qed
  3635 
  3635 
  3636   have "\<exists>surf. homeomorphism (frontier s) sphere pi surf"
  3636   have "\<exists>surf. homeomorphism (frontier s) sphere pi surf"
  3637     apply(rule homeomorphism_compact) apply(rule compact_frontier[OF assms(1)])
  3637     apply(rule homeomorphism_compact) apply(rule compact_frontier[OF assms(1)])
  3638     apply(rule continuous_on_subset[OF contpi]) defer apply(rule set_eqI,rule) 
  3638     apply(rule continuous_on_subset[OF contpi]) defer apply(rule set_eqI,rule)
  3639     unfolding inj_on_def prefer 3 apply(rule,rule,rule)
  3639     unfolding inj_on_def prefer 3 apply(rule,rule,rule)
  3640   proof- fix x assume "x\<in>pi ` frontier s" then obtain y where "y\<in>frontier s" "x = pi y" by auto
  3640   proof- fix x assume "x\<in>pi ` frontier s" then obtain y where "y\<in>frontier s" "x = pi y" by auto
  3641     thus "x \<in> sphere" using pi(1)[of y] and `0 \<notin> frontier s` by auto
  3641     thus "x \<in> sphere" using pi(1)[of y] and `0 \<notin> frontier s` by auto
  3642   next fix x assume "x\<in>sphere" hence "norm x = 1" "x\<noteq>0" unfolding sphere_def by auto
  3642   next fix x assume "x\<in>sphere" hence "norm x = 1" "x\<noteq>0" unfolding sphere_def by auto
  3643     then obtain u where "0 \<le> u" "u *\<^sub>R x \<in> frontier s" "\<forall>v>u. v *\<^sub>R x \<notin> s"
  3643     then obtain u where "0 \<le> u" "u *\<^sub>R x \<in> frontier s" "\<forall>v>u. v *\<^sub>R x \<notin> s"
  3644       using compact_frontier_line_lemma[OF assms(1) `0\<in>s`, of x] by auto
  3644       using compact_frontier_line_lemma[OF assms(1) `0\<in>s`, of x] by auto
  3645     thus "x \<in> pi ` frontier s" unfolding image_iff le_less pi_def apply(rule_tac x="u *\<^sub>R x" in bexI) using `norm x = 1` `0\<notin>frontier s` by auto
  3645     thus "x \<in> pi ` frontier s" unfolding image_iff le_less pi_def apply(rule_tac x="u *\<^sub>R x" in bexI) using `norm x = 1` `0\<notin>frontier s` by auto
  3646   next fix x y assume as:"x \<in> frontier s" "y \<in> frontier s" "pi x = pi y"
  3646   next fix x y assume as:"x \<in> frontier s" "y \<in> frontier s" "pi x = pi y"
  3647     hence xys:"x\<in>s" "y\<in>s" using fs by auto
  3647     hence xys:"x\<in>s" "y\<in>s" using fs by auto
  3648     from as(1,2) have nor:"norm x \<noteq> 0" "norm y \<noteq> 0" using `0\<notin>frontier s` by auto 
  3648     from as(1,2) have nor:"norm x \<noteq> 0" "norm y \<noteq> 0" using `0\<notin>frontier s` by auto
  3649     from nor have x:"x = norm x *\<^sub>R ((inverse (norm y)) *\<^sub>R y)" unfolding as(3)[unfolded pi_def, symmetric] by auto 
  3649     from nor have x:"x = norm x *\<^sub>R ((inverse (norm y)) *\<^sub>R y)" unfolding as(3)[unfolded pi_def, symmetric] by auto
  3650     from nor have y:"y = norm y *\<^sub>R ((inverse (norm x)) *\<^sub>R x)" unfolding as(3)[unfolded pi_def] by auto 
  3650     from nor have y:"y = norm y *\<^sub>R ((inverse (norm x)) *\<^sub>R x)" unfolding as(3)[unfolded pi_def] by auto
  3651     have "0 \<le> norm y * inverse (norm x)" "0 \<le> norm x * inverse (norm y)"
  3651     have "0 \<le> norm y * inverse (norm x)" "0 \<le> norm x * inverse (norm y)"
  3652       unfolding divide_inverse[symmetric] apply(rule_tac[!] divide_nonneg_pos) using nor by auto
  3652       unfolding divide_inverse[symmetric] apply(rule_tac[!] divide_nonneg_pos) using nor by auto
  3653     hence "norm x = norm y" apply(rule_tac ccontr) unfolding neq_iff
  3653     hence "norm x = norm y" apply(rule_tac ccontr) unfolding neq_iff
  3654       using x y and front_smul[THEN bspec, OF as(1), THEN spec[where x="norm y * (inverse (norm x))"]]
  3654       using x y and front_smul[THEN bspec, OF as(1), THEN spec[where x="norm y * (inverse (norm x))"]]
  3655       using front_smul[THEN bspec, OF as(2), THEN spec[where x="norm x * (inverse (norm y))"]]
  3655       using front_smul[THEN bspec, OF as(2), THEN spec[where x="norm x * (inverse (norm y))"]]
  3656       using xys nor by(auto simp add:field_simps divide_le_eq_1 divide_inverse[symmetric])
  3656       using xys nor by(auto simp add:field_simps divide_le_eq_1 divide_inverse[symmetric])
  3657     thus "x = y" apply(subst injpi[symmetric]) using as(3) by auto
  3657     thus "x = y" apply(subst injpi[symmetric]) using as(3) by auto
  3658   qed(insert `0 \<notin> frontier s`, auto)
  3658   qed(insert `0 \<notin> frontier s`, auto)
  3659   then obtain surf where surf:"\<forall>x\<in>frontier s. surf (pi x) = x"  "pi ` frontier s = sphere" "continuous_on (frontier s) pi"
  3659   then obtain surf where surf:"\<forall>x\<in>frontier s. surf (pi x) = x"  "pi ` frontier s = sphere" "continuous_on (frontier s) pi"
  3660     "\<forall>y\<in>sphere. pi (surf y) = y" "surf ` sphere = frontier s" "continuous_on sphere surf" unfolding homeomorphism_def by auto
  3660     "\<forall>y\<in>sphere. pi (surf y) = y" "surf ` sphere = frontier s" "continuous_on sphere surf" unfolding homeomorphism_def by auto
  3661   
  3661 
  3662   have cont_surfpi:"continuous_on (UNIV -  {0}) (surf \<circ> pi)" apply(rule continuous_on_compose, rule contpi)
  3662   have cont_surfpi:"continuous_on (UNIV -  {0}) (surf \<circ> pi)" apply(rule continuous_on_compose, rule contpi)
  3663     apply(rule continuous_on_subset[of sphere], rule surf(6)) using pi(1) by auto
  3663     apply(rule continuous_on_subset[of sphere], rule surf(6)) using pi(1) by auto
  3664 
  3664 
  3665   { fix x assume as:"x \<in> cball (0::'a) 1"
  3665   { fix x assume as:"x \<in> cball (0::'a) 1"
  3666     have "norm x *\<^sub>R surf (pi x) \<in> s" proof(cases "x=0 \<or> norm x = 1") 
  3666     have "norm x *\<^sub>R surf (pi x) \<in> s" proof(cases "x=0 \<or> norm x = 1")
  3667       case False hence "pi x \<in> sphere" "norm x < 1" using pi(1)[of x] as by(auto simp add: dist_norm)
  3667       case False hence "pi x \<in> sphere" "norm x < 1" using pi(1)[of x] as by(auto simp add: dist_norm)
  3668       thus ?thesis apply(rule_tac assms(3)[rule_format, THEN DiffD1])
  3668       thus ?thesis apply(rule_tac assms(3)[rule_format, THEN DiffD1])
  3669         apply(rule_tac fs[unfolded subset_eq, rule_format])
  3669         apply(rule_tac fs[unfolded subset_eq, rule_format])
  3670         unfolding surf(5)[symmetric] by auto
  3670         unfolding surf(5)[symmetric] by auto
  3671     next case True thus ?thesis apply rule defer unfolding pi_def apply(rule fs[unfolded subset_eq, rule_format])
  3671     next case True thus ?thesis apply rule defer unfolding pi_def apply(rule fs[unfolded subset_eq, rule_format])
  3682       hence *:"?a * norm x > 0" and"?a > 0" "?a \<noteq> 0" using surf(5) `0\<notin>frontier s` apply -
  3682       hence *:"?a * norm x > 0" and"?a > 0" "?a \<noteq> 0" using surf(5) `0\<notin>frontier s` apply -
  3683         apply(rule_tac mult_pos_pos) using False[unfolded zero_less_norm_iff[symmetric]] by auto
  3683         apply(rule_tac mult_pos_pos) using False[unfolded zero_less_norm_iff[symmetric]] by auto
  3684       have "norm (surf (pi x)) \<noteq> 0" using ** False by auto
  3684       have "norm (surf (pi x)) \<noteq> 0" using ** False by auto
  3685       hence "norm x = norm ((?a * norm x) *\<^sub>R surf (pi x))"
  3685       hence "norm x = norm ((?a * norm x) *\<^sub>R surf (pi x))"
  3686         unfolding norm_scaleR abs_mult abs_norm_cancel abs_of_pos[OF `?a > 0`] by auto
  3686         unfolding norm_scaleR abs_mult abs_norm_cancel abs_of_pos[OF `?a > 0`] by auto
  3687       moreover have "pi x = pi ((inverse (norm (surf (pi x))) * norm x) *\<^sub>R surf (pi x))" 
  3687       moreover have "pi x = pi ((inverse (norm (surf (pi x))) * norm x) *\<^sub>R surf (pi x))"
  3688         unfolding pi(2)[OF *] surf(4)[rule_format, OF pix] ..
  3688         unfolding pi(2)[OF *] surf(4)[rule_format, OF pix] ..
  3689       moreover have "surf (pi x) \<in> frontier s" using surf(5) pix by auto
  3689       moreover have "surf (pi x) \<in> frontier s" using surf(5) pix by auto
  3690       hence "dist 0 (inverse (norm (surf (pi x))) *\<^sub>R x) \<le> 1" unfolding dist_norm
  3690       hence "dist 0 (inverse (norm (surf (pi x))) *\<^sub>R x) \<le> 1" unfolding dist_norm
  3691         using ** and * using front_smul[THEN bspec[where x="surf (pi x)"], THEN spec[where x="norm x * ?a"]]
  3691         using ** and * using front_smul[THEN bspec[where x="surf (pi x)"], THEN spec[where x="norm x * ?a"]]
  3692         using False `x\<in>s` by(auto simp add:field_simps)
  3692         using False `x\<in>s` by(auto simp add:field_simps)
  3701     fix x::"'a" assume as:"x \<in> cball 0 1"
  3701     fix x::"'a" assume as:"x \<in> cball 0 1"
  3702     thus "continuous (at x) (\<lambda>x. norm x *\<^sub>R surf (pi x))" proof(cases "x=0")
  3702     thus "continuous (at x) (\<lambda>x. norm x *\<^sub>R surf (pi x))" proof(cases "x=0")
  3703       case False thus ?thesis apply (intro continuous_intros)
  3703       case False thus ?thesis apply (intro continuous_intros)
  3704         using cont_surfpi unfolding continuous_on_eq_continuous_at[OF open_delete[OF open_UNIV]] o_def by auto
  3704         using cont_surfpi unfolding continuous_on_eq_continuous_at[OF open_delete[OF open_UNIV]] o_def by auto
  3705     next obtain B where B:"\<forall>x\<in>s. norm x \<le> B" using compact_imp_bounded[OF assms(1)] unfolding bounded_iff by auto
  3705     next obtain B where B:"\<forall>x\<in>s. norm x \<le> B" using compact_imp_bounded[OF assms(1)] unfolding bounded_iff by auto
  3706       hence "B > 0" using assms(2) unfolding subset_eq apply(erule_tac x="basis 0" in ballE) defer 
  3706       hence "B > 0" using assms(2) unfolding subset_eq apply(erule_tac x="basis 0" in ballE) defer
  3707         apply(erule_tac x="basis 0" in ballE)
  3707         apply(erule_tac x="basis 0" in ballE)
  3708         unfolding Ball_def mem_cball dist_norm using DIM_positive[where 'a='a]
  3708         unfolding Ball_def mem_cball dist_norm using DIM_positive[where 'a='a]
  3709         by auto
  3709         by auto
  3710       case True show ?thesis unfolding True continuous_at Lim_at apply(rule,rule) apply(rule_tac x="e / B" in exI)
  3710       case True show ?thesis unfolding True continuous_at Lim_at apply(rule,rule) apply(rule_tac x="e / B" in exI)
  3711         apply(rule) apply(rule divide_pos_pos) prefer 3 apply(rule,rule,erule conjE)
  3711         apply(rule) apply(rule divide_pos_pos) prefer 3 apply(rule,rule,erule conjE)
  3724         hence "surf (pi x) \<in> frontier s" using surf(5) by auto
  3724         hence "surf (pi x) \<in> frontier s" using surf(5) by auto
  3725         thus False using `0\<notin>frontier s` unfolding as by simp qed
  3725         thus False using `0\<notin>frontier s` unfolding as by simp qed
  3726     } note surf_0 = this
  3726     } note surf_0 = this
  3727     show "inj_on (\<lambda>x. norm x *\<^sub>R surf (pi x)) (cball 0 1)" unfolding inj_on_def proof(rule,rule,rule)
  3727     show "inj_on (\<lambda>x. norm x *\<^sub>R surf (pi x)) (cball 0 1)" unfolding inj_on_def proof(rule,rule,rule)
  3728       fix x y assume as:"x \<in> cball 0 1" "y \<in> cball 0 1" "norm x *\<^sub>R surf (pi x) = norm y *\<^sub>R surf (pi y)"
  3728       fix x y assume as:"x \<in> cball 0 1" "y \<in> cball 0 1" "norm x *\<^sub>R surf (pi x) = norm y *\<^sub>R surf (pi y)"
  3729       thus "x=y" proof(cases "x=0 \<or> y=0") 
  3729       thus "x=y" proof(cases "x=0 \<or> y=0")
  3730         case True thus ?thesis using as by(auto elim: surf_0) next
  3730         case True thus ?thesis using as by(auto elim: surf_0) next
  3731         case False
  3731         case False
  3732         hence "pi (surf (pi x)) = pi (surf (pi y))" using as(3)
  3732         hence "pi (surf (pi x)) = pi (surf (pi y))" using as(3)
  3733           using pi(2)[of "norm x" "surf (pi x)"] pi(2)[of "norm y" "surf (pi y)"] by auto
  3733           using pi(2)[of "norm x" "surf (pi x)"] pi(2)[of "norm y" "surf (pi y)"] by auto
  3734         moreover have "pi x \<in> sphere" "pi y \<in> sphere" using pi(1) False by auto
  3734         moreover have "pi x \<in> sphere" "pi y \<in> sphere" using pi(1) False by auto
  3735         ultimately have *:"pi x = pi y" using surf(4)[THEN bspec[where x="pi x"]] surf(4)[THEN bspec[where x="pi y"]] by auto 
  3735         ultimately have *:"pi x = pi y" using surf(4)[THEN bspec[where x="pi x"]] surf(4)[THEN bspec[where x="pi y"]] by auto
  3736         moreover have "norm x = norm y" using as(3)[unfolded *] using False by(auto dest:surf_0)
  3736         moreover have "norm x = norm y" using as(3)[unfolded *] using False by(auto dest:surf_0)
  3737         ultimately show ?thesis using injpi by auto qed qed
  3737         ultimately show ?thesis using injpi by auto qed qed
  3738   qed auto qed
  3738   qed auto qed
  3739 
  3739 
  3740 lemma homeomorphic_convex_compact_lemma:
  3740 lemma homeomorphic_convex_compact_lemma:
  3788 definition "epigraph s (f::_ \<Rightarrow> real) = {xy. fst xy \<in> s \<and> f (fst xy) \<le> snd xy}"
  3788 definition "epigraph s (f::_ \<Rightarrow> real) = {xy. fst xy \<in> s \<and> f (fst xy) \<le> snd xy}"
  3789 
  3789 
  3790 lemma mem_epigraph: "(x, y) \<in> epigraph s f \<longleftrightarrow> x \<in> s \<and> f x \<le> y" unfolding epigraph_def by auto
  3790 lemma mem_epigraph: "(x, y) \<in> epigraph s f \<longleftrightarrow> x \<in> s \<and> f x \<le> y" unfolding epigraph_def by auto
  3791 
  3791 
  3792 (** This might break sooner or later. In fact it did already once. **)
  3792 (** This might break sooner or later. In fact it did already once. **)
  3793 lemma convex_epigraph: 
  3793 lemma convex_epigraph:
  3794   "convex(epigraph s f) \<longleftrightarrow> convex_on s f \<and> convex s"
  3794   "convex(epigraph s f) \<longleftrightarrow> convex_on s f \<and> convex s"
  3795   unfolding convex_def convex_on_def
  3795   unfolding convex_def convex_on_def
  3796   unfolding Ball_def split_paired_All epigraph_def
  3796   unfolding Ball_def split_paired_All epigraph_def
  3797   unfolding mem_Collect_eq fst_conv snd_conv fst_add snd_add fst_scaleR snd_scaleR Ball_def[symmetric]
  3797   unfolding mem_Collect_eq fst_conv snd_conv fst_add snd_add fst_scaleR snd_scaleR Ball_def[symmetric]
  3798   apply safe defer apply(erule_tac x=x in allE,erule_tac x="f x" in allE) apply safe
  3798   apply safe defer apply(erule_tac x=x in allE,erule_tac x="f x" in allE) apply safe
  3875   { fix y assume "y \<in> s" have "y \<in> ?halfr \<union> ?halfl" apply(rule ccontr)
  3875   { fix y assume "y \<in> s" have "y \<in> ?halfr \<union> ?halfl" apply(rule ccontr)
  3876     using as(6) `y\<in>s` by (auto simp add: inner_vector_def) }
  3876     using as(6) `y\<in>s` by (auto simp add: inner_vector_def) }
  3877   moreover have "a\<in>?halfl" "b\<in>?halfr" using * by (auto simp add: inner_vector_def)
  3877   moreover have "a\<in>?halfl" "b\<in>?halfr" using * by (auto simp add: inner_vector_def)
  3878   hence "?halfl \<inter> s \<noteq> {}" "?halfr \<inter> s \<noteq> {}"  using as(2-3) by auto
  3878   hence "?halfl \<inter> s \<noteq> {}" "?halfr \<inter> s \<noteq> {}"  using as(2-3) by auto
  3879   ultimately show False apply(rule_tac notE[OF as(1)[unfolded connected_def]])
  3879   ultimately show False apply(rule_tac notE[OF as(1)[unfolded connected_def]])
  3880     apply(rule_tac x="?halfl" in exI, rule_tac x="?halfr" in exI) 
  3880     apply(rule_tac x="?halfl" in exI, rule_tac x="?halfr" in exI)
  3881     apply(rule, rule open_halfspace_lt, rule, rule open_halfspace_gt)
  3881     apply(rule, rule open_halfspace_lt, rule, rule open_halfspace_gt)
  3882     by(auto simp add: field_simps) qed
  3882     by(auto simp add: field_simps) qed
  3883 
  3883 
  3884 lemma is_interval_convex_1:
  3884 lemma is_interval_convex_1:
  3885   "is_interval s \<longleftrightarrow> convex (s::(real^1) set)" 
  3885   "is_interval s \<longleftrightarrow> convex (s::(real^1) set)"
  3886 by(metis is_interval_convex convex_connected is_interval_connected_1)
  3886 by(metis is_interval_convex convex_connected is_interval_connected_1)
  3887 
  3887 
  3888 lemma convex_connected_1:
  3888 lemma convex_connected_1:
  3889   "connected s \<longleftrightarrow> convex (s::(real^1) set)" 
  3889   "connected s \<longleftrightarrow> convex (s::(real^1) set)"
  3890 by(metis is_interval_convex convex_connected is_interval_connected_1)
  3890 by(metis is_interval_convex convex_connected is_interval_connected_1)
  3891 *)
  3891 *)
  3892 subsection {* Another intermediate value theorem formulation *}
  3892 subsection {* Another intermediate value theorem formulation *}
  3893 
  3893 
  3894 lemma ivt_increasing_component_on_1: fixes f::"real \<Rightarrow> 'a::euclidean_space"
  3894 lemma ivt_increasing_component_on_1: fixes f::"real \<Rightarrow> 'a::euclidean_space"
  3895   assumes "a \<le> b" "continuous_on {a .. b} f" "(f a)$$k \<le> y" "y \<le> (f b)$$k"
  3895   assumes "a \<le> b" "continuous_on {a .. b} f" "(f a)$$k \<le> y" "y \<le> (f b)$$k"
  3896   shows "\<exists>x\<in>{a..b}. (f x)$$k = y"
  3896   shows "\<exists>x\<in>{a..b}. (f x)$$k = y"
  3897 proof- have "f a \<in> f ` {a..b}" "f b \<in> f ` {a..b}" apply(rule_tac[!] imageI) 
  3897 proof- have "f a \<in> f ` {a..b}" "f b \<in> f ` {a..b}" apply(rule_tac[!] imageI)
  3898     using assms(1) by auto
  3898     using assms(1) by auto
  3899   thus ?thesis using connected_ivt_component[of "f ` {a..b}" "f a" "f b" k y]
  3899   thus ?thesis using connected_ivt_component[of "f ` {a..b}" "f a" "f b" k y]
  3900     using connected_continuous_image[OF assms(2) convex_connected[OF convex_real_interval(5)]]
  3900     using connected_continuous_image[OF assms(2) convex_connected[OF convex_real_interval(5)]]
  3901     using assms by(auto intro!: imageI) qed
  3901     using assms by(auto intro!: imageI) qed
  3902 
  3902 
  3935 
  3935 
  3936 lemma unit_interval_convex_hull:
  3936 lemma unit_interval_convex_hull:
  3937   "{0::'a::ordered_euclidean_space .. (\<chi>\<chi> i. 1)} = convex hull {x. \<forall>i<DIM('a). (x$$i = 0) \<or> (x$$i = 1)}"
  3937   "{0::'a::ordered_euclidean_space .. (\<chi>\<chi> i. 1)} = convex hull {x. \<forall>i<DIM('a). (x$$i = 0) \<or> (x$$i = 1)}"
  3938   (is "?int = convex hull ?points")
  3938   (is "?int = convex hull ?points")
  3939 proof- have 01:"{0,(\<chi>\<chi> i. 1)} \<subseteq> convex hull ?points" apply rule apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) by auto
  3939 proof- have 01:"{0,(\<chi>\<chi> i. 1)} \<subseteq> convex hull ?points" apply rule apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) by auto
  3940   { fix n x assume "x\<in>{0::'a::ordered_euclidean_space .. \<chi>\<chi> i. 1}" "n \<le> DIM('a)" "card {i. i<DIM('a) \<and> x$$i \<noteq> 0} \<le> n" 
  3940   { fix n x assume "x\<in>{0::'a::ordered_euclidean_space .. \<chi>\<chi> i. 1}" "n \<le> DIM('a)" "card {i. i<DIM('a) \<and> x$$i \<noteq> 0} \<le> n"
  3941   hence "x\<in>convex hull ?points" proof(induct n arbitrary: x)
  3941   hence "x\<in>convex hull ?points" proof(induct n arbitrary: x)
  3942     case 0 hence "x = 0" apply(subst euclidean_eq) apply rule by auto
  3942     case 0 hence "x = 0" apply(subst euclidean_eq) apply rule by auto
  3943     thus "x\<in>convex hull ?points" using 01 by auto
  3943     thus "x\<in>convex hull ?points" using 01 by auto
  3944   next
  3944   next
  3945     case (Suc n) show "x\<in>convex hull ?points" proof(cases "{i. i<DIM('a) \<and> x$$i \<noteq> 0} = {}")
  3945     case (Suc n) show "x\<in>convex hull ?points" proof(cases "{i. i<DIM('a) \<and> x$$i \<noteq> 0} = {}")
  3956         using i'(2-) `x$$i \<noteq> 0` by auto
  3956         using i'(2-) `x$$i \<noteq> 0` by auto
  3957       show ?thesis proof(cases "x$$i=1")
  3957       show ?thesis proof(cases "x$$i=1")
  3958         case True have "\<forall>j\<in>{i. i<DIM('a) \<and> x$$i \<noteq> 0}. x$$j = 1" apply(rule, rule ccontr) unfolding mem_Collect_eq
  3958         case True have "\<forall>j\<in>{i. i<DIM('a) \<and> x$$i \<noteq> 0}. x$$j = 1" apply(rule, rule ccontr) unfolding mem_Collect_eq
  3959         proof(erule conjE) fix j assume as:"x $$ j \<noteq> 0" "x $$ j \<noteq> 1" "j<DIM('a)"
  3959         proof(erule conjE) fix j assume as:"x $$ j \<noteq> 0" "x $$ j \<noteq> 1" "j<DIM('a)"
  3960           hence j:"x$$j \<in> {0<..<1}" using Suc(2) by(auto simp add: eucl_le[where 'a='a] elim!:allE[where x=j])
  3960           hence j:"x$$j \<in> {0<..<1}" using Suc(2) by(auto simp add: eucl_le[where 'a='a] elim!:allE[where x=j])
  3961           hence "x$$j \<in> op $$ x ` {i. i<DIM('a) \<and> x $$ i \<noteq> 0}" using as(3) by auto 
  3961           hence "x$$j \<in> op $$ x ` {i. i<DIM('a) \<and> x $$ i \<noteq> 0}" using as(3) by auto
  3962           hence "x$$j \<ge> x$$i" unfolding i'(1) xi_def apply(rule_tac Min_le) by auto
  3962           hence "x$$j \<ge> x$$i" unfolding i'(1) xi_def apply(rule_tac Min_le) by auto
  3963           thus False using True Suc(2) j by(auto simp add: elim!:ballE[where x=j]) qed
  3963           thus False using True Suc(2) j by(auto simp add: elim!:ballE[where x=j]) qed
  3964         thus "x\<in>convex hull ?points" apply(rule_tac hull_subset[unfolded subset_eq, rule_format])
  3964         thus "x\<in>convex hull ?points" apply(rule_tac hull_subset[unfolded subset_eq, rule_format])
  3965           by auto
  3965           by auto
  3966       next let ?y = "\<lambda>j. if x$$j = 0 then 0 else (x$$j - x$$i) / (1 - x$$i)"
  3966       next let ?y = "\<lambda>j. if x$$j = 0 then 0 else (x$$j - x$$i) / (1 - x$$i)"
  3973             by(auto simp add:field_simps)
  3973             by(auto simp add:field_simps)
  3974           hence "0 \<le> ?y j \<and> ?y j \<le> 1" by auto }
  3974           hence "0 \<le> ?y j \<and> ?y j \<le> 1" by auto }
  3975         moreover have "i\<in>{j. j<DIM('a) \<and> x$$j \<noteq> 0} - {j. j<DIM('a) \<and> ((\<chi>\<chi> j. ?y j)::'a) $$ j \<noteq> 0}"
  3975         moreover have "i\<in>{j. j<DIM('a) \<and> x$$j \<noteq> 0} - {j. j<DIM('a) \<and> ((\<chi>\<chi> j. ?y j)::'a) $$ j \<noteq> 0}"
  3976           using i01 using i'(3) by auto
  3976           using i01 using i'(3) by auto
  3977         hence "{j. j<DIM('a) \<and> x$$j \<noteq> 0} \<noteq> {j. j<DIM('a) \<and> ((\<chi>\<chi> j. ?y j)::'a) $$ j \<noteq> 0}" using i'(3) by blast
  3977         hence "{j. j<DIM('a) \<and> x$$j \<noteq> 0} \<noteq> {j. j<DIM('a) \<and> ((\<chi>\<chi> j. ?y j)::'a) $$ j \<noteq> 0}" using i'(3) by blast
  3978         hence **:"{j. j<DIM('a) \<and> ((\<chi>\<chi> j. ?y j)::'a) $$ j \<noteq> 0} \<subset> {j. j<DIM('a) \<and> x$$j \<noteq> 0}" apply - apply rule 
  3978         hence **:"{j. j<DIM('a) \<and> ((\<chi>\<chi> j. ?y j)::'a) $$ j \<noteq> 0} \<subset> {j. j<DIM('a) \<and> x$$j \<noteq> 0}" apply - apply rule
  3979           by auto
  3979           by auto
  3980         have "card {j. j<DIM('a) \<and> ((\<chi>\<chi> j. ?y j)::'a) $$ j \<noteq> 0} \<le> n"
  3980         have "card {j. j<DIM('a) \<and> ((\<chi>\<chi> j. ?y j)::'a) $$ j \<noteq> 0} \<le> n"
  3981           using less_le_trans[OF psubset_card_mono[OF _ **] Suc(4)] by auto
  3981           using less_le_trans[OF psubset_card_mono[OF _ **] Suc(4)] by auto
  3982         ultimately show ?thesis apply(subst *) apply(rule convex_convex_hull[unfolded convex_def, rule_format])
  3982         ultimately show ?thesis apply(subst *) apply(rule convex_convex_hull[unfolded convex_def, rule_format])
  3983           apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) defer apply(rule Suc(1))
  3983           apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) defer apply(rule Suc(1))
  3984           unfolding mem_interval using i01 Suc(3) by auto
  3984           unfolding mem_interval using i01 Suc(3) by auto
  3985       qed qed qed } note * = this
  3985       qed qed qed } note * = this
  3986   have **:"DIM('a) = card {..<DIM('a)}" by auto
  3986   have **:"DIM('a) = card {..<DIM('a)}" by auto
  3987   show ?thesis apply rule defer apply(rule hull_minimal) unfolding subset_eq prefer 3 apply rule 
  3987   show ?thesis apply rule defer apply(rule hull_minimal) unfolding subset_eq prefer 3 apply rule
  3988     apply(rule_tac n2="DIM('a)" in *) prefer 3 apply(subst(2) **) 
  3988     apply(rule_tac n2="DIM('a)" in *) prefer 3 apply(subst(2) **)
  3989     apply(rule card_mono) using 01 and convex_interval(1) prefer 5 apply - apply rule
  3989     apply(rule card_mono) using 01 and convex_interval(1) prefer 5 apply - apply rule
  3990     unfolding mem_interval apply rule unfolding mem_Collect_eq apply(erule_tac x=i in allE)
  3990     unfolding mem_interval apply rule unfolding mem_Collect_eq apply(erule_tac x=i in allE)
  3991     by auto qed
  3991     by auto qed
  3992 
  3992 
  3993 text {* And this is a finite set of vertices. *}
  3993 text {* And this is a finite set of vertices. *}
  4018         using assms by(auto simp add: field_simps)
  4018         using assms by(auto simp add: field_simps)
  4019       hence "inverse d * (x $$ i * 2) \<le> 2 + inverse d * (y $$ i * 2)"
  4019       hence "inverse d * (x $$ i * 2) \<le> 2 + inverse d * (y $$ i * 2)"
  4020             "inverse d * (y $$ i * 2) \<le> 2 + inverse d * (x $$ i * 2)" by(auto simp add:field_simps) }
  4020             "inverse d * (y $$ i * 2) \<le> 2 + inverse d * (x $$ i * 2)" by(auto simp add:field_simps) }
  4021     hence "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \<in> {0..\<chi>\<chi> i.1}" unfolding mem_interval using assms
  4021     hence "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \<in> {0..\<chi>\<chi> i.1}" unfolding mem_interval using assms
  4022       by(auto simp add: field_simps)
  4022       by(auto simp add: field_simps)
  4023     thus "\<exists>z\<in>{0..\<chi>\<chi> i.1}. y = x - ?d + (2 * d) *\<^sub>R z" apply- apply(rule_tac x="inverse (2 * d) *\<^sub>R (y - (x - ?d))" in bexI) 
  4023     thus "\<exists>z\<in>{0..\<chi>\<chi> i.1}. y = x - ?d + (2 * d) *\<^sub>R z" apply- apply(rule_tac x="inverse (2 * d) *\<^sub>R (y - (x - ?d))" in bexI)
  4024       using assms by auto
  4024       using assms by auto
  4025   next
  4025   next
  4026     fix y z assume as:"z\<in>{0..\<chi>\<chi> i.1}" "y = x - ?d + (2*d) *\<^sub>R z" 
  4026     fix y z assume as:"z\<in>{0..\<chi>\<chi> i.1}" "y = x - ?d + (2*d) *\<^sub>R z"
  4027     have "\<And>i. i<DIM('a) \<Longrightarrow> 0 \<le> d * z $$ i \<and> d * z $$ i \<le> d"
  4027     have "\<And>i. i<DIM('a) \<Longrightarrow> 0 \<le> d * z $$ i \<and> d * z $$ i \<le> d"
  4028       using assms as(1)[unfolded mem_interval] apply(erule_tac x=i in allE)
  4028       using assms as(1)[unfolded mem_interval] apply(erule_tac x=i in allE)
  4029       apply rule apply(rule mult_nonneg_nonneg) prefer 3 apply(rule mult_right_le_one_le)
  4029       apply rule apply(rule mult_nonneg_nonneg) prefer 3 apply(rule mult_right_le_one_le)
  4030       using assms by auto
  4030       using assms by auto
  4031     thus "y \<in> {x - ?d..x + ?d}" unfolding as(2) mem_interval apply- apply rule using as(1)[unfolded mem_interval]
  4031     thus "y \<in> {x - ?d..x + ?d}" unfolding as(2) mem_interval apply- apply rule using as(1)[unfolded mem_interval]
  4045   have B:"0 < B" "\<And>x. x\<in>s \<Longrightarrow> abs (f x) \<le> B"
  4045   have B:"0 < B" "\<And>x. x\<in>s \<Longrightarrow> abs (f x) \<le> B"
  4046     unfolding B_def defer apply(drule assms(3)[rule_format]) by auto
  4046     unfolding B_def defer apply(drule assms(3)[rule_format]) by auto
  4047   obtain k where "k>0"and k:"cball x k \<subseteq> s" using assms(1)[unfolded open_contains_cball, THEN bspec[where x=x]] using `x\<in>s` by auto
  4047   obtain k where "k>0"and k:"cball x k \<subseteq> s" using assms(1)[unfolded open_contains_cball, THEN bspec[where x=x]] using `x\<in>s` by auto
  4048   show "\<exists>d>0. \<forall>x'. norm (x' - x) < d \<longrightarrow> \<bar>f x' - f x\<bar> < e"
  4048   show "\<exists>d>0. \<forall>x'. norm (x' - x) < d \<longrightarrow> \<bar>f x' - f x\<bar> < e"
  4049     apply(rule_tac x="min (k / 2) (e / (2 * B) * k)" in exI) apply rule defer proof(rule,rule)
  4049     apply(rule_tac x="min (k / 2) (e / (2 * B) * k)" in exI) apply rule defer proof(rule,rule)
  4050     fix y assume as:"norm (y - x) < min (k / 2) (e / (2 * B) * k)" 
  4050     fix y assume as:"norm (y - x) < min (k / 2) (e / (2 * B) * k)"
  4051     show "\<bar>f y - f x\<bar> < e" proof(cases "y=x")
  4051     show "\<bar>f y - f x\<bar> < e" proof(cases "y=x")
  4052       case False def t \<equiv> "k / norm (y - x)"
  4052       case False def t \<equiv> "k / norm (y - x)"
  4053       have "2 < t" "0<t" unfolding t_def using as False and `k>0` by(auto simp add:field_simps)
  4053       have "2 < t" "0<t" unfolding t_def using as False and `k>0` by(auto simp add:field_simps)
  4054       have "y\<in>s" apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_norm
  4054       have "y\<in>s" apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_norm
  4055         apply(rule order_trans[of _ "2 * norm (x - y)"]) using as by(auto simp add: field_simps norm_minus_commute) 
  4055         apply(rule order_trans[of _ "2 * norm (x - y)"]) using as by(auto simp add: field_simps norm_minus_commute)
  4056       { def w \<equiv> "x + t *\<^sub>R (y - x)"
  4056       { def w \<equiv> "x + t *\<^sub>R (y - x)"
  4057         have "w\<in>s" unfolding w_def apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_norm 
  4057         have "w\<in>s" unfolding w_def apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_norm
  4058           unfolding t_def using `k>0` by auto
  4058           unfolding t_def using `k>0` by auto
  4059         have "(1 / t) *\<^sub>R x + - x + ((t - 1) / t) *\<^sub>R x = (1 / t - 1 + (t - 1) / t) *\<^sub>R x" by (auto simp add: algebra_simps)
  4059         have "(1 / t) *\<^sub>R x + - x + ((t - 1) / t) *\<^sub>R x = (1 / t - 1 + (t - 1) / t) *\<^sub>R x" by (auto simp add: algebra_simps)
  4060         also have "\<dots> = 0"  using `t>0` by(auto simp add:field_simps)
  4060         also have "\<dots> = 0"  using `t>0` by(auto simp add:field_simps)
  4061         finally have w:"(1 / t) *\<^sub>R w + ((t - 1) / t) *\<^sub>R x = y" unfolding w_def using False and `t>0` by (auto simp add: algebra_simps)
  4061         finally have w:"(1 / t) *\<^sub>R w + ((t - 1) / t) *\<^sub>R x = y" unfolding w_def using False and `t>0` by (auto simp add: algebra_simps)
  4062         have  "2 * B < e * t" unfolding t_def using `0<e` `0<k` `B>0` and as and False by (auto simp add:field_simps) 
  4062         have  "2 * B < e * t" unfolding t_def using `0<e` `0<k` `B>0` and as and False by (auto simp add:field_simps)
  4063         hence "(f w - f x) / t < e"
  4063         hence "(f w - f x) / t < e"
  4064           using B(2)[OF `w\<in>s`] and B(2)[OF `x\<in>s`] using `t>0` by(auto simp add:field_simps) 
  4064           using B(2)[OF `w\<in>s`] and B(2)[OF `x\<in>s`] using `t>0` by(auto simp add:field_simps)
  4065         hence th1:"f y - f x < e" apply- apply(rule le_less_trans) defer apply assumption
  4065         hence th1:"f y - f x < e" apply- apply(rule le_less_trans) defer apply assumption
  4066           using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w]
  4066           using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w]
  4067           using `0<t` `2<t` and `x\<in>s` `w\<in>s` by(auto simp add:field_simps) }
  4067           using `0<t` `2<t` and `x\<in>s` `w\<in>s` by(auto simp add:field_simps) }
  4068       moreover 
  4068       moreover
  4069       { def w \<equiv> "x - t *\<^sub>R (y - x)"
  4069       { def w \<equiv> "x - t *\<^sub>R (y - x)"
  4070         have "w\<in>s" unfolding w_def apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_norm 
  4070         have "w\<in>s" unfolding w_def apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_norm
  4071           unfolding t_def using `k>0` by auto
  4071           unfolding t_def using `k>0` by auto
  4072         have "(1 / (1 + t)) *\<^sub>R x + (t / (1 + t)) *\<^sub>R x = (1 / (1 + t) + t / (1 + t)) *\<^sub>R x" by (auto simp add: algebra_simps)
  4072         have "(1 / (1 + t)) *\<^sub>R x + (t / (1 + t)) *\<^sub>R x = (1 / (1 + t) + t / (1 + t)) *\<^sub>R x" by (auto simp add: algebra_simps)
  4073         also have "\<dots>=x" using `t>0` by (auto simp add:field_simps)
  4073         also have "\<dots>=x" using `t>0` by (auto simp add:field_simps)
  4074         finally have w:"(1 / (1+t)) *\<^sub>R w + (t / (1 + t)) *\<^sub>R y = x" unfolding w_def using False and `t>0` by (auto simp add: algebra_simps)
  4074         finally have w:"(1 / (1+t)) *\<^sub>R w + (t / (1 + t)) *\<^sub>R y = x" unfolding w_def using False and `t>0` by (auto simp add: algebra_simps)
  4075         have  "2 * B < e * t" unfolding t_def using `0<e` `0<k` `B>0` and as and False by (auto simp add:field_simps) 
  4075         have  "2 * B < e * t" unfolding t_def using `0<e` `0<k` `B>0` and as and False by (auto simp add:field_simps)
  4076         hence *:"(f w - f y) / t < e" using B(2)[OF `w\<in>s`] and B(2)[OF `y\<in>s`] using `t>0` by(auto simp add:field_simps) 
  4076         hence *:"(f w - f y) / t < e" using B(2)[OF `w\<in>s`] and B(2)[OF `y\<in>s`] using `t>0` by(auto simp add:field_simps)
  4077         have "f x \<le> 1 / (1 + t) * f w + (t / (1 + t)) * f y" 
  4077         have "f x \<le> 1 / (1 + t) * f w + (t / (1 + t)) * f y"
  4078           using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w]
  4078           using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w]
  4079           using `0<t` `2<t` and `y\<in>s` `w\<in>s` by (auto simp add:field_simps)
  4079           using `0<t` `2<t` and `y\<in>s` `w\<in>s` by (auto simp add:field_simps)
  4080         also have "\<dots> = (f w + t * f y) / (1 + t)" using `t>0` unfolding divide_inverse by (auto simp add:field_simps)
  4080         also have "\<dots> = (f w + t * f y) / (1 + t)" using `t>0` unfolding divide_inverse by (auto simp add:field_simps)
  4081         also have "\<dots> < e + f y" using `t>0` * `e>0` by(auto simp add:field_simps)
  4081         also have "\<dots> < e + f y" using `t>0` * `e>0` by(auto simp add:field_simps)
  4082         finally have "f x - f y < e" by auto }
  4082         finally have "f x - f y < e" by auto }
  4083       ultimately show ?thesis by auto 
  4083       ultimately show ?thesis by auto
  4084     qed(insert `0<e`, auto) 
  4084     qed(insert `0<e`, auto)
  4085   qed(insert `0<e` `0<k` `0<B`, auto simp add:field_simps intro!:mult_pos_pos) qed
  4085   qed(insert `0<e` `0<k` `0<B`, auto simp add:field_simps intro!:mult_pos_pos) qed
  4086 
  4086 
  4087 subsection {* Upper bound on a ball implies upper and lower bounds *}
  4087 subsection {* Upper bound on a ball implies upper and lower bounds *}
  4088 
  4088 
  4089 lemma convex_bounds_lemma:
  4089 lemma convex_bounds_lemma:
  4095   have *:"x - (2 *\<^sub>R x - y) = y - x" by (simp add: scaleR_2)
  4095   have *:"x - (2 *\<^sub>R x - y) = y - x" by (simp add: scaleR_2)
  4096   have z:"z\<in>cball x e" using y unfolding z_def mem_cball dist_norm * by(auto simp add: norm_minus_commute)
  4096   have z:"z\<in>cball x e" using y unfolding z_def mem_cball dist_norm * by(auto simp add: norm_minus_commute)
  4097   have "(1 / 2) *\<^sub>R y + (1 / 2) *\<^sub>R z = x" unfolding z_def by (auto simp add: algebra_simps)
  4097   have "(1 / 2) *\<^sub>R y + (1 / 2) *\<^sub>R z = x" unfolding z_def by (auto simp add: algebra_simps)
  4098   thus "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" using assms(1)[unfolded convex_on_def,rule_format, OF y z, of "1/2" "1/2"]
  4098   thus "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" using assms(1)[unfolded convex_on_def,rule_format, OF y z, of "1/2" "1/2"]
  4099     using assms(2)[rule_format,OF y] assms(2)[rule_format,OF z] by(auto simp add:field_simps)
  4099     using assms(2)[rule_format,OF y] assms(2)[rule_format,OF z] by(auto simp add:field_simps)
  4100 next case False fix y assume "y\<in>cball x e" 
  4100 next case False fix y assume "y\<in>cball x e"
  4101   hence "dist x y < 0" using False unfolding mem_cball not_le by (auto simp del: dist_not_less_zero)
  4101   hence "dist x y < 0" using False unfolding mem_cball not_le by (auto simp del: dist_not_less_zero)
  4102   thus "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" using zero_le_dist[of x y] by auto qed
  4102   thus "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" using zero_le_dist[of x y] by auto qed
  4103 
  4103 
  4104 subsubsection {* Hence a convex function on an open set is continuous *}
  4104 subsubsection {* Hence a convex function on an open set is continuous *}
  4105 
  4105 
  4106 lemma convex_on_continuous:
  4106 lemma convex_on_continuous:
  4107   assumes "open (s::('a::ordered_euclidean_space) set)" "convex_on s f" 
  4107   assumes "open (s::('a::ordered_euclidean_space) set)" "convex_on s f"
  4108   (* FIXME: generalize to euclidean_space *)
  4108   (* FIXME: generalize to euclidean_space *)
  4109   shows "continuous_on s f"
  4109   shows "continuous_on s f"
  4110   unfolding continuous_on_eq_continuous_at[OF assms(1)] proof
  4110   unfolding continuous_on_eq_continuous_at[OF assms(1)] proof
  4111   note dimge1 = DIM_positive[where 'a='a]
  4111   note dimge1 = DIM_positive[where 'a='a]
  4112   fix x assume "x\<in>s"
  4112   fix x assume "x\<in>s"
  4113   then obtain e where e:"cball x e \<subseteq> s" "e>0" using assms(1) unfolding open_contains_cball by auto
  4113   then obtain e where e:"cball x e \<subseteq> s" "e>0" using assms(1) unfolding open_contains_cball by auto
  4114   def d \<equiv> "e / real DIM('a)"
  4114   def d \<equiv> "e / real DIM('a)"
  4115   have "0 < d" unfolding d_def using `e>0` dimge1 by(rule_tac divide_pos_pos, auto) 
  4115   have "0 < d" unfolding d_def using `e>0` dimge1 by(rule_tac divide_pos_pos, auto)
  4116   let ?d = "(\<chi>\<chi> i. d)::'a"
  4116   let ?d = "(\<chi>\<chi> i. d)::'a"
  4117   obtain c where c:"finite c" "{x - ?d..x + ?d} = convex hull c" using cube_convex_hull[OF `d>0`, of x] by auto
  4117   obtain c where c:"finite c" "{x - ?d..x + ?d} = convex hull c" using cube_convex_hull[OF `d>0`, of x] by auto
  4118   have "x\<in>{x - ?d..x + ?d}" using `d>0` unfolding mem_interval by auto
  4118   have "x\<in>{x - ?d..x + ?d}" using `d>0` unfolding mem_interval by auto
  4119   hence "c\<noteq>{}" using c by auto
  4119   hence "c\<noteq>{}" using c by auto
  4120   def k \<equiv> "Max (f ` c)"
  4120   def k \<equiv> "Max (f ` c)"
  4121   have "convex_on {x - ?d..x + ?d} f" apply(rule convex_on_subset[OF assms(2)])
  4121   have "convex_on {x - ?d..x + ?d} f" apply(rule convex_on_subset[OF assms(2)])
  4122     apply(rule subset_trans[OF _ e(1)]) unfolding subset_eq mem_cball proof 
  4122     apply(rule subset_trans[OF _ e(1)]) unfolding subset_eq mem_cball proof
  4123     fix z assume z:"z\<in>{x - ?d..x + ?d}"
  4123     fix z assume z:"z\<in>{x - ?d..x + ?d}"
  4124     have e:"e = setsum (\<lambda>i. d) {..<DIM('a)}" unfolding setsum_constant d_def using dimge1
  4124     have e:"e = setsum (\<lambda>i. d) {..<DIM('a)}" unfolding setsum_constant d_def using dimge1
  4125       unfolding real_eq_of_nat by auto
  4125       unfolding real_eq_of_nat by auto
  4126     show "dist x z \<le> e" unfolding dist_norm e apply(rule_tac order_trans[OF norm_le_l1], rule setsum_mono)
  4126     show "dist x z \<le> e" unfolding dist_norm e apply(rule_tac order_trans[OF norm_le_l1], rule setsum_mono)
  4127       using z[unfolded mem_interval] apply(erule_tac x=i in allE) by auto qed
  4127       using z[unfolded mem_interval] apply(erule_tac x=i in allE) by auto qed
  4130   have "d \<le> e" unfolding d_def apply(rule mult_imp_div_pos_le) using `e>0` dimge1 unfolding mult_le_cancel_left1 by auto
  4130   have "d \<le> e" unfolding d_def apply(rule mult_imp_div_pos_le) using `e>0` dimge1 unfolding mult_le_cancel_left1 by auto
  4131   hence dsube:"cball x d \<subseteq> cball x e" unfolding subset_eq Ball_def mem_cball by auto
  4131   hence dsube:"cball x d \<subseteq> cball x e" unfolding subset_eq Ball_def mem_cball by auto
  4132   have conv:"convex_on (cball x d) f" apply(rule convex_on_subset, rule convex_on_subset[OF assms(2)]) apply(rule e(1)) using dsube by auto
  4132   have conv:"convex_on (cball x d) f" apply(rule convex_on_subset, rule convex_on_subset[OF assms(2)]) apply(rule e(1)) using dsube by auto
  4133   hence "\<forall>y\<in>cball x d. abs (f y) \<le> k + 2 * abs (f x)" apply(rule_tac convex_bounds_lemma) apply assumption proof
  4133   hence "\<forall>y\<in>cball x d. abs (f y) \<le> k + 2 * abs (f x)" apply(rule_tac convex_bounds_lemma) apply assumption proof
  4134     fix y assume y:"y\<in>cball x d"
  4134     fix y assume y:"y\<in>cball x d"
  4135     { fix i assume "i<DIM('a)" hence "x $$ i - d \<le> y $$ i"  "y $$ i \<le> x $$ i + d" 
  4135     { fix i assume "i<DIM('a)" hence "x $$ i - d \<le> y $$ i"  "y $$ i \<le> x $$ i + d"
  4136         using order_trans[OF component_le_norm y[unfolded mem_cball dist_norm], of i] by auto  }
  4136         using order_trans[OF component_le_norm y[unfolded mem_cball dist_norm], of i] by auto  }
  4137     thus "f y \<le> k" apply(rule_tac k[rule_format]) unfolding mem_cball mem_interval dist_norm 
  4137     thus "f y \<le> k" apply(rule_tac k[rule_format]) unfolding mem_cball mem_interval dist_norm
  4138       by auto qed
  4138       by auto qed
  4139   hence "continuous_on (ball x d) f" apply(rule_tac convex_on_bounded_continuous)
  4139   hence "continuous_on (ball x d) f" apply(rule_tac convex_on_bounded_continuous)
  4140     apply(rule open_ball, rule convex_on_subset[OF conv], rule ball_subset_cball)
  4140     apply(rule open_ball, rule convex_on_subset[OF conv], rule ball_subset_cball)
  4141     apply force
  4141     apply force
  4142     done
  4142     done
  4143   thus "continuous (at x) f" unfolding continuous_on_eq_continuous_at[OF open_ball]
  4143   thus "continuous (at x) f" unfolding continuous_on_eq_continuous_at[OF open_ball]
  4144     using `d>0` by auto 
  4144     using `d>0` by auto
  4145 qed
  4145 qed
  4146 
  4146 
  4147 subsection {* Line segments, Starlike Sets, etc. *}
  4147 subsection {* Line segments, Starlike Sets, etc. *}
  4148 
  4148 
  4149 (* Use the same overloading tricks as for intervals, so that 
  4149 (* Use the same overloading tricks as for intervals, so that
  4150    segment[a,b] is closed and segment(a,b) is open relative to affine hull. *)
  4150    segment[a,b] is closed and segment(a,b) is open relative to affine hull. *)
  4151 
  4151 
  4152 definition
  4152 definition
  4153   midpoint :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a" where
  4153   midpoint :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a" where
  4154   "midpoint a b = (inverse (2::real)) *\<^sub>R (a + b)"
  4154   "midpoint a b = (inverse (2::real)) *\<^sub>R (a + b)"
  4216 lemma segment_convex_hull:
  4216 lemma segment_convex_hull:
  4217  "closed_segment a b = convex hull {a,b}" proof-
  4217  "closed_segment a b = convex hull {a,b}" proof-
  4218   have *:"\<And>x. {x} \<noteq> {}" by auto
  4218   have *:"\<And>x. {x} \<noteq> {}" by auto
  4219   have **:"\<And>u v. u + v = 1 \<longleftrightarrow> u = 1 - (v::real)" by auto
  4219   have **:"\<And>u v. u + v = 1 \<longleftrightarrow> u = 1 - (v::real)" by auto
  4220   show ?thesis unfolding segment convex_hull_insert[OF *] convex_hull_singleton apply(rule set_eqI)
  4220   show ?thesis unfolding segment convex_hull_insert[OF *] convex_hull_singleton apply(rule set_eqI)
  4221     unfolding mem_Collect_eq apply(rule,erule exE) 
  4221     unfolding mem_Collect_eq apply(rule,erule exE)
  4222     apply(rule_tac x="1 - u" in exI) apply rule defer apply(rule_tac x=u in exI) defer
  4222     apply(rule_tac x="1 - u" in exI) apply rule defer apply(rule_tac x=u in exI) defer
  4223     apply(erule exE, (erule conjE)?)+ apply(rule_tac x="1 - u" in exI) unfolding ** by auto qed
  4223     apply(erule exE, (erule conjE)?)+ apply(rule_tac x="1 - u" in exI) unfolding ** by auto qed
  4224 
  4224 
  4225 lemma convex_segment: "convex (closed_segment a b)"
  4225 lemma convex_segment: "convex (closed_segment a b)"
  4226   unfolding segment_convex_hull by(rule convex_convex_hull)
  4226   unfolding segment_convex_hull by(rule convex_convex_hull)
  4239   fixes x a b :: "'a::euclidean_space"
  4239   fixes x a b :: "'a::euclidean_space"
  4240   assumes "x \<in> closed_segment a b"
  4240   assumes "x \<in> closed_segment a b"
  4241   shows "norm(x - a) \<le> norm(b - a)" "norm(x - b) \<le> norm(b - a)"
  4241   shows "norm(x - a) \<le> norm(b - a)" "norm(x - b) \<le> norm(b - a)"
  4242   using segment_furthest_le[OF assms, of a]
  4242   using segment_furthest_le[OF assms, of a]
  4243   using segment_furthest_le[OF assms, of b]
  4243   using segment_furthest_le[OF assms, of b]
  4244   by (auto simp add:norm_minus_commute) 
  4244   by (auto simp add:norm_minus_commute)
  4245 
  4245 
  4246 lemma segment_refl:"closed_segment a a = {a}" unfolding segment by (auto simp add: algebra_simps)
  4246 lemma segment_refl:"closed_segment a a = {a}" unfolding segment by (auto simp add: algebra_simps)
  4247 
  4247 
  4248 lemma between_mem_segment: "between (a,b) x \<longleftrightarrow> x \<in> closed_segment a b"
  4248 lemma between_mem_segment: "between (a,b) x \<longleftrightarrow> x \<in> closed_segment a b"
  4249   unfolding between_def by auto
  4249   unfolding between_def by auto
  4250 
  4250 
  4251 lemma between:"between (a,b) (x::'a::euclidean_space) \<longleftrightarrow> dist a b = (dist a x) + (dist x b)"
  4251 lemma between:"between (a,b) (x::'a::euclidean_space) \<longleftrightarrow> dist a b = (dist a x) + (dist x b)"
  4252 proof(cases "a = b")
  4252 proof(cases "a = b")
  4253   case True thus ?thesis unfolding between_def split_conv
  4253   case True thus ?thesis unfolding between_def split_conv
  4254     by(auto simp add:segment_refl dist_commute) next
  4254     by(auto simp add:segment_refl dist_commute) next
  4255   case False hence Fal:"norm (a - b) \<noteq> 0" and Fal2: "norm (a - b) > 0" by auto 
  4255   case False hence Fal:"norm (a - b) \<noteq> 0" and Fal2: "norm (a - b) > 0" by auto
  4256   have *:"\<And>u. a - ((1 - u) *\<^sub>R a + u *\<^sub>R b) = u *\<^sub>R (a - b)" by (auto simp add: algebra_simps)
  4256   have *:"\<And>u. a - ((1 - u) *\<^sub>R a + u *\<^sub>R b) = u *\<^sub>R (a - b)" by (auto simp add: algebra_simps)
  4257   show ?thesis unfolding between_def split_conv closed_segment_def mem_Collect_eq
  4257   show ?thesis unfolding between_def split_conv closed_segment_def mem_Collect_eq
  4258     apply rule apply(erule exE, (erule conjE)+) apply(subst dist_triangle_eq) proof-
  4258     apply rule apply(erule exE, (erule conjE)+) apply(subst dist_triangle_eq) proof-
  4259       fix u assume as:"x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \<le> u" "u \<le> 1" 
  4259       fix u assume as:"x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \<le> u" "u \<le> 1"
  4260       hence *:"a - x = u *\<^sub>R (a - b)" "x - b = (1 - u) *\<^sub>R (a - b)"
  4260       hence *:"a - x = u *\<^sub>R (a - b)" "x - b = (1 - u) *\<^sub>R (a - b)"
  4261         unfolding as(1) by(auto simp add:algebra_simps)
  4261         unfolding as(1) by(auto simp add:algebra_simps)
  4262       show "norm (a - x) *\<^sub>R (x - b) = norm (x - b) *\<^sub>R (a - x)"
  4262       show "norm (a - x) *\<^sub>R (x - b) = norm (x - b) *\<^sub>R (a - x)"
  4263         unfolding norm_minus_commute[of x a] * using as(2,3)
  4263         unfolding norm_minus_commute[of x a] * using as(2,3)
  4264         by(auto simp add: field_simps)
  4264         by(auto simp add: field_simps)
  4265     next assume as:"dist a b = dist a x + dist x b"
  4265     next assume as:"dist a b = dist a x + dist x b"
  4266       have "norm (a - x) / norm (a - b) \<le> 1" unfolding divide_le_eq_1_pos[OF Fal2]
  4266       have "norm (a - x) / norm (a - b) \<le> 1" unfolding divide_le_eq_1_pos[OF Fal2]
  4267         unfolding as[unfolded dist_norm] norm_ge_zero by auto 
  4267         unfolding as[unfolded dist_norm] norm_ge_zero by auto
  4268       thus "\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1" apply(rule_tac x="dist a x / dist a b" in exI)
  4268       thus "\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1" apply(rule_tac x="dist a x / dist a b" in exI)
  4269         unfolding dist_norm apply(subst euclidean_eq) apply rule defer apply(rule, rule divide_nonneg_pos) prefer 4
  4269         unfolding dist_norm apply(subst euclidean_eq) apply rule defer apply(rule, rule divide_nonneg_pos) prefer 4
  4270       proof(rule,rule) fix i assume i:"i<DIM('a)"
  4270       proof(rule,rule) fix i assume i:"i<DIM('a)"
  4271           have "((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) $$ i =
  4271           have "((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) $$ i =
  4272             ((norm (a - b) - norm (a - x)) * (a $$ i) + norm (a - x) * (b $$ i)) / norm (a - b)"
  4272             ((norm (a - b) - norm (a - x)) * (a $$ i) + norm (a - x) * (b $$ i)) / norm (a - b)"
  4273             using Fal by(auto simp add: field_simps)
  4273             using Fal by(auto simp add: field_simps)
  4274           also have "\<dots> = x$$i" apply(rule divide_eq_imp[OF Fal])
  4274           also have "\<dots> = x$$i" apply(rule divide_eq_imp[OF Fal])
  4275             unfolding as[unfolded dist_norm] using as[unfolded dist_triangle_eq] apply-
  4275             unfolding as[unfolded dist_norm] using as[unfolded dist_triangle_eq] apply-
  4276             apply(subst (asm) euclidean_eq) using i apply(erule_tac x=i in allE) by(auto simp add:field_simps)
  4276             apply(subst (asm) euclidean_eq) using i apply(erule_tac x=i in allE) by(auto simp add:field_simps)
  4277           finally show "x $$ i = ((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) $$ i" 
  4277           finally show "x $$ i = ((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) $$ i"
  4278             by auto
  4278             by auto
  4279         qed(insert Fal2, auto) qed qed
  4279         qed(insert Fal2, auto) qed qed
  4280 
  4280 
  4281 lemma between_midpoint: fixes a::"'a::euclidean_space" shows
  4281 lemma between_midpoint: fixes a::"'a::euclidean_space" shows
  4282   "between (a,b) (midpoint a b)" (is ?t1) 
  4282   "between (a,b) (midpoint a b)" (is ?t1)
  4283   "between (b,a) (midpoint a b)" (is ?t2)
  4283   "between (b,a) (midpoint a b)" (is ?t2)
  4284 proof- 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" by auto
  4284 proof- 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" by auto
  4285   show ?t1 ?t2 unfolding between midpoint_def dist_norm apply(rule_tac[!] *)
  4285   show ?t1 ?t2 unfolding between midpoint_def dist_norm apply(rule_tac[!] *)
  4286     unfolding euclidean_eq[where 'a='a]
  4286     unfolding euclidean_eq[where 'a='a]
  4287     by(auto simp add:field_simps) qed
  4287     by(auto simp add:field_simps) qed
  4301     apply(rule) defer unfolding subset_eq Ball_def mem_ball proof(rule,rule)
  4301     apply(rule) defer unfolding subset_eq Ball_def mem_ball proof(rule,rule)
  4302     fix y assume as:"dist (x - e *\<^sub>R (x - c)) y < e * d"
  4302     fix y assume as:"dist (x - e *\<^sub>R (x - c)) y < e * d"
  4303     have *:"y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" using `e>0` by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib)
  4303     have *:"y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" using `e>0` by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib)
  4304     have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = abs(1/e) * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)"
  4304     have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = abs(1/e) * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)"
  4305       unfolding dist_norm unfolding norm_scaleR[symmetric] apply(rule arg_cong[where f=norm]) using `e>0`
  4305       unfolding dist_norm unfolding norm_scaleR[symmetric] apply(rule arg_cong[where f=norm]) using `e>0`
  4306       by(auto simp add: euclidean_eq[where 'a='a] field_simps) 
  4306       by(auto simp add: euclidean_eq[where 'a='a] field_simps)
  4307     also have "\<dots> = abs(1/e) * norm (x - e *\<^sub>R (x - c) - y)" by(auto intro!:arg_cong[where f=norm] simp add: algebra_simps)
  4307     also have "\<dots> = abs(1/e) * norm (x - e *\<^sub>R (x - c) - y)" by(auto intro!:arg_cong[where f=norm] simp add: algebra_simps)
  4308     also have "\<dots> < d" using as[unfolded dist_norm] and `e>0`
  4308     also have "\<dots> < d" using as[unfolded dist_norm] and `e>0`
  4309       by(auto simp add:pos_divide_less_eq[OF `e>0`] mult_commute)
  4309       by(auto simp add:pos_divide_less_eq[OF `e>0`] mult_commute)
  4310     finally show "y \<in> s" apply(subst *) apply(rule assms(1)[unfolded convex_alt,rule_format])
  4310     finally show "y \<in> s" apply(subst *) apply(rule assms(1)[unfolded convex_alt,rule_format])
  4311       apply(rule d[unfolded subset_eq,rule_format]) unfolding mem_ball using assms(3-5) by auto
  4311       apply(rule d[unfolded subset_eq,rule_format]) unfolding mem_ball using assms(3-5) by auto
  4332   def z \<equiv> "c + ((1 - e) / e) *\<^sub>R (x - y)"
  4332   def z \<equiv> "c + ((1 - e) / e) *\<^sub>R (x - y)"
  4333   have *:"x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" unfolding z_def using `e>0` by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib)
  4333   have *:"x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" unfolding z_def using `e>0` by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib)
  4334   have "z\<in>interior s" apply(rule interior_mono[OF d,unfolded subset_eq,rule_format])
  4334   have "z\<in>interior s" apply(rule interior_mono[OF d,unfolded subset_eq,rule_format])
  4335     unfolding interior_open[OF open_ball] mem_ball z_def dist_norm using y and assms(4,5)
  4335     unfolding interior_open[OF open_ball] mem_ball z_def dist_norm using y and assms(4,5)
  4336     by(auto simp add:field_simps norm_minus_commute)
  4336     by(auto simp add:field_simps norm_minus_commute)
  4337   thus ?thesis unfolding * apply - apply(rule mem_interior_convex_shrink) 
  4337   thus ?thesis unfolding * apply - apply(rule mem_interior_convex_shrink)
  4338     using assms(1,4-5) `y\<in>s` by auto qed
  4338     using assms(1,4-5) `y\<in>s` by auto qed
  4339 
  4339 
  4340 subsection {* Some obvious but surprisingly hard simplex lemmas *}
  4340 subsection {* Some obvious but surprisingly hard simplex lemmas *}
  4341 
  4341 
  4342 lemma simplex:
  4342 lemma simplex:
  4348   unfolding if_smult and setsum_delta_notmem[OF assms(2)] by auto
  4348   unfolding if_smult and setsum_delta_notmem[OF assms(2)] by auto
  4349 
  4349 
  4350 lemma substd_simplex: assumes "d\<subseteq>{..<DIM('a::euclidean_space)}"
  4350 lemma substd_simplex: assumes "d\<subseteq>{..<DIM('a::euclidean_space)}"
  4351   shows "convex hull (insert 0 { basis i | i. i : d}) =
  4351   shows "convex hull (insert 0 { basis i | i. i : d}) =
  4352         {x::'a::euclidean_space . (!i<DIM('a). 0 <= x$$i) & setsum (%i. x$$i) d <= 1 &
  4352         {x::'a::euclidean_space . (!i<DIM('a). 0 <= x$$i) & setsum (%i. x$$i) d <= 1 &
  4353   (!i<DIM('a). i ~: d --> x$$i = 0)}" 
  4353   (!i<DIM('a). i ~: d --> x$$i = 0)}"
  4354   (is "convex hull (insert 0 ?p) = ?s")
  4354   (is "convex hull (insert 0 ?p) = ?s")
  4355 (* Proof is a modified copy of the proof of similar lemma std_simplex in Convex_Euclidean_Space.thy *)
  4355 (* Proof is a modified copy of the proof of similar lemma std_simplex in Convex_Euclidean_Space.thy *)
  4356 proof- let ?D = d (*"{..<DIM('a::euclidean_space)}"*)
  4356 proof- let ?D = d (*"{..<DIM('a::euclidean_space)}"*)
  4357   have "0 ~: ?p" using assms by (auto simp: image_def)
  4357   have "0 ~: ?p" using assms by (auto simp: image_def)
  4358   have "{(basis i)::'n::euclidean_space |i. i \<in> ?D} = basis ` ?D" by auto
  4358   have "{(basis i)::'n::euclidean_space |i. i \<in> ?D} = basis ` ?D" by auto
  4359   note sumbas = this setsum_reindex[OF basis_inj_on[of d], unfolded o_def, OF assms]
  4359   note sumbas = this setsum_reindex[OF basis_inj_on[of d], unfolded o_def, OF assms]
  4360   show ?thesis unfolding simplex[OF finite_substdbasis `0 ~: ?p`] 
  4360   show ?thesis unfolding simplex[OF finite_substdbasis `0 ~: ?p`]
  4361     apply(rule set_eqI) unfolding mem_Collect_eq apply rule
  4361     apply(rule set_eqI) unfolding mem_Collect_eq apply rule
  4362     apply(erule exE, (erule conjE)+) apply(erule_tac[2] conjE)+ proof-
  4362     apply(erule exE, (erule conjE)+) apply(erule_tac[2] conjE)+ proof-
  4363     fix x::"'a::euclidean_space" and u assume as: "\<forall>x\<in>{basis i |i. i \<in>?D}. 0 \<le> u x"
  4363     fix x::"'a::euclidean_space" and u assume as: "\<forall>x\<in>{basis i |i. i \<in>?D}. 0 \<le> u x"
  4364       "setsum u {basis i |i. i \<in> ?D} \<le> 1" "(\<Sum>x\<in>{basis i |i. i \<in>?D}. u x *\<^sub>R x) = x"
  4364       "setsum u {basis i |i. i \<in> ?D} \<le> 1" "(\<Sum>x\<in>{basis i |i. i \<in>?D}. u x *\<^sub>R x) = x"
  4365     have *:"\<forall>i<DIM('a). i:d --> u (basis i) = x$$i" and "(!i<DIM('a). i ~: d --> x $$ i = 0)" using as(3) 
  4365     have *:"\<forall>i<DIM('a). i:d --> u (basis i) = x$$i" and "(!i<DIM('a). i ~: d --> x $$ i = 0)" using as(3)
  4366       unfolding sumbas unfolding substdbasis_expansion_unique[OF assms] by auto
  4366       unfolding sumbas unfolding substdbasis_expansion_unique[OF assms] by auto
  4367     hence **:"setsum u {basis i |i. i \<in> ?D} = setsum (op $$ x) ?D" unfolding sumbas 
  4367     hence **:"setsum u {basis i |i. i \<in> ?D} = setsum (op $$ x) ?D" unfolding sumbas
  4368       apply-apply(rule setsum_cong2) using assms by auto
  4368       apply-apply(rule setsum_cong2) using assms by auto
  4369     have " (\<forall>i<DIM('a). 0 \<le> x$$i) \<and> setsum (op $$ x) ?D \<le> 1" 
  4369     have " (\<forall>i<DIM('a). 0 \<le> x$$i) \<and> setsum (op $$ x) ?D \<le> 1"
  4370       apply - proof(rule,rule,rule)
  4370       apply - proof(rule,rule,rule)
  4371       fix i assume i:"i<DIM('a)" have "i : d ==> 0 \<le> x$$i" unfolding *[rule_format,OF i,symmetric] 
  4371       fix i assume i:"i<DIM('a)" have "i : d ==> 0 \<le> x$$i" unfolding *[rule_format,OF i,symmetric]
  4372          apply(rule_tac as(1)[rule_format]) by auto
  4372          apply(rule_tac as(1)[rule_format]) by auto
  4373       moreover have "i ~: d ==> 0 \<le> x$$i" 
  4373       moreover have "i ~: d ==> 0 \<le> x$$i"
  4374         using `(!i<DIM('a). i ~: d --> x $$ i = 0)`[rule_format, OF i] by auto
  4374         using `(!i<DIM('a). i ~: d --> x $$ i = 0)`[rule_format, OF i] by auto
  4375       ultimately show "0 \<le> x$$i" by auto
  4375       ultimately show "0 \<le> x$$i" by auto
  4376     qed(insert as(2)[unfolded **], auto)
  4376     qed(insert as(2)[unfolded **], auto)
  4377     from this show " (\<forall>i<DIM('a). 0 \<le> x$$i) \<and> setsum (op $$ x) ?D \<le> 1 & (!i<DIM('a). i ~: d --> x $$ i = 0)" 
  4377     from this show " (\<forall>i<DIM('a). 0 \<le> x$$i) \<and> setsum (op $$ x) ?D \<le> 1 & (!i<DIM('a). i ~: d --> x $$ i = 0)"
  4378       using `(!i<DIM('a). i ~: d --> x $$ i = 0)` by auto
  4378       using `(!i<DIM('a). i ~: d --> x $$ i = 0)` by auto
  4379   next fix x::"'a::euclidean_space" assume as:"\<forall>i<DIM('a). 0 \<le> x $$ i" "setsum (op $$ x) ?D \<le> 1"
  4379   next fix x::"'a::euclidean_space" assume as:"\<forall>i<DIM('a). 0 \<le> x $$ i" "setsum (op $$ x) ?D \<le> 1"
  4380       "(!i<DIM('a). i ~: d --> x $$ i = 0)"
  4380       "(!i<DIM('a). i ~: d --> x $$ i = 0)"
  4381     show "\<exists>u. (\<forall>x\<in>{basis i |i. i \<in> ?D}. 0 \<le> u x) \<and>
  4381     show "\<exists>u. (\<forall>x\<in>{basis i |i. i \<in> ?D}. 0 \<le> u x) \<and>
  4382       setsum u {basis i |i. i \<in> ?D} \<le> 1 \<and> (\<Sum>x\<in>{basis i |i. i \<in> ?D}. u x *\<^sub>R x) = x"
  4382       setsum u {basis i |i. i \<in> ?D} \<le> 1 \<and> (\<Sum>x\<in>{basis i |i. i \<in> ?D}. u x *\<^sub>R x) = x"
  4422       fix i assume "i\<in>{..<DIM('a)}" hence "abs (y$$i - x$$i) < ?d" apply-apply(rule le_less_trans)
  4422       fix i assume "i\<in>{..<DIM('a)}" hence "abs (y$$i - x$$i) < ?d" apply-apply(rule le_less_trans)
  4423         using component_le_norm[of "y - x" i]
  4423         using component_le_norm[of "y - x" i]
  4424         using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] by(auto simp add: norm_minus_commute)
  4424         using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] by(auto simp add: norm_minus_commute)
  4425       thus "y $$ i \<le> x $$ i + ?d" by auto qed
  4425       thus "y $$ i \<le> x $$ i + ?d" by auto qed
  4426     also have "\<dots> \<le> 1" unfolding setsum_addf setsum_constant real_eq_of_nat by(auto simp add: Suc_le_eq)
  4426     also have "\<dots> \<le> 1" unfolding setsum_addf setsum_constant real_eq_of_nat by(auto simp add: Suc_le_eq)
  4427     finally show "(\<forall>i<DIM('a). 0 \<le> y $$ i) \<and> setsum (op $$ y) {..<DIM('a)} \<le> 1" 
  4427     finally show "(\<forall>i<DIM('a). 0 \<le> y $$ i) \<and> setsum (op $$ y) {..<DIM('a)} \<le> 1"
  4428     proof safe fix i assume i:"i<DIM('a)"
  4428     proof safe fix i assume i:"i<DIM('a)"
  4429       have "norm (x - y) < x$$i" apply(rule less_le_trans) 
  4429       have "norm (x - y) < x$$i" apply(rule less_le_trans)
  4430         apply(rule y[unfolded min_less_iff_conj dist_norm, THEN conjunct1]) using i by auto
  4430         apply(rule y[unfolded min_less_iff_conj dist_norm, THEN conjunct1]) using i by auto
  4431       thus "0 \<le> y$$i" using component_le_norm[of "x - y" i] and as(1)[rule_format, of i] by auto
  4431       thus "0 \<le> y$$i" using component_le_norm[of "x - y" i] and as(1)[rule_format, of i] by auto
  4432     qed qed auto qed
  4432     qed qed auto qed
  4433 
  4433 
  4434 lemma interior_std_simplex_nonempty: obtains a::"'a::euclidean_space" where
  4434 lemma interior_std_simplex_nonempty: obtains a::"'a::euclidean_space" where
  4454 proof-
  4454 proof-
  4455 have "finite d" apply(rule finite_subset) using assms by auto
  4455 have "finite d" apply(rule finite_subset) using assms by auto
  4456 { assume "d={}" hence ?thesis using rel_interior_sing using euclidean_eq[of _ 0] by auto }
  4456 { assume "d={}" hence ?thesis using rel_interior_sing using euclidean_eq[of _ 0] by auto }
  4457 moreover
  4457 moreover
  4458 { assume "d~={}"
  4458 { assume "d~={}"
  4459 have h0: "affine hull (convex hull (insert 0 ?p))={x::'a::euclidean_space. (!i<DIM('a). i ~: d --> x$$i = 0)}" 
  4459 have h0: "affine hull (convex hull (insert 0 ?p))={x::'a::euclidean_space. (!i<DIM('a). i ~: d --> x$$i = 0)}"
  4460    using affine_hull_convex_hull affine_hull_substd_basis assms by auto 
  4460    using affine_hull_convex_hull affine_hull_substd_basis assms by auto
  4461 have aux: "!x::'n::euclidean_space. !i. ((! i:d. 0 <= x$$i) & (!i. i ~: d --> x$$i = 0))--> 0 <= x$$i" by auto
  4461 have aux: "!x::'n::euclidean_space. !i. ((! i:d. 0 <= x$$i) & (!i. i ~: d --> x$$i = 0))--> 0 <= x$$i" by auto
  4462 { fix x::"'a::euclidean_space" assume x_def: "x : rel_interior (convex hull (insert 0 ?p))"
  4462 { fix x::"'a::euclidean_space" assume x_def: "x : rel_interior (convex hull (insert 0 ?p))"
  4463   from this obtain e where e0: "e>0" and 
  4463   from this obtain e where e0: "e>0" and
  4464        "ball x e Int {xa. (!i<DIM('a). i ~: d --> xa$$i = 0)} <= convex hull (insert 0 ?p)" 
  4464        "ball x e Int {xa. (!i<DIM('a). i ~: d --> xa$$i = 0)} <= convex hull (insert 0 ?p)"
  4465        using mem_rel_interior_ball[of x "convex hull (insert 0 ?p)"] h0 by auto   
  4465        using mem_rel_interior_ball[of x "convex hull (insert 0 ?p)"] h0 by auto
  4466   hence as: "ALL xa. (dist x xa < e & (!i<DIM('a). i ~: d --> xa$$i = 0)) -->
  4466   hence as: "ALL xa. (dist x xa < e & (!i<DIM('a). i ~: d --> xa$$i = 0)) -->
  4467     (!i : d. 0 <= xa $$ i) & setsum (op $$ xa) d <= 1"
  4467     (!i : d. 0 <= xa $$ i) & setsum (op $$ xa) d <= 1"
  4468     unfolding ball_def unfolding substd_simplex[OF assms] using assms by auto
  4468     unfolding ball_def unfolding substd_simplex[OF assms] using assms by auto
  4469   have x0: "(!i<DIM('a). i ~: d --> x$$i = 0)" 
  4469   have x0: "(!i<DIM('a). i ~: d --> x$$i = 0)"
  4470     using x_def rel_interior_subset  substd_simplex[OF assms] by auto
  4470     using x_def rel_interior_subset  substd_simplex[OF assms] by auto
  4471   have "(!i : d. 0 < x $$ i) & setsum (op $$ x) d < 1 & (!i<DIM('a). i ~: d --> x$$i = 0)" apply(rule,rule) 
  4471   have "(!i : d. 0 < x $$ i) & setsum (op $$ x) d < 1 & (!i<DIM('a). i ~: d --> x$$i = 0)" apply(rule,rule)
  4472   proof-
  4472   proof-
  4473     fix i::nat assume "i:d" 
  4473     fix i::nat assume "i:d"
  4474     hence "\<forall>ia\<in>d. 0 \<le> (x - (e / 2) *\<^sub>R basis i) $$ ia" apply-apply(rule as[rule_format,THEN conjunct1])
  4474     hence "\<forall>ia\<in>d. 0 \<le> (x - (e / 2) *\<^sub>R basis i) $$ ia" apply-apply(rule as[rule_format,THEN conjunct1])
  4475       unfolding dist_norm using assms `e>0` x0 by auto
  4475       unfolding dist_norm using assms `e>0` x0 by auto
  4476     thus "0 < x $$ i" apply(erule_tac x=i in ballE) using `e>0` `i\<in>d` assms by auto
  4476     thus "0 < x $$ i" apply(erule_tac x=i in ballE) using `e>0` `i\<in>d` assms by auto
  4477   next obtain a where a:"a:d" using `d ~= {}` by auto
  4477   next obtain a where a:"a:d" using `d ~= {}` by auto
  4478     have **:"dist x (x + (e / 2) *\<^sub>R basis a) < e"
  4478     have **:"dist x (x + (e / 2) *\<^sub>R basis a) < e"
  4486       using as[THEN spec[where x="x + (e / 2) *\<^sub>R basis a"]] `a:d` using x0
  4486       using as[THEN spec[where x="x + (e / 2) *\<^sub>R basis a"]] `a:d` using x0
  4487       by(auto elim:allE[where x=a])
  4487       by(auto elim:allE[where x=a])
  4488     have "setsum (op $$ x) d < setsum (op $$ (x + (e / 2) *\<^sub>R basis a)) d" unfolding * setsum_addf
  4488     have "setsum (op $$ x) d < setsum (op $$ (x + (e / 2) *\<^sub>R basis a)) d" unfolding * setsum_addf
  4489       using `0<e` `a:d` using `finite d` by(auto simp add: setsum_delta')
  4489       using `0<e` `a:d` using `finite d` by(auto simp add: setsum_delta')
  4490     also have "\<dots> \<le> 1" using ** h1 as[rule_format, of "x + (e / 2) *\<^sub>R basis a"] by auto
  4490     also have "\<dots> \<le> 1" using ** h1 as[rule_format, of "x + (e / 2) *\<^sub>R basis a"] by auto
  4491     finally show "setsum (op $$ x) d < 1 & (!i<DIM('a). i ~: d --> x$$i = 0)" using x0 by auto 
  4491     finally show "setsum (op $$ x) d < 1 & (!i<DIM('a). i ~: d --> x$$i = 0)" using x0 by auto
  4492   qed
  4492   qed
  4493 }
  4493 }
  4494 moreover
  4494 moreover
  4495 {
  4495 {
  4496   fix x::"'a::euclidean_space" assume as: "x : ?s"
  4496   fix x::"'a::euclidean_space" assume as: "x : ?s"
  4497   have "!i. ((0<x$$i) | (0=x$$i) --> 0<=x$$i)" by auto
  4497   have "!i. ((0<x$$i) | (0=x$$i) --> 0<=x$$i)" by auto
  4498   moreover have "!i. (i:d) | (i ~: d)" by auto
  4498   moreover have "!i. (i:d) | (i ~: d)" by auto
  4499   ultimately 
  4499   ultimately
  4500   have "!i. ( (ALL i:d. 0 < x$$i) & (ALL i. i ~: d --> x$$i = 0) ) --> 0 <= x$$i" by metis
  4500   have "!i. ( (ALL i:d. 0 < x$$i) & (ALL i. i ~: d --> x$$i = 0) ) --> 0 <= x$$i" by metis
  4501   hence h2: "x : convex hull (insert 0 ?p)" using as assms 
  4501   hence h2: "x : convex hull (insert 0 ?p)" using as assms
  4502     unfolding substd_simplex[OF assms] by fastforce 
  4502     unfolding substd_simplex[OF assms] by fastforce
  4503   obtain a where a:"a:d" using `d ~= {}` by auto
  4503   obtain a where a:"a:d" using `d ~= {}` by auto
  4504   let ?d = "(1 - setsum (op $$ x) d) / real (card d)"
  4504   let ?d = "(1 - setsum (op $$ x) d) / real (card d)"
  4505   have "0 < card d" using `d ~={}` `finite d` by (simp add: card_gt_0_iff)
  4505   have "0 < card d" using `d ~={}` `finite d` by (simp add: card_gt_0_iff)
  4506   have "Min ((op $$ x) ` d) > 0" using as `d \<noteq> {}` `finite d` by (simp add: Min_grI)
  4506   have "Min ((op $$ x) ` d) > 0" using as `d \<noteq> {}` `finite d` by (simp add: Min_grI)
  4507   moreover have "?d > 0" apply(rule divide_pos_pos) using as using `0 < card d` by auto
  4507   moreover have "?d > 0" apply(rule divide_pos_pos) using as using `0 < card d` by auto
  4519         by(auto simp add: norm_minus_commute)
  4519         by(auto simp add: norm_minus_commute)
  4520       thus "y $$ i \<le> x $$ i + ?d" by auto qed
  4520       thus "y $$ i \<le> x $$ i + ?d" by auto qed
  4521     also have "\<dots> \<le> 1" unfolding setsum_addf setsum_constant real_eq_of_nat
  4521     also have "\<dots> \<le> 1" unfolding setsum_addf setsum_constant real_eq_of_nat
  4522       using `0 < card d` by auto
  4522       using `0 < card d` by auto
  4523     finally show "setsum (op $$ y) d \<le> 1" .
  4523     finally show "setsum (op $$ y) d \<le> 1" .
  4524      
  4524 
  4525     fix i assume "i<DIM('a)" thus "0 \<le> y$$i" 
  4525     fix i assume "i<DIM('a)" thus "0 \<le> y$$i"
  4526     proof(cases "i\<in>d") case True
  4526     proof(cases "i\<in>d") case True
  4527       have "norm (x - y) < x$$i" using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1]
  4527       have "norm (x - y) < x$$i" using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1]
  4528         using Min_gr_iff[of "op $$ x ` d" "norm (x - y)"] `0 < card d` `i:d`
  4528         using Min_gr_iff[of "op $$ x ` d" "norm (x - y)"] `0 < card d` `i:d`
  4529         by (simp add: card_gt_0_iff)
  4529         by (simp add: card_gt_0_iff)
  4530       thus "0 \<le> y$$i" using component_le_norm[of "x - y" i] and as(1)[rule_format] by auto
  4530       thus "0 \<le> y$$i" using component_le_norm[of "x - y" i] and as(1)[rule_format] by auto
  4546   have *:"{basis i :: 'a | i. i \<in> ?D} = basis ` ?D" by auto
  4546   have *:"{basis i :: 'a | i. i \<in> ?D} = basis ` ?D" by auto
  4547   have "finite d" apply(rule finite_subset) using assms(2) by auto
  4547   have "finite d" apply(rule finite_subset) using assms(2) by auto
  4548   hence d1: "0 < real(card d)" using `d ~={}` by auto
  4548   hence d1: "0 < real(card d)" using `d ~={}` by auto
  4549   { fix i assume "i:d" have "?a $$ i = inverse (2 * real (card d))"
  4549   { fix i assume "i:d" have "?a $$ i = inverse (2 * real (card d))"
  4550       unfolding * setsum_reindex[OF basis_inj_on, OF assms(2)] o_def
  4550       unfolding * setsum_reindex[OF basis_inj_on, OF assms(2)] o_def
  4551       apply(rule trans[of _ "setsum (\<lambda>j. if i = j then inverse (2 * real (card d)) else 0) ?D"]) 
  4551       apply(rule trans[of _ "setsum (\<lambda>j. if i = j then inverse (2 * real (card d)) else 0) ?D"])
  4552       unfolding euclidean_component_setsum
  4552       unfolding euclidean_component_setsum
  4553       apply(rule setsum_cong2)
  4553       apply(rule setsum_cong2)
  4554       using `i:d` `finite d` setsum_delta'[of d i "(%k. inverse (2 * real (card d)))"] d1 assms(2)
  4554       using `i:d` `finite d` setsum_delta'[of d i "(%k. inverse (2 * real (card d)))"] d1 assms(2)
  4555       by (auto simp add: Euclidean_Space.basis_component[of i])}
  4555       by (auto simp add: Euclidean_Space.basis_component[of i])}
  4556   note ** = this
  4556   note ** = this
  4557   show ?thesis apply(rule that[of ?a]) unfolding rel_interior_substd_simplex[OF assms(2)] mem_Collect_eq
  4557   show ?thesis apply(rule that[of ?a]) unfolding rel_interior_substd_simplex[OF assms(2)] mem_Collect_eq
  4558   proof safe fix i assume "i:d" 
  4558   proof safe fix i assume "i:d"
  4559     have "0 < inverse (2 * real (card d))" using d1 by auto
  4559     have "0 < inverse (2 * real (card d))" using d1 by auto
  4560     also have "...=?a $$ i" using **[of i] `i:d` by auto
  4560     also have "...=?a $$ i" using **[of i] `i:d` by auto
  4561     finally show "0 < ?a $$ i" by auto
  4561     finally show "0 < ?a $$ i" by auto
  4562   next have "setsum (op $$ ?a) ?D = setsum (\<lambda>i. inverse (2 * real (card d))) ?D" 
  4562   next have "setsum (op $$ ?a) ?D = setsum (\<lambda>i. inverse (2 * real (card d))) ?D"
  4563       by(rule setsum_cong2, rule **) 
  4563       by(rule setsum_cong2, rule **)
  4564     also have "\<dots> < 1" unfolding setsum_constant real_eq_of_nat divide_real_def[symmetric]
  4564     also have "\<dots> < 1" unfolding setsum_constant real_eq_of_nat divide_real_def[symmetric]
  4565       by (auto simp add:field_simps)
  4565       by (auto simp add:field_simps)
  4566     finally show "setsum (op $$ ?a) ?D < 1" by auto
  4566     finally show "setsum (op $$ ?a) ?D < 1" by auto
  4567   next fix i assume "i<DIM('a)" and "i~:d"
  4567   next fix i assume "i<DIM('a)" and "i~:d"
  4568     have "?a : (span {basis i | i. i : d})" 
  4568     have "?a : (span {basis i | i. i : d})"
  4569       apply (rule span_setsum[of "{basis i |i. i : d}" "(%b. b /\<^sub>R (2 * real (card d)))" "{basis i |i. i : d}"]) 
  4569       apply (rule span_setsum[of "{basis i |i. i : d}" "(%b. b /\<^sub>R (2 * real (card d)))" "{basis i |i. i : d}"])
  4570       using finite_substdbasis[of d] apply blast 
  4570       using finite_substdbasis[of d] apply blast
  4571     proof-
  4571     proof-
  4572       { fix x assume "(x :: 'a::euclidean_space): {basis i |i. i : d}"
  4572       { fix x assume "(x :: 'a::euclidean_space): {basis i |i. i : d}"
  4573         hence "x : span {basis i |i. i : d}" 
  4573         hence "x : span {basis i |i. i : d}"
  4574           using span_superset[of _ "{basis i |i. i : d}"] by auto
  4574           using span_superset[of _ "{basis i |i. i : d}"] by auto
  4575         hence "(x /\<^sub>R (2 * real (card d))) : (span {basis i |i. i : d})"
  4575         hence "(x /\<^sub>R (2 * real (card d))) : (span {basis i |i. i : d})"
  4576           using span_mul[of x "{basis i |i. i : d}" "(inverse (real (card d)) / 2)"] by auto
  4576           using span_mul[of x "{basis i |i. i : d}" "(inverse (real (card d)) / 2)"] by auto
  4577       } thus "\<forall>x\<in>{basis i |i. i \<in> d}. x /\<^sub>R (2 * real (card d)) \<in> span {basis i ::'a |i. i \<in> d}" by auto
  4577       } thus "\<forall>x\<in>{basis i |i. i \<in> d}. x /\<^sub>R (2 * real (card d)) \<in> span {basis i ::'a |i. i \<in> d}" by auto
  4578     qed
  4578     qed
  4580   qed
  4580   qed
  4581 qed
  4581 qed
  4582 
  4582 
  4583 subsection {* Relative interior of convex set *}
  4583 subsection {* Relative interior of convex set *}
  4584 
  4584 
  4585 lemma rel_interior_convex_nonempty_aux: 
  4585 lemma rel_interior_convex_nonempty_aux:
  4586 fixes S :: "('n::euclidean_space) set" 
  4586 fixes S :: "('n::euclidean_space) set"
  4587 assumes "convex S" and "0 : S"
  4587 assumes "convex S" and "0 : S"
  4588 shows "rel_interior S ~= {}"
  4588 shows "rel_interior S ~= {}"
  4589 proof-
  4589 proof-
  4590 { assume "S = {0}" hence ?thesis using rel_interior_sing by auto }
  4590 { assume "S = {0}" hence ?thesis using rel_interior_sing by auto }
  4591 moreover { 
  4591 moreover {
  4592 assume "S ~= {0}"
  4592 assume "S ~= {0}"
  4593 obtain B where B_def: "independent B & B<=S & (S <= span B) & card B = dim S" using basis_exists[of S] by auto
  4593 obtain B where B_def: "independent B & B<=S & (S <= span B) & card B = dim S" using basis_exists[of S] by auto
  4594 hence "B~={}" using B_def assms `S ~= {0}` span_empty by auto
  4594 hence "B~={}" using B_def assms `S ~= {0}` span_empty by auto
  4595 have "insert 0 B <= span B" using subspace_span[of B] subspace_0[of "span B"] span_inc by auto
  4595 have "insert 0 B <= span B" using subspace_span[of B] subspace_0[of "span B"] span_inc by auto
  4596 hence "span (insert 0 B) <= span B" 
  4596 hence "span (insert 0 B) <= span B"
  4597     using span_span[of B] span_mono[of "insert 0 B" "span B"] by blast
  4597     using span_span[of B] span_mono[of "insert 0 B" "span B"] by blast
  4598 hence "convex hull insert 0 B <= span B" 
  4598 hence "convex hull insert 0 B <= span B"
  4599     using convex_hull_subset_span[of "insert 0 B"] by auto
  4599     using convex_hull_subset_span[of "insert 0 B"] by auto
  4600 hence "span (convex hull insert 0 B) <= span B"
  4600 hence "span (convex hull insert 0 B) <= span B"
  4601     using span_span[of B] span_mono[of "convex hull insert 0 B" "span B"] by blast
  4601     using span_span[of B] span_mono[of "convex hull insert 0 B" "span B"] by blast
  4602 hence *: "span (convex hull insert 0 B) = span B" 
  4602 hence *: "span (convex hull insert 0 B) = span B"
  4603     using span_mono[of B "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto
  4603     using span_mono[of B "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto
  4604 hence "span (convex hull insert 0 B) = span S"
  4604 hence "span (convex hull insert 0 B) = span S"
  4605     using B_def span_mono[of B S] span_mono[of S "span B"] span_span[of B] by auto
  4605     using B_def span_mono[of B S] span_mono[of S "span B"] span_span[of B] by auto
  4606 moreover have "0 : affine hull (convex hull insert 0 B)"
  4606 moreover have "0 : affine hull (convex hull insert 0 B)"
  4607     using hull_subset[of "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto
  4607     using hull_subset[of "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto
  4608 ultimately have **: "affine hull (convex hull insert 0 B) = affine hull S"
  4608 ultimately have **: "affine hull (convex hull insert 0 B) = affine hull S"
  4609     using affine_hull_span_0[of "convex hull insert 0 B"] affine_hull_span_0[of "S"] 
  4609     using affine_hull_span_0[of "convex hull insert 0 B"] affine_hull_span_0[of "S"]
  4610     assms  hull_subset[of S] by auto
  4610     assms  hull_subset[of S] by auto
  4611 obtain d and f::"'n=>'n" where fd: "card d = card B & linear f & f ` B = {basis i |i. i : (d :: nat set)} & 
  4611 obtain d and f::"'n=>'n" where fd: "card d = card B & linear f & f ` B = {basis i |i. i : (d :: nat set)} &
  4612        f ` span B = {x. ALL i<DIM('n). i ~: d --> x $$ i = (0::real)} &  inj_on f (span B)" and d:"d\<subseteq>{..<DIM('n)}"
  4612        f ` span B = {x. ALL i<DIM('n). i ~: d --> x $$ i = (0::real)} &  inj_on f (span B)" and d:"d\<subseteq>{..<DIM('n)}"
  4613     using basis_to_substdbasis_subspace_isomorphism[of B,OF _ ] B_def by auto
  4613     using basis_to_substdbasis_subspace_isomorphism[of B,OF _ ] B_def by auto
  4614 hence "bounded_linear f" using linear_conv_bounded_linear by auto
  4614 hence "bounded_linear f" using linear_conv_bounded_linear by auto
  4615 have "d ~={}" using fd B_def `B ~={}` by auto
  4615 have "d ~={}" using fd B_def `B ~={}` by auto
  4616 have "(insert 0 {basis i |i. i : d}) = f ` (insert 0 B)" using fd linear_0 by auto
  4616 have "(insert 0 {basis i |i. i : d}) = f ` (insert 0 B)" using fd linear_0 by auto
  4617 hence "(convex hull (insert 0 {basis i |i. i : d})) = f ` (convex hull (insert 0 B))"
  4617 hence "(convex hull (insert 0 {basis i |i. i : d})) = f ` (convex hull (insert 0 B))"
  4618    using convex_hull_linear_image[of f "(insert 0 {basis i |i. i : d})"] 
  4618    using convex_hull_linear_image[of f "(insert 0 {basis i |i. i : d})"]
  4619    convex_hull_linear_image[of f "(insert 0 B)"] `bounded_linear f` by auto
  4619    convex_hull_linear_image[of f "(insert 0 B)"] `bounded_linear f` by auto
  4620 moreover have "rel_interior (f ` (convex hull insert 0 B)) = 
  4620 moreover have "rel_interior (f ` (convex hull insert 0 B)) =
  4621    f ` rel_interior (convex hull insert 0 B)"
  4621    f ` rel_interior (convex hull insert 0 B)"
  4622    apply (rule  rel_interior_injective_on_span_linear_image[of f "(convex hull insert 0 B)"])
  4622    apply (rule  rel_interior_injective_on_span_linear_image[of f "(convex hull insert 0 B)"])
  4623    using `bounded_linear f` fd * by auto
  4623    using `bounded_linear f` fd * by auto
  4624 ultimately have "rel_interior (convex hull insert 0 B) ~= {}"
  4624 ultimately have "rel_interior (convex hull insert 0 B) ~= {}"
  4625    using rel_interior_substd_simplex_nonempty[OF `d~={}` d] apply auto by blast 
  4625    using rel_interior_substd_simplex_nonempty[OF `d~={}` d] apply auto by blast
  4626 moreover have "convex hull (insert 0 B) <= S"
  4626 moreover have "convex hull (insert 0 B) <= S"
  4627    using B_def assms hull_mono[of "insert 0 B" "S" "convex"] convex_hull_eq by auto
  4627    using B_def assms hull_mono[of "insert 0 B" "S" "convex"] convex_hull_eq by auto
  4628 ultimately have ?thesis using subset_rel_interior[of "convex hull insert 0 B" S] ** by auto
  4628 ultimately have ?thesis using subset_rel_interior[of "convex hull insert 0 B" S] ** by auto
  4629 } ultimately show ?thesis by auto
  4629 } ultimately show ?thesis by auto
  4630 qed
  4630 qed
  4634 assumes "convex S"
  4634 assumes "convex S"
  4635 shows "rel_interior S = {} <-> S = {}"
  4635 shows "rel_interior S = {} <-> S = {}"
  4636 proof-
  4636 proof-
  4637 { assume "S ~= {}" from this obtain a where "a : S" by auto
  4637 { assume "S ~= {}" from this obtain a where "a : S" by auto
  4638   hence "0 : op + (-a) ` S" using assms exI[of "(%x. x:S & -a+x=0)" a] by auto
  4638   hence "0 : op + (-a) ` S" using assms exI[of "(%x. x:S & -a+x=0)" a] by auto
  4639   hence "rel_interior (op + (-a) ` S) ~= {}"  
  4639   hence "rel_interior (op + (-a) ` S) ~= {}"
  4640     using rel_interior_convex_nonempty_aux[of "op + (-a) ` S"] 
  4640     using rel_interior_convex_nonempty_aux[of "op + (-a) ` S"]
  4641           convex_translation[of S "-a"] assms by auto 
  4641           convex_translation[of S "-a"] assms by auto
  4642   hence "rel_interior S ~= {}" using rel_interior_translation by auto
  4642   hence "rel_interior S ~= {}" using rel_interior_translation by auto
  4643 } from this show ?thesis using rel_interior_empty by auto
  4643 } from this show ?thesis using rel_interior_empty by auto
  4644 qed
  4644 qed
  4645 
  4645 
  4646 lemma convex_rel_interior:
  4646 lemma convex_rel_interior:
  4651 { fix "x" "y" "u"
  4651 { fix "x" "y" "u"
  4652   assume assm: "x:rel_interior S" "y:rel_interior S" "0<=u" "(u :: real) <= 1"
  4652   assume assm: "x:rel_interior S" "y:rel_interior S" "0<=u" "(u :: real) <= 1"
  4653   hence "x:S" using rel_interior_subset by auto
  4653   hence "x:S" using rel_interior_subset by auto
  4654   have "x - u *\<^sub>R (x-y) : rel_interior S"
  4654   have "x - u *\<^sub>R (x-y) : rel_interior S"
  4655   proof(cases "0=u")
  4655   proof(cases "0=u")
  4656      case False hence "0<u" using assm by auto 
  4656      case False hence "0<u" using assm by auto
  4657         thus ?thesis
  4657         thus ?thesis
  4658         using assm rel_interior_convex_shrink[of S y x u] assms `x:S` by auto
  4658         using assm rel_interior_convex_shrink[of S y x u] assms `x:S` by auto
  4659      next
  4659      next
  4660      case True thus ?thesis using assm by auto
  4660      case True thus ?thesis using assm by auto
  4661   qed
  4661   qed
  4662   hence "(1-u) *\<^sub>R x + u *\<^sub>R y : rel_interior S" by (simp add: algebra_simps)
  4662   hence "(1-u) *\<^sub>R x + u *\<^sub>R y : rel_interior S" by (simp add: algebra_simps)
  4663 } from this show ?thesis unfolding convex_alt by auto
  4663 } from this show ?thesis unfolding convex_alt by auto
  4664 qed
  4664 qed
  4665 
  4665 
  4666 lemma convex_closure_rel_interior: 
  4666 lemma convex_closure_rel_interior:
  4667 fixes S :: "('n::euclidean_space) set" 
  4667 fixes S :: "('n::euclidean_space) set"
  4668 assumes "convex S"
  4668 assumes "convex S"
  4669 shows "closure(rel_interior S) = closure S"
  4669 shows "closure(rel_interior S) = closure S"
  4670 proof-
  4670 proof-
  4671 have h1: "closure(rel_interior S) <= closure S" 
  4671 have h1: "closure(rel_interior S) <= closure S"
  4672    using closure_mono[of "rel_interior S" S] rel_interior_subset[of S] by auto
  4672    using closure_mono[of "rel_interior S" S] rel_interior_subset[of S] by auto
  4673 { assume "S ~= {}" from this obtain a where a_def: "a : rel_interior S" 
  4673 { assume "S ~= {}" from this obtain a where a_def: "a : rel_interior S"
  4674     using rel_interior_convex_nonempty assms by auto
  4674     using rel_interior_convex_nonempty assms by auto
  4675   { fix x assume x_def: "x : closure S"
  4675   { fix x assume x_def: "x : closure S"
  4676     { assume "x=a" hence "x : closure(rel_interior S)" using a_def unfolding closure_def by auto }
  4676     { assume "x=a" hence "x : closure(rel_interior S)" using a_def unfolding closure_def by auto }
  4677     moreover
  4677     moreover
  4678     { assume "x ~= a"
  4678     { assume "x ~= a"
  4679        { fix e :: real assume e_def: "e>0" 
  4679        { fix e :: real assume e_def: "e>0"
  4680          def e1 == "min 1 (e/norm (x - a))" hence e1_def: "e1>0 & e1<=1 & e1*norm(x-a)<=e"
  4680          def e1 == "min 1 (e/norm (x - a))" hence e1_def: "e1>0 & e1<=1 & e1*norm(x-a)<=e"
  4681             using `x ~= a` `e>0` divide_pos_pos[of e] le_divide_eq[of e1 e "norm(x-a)"] by simp 
  4681             using `x ~= a` `e>0` divide_pos_pos[of e] le_divide_eq[of e1 e "norm(x-a)"] by simp
  4682          hence *: "x - e1 *\<^sub>R (x - a) : rel_interior S"
  4682          hence *: "x - e1 *\<^sub>R (x - a) : rel_interior S"
  4683             using rel_interior_closure_convex_shrink[of S a x e1] assms x_def a_def e1_def by auto
  4683             using rel_interior_closure_convex_shrink[of S a x e1] assms x_def a_def e1_def by auto
  4684          have "EX y. y:rel_interior S & y ~= x & (dist y x) <= e"
  4684          have "EX y. y:rel_interior S & y ~= x & (dist y x) <= e"
  4685             apply (rule_tac x="x - e1 *\<^sub>R (x - a)" in exI)
  4685             apply (rule_tac x="x - e1 *\<^sub>R (x - a)" in exI)
  4686             using * e1_def dist_norm[of "x - e1 *\<^sub>R (x - a)" x] `x ~= a` by simp
  4686             using * e1_def dist_norm[of "x - e1 *\<^sub>R (x - a)" x] `x ~= a` by simp
  4687       } hence "x islimpt rel_interior S" unfolding islimpt_approachable_le by auto 
  4687       } hence "x islimpt rel_interior S" unfolding islimpt_approachable_le by auto
  4688       hence "x : closure(rel_interior S)" unfolding closure_def by auto 
  4688       hence "x : closure(rel_interior S)" unfolding closure_def by auto
  4689     } ultimately have "x : closure(rel_interior S)" by auto
  4689     } ultimately have "x : closure(rel_interior S)" by auto
  4690   } hence ?thesis using h1 by auto
  4690   } hence ?thesis using h1 by auto
  4691 }
  4691 }
  4692 moreover
  4692 moreover
  4693 { assume "S = {}" hence "rel_interior S = {}" using rel_interior_empty by auto
  4693 { assume "S = {}" hence "rel_interior S = {}" using rel_interior_empty by auto
  4694   hence "closure(rel_interior S) = {}" using closure_empty by auto 
  4694   hence "closure(rel_interior S) = {}" using closure_empty by auto
  4695   hence ?thesis using `S={}` by auto 
  4695   hence ?thesis using `S={}` by auto
  4696 } ultimately show ?thesis by blast
  4696 } ultimately show ?thesis by blast
  4697 qed
  4697 qed
  4698 
  4698 
  4699 lemma rel_interior_same_affine_hull:
  4699 lemma rel_interior_same_affine_hull:
  4700   fixes S :: "('n::euclidean_space) set"
  4700   fixes S :: "('n::euclidean_space) set"
  4701   assumes "convex S"
  4701   assumes "convex S"
  4702   shows "affine hull (rel_interior S) = affine hull S"
  4702   shows "affine hull (rel_interior S) = affine hull S"
  4703 by (metis assms closure_same_affine_hull convex_closure_rel_interior)
  4703 by (metis assms closure_same_affine_hull convex_closure_rel_interior)
  4704 
  4704 
  4705 lemma rel_interior_aff_dim: 
  4705 lemma rel_interior_aff_dim:
  4706   fixes S :: "('n::euclidean_space) set"
  4706   fixes S :: "('n::euclidean_space) set"
  4707   assumes "convex S"
  4707   assumes "convex S"
  4708   shows "aff_dim (rel_interior S) = aff_dim S"
  4708   shows "aff_dim (rel_interior S) = aff_dim S"
  4709 by (metis aff_dim_affine_hull2 assms rel_interior_same_affine_hull)
  4709 by (metis aff_dim_affine_hull2 assms rel_interior_same_affine_hull)
  4710 
  4710 
  4739 moreover have "0<e" using e_def assms divide_pos_pos[of a "a+b"] by auto
  4739 moreover have "0<e" using e_def assms divide_pos_pos[of a "a+b"] by auto
  4740 moreover have "e<=1" using e_def assms by auto
  4740 moreover have "e<=1" using e_def assms by auto
  4741 ultimately show ?thesis using that[of e] by auto
  4741 ultimately show ?thesis using that[of e] by auto
  4742 qed
  4742 qed
  4743 
  4743 
  4744 lemma convex_rel_interior_closure: 
  4744 lemma convex_rel_interior_closure:
  4745   fixes S :: "('n::euclidean_space) set" 
  4745   fixes S :: "('n::euclidean_space) set"
  4746   assumes "convex S"
  4746   assumes "convex S"
  4747   shows "rel_interior (closure S) = rel_interior S"
  4747   shows "rel_interior (closure S) = rel_interior S"
  4748 proof-
  4748 proof-
  4749 { assume "S={}" hence ?thesis using assms rel_interior_convex_nonempty by auto }
  4749 { assume "S={}" hence ?thesis using assms rel_interior_convex_nonempty by auto }
  4750 moreover
  4750 moreover
  4751 { assume "S ~= {}"
  4751 { assume "S ~= {}"
  4752   have "rel_interior (closure S) >= rel_interior S" 
  4752   have "rel_interior (closure S) >= rel_interior S"
  4753     using subset_rel_interior[of S "closure S"] closure_same_affine_hull closure_subset by auto
  4753     using subset_rel_interior[of S "closure S"] closure_same_affine_hull closure_subset by auto
  4754   moreover
  4754   moreover
  4755   { fix z assume z_def: "z : rel_interior (closure S)"
  4755   { fix z assume z_def: "z : rel_interior (closure S)"
  4756     obtain x where x_def: "x : rel_interior S" 
  4756     obtain x where x_def: "x : rel_interior S"
  4757       using `S ~= {}` assms rel_interior_convex_nonempty by auto  
  4757       using `S ~= {}` assms rel_interior_convex_nonempty by auto
  4758     { assume "x=z" hence "z : rel_interior S" using x_def by auto }
  4758     { assume "x=z" hence "z : rel_interior S" using x_def by auto }
  4759     moreover
  4759     moreover
  4760     { assume "x ~= z"
  4760     { assume "x ~= z"
  4761       obtain e where e_def: "e > 0 & cball z e Int affine hull closure S <= closure S" 
  4761       obtain e where e_def: "e > 0 & cball z e Int affine hull closure S <= closure S"
  4762         using z_def rel_interior_cball[of "closure S"] by auto
  4762         using z_def rel_interior_cball[of "closure S"] by auto
  4763       hence *: "0 < e/norm(z-x)" using e_def `x ~= z` divide_pos_pos[of e "norm(z-x)"] by auto 
  4763       hence *: "0 < e/norm(z-x)" using e_def `x ~= z` divide_pos_pos[of e "norm(z-x)"] by auto
  4764       def y == "z + (e/norm(z-x)) *\<^sub>R (z-x)"
  4764       def y == "z + (e/norm(z-x)) *\<^sub>R (z-x)"
  4765       have yball: "y : cball z e"
  4765       have yball: "y : cball z e"
  4766         using mem_cball y_def dist_norm[of z y] e_def by auto 
  4766         using mem_cball y_def dist_norm[of z y] e_def by auto
  4767       have "x : affine hull closure S" 
  4767       have "x : affine hull closure S"
  4768         using x_def rel_interior_subset_closure hull_inc[of x "closure S"] by auto
  4768         using x_def rel_interior_subset_closure hull_inc[of x "closure S"] by auto
  4769       moreover have "z : affine hull closure S" 
  4769       moreover have "z : affine hull closure S"
  4770         using z_def rel_interior_subset hull_subset[of "closure S"] by auto
  4770         using z_def rel_interior_subset hull_subset[of "closure S"] by auto
  4771       ultimately have "y : affine hull closure S" 
  4771       ultimately have "y : affine hull closure S"
  4772         using y_def affine_affine_hull[of "closure S"] 
  4772         using y_def affine_affine_hull[of "closure S"]
  4773           mem_affine_3_minus [of "affine hull closure S" z z x "e/norm(z-x)"] by auto
  4773           mem_affine_3_minus [of "affine hull closure S" z z x "e/norm(z-x)"] by auto
  4774       hence "y : closure S" using e_def yball by auto
  4774       hence "y : closure S" using e_def yball by auto
  4775       have "(1+(e/norm(z-x))) *\<^sub>R z = (e/norm(z-x)) *\<^sub>R x + y"
  4775       have "(1+(e/norm(z-x))) *\<^sub>R z = (e/norm(z-x)) *\<^sub>R x + y"
  4776         using y_def by (simp add: algebra_simps) 
  4776         using y_def by (simp add: algebra_simps)
  4777       from this obtain e1 where "0 < e1 & e1 <= 1 & z = y - e1 *\<^sub>R (y - x)"
  4777       from this obtain e1 where "0 < e1 & e1 <= 1 & z = y - e1 *\<^sub>R (y - x)"
  4778         using * convex_rel_interior_closure_aux[of "e / norm (z - x)" 1 z x y] 
  4778         using * convex_rel_interior_closure_aux[of "e / norm (z - x)" 1 z x y]
  4779           by (auto simp add: algebra_simps)
  4779           by (auto simp add: algebra_simps)
  4780       hence "z : rel_interior S" 
  4780       hence "z : rel_interior S"
  4781         using rel_interior_closure_convex_shrink assms x_def `y : closure S` by auto
  4781         using rel_interior_closure_convex_shrink assms x_def `y : closure S` by auto
  4782     } ultimately have "z : rel_interior S" by auto
  4782     } ultimately have "z : rel_interior S" by auto
  4783   } ultimately have ?thesis by auto
  4783   } ultimately have ?thesis by auto
  4784 } ultimately show ?thesis by blast
  4784 } ultimately show ?thesis by blast
  4785 qed
  4785 qed
  4786 
  4786 
  4787 lemma convex_interior_closure: 
  4787 lemma convex_interior_closure:
  4788 fixes S :: "('n::euclidean_space) set" 
  4788 fixes S :: "('n::euclidean_space) set"
  4789 assumes "convex S"
  4789 assumes "convex S"
  4790 shows "interior (closure S) = interior S"
  4790 shows "interior (closure S) = interior S"
  4791 using closure_aff_dim[of S] interior_rel_interior_gen[of S] interior_rel_interior_gen[of "closure S"] 
  4791 using closure_aff_dim[of S] interior_rel_interior_gen[of S] interior_rel_interior_gen[of "closure S"]
  4792       convex_rel_interior_closure[of S] assms by auto
  4792       convex_rel_interior_closure[of S] assms by auto
  4793 
  4793 
  4794 lemma closure_eq_rel_interior_eq:
  4794 lemma closure_eq_rel_interior_eq:
  4795 fixes S1 S2 ::  "('n::euclidean_space) set" 
  4795 fixes S1 S2 ::  "('n::euclidean_space) set"
  4796 assumes "convex S1" "convex S2"
  4796 assumes "convex S1" "convex S2"
  4797 shows "(closure S1 = closure S2) <-> (rel_interior S1 = rel_interior S2)"
  4797 shows "(closure S1 = closure S2) <-> (rel_interior S1 = rel_interior S2)"
  4798  by (metis convex_rel_interior_closure convex_closure_rel_interior assms)
  4798  by (metis convex_rel_interior_closure convex_closure_rel_interior assms)
  4799 
  4799 
  4800 
  4800 
  4801 lemma closure_eq_between:
  4801 lemma closure_eq_between:
  4802 fixes S1 S2 ::  "('n::euclidean_space) set" 
  4802 fixes S1 S2 ::  "('n::euclidean_space) set"
  4803 assumes "convex S1" "convex S2"
  4803 assumes "convex S1" "convex S2"
  4804 shows "(closure S1 = closure S2) <-> 
  4804 shows "(closure S1 = closure S2) <->
  4805       ((rel_interior S1 <= S2) & (S2 <= closure S1))" (is "?A <-> ?B")
  4805       ((rel_interior S1 <= S2) & (S2 <= closure S1))" (is "?A <-> ?B")
  4806 proof-
  4806 proof-
  4807 have "?A --> ?B" by (metis assms closure_subset convex_rel_interior_closure rel_interior_subset)
  4807 have "?A --> ?B" by (metis assms closure_subset convex_rel_interior_closure rel_interior_subset)
  4808 moreover have "?B --> (closure S1 <= closure S2)" 
  4808 moreover have "?B --> (closure S1 <= closure S2)"
  4809      by (metis assms(1) convex_closure_rel_interior closure_mono)
  4809      by (metis assms(1) convex_closure_rel_interior closure_mono)
  4810 moreover have "?B --> (closure S1 >= closure S2)" by (metis closed_closure closure_minimal)
  4810 moreover have "?B --> (closure S1 >= closure S2)" by (metis closed_closure closure_minimal)
  4811 ultimately show ?thesis by blast
  4811 ultimately show ?thesis by blast
  4812 qed
  4812 qed
  4813 
  4813 
  4814 lemma open_inter_closure_rel_interior:
  4814 lemma open_inter_closure_rel_interior:
  4815 fixes S A ::  "('n::euclidean_space) set" 
  4815 fixes S A ::  "('n::euclidean_space) set"
  4816 assumes "convex S" "open A"
  4816 assumes "convex S" "open A"
  4817 shows "((A Int closure S) = {}) <-> ((A Int rel_interior S) = {})"
  4817 shows "((A Int closure S) = {}) <-> ((A Int rel_interior S) = {})"
  4818 by (metis assms convex_closure_rel_interior open_inter_closure_eq_empty) 
  4818 by (metis assms convex_closure_rel_interior open_inter_closure_eq_empty)
  4819 
  4819 
  4820 definition "rel_frontier S = closure S - rel_interior S"
  4820 definition "rel_frontier S = closure S - rel_interior S"
  4821 
  4821 
  4822 lemma closed_affine_hull: "closed (affine hull ((S :: ('n::euclidean_space) set)))"
  4822 lemma closed_affine_hull: "closed (affine hull ((S :: ('n::euclidean_space) set)))"
  4823 by (metis affine_affine_hull affine_closed)
  4823 by (metis affine_affine_hull affine_closed)
  4824 
  4824 
  4825 lemma closed_rel_frontier: "closed(rel_frontier (S :: ('n::euclidean_space) set))"
  4825 lemma closed_rel_frontier: "closed(rel_frontier (S :: ('n::euclidean_space) set))"
  4826 proof-
  4826 proof-
  4827 have *: "closedin (subtopology euclidean (affine hull S)) (closure S - rel_interior S)" 
  4827 have *: "closedin (subtopology euclidean (affine hull S)) (closure S - rel_interior S)"
  4828 apply (rule closedin_diff[of "subtopology euclidean (affine hull S)""closure S" "rel_interior S"])  using closed_closedin_trans[of "affine hull S" "closure S"] closed_affine_hull[of S] 
  4828 apply (rule closedin_diff[of "subtopology euclidean (affine hull S)""closure S" "rel_interior S"])  using closed_closedin_trans[of "affine hull S" "closure S"] closed_affine_hull[of S]
  4829   closure_affine_hull[of S] opein_rel_interior[of S] by auto 
  4829   closure_affine_hull[of S] opein_rel_interior[of S] by auto
  4830 show ?thesis apply (rule closedin_closed_trans[of "affine hull S" "rel_frontier S"]) 
  4830 show ?thesis apply (rule closedin_closed_trans[of "affine hull S" "rel_frontier S"])
  4831   unfolding rel_frontier_def using * closed_affine_hull by auto 
  4831   unfolding rel_frontier_def using * closed_affine_hull by auto
  4832 qed
  4832 qed
  4833  
  4833 
  4834 
  4834 
  4835 lemma convex_rel_frontier_aff_dim:
  4835 lemma convex_rel_frontier_aff_dim:
  4836 fixes S1 S2 ::  "('n::euclidean_space) set" 
  4836 fixes S1 S2 ::  "('n::euclidean_space) set"
  4837 assumes "convex S1" "convex S2" "S2 ~= {}"
  4837 assumes "convex S1" "convex S2" "S2 ~= {}"
  4838 assumes "S1 <= rel_frontier S2"
  4838 assumes "S1 <= rel_frontier S2"
  4839 shows "aff_dim S1 < aff_dim S2" 
  4839 shows "aff_dim S1 < aff_dim S2"
  4840 proof-
  4840 proof-
  4841 have "S1 <= closure S2" using assms unfolding rel_frontier_def by auto
  4841 have "S1 <= closure S2" using assms unfolding rel_frontier_def by auto
  4842 hence *: "affine hull S1 <= affine hull S2" 
  4842 hence *: "affine hull S1 <= affine hull S2"
  4843    using hull_mono[of "S1" "closure S2"] closure_same_affine_hull[of S2] by auto
  4843    using hull_mono[of "S1" "closure S2"] closure_same_affine_hull[of S2] by auto
  4844 hence "aff_dim S1 <= aff_dim S2" using * aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2] 
  4844 hence "aff_dim S1 <= aff_dim S2" using * aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2]
  4845     aff_dim_subset[of "affine hull S1" "affine hull S2"] by auto
  4845     aff_dim_subset[of "affine hull S1" "affine hull S2"] by auto
  4846 moreover
  4846 moreover
  4847 { assume eq: "aff_dim S1 = aff_dim S2"
  4847 { assume eq: "aff_dim S1 = aff_dim S2"
  4848   hence "S1 ~= {}" using aff_dim_empty[of S1] aff_dim_empty[of S2] `S2 ~= {}` by auto
  4848   hence "S1 ~= {}" using aff_dim_empty[of S1] aff_dim_empty[of S2] `S2 ~= {}` by auto
  4849   have **: "affine hull S1 = affine hull S2" 
  4849   have **: "affine hull S1 = affine hull S2"
  4850      apply (rule affine_dim_equal) using * affine_affine_hull apply auto
  4850      apply (rule affine_dim_equal) using * affine_affine_hull apply auto
  4851      using `S1 ~= {}` hull_subset[of S1] apply auto
  4851      using `S1 ~= {}` hull_subset[of S1] apply auto
  4852      using eq aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2] by auto
  4852      using eq aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2] by auto
  4853   obtain a where a_def: "a : rel_interior S1"
  4853   obtain a where a_def: "a : rel_interior S1"
  4854      using  `S1 ~= {}` rel_interior_convex_nonempty assms by auto
  4854      using  `S1 ~= {}` rel_interior_convex_nonempty assms by auto
  4855   obtain T where T_def: "open T & a : T Int S1 & T Int affine hull S1 <= S1"
  4855   obtain T where T_def: "open T & a : T Int S1 & T Int affine hull S1 <= S1"
  4856      using mem_rel_interior[of a S1] a_def by auto
  4856      using mem_rel_interior[of a S1] a_def by auto
  4857   hence "a : T Int closure S2" using a_def assms unfolding rel_frontier_def by auto
  4857   hence "a : T Int closure S2" using a_def assms unfolding rel_frontier_def by auto
  4858   from this obtain b where b_def: "b : T Int rel_interior S2" 
  4858   from this obtain b where b_def: "b : T Int rel_interior S2"
  4859      using open_inter_closure_rel_interior[of S2 T] assms T_def by auto
  4859      using open_inter_closure_rel_interior[of S2 T] assms T_def by auto
  4860   hence "b : affine hull S1" using rel_interior_subset hull_subset[of S2] ** by auto
  4860   hence "b : affine hull S1" using rel_interior_subset hull_subset[of S2] ** by auto
  4861   hence "b : S1" using T_def b_def by auto
  4861   hence "b : S1" using T_def b_def by auto
  4862   hence False using b_def assms unfolding rel_frontier_def by auto
  4862   hence False using b_def assms unfolding rel_frontier_def by auto
  4863 } ultimately show ?thesis using less_le by auto
  4863 } ultimately show ?thesis using less_le by auto
  4864 qed
  4864 qed
  4865 
  4865 
  4866 
  4866 
  4867 lemma convex_rel_interior_if:
  4867 lemma convex_rel_interior_if:
  4868 fixes S ::  "('n::euclidean_space) set" 
  4868 fixes S ::  "('n::euclidean_space) set"
  4869 assumes "convex S"
  4869 assumes "convex S"
  4870 assumes "z : rel_interior S"
  4870 assumes "z : rel_interior S"
  4871 shows "(!x:affine hull S. EX m. m>1 & (!e. (e>1 & e<=m) --> (1-e)*\<^sub>R x+ e *\<^sub>R z : S ))"
  4871 shows "(!x:affine hull S. EX m. m>1 & (!e. (e>1 & e<=m) --> (1-e)*\<^sub>R x+ e *\<^sub>R z : S ))"
  4872 proof-
  4872 proof-
  4873 obtain e1 where e1_def: "e1>0 & cball z e1 Int affine hull S <= S" 
  4873 obtain e1 where e1_def: "e1>0 & cball z e1 Int affine hull S <= S"
  4874     using mem_rel_interior_cball[of z S] assms by auto
  4874     using mem_rel_interior_cball[of z S] assms by auto
  4875 { fix x assume x_def: "x:affine hull S"
  4875 { fix x assume x_def: "x:affine hull S"
  4876   { assume "x ~= z"
  4876   { assume "x ~= z"
  4877     def m == "1+e1/norm(x-z)" 
  4877     def m == "1+e1/norm(x-z)"
  4878     hence "m>1" using e1_def `x ~= z` divide_pos_pos[of e1 "norm (x - z)"] by auto 
  4878     hence "m>1" using e1_def `x ~= z` divide_pos_pos[of e1 "norm (x - z)"] by auto
  4879     { fix e assume e_def: "e>1 & e<=m"
  4879     { fix e assume e_def: "e>1 & e<=m"
  4880       have "z : affine hull S" using assms rel_interior_subset hull_subset[of S] by auto
  4880       have "z : affine hull S" using assms rel_interior_subset hull_subset[of S] by auto
  4881       hence *: "(1-e)*\<^sub>R x+ e *\<^sub>R z : affine hull S"
  4881       hence *: "(1-e)*\<^sub>R x+ e *\<^sub>R z : affine hull S"
  4882          using mem_affine[of "affine hull S" x z "(1-e)" e] affine_affine_hull[of S] x_def by auto
  4882          using mem_affine[of "affine hull S" x z "(1-e)" e] affine_affine_hull[of S] x_def by auto
  4883       have "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) = norm ((e - 1) *\<^sub>R (x-z))" by (simp add: algebra_simps)
  4883       have "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) = norm ((e - 1) *\<^sub>R (x-z))" by (simp add: algebra_simps)
  4884       also have "...= (e - 1) * norm(x-z)" using norm_scaleR e_def by auto  
  4884       also have "...= (e - 1) * norm(x-z)" using norm_scaleR e_def by auto
  4885       also have "...<=(m - 1) * norm(x-z)" using e_def mult_right_mono[of _ _ "norm(x-z)"] by auto
  4885       also have "...<=(m - 1) * norm(x-z)" using e_def mult_right_mono[of _ _ "norm(x-z)"] by auto
  4886       also have "...= (e1 / norm (x - z)) * norm (x - z)" using m_def by auto
  4886       also have "...= (e1 / norm (x - z)) * norm (x - z)" using m_def by auto
  4887       also have "...=e1" using `x ~= z` e1_def by simp
  4887       also have "...=e1" using `x ~= z` e1_def by simp
  4888       finally have **: "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) <= e1" by auto
  4888       finally have **: "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) <= e1" by auto
  4889       have "(1-e)*\<^sub>R x+ e *\<^sub>R z : cball z e1"
  4889       have "(1-e)*\<^sub>R x+ e *\<^sub>R z : cball z e1"
  4897       hence "(1-e)*\<^sub>R x+ e *\<^sub>R z : S"
  4897       hence "(1-e)*\<^sub>R x+ e *\<^sub>R z : S"
  4898         using e1_def x_def `x=z` by (auto simp add: algebra_simps)
  4898         using e1_def x_def `x=z` by (auto simp add: algebra_simps)
  4899       hence "(1-e)*\<^sub>R x+ e *\<^sub>R z : S" using e_def by auto
  4899       hence "(1-e)*\<^sub>R x+ e *\<^sub>R z : S" using e_def by auto
  4900     } hence "EX m. m>1 & (!e. (e>1 & e<=m) --> (1-e)*\<^sub>R x+ e *\<^sub>R z : S )" using `m>1` by auto
  4900     } hence "EX m. m>1 & (!e. (e>1 & e<=m) --> (1-e)*\<^sub>R x+ e *\<^sub>R z : S )" using `m>1` by auto
  4901   } ultimately have "EX m. m>1 & (!e. (e>1 & e<=m) --> (1-e)*\<^sub>R x+ e *\<^sub>R z : S )" by auto
  4901   } ultimately have "EX m. m>1 & (!e. (e>1 & e<=m) --> (1-e)*\<^sub>R x+ e *\<^sub>R z : S )" by auto
  4902 } from this show ?thesis by auto 
  4902 } from this show ?thesis by auto
  4903 qed
  4903 qed
  4904 
  4904 
  4905 lemma convex_rel_interior_if2:
  4905 lemma convex_rel_interior_if2:
  4906 fixes S ::  "('n::euclidean_space) set" 
  4906 fixes S ::  "('n::euclidean_space) set"
  4907 assumes "convex S"
  4907 assumes "convex S"
  4908 assumes "z : rel_interior S"
  4908 assumes "z : rel_interior S"
  4909 shows "(!x:affine hull S. EX e. e>1 & (1-e)*\<^sub>R x+ e *\<^sub>R z : S)"
  4909 shows "(!x:affine hull S. EX e. e>1 & (1-e)*\<^sub>R x+ e *\<^sub>R z : S)"
  4910 using convex_rel_interior_if[of S z] assms by auto
  4910 using convex_rel_interior_if[of S z] assms by auto
  4911 
  4911 
  4912 lemma convex_rel_interior_only_if:
  4912 lemma convex_rel_interior_only_if:
  4913 fixes S ::  "('n::euclidean_space) set" 
  4913 fixes S ::  "('n::euclidean_space) set"
  4914 assumes "convex S" "S ~= {}"
  4914 assumes "convex S" "S ~= {}"
  4915 assumes "(!x:S. EX e. e>1 & (1-e)*\<^sub>R x+ e *\<^sub>R z : S)"
  4915 assumes "(!x:S. EX e. e>1 & (1-e)*\<^sub>R x+ e *\<^sub>R z : S)"
  4916 shows "z : rel_interior S" 
  4916 shows "z : rel_interior S"
  4917 proof-
  4917 proof-
  4918 obtain x where x_def: "x : rel_interior S" using rel_interior_convex_nonempty assms by auto
  4918 obtain x where x_def: "x : rel_interior S" using rel_interior_convex_nonempty assms by auto
  4919 hence "x:S" using rel_interior_subset by auto
  4919 hence "x:S" using rel_interior_subset by auto
  4920 from this obtain e where e_def: "e>1 & (1 - e) *\<^sub>R x + e *\<^sub>R z : S" using assms by auto
  4920 from this obtain e where e_def: "e>1 & (1 - e) *\<^sub>R x + e *\<^sub>R z : S" using assms by auto
  4921 def y == "(1 - e) *\<^sub>R x + e *\<^sub>R z" hence "y:S" using e_def by auto
  4921 def y == "(1 - e) *\<^sub>R x + e *\<^sub>R z" hence "y:S" using e_def by auto
  4922 def e1 == "1/e" hence "0<e1 & e1<1" using e_def by auto
  4922 def e1 == "1/e" hence "0<e1 & e1<1" using e_def by auto
  4923 hence "z=y-(1-e1)*\<^sub>R (y-x)" using e1_def y_def by (auto simp add: algebra_simps)
  4923 hence "z=y-(1-e1)*\<^sub>R (y-x)" using e1_def y_def by (auto simp add: algebra_simps)
  4924 from this show ?thesis 
  4924 from this show ?thesis
  4925     using rel_interior_convex_shrink[of S x y "1-e1"] `0<e1 & e1<1` `y:S` x_def assms by auto
  4925     using rel_interior_convex_shrink[of S x y "1-e1"] `0<e1 & e1<1` `y:S` x_def assms by auto
  4926 qed
  4926 qed
  4927 
  4927 
  4928 lemma convex_rel_interior_iff:
  4928 lemma convex_rel_interior_iff:
  4929 fixes S ::  "('n::euclidean_space) set" 
  4929 fixes S ::  "('n::euclidean_space) set"
  4930 assumes "convex S" "S ~= {}"
  4930 assumes "convex S" "S ~= {}"
  4931 shows "z : rel_interior S <-> (!x:S. EX e. e>1 & (1-e)*\<^sub>R x+ e *\<^sub>R z : S)"
  4931 shows "z : rel_interior S <-> (!x:S. EX e. e>1 & (1-e)*\<^sub>R x+ e *\<^sub>R z : S)"
  4932 using assms hull_subset[of S "affine"] 
  4932 using assms hull_subset[of S "affine"]
  4933       convex_rel_interior_if[of S z] convex_rel_interior_only_if[of S z] by auto
  4933       convex_rel_interior_if[of S z] convex_rel_interior_only_if[of S z] by auto
  4934 
  4934 
  4935 lemma convex_rel_interior_iff2:
  4935 lemma convex_rel_interior_iff2:
  4936 fixes S ::  "('n::euclidean_space) set" 
  4936 fixes S ::  "('n::euclidean_space) set"
  4937 assumes "convex S" "S ~= {}"
  4937 assumes "convex S" "S ~= {}"
  4938 shows "z : rel_interior S <-> (!x:affine hull S. EX e. e>1 & (1-e)*\<^sub>R x+ e *\<^sub>R z : S)"
  4938 shows "z : rel_interior S <-> (!x:affine hull S. EX e. e>1 & (1-e)*\<^sub>R x+ e *\<^sub>R z : S)"
  4939 using assms hull_subset[of S] 
  4939 using assms hull_subset[of S]
  4940       convex_rel_interior_if2[of S z] convex_rel_interior_only_if[of S z] by auto
  4940       convex_rel_interior_if2[of S z] convex_rel_interior_only_if[of S z] by auto
  4941 
  4941 
  4942 
  4942 
  4943 lemma convex_interior_iff:
  4943 lemma convex_interior_iff:
  4944 fixes S ::  "('n::euclidean_space) set" 
  4944 fixes S ::  "('n::euclidean_space) set"
  4945 assumes "convex S"
  4945 assumes "convex S"
  4946 shows "z : interior S <-> (!x. EX e. e>0 & z+ e *\<^sub>R x : S)"
  4946 shows "z : interior S <-> (!x. EX e. e>0 & z+ e *\<^sub>R x : S)"
  4947 proof-
  4947 proof-
  4948 { assume a: "~(aff_dim S = int DIM('n))"
  4948 { assume a: "~(aff_dim S = int DIM('n))"
  4949   { assume "z : interior S"
  4949   { assume "z : interior S"
  4959          hence x2: "x2 : affine hull S" using e2_def hull_subset[of S] by auto
  4959          hence x2: "x2 : affine hull S" using e2_def hull_subset[of S] by auto
  4960       have *: "e1/(e1+e2) + e2/(e1+e2) = 1" using add_divide_distrib[of e1 e2 "e1+e2"] e1_def e2_def by simp
  4960       have *: "e1/(e1+e2) + e2/(e1+e2) = 1" using add_divide_distrib[of e1 e2 "e1+e2"] e1_def e2_def by simp
  4961       hence "z = (e2/(e1+e2)) *\<^sub>R x1 + (e1/(e1+e2)) *\<^sub>R x2"
  4961       hence "z = (e2/(e1+e2)) *\<^sub>R x1 + (e1/(e1+e2)) *\<^sub>R x2"
  4962          using x1_def x2_def apply (auto simp add: algebra_simps)
  4962          using x1_def x2_def apply (auto simp add: algebra_simps)
  4963          using scaleR_left_distrib[of "e1/(e1+e2)" "e2/(e1+e2)" z] by auto
  4963          using scaleR_left_distrib[of "e1/(e1+e2)" "e2/(e1+e2)" z] by auto
  4964       hence z: "z : affine hull S" 
  4964       hence z: "z : affine hull S"
  4965          using mem_affine[of "affine hull S" x1 x2 "e2/(e1+e2)" "e1/(e1+e2)"]  
  4965          using mem_affine[of "affine hull S" x1 x2 "e2/(e1+e2)" "e1/(e1+e2)"]
  4966          x1 x2 affine_affine_hull[of S] * by auto
  4966          x1 x2 affine_affine_hull[of S] * by auto
  4967       have "x1-x2 = (e1+e2) *\<^sub>R (x-z)"
  4967       have "x1-x2 = (e1+e2) *\<^sub>R (x-z)"
  4968          using x1_def x2_def by (auto simp add: algebra_simps)
  4968          using x1_def x2_def by (auto simp add: algebra_simps)
  4969       hence "x=z+(1/(e1+e2)) *\<^sub>R (x1-x2)" using e1_def e2_def by simp
  4969       hence "x=z+(1/(e1+e2)) *\<^sub>R (x1-x2)" using e1_def e2_def by simp
  4970       hence "x : affine hull S" using mem_affine_3_minus[of "affine hull S" z x1 x2 "1/(e1+e2)"] 
  4970       hence "x : affine hull S" using mem_affine_3_minus[of "affine hull S" z x1 x2 "1/(e1+e2)"]
  4971           x1 x2 z affine_affine_hull[of S] by auto
  4971           x1 x2 z affine_affine_hull[of S] by auto
  4972     } hence "affine hull S = UNIV" by auto
  4972     } hence "affine hull S = UNIV" by auto
  4973     hence "aff_dim S = int DIM('n)" using aff_dim_affine_hull[of S] by (simp add: aff_dim_univ)
  4973     hence "aff_dim S = int DIM('n)" using aff_dim_affine_hull[of S] by (simp add: aff_dim_univ)
  4974     hence False using a by auto
  4974     hence False using a by auto
  4975   } ultimately have ?thesis by auto
  4975   } ultimately have ?thesis by auto
  5005 qed
  5005 qed
  5006 
  5006 
  5007 subsubsection {* Relative interior and closure under common operations *}
  5007 subsubsection {* Relative interior and closure under common operations *}
  5008 
  5008 
  5009 lemma rel_interior_inter_aux: "Inter {rel_interior S |S. S : I} <= Inter I"
  5009 lemma rel_interior_inter_aux: "Inter {rel_interior S |S. S : I} <= Inter I"
  5010 proof- 
  5010 proof-
  5011 { fix y assume "y : Inter {rel_interior S |S. S : I}"
  5011 { fix y assume "y : Inter {rel_interior S |S. S : I}"
  5012   hence y_def: "!S : I. y : rel_interior S" by auto
  5012   hence y_def: "!S : I. y : rel_interior S" by auto
  5013   { fix S assume "S : I" hence "y : S" using rel_interior_subset y_def by auto }
  5013   { fix S assume "S : I" hence "y : S" using rel_interior_subset y_def by auto }
  5014   hence "y : Inter I" by auto
  5014   hence "y : Inter I" by auto
  5015 } thus ?thesis by auto
  5015 } thus ?thesis by auto
  5016 qed
  5016 qed
  5017 
  5017 
  5018 lemma closure_inter: "closure (Inter I) <= Inter {closure S |S. S : I}"
  5018 lemma closure_inter: "closure (Inter I) <= Inter {closure S |S. S : I}"
  5019 proof- 
  5019 proof-
  5020 { fix y assume "y : Inter I" hence y_def: "!S : I. y : S" by auto
  5020 { fix y assume "y : Inter I" hence y_def: "!S : I. y : S" by auto
  5021   { fix S assume "S : I" hence "y : closure S" using closure_subset y_def by auto }
  5021   { fix S assume "S : I" hence "y : closure S" using closure_subset y_def by auto }
  5022   hence "y : Inter {closure S |S. S : I}" by auto
  5022   hence "y : Inter {closure S |S. S : I}" by auto
  5023 } hence "Inter I <= Inter {closure S |S. S : I}" by auto
  5023 } hence "Inter I <= Inter {closure S |S. S : I}" by auto
  5024 moreover have "closed (Inter {closure S |S. S : I})"
  5024 moreover have "closed (Inter {closure S |S. S : I})"
  5025   unfolding closed_Inter closed_closure by auto
  5025   unfolding closed_Inter closed_closure by auto
  5026 ultimately show ?thesis using closure_hull[of "Inter I"]
  5026 ultimately show ?thesis using closure_hull[of "Inter I"]
  5027   hull_minimal[of "Inter I" "Inter {closure S |S. S : I}" "closed"] by auto
  5027   hull_minimal[of "Inter I" "Inter {closure S |S. S : I}" "closed"] by auto
  5028 qed
  5028 qed
  5029 
  5029 
  5030 lemma convex_closure_rel_interior_inter: 
  5030 lemma convex_closure_rel_interior_inter:
  5031 assumes "!S : I. convex (S :: ('n::euclidean_space) set)"
  5031 assumes "!S : I. convex (S :: ('n::euclidean_space) set)"
  5032 assumes "Inter {rel_interior S |S. S : I} ~= {}"
  5032 assumes "Inter {rel_interior S |S. S : I} ~= {}"
  5033 shows "Inter {closure S |S. S : I} <= closure (Inter {rel_interior S |S. S : I})"
  5033 shows "Inter {closure S |S. S : I} <= closure (Inter {rel_interior S |S. S : I})"
  5034 proof-
  5034 proof-
  5035 obtain x where x_def: "!S : I. x : rel_interior S" using assms by auto
  5035 obtain x where x_def: "!S : I. x : rel_interior S" using assms by auto
  5036 { fix y assume "y : Inter {closure S |S. S : I}" hence y_def: "!S : I. y : closure S" by auto
  5036 { fix y assume "y : Inter {closure S |S. S : I}" hence y_def: "!S : I. y : closure S" by auto
  5037   { assume "y = x" 
  5037   { assume "y = x"
  5038     hence "y : closure (Inter {rel_interior S |S. S : I})"
  5038     hence "y : closure (Inter {rel_interior S |S. S : I})"
  5039        using x_def closure_subset[of "Inter {rel_interior S |S. S : I}"] by auto
  5039        using x_def closure_subset[of "Inter {rel_interior S |S. S : I}"] by auto
  5040   }
  5040   }
  5041   moreover
  5041   moreover
  5042   { assume "y ~= x"
  5042   { assume "y ~= x"
  5043     { fix e :: real assume e_def: "0 < e"
  5043     { fix e :: real assume e_def: "0 < e"
  5044       def e1 == "min 1 (e/norm (y - x))" hence e1_def: "e1>0 & e1<=1 & e1*norm(y-x)<=e"
  5044       def e1 == "min 1 (e/norm (y - x))" hence e1_def: "e1>0 & e1<=1 & e1*norm(y-x)<=e"
  5045         using `y ~= x` `e>0` divide_pos_pos[of e] le_divide_eq[of e1 e "norm(y-x)"] by simp 
  5045         using `y ~= x` `e>0` divide_pos_pos[of e] le_divide_eq[of e1 e "norm(y-x)"] by simp
  5046       def z == "y - e1 *\<^sub>R (y - x)"
  5046       def z == "y - e1 *\<^sub>R (y - x)"
  5047       { fix S assume "S : I" 
  5047       { fix S assume "S : I"
  5048         hence "z : rel_interior S" using rel_interior_closure_convex_shrink[of S x y e1] 
  5048         hence "z : rel_interior S" using rel_interior_closure_convex_shrink[of S x y e1]
  5049            assms x_def y_def e1_def z_def by auto
  5049            assms x_def y_def e1_def z_def by auto
  5050       } hence *: "z : Inter {rel_interior S |S. S : I}" by auto
  5050       } hence *: "z : Inter {rel_interior S |S. S : I}" by auto
  5051       have "EX z. z:Inter {rel_interior S |S. S : I} & z ~= y & (dist z y) <= e"
  5051       have "EX z. z:Inter {rel_interior S |S. S : I} & z ~= y & (dist z y) <= e"
  5052            apply (rule_tac x="z" in exI) using `y ~= x` z_def * e1_def e_def dist_norm[of z y] by simp
  5052            apply (rule_tac x="z" in exI) using `y ~= x` z_def * e1_def e_def dist_norm[of z y] by simp
  5053     } hence "y islimpt Inter {rel_interior S |S. S : I}" unfolding islimpt_approachable_le by blast 
  5053     } hence "y islimpt Inter {rel_interior S |S. S : I}" unfolding islimpt_approachable_le by blast
  5054     hence "y : closure (Inter {rel_interior S |S. S : I})" unfolding closure_def by auto
  5054     hence "y : closure (Inter {rel_interior S |S. S : I})" unfolding closure_def by auto
  5055   } ultimately have "y : closure (Inter {rel_interior S |S. S : I})" by auto
  5055   } ultimately have "y : closure (Inter {rel_interior S |S. S : I})" by auto
  5056 } from this show ?thesis by auto
  5056 } from this show ?thesis by auto
  5057 qed
  5057 qed
  5058 
  5058 
  5059 
  5059 
  5060 lemma convex_closure_inter: 
  5060 lemma convex_closure_inter:
  5061 assumes "!S : I. convex (S :: ('n::euclidean_space) set)"
  5061 assumes "!S : I. convex (S :: ('n::euclidean_space) set)"
  5062 assumes "Inter {rel_interior S |S. S : I} ~= {}"
  5062 assumes "Inter {rel_interior S |S. S : I} ~= {}"
  5063 shows "closure (Inter I) = Inter {closure S |S. S : I}"
  5063 shows "closure (Inter I) = Inter {closure S |S. S : I}"
  5064 proof-
  5064 proof-
  5065 have "Inter {closure S |S. S : I} <= closure (Inter {rel_interior S |S. S : I})" 
  5065 have "Inter {closure S |S. S : I} <= closure (Inter {rel_interior S |S. S : I})"
  5066   using convex_closure_rel_interior_inter assms by auto
  5066   using convex_closure_rel_interior_inter assms by auto
  5067 moreover have "closure (Inter {rel_interior S |S. S : I}) <= closure (Inter I)" 
  5067 moreover have "closure (Inter {rel_interior S |S. S : I}) <= closure (Inter I)"
  5068     using rel_interior_inter_aux 
  5068     using rel_interior_inter_aux
  5069           closure_mono[of "Inter {rel_interior S |S. S : I}" "Inter I"] by auto
  5069           closure_mono[of "Inter {rel_interior S |S. S : I}" "Inter I"] by auto
  5070 ultimately show ?thesis using closure_inter[of I] by auto
  5070 ultimately show ?thesis using closure_inter[of I] by auto
  5071 qed
  5071 qed
  5072 
  5072 
  5073 lemma convex_inter_rel_interior_same_closure: 
  5073 lemma convex_inter_rel_interior_same_closure:
  5074 assumes "!S : I. convex (S :: ('n::euclidean_space) set)"
  5074 assumes "!S : I. convex (S :: ('n::euclidean_space) set)"
  5075 assumes "Inter {rel_interior S |S. S : I} ~= {}"
  5075 assumes "Inter {rel_interior S |S. S : I} ~= {}"
  5076 shows "closure (Inter {rel_interior S |S. S : I}) = closure (Inter I)"
  5076 shows "closure (Inter {rel_interior S |S. S : I}) = closure (Inter I)"
  5077 proof-
  5077 proof-
  5078 have "Inter {closure S |S. S : I} <= closure (Inter {rel_interior S |S. S : I})" 
  5078 have "Inter {closure S |S. S : I} <= closure (Inter {rel_interior S |S. S : I})"
  5079   using convex_closure_rel_interior_inter assms by auto
  5079   using convex_closure_rel_interior_inter assms by auto
  5080 moreover have "closure (Inter {rel_interior S |S. S : I}) <= closure (Inter I)" 
  5080 moreover have "closure (Inter {rel_interior S |S. S : I}) <= closure (Inter I)"
  5081     using rel_interior_inter_aux 
  5081     using rel_interior_inter_aux
  5082           closure_mono[of "Inter {rel_interior S |S. S : I}" "Inter I"] by auto
  5082           closure_mono[of "Inter {rel_interior S |S. S : I}" "Inter I"] by auto
  5083 ultimately show ?thesis using closure_inter[of I] by auto
  5083 ultimately show ?thesis using closure_inter[of I] by auto
  5084 qed
  5084 qed
  5085 
  5085 
  5086 lemma convex_rel_interior_inter: 
  5086 lemma convex_rel_interior_inter:
  5087 assumes "!S : I. convex (S :: ('n::euclidean_space) set)"
  5087 assumes "!S : I. convex (S :: ('n::euclidean_space) set)"
  5088 assumes "Inter {rel_interior S |S. S : I} ~= {}"
  5088 assumes "Inter {rel_interior S |S. S : I} ~= {}"
  5089 shows "rel_interior (Inter I) <= Inter {rel_interior S |S. S : I}"
  5089 shows "rel_interior (Inter I) <= Inter {rel_interior S |S. S : I}"
  5090 proof-
  5090 proof-
  5091 have "convex(Inter I)" using assms convex_Inter by auto
  5091 have "convex(Inter I)" using assms convex_Inter by auto
  5092 moreover have "convex(Inter {rel_interior S |S. S : I})" apply (rule convex_Inter)
  5092 moreover have "convex(Inter {rel_interior S |S. S : I})" apply (rule convex_Inter)
  5093    using assms convex_rel_interior by auto 
  5093    using assms convex_rel_interior by auto
  5094 ultimately have "rel_interior (Inter {rel_interior S |S. S : I}) = rel_interior (Inter I)"
  5094 ultimately have "rel_interior (Inter {rel_interior S |S. S : I}) = rel_interior (Inter I)"
  5095    using convex_inter_rel_interior_same_closure assms 
  5095    using convex_inter_rel_interior_same_closure assms
  5096    closure_eq_rel_interior_eq[of "Inter {rel_interior S |S. S : I}" "Inter I"] by blast
  5096    closure_eq_rel_interior_eq[of "Inter {rel_interior S |S. S : I}" "Inter I"] by blast
  5097 from this show ?thesis using rel_interior_subset[of "Inter {rel_interior S |S. S : I}"] by auto
  5097 from this show ?thesis using rel_interior_subset[of "Inter {rel_interior S |S. S : I}"] by auto
  5098 qed
  5098 qed
  5099 
  5099 
  5100 lemma convex_rel_interior_finite_inter: 
  5100 lemma convex_rel_interior_finite_inter:
  5101 assumes "!S : I. convex (S :: ('n::euclidean_space) set)"
  5101 assumes "!S : I. convex (S :: ('n::euclidean_space) set)"
  5102 assumes "Inter {rel_interior S |S. S : I} ~= {}"
  5102 assumes "Inter {rel_interior S |S. S : I} ~= {}"
  5103 assumes "finite I"
  5103 assumes "finite I"
  5104 shows "rel_interior (Inter I) = Inter {rel_interior S |S. S : I}"
  5104 shows "rel_interior (Inter I) = Inter {rel_interior S |S. S : I}"
  5105 proof-
  5105 proof-
  5108 { assume "I={}" hence ?thesis using Inter_empty rel_interior_univ2 by auto }
  5108 { assume "I={}" hence ?thesis using Inter_empty rel_interior_univ2 by auto }
  5109 moreover
  5109 moreover
  5110 { assume "I ~= {}"
  5110 { assume "I ~= {}"
  5111 { fix z assume z_def: "z : Inter {rel_interior S |S. S : I}"
  5111 { fix z assume z_def: "z : Inter {rel_interior S |S. S : I}"
  5112   { fix x assume x_def: "x : Inter I"
  5112   { fix x assume x_def: "x : Inter I"
  5113     { fix S assume S_def: "S : I" hence "z : rel_interior S" "x : S" using z_def x_def by auto 
  5113     { fix S assume S_def: "S : I" hence "z : rel_interior S" "x : S" using z_def x_def by auto
  5114       (*from this obtain e where e_def: "e>1 & (1 - e) *\<^sub>R x + e *\<^sub>R z : S"*)
  5114       (*from this obtain e where e_def: "e>1 & (1 - e) *\<^sub>R x + e *\<^sub>R z : S"*)
  5115       hence "EX m. m>1 & (!e. (e>1 & e<=m) --> (1-e)*\<^sub>R x+ e *\<^sub>R z : S )"
  5115       hence "EX m. m>1 & (!e. (e>1 & e<=m) --> (1-e)*\<^sub>R x+ e *\<^sub>R z : S )"
  5116          using convex_rel_interior_if[of S z] S_def assms hull_subset[of S] by auto
  5116          using convex_rel_interior_if[of S z] S_def assms hull_subset[of S] by auto
  5117     } from this obtain mS where mS_def: "!S : I. (mS(S) > (1 :: real) & 
  5117     } from this obtain mS where mS_def: "!S : I. (mS(S) > (1 :: real) &
  5118          (!e. (e>1 & e<=mS(S)) --> (1-e)*\<^sub>R x+ e *\<^sub>R z : S))" by metis
  5118          (!e. (e>1 & e<=mS(S)) --> (1-e)*\<^sub>R x+ e *\<^sub>R z : S))" by metis
  5119     obtain e where e_def: "e=Min (mS ` I)" by auto 
  5119     obtain e where e_def: "e=Min (mS ` I)" by auto
  5120     have "e : (mS ` I)" using e_def assms `I ~= {}` by simp
  5120     have "e : (mS ` I)" using e_def assms `I ~= {}` by simp
  5121     hence "e>(1 :: real)" using mS_def by auto
  5121     hence "e>(1 :: real)" using mS_def by auto
  5122     moreover have "!S : I. e<=mS(S)" using e_def assms by auto
  5122     moreover have "!S : I. e<=mS(S)" using e_def assms by auto
  5123     ultimately have "EX e>1. (1 - e) *\<^sub>R x + e *\<^sub>R z : Inter I" using mS_def by auto
  5123     ultimately have "EX e>1. (1 - e) *\<^sub>R x + e *\<^sub>R z : Inter I" using mS_def by auto
  5124   } hence "z : rel_interior (Inter I)" using convex_rel_interior_iff[of "Inter I" z]
  5124   } hence "z : rel_interior (Inter I)" using convex_rel_interior_iff[of "Inter I" z]
  5125        `Inter I ~= {}` `convex (Inter I)` by auto
  5125        `Inter I ~= {}` `convex (Inter I)` by auto
  5126 } from this have ?thesis using convex_rel_interior_inter[of I] assms by auto
  5126 } from this have ?thesis using convex_rel_interior_inter[of I] assms by auto
  5127 } ultimately show ?thesis by blast
  5127 } ultimately show ?thesis by blast
  5128 qed
  5128 qed
  5129 
  5129 
  5130 lemma convex_closure_inter_two: 
  5130 lemma convex_closure_inter_two:
  5131 fixes S T :: "('n::euclidean_space) set"
  5131 fixes S T :: "('n::euclidean_space) set"
  5132 assumes "convex S" "convex T"
  5132 assumes "convex S" "convex T"
  5133 assumes "(rel_interior S) Int (rel_interior T) ~= {}"
  5133 assumes "(rel_interior S) Int (rel_interior T) ~= {}"
  5134 shows "closure (S Int T) = (closure S) Int (closure T)" 
  5134 shows "closure (S Int T) = (closure S) Int (closure T)"
  5135 using convex_closure_inter[of "{S,T}"] assms by auto
  5135 using convex_closure_inter[of "{S,T}"] assms by auto
  5136 
  5136 
  5137 lemma convex_rel_interior_inter_two: 
  5137 lemma convex_rel_interior_inter_two:
  5138 fixes S T :: "('n::euclidean_space) set"
  5138 fixes S T :: "('n::euclidean_space) set"
  5139 assumes "convex S" "convex T"
  5139 assumes "convex S" "convex T"
  5140 assumes "(rel_interior S) Int (rel_interior T) ~= {}"
  5140 assumes "(rel_interior S) Int (rel_interior T) ~= {}"
  5141 shows "rel_interior (S Int T) = (rel_interior S) Int (rel_interior T)" 
  5141 shows "rel_interior (S Int T) = (rel_interior S) Int (rel_interior T)"
  5142 using convex_rel_interior_finite_inter[of "{S,T}"] assms by auto
  5142 using convex_rel_interior_finite_inter[of "{S,T}"] assms by auto
  5143 
  5143 
  5144 
  5144 
  5145 lemma convex_affine_closure_inter: 
  5145 lemma convex_affine_closure_inter:
  5146 fixes S T :: "('n::euclidean_space) set"
  5146 fixes S T :: "('n::euclidean_space) set"
  5147 assumes "convex S" "affine T"
  5147 assumes "convex S" "affine T"
  5148 assumes "(rel_interior S) Int T ~= {}"
  5148 assumes "(rel_interior S) Int T ~= {}"
  5149 shows "closure (S Int T) = (closure S) Int T"
  5149 shows "closure (S Int T) = (closure S) Int T"
  5150 proof- 
  5150 proof-
  5151 have "affine hull T = T" using assms by auto
  5151 have "affine hull T = T" using assms by auto
  5152 hence "rel_interior T = T" using rel_interior_univ[of T] by metis
  5152 hence "rel_interior T = T" using rel_interior_univ[of T] by metis
  5153 moreover have "closure T = T" using assms affine_closed[of T] by auto
  5153 moreover have "closure T = T" using assms affine_closed[of T] by auto
  5154 ultimately show ?thesis using convex_closure_inter_two[of S T] assms affine_imp_convex by auto 
  5154 ultimately show ?thesis using convex_closure_inter_two[of S T] assms affine_imp_convex by auto
  5155 qed
  5155 qed
  5156 
  5156 
  5157 lemma convex_affine_rel_interior_inter: 
  5157 lemma convex_affine_rel_interior_inter:
  5158 fixes S T :: "('n::euclidean_space) set"
  5158 fixes S T :: "('n::euclidean_space) set"
  5159 assumes "convex S" "affine T"
  5159 assumes "convex S" "affine T"
  5160 assumes "(rel_interior S) Int T ~= {}"
  5160 assumes "(rel_interior S) Int T ~= {}"
  5161 shows "rel_interior (S Int T) = (rel_interior S) Int T"
  5161 shows "rel_interior (S Int T) = (rel_interior S) Int T"
  5162 proof- 
  5162 proof-
  5163 have "affine hull T = T" using assms by auto
  5163 have "affine hull T = T" using assms by auto
  5164 hence "rel_interior T = T" using rel_interior_univ[of T] by metis
  5164 hence "rel_interior T = T" using rel_interior_univ[of T] by metis
  5165 moreover have "closure T = T" using assms affine_closed[of T] by auto
  5165 moreover have "closure T = T" using assms affine_closed[of T] by auto
  5166 ultimately show ?thesis using convex_rel_interior_inter_two[of S T] assms affine_imp_convex by auto 
  5166 ultimately show ?thesis using convex_rel_interior_inter_two[of S T] assms affine_imp_convex by auto
  5167 qed
  5167 qed
  5168 
  5168 
  5169 lemma subset_rel_interior_convex:
  5169 lemma subset_rel_interior_convex:
  5170 fixes S T :: "('n::euclidean_space) set"
  5170 fixes S T :: "('n::euclidean_space) set"
  5171 assumes "convex S" "convex T"
  5171 assumes "convex S" "convex T"
  5173 assumes "~(S <= rel_frontier T)"
  5173 assumes "~(S <= rel_frontier T)"
  5174 shows "rel_interior S <= rel_interior T"
  5174 shows "rel_interior S <= rel_interior T"
  5175 proof-
  5175 proof-
  5176 have *: "S Int closure T = S" using assms by auto
  5176 have *: "S Int closure T = S" using assms by auto
  5177 have "~(rel_interior S <= rel_frontier T)"
  5177 have "~(rel_interior S <= rel_frontier T)"
  5178      using closure_mono[of "rel_interior S" "rel_frontier T"] closed_rel_frontier[of T] 
  5178      using closure_mono[of "rel_interior S" "rel_frontier T"] closed_rel_frontier[of T]
  5179      closure_closed convex_closure_rel_interior[of S] closure_subset[of S] assms by auto
  5179      closure_closed convex_closure_rel_interior[of S] closure_subset[of S] assms by auto
  5180 hence "(rel_interior S) Int (rel_interior (closure T)) ~= {}" 
  5180 hence "(rel_interior S) Int (rel_interior (closure T)) ~= {}"
  5181      using assms rel_frontier_def[of T] rel_interior_subset convex_rel_interior_closure[of T] by auto
  5181      using assms rel_frontier_def[of T] rel_interior_subset convex_rel_interior_closure[of T] by auto
  5182 hence "rel_interior S Int rel_interior T = rel_interior (S Int closure T)" using assms convex_closure  
  5182 hence "rel_interior S Int rel_interior T = rel_interior (S Int closure T)" using assms convex_closure
  5183      convex_rel_interior_inter_two[of S "closure T"] convex_rel_interior_closure[of T] by auto
  5183      convex_rel_interior_inter_two[of S "closure T"] convex_rel_interior_closure[of T] by auto
  5184 also have "...=rel_interior (S)" using * by auto
  5184 also have "...=rel_interior (S)" using * by auto
  5185 finally show ?thesis by auto
  5185 finally show ?thesis by auto
  5186 qed
  5186 qed
  5187 
  5187 
  5195 { assume "S = {}" hence ?thesis using assms rel_interior_empty rel_interior_convex_nonempty by auto }
  5195 { assume "S = {}" hence ?thesis using assms rel_interior_empty rel_interior_convex_nonempty by auto }
  5196 moreover
  5196 moreover
  5197 { assume "S ~= {}"
  5197 { assume "S ~= {}"
  5198 have *: "f ` (rel_interior S) <= f ` S" unfolding image_mono using rel_interior_subset by auto
  5198 have *: "f ` (rel_interior S) <= f ` S" unfolding image_mono using rel_interior_subset by auto
  5199 have "f ` S <= f ` (closure S)" unfolding image_mono using closure_subset by auto
  5199 have "f ` S <= f ` (closure S)" unfolding image_mono using closure_subset by auto
  5200 also have "... = f ` (closure (rel_interior S))" using convex_closure_rel_interior assms by auto  
  5200 also have "... = f ` (closure (rel_interior S))" using convex_closure_rel_interior assms by auto
  5201 also have "... <= closure (f ` (rel_interior S))" using closure_linear_image assms by auto 
  5201 also have "... <= closure (f ` (rel_interior S))" using closure_linear_image assms by auto
  5202 finally have "closure (f ` S) = closure (f ` rel_interior S)"
  5202 finally have "closure (f ` S) = closure (f ` rel_interior S)"
  5203    using closure_mono[of "f ` S" "closure (f ` rel_interior S)"] closure_closure 
  5203    using closure_mono[of "f ` S" "closure (f ` rel_interior S)"] closure_closure
  5204          closure_mono[of "f ` rel_interior S" "f ` S"] * by auto
  5204          closure_mono[of "f ` rel_interior S" "f ` S"] * by auto
  5205 hence "rel_interior (f ` S) = rel_interior (f ` rel_interior S)" using assms convex_rel_interior
  5205 hence "rel_interior (f ` S) = rel_interior (f ` rel_interior S)" using assms convex_rel_interior
  5206    linear_conv_bounded_linear[of f] convex_linear_image[of S] convex_linear_image[of "rel_interior S"] 
  5206    linear_conv_bounded_linear[of f] convex_linear_image[of S] convex_linear_image[of "rel_interior S"]
  5207    closure_eq_rel_interior_eq[of "f ` S" "f ` rel_interior S"] by auto
  5207    closure_eq_rel_interior_eq[of "f ` S" "f ` rel_interior S"] by auto
  5208 hence "rel_interior (f ` S) <= f ` rel_interior S" using rel_interior_subset by auto
  5208 hence "rel_interior (f ` S) <= f ` rel_interior S" using rel_interior_subset by auto
  5209 moreover
  5209 moreover
  5210 { fix z assume z_def: "z : f ` rel_interior S"
  5210 { fix z assume z_def: "z : f ` rel_interior S"
  5211   from this obtain z1 where z1_def: "z1 : rel_interior S & (f z1 = z)" by auto
  5211   from this obtain z1 where z1_def: "z1 : rel_interior S & (f z1 = z)" by auto
  5216     moreover have "f ((1 - e) *\<^sub>R x1 + e *\<^sub>R z1) = (1 - e) *\<^sub>R x + e *\<^sub>R z"
  5216     moreover have "f ((1 - e) *\<^sub>R x1 + e *\<^sub>R z1) = (1 - e) *\<^sub>R x + e *\<^sub>R z"
  5217         using x1_def z1_def `linear f` by (simp add: linear_add_cmul)
  5217         using x1_def z1_def `linear f` by (simp add: linear_add_cmul)
  5218     ultimately have "(1 - e) *\<^sub>R x + e *\<^sub>R z : f ` S"
  5218     ultimately have "(1 - e) *\<^sub>R x + e *\<^sub>R z : f ` S"
  5219         using imageI[of "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1" S f] by auto
  5219         using imageI[of "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1" S f] by auto
  5220     hence "EX e. (e>1 & (1 - e) *\<^sub>R x + e *\<^sub>R z : f ` S)" using e_def by auto
  5220     hence "EX e. (e>1 & (1 - e) *\<^sub>R x + e *\<^sub>R z : f ` S)" using e_def by auto
  5221   } from this have "z : rel_interior (f ` S)" using convex_rel_interior_iff[of "f ` S" z] `convex S` 
  5221   } from this have "z : rel_interior (f ` S)" using convex_rel_interior_iff[of "f ` S" z] `convex S`
  5222        `linear f` `S ~= {}` convex_linear_image[of S f]  linear_conv_bounded_linear[of f] by auto
  5222        `linear f` `S ~= {}` convex_linear_image[of S f]  linear_conv_bounded_linear[of f] by auto
  5223 } ultimately have ?thesis by auto
  5223 } ultimately have ?thesis by auto
  5224 } ultimately show ?thesis by blast
  5224 } ultimately show ?thesis by blast
  5225 qed
  5225 qed
  5226 
  5226 
  5244 assumes "convex S"
  5244 assumes "convex S"
  5245 assumes "f -` (rel_interior S) ~= {}"
  5245 assumes "f -` (rel_interior S) ~= {}"
  5246 shows "rel_interior (f -` S) = f -` (rel_interior S)"
  5246 shows "rel_interior (f -` S) = f -` (rel_interior S)"
  5247 proof-
  5247 proof-
  5248 have "S ~= {}" using assms rel_interior_empty by auto
  5248 have "S ~= {}" using assms rel_interior_empty by auto
  5249 have nonemp: "f -` S ~= {}" by (metis assms(3) rel_interior_subset subset_empty vimage_mono) 
  5249 have nonemp: "f -` S ~= {}" by (metis assms(3) rel_interior_subset subset_empty vimage_mono)
  5250 hence "S Int (range f) ~= {}" by auto
  5250 hence "S Int (range f) ~= {}" by auto
  5251 have conv: "convex (f -` S)" using convex_linear_preimage assms linear_conv_bounded_linear by auto
  5251 have conv: "convex (f -` S)" using convex_linear_preimage assms linear_conv_bounded_linear by auto
  5252 hence "convex (S Int (range f))"
  5252 hence "convex (S Int (range f))"
  5253   by (metis assms(1) assms(2) convex_Int subspace_UNIV subspace_imp_convex subspace_linear_image)
  5253   by (metis assms(1) assms(2) convex_Int subspace_UNIV subspace_imp_convex subspace_linear_image)
  5254 { fix z assume "z : f -` (rel_interior S)"
  5254 { fix z assume "z : f -` (rel_interior S)"
  5257     from this obtain e where e_def: "e>1 & (1-e)*\<^sub>R (f x)+ e *\<^sub>R (f z) : S"
  5257     from this obtain e where e_def: "e>1 & (1-e)*\<^sub>R (f x)+ e *\<^sub>R (f z) : S"
  5258       using convex_rel_interior_iff[of S "f z"] z_def assms `S ~= {}` by auto
  5258       using convex_rel_interior_iff[of S "f z"] z_def assms `S ~= {}` by auto
  5259     moreover have "(1-e)*\<^sub>R (f x)+ e *\<^sub>R (f z) = f ((1-e)*\<^sub>R x + e *\<^sub>R z)"
  5259     moreover have "(1-e)*\<^sub>R (f x)+ e *\<^sub>R (f z) = f ((1-e)*\<^sub>R x + e *\<^sub>R z)"
  5260       using `linear f` by (simp add: linear_def)
  5260       using `linear f` by (simp add: linear_def)
  5261     ultimately have "EX e. e>1 & (1-e)*\<^sub>R x + e *\<^sub>R z : f -` S" using e_def by auto
  5261     ultimately have "EX e. e>1 & (1-e)*\<^sub>R x + e *\<^sub>R z : f -` S" using e_def by auto
  5262   } hence "z : rel_interior (f -` S)" 
  5262   } hence "z : rel_interior (f -` S)"
  5263        using convex_rel_interior_iff[of "f -` S" z] conv nonemp by auto
  5263        using convex_rel_interior_iff[of "f -` S" z] conv nonemp by auto
  5264 } 
  5264 }
  5265 moreover
  5265 moreover
  5266 { fix z assume z_def: "z : rel_interior (f -` S)" 
  5266 { fix z assume z_def: "z : rel_interior (f -` S)"
  5267   { fix x assume x_def: "x: S Int (range f)"
  5267   { fix x assume x_def: "x: S Int (range f)"
  5268     from this obtain y where y_def: "(f y = x) & (y : f -` S)" by auto
  5268     from this obtain y where y_def: "(f y = x) & (y : f -` S)" by auto
  5269     from this obtain e where e_def: "e>1 & (1-e)*\<^sub>R y+ e *\<^sub>R z : f -` S"
  5269     from this obtain e where e_def: "e>1 & (1-e)*\<^sub>R y+ e *\<^sub>R z : f -` S"
  5270       using convex_rel_interior_iff[of "f -` S" z] z_def conv by auto
  5270       using convex_rel_interior_iff[of "f -` S" z] z_def conv by auto
  5271     moreover have "(1-e)*\<^sub>R x+ e *\<^sub>R (f z) = f ((1-e)*\<^sub>R y + e *\<^sub>R z)"
  5271     moreover have "(1-e)*\<^sub>R x+ e *\<^sub>R (f z) = f ((1-e)*\<^sub>R y + e *\<^sub>R z)"
  5274       using e_def by auto
  5274       using e_def by auto
  5275   } hence "f z : rel_interior (S Int (range f))" using `convex (S Int (range f))`
  5275   } hence "f z : rel_interior (S Int (range f))" using `convex (S Int (range f))`
  5276     `S Int (range f) ~= {}` convex_rel_interior_iff[of "S Int (range f)" "f z"] by auto
  5276     `S Int (range f) ~= {}` convex_rel_interior_iff[of "S Int (range f)" "f z"] by auto
  5277   moreover have "affine (range f)"
  5277   moreover have "affine (range f)"
  5278     by (metis assms(1) subspace_UNIV subspace_imp_affine subspace_linear_image)
  5278     by (metis assms(1) subspace_UNIV subspace_imp_affine subspace_linear_image)
  5279   ultimately have "f z : rel_interior S" 
  5279   ultimately have "f z : rel_interior S"
  5280     using convex_affine_rel_interior_inter[of S "range f"] assms by auto
  5280     using convex_affine_rel_interior_inter[of S "range f"] assms by auto
  5281   hence "z : f -` (rel_interior S)" by auto
  5281   hence "z : f -` (rel_interior S)" by auto
  5282 }
  5282 }
  5283 ultimately show ?thesis by auto
  5283 ultimately show ?thesis by auto
  5284 qed
  5284 qed
  5285     
  5285 
  5286 
  5286 
  5287 lemma convex_direct_sum:
  5287 lemma convex_direct_sum:
  5288 fixes S :: "('n::euclidean_space) set"
  5288 fixes S :: "('n::euclidean_space) set"
  5289 fixes T :: "('m::euclidean_space) set"
  5289 fixes T :: "('m::euclidean_space) set"
  5290 assumes "convex S" "convex T"
  5290 assumes "convex S" "convex T"
  5311 fixes T :: "('m::euclidean_space) set"
  5311 fixes T :: "('m::euclidean_space) set"
  5312 shows "convex hull (S <*> T) = (convex hull S) <*> (convex hull T)"
  5312 shows "convex hull (S <*> T) = (convex hull S) <*> (convex hull T)"
  5313 proof-
  5313 proof-
  5314 { fix x assume "x : (convex hull S) <*> (convex hull T)"
  5314 { fix x assume "x : (convex hull S) <*> (convex hull T)"
  5315   from this obtain xs xt where xst_def: "xs : convex hull S & xt : convex hull T & (xs,xt) = x" by auto
  5315   from this obtain xs xt where xst_def: "xs : convex hull S & xt : convex hull T & (xs,xt) = x" by auto
  5316   from xst_def obtain sI su where s: "finite sI & sI <= S & (ALL x:sI. 0 <= su x) & setsum su sI = 1 
  5316   from xst_def obtain sI su where s: "finite sI & sI <= S & (ALL x:sI. 0 <= su x) & setsum su sI = 1
  5317      & (SUM v:sI. su v *\<^sub>R v) = xs" using convex_hull_explicit[of S] by auto
  5317      & (SUM v:sI. su v *\<^sub>R v) = xs" using convex_hull_explicit[of S] by auto
  5318   from xst_def obtain tI tu where t: "finite tI & tI <= T & (ALL x:tI. 0 <= tu x) & setsum tu tI = 1 
  5318   from xst_def obtain tI tu where t: "finite tI & tI <= T & (ALL x:tI. 0 <= tu x) & setsum tu tI = 1
  5319      & (SUM v:tI. tu v *\<^sub>R v) = xt" using convex_hull_explicit[of T] by auto
  5319      & (SUM v:tI. tu v *\<^sub>R v) = xt" using convex_hull_explicit[of T] by auto
  5320   def I == "(sI <*> tI)"
  5320   def I == "(sI <*> tI)"
  5321   def u == "(%i. (su (fst i))*(tu(snd i)))"
  5321   def u == "(%i. (su (fst i))*(tu(snd i)))"
  5322   have "fst (SUM v:sI <*> tI. (su (fst v) * tu (snd v)) *\<^sub>R v)=
  5322   have "fst (SUM v:sI <*> tI. (su (fst v) * tu (snd v)) *\<^sub>R v)=
  5323      (SUM vs:sI. SUM vt:tI. (su vs * tu vt) *\<^sub>R vs)"
  5323      (SUM vs:sI. SUM vt:tI. (su vs * tu vt) *\<^sub>R vs)"
  5340   also have "...=(SUM vs:sI. su vs) *\<^sub>R xt" by (simp add: scaleR_left.setsum)
  5340   also have "...=(SUM vs:sI. su vs) *\<^sub>R xt" by (simp add: scaleR_left.setsum)
  5341   also have "...=xt" using s by auto
  5341   also have "...=xt" using s by auto
  5342   finally have h2: "snd (SUM v:sI <*> tI. (su (fst v) * tu (snd v)) *\<^sub>R v)=xt" by auto
  5342   finally have h2: "snd (SUM v:sI <*> tI. (su (fst v) * tu (snd v)) *\<^sub>R v)=xt" by auto
  5343   from h1 h2 have "(SUM v:sI <*> tI. (su (fst v) * tu (snd v)) *\<^sub>R v) = x" using xst_def by auto
  5343   from h1 h2 have "(SUM v:sI <*> tI. (su (fst v) * tu (snd v)) *\<^sub>R v) = x" using xst_def by auto
  5344 
  5344 
  5345   moreover have "finite I & (I <= S <*> T)" using s t I_def by auto 
  5345   moreover have "finite I & (I <= S <*> T)" using s t I_def by auto
  5346   moreover have "!i:I. 0 <= u i" using s t I_def u_def by (simp add: mult_nonneg_nonneg)
  5346   moreover have "!i:I. 0 <= u i" using s t I_def u_def by (simp add: mult_nonneg_nonneg)
  5347   moreover have "setsum u I = 1" using u_def I_def setsum_cartesian_product[of "(% x y. (su x)*(tu y))"] 
  5347   moreover have "setsum u I = 1" using u_def I_def setsum_cartesian_product[of "(% x y. (su x)*(tu y))"]
  5348      s t setsum_product[of su sI tu tI] by (auto simp add: split_def)
  5348      s t setsum_product[of su sI tu tI] by (auto simp add: split_def)
  5349   ultimately have "x : convex hull (S <*> T)" 
  5349   ultimately have "x : convex hull (S <*> T)"
  5350      apply (subst convex_hull_explicit[of "S <*> T"]) apply rule
  5350      apply (subst convex_hull_explicit[of "S <*> T"]) apply rule
  5351      apply (rule_tac x="I" in exI) apply (rule_tac x="u" in exI)
  5351      apply (rule_tac x="I" in exI) apply (rule_tac x="u" in exI)
  5352      using I_def u_def by auto
  5352      using I_def u_def by auto
  5353 }
  5353 }
  5354 hence "convex hull (S <*> T) >= (convex hull S) <*> (convex hull T)" by auto
  5354 hence "convex hull (S <*> T) >= (convex hull S) <*> (convex hull T)" by auto
  5355 moreover have "convex ((convex hull S) <*> (convex hull T))" 
  5355 moreover have "convex ((convex hull S) <*> (convex hull T))"
  5356    by (simp add: convex_direct_sum convex_convex_hull)
  5356    by (simp add: convex_direct_sum convex_convex_hull)
  5357 ultimately show ?thesis 
  5357 ultimately show ?thesis
  5358    using hull_minimal[of "S <*> T" "(convex hull S) <*> (convex hull T)" "convex"] 
  5358    using hull_minimal[of "S <*> T" "(convex hull S) <*> (convex hull T)" "convex"]
  5359          hull_subset[of S convex] hull_subset[of T convex] by auto
  5359          hull_subset[of S convex] hull_subset[of T convex] by auto
  5360 qed
  5360 qed
  5361 
  5361 
  5362 lemma rel_interior_direct_sum:
  5362 lemma rel_interior_direct_sum:
  5363 fixes S :: "('n::euclidean_space) set"
  5363 fixes S :: "('n::euclidean_space) set"
  5371 moreover {
  5371 moreover {
  5372 assume "S ~={}" "T ~={}"
  5372 assume "S ~={}" "T ~={}"
  5373 hence ri: "rel_interior S ~= {}" "rel_interior T ~= {}" using rel_interior_convex_nonempty assms by auto
  5373 hence ri: "rel_interior S ~= {}" "rel_interior T ~= {}" using rel_interior_convex_nonempty assms by auto
  5374 hence "fst -` rel_interior S ~= {}" using fst_vimage_eq_Times[of "rel_interior S"] by auto
  5374 hence "fst -` rel_interior S ~= {}" using fst_vimage_eq_Times[of "rel_interior S"] by auto
  5375 hence "rel_interior ((fst :: 'n * 'm => 'n) -` S) = fst -` rel_interior S"
  5375 hence "rel_interior ((fst :: 'n * 'm => 'n) -` S) = fst -` rel_interior S"
  5376   using fst_linear `convex S` rel_interior_convex_linear_preimage[of fst S] by auto 
  5376   using fst_linear `convex S` rel_interior_convex_linear_preimage[of fst S] by auto
  5377 hence s: "rel_interior (S <*> (UNIV :: 'm set)) = rel_interior S <*> UNIV" by (simp add: fst_vimage_eq_Times)
  5377 hence s: "rel_interior (S <*> (UNIV :: 'm set)) = rel_interior S <*> UNIV" by (simp add: fst_vimage_eq_Times)
  5378 from ri have "snd -` rel_interior T ~= {}" using snd_vimage_eq_Times[of "rel_interior T"] by auto
  5378 from ri have "snd -` rel_interior T ~= {}" using snd_vimage_eq_Times[of "rel_interior T"] by auto
  5379 hence "rel_interior ((snd :: 'n * 'm => 'm) -` T) = snd -` rel_interior T"
  5379 hence "rel_interior ((snd :: 'n * 'm => 'm) -` T) = snd -` rel_interior T"
  5380   using snd_linear `convex T` rel_interior_convex_linear_preimage[of snd T] by auto 
  5380   using snd_linear `convex T` rel_interior_convex_linear_preimage[of snd T] by auto
  5381 hence t: "rel_interior ((UNIV :: 'n set) <*> T) = UNIV <*> rel_interior T" by (simp add: snd_vimage_eq_Times)
  5381 hence t: "rel_interior ((UNIV :: 'n set) <*> T) = UNIV <*> rel_interior T" by (simp add: snd_vimage_eq_Times)
  5382 from s t have *: "rel_interior (S <*> (UNIV :: 'm set)) Int rel_interior ((UNIV :: 'n set) <*> T) 
  5382 from s t have *: "rel_interior (S <*> (UNIV :: 'm set)) Int rel_interior ((UNIV :: 'n set) <*> T)
  5383   = rel_interior S <*> rel_interior T" by auto
  5383   = rel_interior S <*> rel_interior T" by auto
  5384 have "(S <*> T) = (S <*> (UNIV :: 'm set)) Int ((UNIV :: 'n set) <*> T)" by auto
  5384 have "(S <*> T) = (S <*> (UNIV :: 'm set)) Int ((UNIV :: 'n set) <*> T)" by auto
  5385 hence "rel_interior (S <*> T) = rel_interior ((S <*> (UNIV :: 'm set)) Int ((UNIV :: 'n set) <*> T))" by auto
  5385 hence "rel_interior (S <*> T) = rel_interior ((S <*> (UNIV :: 'm set)) Int ((UNIV :: 'n set) <*> T))" by auto
  5386 also have "...=rel_interior (S <*> (UNIV :: 'm set)) Int rel_interior ((UNIV :: 'n set) <*> T)" 
  5386 also have "...=rel_interior (S <*> (UNIV :: 'm set)) Int rel_interior ((UNIV :: 'n set) <*> T)"
  5387    apply (subst convex_rel_interior_inter_two[of "S <*> (UNIV :: 'm set)" "(UNIV :: 'n set) <*> T"]) 
  5387    apply (subst convex_rel_interior_inter_two[of "S <*> (UNIV :: 'm set)" "(UNIV :: 'n set) <*> T"])
  5388    using * ri assms convex_direct_sum by auto
  5388    using * ri assms convex_direct_sum by auto
  5389 finally have ?thesis using * by auto
  5389 finally have ?thesis using * by auto
  5390 }
  5390 }
  5391 ultimately show ?thesis by blast
  5391 ultimately show ?thesis by blast
  5392 qed
  5392 qed
  5393 
  5393 
  5394 lemma rel_interior_scaleR: 
  5394 lemma rel_interior_scaleR:
  5395 fixes S :: "('n::euclidean_space) set"
  5395 fixes S :: "('n::euclidean_space) set"
  5396 assumes "c ~= 0"
  5396 assumes "c ~= 0"
  5397 shows "(op *\<^sub>R c) ` (rel_interior S) = rel_interior ((op *\<^sub>R c) ` S)"
  5397 shows "(op *\<^sub>R c) ` (rel_interior S) = rel_interior ((op *\<^sub>R c) ` S)"
  5398 using rel_interior_injective_linear_image[of "(op *\<^sub>R c)" S]
  5398 using rel_interior_injective_linear_image[of "(op *\<^sub>R c)" S]
  5399       linear_conv_bounded_linear[of "op *\<^sub>R c"] linear_scaleR injective_scaleR[of c] assms by auto
  5399       linear_conv_bounded_linear[of "op *\<^sub>R c"] linear_scaleR injective_scaleR[of c] assms by auto
  5400 
  5400 
  5401 lemma rel_interior_convex_scaleR: 
  5401 lemma rel_interior_convex_scaleR:
  5402 fixes S :: "('n::euclidean_space) set"
  5402 fixes S :: "('n::euclidean_space) set"
  5403 assumes "convex S"
  5403 assumes "convex S"
  5404 shows "(op *\<^sub>R c) ` (rel_interior S) = rel_interior ((op *\<^sub>R c) ` S)"
  5404 shows "(op *\<^sub>R c) ` (rel_interior S) = rel_interior ((op *\<^sub>R c) ` S)"
  5405 by (metis assms linear_scaleR rel_interior_convex_linear_image)
  5405 by (metis assms linear_scaleR rel_interior_convex_linear_image)
  5406 
  5406 
  5407 lemma convex_rel_open_scaleR: 
  5407 lemma convex_rel_open_scaleR:
  5408 fixes S :: "('n::euclidean_space) set"
  5408 fixes S :: "('n::euclidean_space) set"
  5409 assumes "convex S" "rel_open S"
  5409 assumes "convex S" "rel_open S"
  5410 shows "convex ((op *\<^sub>R c) ` S) & rel_open ((op *\<^sub>R c) ` S)"
  5410 shows "convex ((op *\<^sub>R c) ` S) & rel_open ((op *\<^sub>R c) ` S)"
  5411 by (metis assms convex_scaling rel_interior_convex_scaleR rel_open_def)
  5411 by (metis assms convex_scaling rel_interior_convex_scaleR rel_open_def)
  5412 
  5412 
  5413 
  5413 
  5414 lemma convex_rel_open_finite_inter: 
  5414 lemma convex_rel_open_finite_inter:
  5415 assumes "!S : I. (convex (S :: ('n::euclidean_space) set) & rel_open S)"
  5415 assumes "!S : I. (convex (S :: ('n::euclidean_space) set) & rel_open S)"
  5416 assumes "finite I"
  5416 assumes "finite I"
  5417 shows "convex (Inter I) & rel_open (Inter I)"
  5417 shows "convex (Inter I) & rel_open (Inter I)"
  5418 proof-
  5418 proof-
  5419 { assume "Inter {rel_interior S |S. S : I} = {}"
  5419 { assume "Inter {rel_interior S |S. S : I} = {}"
  5431 lemma convex_rel_open_linear_image:
  5431 lemma convex_rel_open_linear_image:
  5432 fixes f :: "('m::euclidean_space) => ('n::euclidean_space)"
  5432 fixes f :: "('m::euclidean_space) => ('n::euclidean_space)"
  5433 assumes "linear f"
  5433 assumes "linear f"
  5434 assumes "convex S" "rel_open S"
  5434 assumes "convex S" "rel_open S"
  5435 shows "convex (f ` S) & rel_open (f ` S)"
  5435 shows "convex (f ` S) & rel_open (f ` S)"
  5436 by (metis assms convex_linear_image rel_interior_convex_linear_image 
  5436 by (metis assms convex_linear_image rel_interior_convex_linear_image
  5437    linear_conv_bounded_linear rel_open_def)
  5437    linear_conv_bounded_linear rel_open_def)
  5438 
  5438 
  5439 lemma convex_rel_open_linear_preimage:
  5439 lemma convex_rel_open_linear_preimage:
  5440 fixes f :: "('m::euclidean_space) => ('n::euclidean_space)"
  5440 fixes f :: "('m::euclidean_space) => ('n::euclidean_space)"
  5441 assumes "linear f"
  5441 assumes "linear f"
  5442 assumes "convex S" "rel_open S"
  5442 assumes "convex S" "rel_open S"
  5443 shows "convex (f -` S) & rel_open (f -` S)" 
  5443 shows "convex (f -` S) & rel_open (f -` S)"
  5444 proof-
  5444 proof-
  5445 { assume "f -` (rel_interior S) = {}"
  5445 { assume "f -` (rel_interior S) = {}"
  5446   hence "f -` S = {}" using assms unfolding rel_open_def by auto
  5446   hence "f -` S = {}" using assms unfolding rel_open_def by auto
  5447   hence ?thesis unfolding rel_open_def using rel_interior_empty by auto
  5447   hence ?thesis unfolding rel_open_def using rel_interior_empty by auto
  5448 }
  5448 }
  5481     moreover have "x = snd (y,x)" by auto
  5481     moreover have "x = snd (y,x)" by auto
  5482     ultimately have "x : snd ` (S Int fst -` {y})" by blast
  5482     ultimately have "x : snd ` (S Int fst -` {y})" by blast
  5483   }
  5483   }
  5484   hence "snd ` (S Int fst -` {y}) = f y" using assms by auto
  5484   hence "snd ` (S Int fst -` {y}) = f y" using assms by auto
  5485   hence ***: "rel_interior (f y) = snd ` rel_interior (S Int fst -` {y})"
  5485   hence ***: "rel_interior (f y) = snd ` rel_interior (S Int fst -` {y})"
  5486     using rel_interior_convex_linear_image[of snd "S Int fst -` {y}"] snd_linear conv by auto 
  5486     using rel_interior_convex_linear_image[of snd "S Int fst -` {y}"] snd_linear conv by auto
  5487   { fix z assume "z : rel_interior (f y)"
  5487   { fix z assume "z : rel_interior (f y)"
  5488     hence "z : snd ` rel_interior (S Int fst -` {y})" using *** by auto
  5488     hence "z : snd ` rel_interior (S Int fst -` {y})" using *** by auto
  5489     moreover have "{y} = fst ` rel_interior (S Int fst -` {y})" using * ** rel_interior_subset by auto   
  5489     moreover have "{y} = fst ` rel_interior (S Int fst -` {y})" using * ** rel_interior_subset by auto
  5490     ultimately have "(y,z) : rel_interior (S Int fst -` {y})" by force
  5490     ultimately have "(y,z) : rel_interior (S Int fst -` {y})" by force
  5491     hence "(y,z) : rel_interior S" using ** by auto
  5491     hence "(y,z) : rel_interior S" using ** by auto
  5492   }
  5492   }
  5493   moreover
  5493   moreover
  5494   { fix z assume "(y,z) : rel_interior S"
  5494   { fix z assume "(y,z) : rel_interior S"
  5495     hence "(y,z) : rel_interior (S Int fst -` {y})" using ** by auto
  5495     hence "(y,z) : rel_interior (S Int fst -` {y})" using ** by auto
  5496     hence "z : snd ` rel_interior (S Int fst -` {y})" by (metis Range_iff snd_eq_Range) 
  5496     hence "z : snd ` rel_interior (S Int fst -` {y})" by (metis Range_iff snd_eq_Range)
  5497     hence "z : rel_interior (f y)" using *** by auto
  5497     hence "z : rel_interior (f y)" using *** by auto
  5498   }
  5498   }
  5499   ultimately have "!!z. (y,z) : rel_interior S <-> z : rel_interior (f y)" by auto
  5499   ultimately have "!!z. (y,z) : rel_interior S <-> z : rel_interior (f y)" by auto
  5500 } 
  5500 }
  5501 hence h2: "!!y z. y : rel_interior {t. f t ~= {}} ==> ((y, z) : rel_interior S) = (z : rel_interior (f y))"
  5501 hence h2: "!!y z. y : rel_interior {t. f t ~= {}} ==> ((y, z) : rel_interior S) = (z : rel_interior (f y))"
  5502   by auto
  5502   by auto
  5503 { fix y z assume asm: "(y, z) : rel_interior S"
  5503 { fix y z assume asm: "(y, z) : rel_interior S"
  5504   hence "y : fst ` rel_interior S" by (metis Domain_iff fst_eq_Domain)
  5504   hence "y : fst ` rel_interior S" by (metis Domain_iff fst_eq_Domain)
  5505   hence "y : rel_interior {t. f t ~= {}}" using h1 by auto
  5505   hence "y : rel_interior {t. f t ~= {}}" using h1 by auto
  5526 qed
  5526 qed
  5527 
  5527 
  5528 lemma rel_interior_convex_cone_aux:
  5528 lemma rel_interior_convex_cone_aux:
  5529 fixes S :: "('m::euclidean_space) set"
  5529 fixes S :: "('m::euclidean_space) set"
  5530 assumes "convex S"
  5530 assumes "convex S"
  5531 shows "(c,x) : rel_interior (cone hull ({(1 :: real)} <*> S)) <-> 
  5531 shows "(c,x) : rel_interior (cone hull ({(1 :: real)} <*> S)) <->
  5532        c>0 & x : ((op *\<^sub>R c) ` (rel_interior S))"
  5532        c>0 & x : ((op *\<^sub>R c) ` (rel_interior S))"
  5533 proof-
  5533 proof-
  5534 { assume "S={}" hence ?thesis by (simp add: rel_interior_empty cone_hull_empty) } 
  5534 { assume "S={}" hence ?thesis by (simp add: rel_interior_empty cone_hull_empty) }
  5535 moreover
  5535 moreover
  5536 { assume "S ~= {}" from this obtain s where "s : S" by auto
  5536 { assume "S ~= {}" from this obtain s where "s : S" by auto
  5537 have conv: "convex ({(1 :: real)} <*> S)" using convex_direct_sum[of "{(1 :: real)}" S] 
  5537 have conv: "convex ({(1 :: real)} <*> S)" using convex_direct_sum[of "{(1 :: real)}" S]
  5538    assms convex_singleton[of "1 :: real"] by auto
  5538    assms convex_singleton[of "1 :: real"] by auto
  5539 def f == "(%y. {z. (y,z) : cone hull ({(1 :: real)} <*> S)})"
  5539 def f == "(%y. {z. (y,z) : cone hull ({(1 :: real)} <*> S)})"
  5540 hence *: "(c, x) : rel_interior (cone hull ({(1 :: real)} <*> S)) =
  5540 hence *: "(c, x) : rel_interior (cone hull ({(1 :: real)} <*> S)) =
  5541       (c : rel_interior {y. f y ~= {}} & x : rel_interior (f c))"
  5541       (c : rel_interior {y. f y ~= {}} & x : rel_interior (f c))"
  5542   apply (subst rel_interior_projection[of "cone hull ({(1 :: real)} <*> S)" f c x])
  5542   apply (subst rel_interior_projection[of "cone hull ({(1 :: real)} <*> S)" f c x])
  5559 
  5559 
  5560 
  5560 
  5561 lemma rel_interior_convex_cone:
  5561 lemma rel_interior_convex_cone:
  5562 fixes S :: "('m::euclidean_space) set"
  5562 fixes S :: "('m::euclidean_space) set"
  5563 assumes "convex S"
  5563 assumes "convex S"
  5564 shows "rel_interior (cone hull ({(1 :: real)} <*> S)) = 
  5564 shows "rel_interior (cone hull ({(1 :: real)} <*> S)) =
  5565        {(c,c *\<^sub>R x) |c x. c>0 & x : (rel_interior S)}"
  5565        {(c,c *\<^sub>R x) |c x. c>0 & x : (rel_interior S)}"
  5566 (is "?lhs=?rhs")
  5566 (is "?lhs=?rhs")
  5567 proof-
  5567 proof-
  5568 { fix z assume "z:?lhs" 
  5568 { fix z assume "z:?lhs"
  5569   have *: "z=(fst z,snd z)" by auto 
  5569   have *: "z=(fst z,snd z)" by auto
  5570   have "z:?rhs" using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms `z:?lhs` apply auto
  5570   have "z:?rhs" using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms `z:?lhs` apply auto
  5571      apply (rule_tac x="fst z" in exI) apply (rule_tac x="x" in exI) using * by auto
  5571      apply (rule_tac x="fst z" in exI) apply (rule_tac x="x" in exI) using * by auto
  5572 }
  5572 }
  5573 moreover
  5573 moreover
  5574 { fix z assume "z:?rhs" hence "z:?lhs" 
  5574 { fix z assume "z:?rhs" hence "z:?lhs"
  5575   using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms by auto
  5575   using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms by auto
  5576 }
  5576 }
  5577 ultimately show ?thesis by blast
  5577 ultimately show ?thesis by blast
  5578 qed
  5578 qed
  5579 
  5579 
  5580 lemma convex_hull_finite_union:
  5580 lemma convex_hull_finite_union:
  5581 assumes "finite I"
  5581 assumes "finite I"
  5582 assumes "!i:I. (convex (S i) & (S i) ~= {})"
  5582 assumes "!i:I. (convex (S i) & (S i) ~= {})"
  5583 shows "convex hull (Union (S ` I)) = 
  5583 shows "convex hull (Union (S ` I)) =
  5584        {setsum (%i. c i *\<^sub>R s i) I |c s. (!i:I. c i >= 0) & (setsum c I = 1) & (!i:I. s i : S i)}"
  5584        {setsum (%i. c i *\<^sub>R s i) I |c s. (!i:I. c i >= 0) & (setsum c I = 1) & (!i:I. s i : S i)}"
  5585   (is "?lhs = ?rhs")
  5585   (is "?lhs = ?rhs")
  5586 proof-
  5586 proof-
  5587 { fix x assume "x : ?rhs" 
  5587 { fix x assume "x : ?rhs"
  5588   from this obtain c s 
  5588   from this obtain c s
  5589     where *: "setsum (%i. c i *\<^sub>R s i) I=x" "(setsum c I = 1)"
  5589     where *: "setsum (%i. c i *\<^sub>R s i) I=x" "(setsum c I = 1)"
  5590      "(!i:I. c i >= 0) & (!i:I. s i : S i)" by auto
  5590      "(!i:I. c i >= 0) & (!i:I. s i : S i)" by auto
  5591   hence "!i:I. s i : convex hull (Union (S ` I))" using hull_subset[of "Union (S ` I)" convex] by auto
  5591   hence "!i:I. s i : convex hull (Union (S ` I))" using hull_subset[of "Union (S ` I)" convex] by auto
  5592   hence "x : ?lhs" unfolding *(1)[symmetric]
  5592   hence "x : ?lhs" unfolding *(1)[symmetric]
  5593      apply (subst convex_setsum[of I "convex hull Union (S ` I)" c s])
  5593      apply (subst convex_setsum[of I "convex hull Union (S ` I)" c s])
  5594      using * assms convex_convex_hull by auto
  5594      using * assms convex_convex_hull by auto
  5595 } hence "?lhs >= ?rhs" by auto
  5595 } hence "?lhs >= ?rhs" by auto
  5596 
  5596 
  5597 { fix i assume "i:I"
  5597 { fix i assume "i:I"
  5598     from this assms have "EX p. p : S i" by auto
  5598     from this assms have "EX p. p : S i" by auto
  5599 } 
  5599 }
  5600 from this obtain p where p_def: "!i:I. p i : S i" by metis
  5600 from this obtain p where p_def: "!i:I. p i : S i" by metis
  5601 
  5601 
  5602 { fix i assume "i:I"
  5602 { fix i assume "i:I"
  5603   { fix x assume "x : S i"
  5603   { fix x assume "x : S i"
  5604     def c == "(%j. if (j=i) then (1::real) else 0)"
  5604     def c == "(%j. if (j=i) then (1::real) else 0)"
  5605     hence *: "setsum c I = 1" using `finite I` `i:I` setsum_delta[of I i "(%(j::'a). (1::real))"] by auto
  5605     hence *: "setsum c I = 1" using `finite I` `i:I` setsum_delta[of I i "(%(j::'a). (1::real))"] by auto
  5606     def s == "(%j. if (j=i) then x else p j)"
  5606     def s == "(%j. if (j=i) then x else p j)"
  5607     hence "!j. c j *\<^sub>R s j = (if (j=i) then x else 0)" using c_def by (auto simp add: algebra_simps)
  5607     hence "!j. c j *\<^sub>R s j = (if (j=i) then x else 0)" using c_def by (auto simp add: algebra_simps)
  5608     hence "x = setsum (%i. c i *\<^sub>R s i) I"
  5608     hence "x = setsum (%i. c i *\<^sub>R s i) I"
  5609        using s_def c_def `finite I` `i:I` setsum_delta[of I i "(%(j::'a). x)"] by auto 
  5609        using s_def c_def `finite I` `i:I` setsum_delta[of I i "(%(j::'a). x)"] by auto
  5610     hence "x : ?rhs" apply auto
  5610     hence "x : ?rhs" apply auto
  5611       apply (rule_tac x="c" in exI) 
  5611       apply (rule_tac x="c" in exI)
  5612       apply (rule_tac x="s" in exI) using * c_def s_def p_def `x : S i` by auto 
  5612       apply (rule_tac x="s" in exI) using * c_def s_def p_def `x : S i` by auto
  5613   } hence "?rhs >= S i" by auto
  5613   } hence "?rhs >= S i" by auto
  5614 } hence *: "?rhs >= Union (S ` I)" by auto
  5614 } hence *: "?rhs >= Union (S ` I)" by auto
  5615 
  5615 
  5616 { fix u v assume uv: "(u :: real)>=0 & v>=0 & u+v=1"
  5616 { fix u v assume uv: "(u :: real)>=0 & v>=0 & u+v=1"
  5617   fix x y assume xy: "(x : ?rhs) & (y : ?rhs)"
  5617   fix x y assume xy: "(x : ?rhs) & (y : ?rhs)"
  5622   def e == "(%i. u * (c i)+v * (d i))"
  5622   def e == "(%i. u * (c i)+v * (d i))"
  5623   have ge0: "!i:I. e i >= 0"  using e_def xc yc uv by (simp add: mult_nonneg_nonneg)
  5623   have ge0: "!i:I. e i >= 0"  using e_def xc yc uv by (simp add: mult_nonneg_nonneg)
  5624   have "setsum (%i. u * c i) I = u * setsum c I" by (simp add: setsum_right_distrib)
  5624   have "setsum (%i. u * c i) I = u * setsum c I" by (simp add: setsum_right_distrib)
  5625   moreover have "setsum (%i. v * d i) I = v * setsum d I" by (simp add: setsum_right_distrib)
  5625   moreover have "setsum (%i. v * d i) I = v * setsum d I" by (simp add: setsum_right_distrib)
  5626   ultimately have sum1: "setsum e I = 1" using e_def xc yc uv by (simp add: setsum_addf)
  5626   ultimately have sum1: "setsum e I = 1" using e_def xc yc uv by (simp add: setsum_addf)
  5627   def q == "(%i. if (e i = 0) then (p i) 
  5627   def q == "(%i. if (e i = 0) then (p i)
  5628                  else (u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i))"
  5628                  else (u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i))"
  5629   { fix i assume "i:I"
  5629   { fix i assume "i:I"
  5630     { assume "e i = 0" hence "q i : S i" using `i:I` p_def q_def by auto }
  5630     { assume "e i = 0" hence "q i : S i" using `i:I` p_def q_def by auto }
  5631     moreover
  5631     moreover
  5632     { assume "e i ~= 0" 
  5632     { assume "e i ~= 0"
  5633       hence "q i : S i" using mem_convex_alt[of "S i" "s i" "t i" "u * (c i)" "v * (d i)"] 
  5633       hence "q i : S i" using mem_convex_alt[of "S i" "s i" "t i" "u * (c i)" "v * (d i)"]
  5634          mult_nonneg_nonneg[of u "c i"] mult_nonneg_nonneg[of v "d i"]
  5634          mult_nonneg_nonneg[of u "c i"] mult_nonneg_nonneg[of v "d i"]
  5635          assms q_def e_def `i:I` `e i ~= 0` xc yc uv by auto
  5635          assms q_def e_def `i:I` `e i ~= 0` xc yc uv by auto
  5636     } ultimately have "q i : S i" by auto
  5636     } ultimately have "q i : S i" by auto
  5637   } hence qs: "!i:I. q i : S i" by auto
  5637   } hence qs: "!i:I. q i : S i" by auto
  5638   { fix i assume "i:I"
  5638   { fix i assume "i:I"
  5639     { assume "e i = 0" 
  5639     { assume "e i = 0"
  5640       have ge: "u * (c i) >= 0 & v * (d i) >= 0" using xc yc uv `i:I` by (simp add: mult_nonneg_nonneg)
  5640       have ge: "u * (c i) >= 0 & v * (d i) >= 0" using xc yc uv `i:I` by (simp add: mult_nonneg_nonneg)
  5641       moreover hence "u * (c i) <= 0 & v * (d i) <= 0" using `e i = 0` e_def `i:I` by simp 
  5641       moreover hence "u * (c i) <= 0 & v * (d i) <= 0" using `e i = 0` e_def `i:I` by simp
  5642       ultimately have "u * (c i) = 0 & v * (d i) = 0" by auto
  5642       ultimately have "u * (c i) = 0 & v * (d i) = 0" by auto
  5643       hence "(u * (c i))*\<^sub>R (s i)+(v * (d i))*\<^sub>R (t i) = (e i) *\<^sub>R (q i)"
  5643       hence "(u * (c i))*\<^sub>R (s i)+(v * (d i))*\<^sub>R (t i) = (e i) *\<^sub>R (q i)"
  5644          using `e i = 0` by auto
  5644          using `e i = 0` by auto
  5645     }
  5645     }
  5646     moreover
  5646     moreover
  5649          using q_def by auto
  5649          using q_def by auto
  5650       hence "e i *\<^sub>R ((u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i))
  5650       hence "e i *\<^sub>R ((u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i))
  5651              = (e i) *\<^sub>R (q i)" by auto
  5651              = (e i) *\<^sub>R (q i)" by auto
  5652       hence "(u * (c i))*\<^sub>R (s i)+(v * (d i))*\<^sub>R (t i) = (e i) *\<^sub>R (q i)"
  5652       hence "(u * (c i))*\<^sub>R (s i)+(v * (d i))*\<^sub>R (t i) = (e i) *\<^sub>R (q i)"
  5653          using `e i ~= 0` by (simp add: algebra_simps)
  5653          using `e i ~= 0` by (simp add: algebra_simps)
  5654     } ultimately have 
  5654     } ultimately have
  5655       "(u * (c i))*\<^sub>R (s i)+(v * (d i))*\<^sub>R (t i) = (e i) *\<^sub>R (q i)" by blast
  5655       "(u * (c i))*\<^sub>R (s i)+(v * (d i))*\<^sub>R (t i) = (e i) *\<^sub>R (q i)" by blast
  5656   } hence *: "!i:I.
  5656   } hence *: "!i:I.
  5657     (u * (c i))*\<^sub>R (s i)+(v * (d i))*\<^sub>R (t i) = (e i) *\<^sub>R (q i)" by auto
  5657     (u * (c i))*\<^sub>R (s i)+(v * (d i))*\<^sub>R (t i) = (e i) *\<^sub>R (q i)" by auto
  5658   have "u *\<^sub>R x + v *\<^sub>R y =
  5658   have "u *\<^sub>R x + v *\<^sub>R y =
  5659        setsum (%i. (u * (c i))*\<^sub>R (s i)+(v * (d i))*\<^sub>R (t i)) I"
  5659        setsum (%i. (u * (c i))*\<^sub>R (s i)+(v * (d i))*\<^sub>R (t i)) I"
  5660           using xc yc by (simp add: algebra_simps scaleR_right.setsum setsum_addf)
  5660           using xc yc by (simp add: algebra_simps scaleR_right.setsum setsum_addf)
  5661   also have "...=setsum (%i. (e i) *\<^sub>R (q i)) I" using * by auto
  5661   also have "...=setsum (%i. (e i) *\<^sub>R (q i)) I" using * by auto
  5662   finally have "u *\<^sub>R x + v *\<^sub>R y = setsum (%i. (e i) *\<^sub>R (q i)) I" by auto
  5662   finally have "u *\<^sub>R x + v *\<^sub>R y = setsum (%i. (e i) *\<^sub>R (q i)) I" by auto
  5663   hence "u *\<^sub>R x + v *\<^sub>R y : ?rhs" using ge0 sum1 qs by auto
  5663   hence "u *\<^sub>R x + v *\<^sub>R y : ?rhs" using ge0 sum1 qs by auto
  5664 } hence "convex ?rhs" unfolding convex_def by auto
  5664 } hence "convex ?rhs" unfolding convex_def by auto
  5665 from this show ?thesis using `?lhs >= ?rhs` * 
  5665 from this show ?thesis using `?lhs >= ?rhs` *
  5666    hull_minimal[of "Union (S ` I)" "?rhs" "convex"] by blast
  5666    hull_minimal[of "Union (S ` I)" "?rhs" "convex"] by blast
  5667 qed
  5667 qed
  5668 
  5668 
  5669 lemma convex_hull_union_two:
  5669 lemma convex_hull_union_two:
  5670 fixes S T :: "('m::euclidean_space) set"
  5670 fixes S T :: "('m::euclidean_space) set"
  5677 have "Union (s ` I) = S Un T" using s_def I_def by auto
  5677 have "Union (s ` I) = S Un T" using s_def I_def by auto
  5678 hence "convex hull (Union (s ` I)) = convex hull (S Un T)" by auto
  5678 hence "convex hull (Union (s ` I)) = convex hull (S Un T)" by auto
  5679 moreover have "convex hull Union (s ` I) =
  5679 moreover have "convex hull Union (s ` I) =
  5680     {SUM i:I. c i *\<^sub>R sa i |c sa. (ALL i:I. 0 <= c i) & setsum c I = 1 & (ALL i:I. sa i : s i)}"
  5680     {SUM i:I. c i *\<^sub>R sa i |c sa. (ALL i:I. 0 <= c i) & setsum c I = 1 & (ALL i:I. sa i : s i)}"
  5681     apply (subst convex_hull_finite_union[of I s]) using assms s_def I_def by auto
  5681     apply (subst convex_hull_finite_union[of I s]) using assms s_def I_def by auto
  5682 moreover have 
  5682 moreover have
  5683   "{SUM i:I. c i *\<^sub>R sa i |c sa. (ALL i:I. 0 <= c i) & setsum c I = 1 & (ALL i:I. sa i : s i)} <=
  5683   "{SUM i:I. c i *\<^sub>R sa i |c sa. (ALL i:I. 0 <= c i) & setsum c I = 1 & (ALL i:I. sa i : s i)} <=
  5684   ?rhs"
  5684   ?rhs"
  5685   using s_def I_def by auto
  5685   using s_def I_def by auto
  5686 ultimately have "?lhs<=?rhs" by auto 
  5686 ultimately have "?lhs<=?rhs" by auto
  5687 { fix x assume "x : ?rhs" 
  5687 { fix x assume "x : ?rhs"
  5688   from this obtain u v s t 
  5688   from this obtain u v s t
  5689     where *: "x=u *\<^sub>R s + v *\<^sub>R t & u>=0 & v>=0 & u+v=1 & s:S & t:T" by auto
  5689     where *: "x=u *\<^sub>R s + v *\<^sub>R t & u>=0 & v>=0 & u+v=1 & s:S & t:T" by auto
  5690   hence "x : convex hull {s,t}" using convex_hull_2[of s t] by auto
  5690   hence "x : convex hull {s,t}" using convex_hull_2[of s t] by auto
  5691   hence "x : convex hull (S Un T)" using * hull_mono[of "{s, t}" "S Un T"] by auto
  5691   hence "x : convex hull (S Un T)" using * hull_mono[of "{s, t}" "S Un T"] by auto
  5692 } hence "?lhs >= ?rhs" by blast
  5692 } hence "?lhs >= ?rhs" by blast
  5693 from this show ?thesis using `?lhs<=?rhs` by auto
  5693 from this show ?thesis using `?lhs<=?rhs` by auto