# HG changeset patch # User himmelma # Date 1243496805 -7200 # Node ID f6427bc40421ff0f383731a248da6d71577c73b5 # Parent 1ba01cdd9a9a068afd6631cea224c769530cc9c8 Added Convex_Euclidean_Space.thy diff -r 1ba01cdd9a9a -r f6427bc40421 src/HOL/Library/Convex_Euclidean_Space.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Convex_Euclidean_Space.thy Thu May 28 09:46:45 2009 +0200 @@ -0,0 +1,3243 @@ +(* Title: Convex + ID: $Id: + Author: Robert Himmelmann, TU Muenchen*) + +header {* Convex sets, functions and related things. *} + +theory Convex + imports Topology_Euclidean_Space +begin + + +(* ------------------------------------------------------------------------- *) +(* To be moved elsewhere *) +(* ------------------------------------------------------------------------- *) + +declare vector_add_ldistrib[simp] vector_ssub_ldistrib[simp] vector_smult_assoc[simp] vector_smult_rneg[simp] +declare vector_sadd_rdistrib[simp] vector_sub_rdistrib[simp] +declare dot_ladd[simp] dot_radd[simp] dot_lsub[simp] dot_rsub[simp] +declare dot_lmult[simp] dot_rmult[simp] dot_lneg[simp] dot_rneg[simp] +declare UNIV_1[simp] + +term "(x::real^'n \ real) 0" + +lemma dim1in[intro]:"Suc 0 \ {1::nat .. CARD(1)}" by auto + +lemmas vector_component_simps = vector_minus_component vector_smult_component vector_add_component vector_less_eq_def Cart_lambda_beta dest_vec1_def basis_component vector_uminus_component + +lemmas continuous_intros = continuous_add continuous_vmul continuous_cmul continuous_const continuous_sub continuous_at_id continuous_within_id + +lemmas continuous_on_intros = continuous_on_add continuous_on_const continuous_on_id continuous_on_compose continuous_on_cmul continuous_on_neg continuous_on_sub + uniformly_continuous_on_add uniformly_continuous_on_const uniformly_continuous_on_id uniformly_continuous_on_compose uniformly_continuous_on_cmul uniformly_continuous_on_neg uniformly_continuous_on_sub + +lemma dest_vec1_simps[simp]: fixes a::"real^1" + shows "a$1 = 0 \ a = 0" (*"a \ 1 \ dest_vec1 a \ 1" "0 \ a \ 0 \ dest_vec1 a"*) + "a \ b \ dest_vec1 a \ dest_vec1 b" "dest_vec1 (1::real^1) = 1" + by(auto simp add:vector_component_simps all_1 Cart_eq) + +definition + "is_interval_new (s::((real^'n::finite) set)) \ + (\a\s. \b\s. \x. (\i. ((a$i \ x$i \ x$i \ b$i) \ (b$i \ x$i \ x$i \ a$i))) \ x \ s)" + +lemma is_interval_interval_new: "is_interval_new {a .. b}" (is ?th1) "is_interval_new {a<..x y z::real. x < y \ y < z \ x < z" by auto + show ?th1 ?th2 unfolding is_interval_new_def mem_interval Ball_def atLeastAtMost_iff + by(meson real_le_trans le_less_trans less_le_trans *)+ qed + +lemma nequals0I:"x\A \ A \ {}" by auto + +lemma norm_not_0:"(x::real^'n::finite)\0 \ norm x \ 0" by auto + +lemma vector_unminus_smult[simp]: "(-1::real) *s x = -x" unfolding pth_3[symmetric] by simp + +lemma setsum_delta_notmem: assumes "x\s" + shows "setsum (\y. if (y = x) then P x else Q y) s = setsum Q s" + "setsum (\y. if (x = y) then P x else Q y) s = setsum Q s" + "setsum (\y. if (y = x) then P y else Q y) s = setsum Q s" + "setsum (\y. if (x = y) then P y else Q y) s = setsum Q s" + apply(rule_tac [!] setsum_cong2) using assms by auto + +lemma setsum_diff1'':assumes "finite A" "a \ A" + shows "setsum f (A - {a}) = setsum f A - (f a::'a::ring)" unfolding setsum_diff1'[OF assms] by auto + +lemma setsum_delta'': fixes s::"(real^'n) set" assumes "finite s" + shows "(\x\s. (if y = x then f x else 0) *s x) = (if y\s then (f y) *s y else 0)" +proof- + have *:"\x y. (if y = x then f x else (0::real)) *s x = (if x=y then (f x) *s x else 0)" by auto + show ?thesis unfolding * using setsum_delta[OF assms, of y "\x. f x *s x"] by auto +qed + +lemma not_disjointI:"x\A \ x\B \ A \ B \ {}" by blast + +lemma if_smult:"(if P then x else (y::real)) *s v = (if P then x *s v else y *s v)" by auto + +lemma ex_bij_betw_nat_finite_1: + assumes "finite M" + shows "\h. bij_betw h {1 .. card M} M" +proof- + obtain h where h:"bij_betw h {0..i. i - 1) ` {1..card M} = {0..i. i - Suc 0) {Suc 0..card M}" unfolding inj_on_def by auto + hence "inj_on ?h {1..card M}" apply(rule_tac comp_inj_on) unfolding * using h[unfolded bij_betw_def] by auto + ultimately show ?thesis apply(rule_tac x="h \ (\i. i - 1)" in exI) unfolding o_def and bij_betw_def by auto +qed + +lemma finite_subset_image: + assumes "B \ f ` A" "finite B" + shows "\C\A. finite C \ B = f ` C" +proof- from assms(1) have "\x\B. \y\A. x = f y" by auto + then obtain c where "\x\B. c x \ A \ x = f (c x)" + using bchoice[of B "\x y. y\A \ x = f y"] by auto + thus ?thesis apply(rule_tac x="c ` B" in exI) using assms(2) by auto qed + +lemma inj_on_image_eq_iff: assumes "inj_on f (A \ B)" + shows "f ` A = f ` B \ A = B" + using assms by(blast dest: inj_onD) + + +lemma mem_interval_1: fixes x :: "real^1" shows + "(x \ {a .. b} \ dest_vec1 a \ dest_vec1 x \ dest_vec1 x \ dest_vec1 b)" + "(x \ {a<.. dest_vec1 a < dest_vec1 x \ dest_vec1 x < dest_vec1 b)" +by(simp_all add: Cart_eq vector_less_def vector_less_eq_def dest_vec1_def all_1) + +lemma image_smult_interval:"(\x. m *s (x::real^'n::finite)) ` {a..b} = + (if {a..b} = {} then {} else if 0 \ m then {m *s a..m *s b} else {m *s b..m *s a})" + using image_affinity_interval[of m 0 a b] by auto + +lemma dest_vec1_inverval: + "dest_vec1 ` {a .. b} = {dest_vec1 a .. dest_vec1 b}" + "dest_vec1 ` {a<.. b} = {dest_vec1 a<.. dest_vec1 b}" + "dest_vec1 ` {a ..x. dest_vec1 (f x)) S" + using dest_vec1_sum[OF assms] by auto + +lemma dist_triangle_eq:"dist x z = dist x y + dist y z \ norm (x - y) *s (y - z) = norm (y - z) *s (x - y)" +proof- have *:"x - y + (y - z) = x - z" by auto + show ?thesis unfolding dist_def norm_triangle_eq[of "x - y" "y - z", unfolded *] + by(auto simp add:norm_minus_commute) qed + +lemma norm_eqI:"x = y \ norm x = norm y" by auto +lemma norm_minus_eqI:"(x::real^'n::finite) = - y \ norm x = norm y" by auto + +lemma Min_grI: assumes "finite A" "A \ {}" "\a\A. x < a" shows "x < Min A" + unfolding Min_gr_iff[OF assms(1,2)] using assms(3) by auto + +lemma dimindex_ge_1:"CARD(_::finite) \ 1" + using one_le_card_finite by auto + +lemma real_dimindex_ge_1:"real (CARD('n::finite)) \ 1" + by(metis dimindex_ge_1 linorder_not_less real_eq_of_nat real_le_trans real_of_nat_1 real_of_nat_le_iff) + +lemma real_dimindex_gt_0:"real (CARD('n::finite)) > 0" apply(rule less_le_trans[OF _ real_dimindex_ge_1]) by auto + +subsection {* Affine set and affine hull.*} + +definition "affine s \ (\x\s. \y\s. \u v::real. u + v = 1 \ (u *s x + v *s y) \ s)" + +lemma affine_alt: "affine s \ (\x\s. \y\s. \u::real. (1 - u) *s x + u *s y \ s)" +proof- have *:"\u v ::real. u + v = 1 \ v = 1 - u" by auto + { fix x y assume "x\s" "y\s" + hence "(\u v::real. u + v = 1 \ u *s x + v *s y \ s) \ (\u::real. (1 - u) *s x + u *s y \ s)" apply auto + apply(erule_tac[!] x="1 - u" in allE) unfolding * by auto } + thus ?thesis unfolding affine_def by auto qed + +lemma affine_empty[intro]: "affine {}" + unfolding affine_def by auto + +lemma affine_sing[intro]: "affine {x}" + unfolding affine_alt by (auto simp add: vector_sadd_rdistrib[THEN sym]) + +lemma affine_UNIV[intro]: "affine UNIV" + unfolding affine_def by auto + +lemma affine_Inter: "(\s\f. affine s) \ affine (\ f)" + unfolding affine_def by auto + +lemma affine_Int: "affine s \ affine t \ affine (s \ t)" + unfolding affine_def by auto + +lemma affine_affine_hull: "affine(affine hull s)" + unfolding hull_def using affine_Inter[of "{t \ affine. s \ t}"] + unfolding mem_def by auto + +lemma affine_hull_eq[simp]: "(affine hull s = s) \ affine s" +proof- + { fix f assume "f \ affine" + hence "affine (\f)" using affine_Inter[of f] unfolding subset_eq mem_def by auto } + thus ?thesis using hull_eq[unfolded mem_def, of affine s] by auto +qed + +lemma setsum_restrict_set'': assumes "finite A" + shows "setsum f {x \ A. P x} = (\x\A. if P x then f x else 0)" + unfolding mem_def[of _ P, symmetric] unfolding setsum_restrict_set'[OF assms] .. + +subsection {* Some explicit formulations (from Lars Schewe). *} + +lemma affine: fixes V::"(real^'n) set" + shows "affine V \ (\s u. finite s \ s \ {} \ s \ V \ setsum u s = 1 \ (setsum (\x. (u x) *s x)) s \ V)" +unfolding affine_def apply rule apply(rule, rule, rule) apply(erule conjE)+ +defer apply(rule, rule, rule, rule, rule) proof- + fix x y u v assume as:"x \ V" "y \ V" "u + v = (1::real)" + "\s u. finite s \ s \ {} \ s \ V \ setsum u s = 1 \ (\x\s. u x *s x) \ V" + thus "u *s x + v *s 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) + by(auto simp add: vector_sadd_rdistrib[THEN sym]) +next + fix s u assume as:"\x\V. \y\V. \u v. u + v = 1 \ u *s x + v *s y \ V" + "finite s" "s \ {}" "s \ V" "setsum u s = (1::real)" + def n \ "card s" + have "card s = 0 \ card s = 1 \ card s = 2 \ card s > 2" by auto + thus "(\x\s. u x *s x) \ V" proof(auto simp only: disjE) + assume "card s = 2" hence "card s = Suc (Suc 0)" by auto + then obtain a b where "s = {a, b}" unfolding card_Suc_eq by auto + thus ?thesis using as(1)[THEN bspec[where x=a], THEN bspec[where x=b]] using as(4,5) + by(auto simp add: setsum_clauses(2)) + next assume "card s > 2" thus ?thesis using as and n_def proof(induct n arbitrary: u s) + case (Suc n) fix s::"(real^'n) set" and u::"real^'n\ real" + assume IA:"\u s. \2 < card s; \x\V. \y\V. \u v. u + v = 1 \ u *s x + v *s y \ V; finite s; + s \ {}; s \ V; setsum u s = 1; n \ card s \ \ (\x\s. u x *s x) \ V" and + as:"Suc n \ card s" "2 < card s" "\x\V. \y\V. \u v. u + v = 1 \ u *s x + v *s y \ V" + "finite s" "s \ {}" "s \ V" "setsum u s = 1" + have "\x\s. u x \ 1" proof(rule_tac ccontr) + assume " \ (\x\s. u x \ 1)" hence "setsum u s = real_of_nat (card s)" unfolding card_eq_setsum by auto + thus False using as(7) and `card s > 2` by (metis Numeral1_eq1_nat less_0_number_of less_int_code(15) + less_nat_number_of not_less_iff_gr_or_eq of_nat_1 of_nat_eq_iff pos2 rel_simps(4)) 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) by auto + have *:"s = insert x (s - {x})" "finite (s - {x})" using `x\s` and as(4) by auto + have **:"setsum u (s - {x}) = 1 - u x" + using setsum_clauses(2)[OF *(2), of u x, unfolded *(1)[THEN sym] as(7)] by auto + have ***:"inverse (1 - u x) * setsum u (s - {x}) = 1" unfolding ** using `u x \ 1` by auto + have "(\xa\s - {x}. (inverse (1 - u x) * u xa) *s xa) \ V" proof(cases "card (s - {x}) > 2") + case True hence "s - {x} \ {}" "card (s - {x}) = n" unfolding c and as(1)[symmetric] proof(rule_tac ccontr) + assume "\ s - {x} \ {}" hence "card (s - {x}) = 0" unfolding card_0_eq[OF *(2)] by simp + thus False using True by auto qed auto + thus ?thesis apply(rule_tac IA[of "s - {x}" "\y. (inverse (1 - u x) * u y)"]) + unfolding setsum_right_distrib[THEN sym] using as and *** and True by auto + next case False hence "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 + thus ?thesis using as(3)[THEN bspec[where x=a], THEN bspec[where x=b]] + using *** *(2) and `s \ V` unfolding setsum_right_distrib by(auto simp add: setsum_clauses(2)) qed + thus "(\x\s. u x *s x) \ V" unfolding vector_smult_assoc[THEN sym] and setsum_cmul + apply(subst *) unfolding setsum_clauses(2)[OF *(2)] + using as(3)[THEN bspec[where x=x], THEN bspec[where x="(inverse (1 - u x)) *s (\xa\s - {x}. u xa *s xa)"], + THEN spec[where x="u x"], THEN spec[where x="1 - u x"]] and rev_subsetD[OF `x\s` `s\V`] and `u x \ 1` by auto + qed auto + next assume "card s = 1" then obtain a where "s={a}" by(auto simp add: card_Suc_eq) + thus ?thesis using as(4,5) by simp + qed(insert `s\{}` `finite s`, auto) +qed + +lemma affine_hull_explicit: + "affine hull p = {y. \s u. finite s \ s \ {} \ s \ p \ setsum u s = 1 \ setsum (\v. (u v) *s v) s = y}" + apply(rule hull_unique) apply(subst subset_eq) prefer 3 apply rule unfolding mem_Collect_eq and mem_def[of _ affine] + apply (erule exE)+ apply(erule conjE)+ prefer 2 apply rule proof- + fix x assume "x\p" thus "\s u. finite s \ s \ {} \ s \ p \ setsum u s = 1 \ (\v\s. u v *s v) = x" + apply(rule_tac x="{x}" in exI, rule_tac x="\x. 1" in exI) by auto +next + fix t x s u assume as:"p \ t" "affine t" "finite s" "s \ {}" "s \ p" "setsum u s = 1" "(\v\s. u v *s v) = x" + thus "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 \ setsum u s = 1 \ (\v\s. u v *s v) = y}" unfolding affine_def + apply(rule,rule,rule,rule,rule) unfolding mem_Collect_eq proof- + fix u v ::real assume uv:"u + v = 1" + fix x assume "\s u. finite s \ s \ {} \ s \ p \ setsum u s = 1 \ (\v\s. u v *s v) = x" + then obtain sx ux where x:"finite sx" "sx \ {}" "sx \ p" "setsum ux sx = 1" "(\v\sx. ux v *s v) = x" by auto + fix y assume "\s u. finite s \ s \ {} \ s \ p \ setsum u s = 1 \ (\v\s. u v *s v) = y" + then obtain sy uy where y:"finite sy" "sy \ {}" "sy \ p" "setsum uy sy = 1" "(\v\sy. uy v *s v) = y" by auto + have xy:"finite (sx \ sy)" using x(1) y(1) by auto + have **:"(sx \ sy) \ sx = sx" "(sx \ sy) \ sy = sy" by auto + show "\s ua. finite s \ s \ {} \ s \ p \ setsum ua s = 1 \ (\v\s. ua v *s v) = u *s x + v *s 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 vector_sadd_rdistrib setsum_addf if_smult vector_smult_lzero ** setsum_restrict_set[OF xy, THEN sym] + unfolding vector_smult_assoc[THEN sym] setsum_cmul and setsum_right_distrib[THEN sym] + unfolding x y using x(1-3) y(1-3) uv by simp qed qed + +lemma affine_hull_finite: + assumes "finite s" + shows "affine hull s = {y. \u. setsum u s = 1 \ setsum (\v. u v *s v) s = y}" + unfolding affine_hull_explicit and expand_set_eq and mem_Collect_eq apply (rule,rule) + apply(erule exE)+ apply(erule conjE)+ defer apply(erule exE) apply(erule conjE) proof- + fix x u assume "setsum u s = 1" "(\v\s. u v *s v) = x" + thus "\sa u. finite sa \ \ (\x. (x \ sa) = (x \ {})) \ sa \ s \ setsum u sa = 1 \ (\v\sa. u v *s v) = x" + apply(rule_tac x=s in exI, rule_tac x=u in exI) using assms by auto +next + fix x t u assume "t \ s" hence *:"s \ t = t" by auto + assume "finite t" "\ (\x. (x \ t) = (x \ {}))" "setsum u t = 1" "(\v\t. u v *s v) = x" + thus "\u. setsum u s = 1 \ (\v\s. u v *s v) = x" apply(rule_tac x="\x. if x\t then u x else 0" in exI) + unfolding if_smult vector_smult_lzero and setsum_restrict_set[OF assms, THEN sym] and * by auto qed + +subsection {* Stepping theorems and hence small special cases. *} + +lemma affine_hull_empty[simp]: "affine hull {} = {}" + apply(rule hull_unique) unfolding mem_def by auto + +lemma affine_hull_finite_step: + shows "(\u::real^'n=>real. setsum u {} = w \ setsum (\x. u x *s x) {} = y) \ w = 0 \ y = 0" (is ?th1) + "finite s \ (\u. setsum u (insert a s) = w \ setsum (\x. u x *s x) (insert a s) = y) \ + (\v u. setsum u s = w - v \ setsum (\x. u x *s x) s = y - v *s a)" (is "?as \ (?lhs = ?rhs)") +proof- + show ?th1 by simp + assume ?as + { assume ?lhs + then obtain u where u:"setsum u (insert a s) = w \ (\x\insert a s. u x *s x) = y" by auto + have ?rhs proof(cases "a\s") + case True hence *:"insert a s = s" by auto + show ?thesis using u[unfolded *] apply(rule_tac x=0 in exI) by auto + next + case False thus ?thesis apply(rule_tac x="u a" in exI) using u and `?as` by auto + qed } moreover + { assume ?rhs + then obtain v u where vu:"setsum u s = w - v" "(\x\s. u x *s x) = y - v *s a" by auto + have *:"\x M. (if x = a then v else M) *s x = (if x = a then v *s x else M *s x)" by auto + have ?lhs proof(cases "a\s") + case True thus ?thesis + apply(rule_tac x="\x. (if x=a then v else 0) + u x" in exI) + unfolding setsum_clauses(2)[OF `?as`] apply simp + unfolding vector_sadd_rdistrib and setsum_addf + unfolding vu and * and pth_4(1) + by (auto simp add: setsum_delta[OF `?as`]) + next + case False + hence **:"\x. x \ s \ u x = (if x = a then v else u x)" + "\x. x \ s \ u x *s x = (if x = a then v *s x else u x *s x)" by auto + from False show ?thesis + apply(rule_tac x="\x. if x=a then v else u x" in exI) + unfolding setsum_clauses(2)[OF `?as`] and * using vu + using setsum_cong2[of s "\x. u x *s x" "\x. if x = a then v *s x else u x *s x", OF **(2)] + using setsum_cong2[of s u "\x. if x = a then v else u x", OF **(1)] by auto + qed } + ultimately show "?lhs = ?rhs" by blast +qed + +lemma affine_hull_2: "affine hull {a,b::real^'n} = {u *s a + v *s b| u v. (u + v = 1)}" (is "?lhs = ?rhs") +proof- + have *:"\x y z. z = x - y \ y + z = (x::real)" + "\x y z. z = x - y \ y + z = (x::real^'n)" by auto + have "?lhs = {y. \u. setsum u {a, b} = 1 \ (\v\{a, b}. u v *s v) = y}" + using affine_hull_finite[of "{a,b}"] by auto + also have "\ = {y. \v u. u b = 1 - v \ u b *s b = y - v *s a}" + by(simp add: affine_hull_finite_step(2)[of "{b}" a]) + also have "\ = ?rhs" unfolding * by auto + finally show ?thesis by auto +qed + +lemma affine_hull_3: "affine hull {a,b,c::real^'n} = { u *s a + v *s b + w *s c| u v w. u + v + w = 1}" (is "?lhs = ?rhs") +proof- + have *:"\x y z. z = x - y \ y + z = (x::real)" + "\x y z. z = x - y \ y + z = (x::real^'n)" by auto + 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) by(auto intro!: exI) +qed + +subsection {* 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)+ proof- + fix x t u assume as:"finite t" "t \ {}" "t \ insert a s" "setsum u t = 1" "(\v\t. u v *s v) = x" + have "(\x. x - a) ` (t - {a}) \ {x - a |x. x \ s}" using as(3) by auto + thus "\v. x = a + v \ (\S u. finite S \ S \ {x - a |x. x \ s} \ (\v\S. u v *s v) = v)" + apply(rule_tac x="x - a" in exI) apply rule defer apply(rule_tac x="(\x. x - a) ` (t - {a})" in exI) + apply(rule_tac x="\x. u (x + a)" in exI) using as(1) + apply(simp add: setsum_reindex[unfolded inj_on_def] setsum_subtractf setsum_diff1 setsum_vmul[THEN sym]) + unfolding as by simp_all 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 *s v) = y" unfolding span_explicit by auto + def f \ "(\x. x + a) ` t" + have f:"finite f" "f \ s" "(\v\f. u (v - a) *s (v - a)) = y - a" unfolding f_def using obt + by(auto simp add: setsum_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 \ setsum u sa = 1 \ (\v\sa. u v *s v) = y" + apply(rule_tac x="insert a f" in exI) + apply(rule_tac x="\x. if x=a then 1 - setsum (\x. u (x - a)) f else u (x - a)" in exI) + using assms and f unfolding setsum_clauses(2)[OF f(1)] and if_smult + unfolding setsum_cases[OF f(1), of "{a}", unfolded singleton_iff] and * + by (auto simp add: setsum_subtractf setsum_vmul field_simps) 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 + +subsection {* Convexity. *} + +definition "convex (s::(real^'n) set) \ + (\x\s. \y\s. \u\0. \v\0. (u + v = 1) \ (u *s x + v *s y) \ s)" + +lemma convex_alt: "convex s \ (\x\s. \y\s. \u. 0 \ u \ u \ 1 \ ((1 - u) *s x + u *s y) \ s)" +proof- have *:"\u v::real. u + v = 1 \ u = 1 - v" by auto + show ?thesis unfolding convex_def apply auto + apply(erule_tac x=x in ballE) apply(erule_tac x=y in ballE) apply(erule_tac x="1 - u" in allE) + by (auto simp add: *) qed + +lemma mem_convex: + assumes "convex s" "a \ s" "b \ s" "0 \ u" "u \ 1" + shows "((1 - u) *s a + u *s b) \ s" + using assms unfolding convex_alt by auto + +lemma convex_empty[intro]: "convex {}" + unfolding convex_def by simp + +lemma convex_singleton[intro]: "convex {a}" + unfolding convex_def by (auto simp add:vector_sadd_rdistrib[THEN sym]) + +lemma convex_UNIV[intro]: "convex UNIV" + unfolding convex_def by auto + +lemma convex_Inter: "(\s\f. convex s) ==> convex(\ f)" + unfolding convex_def by auto + +lemma convex_Int: "convex s \ convex t \ convex (s \ t)" + unfolding convex_def by auto + +lemma convex_halfspace_le: "convex {x. a \ x \ b}" + unfolding convex_def apply auto + unfolding dot_radd dot_rmult by (metis real_convex_bound_le) + +lemma convex_halfspace_ge: "convex {x. a \ x \ b}" +proof- have *:"{x. a \ x \ b} = {x. -a \ x \ -b}" by auto + show ?thesis apply(unfold *) using convex_halfspace_le[of "-a" "-b"] by auto qed + +lemma convex_hyperplane: "convex {x. a \ x = b}" +proof- + have *:"{x. a \ x = b} = {x. a \ x \ b} \ {x. a \ x \ b}" by auto + show ?thesis unfolding * apply(rule convex_Int) + using convex_halfspace_le convex_halfspace_ge by auto +qed + +lemma convex_halfspace_lt: "convex {x. a \ x < b}" + unfolding convex_def by(auto simp add: real_convex_bound_lt dot_radd dot_rmult) + +lemma convex_halfspace_gt: "convex {x. a \ x > b}" + using convex_halfspace_lt[of "-a" "-b"] by(auto simp add: dot_lneg neg_less_iff_less) + +lemma convex_positive_orthant: "convex {x::real^'n. (\i. 0 \ x$i)}" + unfolding convex_def apply auto apply(erule_tac x=i in allE)+ + apply(rule add_nonneg_nonneg) by(auto simp add: mult_nonneg_nonneg) + +subsection {* Explicit expressions for convexity in terms of arbitrary sums. *} + +lemma convex: "convex s \ + (\(k::nat) u x. (\i. 1\i \ i\k \ 0 \ u i \ x i \s) \ (setsum u {1..k} = 1) + \ setsum (\i. u i *s x i) {1..k} \ s)" + unfolding convex_def apply rule apply(rule allI)+ defer apply(rule ballI)+ apply(rule allI)+ proof(rule,rule,rule,rule) + fix x y u v assume as:"\(k::nat) u x. (\i. 1 \ i \ i \ k \ 0 \ u i \ x i \ s) \ setsum u {1..k} = 1 \ (\i = 1..k. u i *s x i) \ s" + "x \ s" "y \ s" "0 \ u" "0 \ v" "u + v = (1::real)" + show "u *s x + v *s y \ s" using as(1)[THEN spec[where x=2], THEN spec[where x="\n. if n=1 then u else v"], THEN spec[where x="\n. if n=1 then x else y"]] and as(2-) + by (auto simp add: setsum_head_Suc) +next + fix k u x assume as:"\x\s. \y\s. \u\0. \v\0. u + v = 1 \ u *s x + v *s y \ s" + show "(\i::nat. 1 \ i \ i \ k \ 0 \ u i \ x i \ s) \ setsum u {1..k} = 1 \ (\i = 1..k. u i *s x i) \ s" apply(rule,erule conjE) proof(induct k arbitrary: u) + case (Suc k) show ?case proof(cases "u (Suc k) = 1") + case True hence "(\i = Suc 0..k. u i *s x i) = 0" apply(rule_tac setsum_0') apply(rule ccontr) unfolding ball_simps apply(erule bexE) proof- + fix i assume i:"i \ {Suc 0..k}" "u i *s x i \ 0" + hence ui:"u i \ 0" by auto + hence "setsum (\k. if k=i then u i else 0) {1 .. k} \ setsum u {1 .. k}" apply(rule_tac setsum_mono) using Suc(2) by auto + hence "setsum u {1 .. k} \ u i" using i(1) by(auto simp add: setsum_delta) + hence "setsum u {1 .. k} > 0" using ui apply(rule_tac less_le_trans[of _ "u i"]) using Suc(2)[THEN spec[where x=i]] and i(1) by auto + thus False using Suc(3) unfolding setsum_cl_ivl_Suc and True by simp qed + thus ?thesis unfolding setsum_cl_ivl_Suc using True and Suc(2) by auto + next + have *:"setsum u {1..k} = 1 - u (Suc k)" using Suc(3)[unfolded setsum_cl_ivl_Suc] by auto + have **:"u (Suc k) \ 1" apply(rule ccontr) unfolding not_le using Suc(3) using setsum_nonneg[of "{1..k}" u] using Suc(2) by auto + have ***:"\i k. (u i / (1 - u (Suc k))) *s x i = (inverse (1 - u (Suc k))) *s (u i *s x i)" unfolding real_divide_def by auto + case False hence nn:"1 - u (Suc k) \ 0" by auto + have "(\i = 1..k. (u i / (1 - u (Suc k))) *s x i) \ s" apply(rule Suc(1)) unfolding setsum_divide_distrib[THEN sym] and * + apply(rule_tac allI) apply(rule,rule) apply(rule divide_nonneg_pos) using nn Suc(2) ** by auto + hence "(1 - u (Suc k)) *s (\i = 1..k. (u i / (1 - u (Suc k))) *s x i) + u (Suc k) *s x (Suc k) \ s" + apply(rule as[THEN bspec, THEN bspec, THEN spec, THEN mp, THEN spec, THEN mp, THEN mp]) using Suc(2)[THEN spec[where x="Suc k"]] and ** by auto + thus ?thesis unfolding setsum_cl_ivl_Suc and *** and setsum_cmul using nn by auto qed qed auto qed + + +lemma convex_explicit: "convex (s::(real^'n) set) \ + (\t u. finite t \ t \ s \ (\x\t. 0 \ u x) \ setsum u t = 1 \ setsum (\x. u x *s x) t \ s)" + unfolding convex_def apply(rule,rule,rule) apply(subst imp_conjL,rule) defer apply(rule,rule,rule,rule,rule,rule,rule) proof- + fix x y u v assume as:"\t u. finite t \ t \ s \ (\x\t. 0 \ u x) \ setsum u t = 1 \ (\x\t. u x *s x) \ s" "x \ s" "y \ s" "0 \ u" "0 \ v" "u + v = (1::real)" + show "u *s x + v *s y \ s" proof(cases "x=y") + case True show ?thesis unfolding True and vector_sadd_rdistrib[THEN sym] using as(3,6) by auto next + case False thus ?thesis using as(1)[THEN spec[where x="{x,y}"], THEN spec[where x="\z. if z=x then u else v"]] and as(2-) by auto qed +next + fix t u assume asm:"\x\s. \y\s. \u\0. \v\0. u + v = 1 \ u *s x + v *s y \ s" "finite (t::(real^'n) set)" + (*"finite t" "t \ s" "\x\t. (0::real) \ u x" "setsum u t = 1"*) + from this(2) have "\u. t \ s \ (\x\t. 0 \ u x) \ setsum u t = 1 \ (\x\t. u x *s x) \ s" apply(induct_tac t rule:finite_induct) + prefer 3 apply (rule,rule) apply(erule conjE)+ proof- + fix x f u assume ind:"\u. f \ s \ (\x\f. 0 \ u x) \ setsum u f = 1 \ (\x\f. u x *s x) \ s" + assume as:"finite f" "x \ f" "insert x f \ s" "\x\insert x f. 0 \ u x" "setsum u (insert x f) = (1::real)" + show "(\x\insert x f. u x *s x) \ s" proof(cases "u x = 1") + case True hence "setsum (\x. u x *s x) f = 0" apply(rule_tac setsum_0') apply(rule ccontr) unfolding ball_simps apply(erule bexE) proof- + fix y assume y:"y \ f" "u y *s y \ 0" + hence uy:"u y \ 0" by auto + hence "setsum (\k. if k=y then u y else 0) f \ setsum u f" apply(rule_tac setsum_mono) using as(4) by auto + hence "setsum u f \ u y" using y(1) and as(1) by(auto simp add: setsum_delta) + hence "setsum u f > 0" using uy apply(rule_tac less_le_trans[of _ "u y"]) using as(4) and y(1) by auto + thus False using as(2,5) unfolding setsum_clauses(2)[OF as(1)] and True by auto qed + thus ?thesis unfolding setsum_clauses(2)[OF as(1)] using as(2,3) unfolding True by auto + next + have *:"setsum u f = setsum u (insert x f) - u x" using as(2) unfolding setsum_clauses(2)[OF as(1)] by auto + have **:"u x \ 1" apply(rule ccontr) unfolding not_le using as(5)[unfolded setsum_clauses(2)[OF as(1)]] and as(2) + using setsum_nonneg[of f u] and as(4) by auto + case False hence "inverse (1 - u x) *s (\x\f. u x *s x) \ s" unfolding setsum_cmul[THEN sym] and vector_smult_assoc + apply(rule_tac ind[THEN spec, THEN mp]) apply rule defer apply rule apply rule apply(rule mult_nonneg_nonneg) + unfolding setsum_right_distrib[THEN sym] and * using as and ** by auto + hence "u x *s x + (1 - u x) *s ((inverse (1 - u x)) *s setsum (\x. u x *s x) f) \s" + apply(rule_tac asm(1)[THEN bspec, THEN bspec, THEN spec, THEN mp, THEN spec, THEN mp, THEN mp]) using as and ** False by auto + thus ?thesis unfolding setsum_clauses(2)[OF as(1)] using as(2) and False by auto qed + qed auto thus "t \ s \ (\x\t. 0 \ u x) \ setsum u t = 1 \ (\x\t. u x *s x) \ s" by auto +qed + +lemma convex_finite: assumes "finite s" + shows "convex s \ (\u. (\x\s. 0 \ u x) \ setsum u s = 1 + \ setsum (\x. u x *s x) s \ s)" + unfolding convex_explicit apply(rule, rule, rule) defer apply(rule,rule,rule)apply(erule conjE)+ proof- + fix t u assume as:"\u. (\x\s. 0 \ u x) \ setsum u s = 1 \ (\x\s. u x *s x) \ s" " finite t" "t \ s" "\x\t. 0 \ u x" "setsum u t = (1::real)" + have *:"s \ t = t" using as(3) by auto + show "(\x\t. u x *s x) \ s" using as(1)[THEN spec[where x="\x. if x\t then u x else 0"]] + unfolding if_smult and setsum_cases[OF assms] and * using as(2-) by auto +qed (erule_tac x=s in allE, erule_tac x=u in allE, auto) + +subsection {* Cones. *} + +definition "cone (s::(real^'n) set) \ (\x\s. \c\0. (c *s x) \ s)" + +lemma cone_empty[intro, simp]: "cone {}" + unfolding cone_def by auto + +lemma cone_univ[intro, simp]: "cone UNIV" + unfolding cone_def by auto + +lemma cone_Inter[intro]: "(\s\f. cone s) \ cone(\ f)" + unfolding cone_def by auto + +subsection {* Conic hull. *} + +lemma cone_cone_hull: "cone (cone hull s)" + unfolding hull_def using cone_Inter[of "{t \ conic. s \ t}"] + by (auto simp add: mem_def) + +lemma cone_hull_eq: "(cone hull s = s) \ cone s" + apply(rule hull_eq[unfolded mem_def]) + using cone_Inter unfolding subset_eq by (auto simp add: mem_def) + +subsection {* Affine dependence and consequential theorems (from Lars Schewe). *} + +definition "affine_dependent (s::(real^'n) set) \ (\x\s. x \ (affine hull (s - {x})))" + +lemma affine_dependent_explicit: + "affine_dependent p \ + (\s u. finite s \ s \ p \ setsum u s = 0 \ + (\v\s. u v \ 0) \ setsum (\v. u v *s v) s = 0)" + unfolding 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) +proof- + fix x s u assume as:"x \ p" "finite s" "s \ {}" "s \ p - {x}" "setsum u s = 1" "(\v\s. u v *s v) = x" + have "x\s" using as(1,4) by auto + show "\s u. finite s \ s \ p \ setsum u s = 0 \ (\v\s. u v \ 0) \ (\v\s. u v *s 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 setsum_clauses(2)[OF as(2)] and setsum_delta_notmem[OF `x\s`] and as using as by auto +next + fix s u v assume as:"finite s" "s \ p" "setsum u s = 0" "(\v\s. u v *s v) = 0" "v \ s" "u v \ 0" + have "s \ {v}" using as(3,6) by auto + thus "\x\p. \s u. finite s \ s \ {} \ s \ p - {x} \ setsum u s = 1 \ (\v\s. u v *s v) = x" + apply(rule_tac x=v in bexI, rule_tac x="s - {v}" in exI, rule_tac x="\x. - (1 / u v) * u x" in exI) + unfolding vector_smult_assoc[THEN sym] and setsum_cmul unfolding setsum_right_distrib[THEN sym] and setsum_diff1''[OF as(1,5)] using as by auto +qed + +lemma affine_dependent_explicit_finite: + assumes "finite (s::(real^'n) set)" + shows "affine_dependent s \ (\u. setsum u s = 0 \ (\v\s. u v \ 0) \ setsum (\v. u v *s v) s = 0)" + (is "?lhs = ?rhs") +proof + have *:"\vt u v. (if vt then u v else 0) *s v = (if vt then (u v) *s v else (0::real^'n))" by auto + assume ?lhs + then obtain t u v where "finite t" "t \ s" "setsum u t = 0" "v\t" "u v \ 0" "(\v\t. u v *s v) = 0" + unfolding affine_dependent_explicit by auto + thus ?rhs apply(rule_tac x="\x. if x\t then u x else 0" in exI) + apply auto unfolding * and setsum_restrict_set[OF assms, THEN sym] + unfolding Int_absorb2[OF `t\s`, unfolded Int_commute] by auto +next + assume ?rhs + then obtain u v where "setsum u s = 0" "v\s" "u v \ 0" "(\v\s. u v *s v) = 0" by auto + thus ?lhs unfolding affine_dependent_explicit using assms by auto +qed + +subsection {* A general lemma. *} + +lemma convex_connected: + assumes "convex s" shows "connected s" +proof- + { fix e1 e2 assume as:"open e1" "open e2" "e1 \ e2 \ s = {}" "s \ e1 \ e2" + assume "e1 \ s \ {}" "e2 \ s \ {}" + then obtain x1 x2 where x1:"x1\e1" "x1\s" and x2:"x2\e2" "x2\s" by auto + hence n:"norm (x1 - x2) > 0" unfolding zero_less_norm_iff using as(3) by auto + + { fix x e::real assume as:"0 \ x" "x \ 1" "0 < e" + { fix y have *:"(1 - x) *s x1 + x *s x2 - ((1 - y) *s x1 + y *s x2) = (y - x) *s x1 - (y - x) *s x2" + by(simp add: ring_simps vector_sadd_rdistrib vector_sub_rdistrib) + assume "\y - x\ < e / norm (x1 - x2)" + hence "norm ((1 - x) *s x1 + x *s x2 - ((1 - y) *s x1 + y *s x2)) < e" + unfolding * and vector_ssub_ldistrib[THEN sym] and norm_mul + unfolding less_divide_eq using n by auto } + hence "\d>0. \y. \y - x\ < d \ norm ((1 - x) *s x1 + x *s x2 - ((1 - y) *s x1 + y *s x2)) < e" + apply(rule_tac x="e / norm (x1 - x2)" in exI) using as + apply auto unfolding zero_less_divide_iff using n by simp } note * = this + + have "\x\0. x \ 1 \ (1 - x) *s x1 + x *s x2 \ e1 \ (1 - x) *s x1 + x *s x2 \ e2" + apply(rule connected_real_lemma) apply (simp add: `x1\e1` `x2\e2` dist_sym)+ + using * apply(simp add: dist_def) + using as(1,2)[unfolded open_def] apply simp + using as(1,2)[unfolded open_def] apply simp + using assms[unfolded convex_alt, THEN bspec[where x=x1], THEN bspec[where x=x2]] using x1 x2 + using as(3) by auto + then obtain x where "x\0" "x\1" "(1 - x) *s x1 + x *s x2 \ e1" "(1 - x) *s x1 + x *s x2 \ e2" by auto + hence False using as(4) + using assms[unfolded convex_alt, THEN bspec[where x=x1], THEN bspec[where x=x2]] + using x1(2) x2(2) by auto } + thus ?thesis unfolding connected_def by auto +qed + +subsection {* One rather trivial consequence. *} + +lemma connected_UNIV: "connected UNIV" + by(simp add: convex_connected convex_UNIV) + +subsection {* Convex functions into the reals. *} + +definition "convex_on s (f::real^'n \ real) = + (\x\s. \y\s. \u\0. \v\0. u + v = 1 \ f (u *s x + v *s y) \ u * f x + v * f y)" + +lemma convex_on_subset: "convex_on t f \ s \ t \ convex_on s f" + unfolding convex_on_def by auto + +lemma convex_add: + assumes "convex_on s f" "convex_on s g" + shows "convex_on s (\x. f x + g x)" +proof- + { fix x y assume "x\s" "y\s" moreover + fix u v ::real assume "0 \ u" "0 \ v" "u + v = 1" + ultimately have "f (u *s x + v *s y) + g (u *s x + v *s y) \ (u * f x + v * f y) + (u * g x + v * g y)" + using assms(1)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x=u]] + using assms(2)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x=u]] + apply - apply(rule add_mono) by auto + hence "f (u *s x + v *s y) + g (u *s x + v *s y) \ u * (f x + g x) + v * (f y + g y)" by (simp add: ring_simps) } + thus ?thesis unfolding convex_on_def by auto +qed + +lemma convex_cmul: + assumes "0 \ (c::real)" "convex_on s f" + shows "convex_on s (\x. c * f x)" +proof- + have *:"\u c fx v fy ::real. u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" by (simp add: ring_simps) + show ?thesis using assms(2) and mult_mono1[OF _ assms(1)] unfolding convex_on_def and * by auto +qed + +lemma convex_lower: + assumes "convex_on s f" "x\s" "y \ s" "0 \ u" "0 \ v" "u + v = 1" + shows "f (u *s x + v *s y) \ max (f x) (f y)" +proof- + let ?m = "max (f x) (f y)" + have "u * f x + v * f y \ u * max (f x) (f y) + v * max (f x) (f y)" apply(rule add_mono) + using assms(4,5) by(auto simp add: mult_mono1) + also have "\ = max (f x) (f y)" using assms(6) unfolding distrib[THEN sym] by auto + finally show ?thesis using assms(1)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x=u]] + using assms(2-6) by auto +qed + +lemma convex_local_global_minimum: + assumes "0 s" "\y\ball x e. f x \ f y" + shows "\y\s. f x \ f y" +proof(rule ccontr) + have "x\s" using assms(1,3) by auto + assume "\ (\y\s. f x \ f y)" + then obtain y where "y\s" and y:"f x > f y" by auto + hence xy:"0 < dist x y" by (auto simp add: dist_nz[THEN sym]) + + then obtain u where "0 < u" "u \ 1" and u:"u < e / dist x y" + using real_lbound_gt_zero[of 1 "e / dist x y"] using xy `e>0` and divide_pos_pos[of e "dist x y"] by auto + hence "f ((1-u) *s x + u *s y) \ (1-u) * f x + u * f y" using `x\s` `y\s` + using assms(2)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x="1-u"]] by auto + moreover + have *:"x - ((1 - u) *s x + u *s y) = u *s (x - y)" by (simp add: vector_ssub_ldistrib vector_sub_rdistrib) + have "(1 - u) *s x + u *s y \ ball x e" unfolding mem_ball dist_def unfolding * and norm_mul and abs_of_pos[OF `0 f ((1 - u) *s x + u *s y)" using assms(4) by auto + ultimately show False using mult_strict_left_mono[OF y `u>0`] unfolding left_diff_distrib by auto +qed + +lemma convex_distance: "convex_on s (\x. dist a x)" +proof(auto simp add: convex_on_def dist_def) + fix x y assume "x\s" "y\s" + fix u v ::real assume "0 \ u" "0 \ v" "u + v = 1" + have "a = u *s a + v *s a" unfolding vector_sadd_rdistrib[THEN sym] and `u+v=1` by simp + hence *:"a - (u *s x + v *s y) = (u *s (a - x)) + (v *s (a - y))" by auto + show "norm (a - (u *s x + v *s y)) \ u * norm (a - x) + v * norm (a - y)" + unfolding * using norm_triangle_ineq[of "u *s (a - x)" "v *s (a - y)"] unfolding norm_mul + using `0 \ u` `0 \ v` by auto +qed + +subsection {* Arithmetic operations on sets preserve convexity. *} + +lemma convex_scaling: "convex s \ convex ((\x. c *s x) ` s)" + unfolding convex_def and image_iff apply auto + apply (rule_tac x="u *s x+v *s y" in bexI) by (auto simp add: field_simps) + +lemma convex_negations: "convex s \ convex ((\x. -x)` s)" + unfolding convex_def and image_iff apply auto + apply (rule_tac x="u *s x+v *s y" in bexI) by auto + +lemma convex_sums: + assumes "convex s" "convex t" + shows "convex {x + y| x y. x \ s \ y \ t}" +proof(auto simp add: convex_def image_iff) + fix xa xb ya yb assume xy:"xa\s" "xb\s" "ya\t" "yb\t" + fix u v ::real assume uv:"0 \ u" "0 \ v" "u + v = 1" + show "\x y. u *s xa + u *s ya + (v *s xb + v *s yb) = x + y \ x \ s \ y \ t" + apply(rule_tac x="u *s xa + v *s xb" in exI) apply(rule_tac x="u *s ya + v *s yb" in exI) + using assms(1)[unfolded convex_def, THEN bspec[where x=xa], THEN bspec[where x=xb]] + using assms(2)[unfolded convex_def, THEN bspec[where x=ya], THEN bspec[where x=yb]] + using uv xy by auto +qed + +lemma convex_differences: + assumes "convex s" "convex t" + shows "convex {x - y| x y. x \ s \ y \ t}" +proof- + have "{x - y| x y. x \ s \ y \ t} = {x + y |x y. x \ s \ y \ uminus ` t}" unfolding image_iff apply auto + apply(rule_tac x=xa in exI) apply(rule_tac x="-y" in exI) apply simp + apply(rule_tac x=xa in exI) apply(rule_tac x=xb in exI) by simp + thus ?thesis using convex_sums[OF assms(1) convex_negations[OF assms(2)]] by auto +qed + +lemma convex_translation: assumes "convex s" shows "convex ((\x. a + x) ` s)" +proof- have "{a + y |y. y \ s} = (\x. a + x) ` s" by auto + thus ?thesis using convex_sums[OF convex_singleton[of a] assms] by auto qed + +lemma convex_affinity: assumes "convex (s::(real^'n) set)" shows "convex ((\x. a + c *s x) ` s)" +proof- have "(\x. a + c *s x) ` s = op + a ` op *s c ` s" by auto + thus ?thesis using convex_translation[OF convex_scaling[OF assms], of a c] by auto qed + +lemma convex_linear_image: assumes c:"convex s" and l:"linear f" shows "convex(f ` s)" +proof(auto simp add: convex_def) + fix x y assume xy:"x \ s" "y \ s" + fix u v ::real assume uv:"0 \ u" "0 \ v" "u + v = 1" + show "u *s f x + v *s f y \ f ` s" unfolding image_iff + apply(rule_tac x="u *s x + v *s y" in bexI) + unfolding linear_add[OF l] linear_cmul[OF l] + using c[unfolded convex_def] xy uv by auto +qed + +subsection {* Balls, being convex, are connected. *} + +lemma convex_ball: "convex (ball x e)" +proof(auto simp add: convex_def) + fix y z assume yz:"dist x y < e" "dist x z < e" + fix u v ::real assume uv:"0 \ u" "0 \ v" "u + v = 1" + have "dist x (u *s y + v *s z) \ u * dist x y + v * dist x z" using uv yz + using convex_distance[of "ball x e" x, unfolded convex_on_def, THEN bspec[where x=y], THEN bspec[where x=z]] by auto + thus "dist x (u *s y + v *s z) < e" using real_convex_bound_lt[OF yz uv] by auto +qed + +lemma convex_cball: "convex(cball x e)" +proof(auto simp add: convex_def Ball_def mem_cball) + fix y z assume yz:"dist x y \ e" "dist x z \ e" + fix u v ::real assume uv:" 0 \ u" "0 \ v" "u + v = 1" + have "dist x (u *s y + v *s z) \ u * dist x y + v * dist x z" using uv yz + using convex_distance[of "cball x e" x, unfolded convex_on_def, THEN bspec[where x=y], THEN bspec[where x=z]] by auto + thus "dist x (u *s y + v *s z) \ e" using real_convex_bound_le[OF yz uv] by auto +qed + +lemma connected_ball: "connected(ball x e)" + using convex_connected convex_ball by auto + +lemma connected_cball: "connected(cball x e)" + using convex_connected convex_cball by auto + +subsection {* Convex hull. *} + +lemma convex_convex_hull: "convex(convex hull s)" + unfolding hull_def using convex_Inter[of "{t\convex. s\t}"] + unfolding mem_def by auto + +lemma convex_hull_eq: "(convex hull s = s) \ convex s" apply(rule hull_eq[unfolded mem_def]) + using convex_Inter[unfolded Ball_def mem_def] by auto + +lemma bounded_convex_hull: assumes "bounded s" shows "bounded(convex hull s)" +proof- from assms obtain B where B:"\x\s. norm x \ B" unfolding bounded_def by auto + show ?thesis apply(rule bounded_subset[OF bounded_cball, of _ 0 B]) + unfolding subset_hull[unfolded mem_def, of convex, OF convex_cball] + unfolding subset_eq mem_cball dist_def using B by auto qed + +lemma finite_imp_bounded_convex_hull: + "finite s \ bounded(convex hull s)" + using bounded_convex_hull finite_imp_bounded by auto + +subsection {* Stepping theorems for convex hulls of finite sets. *} + +lemma convex_hull_empty[simp]: "convex hull {} = {}" + apply(rule hull_unique) unfolding mem_def by auto + +lemma convex_hull_singleton[simp]: "convex hull {a} = {a}" + apply(rule hull_unique) unfolding mem_def by auto + +lemma convex_hull_insert: + assumes "s \ {}" + shows "convex hull (insert a s) = {x. \u\0. \v\0. \b. (u + v = 1) \ + b \ (convex hull s) \ (x = u *s a + v *s b)}" (is "?xyz = ?hull") + apply(rule,rule hull_minimal,rule) unfolding mem_def[of _ convex] and insert_iff prefer 3 apply rule proof- + fix x assume x:"x = a \ x \ s" + thus "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] by auto +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 *s a + v *s 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) by auto + thus "x\ convex hull insert a s" unfolding obt(5) using convex_convex_hull[of "insert a s", unfolded convex_def] + apply(erule_tac x=a in ballE) apply(erule_tac x=b in ballE) apply(erule_tac x=u in allE) using obt by auto +next + show "convex ?hull" unfolding convex_def apply(rule,rule,rule,rule,rule,rule,rule) proof- + 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 *s a + v1 *s 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 *s a + v2 *s b2" by auto + have *:"\x s1 s2. x - s1 *s x - s2 *s x = ((1::real) - (s1 + s2)) *s x" by auto + have "\b \ convex hull s. u *s x + v *s y = (u * u1) *s a + (v * u2) *s a + (b - (u * u1) *s b - (v * u2) *s b)" + proof(cases "u * v1 + v * v2 = 0") + have *:"\x s1 s2. x - s1 *s x - s2 *s x = ((1::real) - (s1 + s2)) *s x" by auto + case True hence **:"u * v1 = 0" "v * v2 = 0" apply- apply(rule_tac [!] ccontr) + using mult_nonneg_nonneg[OF `u\0` `v1\0`] mult_nonneg_nonneg[OF `v\0` `v2\0`] by auto + hence "u * u1 + v * u2 = 1" using as(3) obt1(3) obt2(3) by auto + thus ?thesis unfolding obt1(5) obt2(5) * using assms hull_subset[of s convex] by(auto simp add: **) + next + 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) + also have "\ = u * (v1 + u1 - u1) + v * (v2 + u2 - u2)" using as(3) obt1(3) obt2(3) by (auto simp add: field_simps) + also have "\ = u * v1 + v * v2" by simp finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" by auto + case False have "0 \ u * v1 + v * v2" "0 \ u * v1" "0 \ u * v1 + v * v2" "0 \ v * v2" apply - + apply(rule add_nonneg_nonneg) prefer 4 apply(rule add_nonneg_nonneg) apply(rule_tac [!] mult_nonneg_nonneg) + using as(1,2) obt1(1,2) obt2(1,2) by auto + thus ?thesis unfolding obt1(5) obt2(5) unfolding * and ** using False + apply(rule_tac x="((u * v1) / (u * v1 + v * v2)) *s b1 + ((v * v2) / (u * v1 + v * v2)) *s b2" in bexI) defer + apply(rule convex_convex_hull[of s, unfolded convex_def, rule_format]) using obt1(4) obt2(4) + unfolding add_divide_distrib[THEN sym] and real_0_le_divide_iff by auto + qed note * = this + have u1:"u1 \ 1" apply(rule ccontr) unfolding obt1(3)[THEN sym] and not_le using obt1(2) by auto + have u2:"u2 \ 1" apply(rule ccontr) unfolding obt2(3)[THEN sym] 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) by auto + also have "\ \ 1" unfolding mult.add_right[THEN sym] and as(3) using u1 u2 by auto + finally + show "u *s x + v *s 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) * by(auto intro!: mult_nonneg_nonneg add_nonneg_nonneg simp add:field_simps) + qed +qed + + +subsection {* Explicit expression for convex hull. *} + +lemma convex_hull_indexed: + "convex hull s = {y. \k u x. (\i\{1::nat .. k}. 0 \ u i \ x i \ s) \ + (setsum u {1..k} = 1) \ + (setsum (\i. u i *s x i) {1..k} = y)}" (is "?xyz = ?hull") + apply(rule hull_unique) unfolding mem_def[of _ convex] apply(rule) defer + apply(subst convex_def) apply(rule,rule,rule,rule,rule,rule,rule) +proof- + fix x assume "x\s" + thus "x \ ?hull" unfolding mem_Collect_eq apply(rule_tac x=1 in exI, rule_tac x="\x. 1" in exI) by auto +next + fix t assume as:"s \ t" "convex t" + show "?hull \ t" apply(rule) unfolding mem_Collect_eq apply(erule exE | erule conjE)+ proof- + fix x k u y assume assm:"\i\{1::nat..k}. 0 \ u i \ y i \ s" "setsum u {1..k} = 1" "(\i = 1..k. u i *s y i) = x" + show "x\t" unfolding assm(3)[THEN sym] apply(rule as(2)[unfolded convex, rule_format]) + using assm(1,2) as(1) by auto qed +next + fix x y u v assume uv:"0\u" "0\v" "u+v=(1::real)" and xy:"x\?hull" "y\?hull" + from xy obtain k1 u1 x1 where x:"\i\{1::nat..k1}. 0\u1 i \ x1 i \ s" "setsum u1 {Suc 0..k1} = 1" "(\i = Suc 0..k1. u1 i *s x1 i) = x" by auto + from xy obtain k2 u2 x2 where y:"\i\{1::nat..k2}. 0\u2 i \ x2 i \ s" "setsum u2 {Suc 0..k2} = 1" "(\i = Suc 0..k2. u2 i *s x2 i) = y" by auto + have *:"\P x1 x2 s1 s2 i.(if P i then s1 else s2) *s (if P i then x1 else x2) = (if P i then s1 *s x1 else s2 *s 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) by(auto simp add: not_le) + have inj:"inj_on (\i. i + k1) {1..k2}" unfolding inj_on_def by auto + show "u *s x + v *s y \ ?hull" apply(rule) + apply(rule_tac x="k1 + k2" in exI, 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 setsum_cases[OF finite_atLeastAtMost[of 1 "k1 + k2"]] and setsum_reindex[OF inj] and o_def + unfolding vector_smult_assoc[THEN sym] setsum_cmul setsum_right_distrib[THEN sym] 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 thus ?thesis using mult_nonneg_nonneg[of u "u1 i"] and uv(1) x(1)[THEN bspec[where x=i]] by auto + next def j \ "i - k1" + case False with i have "j \ {1..k2}" unfolding j_def by auto + thus ?thesis unfolding j_def[symmetric] using False + using mult_nonneg_nonneg[of v "u2 j"] and uv(2) y(1)[THEN bspec[where x=j]] by auto qed + qed(auto simp add: not_le x(2,3) y(2,3) uv(3)) +qed + +lemma convex_hull_finite: + assumes "finite (s::(real^'n)set)" + shows "convex hull s = {y. \u. (\x\s. 0 \ u x) \ + setsum u s = 1 \ setsum (\x. u x *s x) s = y}" (is "?HULL = ?set") +proof(rule hull_unique, auto simp add: mem_def[of _ convex] convex_def[of ?set]) + fix x assume "x\s" thus " \u. (\x\s. 0 \ u x) \ setsum u s = 1 \ (\x\s. u x *s x) = x" + apply(rule_tac x="\y. if x=y then 1 else 0" in exI) apply auto + unfolding setsum_delta'[OF assms] and setsum_delta''[OF assms] by auto +next + fix u v ::real assume uv:"0 \ u" "0 \ v" "u + v = 1" + fix ux assume ux:"\x\s. 0 \ ux x" "setsum ux s = (1::real)" + fix uy assume uy:"\x\s. 0 \ uy x" "setsum uy s = (1::real)" + { fix x assume "x\s" + hence "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, metis add_nonneg_nonneg mult_nonneg_nonneg uv(1) uv(2)) } + moreover have "(\x\s. u * ux x + v * uy x) = 1" + unfolding setsum_addf and setsum_right_distrib[THEN sym] and ux(2) uy(2) using uv(3) by auto + moreover have "(\x\s. (u * ux x + v * uy x) *s x) = u *s (\x\s. ux x *s x) + v *s (\x\s. uy x *s x)" + unfolding vector_sadd_rdistrib and setsum_addf and vector_smult_assoc[THEN sym] and setsum_cmul by auto + ultimately show "\uc. (\x\s. 0 \ uc x) \ setsum uc s = 1 \ (\x\s. uc x *s x) = u *s (\x\s. ux x *s x) + v *s (\x\s. uy x *s x)" + apply(rule_tac x="\x. u * ux x + v * uy x" in exI) by auto +next + fix t assume t:"s \ t" "convex t" + fix u assume u:"\x\s. 0 \ u x" "setsum u s = (1::real)" + thus "(\x\s. u x *s 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 + +subsection {* Another formulation from Lars Schewe. *} + +lemma convex_hull_explicit: + "convex hull p = {y. \s u. finite s \ s \ p \ + (\x\s. 0 \ u x) \ setsum u s = 1 \ setsum (\v. u v *s v) s = y}" (is "?lhs = ?rhs") +proof- + { fix x assume "x\?lhs" + then obtain k u y where obt:"\i\{1::nat..k}. 0 \ u i \ y i \ p" "setsum u {1..k} = 1" "(\i = 1..k. u i *s y i) = x" + unfolding convex_hull_indexed by auto + + have fin:"finite {1..k}" by auto + have fin':"\v. finite {i \ {1..k}. y i = v}" by auto + { fix j assume "j\{1..k}" + hence "y j \ p" "0 \ setsum u {i. Suc 0 \ i \ i \ k \ y i = y j}" + using obt(1)[THEN bspec[where x=j]] and obt(2) apply simp + apply(rule setsum_nonneg) using obt(1) by auto } + moreover + have "(\v\y ` {1..k}. setsum u {i \ {1..k}. y i = v}) = 1" + unfolding setsum_image_gen[OF fin, THEN sym] using obt(2) by auto + moreover have "(\v\y ` {1..k}. setsum u {i \ {1..k}. y i = v} *s v) = x" + using setsum_image_gen[OF fin, of "\i. u i *s y i" y, THEN sym] + unfolding setsum_vmul[OF fin'] using obt(3) by auto + ultimately have "\s u. finite s \ s \ p \ (\x\s. 0 \ u x) \ setsum u s = 1 \ (\v\s. u v *s v) = x" + apply(rule_tac x="y ` {1..k}" in exI) + apply(rule_tac x="\v. setsum u {i\{1..k}. y i = v}" in exI) by auto + hence "x\?rhs" by auto } + moreover + { fix y assume "y\?rhs" + then obtain s u where obt:"finite s" "s \ p" "\x\s. 0 \ u x" "setsum u s = 1" "(\v\s. u v *s v) = y" by auto + + 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}" + hence "f i \ s" apply(subst f(2)[THEN sym]) by auto + hence "0 \ u (f i)" "f i \ p" using obt(2,3) 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}"] by auto + hence "{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) by auto + hence "card {x. Suc 0 \ x \ x \ card s \ f x = y} = 1" by auto + hence "(\x\{x \ {1..card s}. f x = y}. u (f x)) = u y" "(\x\{x \ {1..card s}. f x = y}. u (f x) *s f x) = u y *s y" by auto } + + hence "(\x = 1..card s. u (f x)) = 1" "(\i = 1..card s. u (f i) *s f i) = y" + unfolding setsum_image_gen[OF *(1), of "\x. u (f x) *s f x" f] and setsum_image_gen[OF *(1), of "\x. u (f x)" f] + unfolding f using setsum_cong2[of s "\y. (\x\{x \ {1..card s}. f x = y}. u (f x) *s f x)" "\v. u v *s v"] + using setsum_cong2 [of 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) \ setsum u {1..k} = 1 \ (\i::nat = 1..k. u i *s x i) = y" + apply(rule_tac x="card s" in exI) apply(rule_tac x="u \ f" in exI) apply(rule_tac x=f in exI) by fastsimp + hence "y \ ?lhs" unfolding convex_hull_indexed by auto } + ultimately show ?thesis unfolding expand_set_eq by blast +qed + +subsection {* A stepping theorem for that expansion. *} + +lemma convex_hull_finite_step: + assumes "finite (s::(real^'n) set)" + shows "(\u. (\x\insert a s. 0 \ u x) \ setsum u (insert a s) = w \ setsum (\x. u x *s x) (insert a s) = y) + \ (\v\0. \u. (\x\s. 0 \ u x) \ setsum u s = w - v \ setsum (\x. u x *s x) s = y - v *s a)" (is "?lhs = ?rhs") +proof(rule, case_tac[!] "a\s") + assume "a\s" hence *:"insert a s = s" by auto + assume ?lhs thus ?rhs unfolding * apply(rule_tac x=0 in exI) by auto +next + assume ?lhs then obtain u where u:"\x\insert a s. 0 \ u x" "setsum u (insert a s) = w" "(\x\insert a s. u x *s x) = y" by auto + assume "a\s" thus ?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 setsum_clauses(2)[OF assms]] and `a\s` by auto +next + assume "a\s" hence *:"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" "setsum u s = w - v" "(\x\s. u x *s x) = y - v *s a" by auto + show ?lhs apply(rule_tac x="\x. (if a = x then v else 0) + u x" in exI) unfolding vector_sadd_rdistrib and setsum_addf and setsum_delta''[OF fin] and setsum_delta'[OF fin] + unfolding setsum_clauses(2)[OF assms] using uv and uv(2)[THEN bspec[where x=a]] and `a\s` by auto +next + assume ?rhs then obtain v u where uv:"v\0" "\x\s. 0 \ u x" "setsum u s = w - v" "(\x\s. u x *s x) = y - v *s a" by auto + moreover assume "a\s" moreover have "(\x\s. if a = x then v else u x) = setsum u s" "(\x\s. (if a = x then v else u x) *s x) = (\x\s. u x *s x)" + apply(rule_tac setsum_cong2) defer apply(rule_tac setsum_cong2) using `a\s` by auto + ultimately show ?lhs apply(rule_tac x="\x. if a = x then v else u x" in exI) unfolding setsum_clauses(2)[OF assms] by auto +qed + +subsection {* Hence some special cases. *} + +lemma convex_hull_2: + "convex hull {a,b} = {u *s a + v *s b | u v. 0 \ u \ 0 \ v \ u + v = 1}" +proof- have *:"\u. (\x\{a, b}. 0 \ u x) \ 0 \ u a \ 0 \ u b" by auto have **:"finite {b}" by auto +show ?thesis apply(simp add: convex_hull_finite) 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) by simp qed + +lemma convex_hull_2_alt: "convex hull {a,b} = {a + u *s (b - a) | u. 0 \ u \ u \ 1}" + unfolding convex_hull_2 unfolding Collect_def +proof(rule ext) have *:"\x y ::real. x + y = 1 \ x = 1 - y" by auto + fix x show "(\v u. x = v *s a + u *s b \ 0 \ v \ 0 \ u \ v + u = 1) = (\u. x = a + u *s (b - a) \ 0 \ u \ u \ 1)" + unfolding * apply auto apply(rule_tac[!] x=u in exI) by auto qed + +lemma convex_hull_3: + "convex hull {a::real^'n,b,c} = { u *s a + v *s b + w *s c | u v w. 0 \ u \ 0 \ v \ 0 \ w \ u + v + w = 1}" +proof- + 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" + "\x y z ::real^'n. x + y + z = 1 \ x = 1 - y - z" by (auto simp add: ring_simps) + show ?thesis unfolding convex_hull_finite[OF fin(1)] and Collect_def and convex_hull_finite_step[OF fin(2)] and * + unfolding convex_hull_finite_step[OF fin(3)] apply(rule ext) apply 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) by simp qed + +lemma convex_hull_3_alt: + "convex hull {a,b,c} = {a + u *s (b - a) + v *s (c - a) | u v. 0 \ u \ 0 \ v \ u + v \ 1}" +proof- have *:"\x y z ::real. x + y + z = 1 \ x = 1 - y - z" by auto + show ?thesis unfolding convex_hull_3 apply (auto simp add: *) apply(rule_tac x=v in exI) apply(rule_tac x=w in exI) apply simp + apply(rule_tac x=u in exI) apply(rule_tac x=v in exI) by simp qed + +subsection {* Relations among closure notions and corresponding hulls. *} + +lemma subspace_imp_affine: "subspace s \ affine s" + unfolding subspace_def affine_def by auto + +lemma affine_imp_convex: "affine s \ convex s" + unfolding affine_def convex_def by auto + +lemma subspace_imp_convex: "subspace s \ convex s" + using subspace_imp_affine affine_imp_convex by auto + +lemma affine_hull_subset_span: "(affine hull s) \ (span s)" + unfolding span_def apply(rule hull_antimono) unfolding subset_eq Ball_def mem_def + using subspace_imp_affine by auto + +lemma convex_hull_subset_span: "(convex hull s) \ (span s)" + unfolding span_def apply(rule hull_antimono) unfolding subset_eq Ball_def mem_def + using subspace_imp_convex by auto + +lemma convex_hull_subset_affine_hull: "(convex hull s) \ (affine hull s)" + unfolding span_def apply(rule hull_antimono) unfolding subset_eq Ball_def mem_def + using affine_imp_convex by auto + +lemma affine_dependent_imp_dependent: "affine_dependent s \ dependent s" + unfolding affine_dependent_def dependent_def + using affine_hull_subset_span by auto + +lemma dependent_imp_affine_dependent: + assumes "dependent {x - a| x . x \ s}" "a \ s" + shows "affine_dependent (insert a s)" +proof- + from assms(1)[unfolded dependent_explicit] obtain S u v + where obt:"finite S" "S \ {x - a |x. x \ s}" "v\S" "u v \ 0" "(\v\S. u v *s v) = 0" by auto + def t \ "(\x. x + a) ` S" + + have inj:"inj_on (\x. x + a) S" unfolding inj_on_def by auto + have "0\S" using obt(2) assms(2) unfolding subset_eq by auto + have fin:"finite t" and "t\s" unfolding t_def using obt(1,2) by auto + + hence "finite (insert a t)" and "insert a t \ insert a s" by auto + moreover have *:"\P Q. (\x\t. (if x = a then P x else Q x)) = (\x\t. Q x)" + apply(rule setsum_cong2) using `a\s` `t\s` by auto + have "(\x\insert a t. if x = a then - (\x\t. u (x - a)) else u (x - a)) = 0" + unfolding setsum_clauses(2)[OF fin] using `a\s` `t\s` apply auto unfolding * 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 by auto + moreover have *:"\P Q. (\x\t. (if x = a then P x else Q x) *s x) = (\x\t. Q x *s x)" + apply(rule setsum_cong2) using `a\s` `t\s` by auto + have "(\x\t. u (x - a)) *s a = (\v\t. u (v - a) *s v)" + unfolding setsum_vmul[OF fin(1)] unfolding t_def and setsum_reindex[OF inj] and o_def + using obt(5) by (auto simp add: setsum_addf) + hence "(\v\insert a t. (if v = a then - (\x\t. u (x - a)) else u (v - a)) *s v) = 0" + unfolding setsum_clauses(2)[OF fin] using `a\s` `t\s` by (auto simp add: * vector_smult_lneg) + ultimately show ?thesis unfolding affine_dependent_explicit + apply(rule_tac x="insert a t" in exI) by auto +qed + +lemma convex_cone: + "convex s \ cone s \ (\x\s. \y\s. (x + y) \ s) \ (\x\s. \c\0. (c *s x) \ s)" (is "?lhs = ?rhs") +proof- + { fix x y assume "x\s" "y\s" and ?lhs + hence "2 *s x \s" "2 *s y \ s" unfolding cone_def by auto + hence "x + y \ s" using `?lhs`[unfolded convex_def, THEN conjunct1] + apply(erule_tac x="2*s x" in ballE) apply(erule_tac x="2*s y" in ballE) + apply(erule_tac x="1/2" in allE) apply simp apply(erule_tac x="1/2" in allE) by auto } + thus ?thesis unfolding convex_def cone_def by blast +qed + +lemma affine_dependent_biggerset: fixes s::"(real^'n::finite) set" + assumes "finite s" "card s \ CARD('n) + 2" + shows "affine_dependent s" +proof- + have "s\{}" using assms 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 by auto + also have "\ > CARD('n)" using assms(2) + unfolding card_Diff_singleton[OF assms(1) `a\s`] by auto + finally show ?thesis apply(subst insert_Diff[OF `a\s`, THEN sym]) + apply(rule dependent_imp_affine_dependent) + apply(rule dependent_biggerset) by auto qed + +lemma affine_dependent_biggerset_general: + assumes "finite (s::(real^'n::finite) set)" "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})" by auto + have **:"card {x - a |x. x \ s - {a}} = card (s - {a})" unfolding * + apply(rule card_image) unfolding inj_on_def by auto + have "dim {x - a |x. x \ s - {a}} \ dim s" + apply(rule subset_le_dim) unfolding subset_eq + using `a\s` by (auto simp add:span_superset span_sub) + also have "\ < dim s + 1" by auto + also have "\ \ card (s - {a})" using assms + using card_Diff_singleton[OF assms(1) `a\s`] by auto + finally show ?thesis apply(subst insert_Diff[OF `a\s`, THEN sym]) + apply(rule dependent_imp_affine_dependent) apply(rule dependent_biggerset_general) unfolding ** by auto qed + +subsection {* Caratheodory's theorem. *} + +lemma convex_hull_caratheodory: fixes p::"(real^'n::finite) set" + shows "convex hull p = {y. \s u. finite s \ s \ p \ card s \ CARD('n) + 1 \ + (\x\s. 0 \ u x) \ setsum u s = 1 \ setsum (\v. u v *s v) s = y}" + unfolding convex_hull_explicit expand_set_eq mem_Collect_eq +proof(rule,rule) + fix y let ?P = "\n. \s u. finite s \ card s = n \ s \ p \ (\x\s. 0 \ u x) \ setsum u s = 1 \ (\v\s. u v *s v) = y" + assume "\s u. finite s \ s \ p \ (\x\s. 0 \ u x) \ setsum u s = 1 \ (\v\s. u v *s v) = y" + then obtain N where "?P N" by auto + hence "\n\N. (\k ?P k) \ ?P n" apply(rule_tac ex_least_nat_le) by auto + then obtain n where "?P n" and smallest:"\k ?P k" by blast + then obtain s u where obt:"finite s" "card s = n" "s\p" "\x\s. 0 \ u x" "setsum u s = 1" "(\v\s. u v *s v) = y" by auto + + have "card s \ CARD('n) + 1" proof(rule ccontr, simp only: not_le) + assume "CARD('n) + 1 < card s" + hence "affine_dependent s" using affine_dependent_biggerset[OF obt(1)] by auto + then obtain w v where wv:"setsum w s = 0" "v\s" "w v \ 0" "(\v\s. w v *s v) = 0" + using affine_dependent_explicit_finite[OF obt(1)] by auto + def i \ "(\v. (u v) / (- w v)) ` {v\s. w v < 0}" def t \ "Min i" + have "\x\s. w x < 0" proof(rule ccontr, simp add: not_less) + assume as:"\x\s. 0 \ w x" + hence "setsum w (s - {v}) \ 0" apply(rule_tac setsum_nonneg) by auto + hence "setsum w s > 0" unfolding setsum_diff1'[OF obt(1) `v\s`] + using as[THEN bspec[where x=v]] and `v\s` using `w v \ 0` by auto + thus False using wv(1) by auto + qed hence "i\{}" unfolding i_def by auto + + hence "t \ 0" using Min_ge_iff[of i 0 ] and obt(1) unfolding t_def i_def + using obt(4)[unfolded le_less] apply auto unfolding divide_le_0_iff by auto + have t:"\v\s. u v + t * w v \ 0" proof + fix v assume "v\s" hence v:"0\u v" using obt(4)[THEN bspec[where x=v]] by auto + show"0 \ u v + t * w v" proof(cases "w v < 0") + case False thus ?thesis apply(rule_tac add_nonneg_nonneg) + using v apply simp apply(rule mult_nonneg_nonneg) using `t\0` by auto next + case True hence "t \ u v / (- w v)" using `v\s` + unfolding t_def i_def apply(rule_tac Min_le) using obt(1) by auto + thus ?thesis unfolding real_0_le_add_iff + using pos_le_divide_eq[OF True[unfolded neg_0_less_iff_less[THEN sym]]] by auto + qed qed + + obtain a where "a\s" and "t = (\v. (u v) / (- w v)) a" and "w a < 0" + using Min_in[OF _ `i\{}`] and obt(1) unfolding i_def t_def by auto + hence a:"a\s" "u a + t * w a = 0" by auto + have *:"\f. setsum f (s - {a}) = setsum f s - ((f a)::'a::ring)" unfolding setsum_diff1'[OF obt(1) `a\s`] by auto + have "(\v\s. u v + t * w v) = 1" + unfolding setsum_addf wv(1) setsum_right_distrib[THEN sym] obt(5) by auto + moreover have "(\v\s. u v *s v + (t * w v) *s v) - (u a *s a + (t * w a) *s a) = y" + unfolding setsum_addf obt(6) vector_smult_assoc[THEN sym] setsum_cmul wv(4) + by (metis diff_0_right a(2) pth_5 pth_8 pth_d vector_mul_eq_0) + ultimately have "?P (n - 1)" 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 by (auto simp add: *) + thus False using smallest[THEN spec[where x="n - 1"]] by auto qed + thus "\s u. finite s \ s \ p \ card s \ CARD('n) + 1 + \ (\x\s. 0 \ u x) \ setsum u s = 1 \ (\v\s. u v *s v) = y" using obt by auto +qed auto + +lemma caratheodory: + "convex hull p = {x::real^'n::finite. \s. finite s \ s \ p \ + card s \ CARD('n) + 1 \ x \ convex hull s}" + unfolding expand_set_eq apply(rule, rule) unfolding mem_Collect_eq proof- + fix x assume "x \ convex hull p" + then obtain s u where "finite s" "s \ p" "card s \ CARD('n) + 1" + "\x\s. 0 \ u x" "setsum u s = 1" "(\v\s. u v *s v) = x"unfolding convex_hull_caratheodory by auto + thus "\s. finite s \ s \ p \ card s \ CARD('n) + 1 \ x \ convex hull s" + apply(rule_tac x=s in exI) using hull_subset[of s convex] + using convex_convex_hull[unfolded convex_explicit, of s, THEN spec[where x=s], THEN spec[where x=u]] by auto +next + fix x assume "\s. finite s \ s \ p \ card s \ CARD('n) + 1 \ x \ convex hull s" + then obtain s where "finite s" "s \ p" "card s \ CARD('n) + 1" "x \ convex hull s" by auto + thus "x \ convex hull p" using hull_mono[OF `s\p`] by auto +qed + +subsection {* Openness and compactness are preserved by convex hull operation. *} + +lemma open_convex_hull: + assumes "open s" + shows "open(convex hull s)" + unfolding open_contains_cball convex_hull_explicit unfolding mem_Collect_eq ball_simps(10) +proof(rule, rule) fix a + assume "\sa u. finite sa \ sa \ s \ (\x\sa. 0 \ u x) \ setsum u sa = 1 \ (\v\sa. u v *s v) = a" + then obtain t u where obt:"finite t" "t\s" "\x\t. 0 \ u x" "setsum u t = 1" "(\v\t. u v *s v) = a" by auto + + 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\{}" unfolding i_def using obt by auto def i \ "b ` t" + + show "\e>0. cball a e \ {y. \sa u. finite sa \ sa \ s \ (\x\sa. 0 \ u x) \ setsum u sa = 1 \ (\v\sa. u v *s v) = y}" + apply(rule_tac x="Min i" in exI) unfolding subset_eq apply rule defer apply rule unfolding mem_Collect_eq + proof- + 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` by auto + next fix y assume "y \ cball a (Min i)" + hence y:"norm (a - y) \ Min i" unfolding dist_def[THEN sym] by auto + { fix x assume "x\t" + hence "Min i \ b x" unfolding i_def apply(rule_tac Min_le) using obt(1) by auto + hence "x + (y - a) \ cball x (b x)" using y unfolding mem_cball dist_def 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 auto } + moreover + 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 setsum_reindex[OF *] o_def using obt(4) by auto + moreover have "(\v\(\v. v + (y - a)) ` t. u (v - (y - a)) *s v) = y" + unfolding setsum_reindex[OF *] o_def using obt(4,5) + by (simp add: setsum_addf setsum_subtractf setsum_vmul[OF obt(1), THEN sym]) + ultimately show "\sa u. finite sa \ (\x\sa. x \ s) \ (\x\sa. 0 \ u x) \ setsum u sa = 1 \ (\v\sa. u v *s 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) by auto + qed +qed + + +lemma compact_convex_combinations: + assumes "compact s" "compact t" + shows "compact { (1 - u) *s x + u *s y | x y u. 0 \ u \ u \ 1 \ x \ s \ y \ t}" +proof- + let ?X = "{ pastecart u w | u w. u \ {vec1 0 .. vec1 1} \ w \ { pastecart x y |x y. x \ s \ y \ t} }" + let ?h = "(\z. (1 - dest_vec1(fstcart z)) *s fstcart(sndcart z) + dest_vec1(fstcart z) *s sndcart(sndcart z))" + have *:"{ (1 - u) *s x + u *s y | x y u. 0 \ u \ u \ 1 \ x \ s \ y \ t} = ?h ` ?X" + apply(rule set_ext) unfolding image_iff mem_Collect_eq unfolding mem_interval_1 vec1_dest_vec1 + apply rule apply auto apply(rule_tac x="pastecart (vec1 u) (pastecart xa y)" in exI) apply simp + apply(rule_tac x="vec1 u" in exI) apply(rule_tac x="pastecart xa y" in exI) by auto + { fix u::"real^1" fix x y assume as:"0 \ dest_vec1 u" "dest_vec1 u \ 1" "x \ s" "y \ t" + hence "continuous (at (pastecart u (pastecart x y))) + (\z. fstcart (sndcart z) - dest_vec1 (fstcart z) *s fstcart (sndcart z) + + dest_vec1 (fstcart z) *s sndcart (sndcart z))" + apply (auto intro!: continuous_add continuous_sub continuous_mul simp add: o_def vec1_dest_vec1) + using linear_continuous_at linear_fstcart linear_sndcart linear_sndcart + using linear_compose[unfolded o_def] by auto } + hence "continuous_on {pastecart u w |u w. u \ {vec1 0..vec1 1} \ w \ {pastecart x y |x y. x \ s \ y \ t}} + (\z. (1 - dest_vec1 (fstcart z)) *s fstcart (sndcart z) + dest_vec1 (fstcart z) *s sndcart (sndcart z))" + apply(rule_tac continuous_at_imp_continuous_on) unfolding mem_Collect_eq + unfolding mem_interval_1 vec1_dest_vec1 by auto + thus ?thesis unfolding * apply(rule compact_continuous_image) + defer apply(rule compact_pastecart) defer apply(rule compact_pastecart) + using compact_interval assms by auto +qed + +lemma compact_convex_hull: fixes s::"(real^'n::finite) set" + assumes "compact s" shows "compact(convex hull s)" +proof(cases "s={}") + case True thus ?thesis using compact_empty by simp +next + case False then obtain w where "w\s" by auto + show ?thesis unfolding caratheodory[of s] + proof(induct "CARD('n) + 1") + have *:"{x.\sa. finite sa \ sa \ s \ card sa \ 0 \ x \ convex hull sa} = {}" + using compact_empty by (auto simp add: convex_hull_empty) + case 0 thus ?case unfolding * by simp + next + case (Suc n) + show ?case proof(cases "n=0") + case True have "{x. \t. finite t \ t \ s \ card t \ Suc n \ x \ convex hull t} = s" + unfolding expand_set_eq 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" by auto + show "x\s" proof(cases "card t = 0") + case True thus ?thesis using t(4) unfolding card_0_eq[OF t(1)] by(simp add: convex_hull_empty) + next + case False hence "card t = Suc 0" using t(3) `n=0` by auto + then obtain a where "t = {a}" unfolding card_Suc_eq by auto + thus ?thesis using t(2,4) by (simp add: convex_hull_singleton) + qed + next + fix x assume "x\s" + thus "\t. finite t \ t \ s \ card t \ Suc n \ x \ convex hull t" + apply(rule_tac x="{x}" in exI) unfolding convex_hull_singleton by auto + qed thus ?thesis using assms by simp + next + case False have "{x. \t. finite t \ t \ s \ card t \ Suc n \ x \ convex hull t} = + { (1 - u) *s x + u *s y | x y u. + 0 \ u \ u \ 1 \ x \ s \ y \ {x. \t. finite t \ t \ s \ card t \ n \ x \ convex hull t}}" + unfolding expand_set_eq and mem_Collect_eq proof(rule,rule) + fix x assume "\u v c. x = (1 - c) *s u + c *s 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) *s u + c *s v" + "0 \ c \ c \ 1" "u \ s" "finite t" "t \ s" "card t \ n" "v \ convex hull t" by auto + moreover have "(1 - c) *s u + c *s v \ convex hull insert u t" + apply(rule mem_convex) 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"] by auto + ultimately show "\t. finite t \ t \ s \ card t \ Suc n \ x \ convex hull t" + apply(rule_tac x="insert u t" in exI) by (auto simp add: card_insert_if) + 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" by auto + let ?P = "\u v c. x = (1 - c) *s u + c *s v \ + 0 \ c \ c \ 1 \ u \ s \ (\t. finite t \ t \ s \ card t \ n \ v \ convex hull t)" + show ?P proof(cases "card t = Suc n") + case False hence "card t \ n" using t(3) by auto + thus ?P apply(rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI) using `w\s` and t + by(auto intro!: exI[where x=t]) + next + case True then obtain a u where au:"t = insert a u" "a\u" apply(drule_tac card_eq_SucD) by auto + show ?P proof(cases "u={}") + case True hence "x=a" using t(4)[unfolded au] by auto + show ?P unfolding `x=a` apply(rule_tac x=a in exI, rule_tac x=a in exI, rule_tac x=1 in exI) + using t and `n\0` unfolding au by(auto intro!: exI[where x="{a}"] simp add: convex_hull_singleton) + next + case False obtain ux vx b where obt:"ux\0" "vx\0" "ux + vx = 1" "b \ convex hull u" "x = ux *s a + vx *s b" + using t(4)[unfolded au convex_hull_insert[OF False]] by auto + have *:"1 - vx = ux" using obt(3) by auto + show ?P apply(rule_tac x=a in exI, rule_tac x=b in exI, rule_tac x=vx in exI) + using obt and t(1-3) unfolding au and * using card_insert_disjoint[OF _ au(2)] + by(auto intro!: exI[where x=u]) + qed + qed + qed + thus ?thesis using compact_convex_combinations[OF assms Suc] by simp + qed + qed +qed + +lemma finite_imp_compact_convex_hull: + "finite s \ compact(convex hull s)" + apply(drule finite_imp_compact, drule compact_convex_hull) by assumption + +subsection {* Extremal points of a simplex are some vertices. *} + +lemma dist_increases_online: assumes "d \ 0" + shows "dist a (b + d) > dist a b \ dist a (b - d) > dist a b" +proof(cases "a \ d - b \ d > 0") + case True hence "0 < d \ d + (a \ d * 2 - b \ d * 2)" + apply(rule_tac add_pos_pos) using assms by auto + thus ?thesis apply(rule_tac disjI2) unfolding dist_def and real_vector_norm_def and real_sqrt_less_iff + by(simp add: dot_rsub dot_radd dot_lsub dot_ladd dot_sym field_simps) +next + case False hence "0 < d \ d + (b \ d * 2 - a \ d * 2)" + apply(rule_tac add_pos_nonneg) using assms by auto + thus ?thesis apply(rule_tac disjI1) unfolding dist_def and real_vector_norm_def and real_sqrt_less_iff + by(simp add: dot_rsub dot_radd dot_lsub dot_ladd dot_sym field_simps) +qed + +lemma norm_increases_online: + "(d::real^'n::finite) \ 0 \ norm(a + d) > norm a \ norm(a - d) > norm a" + using dist_increases_online[of d a 0] unfolding dist_def by auto + +lemma simplex_furthest_lt: + fixes s::"(real^'n::finite) set" assumes "finite s" + shows "\x \ (convex hull s). x \ s \ (\y\(convex hull s). norm(x - a) < norm(y - a))" +proof(induct_tac rule: finite_induct[of 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(rule,rule,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 *s x + v *s 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") + case True 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 + thus ?thesis apply(rule_tac x=z in bexI) unfolding convex_hull_insert[OF False] by auto + next + case False show ?thesis using obt(3) proof(cases "u=0", case_tac[!] "v=0") + assume "u=0" "v\0" hence "y = b" using obt by auto + thus ?thesis using False and obt(4) by auto + next + assume "u\0" "v=0" hence "y = x" using obt by auto + thus ?thesis using y(2) by auto + next + assume "u\0" "v\0" + then obtain w where w:"w>0" "wb" proof(rule ccontr) + assume "\ x\b" hence "y=b" unfolding obt(5) + using obt(3) by(auto simp add: vector_sadd_rdistrib[THEN sym]) + thus False using obt(4) and False by simp qed + hence *:"w *s (x - b) \ 0" using w(1) by auto + show ?thesis using dist_increases_online[OF *, of a y] + proof(erule_tac disjE) + assume "dist a y < dist a (y + w *s (x - b))" + hence "norm (y - a) < norm ((u + w) *s x + (v - w) *s b - a)" + unfolding dist_sym[of a] unfolding dist_def obt(5) by (simp add: ring_simps) + moreover have "(u + w) *s x + (v - w) *s 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) by auto + ultimately show ?thesis by auto + next + assume "dist a y < dist a (y - w *s (x - b))" + hence "norm (y - a) < norm ((u - w) *s x + (v + w) *s b - a)" + unfolding dist_sym[of a] unfolding dist_def obt(5) by (simp add: ring_simps) + moreover have "(u - w) *s x + (v + w) *s 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) by auto + ultimately show ?thesis by auto + qed + qed auto + qed + qed auto +qed (auto simp add: assms) + +lemma simplex_furthest_le: + assumes "finite s" "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] + unfolding dist_sym[of a] unfolding dist_def by auto + thus ?thesis proof(cases "x\s") + case False 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 + thus ?thesis using x(2)[THEN bspec[where x=y]] by auto + qed auto +qed + +lemma simplex_furthest_le_exists: + "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: + assumes "finite s" "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)" + using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]] by auto + thus ?thesis proof(cases "u\s \ v\s", erule_tac 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 + thus ?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)" + using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2) by auto + thus ?thesis using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1) + by (auto simp add: norm_minus_commute) + qed auto +qed + +lemma simplex_extremal_le_exists: + "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 + "closest_point s a = (SOME x. x \ s \ (\y\s. dist a x \ dist a y))" + +lemma closest_point_exists: + assumes "closed s" "s \ {}" + shows "closest_point s a \ s" "\y\s. dist a (closest_point s a) \ dist a y" + unfolding closest_point_def apply(rule_tac[!] someI2_ex) + using distance_attains_inf[OF assms(1,2), of a] by auto + +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_self: + assumes "x \ s" shows "closest_point s x = x" + unfolding closest_point_def apply(rule some1_equality, rule ex1I[of _ x]) + using assms by auto + +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: fixes y::"real^'n::finite" + assumes "y \ z > 0" + shows "\u>0. \v>0. v \ u \ norm(v *s z - y) < norm y" +proof- have z:"z \ z > 0" unfolding dot_pos_lt using assms by auto + thus ?thesis using assms apply(rule_tac x="(y \ z) / (z \ z)" in exI) apply(rule) defer proof(rule+) + fix v assume "0 y \ z / (z \ z)" + thus "norm (v *s z - y) < norm y" unfolding norm_lt using z and assms + by (simp add: field_simps dot_sym mult_strict_left_mono[OF _ `0 (z - x) > 0" + shows "\u>0. u \ 1 \ dist (x + u *s (z - x)) y < dist x y" +proof- obtain u where "u>0" and u:"\v>0. v \ u \ norm (v *s (z - x) - (y - x)) < norm (y - x)" + using closer_points_lemma[OF assms] by auto + 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_def by(auto simp add: 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" + shows "(a - x) \ (y - x) \ 0" +proof(rule ccontr) assume "\ (a - x) \ (y - x) \ 0" + then obtain u where u:"u>0" "u\1" "dist (x + u *s (y - x)) a < dist x a" using closer_point_lemma[of a x y] by auto + let ?z = "(1 - u) *s x + u *s y" have "?z \ s" using mem_convex[OF assms(1,3,4), of u] using u by auto + thus False using assms(5)[THEN bspec[where x="?z"]] and u(3) by (auto simp add: dist_sym field_simps) qed + +lemma any_closest_point_unique: + 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 + +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"] + using closest_point_exists[OF assms(2)] and assms(3) by auto + +lemma closest_point_dot: + assumes "convex s" "closed s" "x \ s" + shows "(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) by auto + +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" + apply(rule ccontr) apply(rule_tac notE[OF assms(4)]) + apply(rule closest_point_unique[OF assms(1-3), of a]) + using closest_point_le[OF assms(2), of _ a] by fastsimp + +lemma closest_point_lipschitz: + assumes "convex s" "closed s" "s \ {}" + shows "dist (closest_point s x) (closest_point s y) \ dist x y" +proof- + have "(x - closest_point s x) \ (closest_point s y - closest_point s x) \ 0" + "(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)] by auto + thus ?thesis unfolding dist_def and norm_le + using dot_pos_le[of "(x - closest_point s x) - (y - closest_point s y)"] + by (auto simp add: dot_sym dot_ladd dot_radd) qed + +lemma continuous_at_closest_point: + assumes "convex s" "closed s" "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" "closed s" "s \ {}" + shows "continuous_on t (closest_point s)" + apply(rule continuous_at_imp_continuous_on) using continuous_at_closest_point[OF assms] by auto + +subsection {* Various point-to-set separating/supporting hyperplane theorems. *} + +lemma supporting_hyperplane_closed_point: + assumes "convex s" "closed s" "s \ {}" "z \ s" + shows "\a b. \y\s. a \ z < b \ (a \ y = b) \ (\x\s. a \ x \ b)" +proof- + from distance_attains_inf[OF assms(2-3)] obtain y where "y\s" and y:"\x\s. dist z y \ dist z x" by auto + show ?thesis apply(rule_tac x="y - z" in exI, rule_tac x="(y - z) \ y" in exI, rule_tac x=y in bexI) + apply rule defer apply rule defer apply(rule, rule ccontr) using `y\s` proof- + show "(y - z) \ z < (y - z) \ y" apply(subst diff_less_iff(1)[THEN sym]) + unfolding dot_rsub[THEN sym] and dot_pos_lt using `y\s` `z\s` by auto + next + fix x assume "x\s" have *:"\u. 0 \ u \ u \ 1 \ dist z y \ dist z ((1 - u) *s y + u *s 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 *s (x - y)) z < dist y z" using closer_point_lemma[of z y x] by auto + thus False using *[THEN spec[where x=v]] by(auto simp add: dist_sym field_simps) + qed auto +qed + +lemma separating_hyperplane_closed_point: + assumes "convex s" "closed s" "z \ s" + shows "\a b. a \ z < b \ (\x\s. a \ x > b)" +proof(cases "s={}") + case True thus ?thesis apply(rule_tac x="-z" in exI, rule_tac x=1 in exI) + using less_le_trans[OF _ dot_pos_le[of z]] by auto +next + case False obtain y where "y\s" and y:"\x\s. dist z y \ dist z x" + using distance_attains_inf[OF assms(2) False] by auto + show ?thesis apply(rule_tac x="y - z" in exI, rule_tac x="(y - z) \ z + (norm(y - z))\ / 2" in exI) + apply rule defer apply rule proof- + fix x assume "x\s" + have "\ 0 < (z - y) \ (x - y)" apply(rule_tac notI) proof(drule closer_point_lemma) + assume "\u>0. u \ 1 \ dist (y + u *s (x - y)) z < dist y z" + then obtain u where "u>0" "u\1" "dist (y + u *s (x - y)) z < dist y z" by auto + thus False using y[THEN bspec[where x="y + u *s (x - y)"]] + using assms(1)[unfolded convex_alt, THEN bspec[where x=y]] + using `x\s` `y\s` by (auto simp add: dist_sym field_simps) qed + moreover have "0 < norm (y - z) ^ 2" using `y\s` `z\s` by auto + hence "0 < (y - z) \ (y - z)" unfolding norm_pow_2 by simp + ultimately show "(y - z) \ z + (norm (y - z))\ / 2 < (y - z) \ x" + unfolding norm_pow_2 and dlo_simps(3) by (auto simp add: field_simps dot_sym) + qed(insert `y\s` `z\s`, auto) +qed + +lemma separating_hyperplane_closed_0: + assumes "convex (s::(real^'n::finite) set)" "closed s" "0 \ s" + shows "\a b. a \ 0 \ 0 < b \ (\x\s. a \ x > b)" + proof(cases "s={}") guess a using UNIV_witness[where 'a='n] .. + case True have "norm ((basis a)::real^'n::finite) = 1" + using norm_basis and dimindex_ge_1 by auto + thus ?thesis apply(rule_tac x="basis a" in exI, rule_tac x=1 in exI) using True by auto +next case False thus ?thesis using False using separating_hyperplane_closed_point[OF assms] + apply - apply(erule exE)+ unfolding dot_rzero apply(rule_tac x=a in exI, rule_tac x=b in exI) by auto qed + +subsection {* Now set-to-set for closed/compact sets. *} + +lemma separating_hyperplane_closed_compact: + assumes "convex (s::(real^'n::finite) set)" "closed s" "convex t" "compact t" "t \ {}" "s \ t = {}" + shows "\a b. (\x\s. a \ x < b) \ (\x\t. a \ x > b)" +proof(cases "s={}") + case True + obtain b where b:"b>0" "\x\t. norm x \ b" using compact_imp_bounded[OF assms(4)] unfolding bounded_pos by auto + obtain z::"real^'n" where z:"norm z = b + 1" using vector_choose_size[of "b + 1"] and b(1) by auto + hence "z\t" using b(2)[THEN bspec[where x=z]] by auto + then obtain a b where ab:"a \ z < b" "\x\t. b < a \ x" + using separating_hyperplane_closed_point[OF assms(3) compact_imp_closed[OF assms(4)], of z] by auto + thus ?thesis using True by auto +next + case False then obtain y where "y\s" by auto + obtain a b where "0 < b" "\x\{x - y |x y. x \ s \ y \ t}. b < a \ x" + using separating_hyperplane_closed_point[OF convex_differences[OF assms(1,3)], of 0] + using closed_compact_differences[OF assms(2,4)] using assms(6) by(auto, blast) + hence ab:"\x\s. \y\t. b + a \ y < a \ x" apply- apply(rule,rule) apply(erule_tac x="x - y" in ballE) by auto + def k \ "rsup ((\x. a \ x) ` t)" + show ?thesis apply(rule_tac x="-a" in exI, rule_tac x="-(k + b / 2)" in exI) + apply(rule,rule) defer apply(rule) unfolding dot_lneg and neg_less_iff_less proof- + from ab have "((\x. a \ x) ` t) *<= (a \ y - b)" + apply(erule_tac x=y in ballE) apply(rule setleI) using `y\s` by auto + hence k:"isLub UNIV ((\x. a \ x) ` t) k" unfolding k_def apply(rule_tac rsup) using assms(5) by auto + fix x assume "x\t" thus "a \ x < (k + b / 2)" using `0 x"] by auto + next + fix x assume "x\s" + hence "k \ a \ x - b" unfolding k_def apply(rule_tac rsup_le) using assms(5) + unfolding setle_def + using ab[THEN bspec[where x=x]] by auto + thus "k + b / 2 < a \ x" using `0 < b` by auto + qed +qed + +lemma separating_hyperplane_compact_closed: + assumes "convex s" "compact s" "s \ {}" "convex t" "closed t" "s \ t = {}" + shows "\a b. (\x\s. a \ x < b) \ (\x\t. a \ x > b)" +proof- obtain a b where "(\x\t. a \ x < b) \ (\x\s. b < a \ x)" + using separating_hyperplane_closed_compact[OF assms(4-5,1-2,3)] and assms(6) by auto + thus ?thesis apply(rule_tac x="-a" in exI, rule_tac x="-b" in exI) by auto qed + +subsection {* General case without assuming closure and getting non-strict separation. *} + +lemma separating_hyperplane_set_0: + assumes "convex s" "(0::real^'n::finite) \ s" + shows "\a. a \ 0 \ (\x\s. 0 \ a \ x)" +proof- let ?k = "\c. {x::real^'n. 0 \ c \ x}" + have "frontier (cball 0 1) \ (\ (?k ` s)) \ {}" + apply(rule compact_imp_fip) apply(rule compact_frontier[OF compact_cball]) + defer apply(rule,rule,erule conjE) proof- + fix f assume as:"f \ ?k ` s" "finite f" + obtain c where c:"f = ?k ` c" "c\s" "finite c" using finite_subset_image[OF as] by auto + then obtain a b where ab:"a \ 0" "0 < b" "\x\convex hull c. b < a \ x" + using separating_hyperplane_closed_0[OF convex_convex_hull, of c] + using finite_imp_compact_convex_hull[OF c(3), THEN compact_imp_closed] and assms(2) + using subset_hull[unfolded mem_def, of convex, OF assms(1), THEN sym, of c] by auto + hence "\x. norm x = 1 \ (\y\c. 0 \ y \ x)" apply(rule_tac x="inverse(norm a) *s a" in exI) + using hull_subset[of c convex] unfolding subset_eq and dot_rmult + apply- apply rule defer apply rule apply(rule mult_nonneg_nonneg) + by(auto simp add: dot_sym elim!: ballE) + thus "frontier (cball 0 1) \ \f \ {}" unfolding c(1) frontier_cball dist_def by auto + qed(insert closed_halfspace_ge, auto) + then obtain x where "norm x = 1" "\y\s. x\?k y" unfolding frontier_cball dist_def by auto + thus ?thesis apply(rule_tac x=x in exI) by(auto simp add: dot_sym) qed + +lemma separating_hyperplane_sets: + assumes "convex s" "convex (t::(real^'n::finite) set)" "s \ {}" "t \ {}" "s \ t = {}" + shows "\a b. a \ 0 \ (\x\s. a \ x \ b) \ (\x\t. 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 \ a \ x" using assms(3-5) by auto + hence "\x\t. \y\s. a \ y \ a \ x" apply- apply(rule, rule) apply(erule_tac x="x - y" in ballE) by auto + thus ?thesis apply(rule_tac x=a in exI, rule_tac x="rsup ((\x. a \ x) ` s)" in exI) using `a\0` + apply(rule) apply(rule,rule) apply(rule rsup[THEN isLubD2]) prefer 4 apply(rule,rule rsup_le) unfolding setle_def + prefer 4 using assms(3-5) by blast+ qed + +subsection {* More convexity generalities. *} + +lemma convex_closure: assumes "convex s" shows "convex(closure s)" + unfolding convex_def Ball_def closure_sequential + apply(rule,rule,rule,rule,rule,rule,rule,rule,rule) apply(erule_tac exE)+ + apply(rule_tac x="\n. u *s xb n + v *s xc n" in exI) apply(rule,rule) + apply(rule assms[unfolded convex_def, rule_format]) prefer 6 + apply(rule Lim_add) apply(rule_tac [1-2] Lim_cmul) by auto + +lemma convex_interior: assumes "convex s" shows "convex(interior s)" + unfolding convex_alt Ball_def mem_interior apply(rule,rule,rule,rule,rule,rule) apply(erule exE | erule conjE)+ proof- + 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) *s x + u *s y) e \ s" apply(rule_tac x="min d e" in exI) + apply rule unfolding subset_eq defer apply rule proof- + fix z assume "z \ ball ((1 - u) *s x + u *s y) (min d e)" + hence "(1- u) *s (z - u *s (y - x)) + u *s (z + (1 - u) *s (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_def by(auto simp add: ring_simps) + thus "z \ s" using u by (auto simp add: ring_simps) qed(insert u ed(3-4), auto) qed + +lemma convex_hull_eq_empty: "convex hull s = {} \ s = {}" + using hull_subset[of s convex] convex_hull_empty by auto + +subsection {* Moving and scaling convex hulls. *} + +lemma convex_hull_translation_lemma: + "convex hull ((\x. a + x) ` s) \ (\x. a + x) ` (convex hull s)" + apply(rule hull_minimal, rule image_mono, rule hull_subset) unfolding mem_def + using convex_translation[OF convex_convex_hull, of a s] by assumption + +lemma convex_hull_bilemma: fixes neg + assumes "(\s a. (convex hull (up a s)) \ up a (convex hull s))" + shows "(\s. up a (up (neg a) s) = s) \ (\s. up (neg a) (up a s) = s) \ (\s t a. s \ t \ up a s \ up a t) + \ \s. (convex hull (up a s)) = up a (convex hull s)" + using assms by(metis subset_antisym) + +lemma convex_hull_translation: + "convex hull ((\x. a + x) ` s) = (\x. a + x) ` (convex hull s)" + apply(rule convex_hull_bilemma[rule_format, of _ _ "\a. -a"], rule convex_hull_translation_lemma) unfolding image_image by auto + +lemma convex_hull_scaling_lemma: + "(convex hull ((\x. c *s x) ` s)) \ (\x. c *s x) ` (convex hull s)" + apply(rule hull_minimal, rule image_mono, rule hull_subset) + unfolding mem_def by(rule convex_scaling, rule convex_convex_hull) + +lemma convex_hull_scaling: + "convex hull ((\x. c *s x) ` s) = (\x. c *s x) ` (convex hull s)" + apply(cases "c=0") defer apply(rule convex_hull_bilemma[rule_format, of _ _ inverse]) apply(rule convex_hull_scaling_lemma) + unfolding image_image vector_smult_assoc by(auto simp add:image_constant_conv convex_hull_eq_empty) + +lemma convex_hull_affinity: + "convex hull ((\x. a + c *s x) ` s) = (\x. a + c *s x) ` (convex hull s)" + unfolding image_image[THEN sym] convex_hull_scaling convex_hull_translation .. + +subsection {* Convex set as intersection of halfspaces. *} + +lemma convex_halfspace_intersection: + assumes "closed s" "convex s" + shows "s = \ {h. s \ h \ (\a b. h = {x. a \ x \ b})}" + apply(rule set_ext, rule) unfolding Inter_iff Ball_def mem_Collect_eq apply(rule,rule,erule conjE) proof- + fix x assume "\xa. s \ xa \ (\a b. xa = {x. a \ x \ b}) \ x \ xa" + hence "\a b. s \ {x. a \ x \ b} \ x \ {x. a \ x \ b}" by blast + thus "x\s" apply(rule_tac ccontr) apply(drule separating_hyperplane_closed_point[OF assms(2,1)]) + apply(erule exE)+ apply(erule_tac x="-a" in allE, erule_tac x="-b" in allE) by auto +qed auto + +subsection {* Radon's theorem (from Lars Schewe). *} + +lemma radon_ex_lemma: + assumes "finite c" "affine_dependent c" + shows "\u. setsum u c = 0 \ (\v\c. u v \ 0) \ setsum (\v. u v *s v) c = 0" +proof- from assms(2)[unfolded affine_dependent_explicit] guess s .. then guess u .. + thus ?thesis apply(rule_tac x="\v. if v\s then u v else 0" in exI) unfolding if_smult vector_smult_lzero + and setsum_restrict_set[OF assms(1), THEN sym] by(auto simp add: Int_absorb1) qed + +lemma radon_s_lemma: + assumes "finite s" "setsum f s = (0::real)" + shows "setsum f {x\s. 0 < f x} = - setsum f {x\s. f x < 0}" +proof- have *:"\x. (if f x < 0 then f x else 0) + (if 0 < f x then f x else 0) = f x" by auto + show ?thesis unfolding real_add_eq_0_iff[THEN sym] and setsum_restrict_set''[OF assms(1)] and setsum_addf[THEN sym] and * + using assms(2) by assumption qed + +lemma radon_v_lemma: + assumes "finite s" "setsum f s = 0" "\x. g x = (0::real) \ f x = (0::real^'n)" + shows "(setsum f {x\s. 0 < g x}) = - setsum f {x\s. g x < 0}" +proof- + have *:"\x. (if 0 < g x then f x else 0) + (if g x < 0 then f x else 0) = f x" using assms(3) by auto + show ?thesis unfolding eq_neg_iff_add_eq_0 and setsum_restrict_set''[OF assms(1)] and setsum_addf[THEN sym] and * + using assms(2) by assumption qed + +lemma radon_partition: + assumes "finite c" "affine_dependent c" + shows "\m p. m \ p = {} \ m \ p = c \ (convex hull m) \ (convex hull p) \ {}" proof- + obtain u v where uv:"setsum u c = 0" "v\c" "u v \ 0" "(\v\c. u v *s v) = 0" using radon_ex_lemma[OF assms] by auto + have fin:"finite {x \ c. 0 < u x}" "finite {x \ c. 0 > u x}" using assms(1) by auto + def z \ "(inverse (setsum u {x\c. u x > 0})) *s setsum (\x. u x *s x) {x\c. u x > 0}" + have "setsum u {x \ c. 0 < u x} \ 0" proof(cases "u v \ 0") + case False hence "u v < 0" by auto + thus ?thesis proof(cases "\w\{x \ c. 0 < u x}. u w > 0") + case True thus ?thesis using setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] by auto + next + case False hence "setsum u c \ setsum (\x. if x=v then u v else 0) c" apply(rule_tac setsum_mono) by auto + thus ?thesis unfolding setsum_delta[OF assms(1)] using uv(2) and `u v < 0` and uv(1) by auto qed + qed (insert setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] uv(2-3), auto) + + hence *:"setsum u {x\c. u x > 0} > 0" unfolding real_less_def apply(rule_tac conjI, rule_tac setsum_nonneg) by auto + moreover have "setsum u ({x \ c. 0 < u x} \ {x \ c. u x < 0}) = setsum u c" + "(\x\{x \ c. 0 < u x} \ {x \ c. u x < 0}. u x *s x) = (\x\c. u x *s x)" + using assms(1) apply(rule_tac[!] setsum_mono_zero_left) by auto + hence "setsum u {x \ c. 0 < u x} = - setsum u {x \ c. 0 > u x}" + "(\x\{x \ c. 0 < u x}. u x *s x) = - (\x\{x \ c. 0 > u x}. u x *s x)" + unfolding eq_neg_iff_add_eq_0 using uv(1,4) by (auto simp add: setsum_Un_zero[OF fin, THEN sym]) + moreover have "\x\{v \ c. u v < 0}. 0 \ inverse (setsum u {x \ c. 0 < u x}) * - u x" + apply (rule) apply (rule mult_nonneg_nonneg) using * by auto + + ultimately have "z \ convex hull {v \ c. u v \ 0}" unfolding convex_hull_explicit mem_Collect_eq + apply(rule_tac x="{v \ c. u v < 0}" in exI, rule_tac x="\y. inverse (setsum u {x\c. u x > 0}) * - u y" in exI) + using assms(1) unfolding vector_smult_assoc[THEN sym] setsum_cmul and z_def + by(auto simp add: setsum_negf vector_smult_lneg mult_right.setsum[THEN sym]) + moreover have "\x\{v \ c. 0 < u v}. 0 \ inverse (setsum u {x \ c. 0 < u x}) * u x" + apply (rule) apply (rule mult_nonneg_nonneg) using * by auto + hence "z \ convex hull {v \ c. u v > 0}" unfolding convex_hull_explicit mem_Collect_eq + apply(rule_tac x="{v \ c. 0 < u v}" in exI, rule_tac x="\y. inverse (setsum u {x\c. u x > 0}) * u y" in exI) + using assms(1) unfolding vector_smult_assoc[THEN sym] setsum_cmul and z_def using * + by(auto simp add: setsum_negf vector_smult_lneg mult_right.setsum[THEN sym]) + ultimately show ?thesis apply(rule_tac x="{v\c. u v \ 0}" in exI, rule_tac x="{v\c. u v > 0}" in exI) by auto +qed + +lemma radon: assumes "affine_dependent c" + obtains m p where "m\c" "p\c" "m \ p = {}" "(convex hull m) \ (convex hull p) \ {}" +proof- from assms[unfolded affine_dependent_explicit] guess s .. then guess u .. + hence *:"finite s" "affine_dependent s" and s:"s \ c" unfolding affine_dependent_explicit by auto + from radon_partition[OF *] guess m .. then guess p .. + thus ?thesis apply(rule_tac that[of p m]) using s by auto qed + +subsection {* Helly's theorem. *} + +lemma helly_induct: fixes f::"(real^'n::finite) set set" + assumes "f hassize n" "n \ CARD('n) + 1" + "\s\f. convex s" "\t\f. card t = CARD('n) + 1 \ \ t \ {}" + shows "\ f \ {}" + using assms unfolding hassize_def apply(erule_tac conjE) proof(induct n arbitrary: f) +case (Suc n) +show "\ f \ {}" apply(cases "n = CARD('n)") apply(rule Suc(4)[rule_format]) + unfolding card_Diff_singleton_if[OF Suc(5)] and Suc(6) proof- + assume ng:"n \ CARD('n)" hence "\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 Suc(5)] and Suc(6) + defer apply(rule Suc(3)[rule_format]) defer apply(rule Suc(4)[rule_format]) using Suc(2,5) by auto + then obtain X where X:"\s\f. X s \ \(f - {s})" by auto + 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" unfolding inj_on_def by auto + hence *:"\ f = \ (f - {s}) \ \ (f - {t})" by auto + show ?thesis unfolding * unfolding ex_in_conv[THEN sym] apply(rule_tac x="X s" in exI) + apply(rule, rule X[rule_format]) using X st by auto + 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 Suc(6) using Suc(2,5) and ng by auto + have "m \ X ` f" "p \ X ` f" using mp(2) by auto + then obtain g h where gh:"m = X ` g" "p = X ` h" "g \ f" "h \ f" unfolding subset_image_iff by auto + hence "f \ (g \ h) = f" by auto + hence f:"f = g \ h" using inj_on_image_eq_iff[of X f "g \ h"] and True + unfolding mp(2)[unfolded image_Un[THEN sym] gh] by auto + have *:"g \ h = {}" using mp(1) unfolding gh 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(3) gh(3-4) unfolding mem_def 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 guess y unfolding image_iff .. + thus "x\\h" using X[THEN bspec[where x=y]] using * f by auto next + fix x assume "x\X ` h" then guess y unfolding image_iff .. + thus "x\\g" using X[THEN bspec[where x=y]] using * f by auto + qed(auto) + thus ?thesis unfolding f using mp(3)[unfolded gh] by blast qed +qed(insert dimindex_ge_1, auto) qed(auto) + +lemma helly: fixes f::"(real^'n::finite) set set" + assumes "finite f" "card f \ CARD('n) + 1" "\s\f. convex s" + "\t\f. card t = CARD('n) + 1 \ \ t \ {}" + shows "\ f \{}" + apply(rule helly_induct) unfolding hassize_def using assms by auto + +subsection {* Convex hull is "preserved" by a linear function. *} + +lemma convex_hull_linear_image: + assumes "linear f" + shows "f ` (convex hull s) = convex 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- show "convex {x. f x \ convex hull f ` s}" + unfolding convex_def by(auto simp add: linear_cmul[OF assms] linear_add[OF assms] + convex_convex_hull[unfolded convex_def, rule_format]) next + show "convex {x. x \ f ` (convex hull s)}" using convex_convex_hull[unfolded convex_def, of s] + unfolding convex_def by (auto simp add: linear_cmul[OF assms, THEN sym] linear_add[OF assms, THEN sym]) +qed auto + +lemma in_convex_hull_linear_image: + assumes "linear f" "x \ convex hull s" shows "(f x) \ convex hull (f ` s)" +using convex_hull_linear_image[OF assms(1)] assms(2) by auto + +subsection {* Homeomorphism of all convex compact sets with nonempty interior. *} + +lemma compact_frontier_line_lemma: + assumes "compact s" "0 \ s" "x \ 0" + obtains u where "0 \ u" "(u *s x) \ frontier s" "\v>u. (v *s x) \ s" +proof- + obtain b where b:"b>0" "\x\s. norm x \ b" using compact_imp_bounded[OF assms(1), unfolded bounded_pos] by auto + let ?A = "{y. \u. 0 \ u \ u \ b / norm(x) \ (y = u *s x)}" + have A:"?A = (\u. dest_vec1 u *s x) ` {0 .. vec1 (b / norm x)}" + unfolding image_image[of "\u. u *s x" "\x. dest_vec1 x", THEN sym] + unfolding dest_vec1_inverval vec1_dest_vec1 by auto + have "compact ?A" unfolding A apply(rule compact_continuous_image, rule continuous_at_imp_continuous_on) + apply(rule, rule continuous_vmul) unfolding o_def vec1_dest_vec1 apply(rule continuous_at_id) by(rule compact_interval) + moreover have "{y. \u\0. u \ b / norm x \ y = u *s x} \ s \ {}" apply(rule not_disjointI[OF _ assms(2)]) + unfolding mem_Collect_eq using `b>0` assms(3) by(auto intro!: divide_nonneg_pos) + ultimately obtain u y where obt: "u\0" "u \ b / norm x" "y = u *s x" + "y\?A" "y\s" "\z\?A \ s. dist 0 z \ dist 0 y" using distance_attains_sup[OF compact_inter[OF _ assms(1), of ?A], of 0] by auto + + have "norm x > 0" using assms(3)[unfolded zero_less_norm_iff[THEN sym]] by auto + { fix v assume as:"v > u" "v *s x \ s" + hence "v \ b / norm x" using b(2)[rule_format, OF as(2)] + using `u\0` unfolding pos_le_divide_eq[OF `norm x > 0`] and norm_mul by auto + hence "norm (v *s x) \ norm y" apply(rule_tac obt(6)[rule_format, unfolded dist_0_norm]) apply(rule IntI) defer + apply(rule as(2)) unfolding mem_Collect_eq apply(rule_tac x=v in exI) + using as(1) `u\0` by(auto simp add:field_simps) + hence False unfolding obt(3) unfolding norm_mul using `u\0` `norm x > 0` `v>u` by(auto simp add:field_simps) + } note u_max = this + + have "u *s x \ frontier s" unfolding frontier_straddle apply(rule,rule,rule) apply(rule_tac x="u *s x" in bexI) unfolding obt(3)[THEN sym] + prefer 3 apply(rule_tac x="(u + (e / 2) / norm x) *s x" in exI) apply(rule, rule) proof- + fix e assume "0 < e" and as:"(u + e / 2 / norm x) *s x \ s" + hence "u + e / 2 / norm x > u" using`norm x > 0` by(auto simp del:zero_less_norm_iff intro!: divide_pos_pos) + thus False using u_max[OF _ as] by auto + qed(insert `y\s`, auto simp add: dist_def obt(3)) + thus ?thesis apply(rule_tac that[of u]) apply(rule obt(1), assumption) + apply(rule,rule,rule ccontr) apply(rule u_max) by auto qed + +lemma starlike_compact_projective: + assumes "compact s" "cball (0::real^'n::finite) 1 \ s " + "\x\s. \u. 0 \ u \ u < 1 \ (u *s x) \ (s - frontier s )" + shows "s homeomorphic (cball (0::real^'n::finite) 1)" +proof- + have fs:"frontier s \ s" apply(rule frontier_subset_closed) using compact_imp_closed[OF assms(1)] by simp + def pi \ "\x::real^'n. inverse (norm x) *s x" + have "0 \ frontier s" unfolding frontier_straddle apply(rule ccontr) unfolding not_not apply(erule_tac x=1 in allE) + using assms(2)[unfolded subset_eq Ball_def mem_cball] by auto + have injpi:"\x y. pi x = pi y \ norm x = norm y \ x = y" unfolding pi_def by auto + + have contpi:"continuous_on (UNIV - {0}) pi" apply(rule continuous_at_imp_continuous_on) + apply rule unfolding pi_def apply(rule continuous_mul) unfolding o_def + apply(rule continuous_at_inv[unfolded o_def]) unfolding continuous_at_vec1_range[unfolded o_def] + apply(rule,rule) apply(rule_tac x=e in exI) apply(rule,assumption,rule,rule) + proof- fix e x y assume "0 < e" "norm (y - x::real^'n) < e" + thus "\norm y - norm x\ < e" using norm_triangle_ineq3[of y x] by auto + qed(auto intro!:continuous_at_id) + def sphere \ "{x::real^'n. norm x = 1}" + have pi:"\x. x \ 0 \ pi x \ sphere" "\x u. u>0 \ pi (u *s x) = pi x" unfolding pi_def sphere_def by auto + + have "0\s" using assms(2) and centre_in_cball[of 0 1] by auto + have front_smul:"\x\frontier s. \u\0. u *s x \ s \ u \ 1" proof(rule,rule,rule) + fix x u assume x:"x\frontier s" and "(0::real)\u" + hence "x\0" using `0\frontier s` by auto + obtain v where v:"0 \ v" "v *s x \ frontier s" "\w>v. w *s x \ s" + using compact_frontier_line_lemma[OF assms(1) `0\s` `x\0`] by auto + have "v=1" apply(rule ccontr) unfolding neq_iff apply(erule disjE) proof- + assume "v<1" thus False using v(3)[THEN spec[where x=1]] using x and fs by auto next + assume "v>1" thus False using assms(3)[THEN bspec[where x="v *s x"], THEN spec[where x="inverse v"]] + using v and x and fs unfolding inverse_less_1_iff by auto qed + show "u *s x \ s \ u \ 1" apply rule using v(3)[unfolded `v=1`, THEN spec[where x=u]] proof- + assume "u\1" thus "u *s x \ s" apply(cases "u=1") + using assms(3)[THEN bspec[where x=x], THEN spec[where x=u]] using `0\u` and x and fs by auto qed auto qed + + have "\surf. homeomorphism (frontier s) sphere pi surf" + apply(rule homeomorphism_compact) apply(rule compact_frontier[OF assms(1)]) + apply(rule continuous_on_subset[OF contpi]) defer apply(rule set_ext,rule) + unfolding inj_on_def prefer 3 apply(rule,rule,rule) + proof- fix x assume "x\pi ` frontier s" then obtain y where "y\frontier s" "x = pi y" by auto + thus "x \ sphere" using pi(1)[of y] and `0 \ frontier s` by auto + next fix x assume "x\sphere" hence "norm x = 1" "x\0" unfolding sphere_def by auto + then obtain u where "0 \ u" "u *s x \ frontier s" "\v>u. v *s x \ s" + using compact_frontier_line_lemma[OF assms(1) `0\s`, of x] by auto + thus "x \ pi ` frontier s" unfolding image_iff le_less pi_def apply(rule_tac x="u *s x" in bexI) using `norm x = 1` `0\frontier s` by auto + next fix x y assume as:"x \ frontier s" "y \ frontier s" "pi x = pi y" + hence xys:"x\s" "y\s" using fs by auto + from as(1,2) have nor:"norm x \ 0" "norm y \ 0" using `0\frontier s` by auto + from nor have x:"x = norm x *s ((inverse (norm y)) *s y)" unfolding as(3)[unfolded pi_def, THEN sym] by auto + from nor have y:"y = norm y *s ((inverse (norm x)) *s x)" unfolding as(3)[unfolded pi_def] by auto + have "0 \ norm y * inverse (norm x)" "0 \ norm x * inverse (norm y)" + unfolding divide_inverse[THEN sym] apply(rule_tac[!] divide_nonneg_pos) using nor by auto + hence "norm x = norm y" apply(rule_tac ccontr) unfolding neq_iff + using x y and front_smul[THEN bspec, OF as(1), THEN spec[where x="norm y * (inverse (norm x))"]] + using front_smul[THEN bspec, OF as(2), THEN spec[where x="norm x * (inverse (norm y))"]] + using xys nor by(auto simp add:field_simps divide_le_eq_1 divide_inverse[THEN sym]) + thus "x = y" apply(subst injpi[THEN sym]) using as(3) by auto + qed(insert `0 \ frontier s`, auto) + then obtain surf where surf:"\x\frontier s. surf (pi x) = x" "pi ` frontier s = sphere" "continuous_on (frontier s) pi" + "\y\sphere. pi (surf y) = y" "surf ` sphere = frontier s" "continuous_on sphere surf" unfolding homeomorphism_def by auto + + have cont_surfpi:"continuous_on (UNIV - {0}) (surf \ pi)" apply(rule continuous_on_compose, rule contpi) + apply(rule continuous_on_subset[of sphere], rule surf(6)) using pi(1) by auto + + { fix x assume as:"x \ cball (0::real^'n) 1" + have "norm x *s surf (pi x) \ s" proof(cases "x=0 \ norm x = 1") + case False hence "pi x \ sphere" "norm x < 1" using pi(1)[of x] as by(auto simp add: dist_def) + thus ?thesis apply(rule_tac assms(3)[rule_format, THEN DiffD1]) + apply(rule_tac fs[unfolded subset_eq, rule_format]) + unfolding surf(5)[THEN sym] by auto + next case True thus ?thesis apply rule defer unfolding pi_def apply(rule fs[unfolded subset_eq, rule_format]) + unfolding surf(5)[unfolded sphere_def, THEN sym] using `0\s` by auto qed } note hom = this + + { fix x assume "x\s" + hence "x \ (\x. norm x *s surf (pi x)) ` cball 0 1" proof(cases "x=0") + case True show ?thesis unfolding image_iff True apply(rule_tac x=0 in bexI) by auto + next let ?a = "inverse (norm (surf (pi x)))" + case False hence invn:"inverse (norm x) \ 0" by auto + from False have pix:"pi x\sphere" using pi(1) by auto + hence "pi (surf (pi x)) = pi x" apply(rule_tac surf(4)[rule_format]) by assumption + hence **:"norm x *s (?a *s surf (pi x)) = x" apply(rule_tac vector_mul_lcancel_imp[OF invn]) unfolding pi_def by auto + hence *:"?a * norm x > 0" and"?a > 0" "?a \ 0" using surf(5) `0\frontier s` apply - + apply(rule_tac mult_pos_pos) using False[unfolded zero_less_norm_iff[THEN sym]] by auto + have "norm (surf (pi x)) \ 0" using ** False by auto + hence "norm x = norm ((?a * norm x) *s surf (pi x))" + unfolding norm_mul abs_mult abs_norm_cancel abs_of_pos[OF `?a > 0`] by auto + moreover have "pi x = pi ((inverse (norm (surf (pi x))) * norm x) *s surf (pi x))" + unfolding pi(2)[OF *] surf(4)[rule_format, OF pix] .. + moreover have "surf (pi x) \ frontier s" using surf(5) pix by auto + hence "dist 0 (inverse (norm (surf (pi x))) *s x) \ 1" unfolding dist_def + using ** and * using front_smul[THEN bspec[where x="surf (pi x)"], THEN spec[where x="norm x * ?a"]] + using False `x\s` by(auto simp add:field_simps) + ultimately show ?thesis unfolding image_iff apply(rule_tac x="inverse (norm (surf(pi x))) *s x" in bexI) + apply(subst injpi[THEN sym]) unfolding norm_mul abs_mult abs_norm_cancel abs_of_pos[OF `?a > 0`] + unfolding pi(2)[OF `?a > 0`] by auto + qed } note hom2 = this + + show ?thesis apply(subst homeomorphic_sym) apply(rule homeomorphic_compact[where f="\x. norm x *s surf (pi x)"]) + apply(rule compact_cball) defer apply(rule set_ext, rule, erule imageE, drule hom) + prefer 4 apply(rule continuous_at_imp_continuous_on, rule) apply(rule_tac [3] hom2) proof- + fix x::"real^'n" assume as:"x \ cball 0 1" + thus "continuous (at x) (\x. norm x *s surf (pi x))" proof(cases "x=0") + case False thus ?thesis apply(rule_tac continuous_mul, rule_tac continuous_at_vec1_norm[THEN spec]) + using cont_surfpi unfolding continuous_on_eq_continuous_at[OF open_delete[OF open_UNIV]] o_def by auto + next guess a using UNIV_witness[where 'a = 'n] .. + obtain B where B:"\x\s. norm x \ B" using compact_imp_bounded[OF assms(1)] unfolding bounded_def by auto + hence "B > 0" using assms(2) unfolding subset_eq apply(erule_tac x="basis a" in ballE) defer apply(erule_tac x="basis a" in ballE) + unfolding Ball_def mem_cball dist_def by (auto simp add: norm_basis[unfolded One_nat_def]) + case True show ?thesis unfolding True continuous_at Lim_at apply(rule,rule) apply(rule_tac x="e / B" in exI) + apply(rule) apply(rule divide_pos_pos) prefer 3 apply(rule,rule,erule conjE) + unfolding norm_0 vector_smult_lzero dist_def diff_0_right norm_mul abs_norm_cancel proof- + fix e and x::"real^'n" assume as:"norm x < e / B" "0 < norm x" "0 frontier s" using pi(1)[of x] unfolding surf(5)[THEN sym] by auto + hence "norm (surf (pi x)) \ B" using B fs by auto + hence "norm x * norm (surf (pi x)) \ norm x * B" using as(2) by auto + also have "\ < e / B * B" apply(rule mult_strict_right_mono) using as(1) `B>0` by auto + also have "\ = e" using `B>0` by auto + finally show "norm x * norm (surf (pi x)) < e" by assumption + qed(insert `B>0`, auto) qed + next { fix x assume as:"surf (pi x) = 0" + have "x = 0" proof(rule ccontr) + assume "x\0" hence "pi x \ sphere" using pi(1) by auto + hence "surf (pi x) \ frontier s" using surf(5) by auto + thus False using `0\frontier s` unfolding as by simp qed + } note surf_0 = this + show "inj_on (\x. norm x *s surf (pi x)) (cball 0 1)" unfolding inj_on_def proof(rule,rule,rule) + fix x y assume as:"x \ cball 0 1" "y \ cball 0 1" "norm x *s surf (pi x) = norm y *s surf (pi y)" + thus "x=y" proof(cases "x=0 \ y=0") + case True thus ?thesis using as by(auto elim: surf_0) next + case False + hence "pi (surf (pi x)) = pi (surf (pi y))" using as(3) + using pi(2)[of "norm x" "surf (pi x)"] pi(2)[of "norm y" "surf (pi y)"] by auto + moreover have "pi x \ sphere" "pi y \ sphere" using pi(1) False by auto + ultimately have *:"pi x = pi y" using surf(4)[THEN bspec[where x="pi x"]] surf(4)[THEN bspec[where x="pi y"]] by auto + moreover have "norm x = norm y" using as(3)[unfolded *] using False by(auto dest:surf_0) + ultimately show ?thesis using injpi by auto qed qed + qed auto qed + +lemma homeomorphic_convex_compact_lemma: fixes s::"(real^'n::finite) set" + assumes "convex s" "compact s" "cball 0 1 \ s" + shows "s homeomorphic (cball (0::real^'n) 1)" + apply(rule starlike_compact_projective[OF assms(2-3)]) proof(rule,rule,rule,erule conjE) + fix x u assume as:"x \ s" "0 \ u" "u < (1::real)" + hence "u *s x \ interior s" unfolding interior_def mem_Collect_eq + apply(rule_tac x="ball (u *s x) (1 - u)" in exI) apply(rule, rule open_ball) + unfolding centre_in_ball apply rule defer apply(rule) unfolding mem_ball proof- + fix y assume "dist (u *s x) y < 1 - u" + hence "inverse (1 - u) *s (y - u *s x) \ s" + using assms(3) apply(erule_tac subsetD) unfolding mem_cball dist_sym dist_def + unfolding group_add_class.diff_0 group_add_class.diff_0_right norm_minus_cancel norm_mul + apply (rule mult_left_le_imp_le[of "1 - u"]) + unfolding class_semiring.mul_a using `u<1` by auto + thus "y \ s" using assms(1)[unfolded convex_def, rule_format, of "inverse(1 - u) *s (y - u *s x)" x "1 - u" u] + using as unfolding vector_smult_assoc by auto qed auto + thus "u *s x \ s - frontier s" using frontier_def and interior_subset by auto qed + +lemma homeomorphic_convex_compact_cball: fixes e::real and s::"(real^'n::finite) set" + assumes "convex s" "compact s" "interior s \ {}" "0 < e" + shows "s homeomorphic (cball (b::real^'n::finite) e)" +proof- obtain a where "a\interior s" using assms(3) by auto + then obtain d where "d>0" and d:"cball a d \ s" unfolding mem_interior_cball by auto + let ?d = "inverse d" and ?n = "0::real^'n" + have "cball ?n 1 \ (\x. inverse d *s (x - a)) ` s" + apply(rule, rule_tac x="d *s x + a" in image_eqI) defer + apply(rule d[unfolded subset_eq, rule_format]) using `d>0` unfolding mem_cball dist_def + by(auto simp add: mult_right_le_one_le) + hence "(\x. inverse d *s (x - a)) ` s homeomorphic cball ?n 1" + using homeomorphic_convex_compact_lemma[of "(\x. ?d *s -a + ?d *s x) ` s", OF convex_affinity compact_affinity] + using assms(1,2) by(auto simp add: uminus_add_conv_diff) + thus ?thesis apply(rule_tac homeomorphic_trans[OF _ homeomorphic_balls(2)[of 1 _ ?n]]) + apply(rule homeomorphic_trans[OF homeomorphic_affinity[of "?d" s "?d *s -a"]]) + using `d>0` `e>0` by(auto simp add: uminus_add_conv_diff) qed + +lemma homeomorphic_convex_compact: fixes s::"(real^'n::finite) set" and t::"(real^'n) set" + assumes "convex s" "compact s" "interior s \ {}" + "convex t" "compact t" "interior t \ {}" + shows "s homeomorphic t" + using assms by(meson zero_less_one homeomorphic_trans homeomorphic_convex_compact_cball homeomorphic_sym) + +subsection {* Epigraphs of convex functions. *} + +definition "epigraph s (f::real^'n \ real) = {xy. fstcart xy \ s \ f(fstcart xy) \ dest_vec1 (sndcart xy)}" + +lemma mem_epigraph: "(pastecart x (vec1 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 forall_pastecart epigraph_def + unfolding mem_Collect_eq fstcart_pastecart sndcart_pastecart sndcart_add sndcart_cmul fstcart_add fstcart_cmul + unfolding Ball_def[symmetric] unfolding dest_vec1_add dest_vec1_cmul + apply(subst forall_dest_vec1[THEN sym])+ by(meson real_le_refl real_le_trans add_mono mult_left_mono) + +lemma convex_epigraphI: assumes "convex_on s f" "convex s" + shows "convex(epigraph s f)" using assms unfolding convex_epigraph by auto + +lemma convex_epigraph_convex: "convex s \ (convex_on s f \ convex(epigraph s f))" + using convex_epigraph by auto + +subsection {* Use this to derive general bound property of convex function. *} + +lemma forall_of_pastecart: + "(\p. P (\x. fstcart (p x)) (\x. sndcart (p x))) \ (\x y. P x y)" apply meson + apply(erule_tac x="\a. pastecart (x a) (y a)" in allE) unfolding o_def by auto + +lemma forall_of_pastecart': + "(\p. P (fstcart p) (sndcart p)) \ (\x y. P x y)" apply meson + apply(erule_tac x="pastecart x y" in allE) unfolding o_def by auto + +lemma forall_of_dest_vec1: "(\v. P (\x. dest_vec1 (v x))) \ (\x. P x)" + apply rule apply rule apply(erule_tac x="(vec1 \ x)" in allE) unfolding o_def vec1_dest_vec1 by auto + +lemma forall_of_dest_vec1': "(\v. P (dest_vec1 v)) \ (\x. P x)" + apply rule apply rule apply(erule_tac x="(vec1 x)" in allE) defer apply rule + apply(erule_tac x="dest_vec1 v" in allE) unfolding o_def vec1_dest_vec1 by auto + +lemma convex_on: + assumes "convex s" + shows "convex_on s f \ (\k u x. (\i\{1..k::nat}. 0 \ u i \ x i \ s) \ setsum u {1..k} = 1 \ + f (setsum (\i. u i *s x i) {1..k} ) \ setsum (\i. u i * f(x i)) {1..k} ) " + unfolding convex_epigraph_convex[OF assms] convex epigraph_def Ball_def mem_Collect_eq + unfolding sndcart_setsum[OF finite_atLeastAtMost] fstcart_setsum[OF finite_atLeastAtMost] dest_vec1_setsum[OF finite_atLeastAtMost] + unfolding fstcart_pastecart sndcart_pastecart sndcart_add sndcart_cmul fstcart_add fstcart_cmul + unfolding dest_vec1_add dest_vec1_cmul apply(subst forall_of_pastecart)+ apply(subst forall_of_dest_vec1)+ apply rule + using assms[unfolded convex] apply simp apply(rule,rule,rule) + apply(erule_tac x=k in allE, erule_tac x=u in allE, erule_tac x=x in allE) apply rule apply rule apply rule defer + apply(rule_tac j="\i = 1..k. u i * f (x i)" in real_le_trans) + defer apply(rule setsum_mono) apply(erule conjE)+ apply(erule_tac x=i in allE)apply(rule mult_left_mono) + using assms[unfolded convex] by auto + +subsection {* Convexity of general and special intervals. *} + +lemma is_interval_convex: assumes "is_interval_new s" shows "convex s" + unfolding convex_def apply(rule,rule,rule,rule,rule,rule,rule) proof- + fix x y u v assume as:"x \ s" "y \ s" "0 \ u" "0 \ v" "u + v = (1::real)" + hence *:"u = 1 - v" "1 - v \ 0" and **:"v = 1 - u" "1 - u \ 0" by auto + { fix a b assume "\ b \ u * a + v * b" + hence "u * a < (1 - v) * b" unfolding not_le using as(4) by(auto simp add: field_simps) + hence "a < b" unfolding * using as(4) *(2) apply(rule_tac mult_left_less_imp_less[of "1 - v"]) by(auto simp add: field_simps) + hence "a \ u * a + v * b" unfolding * using as(4) by (auto simp add: field_simps intro!:mult_right_mono) + } moreover + { fix a b assume "\ u * a + v * b \ a" + hence "v * b > (1 - u) * a" unfolding not_le using as(4) by(auto simp add: field_simps) + hence "a < b" unfolding * using as(4) apply(rule_tac mult_left_less_imp_less) by(auto simp add: ring_simps) + hence "u * a + v * b \ b" unfolding ** using **(2) as(3) by(auto simp add: field_simps intro!:mult_right_mono) } + ultimately show "u *s x + v *s y \ s" apply- apply(rule assms[unfolded is_interval_new_def, rule_format, OF as(1,2)]) + using as(3-) dimindex_ge_1 apply- by(auto simp add: vector_component) qed + +lemma is_interval_connected: "is_interval_new s \ connected s" + using is_interval_convex convex_connected by auto + +lemma convex_interval: "convex {a .. b}" "convex {a<.. (\a\s. \b\s. \ x. dest_vec1 a \ dest_vec1 x \ dest_vec1 x \ dest_vec1 b \ x \ s)" + unfolding is_interval_new_def dest_vec1_def forall_1 by auto + +lemma is_interval_connected_1: "is_interval_new s \ connected (s::(real^1) set)" + apply(rule, rule is_interval_connected, assumption) unfolding is_interval_1 + apply(rule,rule,rule,rule,erule conjE,rule ccontr) proof- + fix a b x assume as:"connected s" "a \ s" "b \ s" "dest_vec1 a \ dest_vec1 x" "dest_vec1 x \ dest_vec1 b" "x\s" + hence *:"dest_vec1 a < dest_vec1 x" "dest_vec1 x < dest_vec1 b" apply(rule_tac [!] ccontr) unfolding not_less by auto + let ?halfl = "{z. basis 1 \ z < dest_vec1 x} " and ?halfr = "{z. basis 1 \ z > dest_vec1 x} " + { fix y assume "y \ s" have "y \ ?halfr \ ?halfl" apply(rule ccontr) + using as(6) `y\s` by (auto simp add: basis_component field_simps dest_vec1_eq[unfolded dest_vec1_def One_nat_def] dest_vec1_def) } + moreover have "a\?halfl" "b\?halfr" using * by (auto simp add: basis_component field_simps dest_vec1_def) + hence "?halfl \ s \ {}" "?halfr \ s \ {}" using as(2-3) by auto + ultimately show False apply(rule_tac notE[OF as(1)[unfolded connected_def]]) + apply(rule_tac x="?halfl" in exI, rule_tac x="?halfr" in exI) + apply(rule, rule open_halfspace_lt, rule, rule open_halfspace_gt) apply(rule, rule, rule ccontr) + by(auto simp add: basis_component field_simps) qed + +lemma is_interval_convex_1: + "is_interval_new s \ convex (s::(real^1) set)" + using is_interval_convex convex_connected is_interval_connected_1 by auto + +lemma convex_connected_1: + "connected s \ convex (s::(real^1) set)" + using is_interval_convex convex_connected is_interval_connected_1 by auto + +subsection {* Another intermediate value theorem formulation. *} + +lemma ivt_increasing_component_on_1: fixes f::"real^1 \ real^'n::finite" + assumes "dest_vec1 a \ dest_vec1 b" "continuous_on {a .. b} f" "(f a)$k \ y" "y \ (f b)$k" + shows "\x\{a..b}. (f x)$k = y" +proof- have "f a \ f ` {a..b}" "f b \ f ` {a..b}" apply(rule_tac[!] imageI) + using assms(1) by(auto simp add: vector_less_eq_def dest_vec1_def) + thus ?thesis using connected_ivt_component[of "f ` {a..b}" "f a" "f b" k y] + using connected_continuous_image[OF assms(2) convex_connected[OF convex_interval(1)]] + using assms by(auto intro!: imageI) qed + +lemma ivt_increasing_component_1: fixes f::"real^1 \ real^'n::finite" + assumes "dest_vec1 a \ dest_vec1 b" + "\x\{a .. b}. continuous (at x) f" "f a$k \ y" "y \ f b$k" + shows "\x\{a..b}. (f x)$k = y" + apply(rule ivt_increasing_component_on_1) using assms using continuous_at_imp_continuous_on by auto + +lemma ivt_decreasing_component_on_1: fixes f::"real^1 \ real^'n::finite" + assumes "dest_vec1 a \ dest_vec1 b" "continuous_on {a .. b} f" "(f b)$k \ y" "y \ (f a)$k" + shows "\x\{a..b}. (f x)$k = y" + apply(subst neg_equal_iff_equal[THEN sym]) unfolding vector_uminus_component[THEN sym] + apply(rule ivt_increasing_component_on_1) using assms using continuous_on_neg + by(auto simp add:vector_uminus_component) + +lemma ivt_decreasing_component_1: fixes f::"real^1 \ real^'n::finite" + assumes "dest_vec1 a \ dest_vec1 b" "\x\{a .. b}. continuous (at x) f" "f b$k \ y" "y \ f a$k" + shows "\x\{a..b}. (f x)$k = y" + apply(rule ivt_decreasing_component_on_1) using assms using continuous_at_imp_continuous_on by auto + +subsection {* A bound within a convex hull, and so an interval. *} + +lemma convex_on_convex_hull_bound: + assumes "convex_on (convex hull s) f" "\x\s. f x \ b" + shows "\x\ convex hull s. f x \ b" proof + fix x assume "x\convex hull s" + then obtain k u v where obt:"\i\{1..k::nat}. 0 \ u i \ v i \ s" "setsum u {1..k} = 1" "(\i = 1..k. u i *s v i) = x" + unfolding convex_hull_indexed mem_Collect_eq by auto + have "(\i = 1..k. u i * f (v i)) \ b" using setsum_mono[of "{1..k}" "\i. u i * f (v i)" "\i. u i * b"] + unfolding setsum_left_distrib[THEN sym] obt(2) mult_1 apply(drule_tac meta_mp) apply(rule mult_left_mono) + using assms(2) obt(1) by auto + thus "f x \ b" using assms(1)[unfolded convex_on[OF convex_convex_hull], rule_format, of k u v] + unfolding obt(2-3) using obt(1) and hull_subset[unfolded subset_eq, rule_format, of _ s] by auto qed + +lemma unit_interval_convex_hull: + "{0::real^'n::finite .. 1} = convex hull {x. \i. (x$i = 0) \ (x$i = 1)}" (is "?int = convex hull ?points") +proof- have 01:"{0,1} \ convex hull ?points" apply rule apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) by auto + { fix n x assume "x\{0::real^'n .. 1}" "n \ CARD('n)" "card {i. x$i \ 0} \ n" + hence "x\convex hull ?points" proof(induct n arbitrary: x) + case 0 hence "x = 0" apply(subst Cart_eq) apply rule by auto + thus "x\convex hull ?points" using 01 by auto + next + case (Suc n) show "x\convex hull ?points" proof(cases "{i. x$i \ 0} = {}") + case True hence "x = 0" unfolding Cart_eq by auto + thus "x\convex hull ?points" using 01 by auto + next + case False def xi \ "Min ((\i. x$i) ` {i. x$i \ 0})" + have "xi \ (\i. x$i) ` {i. x$i \ 0}" unfolding xi_def apply(rule Min_in) using False by auto + then obtain i where i':"x$i = xi" "x$i \ 0" by auto + have i:"\j. x$j > 0 \ x$i \ x$j" + unfolding i'(1) xi_def apply(rule_tac Min_le) unfolding image_iff + defer apply(rule_tac x=j in bexI) using i' by auto + have i01:"x$i \ 1" "x$i > 0" using Suc(2)[unfolded mem_interval,rule_format,of i] using i'(2) `x$i \ 0` + by(auto simp add: Cart_lambda_beta) + show ?thesis proof(cases "x$i=1") + case True have "\j\{i. x$i \ 0}. x$j = 1" apply(rule, rule ccontr) unfolding mem_Collect_eq proof- + fix j assume "x $ j \ 0" "x $ j \ 1" + hence j:"x$j \ {0<..<1}" using Suc(2) by(auto simp add: vector_less_eq_def elim!:allE[where x=j]) + hence "x$j \ op $ x ` {i. x $ i \ 0}" by auto + hence "x$j \ x$i" unfolding i'(1) xi_def apply(rule_tac Min_le) by auto + thus False using True Suc(2) j by(auto simp add: vector_less_eq_def elim!:ballE[where x=j]) qed + thus "x\convex hull ?points" apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) + by(auto simp add: Cart_lambda_beta) + next let ?y = "\j. if x$j = 0 then 0 else (x$j - x$i) / (1 - x$i)" + case False hence *:"x = x$i *s (\ j. if x$j = 0 then 0 else 1) + (1 - x$i) *s (\ j. ?y j)" unfolding Cart_eq + by(auto simp add: Cart_lambda_beta vector_add_component vector_smult_component vector_minus_component field_simps) + { fix j have "x$j \ 0 \ 0 \ (x $ j - x $ i) / (1 - x $ i)" "(x $ j - x $ i) / (1 - x $ i) \ 1" + apply(rule_tac divide_nonneg_pos) using i(1)[of j] using False i01 + using Suc(2)[unfolded mem_interval, rule_format, of j] by(auto simp add:field_simps Cart_lambda_beta) + hence "0 \ ?y j \ ?y j \ 1" by auto } + moreover have "i\{j. x$j \ 0} - {j. ((\ j. ?y j)::real^'n) $ j \ 0}" using i01 by(auto simp add: Cart_lambda_beta) + hence "{j. x$j \ 0} \ {j. ((\ j. ?y j)::real^'n::finite) $ j \ 0}" by auto + hence **:"{j. ((\ j. ?y j)::real^'n::finite) $ j \ 0} \ {j. x$j \ 0}" apply - apply rule by(auto simp add: Cart_lambda_beta) + have "card {j. ((\ j. ?y j)::real^'n) $ j \ 0} \ n" using less_le_trans[OF psubset_card_mono[OF _ **] Suc(4)] by auto + ultimately show ?thesis apply(subst *) apply(rule convex_convex_hull[unfolded convex_def, rule_format]) + apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) defer apply(rule Suc(1)) + unfolding mem_interval using i01 Suc(3) by (auto simp add: Cart_lambda_beta) + qed qed qed } note * = this + show ?thesis apply rule defer apply(rule hull_minimal) unfolding subset_eq prefer 3 apply rule + apply(rule_tac n2="CARD('n)" in *) prefer 3 apply(rule card_mono) using 01 and convex_interval(1) prefer 5 apply - apply rule + unfolding mem_interval apply rule unfolding mem_Collect_eq apply(erule_tac x=i in allE) + by(auto simp add: vector_less_eq_def mem_def[of _ convex]) qed + +subsection {* And this is a finite set of vertices. *} + +lemma unit_cube_convex_hull: obtains s where "finite s" "{0 .. 1::real^'n::finite} = convex hull s" + apply(rule that[of "{x::real^'n::finite. \i. x$i=0 \ x$i=1}"]) + apply(rule finite_subset[of _ "(\s. (\ i. if i\s then 1::real else 0)::real^'n::finite) ` UNIV"]) + prefer 3 apply(rule unit_interval_convex_hull) apply rule unfolding mem_Collect_eq proof- + fix x::"real^'n" assume as:"\i. x $ i = 0 \ x $ i = 1" + show "x \ (\s. \ i. if i \ s then 1 else 0) ` UNIV" apply(rule image_eqI[where x="{i. x$i = 1}"]) + unfolding Cart_eq using as by(auto simp add:Cart_lambda_beta) qed auto + +subsection {* Hence any cube (could do any nonempty interval). *} + +lemma cube_convex_hull: + assumes "0 < d" obtains s::"(real^'n::finite) set" where "finite s" "{x - (\ i. d) .. x + (\ i. d)} = convex hull s" proof- + let ?d = "(\ i. d)::real^'n" + have *:"{x - ?d .. x + ?d} = (\y. x - ?d + (2 * d) *s y) ` {0 .. 1}" apply(rule set_ext, rule) + unfolding image_iff defer apply(erule bexE) proof- + fix y assume as:"y\{x - ?d .. x + ?d}" + { fix i::'n have "x $ i \ d + y $ i" "y $ i \ d + x $ i" using as[unfolded mem_interval, THEN spec[where x=i]] + by(auto simp add: vector_component) + hence "1 \ inverse d * (x $ i - y $ i)" "1 \ inverse d * (y $ i - x $ i)" + apply(rule_tac[!] mult_left_le_imp_le[OF _ assms]) unfolding mult_assoc[THEN sym] + using assms by(auto simp add: field_simps right_inverse) + hence "inverse d * (x $ i * 2) \ 2 + inverse d * (y $ i * 2)" + "inverse d * (y $ i * 2) \ 2 + inverse d * (x $ i * 2)" by(auto simp add:field_simps) } + hence "inverse (2 * d) *s (y - (x - ?d)) \ {0..1}" unfolding mem_interval using assms + by(auto simp add: Cart_eq vector_component_simps field_simps) + thus "\z\{0..1}. y = x - ?d + (2 * d) *s z" apply- apply(rule_tac x="inverse (2 * d) *s (y - (x - ?d))" in bexI) + using assms by(auto simp add: Cart_eq vector_less_eq_def Cart_lambda_beta) + next + fix y z assume as:"z\{0..1}" "y = x - ?d + (2*d) *s z" + have "\i. 0 \ d * z $ i \ d * z $ i \ d" using assms as(1)[unfolded mem_interval] apply(erule_tac x=i in allE) + apply rule apply(rule mult_nonneg_nonneg) prefer 3 apply(rule mult_right_le_one_le) + using assms by(auto simp add: vector_component_simps Cart_eq) + thus "y \ {x - ?d..x + ?d}" unfolding as(2) mem_interval apply- apply rule using as(1)[unfolded mem_interval] + apply(erule_tac x=i in allE) using assms by(auto simp add: vector_component_simps Cart_eq) qed + obtain s where "finite s" "{0..1::real^'n} = convex hull s" using unit_cube_convex_hull by auto + thus ?thesis apply(rule_tac that[of "(\y. x - ?d + (2 * d) *s y)` s"]) unfolding * and convex_hull_affinity by auto qed + +subsection {* Bounded convex function on open set is continuous. *} + +lemma convex_on_bounded_continuous: + assumes "open s" "convex_on s f" "\x\s. abs(f x) \ b" + shows "continuous_on s (vec1 o f)" + apply(rule continuous_at_imp_continuous_on) unfolding continuous_at_vec1_range proof(rule,rule,rule) + fix x e assume "x\s" "(0::real) < e" + def B \ "abs b + 1" + have B:"0 < B" "\x. x\s \ abs (f x) \ B" + unfolding B_def defer apply(drule assms(3)[rule_format]) by auto + 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 + 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) + fix y assume as:"norm (y - x) < min (k / 2) (e / (2 * B) * k)" + show "\f y - f x\ < e" proof(cases "y=x") + case False def t \ "k / norm (y - x)" + have "2 < t" "00` by(auto simp add:field_simps) + have "y\s" apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_def + apply(rule order_trans[of _ "2 * norm (x - y)"]) using as by(auto simp add: field_simps norm_minus_commute) + { def w \ "x + t *s (y - x)" + have "w\s" unfolding w_def apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_def + unfolding t_def using `k>0` by(auto simp add: norm_mul simp del: vector_ssub_ldistrib) + have "(1 / t) *s x + - x + ((t - 1) / t) *s x = (1 / t - 1 + (t - 1) / t) *s x" by auto + also have "\ = 0" using `t>0` by(auto simp add:field_simps simp del:vector_sadd_rdistrib) + finally have w:"(1 / t) *s w + ((t - 1) / t) *s x = y" unfolding w_def using False and `t>0` by auto + have "2 * B < e * t" unfolding t_def using `00` and as and False by (auto simp add:field_simps) + hence "(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) + hence th1:"f y - f x < e" apply- apply(rule le_less_trans) defer apply assumption + using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w] + using `0s` `w\s` by(auto simp add:field_simps) } + moreover + { def w \ "x - t *s (y - x)" + have "w\s" unfolding w_def apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_def + unfolding t_def using `k>0` by(auto simp add: norm_mul simp del: vector_ssub_ldistrib) + have "(1 / (1 + t)) *s x + (t / (1 + t)) *s x = (1 / (1 + t) + t / (1 + t)) *s x" by auto + also have "\=x" using `t>0` by (auto simp add:field_simps simp del:vector_sadd_rdistrib) + finally have w:"(1 / (1+t)) *s w + (t / (1 + t)) *s y = x" unfolding w_def using False and `t>0` by auto + have "2 * B < e * t" unfolding t_def using `00` and as and False by (auto simp add:field_simps) + hence *:"(f w - f y) / t < e" using B(2)[OF `w\s`] and B(2)[OF `y\s`] using `t>0` by(auto simp add: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 `0s` `w\s` by (auto simp add:field_simps) + also have "\ = (f w + t * f y) / (1 + t)" using `t>0` unfolding real_divide_def by (auto simp add:field_simps) + also have "\ < e + f y" using `t>0` * `e>0` by(auto simp add:field_simps) + finally have "f x - f y < e" by auto } + ultimately show ?thesis by auto + qed(insert `0y \ cball x e. f y \ b" + shows "\y \ cball x e. abs(f y) \ b + 2 * abs(f x)" + apply(rule) proof(cases "0 \ e") case True + fix y assume y:"y\cball x e" def z \ "2 *s x - y" + have *:"x - (2 *s x - y) = y - x" by vector + have z:"z\cball x e" using y unfolding z_def mem_cball dist_def * by(auto simp add: norm_minus_commute) + have "(1 / 2) *s y + (1 / 2) *s z = x" unfolding z_def by auto + thus "\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) +next case False fix y assume "y\cball x e" + hence "dist x y < 0" using False unfolding mem_cball not_le by auto + thus "\f y\ \ b + 2 * \f x\" using dist_pos_le[of x y] by auto qed + +subsection {* Hence a convex function on an open set is continuous. *} + +lemma convex_on_continuous: + assumes "open (s::(real^'n::finite) set)" "convex_on s f" + shows "continuous_on s (vec1 \ f)" + unfolding continuous_on_eq_continuous_at[OF assms(1)] proof + note dimge1 = dimindex_ge_1[where 'a='n] + fix x assume "x\s" + then obtain e where e:"cball x e \ s" "e>0" using assms(1) unfolding open_contains_cball by auto + def d \ "e / real CARD('n)" + have "0 < d" unfolding d_def using `e>0` dimge1 by(rule_tac divide_pos_pos, auto) + let ?d = "(\ i. d)::real^'n" + obtain c where c:"finite c" "{x - ?d..x + ?d} = convex hull c" using cube_convex_hull[OF `d>0`, of x] by auto + have "x\{x - ?d..x + ?d}" using `d>0` unfolding mem_interval by(auto simp add:vector_component_simps) + hence "c\{}" apply(rule_tac ccontr) using c by(auto simp add:convex_hull_empty) + def k \ "Max (f ` c)" + have "convex_on {x - ?d..x + ?d} f" apply(rule convex_on_subset[OF assms(2)]) + apply(rule subset_trans[OF _ e(1)]) unfolding subset_eq mem_cball proof + fix z assume z:"z\{x - ?d..x + ?d}" + have e:"e = setsum (\i. d) (UNIV::'n set)" unfolding setsum_constant d_def using dimge1 + by (metis card_enum field_simps d_def not_one_le_zero of_nat_le_iff real_eq_of_nat real_of_nat_1) + show "dist x z \ e" unfolding dist_def e apply(rule_tac order_trans[OF norm_le_l1], rule setsum_mono) + using z[unfolded mem_interval] apply(erule_tac x=i in allE) by(auto simp add:field_simps vector_component_simps) qed + hence k:"\y\{x - ?d..x + ?d}. f y \ k" unfolding c(2) apply(rule_tac convex_on_convex_hull_bound) apply assumption + unfolding k_def apply(rule, rule Max_ge) using c(1) by auto + have "d \ e" unfolding d_def apply(rule mult_imp_div_pos_le) using `e>0` dimge1 unfolding mult_le_cancel_left1 using real_dimindex_ge_1 by auto + hence dsube:"cball x d \ cball x e" unfolding subset_eq Ball_def mem_cball by auto + have conv:"convex_on (cball x d) f" apply(rule convex_on_subset, rule convex_on_subset[OF assms(2)]) apply(rule e(1)) using dsube by auto + hence "\y\cball x d. abs (f y) \ k + 2 * abs (f x)" apply(rule_tac convex_bounds_lemma) apply assumption proof + fix y assume y:"y\cball x d" + { fix i::'n have "x $ i - d \ y $ i" "y $ i \ x $ i + d" + using order_trans[OF component_le_norm y[unfolded mem_cball dist_def], of i] by(auto simp add: vector_component) } + thus "f y \ k" apply(rule_tac k[rule_format]) unfolding mem_cball mem_interval dist_def + by(auto simp add: vector_component_simps) qed + hence "continuous_on (ball x d) (vec1 \ f)" apply(rule_tac convex_on_bounded_continuous) + apply(rule open_ball, rule convex_on_subset[OF conv], rule ball_subset_cball) by auto + thus "continuous (at x) (vec1 \ f)" unfolding continuous_on_eq_continuous_at[OF open_ball] using `d>0` by auto qed + +subsection {* Line segments, starlike sets etc. *) +(* Use the same overloading tricks as for intervals, so that *) +(* segment[a,b] is closed and segment(a,b) is open relative to affine hull. *} + +definition "midpoint a b = (inverse (2::real)) *s (a + b)" + +definition "open_segment a b = {(1 - u) *s a + u *s b | u::real. 0 < u \ u < 1}" + +definition "closed_segment a b = {(1 - u) *s a + u *s b | u::real. 0 \ u \ u \ 1}" + +definition "between = (\ (a,b). closed_segment a b)" + +lemmas segment = open_segment_def closed_segment_def + +definition "starlike s \ (\a\s. \x\s. closed_segment a x \ s)" + +lemma midpoint_refl: "midpoint x x = x" + unfolding midpoint_def unfolding vector_add_ldistrib unfolding vector_sadd_rdistrib[THEN sym] by auto + +lemma midpoint_sym: "midpoint a b = midpoint b a" unfolding midpoint_def by auto + +lemma dist_midpoint: + "dist a (midpoint a b) = (dist a b) / 2" (is ?t1) + "dist b (midpoint a b) = (dist a b) / 2" (is ?t2) + "dist (midpoint a b) a = (dist a b) / 2" (is ?t3) + "dist (midpoint a b) b = (dist a b) / 2" (is ?t4) +proof- + have *: "\x y::real^'n::finite. 2 *s x = - y \ norm x = (norm y) / 2" unfolding equation_minus_iff by auto + have **:"\x y::real^'n::finite. 2 *s x = y \ norm x = (norm y) / 2" by auto + show ?t1 unfolding midpoint_def dist_def apply (rule **) by(auto,vector) + show ?t2 unfolding midpoint_def dist_def apply (rule *) by(auto,vector) + show ?t3 unfolding midpoint_def dist_def apply (rule *) by(auto,vector) + show ?t4 unfolding midpoint_def dist_def apply (rule **) by(auto,vector) qed + +lemma midpoint_eq_endpoint: + "midpoint a b = a \ a = (b::real^'n::finite)" + "midpoint a b = b \ a = b" + unfolding dist_eq_0[THEN sym] dist_midpoint by auto + +lemma convex_contains_segment: + "convex s \ (\a\s. \b\s. closed_segment a b \ s)" + unfolding convex_alt closed_segment_def by auto + +lemma convex_imp_starlike: + "convex s \ s \ {} \ starlike s" + unfolding convex_contains_segment starlike_def by auto + +lemma segment_convex_hull: + "closed_segment a b = convex hull {a,b}" proof- + have *:"\x. {x} \ {}" by auto + have **:"\u v. u + v = 1 \ u = 1 - (v::real)" by auto + show ?thesis unfolding segment convex_hull_insert[OF *] convex_hull_singleton apply(rule set_ext) + unfolding mem_Collect_eq apply(rule,erule exE) + apply(rule_tac x="1 - u" in exI) apply rule defer apply(rule_tac x=u in exI) defer + apply(erule exE, (erule conjE)?)+ apply(rule_tac x="1 - u" in exI) unfolding ** by auto qed + +lemma convex_segment: "convex (closed_segment a b)" + unfolding segment_convex_hull by(rule convex_convex_hull) + +lemma ends_in_segment: "a \ closed_segment a b" "b \ closed_segment a b" + unfolding segment_convex_hull apply(rule_tac[!] hull_subset[unfolded subset_eq, rule_format]) by auto + +lemma segment_furthest_le: + assumes "x \ closed_segment a b" shows "norm(y - x) \ norm(y - a) \ norm(y - x) \ norm(y - b)" proof- + obtain z where "z\{a, b}" "norm (x - y) \ norm (z - y)" using simplex_furthest_le[of "{a, b}" y] + using assms[unfolded segment_convex_hull] by auto + thus ?thesis by(auto simp add:norm_minus_commute) qed + +lemma segment_bound: + assumes "x \ closed_segment a b" + shows "norm(x - a) \ norm(b - a)" "norm(x - b) \ norm(b - a)" + using segment_furthest_le[OF assms, of a] + using segment_furthest_le[OF assms, of b] + by (auto simp add:norm_minus_commute) + +lemma segment_refl:"closed_segment a a = {a}" unfolding segment by auto + +lemma between_mem_segment: "between (a,b) x \ x \ closed_segment a b" + unfolding between_def mem_def by auto + +lemma between:"between (a,b) (x::real^'n::finite) \ dist a b = (dist a x) + (dist x b)" +proof(cases "a = b") + case True thus ?thesis unfolding between_def split_conv mem_def[of x, symmetric] + by(auto simp add:segment_refl dist_sym) next + case False hence Fal:"norm (a - b) \ 0" and Fal2: "norm (a - b) > 0" by auto + have *:"\u. a - ((1 - u) *s a + u *s b) = u *s (a - b)" by auto + show ?thesis unfolding between_def split_conv mem_def[of x, symmetric] closed_segment_def mem_Collect_eq + apply rule apply(erule exE, (erule conjE)+) apply(subst dist_triangle_eq) proof- + fix u assume as:"x = (1 - u) *s a + u *s b" "0 \ u" "u \ 1" + hence *:"a - x = u *s (a - b)" "x - b = (1 - u) *s (a - b)" + unfolding as(1) by(auto simp add:field_simps) + show "norm (a - x) *s (x - b) = norm (x - b) *s (a - x)" + unfolding norm_minus_commute[of x a] * norm_mul Cart_eq using as(2,3) + by(auto simp add: vector_component_simps field_simps) + next assume as:"dist a b = dist a x + dist x b" + have "norm (a - x) / norm (a - b) \ 1" unfolding divide_le_eq_1_pos[OF Fal2] unfolding as[unfolded dist_def] norm_ge_zero by auto + thus "\u. x = (1 - u) *s a + u *s b \ 0 \ u \ u \ 1" apply(rule_tac x="dist a x / dist a b" in exI) + unfolding dist_def Cart_eq apply- apply rule defer apply(rule, rule divide_nonneg_pos) prefer 4 proof rule + fix i::'n have "((1 - norm (a - x) / norm (a - b)) *s a + (norm (a - x) / norm (a - b)) *s b) $ i = + ((norm (a - b) - norm (a - x)) * (a $ i) + norm (a - x) * (b $ i)) / norm (a - b)" + using Fal by(auto simp add:vector_component_simps field_simps) + also have "\ = x$i" apply(rule divide_eq_imp[OF Fal]) + unfolding as[unfolded dist_def] using as[unfolded dist_triangle_eq Cart_eq,rule_format, of i] + by(auto simp add:field_simps vector_component_simps) + finally show "x $ i = ((1 - norm (a - x) / norm (a - b)) *s a + (norm (a - x) / norm (a - b)) *s b) $ i" by auto + qed(insert Fal2, auto) qed qed + +lemma between_midpoint: + "between (a,b) (midpoint a b)" (is ?t1) + "between (b,a) (midpoint a b)" (is ?t2) +proof- have *:"\x y z. x = (1/2::real) *s z \ y = (1/2) *s z \ norm z = norm x + norm y" by auto + show ?t1 ?t2 unfolding between midpoint_def dist_def apply(rule_tac[!] *) + by(auto simp add:field_simps Cart_eq vector_component_simps) qed + +lemma between_mem_convex_hull: + "between (a,b) x \ x \ convex hull {a,b}" + unfolding between_mem_segment segment_convex_hull .. + +subsection {* Shrinking towards the interior of a convex set. *} + +lemma mem_interior_convex_shrink: + assumes "convex s" "c \ interior s" "x \ s" "0 < e" "e \ 1" + shows "x - e *s (x - c) \ interior s" +proof- obtain d where "d>0" and d:"ball c d \ s" using assms(2) unfolding mem_interior by auto + show ?thesis unfolding mem_interior apply(rule_tac x="e*d" in exI) + apply(rule) defer unfolding subset_eq Ball_def mem_ball proof(rule,rule) + fix y assume as:"dist (x - e *s (x - c)) y < e * d" + have *:"y = (1 - (1 - e)) *s ((1 / e) *s y - ((1 - e) / e) *s x) + (1 - e) *s x" using `e>0` by auto + have "dist c ((1 / e) *s y - ((1 - e) / e) *s x) = abs(1/e) * norm (e *s c - y + (1 - e) *s x)" + unfolding dist_def unfolding norm_mul[THEN sym] apply(rule norm_eqI) using `e>0` + by(auto simp add:vector_component_simps Cart_eq field_simps) + also have "\ = abs(1/e) * norm (x - e *s (x - c) - y)" by(auto intro!:norm_eqI) + also have "\ < d" using as[unfolded dist_def] and `e>0` + by(auto simp add:pos_divide_less_eq[OF `e>0`] real_mult_commute) + finally show "y \ s" apply(subst *) apply(rule assms(1)[unfolded convex_alt,rule_format]) + apply(rule d[unfolded subset_eq,rule_format]) unfolding mem_ball using assms(3-5) by auto + qed(rule mult_pos_pos, insert `e>0` `d>0`, auto) qed + +lemma mem_interior_closure_convex_shrink: + assumes "convex s" "c \ interior s" "x \ closure s" "0 < e" "e \ 1" + shows "x - e *s (x - c) \ interior s" +proof- obtain d where "d>0" and d:"ball c d \ s" using assms(2) unfolding mem_interior by auto + have "\y\s. norm (y - x) * (1 - e) < e * d" proof(cases "x\s") + case True thus ?thesis using `e>0` `d>0` by(rule_tac bexI[where x=x], auto intro!: mult_pos_pos) next + case False hence x:"x islimpt s" using assms(3)[unfolded closure_def] by auto + show ?thesis proof(cases "e=1") + case True obtain y where "y\s" "y \ x" "dist y x < 1" + using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto + thus ?thesis apply(rule_tac x=y in bexI) unfolding True using `d>0` by auto next + case False hence "0 < e * d / (1 - e)" and *:"1 - e > 0" + using `e\1` `e>0` `d>0` by(auto intro!:mult_pos_pos divide_pos_pos) + 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 + thus ?thesis apply(rule_tac x=y in bexI) unfolding dist_def using pos_less_divide_eq[OF *] by auto qed qed + then obtain y where "y\s" and y:"norm (y - x) * (1 - e) < e * d" by auto + def z \ "c + ((1 - e) / e) *s (x - y)" + have *:"x - e *s (x - c) = y - e *s (y - z)" unfolding z_def using `e>0` by auto + have "z\interior s" apply(rule subset_interior[OF d,unfolded subset_eq,rule_format]) + unfolding interior_open[OF open_ball] mem_ball z_def dist_def using y and assms(4,5) + by(auto simp del:vector_ssub_ldistrib simp add:field_simps norm_minus_commute) + thus ?thesis unfolding * apply - apply(rule mem_interior_convex_shrink) + using assms(1,4-5) `y\s` by auto qed + +subsection {* Some obvious but surprisingly hard simplex lemmas. *} + +lemma simplex: + assumes "finite s" "0 \ s" + shows "convex hull (insert 0 s) = { y. (\u. (\x\s. 0 \ u x) \ setsum u s \ 1 \ setsum (\x. u x *s x) s = y)}" + unfolding convex_hull_finite[OF finite.insertI[OF assms(1)]] apply(rule set_ext, rule) unfolding mem_Collect_eq + apply(erule_tac[!] exE) apply(erule_tac[!] conjE)+ unfolding setsum_clauses(2)[OF assms(1)] + apply(rule_tac x=u in exI) defer apply(rule_tac x="\x. if x = 0 then 1 - setsum u s else u x" in exI) using assms(2) + unfolding if_smult and setsum_delta_notmem[OF assms(2)] by auto + +lemma std_simplex: + "convex hull (insert 0 { basis i | i. i\UNIV}) = + {x::real^'n::finite . (\i. 0 \ x$i) \ setsum (\i. x$i) UNIV \ 1 }" (is "convex hull (insert 0 ?p) = ?s") +proof- let ?D = "UNIV::'n set" + have "0\?p" by(auto simp add: basis_nonzero) + have "{(basis i)::real^'n |i. i \ ?D} = basis ` ?D" by auto + note sumbas = this setsum_reindex[OF basis_inj, unfolded o_def] + show ?thesis unfolding simplex[OF finite_stdbasis `0\?p`] apply(rule set_ext) unfolding mem_Collect_eq apply rule + apply(erule exE, (erule conjE)+) apply(erule_tac[2] conjE)+ proof- + fix x::"real^'n" and u assume as: "\x\{basis i |i. i \?D}. 0 \ u x" "setsum u {basis i |i. i \ ?D} \ 1" "(\x\{basis i |i. i \?D}. u x *s x) = x" + have *:"\i. u (basis i) = x$i" using as(3) unfolding sumbas and basis_expansion_unique by auto + hence **:"setsum u {basis i |i. i \ ?D} = setsum (op $ x) ?D" unfolding sumbas by(rule_tac setsum_cong, auto) + show " (\i. 0 \ x $ i) \ setsum (op $ x) ?D \ 1" apply - proof(rule,rule) + fix i::'n show "0 \ x$i" unfolding *[rule_format,of i,THEN sym] apply(rule_tac as(1)[rule_format]) by auto + qed(insert as(2)[unfolded **], auto) + next fix x::"real^'n" assume as:"\i. 0 \ x $ i" "setsum (op $ x) ?D \ 1" + show "\u. (\x\{basis i |i. i \ ?D}. 0 \ u x) \ setsum u {basis i |i. i \ ?D} \ 1 \ (\x\{basis i |i. i \ ?D}. u x *s x) = x" + apply(rule_tac x="\y. y \ x" in exI) apply(rule,rule) unfolding mem_Collect_eq apply(erule exE) using as(1) apply(erule_tac x=i in allE) + unfolding sumbas using as(2) and basis_expansion_unique by(auto simp add:dot_basis) qed qed + +lemma interior_std_simplex: + "interior (convex hull (insert 0 { basis i| i. i\UNIV})) = + {x::real^'n::finite. (\i. 0 < x$i) \ setsum (\i. x$i) UNIV < 1 }" + apply(rule set_ext) unfolding mem_interior std_simplex unfolding subset_eq mem_Collect_eq Ball_def mem_ball + unfolding Ball_def[symmetric] apply rule apply(erule exE, (erule conjE)+) defer apply(erule conjE) proof- + fix x::"real^'n" and e assume "0xa. dist x xa < e \ (\x. 0 \ xa $ x) \ setsum (op $ xa) UNIV \ 1" + show "(\xa. 0 < x $ xa) \ setsum (op $ x) UNIV < 1" apply(rule,rule) proof- + fix i::'n show "0 < x $ i" using as[THEN spec[where x="x - (e / 2) *s basis i"]] and `e>0` + unfolding dist_def by(auto simp add: norm_basis vector_component_simps basis_component elim:allE[where x=i]) + next guess a using UNIV_witness[where 'a='n] .. + have **:"dist x (x + (e / 2) *s basis a) < e" using `e>0` and norm_basis[of a] + unfolding dist_def by(auto simp add: vector_component_simps basis_component intro!: mult_strict_left_mono_comm) + have "\i. (x + (e / 2) *s basis a) $ i = x$i + (if i = a then e/2 else 0)" by(auto simp add:vector_component_simps) + hence *:"setsum (op $ (x + (e / 2) *s basis a)) UNIV = setsum (\i. x$i + (if a = i then e/2 else 0)) UNIV" by(rule_tac setsum_cong, auto) + have "setsum (op $ x) UNIV < setsum (op $ (x + (e / 2) *s basis a)) UNIV" unfolding * setsum_addf + using `0 \ 1" using ** apply(drule_tac as[rule_format]) by auto + finally show "setsum (op $ x) UNIV < 1" by auto qed +next + fix x::"real^'n::finite" assume as:"\i. 0 < x $ i" "setsum (op $ x) UNIV < 1" + guess a using UNIV_witness[where 'a='b] .. + let ?d = "(1 - setsum (op $ x) UNIV) / real (CARD('n))" + have "Min ((op $ x) ` UNIV) > 0" apply(rule Min_grI) using as(1) dimindex_ge_1 by auto + moreover have"?d > 0" apply(rule divide_pos_pos) using as(2) using dimindex_ge_1 by(auto simp add: Suc_le_eq) + ultimately show "\e>0. \y. dist x y < e \ (\i. 0 \ y $ i) \ setsum (op $ y) UNIV \ 1" + apply(rule_tac x="min (Min ((op $ x) ` UNIV)) ?D" in exI) apply rule defer apply(rule,rule) proof- + fix y assume y:"dist x y < min (Min (op $ x ` UNIV)) ?d" + have "setsum (op $ y) UNIV \ setsum (\i. x$i + ?d) UNIV" proof(rule setsum_mono) + fix i::'n have "abs (y$i - x$i) < ?d" apply(rule le_less_trans) using component_le_norm[of "y - x" i] + using y[unfolded min_less_iff_conj dist_def, THEN conjunct2] by(auto simp add:vector_component_simps norm_minus_commute) + thus "y $ i \ x $ i + ?d" by auto qed + also have "\ \ 1" unfolding setsum_addf setsum_constant card_enum real_eq_of_nat using dimindex_ge_1 by(auto simp add: Suc_le_eq) + finally show "(\i. 0 \ y $ i) \ setsum (op $ y) UNIV \ 1" apply- proof(rule,rule) + fix i::'n have "norm (x - y) < x$i" using y[unfolded min_less_iff_conj dist_def, THEN conjunct1] + using Min_gr_iff[of "op $ x ` dimset x"] dimindex_ge_1 by auto + thus "0 \ y$i" using component_le_norm[of "x - y" i] and as(1)[rule_format, of i] by(auto simp add: vector_component_simps) + qed auto qed auto qed + +lemma interior_std_simplex_nonempty: obtains a::"real^'n::finite" where + "a \ interior(convex hull (insert 0 {basis i | i . i \ UNIV}))" proof- + let ?D = "UNIV::'n set" let ?a = "setsum (\b. inverse (2 * real CARD('n)) *s b) {(basis i) | i. i \ ?D}" + have *:"{basis i | i. i \ ?D} = basis ` ?D" by auto + { fix i have "?a $ i = inverse (2 * real CARD('n))" + unfolding setsum_component vector_smult_component and * and setsum_reindex[OF basis_inj] and o_def + apply(rule trans[of _ "setsum (\j. if i = j then inverse (2 * real CARD('n)) else 0) ?D"]) apply(rule setsum_cong2) + unfolding setsum_delta'[OF finite_UNIV[where 'a='n]] and real_dimindex_ge_1[where 'n='n] by(auto simp add: basis_component[of i]) } + note ** = this + show ?thesis apply(rule that[of ?a]) unfolding interior_std_simplex mem_Collect_eq proof(rule,rule) + fix i::'n show "0 < ?a $ i" unfolding ** using dimindex_ge_1 by(auto simp add: Suc_le_eq) next + have "setsum (op $ ?a) ?D = setsum (\i. inverse (2 * real CARD('n))) ?D" by(rule setsum_cong2, rule **) + also have "\ < 1" unfolding setsum_constant card_enum real_eq_of_nat real_divide_def[THEN sym] by (auto simp add:field_simps) + finally show "setsum (op $ ?a) ?D < 1" by auto qed qed + +subsection {* Paths. *} + +definition "path (g::real^1 \ real^'n::finite) \ continuous_on {0 .. 1} g" + +definition "pathstart (g::real^1 \ real^'n) = g 0" + +definition "pathfinish (g::real^1 \ real^'n) = g 1" + +definition "path_image (g::real^1 \ real^'n) = g ` {0 .. 1}" + +definition "reversepath (g::real^1 \ real^'n) = (\x. g(1 - x))" + +definition joinpaths:: "(real^1 \ real^'n) \ (real^1 \ real^'n) \ (real^1 \ real^'n)" (infixr "+++" 75) + where "joinpaths g1 g2 = (\x. if dest_vec1 x \ ((1 / 2)::real) then g1 (2 *s x) else g2(2 *s x - 1))" +definition "simple_path (g::real^1 \ real^'n) \ + (\x\{0..1}. \y\{0..1}. g x = g y \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0)" + +definition "injective_path (g::real^1 \ real^'n) \ + (\x\{0..1}. \y\{0..1}. g x = g y \ x = y)" + +subsection {* Some lemmas about these concepts. *} + +lemma injective_imp_simple_path: + "injective_path g \ simple_path g" + unfolding injective_path_def simple_path_def by auto + +lemma path_image_nonempty: "path_image g \ {}" + unfolding path_image_def image_is_empty interval_eq_empty by auto + +lemma pathstart_in_path_image[intro]: "(pathstart g) \ path_image g" + unfolding pathstart_def path_image_def apply(rule imageI) + unfolding mem_interval_1 vec_1[THEN sym] dest_vec1_vec by auto + +lemma pathfinish_in_path_image[intro]: "(pathfinish g) \ path_image g" + unfolding pathfinish_def path_image_def apply(rule imageI) + unfolding mem_interval_1 vec_1[THEN sym] dest_vec1_vec by auto + +lemma connected_path_image[intro]: "path g \ connected(path_image g)" + unfolding path_def path_image_def apply(rule connected_continuous_image, assumption) + by(rule convex_connected, rule convex_interval) + +lemma compact_path_image[intro]: "path g \ compact(path_image g)" + unfolding path_def path_image_def apply(rule compact_continuous_image, assumption) + by(rule compact_interval) + +lemma reversepath_reversepath[simp]: "reversepath(reversepath g) = g" + unfolding reversepath_def by auto + +lemma pathstart_reversepath[simp]: "pathstart(reversepath g) = pathfinish g" + unfolding pathstart_def reversepath_def pathfinish_def by auto + +lemma pathfinish_reversepath[simp]: "pathfinish(reversepath g) = pathstart g" + unfolding pathstart_def reversepath_def pathfinish_def by auto + +lemma pathstart_join[simp]: "pathstart(g1 +++ g2) = pathstart g1" + unfolding pathstart_def joinpaths_def pathfinish_def by auto + +lemma pathfinish_join[simp]:"pathfinish(g1 +++ g2) = pathfinish g2" proof- + have "2 *s 1 - 1 = (1::real^1)" unfolding Cart_eq by(auto simp add:vector_component_simps) + thus ?thesis unfolding pathstart_def joinpaths_def pathfinish_def + unfolding vec_1[THEN sym] dest_vec1_vec by auto qed + +lemma path_image_reversepath[simp]: "path_image(reversepath g) = path_image g" proof- + have *:"\g. path_image(reversepath g) \ path_image g" + unfolding path_image_def subset_eq reversepath_def Ball_def image_iff apply(rule,rule,erule bexE) + apply(rule_tac x="1 - xa" in bexI) by(auto simp add:vector_less_eq_def vector_component_simps elim!:ballE) + show ?thesis using *[of g] *[of "reversepath g"] unfolding reversepath_reversepath by auto qed + +lemma path_reversepath[simp]: "path(reversepath g) \ path g" proof- + have *:"\g. path g \ path(reversepath g)" unfolding path_def reversepath_def + apply(rule continuous_on_compose[unfolded o_def, of _ "\x. 1 - x"]) + apply(rule continuous_on_sub, rule continuous_on_const, rule continuous_on_id) + apply(rule continuous_on_subset[of "{0..1}"], assumption) + by (auto, auto simp add:vector_less_eq_def vector_component_simps elim!:ballE) + show ?thesis using *[of g] *[of "reversepath g"] unfolding reversepath_reversepath by auto qed + +lemmas reversepath_simps = path_reversepath path_image_reversepath pathstart_reversepath pathfinish_reversepath + +lemma path_join[simp]: assumes "pathfinish g1 = pathstart g2" shows "path (g1 +++ g2) \ path g1 \ path g2" + unfolding path_def pathfinish_def pathstart_def apply rule defer apply(erule conjE) proof- + assume as:"continuous_on {0..1} (g1 +++ g2)" + have *:"g1 = (\x. g1 (2 *s x)) \ (\x. (1/2) *s x)" + "g2 = (\x. g2 (2 *s x - 1)) \ (\x. (1/2) *s (x + 1))" unfolding o_def by auto + have "op *s (1 / 2) ` {0::real^1..1} \ {0..1}" "(\x. (1 / 2) *s (x + 1)) ` {(0::real^1)..1} \ {0..1}" + unfolding image_smult_interval by (auto, auto simp add:vector_less_eq_def vector_component_simps elim!:ballE) + thus "continuous_on {0..1} g1 \ continuous_on {0..1} g2" apply -apply rule + apply(subst *) defer apply(subst *) apply (rule_tac[!] continuous_on_compose) + apply (rule continuous_on_cmul, rule continuous_on_add, rule continuous_on_id, rule continuous_on_const) defer + apply (rule continuous_on_cmul, rule continuous_on_id) apply(rule_tac[!] continuous_on_eq[of _ "g1 +++ g2"]) defer prefer 3 + apply(rule_tac[1-2] continuous_on_subset[of "{0 .. 1}"]) apply(rule as, assumption, rule as, assumption) + apply(rule) defer apply rule proof- + fix x assume "x \ op *s (1 / 2) ` {0::real^1..1}" + hence "dest_vec1 x \ 1 / 2" unfolding image_iff by(auto simp add: vector_component_simps) + thus "(g1 +++ g2) x = g1 (2 *s x)" unfolding joinpaths_def by auto next + fix x assume "x \ (\x. (1 / 2) *s (x + 1)) ` {0::real^1..1}" + hence "dest_vec1 x \ 1 / 2" unfolding image_iff by(auto simp add: vector_component_simps) + thus "(g1 +++ g2) x = g2 (2 *s x - 1)" proof(cases "dest_vec1 x = 1 / 2") + case True hence "x = (1/2) *s 1" unfolding Cart_eq by(auto simp add: forall_1 vector_component_simps) + thus ?thesis unfolding joinpaths_def using assms[unfolded pathstart_def pathfinish_def] by auto + qed (auto simp add:le_less joinpaths_def) qed +next assume as:"continuous_on {0..1} g1" "continuous_on {0..1} g2" + have *:"{0 .. 1::real^1} = {0.. (1/2)*s 1} \ {(1/2) *s 1 .. 1}" by(auto simp add: vector_component_simps) + have **:"op *s 2 ` {0..(1 / 2) *s 1} = {0..1::real^1}" apply(rule set_ext, rule) unfolding image_iff + defer apply(rule_tac x="(1/2)*s x" in bexI) by(auto simp add: vector_component_simps) + have ***:"(\x. 2 *s x - 1) ` {(1 / 2) *s 1..1} = {0..1::real^1}" + unfolding image_affinity_interval[of _ "- 1", unfolded diff_def[symmetric]] and interval_eq_empty_1 + by(auto simp add: vector_component_simps) + have ****:"\x::real^1. x $ 1 * 2 = 1 \ x = (1/2) *s 1" unfolding Cart_eq by(auto simp add: forall_1 vector_component_simps) + show "continuous_on {0..1} (g1 +++ g2)" unfolding * apply(rule continuous_on_union) apply(rule closed_interval)+ proof- + show "continuous_on {0..(1 / 2) *s 1} (g1 +++ g2)" apply(rule continuous_on_eq[of _ "\x. g1 (2 *s x)"]) defer + unfolding o_def[THEN sym] apply(rule continuous_on_compose) apply(rule continuous_on_cmul, rule continuous_on_id) + unfolding ** apply(rule as(1)) unfolding joinpaths_def by(auto simp add: vector_component_simps) next + show "continuous_on {(1/2)*s1..1} (g1 +++ g2)" apply(rule continuous_on_eq[of _ "g2 \ (\x. 2 *s x - 1)"]) defer + apply(rule continuous_on_compose) apply(rule continuous_on_sub, rule continuous_on_cmul, rule continuous_on_id, rule continuous_on_const) + unfolding *** o_def joinpaths_def apply(rule as(2)) using assms[unfolded pathstart_def pathfinish_def] + by(auto simp add: vector_component_simps ****) qed qed + +lemma path_image_join_subset: "path_image(g1 +++ g2) \ (path_image g1 \ path_image g2)" proof + fix x assume "x \ path_image (g1 +++ g2)" + then obtain y where y:"y\{0..1}" "x = (if dest_vec1 y \ 1 / 2 then g1 (2 *s y) else g2 (2 *s y - 1))" + unfolding path_image_def image_iff joinpaths_def by auto + thus "x \ path_image g1 \ path_image g2" apply(cases "dest_vec1 y \ 1/2") + apply(rule_tac UnI1) defer apply(rule_tac UnI2) unfolding y(2) path_image_def using y(1) + by(auto intro!: imageI simp add: vector_component_simps) qed + +lemma subset_path_image_join: + assumes "path_image g1 \ s" "path_image g2 \ s" shows "path_image(g1 +++ g2) \ s" + using path_image_join_subset[of g1 g2] and assms by auto + +lemma path_image_join: + assumes "path g1" "path g2" "pathfinish g1 = pathstart g2" + shows "path_image(g1 +++ g2) = (path_image g1) \ (path_image g2)" +apply(rule, rule path_image_join_subset, rule) unfolding Un_iff proof(erule disjE) + fix x assume "x \ path_image g1" + then obtain y where y:"y\{0..1}" "x = g1 y" unfolding path_image_def image_iff by auto + thus "x \ path_image (g1 +++ g2)" unfolding joinpaths_def path_image_def image_iff + apply(rule_tac x="(1/2) *s y" in bexI) by(auto simp add: vector_component_simps) next + fix x assume "x \ path_image g2" + then obtain y where y:"y\{0..1}" "x = g2 y" unfolding path_image_def image_iff by auto + moreover have *:"y $ 1 = 0 \ y = 0" unfolding Cart_eq by auto + ultimately show "x \ path_image (g1 +++ g2)" unfolding joinpaths_def path_image_def image_iff + apply(rule_tac x="(1/2) *s (y + 1)" in bexI) using assms(3)[unfolded pathfinish_def pathstart_def] + by(auto simp add: vector_component_simps) qed + +lemma not_in_path_image_join: + assumes "x \ path_image g1" "x \ path_image g2" shows "x \ path_image(g1 +++ g2)" + using assms and path_image_join_subset[of g1 g2] by auto + +lemma simple_path_reversepath: assumes "simple_path g" shows "simple_path (reversepath g)" + using assms unfolding simple_path_def reversepath_def apply- apply(rule ballI)+ + apply(erule_tac x="1-x" in ballE, erule_tac x="1-y" in ballE) + unfolding mem_interval_1 by(auto simp add:vector_component_simps) + +lemma simple_path_join_loop: + assumes "injective_path g1" "injective_path g2" "pathfinish g2 = pathstart g1" + "(path_image g1 \ path_image g2) \ {pathstart g1,pathstart g2}" + shows "simple_path(g1 +++ g2)" +unfolding simple_path_def proof((rule ballI)+, rule impI) let ?g = "g1 +++ g2" + note inj = assms(1,2)[unfolded injective_path_def, rule_format] + fix x y::"real^1" assume xy:"x \ {0..1}" "y \ {0..1}" "?g x = ?g y" + show "x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" proof(case_tac "x$1 \ 1/2",case_tac[!] "y$1 \ 1/2", unfold not_le) + assume as:"x $ 1 \ 1 / 2" "y $ 1 \ 1 / 2" + hence "g1 (2 *s x) = g1 (2 *s y)" using xy(3) unfolding joinpaths_def dest_vec1_def by auto + moreover have "2 *s x \ {0..1}" "2 *s y \ {0..1}" using xy(1,2) as + unfolding mem_interval_1 dest_vec1_def by(auto simp add:vector_component_simps) + ultimately show ?thesis using inj(1)[of "2*s x" "2*s y"] by auto + next assume as:"x $ 1 > 1 / 2" "y $ 1 > 1 / 2" + hence "g2 (2 *s x - 1) = g2 (2 *s y - 1)" using xy(3) unfolding joinpaths_def dest_vec1_def by auto + moreover have "2 *s x - 1 \ {0..1}" "2 *s y - 1 \ {0..1}" using xy(1,2) as + unfolding mem_interval_1 dest_vec1_def by(auto simp add:vector_component_simps) + ultimately show ?thesis using inj(2)[of "2*s x - 1" "2*s y - 1"] by auto + next assume as:"x $ 1 \ 1 / 2" "y $ 1 > 1 / 2" + hence "?g x \ path_image g1" "?g y \ path_image g2" unfolding path_image_def joinpaths_def + using xy(1,2)[unfolded mem_interval_1] by(auto simp add:vector_component_simps intro!: imageI) + moreover have "?g y \ pathstart g2" using as(2) unfolding pathstart_def joinpaths_def + using inj(2)[of "2 *s y - 1" 0] and xy(2)[unfolded mem_interval_1] + apply(rule_tac ccontr) by(auto simp add:vector_component_simps field_simps Cart_eq) + ultimately have *:"?g x = pathstart g1" using assms(4) unfolding xy(3) by auto + hence "x = 0" unfolding pathstart_def joinpaths_def using as(1) and xy(1)[unfolded mem_interval_1] + using inj(1)[of "2 *s x" 0] by(auto simp add:vector_component_simps) + moreover have "y = 1" using * unfolding xy(3) assms(3)[THEN sym] + unfolding joinpaths_def pathfinish_def using as(2) and xy(2)[unfolded mem_interval_1] + using inj(2)[of "2 *s y - 1" 1] by (auto simp add:vector_component_simps Cart_eq) + ultimately show ?thesis by auto + next assume as:"x $ 1 > 1 / 2" "y $ 1 \ 1 / 2" + hence "?g x \ path_image g2" "?g y \ path_image g1" unfolding path_image_def joinpaths_def + using xy(1,2)[unfolded mem_interval_1] by(auto simp add:vector_component_simps intro!: imageI) + moreover have "?g x \ pathstart g2" using as(1) unfolding pathstart_def joinpaths_def + using inj(2)[of "2 *s x - 1" 0] and xy(1)[unfolded mem_interval_1] + apply(rule_tac ccontr) by(auto simp add:vector_component_simps field_simps Cart_eq) + ultimately have *:"?g y = pathstart g1" using assms(4) unfolding xy(3) by auto + hence "y = 0" unfolding pathstart_def joinpaths_def using as(2) and xy(2)[unfolded mem_interval_1] + using inj(1)[of "2 *s y" 0] by(auto simp add:vector_component_simps) + moreover have "x = 1" using * unfolding xy(3)[THEN sym] assms(3)[THEN sym] + unfolding joinpaths_def pathfinish_def using as(1) and xy(1)[unfolded mem_interval_1] + using inj(2)[of "2 *s x - 1" 1] by(auto simp add:vector_component_simps Cart_eq) + ultimately show ?thesis by auto qed qed + +lemma injective_path_join: + assumes "injective_path g1" "injective_path g2" "pathfinish g1 = pathstart g2" + "(path_image g1 \ path_image g2) \ {pathstart g2}" + shows "injective_path(g1 +++ g2)" + unfolding injective_path_def proof(rule,rule,rule) let ?g = "g1 +++ g2" + note inj = assms(1,2)[unfolded injective_path_def, rule_format] + fix x y assume xy:"x \ {0..1}" "y \ {0..1}" "(g1 +++ g2) x = (g1 +++ g2) y" + show "x = y" proof(cases "x$1 \ 1/2", case_tac[!] "y$1 \ 1/2", unfold not_le) + assume "x $ 1 \ 1 / 2" "y $ 1 \ 1 / 2" thus ?thesis using inj(1)[of "2*s x" "2*s y"] and xy + unfolding mem_interval_1 joinpaths_def by(auto simp add:vector_component_simps) + next assume "x $ 1 > 1 / 2" "y $ 1 > 1 / 2" thus ?thesis using inj(2)[of "2*s x - 1" "2*s y - 1"] and xy + unfolding mem_interval_1 joinpaths_def by(auto simp add:vector_component_simps) + next assume as:"x $ 1 \ 1 / 2" "y $ 1 > 1 / 2" + hence "?g x \ path_image g1" "?g y \ path_image g2" unfolding path_image_def joinpaths_def + using xy(1,2)[unfolded mem_interval_1] by(auto simp add:vector_component_simps intro!: imageI) + hence "?g x = pathfinish g1" "?g y = pathstart g2" using assms(4) unfolding assms(3) xy(3) by auto + thus ?thesis using as and inj(1)[of "2 *s x" 1] inj(2)[of "2 *s y - 1" 0] and xy(1,2) + unfolding pathstart_def pathfinish_def joinpaths_def mem_interval_1 + by(auto simp add:vector_component_simps Cart_eq forall_1) + next assume as:"x $ 1 > 1 / 2" "y $ 1 \ 1 / 2" + hence "?g x \ path_image g2" "?g y \ path_image g1" unfolding path_image_def joinpaths_def + using xy(1,2)[unfolded mem_interval_1] by(auto simp add:vector_component_simps intro!: imageI) + hence "?g x = pathstart g2" "?g y = pathfinish g1" using assms(4) unfolding assms(3) xy(3) by auto + thus ?thesis using as and inj(2)[of "2 *s x - 1" 0] inj(1)[of "2 *s y" 1] and xy(1,2) + unfolding pathstart_def pathfinish_def joinpaths_def mem_interval_1 + by(auto simp add:vector_component_simps forall_1 Cart_eq) qed qed + +lemmas join_paths_simps = path_join path_image_join pathstart_join pathfinish_join + +subsection {* Reparametrizing a closed curve to start at some chosen point. *} + +definition "shiftpath a (f::real^1 \ real^'n) = + (\x. if dest_vec1 (a + x) \ 1 then f(a + x) else f(a + x - 1))" + +lemma pathstart_shiftpath: "a \ 1 \ pathstart(shiftpath a g) = g a" + unfolding pathstart_def shiftpath_def by auto + +(** move this **) +declare forall_1[simp] ex_1[simp] + +lemma pathfinish_shiftpath: assumes "0 \ a" "pathfinish g = pathstart g" + shows "pathfinish(shiftpath a g) = g a" + using assms unfolding pathstart_def pathfinish_def shiftpath_def + by(auto simp add: vector_component_simps) + +lemma endpoints_shiftpath: + assumes "pathfinish g = pathstart g" "a \ {0 .. 1}" + shows "pathfinish(shiftpath a g) = g a" "pathstart(shiftpath a g) = g a" + using assms by(auto intro!:pathfinish_shiftpath pathstart_shiftpath) + +lemma closed_shiftpath: + assumes "pathfinish g = pathstart g" "a \ {0..1}" + shows "pathfinish(shiftpath a g) = pathstart(shiftpath a g)" + using endpoints_shiftpath[OF assms] by auto + +lemma path_shiftpath: + assumes "path g" "pathfinish g = pathstart g" "a \ {0..1}" + shows "path(shiftpath a g)" proof- + have *:"{0 .. 1} = {0 .. 1-a} \ {1-a .. 1}" using assms(3) by(auto simp add: vector_component_simps) + have **:"\x. x + a = 1 \ g (x + a - 1) = g (x + a)" + using assms(2)[unfolded pathfinish_def pathstart_def] by auto + show ?thesis unfolding path_def shiftpath_def * apply(rule continuous_on_union) + apply(rule closed_interval)+ apply(rule continuous_on_eq[of _ "g \ (\x. a + x)"]) prefer 3 + apply(rule continuous_on_eq[of _ "g \ (\x. a - 1 + x)"]) defer prefer 3 + apply(rule continuous_on_intros)+ prefer 2 apply(rule continuous_on_intros)+ + apply(rule_tac[1-2] continuous_on_subset[OF assms(1)[unfolded path_def]]) + using assms(3) and ** by(auto simp add:vector_component_simps field_simps Cart_eq) qed + +lemma shiftpath_shiftpath: assumes "pathfinish g = pathstart g" "a \ {0..1}" "x \ {0..1}" + shows "shiftpath (1 - a) (shiftpath a g) x = g x" + using assms unfolding pathfinish_def pathstart_def shiftpath_def + by(auto simp add: vector_component_simps) + +lemma path_image_shiftpath: + assumes "a \ {0..1}" "pathfinish g = pathstart g" + shows "path_image(shiftpath a g) = path_image g" proof- + { fix x assume as:"g 1 = g 0" "x \ {0..1::real^1}" " \y\{0..1} \ {x. \ a $ 1 + x $ 1 \ 1}. g x \ g (a + y - 1)" + hence "\y\{0..1} \ {x. a $ 1 + x $ 1 \ 1}. g x = g (a + y)" proof(cases "a \ x") + case False thus ?thesis apply(rule_tac x="1 + x - a" in bexI) + using as(1,2) and as(3)[THEN bspec[where x="1 + x - a"]] and assms(1) + by(auto simp add:vector_component_simps field_simps atomize_not) next + case True thus ?thesis using as(1-2) and assms(1) apply(rule_tac x="x - a" in bexI) + by(auto simp add:vector_component_simps field_simps) qed } + thus ?thesis using assms unfolding shiftpath_def path_image_def pathfinish_def pathstart_def + by(auto simp add:vector_component_simps image_iff) qed + +subsection {* Special case of straight-line paths. *} + +definition "linepath a b = (\x. (1 - dest_vec1 x) *s a + dest_vec1 x *s b)" + +lemma pathstart_linepath[simp]: "pathstart(linepath a b) = a" + unfolding pathstart_def linepath_def by auto + +lemma pathfinish_linepath[simp]: "pathfinish(linepath a b) = b" + unfolding pathfinish_def linepath_def by auto + +lemma continuous_linepath_at[intro]: "continuous (at x) (linepath a b)" + unfolding linepath_def by(auto simp add: vec1_dest_vec1 o_def intro!: continuous_intros) + +lemma continuous_on_linepath[intro]: "continuous_on s (linepath a b)" + using continuous_linepath_at by(auto intro!: continuous_at_imp_continuous_on) + +lemma path_linepath[intro]: "path(linepath a b)" + unfolding path_def by(rule continuous_on_linepath) + +lemma path_image_linepath[simp]: "path_image(linepath a b) = (closed_segment a b)" + unfolding path_image_def segment linepath_def apply (rule set_ext, rule) defer + unfolding mem_Collect_eq image_iff apply(erule exE) apply(rule_tac x="u *s 1" in bexI) + by(auto simp add:vector_component_simps) + +lemma reversepath_linepath[simp]: "reversepath(linepath a b) = linepath b a" + unfolding reversepath_def linepath_def by(rule ext, auto simp add:vector_component_simps) + +lemma injective_path_linepath: assumes "a \ b" shows "injective_path(linepath a b)" proof- + { obtain i where i:"a$i \ b$i" using assms[unfolded Cart_eq] by auto + fix x y::"real^1" assume "x $ 1 *s b + y $ 1 *s a = x $ 1 *s a + y $ 1 *s b" + hence "x$1 * (b$i - a$i) = y$1 * (b$i - a$i)" unfolding Cart_eq by(auto simp add:field_simps vector_component_simps) + hence "x = y" unfolding mult_cancel_right Cart_eq using i(1) by(auto simp add:field_simps) } + thus ?thesis unfolding injective_path_def linepath_def by(auto simp add:vector_component_simps field_simps) qed + +lemma simple_path_linepath[intro]: "a \ b \ simple_path(linepath a b)" by(auto intro!: injective_imp_simple_path injective_path_linepath) + +subsection {* Bounding a point away from a path. *} + +lemma not_on_path_ball: assumes "path g" "z \ path_image g" + shows "\e>0. ball z e \ (path_image g) = {}" proof- + obtain a where "a\path_image g" "\y\path_image g. dist z a \ dist z y" + using distance_attains_inf[OF _ path_image_nonempty, of g z] + using compact_path_image[THEN compact_imp_closed, OF assms(1)] by auto + thus ?thesis apply(rule_tac x="dist z a" in exI) using assms(2) by(auto intro!: dist_pos_lt) qed + +lemma not_on_path_cball: assumes "path g" "z \ path_image g" + shows "\e>0. cball z e \ (path_image g) = {}" proof- + obtain e where "ball z e \ path_image g = {}" "e>0" using not_on_path_ball[OF assms] by auto + moreover have "cball z (e/2) \ ball z e" using `e>0` by auto + ultimately show ?thesis apply(rule_tac x="e/2" in exI) by auto qed + +subsection {* Path component, considered as a "joinability" relation (from Tom Hales). *} + +definition "path_component s x y \ (\g. path g \ path_image g \ s \ pathstart g = x \ pathfinish g = y)" + +lemmas path_defs = path_def pathstart_def pathfinish_def path_image_def path_component_def + +lemma path_component_mem: assumes "path_component s x y" shows "x \ s" "y \ s" + using assms unfolding path_defs by auto + +lemma path_component_refl: assumes "x \ s" shows "path_component s x x" + unfolding path_defs apply(rule_tac x="\u. x" in exI) using assms + by(auto intro!:continuous_on_intros) + +lemma path_component_refl_eq: "path_component s x x \ x \ s" + by(auto intro!: path_component_mem path_component_refl) + +lemma path_component_sym: "path_component s x y \ path_component s y x" + using assms unfolding path_component_def apply(erule exE) apply(rule_tac x="reversepath g" in exI) + by(auto simp add: reversepath_simps) + +lemma path_component_trans: assumes "path_component s x y" "path_component s y z" shows "path_component s x z" + using assms unfolding path_component_def apply- apply(erule exE)+ apply(rule_tac x="g +++ ga" in exI) by(auto simp add: path_image_join) + +lemma path_component_of_subset: "s \ t \ path_component s x y \ path_component t x y" + unfolding path_component_def by auto + +subsection {* Can also consider it as a set, as the name suggests. *} + +lemma path_component_set: "path_component s x = { y. (\g. path g \ path_image g \ s \ pathstart g = x \ pathfinish g = y )}" + apply(rule set_ext) unfolding mem_Collect_eq unfolding mem_def path_component_def by auto + +lemma mem_path_component_set:"x \ path_component s y \ path_component s y x" unfolding mem_def by auto + +lemma path_component_subset: "(path_component s x) \ s" + apply(rule, rule path_component_mem(2)) by(auto simp add:mem_def) + +lemma path_component_eq_empty: "path_component s x = {} \ x \ s" + apply rule apply(drule equals0D[of _ x]) defer apply(rule equals0I) unfolding mem_path_component_set + apply(drule path_component_mem(1)) using path_component_refl by auto + +subsection {* Path connectedness of a space. *} + +definition "path_connected s \ (\x\s. \y\s. \g. path g \ (path_image g) \ s \ pathstart g = x \ pathfinish g = y)" + +lemma path_connected_component: "path_connected s \ (\x\s. \y\s. path_component s x y)" + unfolding path_connected_def path_component_def by auto + +lemma path_connected_component_set: "path_connected s \ (\x\s. path_component s x = s)" + unfolding path_connected_component apply(rule, rule, rule, rule path_component_subset) + unfolding subset_eq mem_path_component_set Ball_def mem_def by auto + +subsection {* Some useful lemmas about path-connectedness. *} + +lemma convex_imp_path_connected: assumes "convex s" shows "path_connected s" + unfolding path_connected_def apply(rule,rule,rule_tac x="linepath x y" in exI) + unfolding path_image_linepath using assms[unfolded convex_contains_segment] by auto + +lemma path_connected_imp_connected: assumes "path_connected s" shows "connected s" + unfolding connected_def not_ex apply(rule,rule,rule ccontr) unfolding not_not apply(erule conjE)+ proof- + fix e1 e2 assume as:"open e1" "open e2" "s \ e1 \ e2" "e1 \ e2 \ s = {}" "e1 \ s \ {}" "e2 \ s \ {}" + then obtain x1 x2 where obt:"x1\e1\s" "x2\e2\s" by auto + then obtain g where g:"path g" "path_image g \ s" "pathstart g = x1" "pathfinish g = x2" + using assms[unfolded path_connected_def,rule_format,of x1 x2] by auto + have *:"connected {0..1::real^1}" by(auto intro!: convex_connected convex_interval) + have "{0..1} \ {x \ {0..1}. g x \ e1} \ {x \ {0..1}. g x \ e2}" using as(3) g(2)[unfolded path_defs] by blast + moreover have "{x \ {0..1}. g x \ e1} \ {x \ {0..1}. g x \ e2} = {}" using as(4) g(2)[unfolded path_defs] unfolding subset_eq by auto + moreover have "{x \ {0..1}. g x \ e1} \ {} \ {x \ {0..1}. g x \ e2} \ {}" using g(3,4)[unfolded path_defs] using obt by(auto intro!: exI) + ultimately show False using *[unfolded connected_local not_ex,rule_format, of "{x\{0..1}. g x \ e1}" "{x\{0..1}. g x \ e2}"] + using continuous_open_in_preimage[OF g(1)[unfolded path_def] as(1)] + using continuous_open_in_preimage[OF g(1)[unfolded path_def] as(2)] by auto qed + +lemma open_path_component: assumes "open s" shows "open(path_component s x)" + unfolding open_contains_ball proof + fix y assume as:"y \ path_component s x" + hence "y\s" apply- apply(rule path_component_mem(2)) unfolding mem_def by auto + then obtain e where e:"e>0" "ball y e \ s" using assms[unfolded open_contains_ball] by auto + show "\e>0. ball y e \ path_component s x" apply(rule_tac x=e in exI) apply(rule,rule `e>0`,rule) unfolding mem_ball mem_path_component_set proof- + fix z assume "dist y z < e" thus "path_component s x z" apply(rule_tac path_component_trans[of _ _ y]) defer + apply(rule path_component_of_subset[OF e(2)]) apply(rule convex_imp_path_connected[OF convex_ball, unfolded path_connected_component, rule_format]) using `e>0` + using as[unfolded mem_def] by auto qed qed + +lemma open_non_path_component: assumes "open s" shows "open(s - path_component s x)" unfolding open_contains_ball proof + fix y assume as:"y\s - path_component s x" + then obtain e where e:"e>0" "ball y e \ s" using assms[unfolded open_contains_ball] by auto + show "\e>0. ball y e \ s - path_component s x" apply(rule_tac x=e in exI) apply(rule,rule `e>0`,rule,rule) defer proof(rule ccontr) + fix z assume "z\ball y e" "\ z \ path_component s x" + hence "y \ path_component s x" unfolding not_not mem_path_component_set using `e>0` + apply- apply(rule path_component_trans,assumption) apply(rule path_component_of_subset[OF e(2)]) + apply(rule convex_imp_path_connected[OF convex_ball, unfolded path_connected_component, rule_format]) by auto + thus False using as by auto qed(insert e(2), auto) qed + +lemma connected_open_path_connected: assumes "open s" "connected s" shows "path_connected s" + unfolding path_connected_component_set proof(rule,rule,rule path_component_subset, rule) + fix x y assume "x \ s" "y \ s" show "y \ path_component s x" proof(rule ccontr) + assume "y \ path_component s x" moreover + have "path_component s x \ s \ {}" using `x\s` path_component_eq_empty path_component_subset[of s x] by auto + ultimately show False using `y\s` open_non_path_component[OF assms(1)] open_path_component[OF assms(1)] + using assms(2)[unfolded connected_def not_ex, rule_format, of"path_component s x" "s - path_component s x"] by auto +qed qed + +lemma path_connected_continuous_image: + assumes "continuous_on s f" "path_connected s" shows "path_connected (f ` s)" + unfolding path_connected_def proof(rule,rule) + fix x' y' assume "x' \ f ` s" "y' \ f ` s" + then obtain x y where xy:"x\s" "y\s" "x' = f x" "y' = f y" by auto + guess g using assms(2)[unfolded path_connected_def,rule_format,OF xy(1,2)] .. + thus "\g. path g \ path_image g \ f ` s \ pathstart g = x' \ pathfinish g = y'" + unfolding xy apply(rule_tac x="f \ g" in exI) unfolding path_defs + using assms(1) by(auto intro!: continuous_on_compose continuous_on_subset[of _ _ "g ` {0..1}"]) qed + +lemma homeomorphic_path_connectedness: + "s homeomorphic t \ (path_connected s \ path_connected t)" + unfolding homeomorphic_def homeomorphism_def apply(erule exE|erule conjE)+ apply rule + apply(drule_tac f=f in path_connected_continuous_image) prefer 3 + apply(drule_tac f=g in path_connected_continuous_image) by auto + +lemma path_connected_empty: "path_connected {}" + unfolding path_connected_def by auto + +lemma path_connected_singleton: "path_connected {a}" + unfolding path_connected_def apply(rule,rule) + apply(rule_tac x="linepath a a" in exI) by(auto simp add:segment) + +lemma path_connected_Un: assumes "path_connected s" "path_connected t" "s \ t \ {}" + shows "path_connected (s \ t)" unfolding path_connected_component proof(rule,rule) + fix x y assume as:"x \ s \ t" "y \ s \ t" + from assms(3) obtain z where "z \ s \ t" by auto + thus "path_component (s \ t) x y" using as using assms(1-2)[unfolded path_connected_component] apply- + apply(erule_tac[!] UnE)+ apply(rule_tac[2-3] path_component_trans[of _ _ z]) + by(auto simp add:path_component_of_subset [OF Un_upper1] path_component_of_subset[OF Un_upper2]) qed + +subsection {* sphere is path-connected. *} + +lemma path_connected_punctured_universe: + assumes "2 \ CARD('n::finite)" shows "path_connected((UNIV::(real^'n::finite) set) - {a})" proof- + obtain \ where \:"bij_betw \ {1..CARD('n)} (UNIV::'n set)" using ex_bij_betw_nat_finite_1[OF finite_UNIV] by auto + let ?U = "UNIV::(real^'n) set" let ?u = "?U - {0}" + let ?basis = "\k. basis (\ k)" + let ?A = "\k. {x::real^'n. \i\{1..k}. (basis (\ i)) \ x \ 0}" + have "\k\{2..CARD('n)}. path_connected (?A k)" proof + have *:"\k. ?A (Suc k) = {x. ?basis (Suc k) \ x < 0} \ {x. ?basis (Suc k) \ x > 0} \ ?A k" apply(rule set_ext,rule) defer + apply(erule UnE)+ unfolding mem_Collect_eq apply(rule_tac[1-2] x="Suc k" in bexI) + by(auto elim!: ballE simp add: not_less le_Suc_eq) + fix k assume "k \ {2..CARD('n)}" thus "path_connected (?A k)" proof(induct k) + case (Suc k) show ?case proof(cases "k = 1") + case False from Suc have d:"k \ {1..CARD('n)}" "Suc k \ {1..CARD('n)}" by auto + hence "\ k \ \ (Suc k)" using \[unfolded bij_betw_def inj_on_def, THEN conjunct1, THEN bspec[where x=k]] by auto + hence **:"?basis k + ?basis (Suc k) \ {x. 0 < ?basis (Suc k) \ x} \ (?A k)" + "?basis k - ?basis (Suc k) \ {x. 0 > ?basis (Suc k) \ x} \ ({x. 0 < ?basis (Suc k) \ x} \ (?A k))" using d + by(auto simp add: dot_basis vector_component_simps intro!:bexI[where x=k]) + show ?thesis unfolding * Un_assoc apply(rule path_connected_Un) defer apply(rule path_connected_Un) + prefer 5 apply(rule_tac[1-2] convex_imp_path_connected, rule convex_halfspace_lt, rule convex_halfspace_gt) + apply(rule Suc(1)) apply(rule_tac[2-3] ccontr) using d ** False by auto + next case True hence d:"1\{1..CARD('n)}" "2\{1..CARD('n)}" using Suc(2) by auto + have ***:"Suc 1 = 2" by auto + have **:"\s t P Q. s \ t \ {x. P x \ Q x} = (s \ {x. P x}) \ (t \ {x. Q x})" by auto + have "\ 2 \ \ (Suc 0)" apply(rule ccontr) using \[unfolded bij_betw_def inj_on_def, THEN conjunct1, THEN bspec[where x=2]] using assms by auto + thus ?thesis unfolding * True unfolding ** neq_iff bex_disj_distrib apply - + apply(rule path_connected_Un, rule_tac[1-2] path_connected_Un) defer 3 apply(rule_tac[1-4] convex_imp_path_connected) + apply(rule_tac[5] x=" ?basis 1 + ?basis 2" in nequals0I) + apply(rule_tac[6] x="-?basis 1 + ?basis 2" in nequals0I) + apply(rule_tac[7] x="-?basis 1 - ?basis 2" in nequals0I) + using d unfolding *** by(auto intro!: convex_halfspace_gt convex_halfspace_lt, auto simp add:vector_component_simps dot_basis) + qed qed auto qed note lem = this + + have ***:"\x::real^'n. (\i\{1..CARD('n)}. basis (\ i) \ x \ 0) \ (\i. basis i \ x \ 0)" + apply rule apply(erule bexE) apply(rule_tac x="\ i" in exI) defer apply(erule exE) proof- + fix x::"real^'n" and i assume as:"basis i \ x \ 0" + have "i\\ ` {1..CARD('n)}" using \[unfolded bij_betw_def, THEN conjunct2] by auto + then obtain j where "j\{1..CARD('n)}" "\ j = i" by auto + thus "\i\{1..CARD('n)}. basis (\ i) \ x \ 0" apply(rule_tac x=j in bexI) using as by auto qed auto + have *:"?U - {a} = (\x. x + a) ` {x. x \ 0}" apply(rule set_ext) unfolding image_iff + apply rule apply(rule_tac x="x - a" in bexI) by auto + have **:"\x::real^'n. x\0 \ (\i. basis i \ x \ 0)" unfolding Cart_eq by(auto simp add: dot_basis) + show ?thesis unfolding * apply(rule path_connected_continuous_image) apply(rule continuous_on_intros)+ + unfolding ** apply(rule lem[THEN bspec[where x="CARD('n)"], unfolded ***]) using assms by auto qed + +lemma path_connected_sphere: assumes "2 \ CARD('n::finite)" shows "path_connected {x::real^'n::finite. norm(x - a) = r}" proof(cases "r\0") + case True thus ?thesis proof(cases "r=0") + case False hence "{x::real^'n. norm(x - a) = r} = {}" using True by auto + thus ?thesis using path_connected_empty by auto + qed(auto intro!:path_connected_singleton) next + case False hence *:"{x::real^'n. norm(x - a) = r} = (\x. a + r *s x) ` {x. norm x = 1}" unfolding not_le apply -apply(rule set_ext,rule) + unfolding image_iff apply(rule_tac x="(1/r) *s (x - a)" in bexI) unfolding mem_Collect_eq norm_mul by auto + have ***:"\xa. (if xa = 0 then 0 else 1) \ 1 \ xa = 0" apply(rule ccontr) by auto + have **:"{x::real^'n. norm x = 1} = (\x. (1/norm x) *s x) ` (UNIV - {0})" apply(rule set_ext,rule) + unfolding image_iff apply(rule_tac x=x in bexI) unfolding mem_Collect_eq norm_mul by(auto intro!: ***) + have "continuous_on (UNIV - {0}) (vec1 \ (\x::real^'n. 1 / norm x))" unfolding o_def continuous_on_eq_continuous_within + apply(rule, rule continuous_at_within_inv[unfolded o_def inverse_eq_divide]) apply(rule continuous_at_within) + apply(rule continuous_at_vec1_norm[unfolded o_def,THEN spec]) by auto + thus ?thesis unfolding * ** using path_connected_punctured_universe[OF assms] + by(auto intro!: path_connected_continuous_image continuous_on_intros continuous_on_mul) qed + +lemma connected_sphere: "2 \ CARD('n) \ connected {x::real^'n::finite. norm(x - a) = r}" + using path_connected_sphere path_connected_imp_connected by auto + +(** In continuous_at_vec1_norm : Use \ instead of \. **) + +end \ No newline at end of file