diff -r 493b818e8e10 -r fad29d2a17a5 src/HOL/Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Analysis/Convex_Euclidean_Space.thy Wed May 02 13:49:38 2018 +0200 +++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy Thu May 03 15:07:14 2018 +0200 @@ -68,8 +68,7 @@ then have "span (cball 0 e) = (UNIV :: 'n::euclidean_space set)" by auto then show ?thesis - using dim_span[of "cball (0 :: 'n::euclidean_space) e"] - by auto + using dim_span[of "cball (0 :: 'n::euclidean_space) e"] by (auto simp: dim_UNIV) qed lemma sum_not_0: "sum f A \ 0 \ \a \ A. f a \ 0" @@ -119,8 +118,7 @@ then have "x-a \ S" using assms by auto then have "x \ {a + v |v. v \ S}" apply auto - apply (rule exI[of _ "x-a"]) - apply simp + apply (rule exI[of _ "x-a"], simp) done then have "x \ ((\x. a+x) ` S)" by auto } @@ -1301,7 +1299,7 @@ proof - have *: "x - y + (y - z) = x - z" by auto show ?thesis unfolding dist_norm norm_triangle_eq[of "x - y" "y - z", unfolded *] - by (auto simp add:norm_minus_commute) + by (auto simp:norm_minus_commute) qed @@ -1317,7 +1315,7 @@ unfolding affine_def by auto lemma affine_sing [iff]: "affine {x}" - unfolding affine_alt by (auto simp add: scaleR_left_distrib [symmetric]) + unfolding affine_alt by (auto simp: scaleR_left_distrib [symmetric]) lemma affine_UNIV [iff]: "affine UNIV" unfolding affine_def by auto @@ -1350,293 +1348,200 @@ lemma affine: fixes V::"'a::real_vector set" shows "affine V \ - (\s u. finite s \ s \ {} \ s \ V \ sum u s = 1 \ (sum (\x. (u x) *\<^sub>R x)) s \ V)" - unfolding affine_def - apply rule - apply(rule, rule, rule) - apply(erule conjE)+ - defer - apply (rule, rule, rule, rule, rule) + (\S u. finite S \ S \ {} \ S \ V \ sum u S = 1 \ (\x\S. u x *\<^sub>R x) \ V)" proof - - fix x y u v - assume as: "x \ V" "y \ V" "u + v = (1::real)" - "\s u. finite s \ s \ {} \ s \ V \ sum u s = 1 \ (\x\s. u x *\<^sub>R x) \ V" - then show "u *\<^sub>R x + v *\<^sub>R y \ V" - apply (cases "x = y") - using as(4)[THEN spec[where x="{x,y}"], THEN spec[where x="\w. if w = x then u else v"]] - and as(1-3) - apply (auto simp add: scaleR_left_distrib[symmetric]) - done -next - fix s u - assume as: "\x\V. \y\V. \u v. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ V" - "finite s" "s \ {}" "s \ V" "sum u s = (1::real)" - define n where "n = card s" - have "card s = 0 \ card s = 1 \ card s = 2 \ card s > 2" by auto - then show "(\x\s. u x *\<^sub>R x) \ V" - proof (auto simp only: disjE) - assume "card s = 2" - then have "card s = Suc (Suc 0)" - by auto - then obtain a b where "s = {a, b}" - unfolding card_Suc_eq by auto + have "u *\<^sub>R x + v *\<^sub>R y \ V" if "x \ V" "y \ V" "u + v = (1::real)" + and *: "\S u. \finite S; S \ {}; S \ V; sum u S = 1\ \ (\x\S. u x *\<^sub>R x) \ V" for x y u v + proof (cases "x = y") + case True + then show ?thesis + using that by (metis scaleR_add_left scaleR_one) + next + case False then show ?thesis - using as(1)[THEN bspec[where x=a], THEN bspec[where x=b]] using as(4,5) - by (auto simp add: sum_clauses(2)) - next - assume "card s > 2" - then show ?thesis using as and n_def - proof (induct n arbitrary: u s) - case 0 - then show ?case by auto + using that *[of "{x,y}" "\w. if w = x then u else v"] by auto + qed + moreover have "(\x\S. u x *\<^sub>R x) \ V" + if *: "\x y u v. \x\V; y\V; u + v = 1\ \ u *\<^sub>R x + v *\<^sub>R y \ V" + and "finite S" "S \ {}" "S \ V" "sum u S = 1" for S u + proof - + define n where "n = card S" + consider "card S = 0" | "card S = 1" | "card S = 2" | "card S > 2" by linarith + then show "(\x\S. u x *\<^sub>R x) \ V" + proof cases + assume "card S = 1" + then obtain a where "S={a}" + by (auto simp: card_Suc_eq) + then show ?thesis + using that by simp + next + assume "card S = 2" + then obtain a b where "S = {a, b}" + by (metis Suc_1 card_1_singletonE card_Suc_eq) + then show ?thesis + using *[of a b] that + by (auto simp: sum_clauses(2)) next - case (Suc n) - fix s :: "'a set" and u :: "'a \ real" - assume IA: - "\u s. \2 < card s; \x\V. \y\V. \u v. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ V; finite s; - s \ {}; s \ V; sum u s = 1; n = card s \ \ (\x\s. u x *\<^sub>R x) \ V" - and as: - "Suc n = card s" "2 < card s" "\x\V. \y\V. \u v. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ V" - "finite s" "s \ {}" "s \ V" "sum u s = 1" - have "\x\s. u x \ 1" - proof (rule ccontr) - assume "\ ?thesis" - then have "sum u s = real_of_nat (card s)" - unfolding card_eq_sum by auto - then show False - using as(7) and \card s > 2\ - by (metis One_nat_def less_Suc0 Zero_not_Suc of_nat_1 of_nat_eq_iff numeral_2_eq_2) + assume "card S > 2" + then show ?thesis using that n_def + proof (induct n arbitrary: u S) + case 0 + then show ?case by auto + next + case (Suc n u S) + have "sum u S = card S" if "\ (\x\S. u x \ 1)" + using that unfolding card_eq_sum by auto + with Suc.prems obtain x where "x \ S" and x: "u x \ 1" by force + have c: "card (S - {x}) = card S - 1" + by (simp add: Suc.prems(3) \x \ S\) + have "sum u (S - {x}) = 1 - u x" + by (simp add: Suc.prems sum_diff1_ring \x \ S\) + with x have eq1: "inverse (1 - u x) * sum u (S - {x}) = 1" + by auto + have inV: "(\y\S - {x}. (inverse (1 - u x) * u y) *\<^sub>R y) \ V" + proof (cases "card (S - {x}) > 2") + case True + then have S: "S - {x} \ {}" "card (S - {x}) = n" + using Suc.prems c by force+ + show ?thesis + proof (rule Suc.hyps) + show "(\a\S - {x}. inverse (1 - u x) * u a) = 1" + by (auto simp: eq1 sum_distrib_left[symmetric]) + qed (use S Suc.prems True in auto) + next + case False + then have "card (S - {x}) = Suc (Suc 0)" + using Suc.prems c by auto + then obtain a b where ab: "(S - {x}) = {a, b}" "a\b" + unfolding card_Suc_eq by auto + then show ?thesis + using eq1 \S \ V\ + by (auto simp: sum_distrib_left distrib_left intro!: Suc.prems(2)[of a b]) + qed + have "u x + (1 - u x) = 1 \ + u x *\<^sub>R x + (1 - u x) *\<^sub>R ((\y\S - {x}. u y *\<^sub>R y) /\<^sub>R (1 - u x)) \ V" + by (rule Suc.prems) (use \x \ S\ Suc.prems inV in \auto simp: scaleR_right.sum\) + moreover have "(\a\S. u a *\<^sub>R a) = u x *\<^sub>R x + (\a\S - {x}. u a *\<^sub>R a)" + by (meson Suc.prems(3) sum.remove \x \ S\) + ultimately show "(\x\S. u x *\<^sub>R x) \ V" + by (simp add: x) qed - then obtain x where x:"x \ s" "u x \ 1" by auto - - have c: "card (s - {x}) = card s - 1" - apply (rule card_Diff_singleton) - using \x\s\ as(4) - apply auto - done - have *: "s = insert x (s - {x})" "finite (s - {x})" - using \x\s\ and as(4) by auto - have **: "sum u (s - {x}) = 1 - u x" - using sum_clauses(2)[OF *(2), of u x, unfolded *(1)[symmetric] as(7)] by auto - have ***: "inverse (1 - u x) * sum u (s - {x}) = 1" - unfolding ** using \u x \ 1\ by auto - have "(\xa\s - {x}. (inverse (1 - u x) * u xa) *\<^sub>R xa) \ V" - proof (cases "card (s - {x}) > 2") - case True - then have "s - {x} \ {}" "card (s - {x}) = n" - unfolding c and as(1)[symmetric] - proof (rule_tac ccontr) - assume "\ s - {x} \ {}" - then have "card (s - {x}) = 0" unfolding card_0_eq[OF *(2)] by simp - then show False using True by auto - qed auto - then show ?thesis - apply (rule_tac IA[of "s - {x}" "\y. (inverse (1 - u x) * u y)"]) - unfolding sum_distrib_left[symmetric] - using as and *** and True - apply auto - done - next - case False - then have "card (s - {x}) = Suc (Suc 0)" - using as(2) and c by auto - then obtain a b where "(s - {x}) = {a, b}" "a\b" - unfolding card_Suc_eq by auto - then show ?thesis - using as(3)[THEN bspec[where x=a], THEN bspec[where x=b]] - using *** *(2) and \s \ V\ - unfolding sum_distrib_left - by (auto simp add: sum_clauses(2)) - qed - then have "u x + (1 - u x) = 1 \ - u x *\<^sub>R x + (1 - u x) *\<^sub>R ((\xa\s - {x}. u xa *\<^sub>R xa) /\<^sub>R (1 - u x)) \ V" - apply - - apply (rule as(3)[rule_format]) - unfolding Real_Vector_Spaces.scaleR_right.sum - using x(1) as(6) - apply auto - done - then show "(\x\s. u x *\<^sub>R x) \ V" - unfolding scaleR_scaleR[symmetric] and scaleR_right.sum [symmetric] - apply (subst *) - unfolding sum_clauses(2)[OF *(2)] - using \u x \ 1\ - apply auto - done - qed - next - assume "card s = 1" - then obtain a where "s={a}" - by (auto simp add: card_Suc_eq) - then show ?thesis - using as(4,5) by simp - qed (insert \s\{}\ \finite s\, auto) -qed + qed (use \S\{}\ \finite S\ in auto) + qed + ultimately show ?thesis + unfolding affine_def by meson +qed + lemma affine_hull_explicit: - "affine hull p = - {y. \s u. finite s \ s \ {} \ s \ p \ sum u s = 1 \ sum (\v. (u v) *\<^sub>R v) s = y}" - apply (rule hull_unique) - apply (subst subset_eq) - prefer 3 - apply rule - unfolding mem_Collect_eq - apply (erule exE)+ - apply (erule conjE)+ - prefer 2 - apply rule -proof - - fix x - assume "x\p" - then show "\s u. finite s \ s \ {} \ s \ p \ sum u s = 1 \ (\v\s. u v *\<^sub>R v) = x" - apply (rule_tac x="{x}" in exI) - apply (rule_tac x="\x. 1" in exI) - apply auto - done -next - fix t x s u - assume as: "p \ t" "affine t" "finite s" "s \ {}" - "s \ p" "sum u s = 1" "(\v\s. u v *\<^sub>R v) = x" - then show "x \ t" - using as(2)[unfolded affine, THEN spec[where x=s], THEN spec[where x=u]] - by auto -next - show "affine {y. \s u. finite s \ s \ {} \ s \ p \ sum u s = 1 \ (\v\s. u v *\<^sub>R v) = y}" + "affine hull p = {y. \S u. finite S \ S \ {} \ S \ p \ sum u S = 1 \ sum (\v. u v *\<^sub>R v) S = y}" + (is "_ = ?rhs") +proof (rule hull_unique) + show "p \ ?rhs" + proof (intro subsetI CollectI exI conjI) + show "\x. sum (\z. 1) {x} = 1" + by auto + qed auto + show "?rhs \ T" if "p \ T" "affine T" for T + using that unfolding affine by blast + show "affine ?rhs" unfolding affine_def - apply (rule, rule, rule, rule, rule) - unfolding mem_Collect_eq - proof - - fix u v :: real + proof clarify + fix u v :: real and sx ux sy uy assume uv: "u + v = 1" - fix x - assume "\s u. finite s \ s \ {} \ s \ p \ sum u s = 1 \ (\v\s. u v *\<^sub>R v) = x" - then obtain sx ux where - x: "finite sx" "sx \ {}" "sx \ p" "sum ux sx = 1" "(\v\sx. ux v *\<^sub>R v) = x" - by auto - fix y - assume "\s u. finite s \ s \ {} \ s \ p \ sum u s = 1 \ (\v\s. u v *\<^sub>R v) = y" - then obtain sy uy where - y: "finite sy" "sy \ {}" "sy \ p" "sum uy sy = 1" "(\v\sy. uy v *\<^sub>R v) = y" by auto - have xy: "finite (sx \ sy)" - using x(1) y(1) by auto + and x: "finite sx" "sx \ {}" "sx \ p" "sum ux sx = (1::real)" + and y: "finite sy" "sy \ {}" "sy \ p" "sum uy sy = (1::real)" have **: "(sx \ sy) \ sx = sx" "(sx \ sy) \ sy = sy" by auto - show "\s ua. finite s \ s \ {} \ s \ p \ - sum ua s = 1 \ (\v\s. ua v *\<^sub>R v) = u *\<^sub>R x + v *\<^sub>R y" - apply (rule_tac x="sx \ sy" in exI) - apply (rule_tac x="\a. (if a\sx then u * ux a else 0) + (if a\sy then v * uy a else 0)" in exI) - unfolding scaleR_left_distrib sum.distrib if_smult scaleR_zero_left - ** sum.inter_restrict[OF xy, symmetric] - unfolding scaleR_scaleR[symmetric] Real_Vector_Spaces.scaleR_right.sum [symmetric] - and sum_distrib_left[symmetric] - unfolding x y - using x(1-3) y(1-3) uv - apply simp - done + show "\S w. finite S \ S \ {} \ S \ p \ + sum w S = 1 \ (\v\S. w v *\<^sub>R v) = u *\<^sub>R (\v\sx. ux v *\<^sub>R v) + v *\<^sub>R (\v\sy. uy v *\<^sub>R v)" + proof (intro exI conjI) + show "finite (sx \ sy)" + using x y by auto + show "sum (\i. (if i\sx then u * ux i else 0) + (if i\sy then v * uy i else 0)) (sx \ sy) = 1" + using x y uv + by (simp add: sum_Un sum.distrib sum.inter_restrict[symmetric] sum_distrib_left [symmetric] **) + have "(\i\sx \ sy. ((if i \ sx then u * ux i else 0) + (if i \ sy then v * uy i else 0)) *\<^sub>R i) + = (\i\sx. (u * ux i) *\<^sub>R i) + (\i\sy. (v * uy i) *\<^sub>R i)" + using x y + unfolding scaleR_left_distrib scaleR_zero_left if_smult + by (simp add: sum_Un sum.distrib sum.inter_restrict[symmetric] **) + also have "\ = u *\<^sub>R (\v\sx. ux v *\<^sub>R v) + v *\<^sub>R (\v\sy. uy v *\<^sub>R v)" + unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] by blast + finally show "(\i\sx \ sy. ((if i \ sx then u * ux i else 0) + (if i \ sy then v * uy i else 0)) *\<^sub>R i) + = u *\<^sub>R (\v\sx. ux v *\<^sub>R v) + v *\<^sub>R (\v\sy. uy v *\<^sub>R v)" . + qed (use x y in auto) qed qed lemma affine_hull_finite: - assumes "finite s" - shows "affine hull s = {y. \u. sum u s = 1 \ sum (\v. u v *\<^sub>R v) s = y}" - unfolding affine_hull_explicit and set_eq_iff and mem_Collect_eq - apply (rule, rule) - apply (erule exE)+ - apply (erule conjE)+ - defer - apply (erule exE) - apply (erule conjE) + assumes "finite S" + shows "affine hull S = {y. \u. sum u S = 1 \ sum (\v. u v *\<^sub>R v) S = y}" proof - - fix x u - assume "sum u s = 1" "(\v\s. u v *\<^sub>R v) = x" - then show "\sa u. finite sa \ - \ (\x. (x \ sa) = (x \ {})) \ sa \ s \ sum u sa = 1 \ (\v\sa. u v *\<^sub>R v) = x" - apply (rule_tac x=s in exI, rule_tac x=u in exI) - using assms - apply auto - done -next - fix x t u - assume "t \ s" - then have *: "s \ t = t" - by auto - assume "finite t" "\ (\x. (x \ t) = (x \ {}))" "sum u t = 1" "(\v\t. u v *\<^sub>R v) = x" - then show "\u. sum u s = 1 \ (\v\s. u v *\<^sub>R v) = x" - apply (rule_tac x="\x. if x\t then u x else 0" in exI) - unfolding if_smult scaleR_zero_left and sum.inter_restrict[OF assms, symmetric] and * - apply auto - done + have *: "\h. sum h S = 1 \ (\v\S. h v *\<^sub>R v) = x" + if "F \ S" "finite F" "F \ {}" and sum: "sum u F = 1" and x: "(\v\F. u v *\<^sub>R v) = x" for x F u + proof - + have "S \ F = F" + using that by auto + show ?thesis + proof (intro exI conjI) + show "(\x\S. if x \ F then u x else 0) = 1" + by (metis (mono_tags, lifting) \S \ F = F\ assms sum.inter_restrict sum) + show "(\v\S. (if v \ F then u v else 0) *\<^sub>R v) = x" + by (simp add: if_smult cong: if_cong) (metis (no_types) \S \ F = F\ assms sum.inter_restrict x) + qed + qed + show ?thesis + unfolding affine_hull_explicit using assms + by (fastforce dest: *) qed subsubsection%unimportant \Stepping theorems and hence small special cases\ lemma affine_hull_empty[simp]: "affine hull {} = {}" - by (rule hull_unique) auto - -(*could delete: it simply rewrites sum expressions, but it's used twice*) + by simp + lemma affine_hull_finite_step: fixes y :: "'a::real_vector" - shows - "(\u. sum u {} = w \ sum (\x. u x *\<^sub>R x) {} = y) \ w = 0 \ y = 0" (is ?th1) - and - "finite s \ - (\u. sum u (insert a s) = w \ sum (\x. u x *\<^sub>R x) (insert a s) = y) \ - (\v u. sum u s = w - v \ sum (\x. u x *\<^sub>R x) s = y - v *\<^sub>R a)" (is "_ \ ?lhs = ?rhs") + shows "finite S \ + (\u. sum u (insert a S) = w \ sum (\x. u x *\<^sub>R x) (insert a S) = y) \ + (\v u. sum u S = w - v \ sum (\x. u x *\<^sub>R x) S = y - v *\<^sub>R a)" (is "_ \ ?lhs = ?rhs") proof - - show ?th1 by simp - assume fin: "finite s" + assume fin: "finite S" show "?lhs = ?rhs" proof assume ?lhs - then obtain u where u: "sum u (insert a s) = w \ (\x\insert a s. u x *\<^sub>R x) = y" + then obtain u where u: "sum u (insert a S) = w \ (\x\insert a S. u x *\<^sub>R x) = y" by auto show ?rhs - proof (cases "a \ s") + proof (cases "a \ S") case True - then have *: "insert a s = s" by auto - show ?thesis - using u[unfolded *] - apply(rule_tac x=0 in exI) - apply auto - done + then show ?thesis + using u by (simp add: insert_absorb) (metis diff_zero real_vector.scale_zero_left) next case False - then show ?thesis - apply (rule_tac x="u a" in exI) - using u and fin - apply auto - done + show ?thesis + by (rule exI [where x="u a"]) (use u fin False in auto) qed next assume ?rhs - then obtain v u where vu: "sum u s = w - v" "(\x\s. u x *\<^sub>R x) = y - v *\<^sub>R a" + then obtain v u where vu: "sum u S = w - v" "(\x\S. u x *\<^sub>R x) = y - v *\<^sub>R a" by auto have *: "\x M. (if x = a then v else M) *\<^sub>R x = (if x = a then v *\<^sub>R x else M *\<^sub>R x)" by auto show ?lhs - proof (cases "a \ s") + proof (cases "a \ S") case True - then show ?thesis - apply (rule_tac x="\x. (if x=a then v else 0) + u x" in exI) - unfolding sum_clauses(2)[OF fin] - apply simp - unfolding scaleR_left_distrib and sum.distrib - unfolding vu and * and scaleR_zero_left - apply (auto simp add: sum.delta[OF fin]) - done + show ?thesis + by (rule exI [where x="\x. (if x=a then v else 0) + u x"]) + (simp add: True scaleR_left_distrib sum.distrib sum_clauses fin vu * cong: if_cong) next case False - then have **: - "\x. x \ s \ u x = (if x = a then v else u x)" - "\x. x \ s \ u x *\<^sub>R x = (if x = a then v *\<^sub>R x else u x *\<^sub>R x)" by auto - from False show ?thesis - apply (rule_tac x="\x. if x=a then v else u x" in exI) - unfolding sum_clauses(2)[OF fin] and * using vu - using sum.cong [of s _ "\x. u x *\<^sub>R x" "\x. if x = a then v *\<^sub>R x else u x *\<^sub>R x", OF _ **(2)] - using sum.cong [of s _ u "\x. if x = a then v else u x", OF _ **(1)] - apply auto - done + then show ?thesis + apply (rule_tac x="\x. if x=a then v else u x" in exI) + apply (simp add: vu sum_clauses(2)[OF fin] *) + by (simp add: sum_delta_notmem(3) vu) qed qed qed @@ -1652,7 +1557,7 @@ have "?lhs = {y. \u. sum u {a, b} = 1 \ (\v\{a, b}. u v *\<^sub>R v) = y}" using affine_hull_finite[of "{a,b}"] by auto also have "\ = {y. \v u. u b = 1 - v \ u b *\<^sub>R b = y - v *\<^sub>R a}" - by (simp add: affine_hull_finite_step(2)[of "{b}" a]) + by (simp add: affine_hull_finite_step[of "{b}" a]) also have "\ = ?rhs" unfolding * by auto finally show ?thesis by auto qed @@ -1667,12 +1572,9 @@ show ?thesis apply (simp add: affine_hull_finite affine_hull_finite_step) unfolding * - apply auto - apply (rule_tac x=v in exI) - apply (rule_tac x=va in exI) - apply auto - apply (rule_tac x=u in exI) - apply force + apply safe + apply (metis add.assoc) + apply (rule_tac x=u in exI, force) done qed @@ -1710,56 +1612,57 @@ subsubsection%unimportant \Some relations between affine hull and subspaces\ lemma affine_hull_insert_subset_span: - "affine hull (insert a s) \ {a + v| v . v \ span {x - a | x . x \ s}}" - unfolding subset_eq Ball_def - unfolding affine_hull_explicit span_explicit mem_Collect_eq - apply (rule, rule) - apply (erule exE)+ - apply (erule conjE)+ + "affine hull (insert a S) \ {a + v| v . v \ span {x - a | x . x \ S}}" proof - - fix x t u - assume as: "finite t" "t \ {}" "t \ insert a s" "sum u t = 1" "(\v\t. u v *\<^sub>R v) = x" - have "(\x. x - a) ` (t - {a}) \ {x - a |x. x \ s}" - using as(3) by auto - then show "\v. x = a + v \ (\S u. v = (\v\S. u v *\<^sub>R v) \ finite S \ S \ {x - a |x. x \ s} )" - apply (rule_tac x="x - a" in exI) - apply (rule conjI, simp) - apply (rule_tac x="(\x. x - a) ` (t - {a})" in exI) - apply (rule_tac x="\x. u (x + a)" in exI) - by (simp_all add: as sum.reindex[unfolded inj_on_def] scaleR_right_diff_distrib - sum_subtractf scaleR_left.sum[symmetric] sum_diff1 scaleR_left_diff_distrib) + have "\v T u. x = a + v \ (finite T \ T \ {x - a |x. x \ S} \ (\v\T. u v *\<^sub>R v) = v)" + if "finite F" "F \ {}" "F \ insert a S" "sum u F = 1" "(\v\F. u v *\<^sub>R v) = x" + for x F u + proof - + have *: "(\x. x - a) ` (F - {a}) \ {x - a |x. x \ S}" + using that by auto + show ?thesis + proof (intro exI conjI) + show "finite ((\x. x - a) ` (F - {a}))" + by (simp add: that(1)) + show "(\v\(\x. x - a) ` (F - {a}). u(v+a) *\<^sub>R v) = x-a" + by (simp add: sum.reindex[unfolded inj_on_def] algebra_simps + sum_subtractf scaleR_left.sum[symmetric] sum_diff1 that) + qed (use \F \ insert a S\ in auto) + qed + then show ?thesis + unfolding affine_hull_explicit span_explicit by auto qed lemma affine_hull_insert_span: - assumes "a \ s" - shows "affine hull (insert a s) = {a + v | v . v \ span {x - a | x. x \ s}}" - apply (rule, rule affine_hull_insert_subset_span) - unfolding subset_eq Ball_def - unfolding affine_hull_explicit and mem_Collect_eq -proof (rule, rule, erule exE, erule conjE) - fix y v - assume "y = a + v" "v \ span {x - a |x. x \ s}" - then obtain t u where obt: "finite t" "t \ {x - a |x. x \ s}" "a + (\v\t. u v *\<^sub>R v) = y" - unfolding span_explicit by auto - define f where "f = (\x. x + a) ` t" - have f: "finite f" "f \ s" "(\v\f. u (v - a) *\<^sub>R (v - a)) = y - a" - unfolding f_def using obt by (auto simp add: sum.reindex[unfolded inj_on_def]) - have *: "f \ {a} = {}" "f \ - {a} = f" - using f(2) assms by auto - show "\sa u. finite sa \ sa \ {} \ sa \ insert a s \ sum u sa = 1 \ (\v\sa. u v *\<^sub>R v) = y" - apply (rule_tac x = "insert a f" in exI) - apply (rule_tac x = "\x. if x=a then 1 - sum (\x. u (x - a)) f else u (x - a)" in exI) - using assms and f - unfolding sum_clauses(2)[OF f(1)] and if_smult - unfolding sum.If_cases[OF f(1), of "\x. x = a"] - apply (auto simp add: sum_subtractf scaleR_left.sum algebra_simps *) - done + assumes "a \ S" + shows "affine hull (insert a S) = {a + v | v . v \ span {x - a | x. x \ S}}" +proof - + have *: "\G u. finite G \ G \ {} \ G \ insert a S \ sum u G = 1 \ (\v\G. u v *\<^sub>R v) = y" + if "v \ span {x - a |x. x \ S}" "y = a + v" for y v + proof - + from that + obtain T u where u: "finite T" "T \ {x - a |x. x \ S}" "a + (\v\T. u v *\<^sub>R v) = y" + unfolding span_explicit by auto + define F where "F = (\x. x + a) ` T" + have F: "finite F" "F \ S" "(\v\F. u (v - a) *\<^sub>R (v - a)) = y - a" + unfolding F_def using u by (auto simp: sum.reindex[unfolded inj_on_def]) + have *: "F \ {a} = {}" "F \ - {a} = F" + using F assms by auto + show "\G u. finite G \ G \ {} \ G \ insert a S \ sum u G = 1 \ (\v\G. u v *\<^sub>R v) = y" + apply (rule_tac x = "insert a F" in exI) + apply (rule_tac x = "\x. if x=a then 1 - sum (\x. u (x - a)) F else u (x - a)" in exI) + using assms F + apply (auto simp: sum_clauses sum.If_cases if_smult sum_subtractf scaleR_left.sum algebra_simps *) + done + qed + show ?thesis + by (intro subset_antisym affine_hull_insert_subset_span) (auto simp: affine_hull_explicit dest!: *) qed lemma affine_hull_span: - assumes "a \ s" - shows "affine hull s = {a + v | v. v \ span {x - a | x. x \ s - {a}}}" - using affine_hull_insert_span[of a "s - {a}", unfolded insert_Diff[OF assms]] by auto + assumes "a \ S" + shows "affine hull S = {a + v | v. v \ span {x - a | x. x \ S - {a}}}" + using affine_hull_insert_span[of a "S - {a}", unfolded insert_Diff[OF assms]] by auto subsubsection%unimportant \Parallel affine sets\ @@ -1769,17 +1672,12 @@ lemma affine_parallel_expl_aux: fixes S T :: "'a::real_vector set" - assumes "\x. x \ S \ a + x \ T" + assumes "\x. x \ S \ a + x \ T" shows "T = (\x. a + x) ` S" proof - - { - fix x - assume "x \ T" - then have "( - a) + x \ S" - using assms by auto - then have "x \ ((\x. a + x) ` S)" - using imageI[of "-a+x" S "(\x. a+x)"] by auto - } + have "x \ ((\x. a + x) ` S)" if "x \ T" for x + using that + by (simp add: image_iff) (metis add.commute diff_add_cancel assms) moreover have "T \ (\x. a + x) ` S" using assms by auto ultimately show ?thesis by auto @@ -1791,9 +1689,7 @@ lemma affine_parallel_reflex: "affine_parallel S S" unfolding affine_parallel_def - apply (rule exI[of _ "0"]) - apply auto - done + using image_add_0 by blast lemma affine_parallel_commut: assumes "affine_parallel A B" @@ -2109,7 +2005,7 @@ shows "c *\<^sub>R x \ cone hull S" by (metis assms cone_cone_hull hull_inc mem_cone) -lemma%important cone_hull_expl: "cone hull S = {c *\<^sub>R x | c x. c \ 0 \ x \ S}" +proposition%important cone_hull_expl: "cone hull S = {c *\<^sub>R x | c x. c \ 0 \ x \ S}" (is "?lhs = ?rhs") proof%unimportant - { @@ -2135,8 +2031,7 @@ assume "x \ S" then have "1 *\<^sub>R x \ ?rhs" apply auto - apply (rule_tac x = 1 in exI) - apply auto + apply (rule_tac x = 1 in exI, auto) done then have "x \ ?rhs" by auto } @@ -2169,7 +2064,7 @@ then have "0 \ S \ (\c. c > 0 \ ( *\<^sub>R) c ` S = S)" using cone_iff[of S] assms by auto then have "0 \ closure S \ (\c. c > 0 \ ( *\<^sub>R) c ` closure S = closure S)" - using closure_subset by (auto simp add: closure_scaleR) + using closure_subset by (auto simp: closure_scaleR) then show ?thesis using False cone_iff[of "closure S"] by auto qed @@ -2194,66 +2089,60 @@ "~ affine_dependent s \ ~ affine_dependent(s - t)" by (meson Diff_subset affine_dependent_subset) -lemma%important affine_dependent_explicit: +proposition%important affine_dependent_explicit: "affine_dependent p \ - (\s u. finite s \ s \ p \ sum u s = 0 \ - (\v\s. u v \ 0) \ sum (\v. u v *\<^sub>R v) s = 0)" - unfolding%unimportant affine_dependent_def affine_hull_explicit mem_Collect_eq - apply rule - apply (erule bexE, erule exE, erule exE) - apply (erule conjE)+ - defer - apply (erule exE, erule exE) - apply (erule conjE)+ - apply (erule bexE) + (\S u. finite S \ S \ p \ sum u S = 0 \ (\v\S. u v \ 0) \ sum (\v. u v *\<^sub>R v) S = 0)" proof - - fix x s u - assume as: "x \ p" "finite s" "s \ {}" "s \ p - {x}" "sum u s = 1" "(\v\s. u v *\<^sub>R v) = x" - have "x \ s" using as(1,4) by auto - show "\s u. finite s \ s \ p \ sum u s = 0 \ (\v\s. u v \ 0) \ (\v\s. u v *\<^sub>R v) = 0" - apply (rule_tac x="insert x s" in exI, rule_tac x="\v. if v = x then - 1 else u v" in exI) - unfolding if_smult and sum_clauses(2)[OF as(2)] and sum_delta_notmem[OF \x\s\] and as - using as - apply auto - done -next - fix s u v - assume as: "finite s" "s \ p" "sum u s = 0" "(\v\s. u v *\<^sub>R v) = 0" "v \ s" "u v \ 0" - have "s \ {v}" - using as(3,6) by auto - then show "\x\p. \s u. finite s \ s \ {} \ s \ p - {x} \ sum u s = 1 \ (\v\s. u v *\<^sub>R v) = x" - apply (rule_tac x=v in bexI) - apply (rule_tac x="s - {v}" in exI) - apply (rule_tac x="\x. - (1 / u v) * u x" in exI) - unfolding scaleR_scaleR[symmetric] and scaleR_right.sum [symmetric] - unfolding sum_distrib_left[symmetric] and sum_diff1[OF as(1)] - using as - apply auto - done + have "\S u. finite S \ S \ p \ sum u S = 0 \ (\v\S. u v \ 0) \ (\w\S. u w *\<^sub>R w) = 0" + if "(\w\S. u w *\<^sub>R w) = x" "x \ p" "finite S" "S \ {}" "S \ p - {x}" "sum u S = 1" for x S u + proof (intro exI conjI) + have "x \ S" + using that by auto + then show "(\v \ insert x S. if v = x then - 1 else u v) = 0" + using that by (simp add: sum_delta_notmem) + show "(\w \ insert x S. (if w = x then - 1 else u w) *\<^sub>R w) = 0" + using that \x \ S\ by (simp add: if_smult sum_delta_notmem cong: if_cong) + qed (use that in auto) + moreover have "\x\p. \S u. finite S \ S \ {} \ S \ p - {x} \ sum u S = 1 \ (\v\S. u v *\<^sub>R v) = x" + if "(\v\S. u v *\<^sub>R v) = 0" "finite S" "S \ p" "sum u S = 0" "v \ S" "u v \ 0" for S u v + proof (intro bexI exI conjI) + have "S \ {v}" + using that by auto + then show "S - {v} \ {}" + using that by auto + show "(\x \ S - {v}. - (1 / u v) * u x) = 1" + unfolding sum_distrib_left[symmetric] sum_diff1[OF \finite S\] by (simp add: that) + show "(\x\S - {v}. (- (1 / u v) * u x) *\<^sub>R x) = v" + unfolding sum_distrib_left [symmetric] scaleR_scaleR[symmetric] + scaleR_right.sum [symmetric] sum_diff1[OF \finite S\] + using that by auto + show "S - {v} \ p - {v}" + using that by auto + qed (use that in auto) + ultimately show ?thesis + unfolding affine_dependent_def affine_hull_explicit by auto qed lemma affine_dependent_explicit_finite: - fixes s :: "'a::real_vector set" - assumes "finite s" - shows "affine_dependent s \ - (\u. sum u s = 0 \ (\v\s. u v \ 0) \ sum (\v. u v *\<^sub>R v) s = 0)" + fixes S :: "'a::real_vector set" + assumes "finite S" + shows "affine_dependent S \ + (\u. sum u S = 0 \ (\v\S. u v \ 0) \ sum (\v. u v *\<^sub>R v) S = 0)" (is "?lhs = ?rhs") proof have *: "\vt u v. (if vt then u v else 0) *\<^sub>R v = (if vt then (u v) *\<^sub>R v else 0::'a)" by auto assume ?lhs then obtain t u v where - "finite t" "t \ s" "sum u t = 0" "v\t" "u v \ 0" "(\v\t. u v *\<^sub>R v) = 0" + "finite t" "t \ S" "sum u t = 0" "v\t" "u v \ 0" "(\v\t. u v *\<^sub>R v) = 0" unfolding affine_dependent_explicit by auto then show ?rhs apply (rule_tac x="\x. if x\t then u x else 0" in exI) - apply auto unfolding * and sum.inter_restrict[OF assms, symmetric] - unfolding Int_absorb1[OF \t\s\] - apply auto + apply (auto simp: * sum.inter_restrict[OF assms, symmetric] Int_absorb1[OF \t\S\]) done next assume ?rhs - then obtain u v where "sum u s = 0" "v\s" "u v \ 0" "(\v\s. u v *\<^sub>R v) = 0" + then obtain u v where "sum u S = 0" "v\S" "u v \ 0" "(\v\S. u v *\<^sub>R v) = 0" by auto then show ?lhs unfolding affine_dependent_explicit using assms by auto @@ -2267,15 +2156,15 @@ by (rule Topological_Spaces.topological_space_class.connectedD) lemma convex_connected: - fixes s :: "'a::real_normed_vector set" - assumes "convex s" - shows "connected s" + fixes S :: "'a::real_normed_vector set" + assumes "convex S" + shows "connected S" proof (rule connectedI) fix A B - assume "open A" "open B" "A \ B \ s = {}" "s \ A \ B" + assume "open A" "open B" "A \ B \ S = {}" "S \ A \ B" moreover - assume "A \ s \ {}" "B \ s \ {}" - then obtain a b where a: "a \ A" "a \ s" and b: "b \ B" "b \ s" by auto + assume "A \ S \ {}" "B \ S \ {}" + then obtain a b where a: "a \ A" "a \ S" and b: "b \ B" "b \ S" by auto define f where [abs_def]: "f u = u *\<^sub>R a + (1 - u) *\<^sub>R b" for u then have "continuous_on {0 .. 1} f" by (auto intro!: continuous_intros) @@ -2286,8 +2175,8 @@ using a by (auto intro!: image_eqI[of _ _ 1] simp: f_def) moreover have "b \ B \ f ` {0 .. 1}" using b by (auto intro!: image_eqI[of _ _ 0] simp: f_def) - moreover have "f ` {0 .. 1} \ s" - using \convex s\ a b unfolding convex_def f_def by auto + moreover have "f ` {0 .. 1} \ S" + using \convex S\ a b unfolding convex_def f_def by auto ultimately show False by auto qed @@ -2372,7 +2261,7 @@ lemma convex_ball [iff]: fixes x :: "'a::real_normed_vector" shows "convex (ball x e)" -proof (auto simp add: convex_def) +proof (auto simp: convex_def) fix y z assume yz: "dist x y < e" "dist x z < e" fix u v :: real @@ -2403,7 +2292,7 @@ then have "dist x (u *\<^sub>R y + v *\<^sub>R z) \ e" using convex_bound_le[OF yz uv] by auto } - then show ?thesis by (auto simp add: convex_def Ball_def) + then show ?thesis by (auto simp: convex_def Ball_def) qed lemma connected_ball [iff]: @@ -2481,8 +2370,8 @@ proof show "convex hull (s \ t) \ (convex hull s) \ (convex hull t)" by (intro hull_minimal Sigma_mono hull_subset convex_Times convex_convex_hull) - have "\x\convex hull s. \y\convex hull t. (x, y) \ convex hull (s \ t)" - proof (intro hull_induct) + have "(x, y) \ convex hull (s \ t)" if x: "x \ convex hull s" and y: "y \ convex hull t" for x y + proof (rule hull_induct [OF x], rule hull_induct [OF y]) fix x y assume "x \ s" and "y \ t" then show "(x, y) \ convex hull (s \ t)" by (simp add: hull_inc) @@ -2492,22 +2381,22 @@ by (intro convex_linear_vimage convex_translation convex_convex_hull, simp add: linear_iff) also have "?S = {y. (x, y) \ convex hull (s \ t)}" - by (auto simp add: image_def Bex_def) + by (auto simp: image_def Bex_def) finally show "convex {y. (x, y) \ convex hull (s \ t)}" . next - show "convex {x. \y\convex hull t. (x, y) \ convex hull (s \ t)}" - proof (unfold Collect_ball_eq, rule convex_INT [rule_format]) + show "convex {x. (x, y) \ convex hull s \ t}" + proof - fix y let ?S = "((\x. (x, 0)) -` (\p. (0, - y) + p) ` (convex hull s \ t))" have "convex ?S" by (intro convex_linear_vimage convex_translation convex_convex_hull, simp add: linear_iff) also have "?S = {x. (x, y) \ convex hull (s \ t)}" - by (auto simp add: image_def Bex_def) + by (auto simp: image_def Bex_def) finally show "convex {x. (x, y) \ convex hull (s \ t)}" . qed qed then show "(convex hull s) \ (convex hull t) \ convex hull (s \ t)" - unfolding subset_eq split_paired_Ball_Sigma . + unfolding subset_eq split_paired_Ball_Sigma by blast qed @@ -2520,118 +2409,114 @@ by (rule hull_unique) auto lemma convex_hull_insert: - fixes s :: "'a::real_vector set" - assumes "s \ {}" - shows "convex hull (insert a s) = - {x. \u\0. \v\0. \b. (u + v = 1) \ b \ (convex hull s) \ (x = u *\<^sub>R a + v *\<^sub>R b)}" + fixes S :: "'a::real_vector set" + assumes "S \ {}" + shows "convex hull (insert a S) = + {x. \u\0. \v\0. \b. (u + v = 1) \ b \ (convex hull S) \ (x = u *\<^sub>R a + v *\<^sub>R b)}" (is "_ = ?hull") - apply (rule, rule hull_minimal, rule) - unfolding insert_iff - prefer 3 - apply rule -proof - +proof (intro equalityI hull_minimal subsetI) fix x - assume x: "x = a \ x \ s" + assume "x \ insert a S" + then have "\u\0. \v\0. u + v = 1 \ (\b. b \ convex hull S \ x = u *\<^sub>R a + v *\<^sub>R b)" + unfolding insert_iff + proof + assume "x = a" + then show ?thesis + by (rule_tac x=1 in exI) (use assms hull_subset in fastforce) + next + assume "x \ S" + with hull_subset[of S convex] show ?thesis + by force + qed then show "x \ ?hull" - apply rule - unfolding mem_Collect_eq - apply (rule_tac x=1 in exI) - defer - apply (rule_tac x=0 in exI) - using assms hull_subset[of s convex] - apply auto - done + by simp next fix x assume "x \ ?hull" - then obtain u v b where obt: "u\0" "v\0" "u + v = 1" "b \ convex hull s" "x = u *\<^sub>R a + v *\<^sub>R b" + then obtain u v b where obt: "u\0" "v\0" "u + v = 1" "b \ convex hull S" "x = u *\<^sub>R a + v *\<^sub>R b" by auto - have "a \ convex hull insert a s" "b \ convex hull insert a s" - using hull_mono[of s "insert a s" convex] hull_mono[of "{a}" "insert a s" convex] and obt(4) + have "a \ convex hull insert a S" "b \ convex hull insert a S" + using hull_mono[of S "insert a S" convex] hull_mono[of "{a}" "insert a S" convex] and obt(4) by auto - then show "x \ convex hull insert a s" + then show "x \ convex hull insert a S" unfolding obt(5) using obt(1-3) by (rule convexD [OF convex_convex_hull]) next show "convex ?hull" proof (rule convexI) fix x y u v - assume as: "(0::real) \ u" "0 \ v" "u + v = 1" "x\?hull" "y\?hull" - from as(4) obtain u1 v1 b1 where - obt1: "u1\0" "v1\0" "u1 + v1 = 1" "b1 \ convex hull s" "x = u1 *\<^sub>R a + v1 *\<^sub>R b1" + assume as: "(0::real) \ u" "0 \ v" "u + v = 1" and x: "x \ ?hull" and y: "y \ ?hull" + from x obtain u1 v1 b1 where + obt1: "u1\0" "v1\0" "u1 + v1 = 1" "b1 \ convex hull S" and xeq: "x = u1 *\<^sub>R a + v1 *\<^sub>R b1" by auto - from as(5) obtain u2 v2 b2 where - obt2: "u2\0" "v2\0" "u2 + v2 = 1" "b2 \ convex hull s" "y = u2 *\<^sub>R a + v2 *\<^sub>R b2" + from y obtain u2 v2 b2 where + obt2: "u2\0" "v2\0" "u2 + v2 = 1" "b2 \ convex hull S" and yeq: "y = u2 *\<^sub>R a + v2 *\<^sub>R b2" by auto have *: "\(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" - by (auto simp add: algebra_simps) - have **: "\b \ convex hull s. u *\<^sub>R x + v *\<^sub>R y = + by (auto simp: algebra_simps) + have "\b \ convex hull S. u *\<^sub>R x + v *\<^sub>R y = (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (b - (u * u1) *\<^sub>R b - (v * u2) *\<^sub>R b)" proof (cases "u * v1 + v * v2 = 0") case True have *: "\(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" - by (auto simp add: algebra_simps) - from True have ***: "u * v1 = 0" "v * v2 = 0" - using mult_nonneg_nonneg[OF \u\0\ \v1\0\] mult_nonneg_nonneg[OF \v\0\ \v2\0\] + by (auto simp: algebra_simps) + have eq0: "u * v1 = 0" "v * v2 = 0" + using True mult_nonneg_nonneg[OF \u\0\ \v1\0\] mult_nonneg_nonneg[OF \v\0\ \v2\0\] by arith+ then have "u * u1 + v * u2 = 1" using as(3) obt1(3) obt2(3) by auto then show ?thesis - unfolding obt1(5) obt2(5) * - using assms hull_subset[of s convex] - by (auto simp add: *** scaleR_right_distrib) + using "*" eq0 as obt1(4) xeq yeq by auto next case False 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) + using as(3) obt1(3) obt2(3) by (auto simp: field_simps) also have "\ = u * (v1 + u1 - u1) + v * (v2 + u2 - u2)" - using as(3) obt1(3) obt2(3) by (auto simp add: field_simps) + using as(3) obt1(3) obt2(3) by (auto simp: field_simps) also have "\ = u * v1 + v * v2" by simp finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" by auto - have "0 \ u * v1 + v * v2" "0 \ u * v1" "0 \ u * v1 + v * v2" "0 \ v * v2" + let ?b = "((u * v1) / (u * v1 + v * v2)) *\<^sub>R b1 + ((v * v2) / (u * v1 + v * v2)) *\<^sub>R b2" + have zeroes: "0 \ u * v1 + v * v2" "0 \ u * v1" "0 \ u * v1 + v * v2" "0 \ v * v2" using as(1,2) obt1(1,2) obt2(1,2) by auto - then show ?thesis - unfolding obt1(5) obt2(5) - unfolding * and ** - using False - 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 - apply (rule convexD [OF convex_convex_hull]) - using obt1(4) obt2(4) - unfolding add_divide_distrib[symmetric] and zero_le_divide_iff - apply (auto simp add: scaleR_left_distrib scaleR_right_distrib) - done + show ?thesis + proof + show "u *\<^sub>R x + v *\<^sub>R y = (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (?b - (u * u1) *\<^sub>R ?b - (v * u2) *\<^sub>R ?b)" + unfolding xeq yeq * ** + using False by (auto simp: scaleR_left_distrib scaleR_right_distrib) + show "?b \ convex hull S" + using False zeroes obt1(4) obt2(4) + by (auto simp: convexD [OF convex_convex_hull] scaleR_left_distrib scaleR_right_distrib add_divide_distrib[symmetric] zero_le_divide_iff) + qed qed + then obtain b where b: "b \ convex hull S" + "u *\<^sub>R x + v *\<^sub>R y = (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (b - (u * u1) *\<^sub>R b - (v * u2) *\<^sub>R b)" .. + have u1: "u1 \ 1" unfolding obt1(3)[symmetric] and not_le using obt1(2) by auto have u2: "u2 \ 1" unfolding obt2(3)[symmetric] and not_le using obt2(2) by auto have "u1 * u + u2 * v \ max u1 u2 * u + max u1 u2 * v" - apply (rule add_mono) - apply (rule_tac [!] mult_right_mono) - using as(1,2) obt1(1,2) obt2(1,2) - apply auto - done + proof (rule add_mono) + show "u1 * u \ max u1 u2 * u" "u2 * v \ max u1 u2 * v" + by (simp_all add: as mult_right_mono) + qed also have "\ \ 1" unfolding distrib_left[symmetric] and as(3) using u1 u2 by auto - finally show "u *\<^sub>R x + v *\<^sub>R y \ ?hull" - unfolding mem_Collect_eq - apply (rule_tac x="u * u1 + v * u2" in exI) - apply (rule conjI) - defer - apply (rule_tac x="1 - u * u1 - v * u2" in exI) - unfolding Bex_def - using as(1,2) obt1(1,2) obt2(1,2) ** - apply (auto simp add: algebra_simps) - done + finally have le1: "u1 * u + u2 * v \ 1" . + show "u *\<^sub>R x + v *\<^sub>R y \ ?hull" + proof (intro CollectI exI conjI) + show "0 \ u * u1 + v * u2" + by (simp add: as(1) as(2) obt1(1) obt2(1)) + show "0 \ 1 - u * u1 - v * u2" + by (simp add: le1 diff_diff_add mult.commute) + qed (use b in \auto simp: algebra_simps\) qed qed lemma convex_hull_insert_alt: "convex hull (insert a S) = - (if S = {} then {a} + (if S = {} then {a} else {(1 - u) *\<^sub>R a + u *\<^sub>R x |x u. 0 \ u \ u \ 1 \ x \ convex hull S})" apply (auto simp: convex_hull_insert) using diff_eq_eq apply fastforce @@ -2639,147 +2524,81 @@ subsubsection%unimportant \Explicit expression for convex hull\ -lemma%important convex_hull_indexed: - fixes s :: "'a::real_vector set" - shows "convex hull s = - {y. \k u x. - (\i\{1::nat .. k}. 0 \ u i \ x i \ s) \ - (sum u {1..k} = 1) \ (sum (\i. u i *\<^sub>R x i) {1..k} = y)}" - (is "?xyz = ?hull") - apply%unimportant (rule hull_unique) - apply rule - defer - apply (rule convexI) -proof - - fix x - assume "x\s" - then show "x \ ?hull" - unfolding mem_Collect_eq - apply (rule_tac x=1 in exI, rule_tac x="\x. 1" in exI) - apply auto - done +proposition%important convex_hull_indexed: + fixes S :: "'a::real_vector set" + shows "convex hull S = + {y. \k u x. (\i\{1::nat .. k}. 0 \ u i \ x i \ S) \ + (sum u {1..k} = 1) \ (\i = 1..k. u i *\<^sub>R x i) = y}" + (is "?xyz = ?hull") +proof (rule hull_unique [OF _ convexI]) + show "S \ ?hull" + by (clarsimp, rule_tac x=1 in exI, rule_tac x="\x. 1" in exI, auto) next - fix t - assume as: "s \ t" "convex t" - show "?hull \ t" - apply rule - unfolding mem_Collect_eq - apply (elim exE conjE) - proof - - fix x k u y - assume assm: - "\i\{1::nat..k}. 0 \ u i \ y i \ s" - "sum u {1..k} = 1" "(\i = 1..k. u i *\<^sub>R y i) = x" - show "x\t" - unfolding assm(3) [symmetric] - apply (rule as(2)[unfolded convex, rule_format]) - using assm(1,2) as(1) apply auto - done - qed + fix T + assume "S \ T" "convex T" + then show "?hull \ T" + by (blast intro: convex_sum) next fix x y u v assume uv: "0 \ u" "0 \ v" "u + v = (1::real)" assume xy: "x \ ?hull" "y \ ?hull" from xy obtain k1 u1 x1 where - x: "\i\{1::nat..k1}. 0\u1 i \ x1 i \ s" "sum u1 {Suc 0..k1} = 1" "(\i = Suc 0..k1. u1 i *\<^sub>R x1 i) = x" + x [rule_format]: "\i\{1::nat..k1}. 0\u1 i \ x1 i \ S" + "sum u1 {Suc 0..k1} = 1" "(\i = Suc 0..k1. u1 i *\<^sub>R x1 i) = x" by auto from xy obtain k2 u2 x2 where - y: "\i\{1::nat..k2}. 0\u2 i \ x2 i \ s" "sum u2 {Suc 0..k2} = 1" "(\i = Suc 0..k2. u2 i *\<^sub>R x2 i) = y" + y [rule_format]: "\i\{1::nat..k2}. 0\u2 i \ x2 i \ S" + "sum u2 {Suc 0..k2} = 1" "(\i = Suc 0..k2. u2 i *\<^sub>R x2 i) = y" by auto - have *: "\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)" - "{1..k1 + k2} \ {1..k1} = {1..k1}" "{1..k1 + k2} \ - {1..k1} = (\i. i + k1) ` {1..k2}" - prefer 3 - apply (rule, rule) - unfolding image_iff - apply (rule_tac x = "x - k1" in bexI) - apply (auto simp add: not_le) - done + have *: "\P (x::'a) y s t i. (if P i then s else t) *\<^sub>R (if P i then x else y) = (if P i then s *\<^sub>R x else t *\<^sub>R y)" + "{1..k1 + k2} \ {1..k1} = {1..k1}" "{1..k1 + k2} \ - {1..k1} = (\i. i + k1) ` {1..k2}" + by auto have inj: "inj_on (\i. i + k1) {1..k2}" unfolding inj_on_def by auto + let ?uu = "\i. if i \ {1..k1} then u * u1 i else v * u2 (i - k1)" + let ?xx = "\i. if i \ {1..k1} then x1 i else x2 (i - k1)" show "u *\<^sub>R x + v *\<^sub>R y \ ?hull" - apply rule - apply (rule_tac x="k1 + k2" in exI) - apply (rule_tac x="\i. if i \ {1..k1} then u * u1 i else v * u2 (i - k1)" in exI) - apply (rule_tac x="\i. if i \ {1..k1} then x1 i else x2 (i - k1)" in exI) - apply (rule, rule) - defer - apply rule - unfolding * and sum.If_cases[OF finite_atLeastAtMost[of 1 "k1 + k2"]] and - sum.reindex[OF inj] and o_def Collect_mem_eq - unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] sum_distrib_left[symmetric] - proof - - fix i - assume i: "i \ {1..k1+k2}" - show "0 \ (if i \ {1..k1} then u * u1 i else v * u2 (i - k1)) \ - (if i \ {1..k1} then x1 i else x2 (i - k1)) \ s" - proof (cases "i\{1..k1}") - case True - then show ?thesis - using uv(1) x(1)[THEN bspec[where x=i]] by auto - next - case False - define j where "j = i - k1" - from i False have "j \ {1..k2}" - unfolding j_def by auto - then show ?thesis - using False uv(2) y(1)[THEN bspec[where x=j]] - by (auto simp: j_def[symmetric]) - qed - qed (auto simp add: not_le x(2,3) y(2,3) uv(3)) + proof (intro CollectI exI conjI ballI) + show "0 \ ?uu i" "?xx i \ S" if "i \ {1..k1+k2}" for i + using that by (auto simp add: le_diff_conv uv(1) x(1) uv(2) y(1)) + show "(\i = 1..k1 + k2. ?uu i) = 1" "(\i = 1..k1 + k2. ?uu i *\<^sub>R ?xx i) = u *\<^sub>R x + v *\<^sub>R y" + unfolding * sum.If_cases[OF finite_atLeastAtMost[of 1 "k1 + k2"]] + sum.reindex[OF inj] Collect_mem_eq o_def + unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] sum_distrib_left[symmetric] + by (simp_all add: sum_distrib_left[symmetric] x(2,3) y(2,3) uv(3)) + qed qed lemma convex_hull_finite: - fixes s :: "'a::real_vector set" - assumes "finite s" - shows "convex hull s = {y. \u. (\x\s. 0 \ u x) \ - sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y}" - (is "?HULL = ?set") -proof (rule hull_unique, auto simp add: convex_def[of ?set]) + fixes S :: "'a::real_vector set" + assumes "finite S" + shows "convex hull S = {y. \u. (\x\S. 0 \ u x) \ sum u S = 1 \ sum (\x. u x *\<^sub>R x) S = y}" + (is "?HULL = _") +proof (rule hull_unique [OF _ convexI]; clarify) fix x - assume "x \ s" - then show "\u. (\x\s. 0 \ u x) \ sum u s = 1 \ (\x\s. u x *\<^sub>R x) = x" - apply (rule_tac x="\y. if x=y then 1 else 0" in exI) - apply auto - unfolding sum.delta'[OF assms] and sum_delta''[OF assms] - apply auto - done + assume "x \ S" + then show "\u. (\x\S. 0 \ u x) \ sum u S = 1 \ (\x\S. u x *\<^sub>R x) = x" + by (rule_tac x="\y. if x=y then 1 else 0" in exI) (auto simp: sum.delta'[OF assms] sum_delta''[OF assms]) next fix u v :: real assume uv: "0 \ u" "0 \ v" "u + v = 1" - fix ux assume ux: "\x\s. 0 \ ux x" "sum ux s = (1::real)" - fix uy assume uy: "\x\s. 0 \ uy x" "sum uy s = (1::real)" - { - fix x - assume "x\s" - then have "0 \ u * ux x + v * uy x" - using ux(1)[THEN bspec[where x=x]] uy(1)[THEN bspec[where x=x]] and uv(1,2) - by auto - } + fix ux assume ux [rule_format]: "\x\S. 0 \ ux x" "sum ux S = (1::real)" + fix uy assume uy [rule_format]: "\x\S. 0 \ uy x" "sum uy S = (1::real)" + have "0 \ u * ux x + v * uy x" if "x\S" for x + by (simp add: that uv ux(1) uy(1)) moreover - have "(\x\s. u * ux x + v * uy x) = 1" - unfolding sum.distrib and sum_distrib_left[symmetric] and ux(2) uy(2) + have "(\x\S. u * ux x + v * uy x) = 1" + unfolding sum.distrib and sum_distrib_left[symmetric] ux(2) uy(2) using uv(3) by auto moreover - have "(\x\s. (u * ux x + v * uy x) *\<^sub>R x) = u *\<^sub>R (\x\s. ux x *\<^sub>R x) + v *\<^sub>R (\x\s. uy x *\<^sub>R x)" - unfolding scaleR_left_distrib and sum.distrib and scaleR_scaleR[symmetric] - and scaleR_right.sum [symmetric] + have "(\x\S. (u * ux x + v * uy x) *\<^sub>R x) = u *\<^sub>R (\x\S. ux x *\<^sub>R x) + v *\<^sub>R (\x\S. uy x *\<^sub>R x)" + unfolding scaleR_left_distrib sum.distrib scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] by auto ultimately - show "\uc. (\x\s. 0 \ uc x) \ sum uc s = 1 \ - (\x\s. uc x *\<^sub>R x) = u *\<^sub>R (\x\s. ux x *\<^sub>R x) + v *\<^sub>R (\x\s. uy x *\<^sub>R x)" - apply (rule_tac x="\x. u * ux x + v * uy x" in exI) - apply auto - done -next - fix t - assume t: "s \ t" "convex t" - fix u - assume u: "\x\s. 0 \ u x" "sum u s = (1::real)" - then show "(\x\s. u x *\<^sub>R x) \ t" - using t(2)[unfolded convex_explicit, THEN spec[where x=s], THEN spec[where x=u]] - using assms and t(1) by auto -qed + show "\uc. (\x\S. 0 \ uc x) \ sum uc S = 1 \ + (\x\S. uc x *\<^sub>R x) = u *\<^sub>R (\x\S. ux x *\<^sub>R x) + v *\<^sub>R (\x\S. uy x *\<^sub>R x)" + by (rule_tac x="\x. u * ux x + v * uy x" in exI, auto) +qed (use assms in \auto simp: convex_explicit\) subsubsection%unimportant \Another formulation from Lars Schewe\ @@ -2787,7 +2606,7 @@ lemma convex_hull_explicit: fixes p :: "'a::real_vector set" shows "convex hull p = - {y. \s u. finite s \ s \ p \ (\x\s. 0 \ u x) \ sum u s = 1 \ sum (\v. u v *\<^sub>R v) s = y}" + {y. \S u. finite S \ S \ p \ (\x\S. 0 \ u x) \ sum u S = 1 \ sum (\v. u v *\<^sub>R v) S = y}" (is "?lhs = ?rhs") proof - { @@ -2817,10 +2636,9 @@ using sum_image_gen[OF fin, of "\i. u i *\<^sub>R y i" y, symmetric] unfolding scaleR_left.sum using obt(3) by auto ultimately - have "\s u. finite s \ s \ p \ (\x\s. 0 \ u x) \ sum u s = 1 \ (\v\s. u v *\<^sub>R v) = x" + have "\S u. finite S \ S \ p \ (\x\S. 0 \ u x) \ sum u S = 1 \ (\v\S. u v *\<^sub>R v) = x" apply (rule_tac x="y ` {1..k}" in exI) - apply (rule_tac x="\v. sum u {i\{1..k}. y i = v}" in exI) - apply auto + apply (rule_tac x="\v. sum u {i\{1..k}. y i = v}" in exI, auto) done then have "x\?rhs" by auto } @@ -2828,55 +2646,50 @@ { fix y assume "y\?rhs" - then obtain s u where - obt: "finite s" "s \ p" "\x\s. 0 \ u x" "sum u s = 1" "(\v\s. u v *\<^sub>R v) = y" + then obtain S u where + obt: "finite S" "S \ p" "\x\S. 0 \ u x" "sum u S = 1" "(\v\S. u v *\<^sub>R v) = y" by auto - obtain f where f: "inj_on f {1..card s}" "f ` {1..card s} = s" + 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 { fix i :: nat - assume "i\{1..card s}" - then have "f i \ s" - apply (subst f(2)[symmetric]) - apply auto - done + assume "i\{1..card S}" + then have "f i \ S" + using f(2) by blast then have "0 \ u (f i)" "f i \ p" using obt(2,3) by auto } - moreover have *: "finite {1..card s}" by auto + moreover have *: "finite {1..card S}" by auto { fix y - assume "y\s" - then obtain i where "i\{1..card s}" "f i = y" - using f using image_iff[of y f "{1..card s}"] + assume "y\S" + then obtain i where "i\{1..card S}" "f i = y" + using f using image_iff[of y f "{1..card S}"] by auto - then have "{x. Suc 0 \ x \ x \ card s \ f x = y} = {i}" + then have "{x. Suc 0 \ x \ x \ card S \ f x = y} = {i}" apply auto using f(1)[unfolded inj_on_def] - apply(erule_tac x=x in ballE) - apply auto - done - then have "card {x. Suc 0 \ x \ x \ card s \ f x = y} = 1" by auto - then have "(\x\{x \ {1..card s}. f x = y}. u (f x)) = u y" - "(\x\{x \ {1..card s}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y" - by (auto simp add: sum_constant_scaleR) + by (metis One_nat_def atLeastAtMost_iff) + then have "card {x. Suc 0 \ x \ x \ card S \ f x = y} = 1" by auto + then have "(\x\{x \ {1..card S}. f x = y}. u (f x)) = u y" + "(\x\{x \ {1..card S}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y" + by (auto simp: sum_constant_scaleR) } - then have "(\x = 1..card s. u (f x)) = 1" "(\i = 1..card s. u (f i) *\<^sub>R f i) = y" + then have "(\x = 1..card S. u (f x)) = 1" "(\i = 1..card S. u (f i) *\<^sub>R f i) = y" unfolding sum_image_gen[OF *(1), of "\x. u (f x) *\<^sub>R f x" f] and sum_image_gen[OF *(1), of "\x. u (f x)" f] unfolding f - using sum.cong [of s s "\y. (\x\{x \ {1..card s}. f x = y}. u (f x) *\<^sub>R f x)" "\v. u v *\<^sub>R v"] - using sum.cong [of s s "\y. (\x\{x \ {1..card s}. f x = y}. u (f x))" u] + using sum.cong [of S S "\y. (\x\{x \ {1..card S}. f x = y}. u (f x) *\<^sub>R f x)" "\v. u v *\<^sub>R v"] + using sum.cong [of S S "\y. (\x\{x \ {1..card S}. f x = y}. u (f x))" u] unfolding obt(4,5) by auto ultimately have "\k u x. (\i\{1..k}. 0 \ u i \ x i \ p) \ sum u {1..k} = 1 \ (\i::nat = 1..k. u i *\<^sub>R x i) = y" - apply (rule_tac x="card s" in exI) + apply (rule_tac x="card S" in exI) apply (rule_tac x="u \ f" in exI) - apply (rule_tac x=f in exI) - apply fastforce + apply (rule_tac x=f in exI, fastforce) done then have "y \ ?lhs" unfolding convex_hull_indexed by auto @@ -2889,70 +2702,57 @@ subsubsection%unimportant \A stepping theorem for that expansion\ lemma convex_hull_finite_step: - fixes s :: "'a::real_vector set" - assumes "finite s" + fixes S :: "'a::real_vector set" + assumes "finite S" shows - "(\u. (\x\insert a s. 0 \ u x) \ sum u (insert a s) = w \ sum (\x. u x *\<^sub>R x) (insert a s) = y) - \ (\v\0. \u. (\x\s. 0 \ u x) \ sum u s = w - v \ sum (\x. u x *\<^sub>R x) s = y - v *\<^sub>R a)" + "(\u. (\x\insert a S. 0 \ u x) \ sum u (insert a S) = w \ sum (\x. u x *\<^sub>R x) (insert a S) = y) + \ (\v\0. \u. (\x\S. 0 \ u x) \ sum u S = w - v \ sum (\x. u x *\<^sub>R x) S = y - v *\<^sub>R a)" (is "?lhs = ?rhs") -proof (rule, case_tac[!] "a\s") - assume "a \ s" - then have *: "insert a s = s" by auto +proof (rule, case_tac[!] "a\S") + assume "a \ S" + then have *: "insert a S = S" by auto assume ?lhs then show ?rhs - unfolding * - apply (rule_tac x=0 in exI) - apply auto - done + unfolding * by (rule_tac x=0 in exI, auto) next assume ?lhs then obtain u where - u: "\x\insert a s. 0 \ u x" "sum u (insert a s) = w" "(\x\insert a s. u x *\<^sub>R x) = y" + u: "\x\insert a S. 0 \ u x" "sum u (insert a S) = w" "(\x\insert a S. u x *\<^sub>R x) = y" by auto - assume "a \ s" + assume "a \ S" then show ?rhs apply (rule_tac x="u a" in exI) using u(1)[THEN bspec[where x=a]] apply simp apply (rule_tac x=u in exI) - using u[unfolded sum_clauses(2)[OF assms]] and \a\s\ + using u[unfolded sum_clauses(2)[OF assms]] and \a\S\ apply auto done next - assume "a \ s" - then have *: "insert a s = s" by auto - have fin: "finite (insert a s)" using assms by auto + assume "a \ S" + then have *: "insert a S = S" by auto + have fin: "finite (insert a S)" using assms by auto assume ?rhs - then obtain v u where uv: "v\0" "\x\s. 0 \ u x" "sum u s = w - v" "(\x\s. u x *\<^sub>R x) = y - v *\<^sub>R a" + then obtain v u where uv: "v\0" "\x\S. 0 \ u x" "sum u S = w - v" "(\x\S. u x *\<^sub>R x) = y - v *\<^sub>R a" by auto show ?lhs apply (rule_tac x = "\x. (if a = x then v else 0) + u x" in exI) unfolding scaleR_left_distrib and sum.distrib and sum_delta''[OF fin] and sum.delta'[OF fin] unfolding sum_clauses(2)[OF assms] - using uv and uv(2)[THEN bspec[where x=a]] and \a\s\ + using uv and uv(2)[THEN bspec[where x=a]] and \a\S\ apply auto done next assume ?rhs - then obtain v u where - uv: "v\0" "\x\s. 0 \ u x" "sum u s = w - v" "(\x\s. u x *\<^sub>R x) = y - v *\<^sub>R a" + then obtain v u where uv: "v\0" "\x\S. 0 \ u x" "sum u S = w - v" "(\x\S. u x *\<^sub>R x) = y - v *\<^sub>R a" by auto - moreover - assume "a \ s" + moreover assume "a \ S" moreover - have "(\x\s. if a = x then v else u x) = sum u s" - and "(\x\s. (if a = x then v else u x) *\<^sub>R x) = (\x\s. u x *\<^sub>R x)" - apply (rule_tac sum.cong) apply rule - defer - apply (rule_tac sum.cong) apply rule - using \a \ s\ - apply auto - done + have "(\x\S. if a = x then v else u x) = sum u S" "(\x\S. (if a = x then v else u x) *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" + using \a \ S\ + by (auto simp: intro!: sum.cong) ultimately show ?lhs - apply (rule_tac x="\x. if a = x then v else u x" in exI) - unfolding sum_clauses(2)[OF assms] - apply auto - done + by (rule_tac x="\x. if a = x then v else u x" in exI) (auto simp: sum_clauses(2)[OF assms]) qed @@ -2969,12 +2769,9 @@ unfolding convex_hull_finite_step[OF **, of a 1, unfolded * conj_assoc] apply auto apply (rule_tac x=v in exI) - apply (rule_tac x="1 - v" in exI) - apply simp - apply (rule_tac x=u in exI) - apply simp - apply (rule_tac x="\x. v" in exI) - apply simp + apply (rule_tac x="1 - v" in exI, simp) + apply (rule_tac x=u in exI, simp) + apply (rule_tac x="\x. v" in exI, simp) done qed @@ -2989,7 +2786,7 @@ unfolding * apply auto apply (rule_tac[!] x=u in exI) - apply (auto simp add: algebra_simps) + apply (auto simp: algebra_simps) done qed @@ -2999,22 +2796,17 @@ have fin: "finite {a,b,c}" "finite {b,c}" "finite {c}" by auto have *: "\x y z ::real. x + y + z = 1 \ x = 1 - y - z" - by (auto simp add: field_simps) + by (auto simp: field_simps) show ?thesis unfolding convex_hull_finite[OF fin(1)] and convex_hull_finite_step[OF fin(2)] and * unfolding convex_hull_finite_step[OF fin(3)] - apply (rule Collect_cong) - apply simp + apply (rule Collect_cong, simp) apply auto apply (rule_tac x=va in exI) - apply (rule_tac x="u c" in exI) - apply simp - apply (rule_tac x="1 - v - w" in exI) - apply simp - apply (rule_tac x=v in exI) - apply simp - apply (rule_tac x="\x. w" in exI) - apply simp + apply (rule_tac x="u c" in exI, simp) + apply (rule_tac x="1 - v - w" in exI, simp) + apply (rule_tac x=v in exI, simp) + apply (rule_tac x="\x. w" in exI, simp) done qed @@ -3025,7 +2817,7 @@ by auto show ?thesis unfolding convex_hull_3 - apply (auto simp add: *) + apply (auto simp: *) apply (rule_tac x=v in exI) apply (rule_tac x=w in exI) apply (simp add: algebra_simps) @@ -3084,36 +2876,24 @@ apply auto done have "(\x\insert a t. if x = a then - (\x\t. u (x - a)) else u (x - a)) = 0" - unfolding sum_clauses(2)[OF fin] - using \a\s\ \t\s\ - apply auto - unfolding * - apply auto - done + unfolding sum_clauses(2)[OF fin] * using \a\s\ \t\s\ by auto moreover have "\v\insert a t. (if v = a then - (\x\t. u (x - a)) else u (v - a)) \ 0" - apply (rule_tac x="v + a" in bexI) - using obt(3,4) and \0\S\ - unfolding t_def - apply auto - done + using obt(3,4) \0\S\ + by (rule_tac x="v + a" in bexI) (auto simp: t_def) moreover have *: "\P Q. (\x\t. (if x = a then P x else Q x) *\<^sub>R x) = (\x\t. Q x *\<^sub>R x)" - apply (rule sum.cong) - using \a\s\ \t\s\ - apply auto - done + using \a\s\ \t\s\ by (auto intro!: sum.cong) have "(\x\t. u (x - a)) *\<^sub>R a = (\v\t. u (v - a) *\<^sub>R v)" unfolding scaleR_left.sum unfolding t_def and sum.reindex[OF inj] and o_def using obt(5) - by (auto simp add: sum.distrib scaleR_right_distrib) + by (auto simp: sum.distrib scaleR_right_distrib) then have "(\v\insert a t. (if v = a then - (\x\t. u (x - a)) else u (v - a)) *\<^sub>R v) = 0" unfolding sum_clauses(2)[OF fin] using \a\s\ \t\s\ - by (auto simp add: *) + by (auto simp: *) ultimately show ?thesis unfolding affine_dependent_explicit - apply (rule_tac x="insert a t" in exI) - apply auto + apply (rule_tac x="insert a t" in exI, auto) done qed @@ -3130,10 +2910,8 @@ using \?lhs\[unfolded convex_def, THEN conjunct1] apply (erule_tac x="2*\<^sub>R x" in ballE) apply (erule_tac x="2*\<^sub>R y" in ballE) - apply (erule_tac x="1/2" in allE) - apply simp - apply (erule_tac x="1/2" in allE) - apply auto + apply (erule_tac x="1/2" in allE, simp) + apply (erule_tac x="1/2" in allE, auto) done } then show ?thesis @@ -3150,49 +2928,36 @@ have *: "{x - a |x. x \ s - {a}} = (\x. x - a) ` (s - {a})" by auto have "card {x - a |x. x \ s - {a}} = card (s - {a})" - unfolding * - apply (rule card_image) - unfolding inj_on_def - apply auto - done + unfolding * by (simp add: card_image inj_on_def) also have "\ > DIM('a)" using assms(2) unfolding card_Diff_singleton[OF assms(1) \a\s\] by auto finally show ?thesis apply (subst insert_Diff[OF \a\s\, symmetric]) apply (rule dependent_imp_affine_dependent) - apply (rule dependent_biggerset) - apply auto + apply (rule dependent_biggerset, auto) done qed lemma affine_dependent_biggerset_general: - assumes "finite (s :: 'a::euclidean_space set)" - and "card s \ dim s + 2" - shows "affine_dependent s" + assumes "finite (S :: 'a::euclidean_space set)" + and "card S \ dim S + 2" + shows "affine_dependent S" proof - - from assms(2) have "s \ {}" by auto - then obtain a where "a\s" by auto - have *: "{x - a |x. x \ s - {a}} = (\x. x - a) ` (s - {a})" + from assms(2) have "S \ {}" by auto + then obtain a where "a\S" by auto + have *: "{x - a |x. x \ S - {a}} = (\x. x - a) ` (S - {a})" by auto - have **: "card {x - a |x. x \ s - {a}} = card (s - {a})" - unfolding * - apply (rule card_image) - unfolding inj_on_def - apply auto - done - have "dim {x - a |x. x \ s - {a}} \ dim s" - apply (rule subset_le_dim) - unfolding subset_eq - using \a\s\ - apply (auto simp add:span_base span_diff) - done - also have "\ < dim s + 1" by auto - also have "\ \ card (s - {a})" + have **: "card {x - a |x. x \ S - {a}} = card (S - {a})" + by (metis (no_types, lifting) "*" card_image diff_add_cancel inj_on_def) + have "dim {x - a |x. x \ S - {a}} \ dim S" + using \a\S\ by (auto simp: span_superset span_diff intro: subset_le_dim) + also have "\ < dim S + 1" by auto + also have "\ \ card (S - {a})" using assms - using card_Diff_singleton[OF assms(1) \a\s\] + using card_Diff_singleton[OF assms(1) \a\S\] by auto finally show ?thesis - apply (subst insert_Diff[OF \a\s\, symmetric]) + apply (subst insert_Diff[OF \a\S\, symmetric]) apply (rule dependent_imp_affine_dependent) apply (rule dependent_biggerset_general) unfolding ** @@ -3384,10 +3149,10 @@ using assms by auto then have h0: "independent ((\x. -a + x) ` (S-{a}))" using affine_dependent_iff_dependent2 assms by auto - then obtain B where B: + obtain B where B: "(\x. -a+x) ` (S - {a}) \ B \ B \ (\x. -a+x) ` V \ independent B \ (\x. -a+x) ` V \ span B" - using maximal_independent_subset_extend[of "(\x. -a+x) ` (S-{a})" "(\x. -a + x) ` V"] assms - by blast + using assms + by (blast intro: maximal_independent_subset_extend[OF _ h0, of "(\x. -a + x) ` V"]) define T where "T = (\x. a+x) ` insert 0 B" then have "T = insert a ((\x. a+x) ` B)" by auto @@ -3457,8 +3222,7 @@ some_eq_ex[of "\d. \B. 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"]) - apply auto + apply (rule exI[of _ "B"], auto) done qed @@ -3493,10 +3257,7 @@ then have "card ((\x. -a + x) ` (B - {a})) > 0" using fin by auto moreover have h1: "card ((\x. -a + x) ` (B-{a})) = card (B-{a})" - apply (rule card_image) - using translate_inj_on - apply (auto simp del: uminus_add_conv_diff) - done + by (rule card_image) (use translate_inj_on in blast) ultimately have "card (B-{a}) > 0" by auto then have *: "finite (B - {a})" using card_gt_0_iff[of "(B - {a})"] by auto @@ -3592,23 +3353,10 @@ by auto let ?t = "{x::'a::euclidean_space. \i\Basis. i \ d \ x\i = 0}" have "\f. linear f \ f ` B = d \ f ` span B = ?t \ inj_on f (span B)" - apply (rule basis_to_basis_subspace_isomorphism[of "span B" ?t B "d"]) - apply (rule subspace_span) - apply (rule subspace_substandard) - defer - apply (rule span_superset) - apply (rule assms) - defer - unfolding dim_span[of B] - apply(rule B) - unfolding span_substd_basis[OF d, symmetric] - apply (rule span_superset) - apply (rule independent_substdbasis[OF d]) - apply rule - apply assumption - unfolding t[symmetric] span_substd_basis[OF d] dim_substandard[OF d] - apply auto - done + proof (intro basis_to_basis_subspace_isomorphism subspace_span subspace_substandard span_inc) + show "d \ {x. \i\Basis. i \ d \ x \ i = 0}" + using d inner_not_same_Basis by blast + qed (auto simp: span_substd_basis independent_substdbasis dim_substandard d t B assms) with t \card B = dim B\ d show ?thesis by auto qed @@ -3694,7 +3442,7 @@ assume "a \ b" then have "aff_dim{a,b} = card{a,b} - 1" using affine_independent_2 [of a b] aff_dim_affine_independent by fastforce - also have "... = 1" + also have "\ = 1" using \a \ b\ by simp finally show "aff_dim {a, b} = 1" . qed @@ -3923,9 +3671,9 @@ by blast then have "card {b - a |b. b \ S - {a}} = card ((\b. b-a) ` (S - {a}))" by simp - also have "... = card (S - {a})" + also have "\ = card (S - {a})" by (metis (no_types, lifting) card_image diff_add_cancel inj_onI) - also have "... = card S - 1" + also have "\ = card S - 1" by (simp add: aff_independent_finite assms) finally have 4: "card {b - a |b. b \ S - {a}} = card S - 1" . have "finite S" @@ -4156,8 +3904,7 @@ assume "\s u. finite s \ s \ p \ (\x\s. 0 \ u x) \ sum u s = 1 \ (\v\s. u v *\<^sub>R v) = y" then obtain N where "?P N" by auto then have "\n\N. (\k ?P k) \ ?P n" - apply (rule_tac ex_least_nat_le) - apply auto + apply (rule_tac ex_least_nat_le, auto) done then obtain n where "?P n" and smallest: "\k ?P k" by blast @@ -4178,8 +3925,7 @@ proof (rule ccontr, simp add: not_less) assume as:"\x\s. 0 \ w x" then have "sum w (s - {v}) \ 0" - apply (rule_tac sum_nonneg) - apply auto + apply (rule_tac sum_nonneg, auto) done then have "sum w s > 0" unfolding sum.remove[OF obt(1) \v\s\] @@ -4229,7 +3975,7 @@ apply (rule_tac x="(s - {a})" in exI) apply (rule_tac x="\v. u v + t * w v" in exI) using obt(1-3) and t and a - apply (auto simp add: * scaleR_left_distrib) + apply (auto simp: * scaleR_left_distrib) done then show False using smallest[THEN spec[where x="n - 1"]] by auto @@ -4245,9 +3991,8 @@ (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" - apply (subst convex_hull_caratheodory_aff_dim) - apply clarify - apply (rule_tac x="s" in exI) + apply (subst convex_hull_caratheodory_aff_dim, clarify) + apply (rule_tac x=s in exI) apply (simp add: hull_subset convex_explicit [THEN iffD1, OF convex_convex_hull]) done next @@ -4272,7 +4017,7 @@ next fix x assume "x \ ?rhs" then show "x \ ?lhs" - by (auto simp add: convex_hull_explicit) + by (auto simp: convex_hull_explicit) qed theorem%important caratheodory: @@ -4331,14 +4076,13 @@ qed lemma mem_rel_interior: "x \ rel_interior S \ (\T. open T \ x \ T \ S \ T \ affine hull S \ S)" - by (auto simp add: rel_interior) + by (auto simp: rel_interior) lemma mem_rel_interior_ball: "x \ rel_interior S \ x \ S \ (\e. e > 0 \ ball x e \ affine hull S \ S)" apply (simp add: rel_interior, safe) - apply (force simp add: open_contains_ball) - apply (rule_tac x = "ball x e" in exI) - apply simp + apply (force simp: open_contains_ball) + apply (rule_tac x = "ball x e" in exI, simp) done lemma rel_interior_ball: @@ -4348,10 +4092,9 @@ lemma mem_rel_interior_cball: "x \ rel_interior S \ x \ S \ (\e. e > 0 \ cball x e \ affine hull S \ S)" apply (simp add: rel_interior, safe) - apply (force simp add: open_contains_cball) + apply (force simp: open_contains_cball) apply (rule_tac x = "ball x e" in exI) - apply (simp add: subset_trans [OF ball_subset_cball]) - apply auto + apply (simp add: subset_trans [OF ball_subset_cball], auto) done lemma rel_interior_cball: @@ -4359,7 +4102,7 @@ using mem_rel_interior_cball [of _ S] by auto lemma rel_interior_empty [simp]: "rel_interior {} = {}" - by (auto simp add: rel_interior_def) + by (auto simp: rel_interior_def) lemma affine_hull_sing [simp]: "affine hull {a :: 'n::euclidean_space} = {a}" by (metis affine_hull_eq affine_sing) @@ -4367,8 +4110,7 @@ lemma rel_interior_sing [simp]: fixes a :: "'n::euclidean_space" shows "rel_interior {a} = {a}" apply (auto simp: rel_interior_ball) - apply (rule_tac x=1 in exI) - apply force + apply (rule_tac x=1 in exI, force) done lemma subset_rel_interior: @@ -4376,16 +4118,16 @@ assumes "S \ T" and "affine hull S = affine hull T" shows "rel_interior S \ rel_interior T" - using assms by (auto simp add: rel_interior_def) + using assms by (auto simp: rel_interior_def) lemma rel_interior_subset: "rel_interior S \ S" - by (auto simp add: rel_interior_def) + by (auto simp: rel_interior_def) lemma rel_interior_subset_closure: "rel_interior S \ closure S" - using rel_interior_subset by (auto simp add: closure_def) + using rel_interior_subset by (auto simp: closure_def) lemma interior_subset_rel_interior: "interior S \ rel_interior S" - by (auto simp add: rel_interior interior_def) + by (auto simp: rel_interior interior_def) lemma interior_rel_interior: fixes S :: "'n::euclidean_space set" @@ -4464,7 +4206,7 @@ fix y assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d" "y \ affine hull S" 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) + using \e > 0\ by (auto simp: scaleR_left_diff_distrib scaleR_right_diff_distrib) have "x \ affine hull S" using assms hull_subset[of S] by auto moreover have "1 / e + - ((1 - e) / e) = 1" @@ -4476,17 +4218,17 @@ unfolding dist_norm norm_scaleR[symmetric] apply (rule arg_cong[where f=norm]) using \e > 0\ - apply (auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps) + apply (auto simp: euclidean_eq_iff[where 'a='a] field_simps inner_simps) done also have "\ = \1/e\ * norm (x - e *\<^sub>R (x - c) - y)" by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) also have "\ < d" using as[unfolded dist_norm] and \e > 0\ - by (auto simp add:pos_divide_less_eq[OF \e > 0\] mult.commute) + by (auto simp:pos_divide_less_eq[OF \e > 0\] mult.commute) finally have "y \ S" apply (subst *) apply (rule assms(1)[unfolded convex_alt,rule_format]) - apply (rule d[unfolded subset_eq,rule_format]) + apply (rule d[THEN subsetD]) unfolding mem_ball using assms(3-5) ** apply auto @@ -4518,7 +4260,7 @@ then have "y \ interior {a..}" apply (simp add: mem_interior) apply (rule_tac x="(y-a)" in exI) - apply (auto simp add: dist_norm) + apply (auto simp: dist_norm) done } moreover @@ -4528,7 +4270,7 @@ then obtain e where e: "e > 0" "cball y e \ {a..}" using mem_interior_cball[of y "{a..}"] by auto moreover from e have "y - e \ cball y e" - by (auto simp add: cball_def dist_norm) + by (auto simp: cball_def dist_norm) ultimately have "a \ y - e" by blast then have "a < y" using e by auto } @@ -4558,7 +4300,7 @@ then have "y \ interior {..a}" apply (simp add: mem_interior) apply (rule_tac x="(a-y)" in exI) - apply (auto simp add: dist_norm) + apply (auto simp: dist_norm) done } moreover @@ -4568,7 +4310,7 @@ then obtain e where e: "e > 0" "cball y e \ {..a}" using mem_interior_cball[of y "{..a}"] by auto moreover from e have "y + e \ cball y e" - by (auto simp add: cball_def dist_norm) + by (auto simp: cball_def dist_norm) ultimately have "a \ y + e" by auto then have "a > y" using e by auto } @@ -4578,9 +4320,9 @@ lemma interior_atLeastAtMost_real [simp]: "interior {a..b} = {a<.. {..b}" by auto - also have "interior ... = {a<..} \ {.. = {a<..} \ {.. = {a<.. S") case True then show ?thesis using \e > 0\ \d > 0\ - apply (rule_tac bexI[where x=x]) - apply (auto) + apply (rule_tac bexI[where x=x], auto) done next case False @@ -4739,7 +4479,7 @@ next case False then have "0 < e * d / (1 - e)" and *: "1 - e > 0" - using \e \ 1\ \e > 0\ \d > 0\ by (auto) + using \e \ 1\ \e > 0\ \d > 0\ by auto then obtain y where "y \ S" "y \ x" "dist y x < e * d / (1 - e)" using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto then show ?thesis @@ -4755,11 +4495,11 @@ define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)" 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) + by (auto simp: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) have zball: "z \ ball c d" using mem_ball z_def dist_norm[of c] using y and assms(4,5) - by (auto simp add:field_simps norm_minus_commute) + by (auto simp:field_simps norm_minus_commute) have "x \ affine hull S" using closure_affine_hull assms by auto moreover have "y \ affine hull S" @@ -4770,7 +4510,7 @@ using z_def affine_affine_hull[of S] mem_affine_3_minus [of "affine hull S" c x y "(1 - e) / e"] assms - by (auto simp add: field_simps) + by (auto simp: field_simps) then have "z \ S" using d zball by auto obtain d1 where "d1 > 0" and d1: "ball z d1 \ ball c d" using zball open_ball[of c d] openE[of "ball c d" z] by auto @@ -4868,24 +4608,17 @@ lemma affine_hull_linear_image: assumes "bounded_linear f" shows "f ` (affine hull s) = affine hull f ` s" - apply rule - unfolding subset_eq ball_simps - apply (rule_tac[!] hull_induct, rule hull_inc) - prefer 3 - apply (erule imageE) - apply (rule_tac x=xa in image_eqI) - apply assumption - apply (rule hull_subset[unfolded subset_eq, rule_format]) - apply assumption proof - interpret f: bounded_linear f by fact - show "affine {x. f x \ affine hull f ` s}" + have "affine {x. f x \ affine hull f ` s}" unfolding affine_def - by (auto simp add: f.scaleR f.add affine_affine_hull[unfolded affine_def, rule_format]) - show "affine {x. x \ f ` (affine hull s)}" + by (auto simp: f.scaleR f.add affine_affine_hull[unfolded affine_def, rule_format]) + moreover have "affine {x. x \ f ` (affine hull s)}" using affine_affine_hull[unfolded affine_def, of s] - unfolding affine_def by (auto simp add: f.scaleR [symmetric] f.add [symmetric]) -qed auto + unfolding affine_def by (auto simp: f.scaleR [symmetric] f.add [symmetric]) + ultimately show ?thesis + by (auto simp: hull_inc elim!: hull_induct) +qed lemma rel_interior_injective_on_span_linear_image: @@ -5012,114 +4745,77 @@ subsection%unimportant \Openness and compactness are preserved by convex hull operation\ lemma open_convex_hull[intro]: - fixes s :: "'a::real_normed_vector set" - assumes "open s" - shows "open (convex hull s)" - unfolding open_contains_cball convex_hull_explicit - unfolding mem_Collect_eq ball_simps(8) -proof (rule, rule) - fix a - assume "\sa u. finite sa \ sa \ s \ (\x\sa. 0 \ u x) \ sum u sa = 1 \ (\v\sa. u v *\<^sub>R v) = a" - then obtain t u where obt: "finite t" "t\s" "\x\t. 0 \ u x" "sum u t = 1" "(\v\t. u v *\<^sub>R v) = a" - by auto + fixes S :: "'a::real_normed_vector set" + assumes "open S" + shows "open (convex hull S)" +proof (clarsimp simp: open_contains_cball convex_hull_explicit) + fix T and u :: "'a\real" + assume obt: "finite T" "T\S" "\x\T. 0 \ u x" "sum u T = 1" from assms[unfolded open_contains_cball] obtain b - where b: "\x\s. 0 < b x \ cball x (b x) \ s" - using bchoice[of s "\x e. e > 0 \ cball x e \ s"] by auto - have "b ` t \ {}" + where b: "\x. x\S \ 0 < b x \ cball x (b x) \ S" by metis + have "b ` T \ {}" using obt by auto - define i where "i = b ` t" - - show "\e > 0. - cball a e \ {y. \sa u. finite sa \ sa \ s \ (\x\sa. 0 \ u x) \ sum u sa = 1 \ (\v\sa. u v *\<^sub>R v) = y}" - apply (rule_tac x = "Min i" in exI) - unfolding subset_eq - apply rule - defer - apply rule - unfolding mem_Collect_eq - proof - + define i where "i = b ` T" + let ?\ = "\y. \F. finite F \ F \ S \ (\u. (\x\F. 0 \ u x) \ sum u F = 1 \ (\v\F. u v *\<^sub>R v) = y)" + let ?a = "\v\T. u v *\<^sub>R v" + show "\e > 0. cball ?a e \ {y. ?\ y}" + proof (intro exI subsetI conjI) show "0 < Min i" - unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] \b ` t\{}\] - using b - apply simp - apply rule - apply (erule_tac x=x in ballE) - using \t\s\ - apply auto - done + unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] \b ` T\{}\] + using b \T\S\ by auto next fix y - assume "y \ cball a (Min i)" - then have y: "norm (a - y) \ Min i" + assume "y \ cball ?a (Min i)" + then have y: "norm (?a - y) \ Min i" unfolding dist_norm[symmetric] by auto - { - fix x - assume "x \ t" + { fix x + assume "x \ T" then have "Min i \ b x" - unfolding i_def - apply (rule_tac Min_le) - using obt(1) - apply auto - done - then have "x + (y - a) \ cball x (b x)" + by (simp add: i_def obt(1)) + then have "x + (y - ?a) \ cball x (b x)" using y unfolding mem_cball dist_norm by auto - moreover from \x\t\ have "x \ s" - using obt(2) by auto - ultimately have "x + (y - a) \ s" - using y and b[THEN bspec[where x=x]] unfolding subset_eq by fast + moreover have "x \ S" + using \x\T\ \T\S\ by auto + ultimately have "x + (y - ?a) \ S" + using y b by blast } moreover - have *: "inj_on (\v. v + (y - a)) t" + have *: "inj_on (\v. v + (y - ?a)) T" unfolding inj_on_def by auto - have "(\v\(\v. v + (y - a)) ` t. u (v - (y - a))) = 1" - unfolding sum.reindex[OF *] o_def using obt(4) by auto - moreover have "(\v\(\v. v + (y - a)) ` t. u (v - (y - a)) *\<^sub>R v) = y" - unfolding sum.reindex[OF *] o_def using obt(4,5) + have "(\v\(\v. v + (y - ?a)) ` T. u (v - (y - ?a)) *\<^sub>R v) = y" + unfolding sum.reindex[OF *] o_def using obt(4) by (simp add: sum.distrib sum_subtractf scaleR_left.sum[symmetric] scaleR_right_distrib) - ultimately - show "\sa u. finite sa \ (\x\sa. x \ s) \ (\x\sa. 0 \ u x) \ sum u sa = 1 \ (\v\sa. u v *\<^sub>R v) = y" - apply (rule_tac x="(\v. v + (y - a)) ` t" in exI) - apply (rule_tac x="\v. u (v - (y - a))" in exI) - using obt(1, 3) - apply auto - done + ultimately show "y \ {y. ?\ y}" + proof (intro CollectI exI conjI) + show "finite ((\v. v + (y - ?a)) ` T)" + by (simp add: obt(1)) + show "sum (\v. u (v - (y - ?a))) ((\v. v + (y - ?a)) ` T) = 1" + unfolding sum.reindex[OF *] o_def using obt(4) by auto + qed (use obt(1, 3) in auto) qed qed lemma compact_convex_combinations: - fixes s t :: "'a::real_normed_vector set" - assumes "compact s" "compact t" - shows "compact { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ s \ y \ t}" + fixes S T :: "'a::real_normed_vector set" + assumes "compact S" "compact T" + shows "compact { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ S \ y \ T}" proof - - let ?X = "{0..1} \ s \ t" + let ?X = "{0..1} \ S \ T" let ?h = "(\z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" - have *: "{ (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ s \ y \ t} = ?h ` ?X" - apply (rule set_eqI) - unfolding image_iff mem_Collect_eq - apply rule - apply auto - apply (rule_tac x=u in rev_bexI) - apply simp - apply (erule rev_bexI) - apply (erule rev_bexI) - apply simp - apply auto - done + have *: "{ (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ S \ y \ T} = ?h ` ?X" + by force have "continuous_on ?X (\z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" unfolding continuous_on by (rule ballI) (intro tendsto_intros) - then show ?thesis - unfolding * - apply (rule compact_continuous_image) - apply (intro compact_Times compact_Icc assms) - done + with assms show ?thesis + by (simp add: * compact_Times compact_continuous_image) qed lemma finite_imp_compact_convex_hull: - fixes s :: "'a::real_normed_vector set" - assumes "finite s" - shows "compact (convex hull s)" -proof (cases "s = {}") + fixes S :: "'a::real_normed_vector set" + assumes "finite S" + shows "compact (convex hull S)" +proof (cases "S = {}") case True then show ?thesis by simp next @@ -5142,8 +4838,7 @@ unfolding convex_hull_insert [OF \A \ {}\] apply safe apply (rule_tac x=a in exI, simp) - apply (rule_tac x="1 - a" in exI, simp) - apply fast + apply (rule_tac x="1 - a" in exI, simp, fast) apply (rule_tac x="(u, b)" in image_eqI, simp_all) done finally show "compact (convex hull (insert x A))" . @@ -5151,20 +4846,20 @@ qed lemma compact_convex_hull: - fixes s :: "'a::euclidean_space set" - assumes "compact s" - shows "compact (convex hull s)" -proof (cases "s = {}") + fixes S :: "'a::euclidean_space set" + assumes "compact S" + shows "compact (convex hull S)" +proof (cases "S = {}") case True then show ?thesis using compact_empty by simp next case False - then obtain w where "w \ s" by auto + then obtain w where "w \ S" by auto show ?thesis - unfolding caratheodory[of s] + unfolding caratheodory[of S] proof (induct ("DIM('a) + 1")) case 0 - have *: "{x.\sa. finite sa \ sa \ s \ card sa \ 0 \ x \ convex hull sa} = {}" + have *: "{x.\sa. finite sa \ sa \ S \ card sa \ 0 \ x \ convex hull sa} = {}" using compact_empty by auto from 0 show ?case unfolding * by simp next @@ -5172,27 +4867,27 @@ show ?case proof (cases "n = 0") case True - have "{x. \t. finite t \ t \ s \ card t \ Suc n \ x \ convex hull t} = s" + have "{x. \T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T} = S" unfolding set_eq_iff and mem_Collect_eq proof (rule, rule) fix x - assume "\t. finite t \ t \ s \ card t \ Suc n \ x \ convex hull t" - then obtain t where t: "finite t" "t \ s" "card t \ Suc n" "x \ convex hull t" + assume "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T" + then obtain T where T: "finite T" "T \ S" "card T \ Suc n" "x \ convex hull T" by auto - show "x \ s" - proof (cases "card t = 0") + show "x \ S" + proof (cases "card T = 0") case True then show ?thesis - using t(4) unfolding card_0_eq[OF t(1)] by simp + using T(4) unfolding card_0_eq[OF T(1)] by simp next case False - then have "card t = Suc 0" using t(3) \n=0\ by auto - then obtain a where "t = {a}" unfolding card_Suc_eq by auto - then show ?thesis using t(2,4) by simp + then have "card T = Suc 0" using T(3) \n=0\ by auto + then obtain a where "T = {a}" unfolding card_Suc_eq by auto + then show ?thesis using T(2,4) by simp qed next - fix x assume "x\s" - then show "\t. finite t \ t \ s \ card t \ Suc n \ x \ convex hull t" + fix x assume "x\S" + then show "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T" apply (rule_tac x="{x}" in exI) unfolding convex_hull_singleton apply auto @@ -5201,57 +4896,56 @@ then show ?thesis using assms by simp next case False - have "{x. \t. finite t \ t \ s \ card t \ Suc n \ x \ convex hull t} = + have "{x. \T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T} = {(1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. - 0 \ u \ u \ 1 \ x \ s \ y \ {x. \t. finite t \ t \ s \ card t \ n \ x \ convex hull t}}" + 0 \ u \ u \ 1 \ x \ S \ y \ {x. \T. finite T \ T \ S \ card T \ n \ x \ convex hull T}}" unfolding set_eq_iff and mem_Collect_eq proof (rule, rule) fix x assume "\u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \ - 0 \ c \ c \ 1 \ u \ s \ (\t. finite t \ t \ s \ card t \ n \ v \ convex hull t)" - then obtain u v c t where obt: "x = (1 - c) *\<^sub>R u + c *\<^sub>R v" - "0 \ c \ c \ 1" "u \ s" "finite t" "t \ s" "card t \ n" "v \ convex hull t" + 0 \ c \ c \ 1 \ u \ S \ (\T. finite T \ T \ S \ card T \ n \ v \ convex hull T)" + then obtain u v c T where obt: "x = (1 - c) *\<^sub>R u + c *\<^sub>R v" + "0 \ c \ c \ 1" "u \ S" "finite T" "T \ S" "card T \ n" "v \ convex hull T" by auto - moreover have "(1 - c) *\<^sub>R u + c *\<^sub>R v \ convex hull insert u t" + moreover have "(1 - c) *\<^sub>R u + c *\<^sub>R v \ convex hull insert u T" apply (rule convexD_alt) - using obt(2) and convex_convex_hull and hull_subset[of "insert u t" convex] - using obt(7) and hull_mono[of t "insert u t"] + using obt(2) and convex_convex_hull and hull_subset[of "insert u T" convex] + using obt(7) and hull_mono[of T "insert u T"] apply auto done - ultimately show "\t. finite t \ t \ s \ card t \ Suc n \ x \ convex hull t" - apply (rule_tac x="insert u t" in exI) - apply (auto simp add: card_insert_if) + ultimately show "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T" + apply (rule_tac x="insert u T" in exI) + apply (auto simp: card_insert_if) done next fix x - assume "\t. finite t \ t \ s \ card t \ Suc n \ x \ convex hull t" - then obtain t where t: "finite t" "t \ s" "card t \ Suc n" "x \ convex hull t" + assume "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T" + then obtain T where T: "finite T" "T \ S" "card T \ Suc n" "x \ convex hull T" by auto show "\u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \ - 0 \ c \ c \ 1 \ u \ s \ (\t. finite t \ t \ s \ card t \ n \ v \ convex hull t)" - proof (cases "card t = Suc n") + 0 \ c \ c \ 1 \ u \ S \ (\T. finite T \ T \ S \ card T \ n \ v \ convex hull T)" + proof (cases "card T = Suc n") case False - then have "card t \ n" using t(3) by auto + then have "card T \ n" using T(3) by auto then show ?thesis apply (rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI) - using \w\s\ and t - apply (auto intro!: exI[where x=t]) + using \w\S\ and T + apply (auto intro!: exI[where x=T]) done next case True - then obtain a u where au: "t = insert a u" "a\u" - apply (drule_tac card_eq_SucD) - apply auto + then obtain a u where au: "T = insert a u" "a\u" + apply (drule_tac card_eq_SucD, auto) done show ?thesis proof (cases "u = {}") case True - then have "x = a" using t(4)[unfolded au] by auto + then have "x = a" using T(4)[unfolded au] by auto show ?thesis unfolding \x = a\ apply (rule_tac x=a in exI) apply (rule_tac x=a in exI) apply (rule_tac x=1 in exI) - using t and \n \ 0\ + using T and \n \ 0\ unfolding au apply (auto intro!: exI[where x="{a}"]) done @@ -5259,14 +4953,14 @@ case False obtain ux vx b where obt: "ux\0" "vx\0" "ux + vx = 1" "b \ convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b" - using t(4)[unfolded au convex_hull_insert[OF False]] + using T(4)[unfolded au convex_hull_insert[OF False]] by auto have *: "1 - vx = ux" using obt(3) by auto show ?thesis apply (rule_tac x=a in exI) apply (rule_tac x=b in exI) apply (rule_tac x=vx in exI) - using obt and t(1-3) + using obt and T(1-3) unfolding au and * using card_insert_disjoint[OF _ au(2)] apply (auto intro!: exI[where x=u]) done @@ -5318,25 +5012,25 @@ using dist_increases_online[of d a 0] unfolding dist_norm by auto lemma simplex_furthest_lt: - fixes s :: "'a::real_inner set" - assumes "finite s" - shows "\x \ convex hull s. x \ s \ (\y \ convex hull s. norm (x - a) < norm(y - a))" + fixes S :: "'a::real_inner set" + assumes "finite S" + shows "\x \ convex hull S. x \ S \ (\y \ convex hull S. norm (x - a) < norm(y - a))" using assms proof induct - fix x s - assume as: "finite s" "x\s" "\x\convex hull s. x \ s \ (\y\convex hull s. norm (x - a) < norm (y - a))" - show "\xa\convex hull insert x s. xa \ insert x s \ - (\y\convex hull insert x s. norm (xa - a) < norm (y - a))" - proof (rule, rule, cases "s = {}") + fix x S + assume as: "finite S" "x\S" "\x\convex hull S. x \ S \ (\y\convex hull S. norm (x - a) < norm (y - a))" + show "\xa\convex hull insert x S. xa \ insert x S \ + (\y\convex hull insert x S. norm (xa - a) < norm (y - a))" + proof (intro impI ballI, cases "S = {}") case False fix y - assume y: "y \ convex hull insert x s" "y \ insert x s" - obtain u v b where obt: "u\0" "v\0" "u + v = 1" "b \ convex hull s" "y = u *\<^sub>R x + v *\<^sub>R b" + assume y: "y \ convex hull insert x S" "y \ insert x S" + obtain u v b where obt: "u\0" "v\0" "u + v = 1" "b \ convex hull S" "y = u *\<^sub>R x + v *\<^sub>R b" using y(1)[unfolded convex_hull_insert[OF False]] by auto - show "\z\convex hull insert x s. norm (y - a) < norm (z - a)" - proof (cases "y \ convex hull s") + show "\z\convex hull insert x S. norm (y - a) < norm (z - a)" + proof (cases "y \ convex hull S") case True - then obtain z where "z \ convex hull s" "norm (y - a) < norm (z - a)" + then obtain z where "z \ convex hull S" "norm (y - a) < norm (z - a)" using as(3)[THEN bspec[where x=y]] and y(2) by auto then show ?thesis apply (rule_tac x=z in bexI) @@ -5363,7 +5057,7 @@ proof assume "x = b" then have "y = b" unfolding obt(5) - using obt(3) by (auto simp add: scaleR_left_distrib[symmetric]) + using obt(3) by (auto simp: scaleR_left_distrib[symmetric]) then show False using obt(4) and False by simp qed then have *: "w *\<^sub>R (x - b) \ 0" using w(1) by auto @@ -5375,15 +5069,12 @@ unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps) - moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \ convex hull insert x s" - unfolding convex_hull_insert[OF \s\{}\] and mem_Collect_eq - apply (rule_tac x="u + w" in exI) - apply rule - defer - apply (rule_tac x="v - w" in exI) - using \u \ 0\ and w and obt(3,4) - apply auto - done + moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \ convex hull insert x S" + unfolding convex_hull_insert[OF \S\{}\] + proof (intro CollectI conjI exI) + show "u + w \ 0" "v - w \ 0" + using obt(1) w by auto + qed (use obt in auto) ultimately show ?thesis by auto next assume "dist a y < dist a (y - w *\<^sub>R (x - b))" @@ -5391,39 +5082,36 @@ unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps) - moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \ convex hull insert x s" - unfolding convex_hull_insert[OF \s\{}\] and mem_Collect_eq - apply (rule_tac x="u - w" in exI) - apply rule - defer - apply (rule_tac x="v + w" in exI) - using \u \ 0\ and w and obt(3,4) - apply auto - done + moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \ convex hull insert x S" + unfolding convex_hull_insert[OF \S\{}\] + proof (intro CollectI conjI exI) + show "u - w \ 0" "v + w \ 0" + using obt(1) w by auto + qed (use obt in auto) ultimately show ?thesis by auto qed qed auto qed qed auto -qed (auto simp add: assms) +qed (auto simp: assms) lemma simplex_furthest_le: - fixes s :: "'a::real_inner set" - assumes "finite s" - and "s \ {}" - shows "\y\s. \x\ convex hull s. norm (x - a) \ norm (y - a)" + fixes S :: "'a::real_inner set" + assumes "finite S" + and "S \ {}" + shows "\y\S. \x\ convex hull S. norm (x - a) \ norm (y - a)" proof - - have "convex hull s \ {}" - using hull_subset[of s convex] and assms(2) by auto - then obtain x where x: "x \ convex hull s" "\y\convex hull s. norm (y - a) \ norm (x - a)" - using distance_attains_sup[OF finite_imp_compact_convex_hull[OF assms(1)], of a] + have "convex hull S \ {}" + using hull_subset[of S convex] and assms(2) by auto + then obtain x where x: "x \ convex hull S" "\y\convex hull S. norm (y - a) \ norm (x - a)" + using distance_attains_sup[OF finite_imp_compact_convex_hull[OF \finite S\], of a] unfolding dist_commute[of a] unfolding dist_norm by auto show ?thesis - proof (cases "x \ s") + proof (cases "x \ S") case False - then obtain y where "y \ convex hull s" "norm (x - a) < norm (y - a)" + then obtain y where "y \ convex hull S" "norm (x - a) < norm (y - a)" using simplex_furthest_lt[OF assms(1), THEN bspec[where x=x]] and x(1) by auto then show ?thesis @@ -5435,82 +5123,82 @@ qed lemma simplex_furthest_le_exists: - fixes s :: "('a::real_inner) set" - shows "finite s \ \x\(convex hull s). \y\s. norm (x - a) \ norm (y - a)" - using simplex_furthest_le[of s] by (cases "s = {}") auto + fixes S :: "('a::real_inner) set" + shows "finite S \ \x\(convex hull S). \y\S. norm (x - a) \ norm (y - a)" + using simplex_furthest_le[of S] by (cases "S = {}") auto lemma simplex_extremal_le: - fixes s :: "'a::real_inner set" - assumes "finite s" - and "s \ {}" - shows "\u\s. \v\s. \x\convex hull s. \y \ convex hull s. norm (x - y) \ norm (u - v)" + fixes S :: "'a::real_inner set" + assumes "finite S" + and "S \ {}" + shows "\u\S. \v\S. \x\convex hull S. \y \ convex hull S. norm (x - y) \ norm (u - v)" proof - - have "convex hull s \ {}" - using hull_subset[of s convex] and assms(2) by auto - then obtain u v where obt: "u \ convex hull s" "v \ convex hull s" - "\x\convex hull s. \y\convex hull s. norm (x - y) \ norm (u - v)" + have "convex hull S \ {}" + using hull_subset[of S convex] and assms(2) by auto + then obtain u v where obt: "u \ convex hull S" "v \ convex hull S" + "\x\convex hull S. \y\convex hull S. norm (x - y) \ norm (u - v)" using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]] by (auto simp: dist_norm) then show ?thesis - proof (cases "u\s \ v\s", elim disjE) - assume "u \ s" - then obtain y where "y \ convex hull s" "norm (u - v) < norm (y - v)" + proof (cases "u\S \ v\S", elim disjE) + assume "u \ S" + then obtain y where "y \ convex hull S" "norm (u - v) < norm (y - v)" using simplex_furthest_lt[OF assms(1), THEN bspec[where x=u]] and obt(1) by auto then show ?thesis using obt(3)[THEN bspec[where x=y], THEN bspec[where x=v]] and obt(2) by auto next - assume "v \ s" - then obtain y where "y \ convex hull s" "norm (v - u) < norm (y - u)" + assume "v \ S" + then obtain y where "y \ convex hull S" "norm (v - u) < norm (y - u)" using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2) by auto then show ?thesis using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1) - by (auto simp add: norm_minus_commute) + by (auto simp: norm_minus_commute) qed auto qed lemma simplex_extremal_le_exists: - fixes s :: "'a::real_inner set" - shows "finite s \ x \ convex hull s \ y \ convex hull s \ - \u\s. \v\s. norm (x - y) \ norm (u - v)" - using convex_hull_empty simplex_extremal_le[of s] - by(cases "s = {}") auto + fixes S :: "'a::real_inner set" + shows "finite S \ x \ convex hull S \ y \ convex hull S \ + \u\S. \v\S. norm (x - y) \ norm (u - v)" + using convex_hull_empty simplex_extremal_le[of S] + by(cases "S = {}") auto subsection \Closest point of a convex set is unique, with a continuous projection\ definition%important closest_point :: "'a::{real_inner,heine_borel} set \ 'a \ 'a" - where "closest_point s a = (SOME x. x \ s \ (\y\s. dist a x \ dist a y))" + where "closest_point S a = (SOME x. x \ S \ (\y\S. dist a x \ dist a y))" lemma closest_point_exists: - assumes "closed s" - and "s \ {}" - shows "closest_point s a \ s" - and "\y\s. dist a (closest_point s a) \ dist a y" + assumes "closed S" + and "S \ {}" + shows "closest_point S a \ S" + and "\y\S. dist a (closest_point S a) \ dist a y" unfolding closest_point_def apply(rule_tac[!] someI2_ex) apply (auto intro: distance_attains_inf[OF assms(1,2), of a]) done -lemma closest_point_in_set: "closed s \ s \ {} \ closest_point s a \ s" +lemma closest_point_in_set: "closed S \ S \ {} \ closest_point S a \ S" by (meson closest_point_exists) -lemma closest_point_le: "closed s \ x \ s \ dist a (closest_point s a) \ dist a x" - using closest_point_exists[of s] by auto +lemma closest_point_le: "closed S \ x \ S \ dist a (closest_point S a) \ dist a x" + using closest_point_exists[of S] by auto lemma closest_point_self: - assumes "x \ s" - shows "closest_point s x = x" + assumes "x \ S" + shows "closest_point S x = x" unfolding closest_point_def apply (rule some1_equality, rule ex1I[of _ x]) using assms apply auto done -lemma closest_point_refl: "closed s \ s \ {} \ closest_point s x = x \ x \ s" - using closest_point_in_set[of s x] closest_point_self[of x s] +lemma closest_point_refl: "closed S \ S \ {} \ closest_point S x = x \ x \ S" + using closest_point_in_set[of S x] closest_point_self[of x S] by auto lemma closer_points_lemma: @@ -5519,18 +5207,13 @@ proof - have z: "inner z z > 0" unfolding inner_gt_zero_iff using assms by auto + have "norm (v *\<^sub>R z - y) < norm y" + if "0 < v" and "v \ inner y z / inner z z" for v + unfolding norm_lt using z assms that + by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ \0]) then show ?thesis - using assms - apply (rule_tac x = "inner y z / inner z z" in exI) - apply rule - defer - proof rule+ - fix v - assume "0 < v" and "v \ inner y z / inner z z" - then show "norm (v *\<^sub>R z - y) < norm y" - unfolding norm_lt using z and assms - by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ \0]) - qed auto + using assms z + by (rule_tac x = "inner y z / inner z z" in exI) auto qed lemma closer_point_lemma: @@ -5543,50 +5226,50 @@ show ?thesis apply (rule_tac x="min u 1" in exI) using u[THEN spec[where x="min u 1"]] and \u > 0\ - unfolding dist_norm by (auto simp add: norm_minus_commute field_simps) + unfolding dist_norm by (auto simp: norm_minus_commute field_simps) qed lemma any_closest_point_dot: - assumes "convex s" "closed s" "x \ s" "y \ s" "\z\s. dist a x \ dist a z" + assumes "convex S" "closed S" "x \ S" "y \ S" "\z\S. dist a x \ dist a z" shows "inner (a - x) (y - x) \ 0" proof (rule ccontr) assume "\ ?thesis" then obtain u where u: "u>0" "u\1" "dist (x + u *\<^sub>R (y - x)) a < dist x a" using closer_point_lemma[of a x y] by auto let ?z = "(1 - u) *\<^sub>R x + u *\<^sub>R y" - have "?z \ s" + have "?z \ S" using convexD_alt[OF assms(1,3,4), of u] using u by auto then show False using assms(5)[THEN bspec[where x="?z"]] and u(3) - by (auto simp add: dist_commute algebra_simps) + by (auto simp: dist_commute algebra_simps) qed lemma any_closest_point_unique: fixes x :: "'a::real_inner" - assumes "convex s" "closed s" "x \ s" "y \ s" - "\z\s. dist a x \ dist a z" "\z\s. dist a y \ dist a z" + assumes "convex S" "closed S" "x \ S" "y \ S" + "\z\S. dist a x \ dist a z" "\z\S. dist a y \ dist a z" shows "x = y" using any_closest_point_dot[OF assms(1-4,5)] and any_closest_point_dot[OF assms(1-2,4,3,6)] unfolding norm_pths(1) and norm_le_square - by (auto simp add: algebra_simps) + by (auto simp: algebra_simps) lemma closest_point_unique: - assumes "convex s" "closed s" "x \ s" "\z\s. dist a x \ dist a z" - shows "x = closest_point s a" - using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point s a"] + assumes "convex S" "closed S" "x \ S" "\z\S. dist a x \ dist a z" + shows "x = closest_point S a" + using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point S a"] using closest_point_exists[OF assms(2)] and assms(3) by auto lemma closest_point_dot: - assumes "convex s" "closed s" "x \ s" - shows "inner (a - closest_point s a) (x - closest_point s a) \ 0" + assumes "convex S" "closed S" "x \ S" + shows "inner (a - closest_point S a) (x - closest_point S a) \ 0" apply (rule any_closest_point_dot[OF assms(1,2) _ assms(3)]) using closest_point_exists[OF assms(2)] and assms(3) apply auto done lemma closest_point_lt: - assumes "convex s" "closed s" "x \ s" "x \ closest_point s a" - shows "dist a (closest_point s a) < dist a x" + assumes "convex S" "closed S" "x \ S" "x \ closest_point S a" + shows "dist a (closest_point S a) < dist a x" apply (rule ccontr) apply (rule_tac notE[OF assms(4)]) apply (rule closest_point_unique[OF assms(1-3), of a]) @@ -5595,34 +5278,34 @@ done lemma closest_point_lipschitz: - assumes "convex s" - and "closed s" "s \ {}" - shows "dist (closest_point s x) (closest_point s y) \ dist x y" + assumes "convex S" + and "closed S" "S \ {}" + shows "dist (closest_point S x) (closest_point S y) \ dist x y" proof - - have "inner (x - closest_point s x) (closest_point s y - closest_point s x) \ 0" - and "inner (y - closest_point s y) (closest_point s x - closest_point s y) \ 0" + have "inner (x - closest_point S x) (closest_point S y - closest_point S x) \ 0" + and "inner (y - closest_point S y) (closest_point S x - closest_point S y) \ 0" apply (rule_tac[!] any_closest_point_dot[OF assms(1-2)]) using closest_point_exists[OF assms(2-3)] apply auto done then show ?thesis unfolding dist_norm and norm_le - using inner_ge_zero[of "(x - closest_point s x) - (y - closest_point s y)"] + using inner_ge_zero[of "(x - closest_point S x) - (y - closest_point S y)"] by (simp add: inner_add inner_diff inner_commute) qed lemma continuous_at_closest_point: - assumes "convex s" - and "closed s" - and "s \ {}" - shows "continuous (at x) (closest_point s)" + assumes "convex S" + and "closed S" + and "S \ {}" + shows "continuous (at x) (closest_point S)" unfolding continuous_at_eps_delta using le_less_trans[OF closest_point_lipschitz[OF assms]] by auto lemma continuous_on_closest_point: - assumes "convex s" - and "closed s" - and "s \ {}" - shows "continuous_on t (closest_point s)" + assumes "convex S" + and "closed S" + and "S \ {}" + shows "continuous_on t (closest_point S)" by (metis continuous_at_imp_continuous_on continuous_at_closest_point[OF assms]) proposition closest_point_in_rel_interior: @@ -5647,7 +5330,7 @@ by (simp add: y_def algebra_simps) then have "norm (x - y) = abs ((1 - min 1 (e / norm (closest_point S x - x)))) * norm(x - closest_point S x)" by simp - also have "... < norm(x - closest_point S x)" + also have "\ < norm(x - closest_point S x)" using clo_notx \e > 0\ by (auto simp: mult_less_cancel_right2 divide_simps) finally have no_less: "norm (x - y) < norm (x - closest_point S x)" . @@ -5673,121 +5356,84 @@ lemma supporting_hyperplane_closed_point: fixes z :: "'a::{real_inner,heine_borel}" - assumes "convex s" - and "closed s" - and "s \ {}" - and "z \ s" - shows "\a b. \y\s. inner a z < b \ inner a y = b \ (\x\s. inner a x \ b)" + assumes "convex S" + and "closed S" + and "S \ {}" + and "z \ S" + shows "\a b. \y\S. inner a z < b \ inner a y = b \ (\x\S. inner a x \ b)" proof - - obtain y where "y \ s" and y: "\x\s. dist z y \ dist z x" + obtain y where "y \ S" and y: "\x\S. dist z y \ dist z x" by (metis distance_attains_inf[OF assms(2-3)]) show ?thesis - apply (rule_tac x="y - z" in exI) - apply (rule_tac x="inner (y - z) y" in exI) - apply (rule_tac x=y in bexI) - apply rule - defer - apply rule - defer - apply rule - apply (rule ccontr) - using \y \ s\ - proof - - show "inner (y - z) z < inner (y - z) y" - apply (subst diff_gt_0_iff_gt [symmetric]) - unfolding inner_diff_right[symmetric] and inner_gt_zero_iff - using \y\s\ \z\s\ - apply auto - done - next - fix x - assume "x \ s" - have *: "\u. 0 \ u \ u \ 1 \ dist z y \ dist z ((1 - u) *\<^sub>R y + u *\<^sub>R x)" - using assms(1)[unfolded convex_alt] and y and \x\s\ and \y\s\ by auto - assume "\ inner (y - z) y \ inner (y - z) x" - then obtain v where "v > 0" "v \ 1" "dist (y + v *\<^sub>R (x - y)) z < dist y z" - using closer_point_lemma[of z y x] by (auto simp add: inner_diff) - then show False - using *[THEN spec[where x=v]] by (auto simp add: dist_commute algebra_simps) - qed auto + proof (intro exI bexI conjI ballI) + show "(y - z) \ z < (y - z) \ y" + by (metis \y \ S\ assms(4) diff_gt_0_iff_gt inner_commute inner_diff_left inner_gt_zero_iff right_minus_eq) + show "(y - z) \ y \ (y - z) \ x" if "x \ S" for x + proof (rule ccontr) + have *: "\u. 0 \ u \ u \ 1 \ dist z y \ dist z ((1 - u) *\<^sub>R y + u *\<^sub>R x)" + using assms(1)[unfolded convex_alt] and y and \x\S\ and \y\S\ by auto + assume "\ (y - z) \ y \ (y - z) \ x" + then obtain v where "v > 0" "v \ 1" "dist (y + v *\<^sub>R (x - y)) z < dist y z" + using closer_point_lemma[of z y x] by (auto simp: inner_diff) + then show False + using *[of v] by (auto simp: dist_commute algebra_simps) + qed + qed (use \y \ S\ in auto) qed lemma separating_hyperplane_closed_point: fixes z :: "'a::{real_inner,heine_borel}" - assumes "convex s" - and "closed s" - and "z \ s" - shows "\a b. inner a z < b \ (\x\s. inner a x > b)" -proof (cases "s = {}") + assumes "convex S" + and "closed S" + and "z \ S" + shows "\a b. inner a z < b \ (\x\S. inner a x > b)" +proof (cases "S = {}") case True then show ?thesis - apply (rule_tac x="-z" in exI) - apply (rule_tac x=1 in exI) - using less_le_trans[OF _ inner_ge_zero[of z]] - apply auto - done + by (simp add: gt_ex) next case False - obtain y where "y \ s" and y: "\x\s. dist z y \ dist z x" + obtain y where "y \ S" and y: "\x. x \ S \ dist z y \ dist z x" by (metis distance_attains_inf[OF assms(2) False]) show ?thesis - apply (rule_tac x="y - z" in exI) - apply (rule_tac x="inner (y - z) z + (norm (y - z))\<^sup>2 / 2" in exI) - apply rule - defer - apply rule - proof - + proof (intro exI conjI ballI) + show "(y - z) \ z < inner (y - z) z + (norm (y - z))\<^sup>2 / 2" + using \y\S\ \z\S\ by auto + next fix x - assume "x \ s" - have "\ 0 < inner (z - y) (x - y)" - apply (rule notI) - apply (drule closer_point_lemma) + assume "x \ S" + have "False" if *: "0 < inner (z - y) (x - y)" proof - - assume "\u>0. u \ 1 \ dist (y + u *\<^sub>R (x - y)) z < dist y z" - then obtain u where "u > 0" "u \ 1" "dist (y + u *\<^sub>R (x - y)) z < dist y z" - by auto - then show False using y[THEN bspec[where x="y + u *\<^sub>R (x - y)"]] - using assms(1)[unfolded convex_alt, THEN bspec[where x=y]] - using \x\s\ \y\s\ by (auto simp add: dist_commute algebra_simps) + obtain u where "u > 0" "u \ 1" "dist (y + u *\<^sub>R (x - y)) z < dist y z" + using * closer_point_lemma by blast + then show False using y[of "y + u *\<^sub>R (x - y)"] convexD_alt [OF \convex S\] + using \x\S\ \y\S\ by (auto simp: dist_commute algebra_simps) qed moreover have "0 < (norm (y - z))\<^sup>2" - using \y\s\ \z\s\ by auto + using \y\S\ \z\S\ by auto then have "0 < inner (y - z) (y - z)" unfolding power2_norm_eq_inner by simp - ultimately show "inner (y - z) z + (norm (y - z))\<^sup>2 / 2 < inner (y - z) x" - unfolding power2_norm_eq_inner and not_less - by (auto simp add: field_simps inner_commute inner_diff) - qed (insert \y\s\ \z\s\, auto) + ultimately show "(y - z) \ z + (norm (y - z))\<^sup>2 / 2 < (y - z) \ x" + by (force simp: field_simps power2_norm_eq_inner inner_commute inner_diff) + qed qed lemma separating_hyperplane_closed_0: - assumes "convex (s::('a::euclidean_space) set)" - and "closed s" - and "0 \ s" - shows "\a b. a \ 0 \ 0 < b \ (\x\s. inner a x > b)" -proof (cases "s = {}") + assumes "convex (S::('a::euclidean_space) set)" + and "closed S" + and "0 \ S" + shows "\a b. a \ 0 \ 0 < b \ (\x\S. inner a x > b)" +proof (cases "S = {}") case True - have "norm ((SOME i. i\Basis)::'a) = 1" "(SOME i. i\Basis) \ (0::'a)" - defer - apply (subst norm_le_zero_iff[symmetric]) - apply (auto simp: SOME_Basis) - done + have "(SOME i. i\Basis) \ (0::'a)" + by (metis Basis_zero SOME_Basis) then show ?thesis - apply (rule_tac x="SOME i. i\Basis" in exI) - apply (rule_tac x=1 in exI) - using True using DIM_positive[where 'a='a] - apply auto - done + using True zero_less_one by blast next case False then show ?thesis using False using separating_hyperplane_closed_point[OF assms] - apply (elim exE) - unfolding inner_zero_right - apply (rule_tac x=a in exI) - apply (rule_tac x=b in exI) - apply auto - done + by (metis all_not_in_conv inner_zero_left inner_zero_right less_eq_real_def not_le) qed @@ -5826,7 +5472,7 @@ apply rule apply rule apply (erule_tac x="x - y" in ballE) - apply (auto simp add: inner_diff) + apply (auto simp: inner_diff) done define k where "k = (SUP x:T. a \ x)" show ?thesis @@ -5876,8 +5522,7 @@ by auto then show ?thesis apply (rule_tac x="-a" in exI) - apply (rule_tac x="-b" in exI) - apply auto + apply (rule_tac x="-b" in exI, auto) done qed @@ -5885,13 +5530,13 @@ subsubsection%unimportant \General case without assuming closure and getting non-strict separation\ lemma separating_hyperplane_set_0: - assumes "convex s" "(0::'a::euclidean_space) \ s" - shows "\a. a \ 0 \ (\x\s. 0 \ inner a x)" + assumes "convex S" "(0::'a::euclidean_space) \ S" + shows "\a. a \ 0 \ (\x\S. 0 \ inner a x)" proof - let ?k = "\c. {x::'a. 0 \ inner c x}" - have *: "frontier (cball 0 1) \ \f \ {}" if as: "f \ ?k ` s" "finite f" for f + have *: "frontier (cball 0 1) \ \f \ {}" if as: "f \ ?k ` S" "finite f" for f proof - - obtain c where c: "f = ?k ` c" "c \ s" "finite c" + obtain c where c: "f = ?k ` c" "c \ S" "finite c" using finite_subset_image[OF as(2,1)] by auto then obtain a b where ab: "a \ 0" "0 < b" "\x\convex hull c. b < inner a x" using separating_hyperplane_closed_0[OF convex_convex_hull, of c] @@ -5902,50 +5547,50 @@ apply (rule_tac x = "inverse(norm a) *\<^sub>R a" in exI) using hull_subset[of c convex] unfolding subset_eq and inner_scaleR - by (auto simp add: inner_commute del: ballE elim!: ballE) + by (auto simp: inner_commute del: ballE elim!: ballE) then show "frontier (cball 0 1) \ \f \ {}" unfolding c(1) frontier_cball sphere_def dist_norm by auto qed - have "frontier (cball 0 1) \ (\(?k ` s)) \ {}" + have "frontier (cball 0 1) \ (\(?k ` S)) \ {}" apply (rule compact_imp_fip) apply (rule compact_frontier[OF compact_cball]) using * closed_halfspace_ge by auto - then obtain x where "norm x = 1" "\y\s. x\?k y" + then obtain x where "norm x = 1" "\y\S. x\?k y" unfolding frontier_cball dist_norm sphere_def by auto then show ?thesis by (metis inner_commute mem_Collect_eq norm_eq_zero zero_neq_one) qed lemma separating_hyperplane_sets: - fixes s t :: "'a::euclidean_space set" - assumes "convex s" - and "convex t" - and "s \ {}" - and "t \ {}" - and "s \ t = {}" - shows "\a b. a \ 0 \ (\x\s. inner a x \ b) \ (\x\t. inner a x \ b)" + fixes S T :: "'a::euclidean_space set" + assumes "convex S" + and "convex T" + and "S \ {}" + and "T \ {}" + and "S \ T = {}" + shows "\a b. a \ 0 \ (\x\S. inner a x \ b) \ (\x\T. inner a x \ b)" proof - from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]] - obtain a where "a \ 0" "\x\{x - y |x y. x \ t \ y \ s}. 0 \ inner a x" + obtain a where "a \ 0" "\x\{x - y |x y. x \ T \ y \ S}. 0 \ inner a x" using assms(3-5) by force - then have *: "\x y. x \ t \ y \ s \ inner a y \ inner a x" - by (force simp add: inner_diff) - then have bdd: "bdd_above (((\) a)`s)" - using \t \ {}\ by (auto intro: bdd_aboveI2[OF *]) + then have *: "\x y. x \ T \ y \ S \ inner a y \ inner a x" + by (force simp: inner_diff) + then have bdd: "bdd_above (((\) a)`S)" + using \T \ {}\ by (auto intro: bdd_aboveI2[OF *]) show ?thesis using \a\0\ - by (intro exI[of _ a] exI[of _ "SUP x:s. a \ x"]) - (auto intro!: cSUP_upper bdd cSUP_least \a \ 0\ \s \ {}\ *) + by (intro exI[of _ a] exI[of _ "SUP x:S. a \ x"]) + (auto intro!: cSUP_upper bdd cSUP_least \a \ 0\ \S \ {}\ *) qed subsection%unimportant \More convexity generalities\ lemma convex_closure [intro,simp]: - fixes s :: "'a::real_normed_vector set" - assumes "convex s" - shows "convex (closure s)" + fixes S :: "'a::real_normed_vector set" + assumes "convex S" + shows "convex (closure S)" apply (rule convexI) apply (unfold closure_sequential, elim exE) apply (rule_tac x="\n. u *\<^sub>R xa n + v *\<^sub>R xb n" in exI) @@ -5955,65 +5600,58 @@ done lemma convex_interior [intro,simp]: - fixes s :: "'a::real_normed_vector set" - assumes "convex s" - shows "convex (interior s)" + fixes S :: "'a::real_normed_vector set" + assumes "convex S" + shows "convex (interior S)" unfolding convex_alt Ball_def mem_interior - apply (rule,rule,rule,rule,rule,rule) - apply (elim exE conjE) -proof - +proof clarify fix x y u assume u: "0 \ u" "u \ (1::real)" fix e d - assume ed: "ball x e \ s" "ball y d \ s" "0e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \ s" - apply (rule_tac x="min d e" in exI) - apply rule - unfolding subset_eq - defer - apply rule - proof - + assume ed: "ball x e \ S" "ball y d \ S" "0e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \ S" + proof (intro exI conjI subsetI) fix z assume "z \ ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) (min d e)" - then have "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \ s" + then have "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \ S" apply (rule_tac assms[unfolded convex_alt, rule_format]) using ed(1,2) and u unfolding subset_eq mem_ball Ball_def dist_norm - apply (auto simp add: algebra_simps) + apply (auto simp: algebra_simps) done - then show "z \ s" - using u by (auto simp add: algebra_simps) + then show "z \ S" + using u by (auto simp: algebra_simps) qed(insert u ed(3-4), auto) qed -lemma convex_hull_eq_empty[simp]: "convex hull s = {} \ s = {}" - using hull_subset[of s convex] convex_hull_empty by auto +lemma convex_hull_eq_empty[simp]: "convex hull S = {} \ S = {}" + using hull_subset[of S convex] convex_hull_empty by auto subsection%unimportant \Moving and scaling convex hulls\ lemma convex_hull_set_plus: - "convex hull (s + t) = convex hull s + convex hull t" + "convex hull (S + T) = convex hull S + convex hull T" unfolding set_plus_image apply (subst convex_hull_linear_image [symmetric]) apply (simp add: linear_iff scaleR_right_distrib) apply (simp add: convex_hull_Times) done -lemma translation_eq_singleton_plus: "(\x. a + x) ` t = {a} + t" +lemma translation_eq_singleton_plus: "(\x. a + x) ` T = {a} + T" unfolding set_plus_def by auto lemma convex_hull_translation: - "convex hull ((\x. a + x) ` s) = (\x. a + x) ` (convex hull s)" + "convex hull ((\x. a + x) ` S) = (\x. a + x) ` (convex hull S)" unfolding translation_eq_singleton_plus by (simp only: convex_hull_set_plus convex_hull_singleton) lemma convex_hull_scaling: - "convex hull ((\x. c *\<^sub>R x) ` s) = (\x. c *\<^sub>R x) ` (convex hull s)" - by (simp add: convex_hull_linear_image) + "convex hull ((\x. c *\<^sub>R x) ` S) = (\x. c *\<^sub>R x) ` (convex hull S)" + using linear_scaleR by (rule convex_hull_linear_image [symmetric]) lemma convex_hull_affinity: - "convex hull ((\x. a + c *\<^sub>R x) ` s) = (\x. a + c *\<^sub>R x) ` (convex hull s)" + "convex hull ((\x. a + c *\<^sub>R x) ` S) = (\x. a + c *\<^sub>R x) ` (convex hull S)" by(simp only: image_image[symmetric] convex_hull_scaling convex_hull_translation) @@ -6051,7 +5689,7 @@ using assms mem_convex_alt[of S xx yy cx cy] x y by auto then have "cx *\<^sub>R xx + cy *\<^sub>R yy \ cone hull S" using mem_cone_hull[of "(cx/(cx+cy)) *\<^sub>R xx + (cy/(cx+cy)) *\<^sub>R yy" S "cx+cy"] \cx+cy>0\ - by (auto simp add: scaleR_right_distrib) + by (auto simp: scaleR_right_distrib) then have "u *\<^sub>R x + v *\<^sub>R y \ cone hull S" using x y by auto } @@ -6091,8 +5729,7 @@ fixes s :: "('a::euclidean_space) set" assumes "closed s" "convex s" shows "s = \{h. s \ h \ (\a b. h = {x. inner a x \ b})}" - apply (rule set_eqI) - apply rule + apply (rule set_eqI, rule) unfolding Inter_iff Ball_def mem_Collect_eq apply (rule,rule,erule conjE) proof - @@ -6105,8 +5742,7 @@ apply (drule separating_hyperplane_closed_point[OF assms(2,1)]) apply (erule exE)+ apply (erule_tac x="-a" in allE) - apply (erule_tac x="-b" in allE) - apply auto + apply (erule_tac x="-b" in allE, auto) done qed auto @@ -6124,7 +5760,7 @@ then show ?thesis apply (rule_tac x="\v. if v\s then u v else 0" in exI) unfolding if_smult scaleR_zero_left and sum.inter_restrict[OF assms(1), symmetric] - apply (auto simp add: Int_absorb1) + apply (auto simp: Int_absorb1) done qed @@ -6179,8 +5815,7 @@ next case False then have "sum u c \ sum (\x. if x=v then u v else 0) c" - apply (rule_tac sum_mono) - apply auto + apply (rule_tac sum_mono, auto) done then show ?thesis unfolding sum.delta[OF assms(1)] using uv(2) and \u v < 0\ and uv(1) by auto @@ -6190,20 +5825,18 @@ then have *: "sum u {x\c. u x > 0} > 0" unfolding less_le apply (rule_tac conjI) - apply (rule_tac sum_nonneg) - apply auto + apply (rule_tac sum_nonneg, auto) done moreover have "sum u ({x \ c. 0 < u x} \ {x \ c. u x < 0}) = sum u c" "(\x\{x \ c. 0 < u x} \ {x \ c. u x < 0}. u x *\<^sub>R x) = (\x\c. u x *\<^sub>R x)" using assms(1) - apply (rule_tac[!] sum.mono_neutral_left) - apply auto + apply (rule_tac[!] sum.mono_neutral_left, auto) done then have "sum u {x \ c. 0 < u x} = - sum u {x \ c. 0 > u x}" "(\x\{x \ c. 0 < u x}. u x *\<^sub>R x) = - (\x\{x \ c. 0 > u x}. u x *\<^sub>R x)" unfolding eq_neg_iff_add_eq_0 using uv(1,4) - by (auto simp add: sum.union_inter_neutral[OF fin, symmetric]) + by (auto simp: sum.union_inter_neutral[OF fin, symmetric]) moreover have "\x\{v \ c. u v < 0}. 0 \ inverse (sum u {x \ c. 0 < u x}) * - u x" apply rule apply (rule mult_nonneg_nonneg) @@ -6215,7 +5848,7 @@ apply (rule_tac x="{v \ c. u v < 0}" in exI) apply (rule_tac x="\y. inverse (sum u {x\c. u x > 0}) * - u y" in exI) using assms(1) unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] and z_def - apply (auto simp add: sum_negf sum_distrib_left[symmetric]) + apply (auto simp: sum_negf sum_distrib_left[symmetric]) done moreover have "\x\{v \ c. 0 < u v}. 0 \ inverse (sum u {x \ c. 0 < u x}) * u x" apply rule @@ -6230,12 +5863,11 @@ using assms(1) unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] and z_def using * - apply (auto simp add: sum_negf sum_distrib_left[symmetric]) + apply (auto simp: sum_negf sum_distrib_left[symmetric]) done ultimately show ?thesis apply (rule_tac x="{v\c. u v \ 0}" in exI) - apply (rule_tac x="{v\c. u v > 0}" in exI) - apply auto + apply (rule_tac x="{v\c. u v > 0}" in exI, auto) done qed @@ -6269,7 +5901,7 @@ and "\s\f. convex s" "\t\f. card t = DIM('a) + 1 \ \t \ {}" shows "\f \ {}" using assms -proof (induct n arbitrary: f) +proof (induction n arbitrary: f) case 0 then show ?case by auto next @@ -6277,46 +5909,39 @@ have "finite f" using \card f = Suc n\ by (auto intro: card_ge_0_finite) show "\f \ {}" - apply (cases "n = DIM('a)") - apply (rule Suc(5)[rule_format]) - unfolding \card f = Suc n\ - proof - - assume ng: "n \ DIM('a)" - then have "\X. \s\f. X s \ \(f - {s})" - apply (rule_tac bchoice) - unfolding ex_in_conv - apply (rule, rule Suc(1)[rule_format]) - unfolding card_Diff_singleton_if[OF \finite f\] \card f = Suc n\ - defer - defer - apply (rule Suc(4)[rule_format]) - defer - apply (rule Suc(5)[rule_format]) - using Suc(3) \finite f\ - apply auto - done - then obtain X where X: "\s\f. X s \ \(f - {s})" by auto + proof (cases "n = DIM('a)") + case True + then show ?thesis + by (simp add: Suc.prems(1) Suc.prems(4)) + next + case False + have "\(f - {s}) \ {}" if "s \ f" for s + proof (rule Suc.IH[rule_format]) + show "card (f - {s}) = n" + by (simp add: Suc.prems(1) \finite f\ that) + show "DIM('a) + 1 \ n" + using False Suc.prems(2) by linarith + show "\t. \t \ f - {s}; card t = DIM('a) + 1\ \ \t \ {}" + by (simp add: Suc.prems(4) subset_Diff_insert) + qed (use Suc in auto) + then have "\s\f. \x. x \ \(f - {s})" + by blast + then obtain X where X: "\s. s\f \ X s \ \(f - {s})" + by metis show ?thesis proof (cases "inj_on X f") case False - then obtain s t where st: "s\t" "s\f" "t\f" "X s = X t" + then obtain s t where "s\t" and st: "s\f" "t\f" "X s = X t" unfolding inj_on_def by auto then have *: "\f = \(f - {s}) \ \(f - {t})" by auto show ?thesis - unfolding * - unfolding ex_in_conv[symmetric] - apply (rule_tac x="X s" in exI) - apply rule - apply (rule X[rule_format]) - using X st - apply auto - done + by (metis "*" X disjoint_iff_not_equal st) next case True then obtain m p where mp: "m \ p = {}" "m \ p = X ` f" "convex hull m \ convex hull p \ {}" using radon_partition[of "X ` f"] and affine_dependent_biggerset[of "X ` f"] unfolding card_image[OF True] and \card f = Suc n\ - using Suc(3) \finite f\ and ng + using Suc(3) \finite f\ and False by auto have "m \ X ` f" "p \ X ` f" using mp(2) by auto @@ -6333,38 +5958,17 @@ using inj_on_image_Int[OF True gh(3,4)] by auto have "convex hull (X ` h) \ \g" "convex hull (X ` g) \ \h" - apply (rule_tac [!] hull_minimal) - using Suc gh(3-4) - unfolding subset_eq - apply (rule_tac [2] convex_Inter, rule_tac [4] convex_Inter) - apply rule - prefer 3 - apply rule - proof - - fix x - assume "x \ X ` g" - then obtain y where "y \ g" "x = X y" - unfolding image_iff .. - then show "x \ \h" - using X[THEN bspec[where x=y]] using * f by auto - next - fix x - assume "x \ X ` h" - then obtain y where "y \ h" "x = X y" - unfolding image_iff .. - then show "x \ \g" - using X[THEN bspec[where x=y]] using * f by auto - qed auto + by (rule hull_minimal; use X * f in \auto simp: Suc.prems(3) convex_Inter\)+ then show ?thesis unfolding f using mp(3)[unfolded gh] by blast qed - qed auto + qed qed theorem%important helly: fixes f :: "'a::euclidean_space set set" assumes "card f \ DIM('a) + 1" "\s\f. convex s" - and "\t\f. card t = DIM('a) + 1 \ \t \ {}" + and "\t. \t\f; card t = DIM('a) + 1\ \ \t \ {}" shows "\f \ {}" apply%unimportant (rule helly_induct) using assms @@ -6374,104 +5978,109 @@ subsection \Epigraphs of convex functions\ -definition%important "epigraph s (f :: _ \ real) = {xy. fst xy \ s \ f (fst xy) \ snd xy}" - -lemma mem_epigraph: "(x, y) \ epigraph s f \ x \ s \ f x \ y" +definition%important "epigraph S (f :: _ \ real) = {xy. fst xy \ S \ f (fst xy) \ snd xy}" + +lemma mem_epigraph: "(x, y) \ epigraph S f \ x \ S \ f x \ y" unfolding epigraph_def by auto -lemma convex_epigraph: "convex (epigraph s f) \ convex_on s f \ convex s" - unfolding convex_def convex_on_def - unfolding Ball_def split_paired_All epigraph_def - unfolding mem_Collect_eq fst_conv snd_conv fst_add snd_add fst_scaleR snd_scaleR Ball_def[symmetric] - apply safe - defer - apply (erule_tac x=x in allE) - apply (erule_tac x="f x" in allE) - apply safe - apply (erule_tac x=xa in allE) - apply (erule_tac x="f xa" in allE) - prefer 3 - apply (rule_tac y="u * f a + v * f aa" in order_trans) - defer - apply (auto intro!:mult_left_mono add_mono) - done - -lemma convex_epigraphI: "convex_on s f \ convex s \ convex (epigraph s f)" +lemma convex_epigraph: "convex (epigraph S f) \ convex_on S f \ convex S" +proof safe + assume L: "convex (epigraph S f)" + then show "convex_on S f" + by (auto simp: convex_def convex_on_def epigraph_def) + show "convex S" + using L + apply (clarsimp simp: convex_def convex_on_def epigraph_def) + apply (erule_tac x=x in allE) + apply (erule_tac x="f x" in allE, safe) + apply (erule_tac x=y in allE) + apply (erule_tac x="f y" in allE) + apply (auto simp: ) + done +next + assume "convex_on S f" "convex S" + then show "convex (epigraph S f)" + unfolding convex_def convex_on_def epigraph_def + apply safe + apply (rule_tac [2] y="u * f a + v * f aa" in order_trans) + apply (auto intro!:mult_left_mono add_mono) + done +qed + +lemma convex_epigraphI: "convex_on S f \ convex S \ convex (epigraph S f)" unfolding convex_epigraph by auto -lemma convex_epigraph_convex: "convex s \ convex_on s f \ convex(epigraph s f)" +lemma convex_epigraph_convex: "convex S \ convex_on S f \ convex(epigraph S f)" by (simp add: convex_epigraph) subsubsection%unimportant \Use this to derive general bound property of convex function\ lemma convex_on: - assumes "convex s" - shows "convex_on s f \ - (\k u x. (\i\{1..k::nat}. 0 \ u i \ x i \ s) \ sum u {1..k} = 1 \ - f (sum (\i. u i *\<^sub>R x i) {1..k} ) \ sum (\i. u i * f(x i)) {1..k})" + assumes "convex S" + shows "convex_on S f \ + (\k u x. (\i\{1..k::nat}. 0 \ u i \ x i \ S) \ sum u {1..k} = 1 \ + f (sum (\i. u i *\<^sub>R x i) {1..k}) \ sum (\i. u i * f(x i)) {1..k})" + unfolding convex_epigraph_convex[OF assms] convex epigraph_def Ball_def mem_Collect_eq unfolding fst_sum snd_sum fst_scaleR snd_scaleR apply safe - apply (drule_tac x=k in spec) - apply (drule_tac x=u in spec) - apply (drule_tac x="\i. (x i, f (x i))" in spec) - apply simp - using assms[unfolded convex] - apply simp - apply (rule_tac y="\i = 1..k. u i * f (fst (x i))" in order_trans) - defer - apply (rule sum_mono) - apply (erule_tac x=i in allE) + apply (drule_tac x=k in spec) + apply (drule_tac x=u in spec) + apply (drule_tac x="\i. (x i, f (x i))" in spec) + apply simp + using assms[unfolded convex] apply simp + apply (rule_tac y="\i = 1..k. u i * f (fst (x i))" in order_trans, force) + apply (rule sum_mono) + apply (erule_tac x=i in allE) unfolding real_scaleR_def - apply (rule mult_left_mono) - using assms[unfolded convex] - apply auto + apply (rule mult_left_mono) + using assms[unfolded convex] apply auto done subsection%unimportant \Convexity of general and special intervals\ lemma is_interval_convex: - fixes s :: "'a::euclidean_space set" - assumes "is_interval s" - shows "convex s" + fixes S :: "'a::euclidean_space set" + assumes "is_interval S" + shows "convex S" proof (rule convexI) fix x y and u v :: real - assume as: "x \ s" "y \ s" "0 \ u" "0 \ v" "u + v = 1" + assume as: "x \ S" "y \ S" "0 \ u" "0 \ v" "u + v = 1" then have *: "u = 1 - v" "1 - v \ 0" and **: "v = 1 - u" "1 - u \ 0" by auto { fix a b assume "\ b \ u * a + v * b" then have "u * a < (1 - v) * b" - unfolding not_le using as(4) by (auto simp add: field_simps) + unfolding not_le using as(4) by (auto simp: field_simps) then have "a < b" unfolding * using as(4) *(2) apply (rule_tac mult_left_less_imp_less[of "1 - v"]) - apply (auto simp add: field_simps) + apply (auto simp: field_simps) done then have "a \ u * a + v * b" unfolding * using as(4) - by (auto simp add: field_simps intro!:mult_right_mono) + by (auto simp: field_simps intro!:mult_right_mono) } moreover { fix a b assume "\ u * a + v * b \ a" then have "v * b > (1 - u) * a" - unfolding not_le using as(4) by (auto simp add: field_simps) + unfolding not_le using as(4) by (auto simp: field_simps) then have "a < b" unfolding * using as(4) apply (rule_tac mult_left_less_imp_less) - apply (auto simp add: field_simps) + apply (auto simp: field_simps) done then have "u * a + v * b \ b" unfolding ** using **(2) as(3) - by (auto simp add: field_simps intro!:mult_right_mono) + by (auto simp: field_simps intro!:mult_right_mono) } - ultimately show "u *\<^sub>R x + v *\<^sub>R y \ s" + ultimately show "u *\<^sub>R x + v *\<^sub>R y \ S" apply - apply (rule assms[unfolded is_interval_def, rule_format, OF as(1,2)]) using as(3-) DIM_positive[where 'a='a] @@ -6480,8 +6089,8 @@ qed lemma is_interval_connected: - fixes s :: "'a::euclidean_space set" - shows "is_interval s \ connected s" + fixes S :: "'a::euclidean_space set" + shows "is_interval S \ connected S" using is_interval_convex convex_connected by auto lemma convex_box [simp]: "convex (cbox a b)" "convex (box a (b::'a::euclidean_space))" @@ -6556,12 +6165,9 @@ ultimately show False apply (rule_tac notE[OF as(1)[unfolded connected_def]]) apply (rule_tac x = ?halfl in exI) - apply (rule_tac x = ?halfr in exI) - apply rule - apply (rule open_lessThan) - apply rule - apply (rule open_greaterThan) - apply auto + apply (rule_tac x = ?halfr in exI, rule) + apply (rule open_lessThan, rule) + apply (rule open_greaterThan, auto) done qed @@ -6625,7 +6231,7 @@ fixes f :: "real \ 'a::euclidean_space" shows "a \ b \ \x\{a..b}. continuous (at x) f \ f a\k \ y \ y \ f b\k \ \x\{a..b}. (f x)\k = y" - by (rule ivt_increasing_component_on_1) (auto simp add: continuous_at_imp_continuous_on) + by (rule ivt_increasing_component_on_1) (auto simp: continuous_at_imp_continuous_on) lemma ivt_decreasing_component_on_1: fixes f :: "real \ 'a::euclidean_space" @@ -6699,40 +6305,22 @@ assumes "finite A" and "\i\A. finite (B i)" shows "finite (\i\A. B i)" using assms by (induct set: finite, simp, simp add: finite_set_plus) -lemma set_sum_eq: - "finite A \ (\i\A. B i) = {\i\A. f i |f. \i\A. f i \ B i}" - apply (induct set: finite) - apply simp - apply simp - apply (safe elim!: set_plus_elim) - apply (rule_tac x="fun_upd f x a" in exI) - apply simp - apply (rule_tac f="\x. a + x" in arg_cong) - apply (rule sum.cong [OF refl]) - apply clarsimp - apply fast - done - lemma box_eq_set_sum_Basis: shows "{x. \i\Basis. x\i \ B i} = (\i\Basis. image (\x. x *\<^sub>R i) (B i))" - apply (subst set_sum_eq [OF finite_Basis]) - apply safe + apply (subst set_sum_alt [OF finite_Basis], safe) apply (fast intro: euclidean_representation [symmetric]) apply (subst inner_sum_left) +apply (rename_tac f) apply (subgoal_tac "(\x\Basis. f x \ i) = f i \ i") apply (drule (1) bspec) apply clarsimp apply (frule sum.remove [OF finite_Basis]) - apply (erule trans) - apply simp - apply (rule sum.neutral) - apply clarsimp + apply (erule trans, simp) + apply (rule sum.neutral, clarsimp) apply (frule_tac x=i in bspec, assumption) - apply (drule_tac x=x in bspec, assumption) - apply clarsimp + apply (drule_tac x=x in bspec, assumption, clarsimp) apply (cut_tac u=x and v=i in inner_Basis, assumption+) - apply (rule ccontr) - apply simp + apply (rule ccontr, simp) done lemma convex_hull_set_sum: @@ -6750,8 +6338,8 @@ show "convex (cbox x y)" by (rule convex_box) next - fix s assume "{x, y} \ s" and "convex s" - then show "cbox x y \ s" + fix S assume "{x, y} \ S" and "convex S" + then show "cbox x y \ S" unfolding is_interval_convex_1 [symmetric] is_interval_def Basis_real_def by - (clarify, simp (no_asm_use), fast) qed @@ -6782,74 +6370,53 @@ text \And this is a finite set of vertices.\ lemma unit_cube_convex_hull: - obtains s :: "'a::euclidean_space set" - where "finite s" and "cbox 0 (\Basis) = convex hull s" - apply (rule that[of "{x::'a. \i\Basis. x\i=0 \ x\i=1}"]) - apply (rule finite_subset[of _ "(\s. (\i\Basis. (if i\s then 1 else 0) *\<^sub>R i)::'a) ` Pow Basis"]) - prefer 3 - apply (rule unit_interval_convex_hull) - apply rule - unfolding mem_Collect_eq -proof - - fix x :: 'a - assume as: "\i\Basis. x \ i = 0 \ x \ i = 1" - show "x \ (\s. \i\Basis. (if i\s then 1 else 0) *\<^sub>R i) ` Pow Basis" - apply (rule image_eqI[where x="{i. i\Basis \ x\i = 1}"]) - using as - apply (subst euclidean_eq_iff) - apply auto - done -qed auto + obtains S :: "'a::euclidean_space set" + where "finite S" and "cbox 0 (\Basis) = convex hull S" +proof + show "finite {x::'a. \i\Basis. x \ i = 0 \ x \ i = 1}" + proof (rule finite_subset, clarify) + show "finite ((\S. \i\Basis. (if i \ S then 1 else 0) *\<^sub>R i) ` Pow Basis)" + using finite_Basis by blast + fix x :: 'a + assume as: "\i\Basis. x \ i = 0 \ x \ i = 1" + show "x \ (\S. \i\Basis. (if i\S then 1 else 0) *\<^sub>R i) ` Pow Basis" + apply (rule image_eqI[where x="{i. i\Basis \ x\i = 1}"]) + using as + apply (subst euclidean_eq_iff, auto) + done + qed + show "cbox 0 One = convex hull {x. \i\Basis. x \ i = 0 \ x \ i = 1}" + using unit_interval_convex_hull by blast +qed text \Hence any cube (could do any nonempty interval).\ lemma cube_convex_hull: assumes "d > 0" - obtains s :: "'a::euclidean_space set" where - "finite s" and "cbox (x - (\i\Basis. d*\<^sub>Ri)) (x + (\i\Basis. d*\<^sub>Ri)) = convex hull s" + obtains S :: "'a::euclidean_space set" where + "finite S" and "cbox (x - (\i\Basis. d*\<^sub>Ri)) (x + (\i\Basis. d*\<^sub>Ri)) = convex hull S" proof - - let ?d = "(\i\Basis. d*\<^sub>Ri)::'a" + let ?d = "(\i\Basis. d *\<^sub>R i)::'a" have *: "cbox (x - ?d) (x + ?d) = (\y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 (\Basis)" - apply (rule set_eqI, rule) - unfolding image_iff - defer - apply (erule bexE) - proof - + proof (intro set_eqI iffI) fix y - assume as: "y\cbox (x - ?d) (x + ?d)" + assume "y \ cbox (x - ?d) (x + ?d)" then have "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \ cbox 0 (\Basis)" using assms by (simp add: mem_box field_simps inner_simps) - with \0 < d\ show "\z\cbox 0 (\Basis). y = x - ?d + (2 * d) *\<^sub>R z" - by (intro bexI[of _ "inverse (2 * d) *\<^sub>R (y - (x - ?d))"]) auto + with \0 < d\ show "y \ (\y. x - sum (( *\<^sub>R) d) Basis + (2 * d) *\<^sub>R y) ` cbox 0 One" + by (auto intro: image_eqI[where x= "inverse (2 * d) *\<^sub>R (y - (x - ?d))"]) next - fix y z - assume as: "z\cbox 0 (\Basis)" "y = x - ?d + (2*d) *\<^sub>R z" - have "\i. i\Basis \ 0 \ d * (z \ i) \ d * (z \ i) \ d" - using assms as(1)[unfolded mem_box] - apply (erule_tac x=i in ballE) - apply rule - prefer 2 - apply (rule mult_right_le_one_le) - using assms - apply auto - done + fix y + assume "y \ (\y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 One" + then obtain z where z: "z \ cbox 0 One" "y = x - ?d + (2*d) *\<^sub>R z" + by auto then show "y \ cbox (x - ?d) (x + ?d)" - unfolding as(2) mem_box - apply - - apply rule - using as(1)[unfolded mem_box] - apply (erule_tac x=i in ballE) - using assms - apply (auto simp: inner_simps) - done + using z assms by (auto simp: mem_box inner_simps) qed - obtain s where "finite s" "cbox 0 (\Basis::'a) = convex hull s" + obtain S where "finite S" "cbox 0 (\Basis::'a) = convex hull S" using unit_cube_convex_hull by auto then show ?thesis - apply (rule_tac that[of "(\y. x - ?d + (2 * d) *\<^sub>R y)` s"]) - unfolding * and convex_hull_affinity - apply auto - done + by (rule_tac that[of "(\y. x - ?d + (2 * d) *\<^sub>R y)` S"]) (auto simp: convex_hull_affinity *) qed subsection%unimportant\Representation of any interval as a finite convex hull\ @@ -6877,13 +6444,13 @@ next case False then have *: "\a b. a = m i * b \ b = a / m i" - by (auto simp add: field_simps) + by (auto simp: field_simps) from False have "min (m i * (a \ i)) (m i * (b \ i)) = (if 0 < m i then m i * (a \ i) else m i * (b \ i))" "max (m i * (a \ i)) (m i * (b \ i)) = (if 0 < m i then m i * (b \ i) else m i * (a \ i))" using a_le_b by (auto simp: min_def max_def mult_le_cancel_left) with False show ?thesis using a_le_b - unfolding * by (auto simp add: le_divide_eq divide_le_eq ac_simps) + unfolding * by (auto simp: le_divide_eq divide_le_eq ac_simps) qed qed qed simp @@ -6895,7 +6462,7 @@ lemma cbox_translation: "cbox (c + a) (c + b) = image (\x. c + x) (cbox a b)" using image_affinity_cbox [of 1 c a b] using box_ne_empty [of "a+c" "b+c"] box_ne_empty [of a b] - by (auto simp add: inner_left_distrib add.commute) + by (auto simp: inner_left_distrib add.commute) lemma cbox_image_unit_interval: fixes a :: "'a::euclidean_space" @@ -6909,18 +6476,18 @@ lemma closed_interval_as_convex_hull: fixes a :: "'a::euclidean_space" - obtains s where "finite s" "cbox a b = convex hull s" + obtains S where "finite S" "cbox a b = convex hull S" proof (cases "cbox a b = {}") case True with convex_hull_empty that show ?thesis by blast next case False - obtain s::"'a set" where "finite s" and eq: "cbox 0 One = convex hull s" + obtain S::"'a set" where "finite S" and eq: "cbox 0 One = convex hull S" by (blast intro: unit_cube_convex_hull) have lin: "linear (\x. \k\Basis. ((b \ k - a \ k) * (x \ k)) *\<^sub>R k)" by (rule linear_compose_sum) (auto simp: algebra_simps linearI) - have "finite ((+) a ` (\x. \k\Basis. ((b \ k - a \ k) * (x \ k)) *\<^sub>R k) ` s)" - by (rule finite_imageI \finite s\)+ + have "finite ((+) a ` (\x. \k\Basis. ((b \ k - a \ k) * (x \ k)) *\<^sub>R k) ` S)" + by (rule finite_imageI \finite S\)+ then show ?thesis apply (rule that) apply (simp add: convex_hull_translation convex_hull_linear_image [OF lin, symmetric]) @@ -6932,31 +6499,23 @@ subsection%unimportant \Bounded convex function on open set is continuous\ lemma convex_on_bounded_continuous: - fixes s :: "('a::real_normed_vector) set" - assumes "open s" - and "convex_on s f" - and "\x\s. \f x\ \ b" - shows "continuous_on s f" + fixes S :: "('a::real_normed_vector) set" + assumes "open S" + and "convex_on S f" + and "\x\S. \f x\ \ b" + shows "continuous_on S f" apply (rule continuous_at_imp_continuous_on) unfolding continuous_at_real_range proof (rule,rule,rule) fix x and e :: real - assume "x \ s" "e > 0" + assume "x \ S" "e > 0" define B where "B = \b\ + 1" - have B: "0 < B" "\x. x\s \ \f x\ \ B" - unfolding B_def - defer - apply (drule assms(3)[rule_format]) - apply auto - done - obtain k where "k > 0" and k: "cball x k \ s" - using assms(1)[unfolded open_contains_cball, THEN bspec[where x=x]] - using \x\s\ by auto + then have B: "0 < B""\x. x\S \ \f x\ \ B" + using assms(3) by auto + obtain k where "k > 0" and k: "cball x k \ S" + using \x \ S\ assms(1) open_contains_cball_eq by blast show "\d>0. \x'. norm (x' - x) < d \ \f x' - f x\ < e" - apply (rule_tac x="min (k / 2) (e / (2 * B) * k)" in exI) - apply rule - defer - proof (rule, rule) + proof (intro exI conjI allI impI) fix y assume as: "norm (y - x) < min (k / 2) (e / (2 * B) * k)" show "\f y - f x\ < e" @@ -6965,79 +6524,63 @@ define t where "t = k / norm (y - x)" have "2 < t" "0k>0\ - by (auto simp add:field_simps) - have "y \ s" - apply (rule k[unfolded subset_eq,rule_format]) + by (auto simp:field_simps) + have "y \ S" + apply (rule k[THEN subsetD]) unfolding mem_cball dist_norm apply (rule order_trans[of _ "2 * norm (x - y)"]) using as - by (auto simp add: field_simps norm_minus_commute) + by (auto simp: field_simps norm_minus_commute) { define w where "w = x + t *\<^sub>R (y - x)" - have "w \ s" - unfolding w_def - apply (rule k[unfolded subset_eq,rule_format]) - unfolding mem_cball dist_norm - unfolding t_def - using \k>0\ - apply auto - done + have "w \ S" + using \k>0\ by (auto simp: dist_norm t_def w_def k[THEN subsetD]) 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) + by (auto simp: algebra_simps) also have "\ = 0" - using \t > 0\ by (auto simp add:field_simps) + using \t > 0\ by (auto simp:field_simps) 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) - have "2 * B < e * t" + by (auto simp: algebra_simps) + have 2: "2 * B < e * t" unfolding t_def using \0 < e\ \0 < k\ \B > 0\ and as and False - by (auto simp add:field_simps) - then have "(f w - f x) / t < e" - using B(2)[OF \w\s\] and B(2)[OF \x\s\] - using \t > 0\ by (auto simp add:field_simps) - then have th1: "f y - f x < e" - apply - - apply (rule le_less_trans) - defer - apply assumption + by (auto simp:field_simps) + have "f y - f x \ (f w - f x) / t" using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w] - using \0 < t\ \2 < t\ and \x \ s\ \w \ s\ - by (auto simp add:field_simps) + using \0 < t\ \2 < t\ and \x \ S\ \w \ S\ + by (auto simp:field_simps) + also have "... < e" + using B(2)[OF \w\S\] and B(2)[OF \x\S\] 2 \t > 0\ by (auto simp: field_simps) + finally have th1: "f y - f x < e" . } moreover { define w where "w = x - t *\<^sub>R (y - x)" - have "w \ s" - unfolding w_def - apply (rule k[unfolded subset_eq,rule_format]) - unfolding mem_cball dist_norm - unfolding t_def - using \k > 0\ - apply auto - done + have "w \ S" + using \k > 0\ by (auto simp: dist_norm t_def w_def k[THEN subsetD]) 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) + by (auto simp: algebra_simps) also have "\ = x" - using \t > 0\ by (auto simp add:field_simps) + using \t > 0\ by (auto simp:field_simps) 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) + by (auto simp: algebra_simps) 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) + by (auto simp:field_simps) then have *: "(f w - f y) / t < e" - using B(2)[OF \w\s\] and B(2)[OF \y\s\] + using B(2)[OF \w\S\] and B(2)[OF \y\S\] using \t > 0\ - by (auto simp add:field_simps) + by (auto simp:field_simps) have "f x \ 1 / (1 + t) * f w + (t / (1 + t)) * f y" using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w] - using \0 < t\ \2 < t\ and \y \ s\ \w \ s\ - by (auto simp add:field_simps) + using \0 < t\ \2 < t\ and \y \ S\ \w \ S\ + by (auto simp:field_simps) also have "\ = (f w + t * f y) / (1 + t)" - using \t > 0\ by (auto simp add: divide_simps) + using \t > 0\ by (auto simp: divide_simps) also have "\ < e + f y" - using \t > 0\ * \e > 0\ by (auto simp add: field_simps) + using \t > 0\ * \e > 0\ by (auto simp: field_simps) finally have "f x - f y < e" by auto } ultimately show ?thesis by auto @@ -7062,13 +6605,13 @@ have *: "x - (2 *\<^sub>R x - y) = y - x" by (simp add: scaleR_2) have z: "z \ cball x e" - using y unfolding z_def mem_cball dist_norm * by (auto simp add: norm_minus_commute) + using y unfolding z_def mem_cball dist_norm * by (auto simp: norm_minus_commute) have "(1 / 2) *\<^sub>R y + (1 / 2) *\<^sub>R z = x" - unfolding z_def by (auto simp add: algebra_simps) + unfolding z_def by (auto simp: algebra_simps) then show "\f y\ \ b + 2 * \f x\" using assms(1)[unfolded convex_on_def,rule_format, OF y z, of "1/2" "1/2"] using assms(2)[rule_format,OF y] assms(2)[rule_format,OF z] - by (auto simp add:field_simps) + by (auto simp:field_simps) next case False fix y @@ -7121,12 +6664,7 @@ by (simp add: dist_norm abs_le_iff algebra_simps) show "cball x d \ convex hull c" unfolding 2 - apply clarsimp - apply (simp only: dist_norm) - apply (subst inner_diff_left [symmetric]) - apply simp - apply (erule (1) order_trans [OF Basis_le_norm]) - done + by (clarsimp simp: dist_norm) (metis inner_commute inner_diff_right norm_bound_Basis_le) have e': "e = (\(i::'a)\Basis. d)" by (simp add: d_def DIM_positive) show "convex hull c \ cball x e" @@ -7136,51 +6674,31 @@ apply (rule order_trans [OF L2_set_le_sum]) apply (rule zero_le_dist) unfolding e' - apply (rule sum_mono) - apply simp + apply (rule sum_mono, simp) done qed define k where "k = Max (f ` c)" have "convex_on (convex hull c) f" apply(rule convex_on_subset[OF assms(2)]) - apply(rule subset_trans[OF _ e(1)]) - apply(rule c1) + apply(rule subset_trans[OF c1 e(1)]) done then have k: "\y\convex hull c. f y \ k" - apply (rule_tac convex_on_convex_hull_bound) - apply assumption - unfolding k_def - apply (rule, rule Max_ge) - using c(1) - apply auto - done - have "d \ e" - unfolding d_def - apply (rule mult_imp_div_pos_le) - using \e > 0\ - unfolding mult_le_cancel_left1 - apply (auto simp: real_of_nat_ge_one_iff Suc_le_eq DIM_positive) - done + apply (rule_tac convex_on_convex_hull_bound, assumption) + by (simp add: k_def c) + have "e \ e * real DIM('a)" + using e(2) real_of_nat_ge_one_iff by auto + then have "d \ e" + by (simp add: d_def divide_simps) then have dsube: "cball x d \ cball x e" by (rule subset_cball) have conv: "convex_on (cball x d) f" - apply (rule convex_on_subset) - apply (rule convex_on_subset[OF assms(2)]) - apply (rule e(1)) - apply (rule dsube) - done + using \convex_on (convex hull c) f\ c2 convex_on_subset by blast then have "\y\cball x d. \f y\ \ k + 2 * \f x\" - apply (rule convex_bounds_lemma) - apply (rule ballI) - apply (rule k [rule_format]) - apply (erule rev_subsetD) - apply (rule c2) - done + by (rule convex_bounds_lemma) (use c2 k in blast) then have "continuous_on (ball x d) f" apply (rule_tac convex_on_bounded_continuous) apply (rule open_ball, rule convex_on_subset[OF conv]) - apply (rule ball_subset_cball) - apply force + apply (rule ball_subset_cball, force) done then show "continuous (at x) f" unfolding continuous_on_eq_continuous_at[OF open_ball]