# HG changeset patch # User huffman # Date 1244831017 25200 # Node ID 0e4906ccf2804f0dc24c853c75aa09f0c8fc8714 # Parent 809dfb3d9c76cb1571da166515bb96a7fdb2a79a replace all occurrences of 'op *s' at type real^'n with scaleR diff -r 809dfb3d9c76 -r 0e4906ccf280 src/HOL/Library/Convex_Euclidean_Space.thy --- a/src/HOL/Library/Convex_Euclidean_Space.thy Thu Jun 11 20:19:26 2009 -0700 +++ b/src/HOL/Library/Convex_Euclidean_Space.thy Fri Jun 12 11:23:37 2009 -0700 @@ -39,10 +39,6 @@ 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 vector_sneg_minus1 by simp - (* TODO: move to Euclidean_Space.thy *) - 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" @@ -51,23 +47,23 @@ apply(rule_tac [!] setsum_cong2) using 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)" + shows "(\x\s. (if y = x then f x else 0) *\<^sub>R x) = (if y\s then (f y) *\<^sub>R 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 + have *:"\x y. (if y = x then f x else (0::real)) *\<^sub>R x = (if x=y then (f x) *\<^sub>R x else 0)" by auto + show ?thesis unfolding * using setsum_delta[OF assms, of y "\x. f x *\<^sub>R 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 if_smult:"(if P then x else (y::real)) *\<^sub>R v = (if P then x *\<^sub>R v else y *\<^sub>R v)" by auto 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})" +lemma image_smult_interval:"(\x. m *\<^sub>R (x::real^'n::finite)) ` {a..b} = + (if {a..b} = {} then {} else if 0 \ m then {m *\<^sub>R a..m *\<^sub>R b} else {m *\<^sub>R b..m *\<^sub>R a})" using image_affinity_interval[of m 0 a b] by auto lemma dest_vec1_inverval: @@ -87,9 +83,11 @@ shows " dest_vec1 (setsum f S) = setsum (\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)" +lemma dist_triangle_eq: + fixes x y z :: "real ^ _" + shows "dist x z = dist x y + dist y z \ norm (x - y) *\<^sub>R (y - z) = norm (y - z) *\<^sub>R (x - y)" proof- have *:"x - y + (y - z) = x - z" by auto - show ?thesis unfolding dist_norm norm_triangle_eq[of "x - y" "y - z", unfolded *] + show ?thesis unfolding dist_norm norm_triangle_eq[of "x - y" "y - z", unfolded smult_conv_scaleR *] by(auto simp add:norm_minus_commute) qed lemma norm_eqI:"x = y \ norm x = norm y" by auto @@ -108,12 +106,14 @@ 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)" +definition + affine :: "(real ^ 'n::finite) set \ bool" where + "affine s \ (\x\s. \y\s. \u v::real. u + v = 1 \ (u *\<^sub>R x + v *\<^sub>R y) \ s)" + +lemma affine_alt: "affine s \ (\x\s. \y\s. \u::real. (1 - u) *\<^sub>R x + u *\<^sub>R 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 + hence "(\u v::real. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ s) \ (\u::real. (1 - u) *\<^sub>R x + u *\<^sub>R y \ s)" apply auto apply(erule_tac[!] x="1 - u" in allE) unfolding * by auto } thus ?thesis unfolding affine_def by auto qed @@ -121,7 +121,7 @@ unfolding affine_def by auto lemma affine_sing[intro]: "affine {x}" - unfolding affine_alt by (auto simp add: vector_sadd_rdistrib[THEN sym]) + unfolding affine_alt by (auto simp add: scaleR_left_distrib [symmetric]) lemma affine_UNIV[intro]: "affine UNIV" unfolding affine_def by auto @@ -149,30 +149,30 @@ 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)" +lemma affine: fixes V::"(real^'n::finite) set" + shows "affine V \ (\s u. finite s \ s \ {} \ s \ V \ setsum u s = 1 \ (setsum (\x. (u x) *\<^sub>R x)) s \ V)" unfolding affine_def apply rule apply(rule, rule, rule) apply(erule conjE)+ defer apply(rule, rule, rule, rule, rule) 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") + "\s u. finite s \ s \ {} \ s \ V \ setsum u s = 1 \ (\x\s. u x *\<^sub>R x) \ V" + thus "u *\<^sub>R x + v *\<^sub>R y \ V" apply(cases "x=y") using as(4)[THEN spec[where x="{x,y}"], THEN spec[where x="\w. if w = x then u else v"]] and as(1-3) - by(auto simp add: vector_sadd_rdistrib[THEN sym]) + by(auto simp add: scaleR_left_distrib[THEN sym]) next - fix s u assume as:"\x\V. \y\V. \u v. u + v = 1 \ u *s x + v *s y \ V" + fix s u assume as:"\x\V. \y\V. \u v. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ V" "finite s" "s \ {}" "s \ V" "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) + thus "(\x\s. u x *\<^sub>R 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" + assume IA:"\u s. \2 < card s; \x\V. \y\V. \u v. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ V; finite s; + s \ {}; s \ V; setsum u s = 1; n \ card s \ \ (\x\s. u x *\<^sub>R x) \ V" and + as:"Suc n \ card s" "2 < card s" "\x\V. \y\V. \u v. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ V" "finite s" "s \ {}" "s \ V" "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 @@ -185,7 +185,7 @@ 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") + have "(\xa\s - {x}. (inverse (1 - u x) * u xa) *\<^sub>R 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 @@ -195,9 +195,9 @@ 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 + thus "(\x\s. u x *\<^sub>R x) \ V" unfolding scaleR_scaleR[THEN sym] and scaleR_right.setsum [symmetric] 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)"], + using as(3)[THEN bspec[where x=x], THEN bspec[where x="(inverse (1 - u x)) *\<^sub>R (\xa\s - {x}. u xa *\<^sub>R 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) @@ -206,44 +206,44 @@ 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}" + "affine hull p = {y. \s u. finite s \ s \ {} \ s \ p \ setsum u s = 1 \ setsum (\v. (u v) *\<^sub>R 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" + fix x assume "x\p" thus "\s u. finite s \ s \ {} \ s \ p \ setsum u s = 1 \ (\v\s. u v *\<^sub>R 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" + fix t x s u assume as:"p \ t" "affine t" "finite s" "s \ {}" "s \ p" "setsum u s = 1" "(\v\s. u v *\<^sub>R 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 + show "affine {y. \s u. finite s \ s \ {} \ s \ p \ setsum u s = 1 \ (\v\s. u v *\<^sub>R 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 + fix x assume "\s u. finite s \ s \ {} \ s \ p \ setsum u s = 1 \ (\v\s. u v *\<^sub>R v) = x" + then obtain sx ux where x:"finite sx" "sx \ {}" "sx \ p" "setsum ux sx = 1" "(\v\sx. ux v *\<^sub>R v) = x" by auto + fix y assume "\s u. finite s \ s \ {} \ s \ p \ setsum u s = 1 \ (\v\s. u v *\<^sub>R v) = y" + then obtain sy uy where y:"finite sy" "sy \ {}" "sy \ p" "setsum uy sy = 1" "(\v\sy. uy v *\<^sub>R v) = y" by auto have xy:"finite (sx \ sy)" using x(1) y(1) by auto 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" + show "\s ua. finite s \ s \ {} \ s \ p \ setsum ua s = 1 \ (\v\s. ua v *\<^sub>R v) = u *\<^sub>R x + v *\<^sub>R y" apply(rule_tac x="sx \ sy" in exI) apply(rule_tac x="\a. (if a\sx then u * ux a else 0) + (if a\sy then v * uy a else 0)" in exI) - unfolding 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 scaleR_left_distrib setsum_addf if_smult scaleR_zero_left ** setsum_restrict_set[OF xy, THEN sym] + unfolding scaleR_scaleR[THEN sym] scaleR_right.setsum [symmetric] 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}" + shows "affine hull s = {y. \u. setsum u s = 1 \ setsum (\v. u v *\<^sub>R 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" + fix x u assume "setsum u s = 1" "(\v\s. u v *\<^sub>R v) = x" + thus "\sa u. finite sa \ \ (\x. (x \ sa) = (x \ {})) \ sa \ s \ setsum u sa = 1 \ (\v\sa. u v *\<^sub>R v) = x" apply(rule_tac x=s in exI, rule_tac x=u in exI) using assms 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 + assume "finite t" "\ (\x. (x \ t) = (x \ {}))" "setsum u t = 1" "(\v\t. u v *\<^sub>R v) = x" + thus "\u. setsum u s = 1 \ (\v\s. u v *\<^sub>R v) = x" apply(rule_tac x="\x. if x\t then u x else 0" in exI) + unfolding if_smult scaleR_zero_left and setsum_restrict_set[OF assms, THEN sym] and * by auto qed subsection {* Stepping theorems and hence small special cases. *} @@ -251,14 +251,14 @@ 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)") + shows "(\u::real^'n=>real. setsum u {} = w \ setsum (\x. u x *\<^sub>R x) {} = y) \ w = 0 \ y = 0" (is ?th1) + "finite s \ (\u. setsum u (insert a s) = w \ setsum (\x. u x *\<^sub>R x) (insert a s) = y) \ + (\v u. setsum u s = w - v \ setsum (\x. u x *\<^sub>R x) s = y - v *\<^sub>R 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 + then obtain u where u:"setsum u (insert a s) = w \ (\x\insert a s. u x *\<^sub>R 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 @@ -266,41 +266,41 @@ 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 + then obtain v u where vu:"setsum u s = w - v" "(\x\s. u x *\<^sub>R x) = y - v *\<^sub>R a" by auto + have *:"\x M. (if x = a then v else M) *\<^sub>R x = (if x = a then v *\<^sub>R x else M *\<^sub>R x)" by auto 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 vector_smult_lzero + unfolding scaleR_left_distrib and setsum_addf + unfolding vu and * and scaleR_zero_left 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 + "\x. x \ s \ u x *\<^sub>R x = (if x = a then v *\<^sub>R x else u x *\<^sub>R x)" by auto from False show ?thesis apply(rule_tac x="\x. if x=a then v else u x" in exI) unfolding 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 "\x. u x *\<^sub>R x" "\x. if x = a then v *\<^sub>R x else u x *\<^sub>R 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") +lemma affine_hull_2:"affine hull {a,b::real^'n::finite} = {u *\<^sub>R a + v *\<^sub>R 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}" + have "?lhs = {y. \u. setsum u {a, b} = 1 \ (\v\{a, b}. u v *\<^sub>R v) = y}" using affine_hull_finite[of "{a,b}"] by auto - also have "\ = {y. \v u. u b = 1 - v \ u b *s b = y - v *s a}" + also have "\ = {y. \v u. u b = 1 - v \ u b *\<^sub>R b = y - v *\<^sub>R a}" by(simp add: affine_hull_finite_step(2)[of "{b}" a]) 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") +lemma affine_hull_3: "affine hull {a,b,c::real^'n::finite} = { u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R 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 @@ -314,15 +314,20 @@ 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 + unfolding subset_eq Ball_def unfolding affine_hull_explicit span_explicit mem_Collect_eq smult_conv_scaleR 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" + fix x t u assume as:"finite t" "t \ {}" "t \ insert a s" "setsum u t = 1" "(\v\t. u v *\<^sub>R v) = x" have "(\x. x - a) ` (t - {a}) \ {x - a |x. x \ s}" using as(3) by auto - 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 + thus "\v. x = a + v \ (\S u. finite S \ S \ {x - a |x. x \ s} \ (\v\S. u v *\<^sub>R v) = v)" + apply(rule_tac x="x - a" in exI) + apply (rule conjI, simp) + apply(rule_tac x="(\x. x - a) ` (t - {a})" in exI) + apply(rule_tac x="\x. u (x + a)" in exI) + apply (rule conjI) using as(1) apply simp + apply (erule conjI) + using as(1) + apply (simp add: setsum_reindex[unfolded inj_on_def] scaleR_right_diff_distrib setsum_subtractf scaleR_left.setsum[THEN sym] setsum_diff1 scaleR_left_diff_distrib) + unfolding as by simp qed lemma affine_hull_insert_span: assumes "a \ s" @@ -331,17 +336,17 @@ 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 + then obtain t u where obt:"finite t" "t \ {x - a |x. x \ s}" "a + (\v\t. u v *\<^sub>R v) = y" unfolding span_explicit smult_conv_scaleR 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 + have f:"finite f" "f \ s" "(\v\f. u (v - a) *\<^sub>R (v - a)) = y - a" unfolding f_def using obt by(auto simp add: 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" + show "\sa u. finite sa \ sa \ {} \ sa \ insert a s \ setsum u sa = 1 \ (\v\sa. u v *\<^sub>R v) = y" apply(rule_tac x="insert a f" in exI) apply(rule_tac x="\x. if x=a then 1 - 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 + by (auto simp add: setsum_subtractf scaleR_left.setsum algebra_simps) qed lemma affine_hull_span: assumes "a \ s" @@ -350,10 +355,10 @@ 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)" +definition "convex (s::(real^'n::finite) set) \ + (\x\s. \y\s. \u\0. \v\0. (u + v = 1) \ (u *\<^sub>R x + v *\<^sub>R y) \ s)" + +lemma convex_alt: "convex s \ (\x\s. \y\s. \u. 0 \ u \ u \ 1 \ ((1 - u) *\<^sub>R x + u *\<^sub>R 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) @@ -361,14 +366,14 @@ lemma mem_convex: assumes "convex s" "a \ s" "b \ s" "0 \ u" "u \ 1" - shows "((1 - u) *s a + u *s b) \ s" + shows "((1 - u) *\<^sub>R a + u *\<^sub>R 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]) + unfolding convex_def by (auto simp add:scaleR_left_distrib[THEN sym]) lemma convex_UNIV[intro]: "convex UNIV" unfolding convex_def by auto @@ -381,7 +386,8 @@ 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) + unfolding dot_radd dot_rmult [where 'a=real, unfolded smult_conv_scaleR] + 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 @@ -395,12 +401,13 @@ 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) + unfolding convex_def + by(auto simp add: real_convex_bound_lt dot_radd dot_rmult [where 'a=real, unfolded smult_conv_scaleR]) 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)}" +lemma convex_positive_orthant: "convex {x::real^'n::finite. (\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) @@ -408,18 +415,18 @@ 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)" + \ setsum (\i. u i *\<^sub>R 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" + 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 *\<^sub>R 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-) + show "u *\<^sub>R x + v *\<^sub>R 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) + fix k u x assume as:"\x\s. \y\s. \u\0. \v\0. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R 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 *\<^sub>R 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" + case True hence "(\i = Suc 0..k. u i *\<^sub>R 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 *\<^sub>R 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) @@ -429,32 +436,32 @@ 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 + have ***:"\i k. (u i / (1 - u (Suc k))) *\<^sub>R x i = (inverse (1 - u (Suc k))) *\<^sub>R (u i *\<^sub>R x i)" unfolding real_divide_def by (auto simp add: algebra_simps) 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 * + have "(\i = 1..k. (u i / (1 - u (Suc k))) *\<^sub>R 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" + hence "(1 - u (Suc k)) *\<^sub>R (\i = 1..k. (u i / (1 - u (Suc k))) *\<^sub>R x i) + u (Suc k) *\<^sub>R 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)" + thus ?thesis unfolding setsum_cl_ivl_Suc and *** and scaleR_right.setsum [symmetric] using nn by auto qed qed auto qed + + +lemma convex_explicit: "convex (s::(real^'n::finite) set) \ + (\t u. finite t \ t \ s \ (\x\t. 0 \ u x) \ setsum u t = 1 \ setsum (\x. u x *\<^sub>R 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 + 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 *\<^sub>R x) \ s" "x \ s" "y \ s" "0 \ u" "0 \ v" "u + v = (1::real)" + show "u *\<^sub>R x + v *\<^sub>R y \ s" proof(cases "x=y") + case True show ?thesis unfolding True and scaleR_left_distrib[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)" + fix t u assume asm:"\x\s. \y\s. \u\0. \v\0. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R 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) + from this(2) have "\u. t \ s \ (\x\t. 0 \ u x) \ setsum u t = 1 \ (\x\t. u x *\<^sub>R 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" + fix x f u assume ind:"\u. f \ s \ (\x\f. 0 \ u x) \ setsum u f = 1 \ (\x\f. u x *\<^sub>R 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" + show "(\x\insert x f. u x *\<^sub>R x) \ s" proof(cases "u x = 1") + case True hence "setsum (\x. u x *\<^sub>R 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 *\<^sub>R 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) @@ -465,28 +472,28 @@ 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 + case False hence "inverse (1 - u x) *\<^sub>R (\x\f. u x *\<^sub>R x) \ s" unfolding scaleR_right.setsum and scaleR_scaleR 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" + hence "u x *\<^sub>R x + (1 - u x) *\<^sub>R ((inverse (1 - u x)) *\<^sub>R setsum (\x. u x *\<^sub>R 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 auto thus "t \ s \ (\x\t. 0 \ u x) \ setsum u t = 1 \ (\x\t. u x *\<^sub>R 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)" + \ setsum (\x. u x *\<^sub>R 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)" + fix t u assume as:"\u. (\x\s. 0 \ u x) \ setsum u s = 1 \ (\x\s. u x *\<^sub>R 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"]] + show "(\x\t. u x *\<^sub>R 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)" +definition "cone (s::(real^'n) set) \ (\x\s. \c\0. (c *\<^sub>R x) \ s)" lemma cone_empty[intro, simp]: "cone {}" unfolding cone_def by auto @@ -509,43 +516,45 @@ subsection {* Affine dependence and consequential theorems (from Lars Schewe). *} -definition "affine_dependent (s::(real^'n) set) \ (\x\s. x \ (affine hull (s - {x})))" +definition + affine_dependent :: "(real ^ 'n::finite) set \ bool" where + "affine_dependent s \ (\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)" + (\v\s. u v \ 0) \ setsum (\v. u v *\<^sub>R 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" + fix x s u assume as:"x \ p" "finite s" "s \ {}" "s \ p - {x}" "setsum u s = 1" "(\v\s. u v *\<^sub>R v) = x" have "x\s" using as(1,4) by auto - show "\s u. finite s \ s \ p \ setsum u s = 0 \ (\v\s. u v \ 0) \ (\v\s. u v *s v) = 0" + show "\s u. finite s \ s \ p \ setsum u s = 0 \ (\v\s. u v \ 0) \ (\v\s. u v *\<^sub>R v) = 0" apply(rule_tac x="insert x s" in exI, rule_tac x="\v. if v = x then - 1 else u v" in exI) unfolding if_smult and 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" + fix s u v assume as:"finite s" "s \ p" "setsum u s = 0" "(\v\s. u v *\<^sub>R v) = 0" "v \ s" "u v \ 0" have "s \ {v}" using as(3,6) by auto - thus "\x\p. \s u. finite s \ s \ {} \ s \ p - {x} \ setsum u s = 1 \ (\v\s. u v *s v) = x" + thus "\x\p. \s u. finite s \ s \ {} \ s \ p - {x} \ setsum u s = 1 \ (\v\s. u v *\<^sub>R 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_ring[OF as(1,5)] using as by auto + unfolding scaleR_scaleR[THEN sym] and scaleR_right.setsum [symmetric] unfolding setsum_right_distrib[THEN sym] and setsum_diff1_ring[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)" + assumes "finite (s::(real^'n::finite) set)" + shows "affine_dependent s \ (\u. setsum u s = 0 \ (\v\s. u v \ 0) \ setsum (\v. u v *\<^sub>R 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 + have *:"\vt u v. (if vt then u v else 0) *\<^sub>R v = (if vt then (u v) *\<^sub>R 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" + then obtain t u v where "finite t" "t \ s" "setsum u t = 0" "v\t" "u v \ 0" "(\v\t. u v *\<^sub>R 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 + then obtain u v where "setsum u s = 0" "v\s" "u v \ 0" "(\v\s. u v *\<^sub>R v) = 0" by auto thus ?lhs unfolding affine_dependent_explicit using assms by auto qed @@ -560,24 +569,24 @@ 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) + { fix y have *:"(1 - x) *\<^sub>R x1 + x *\<^sub>R x2 - ((1 - y) *\<^sub>R x1 + y *\<^sub>R x2) = (y - x) *\<^sub>R x1 - (y - x) *\<^sub>R x2" + by (simp add: algebra_simps) 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 + hence "norm ((1 - x) *\<^sub>R x1 + x *\<^sub>R x2 - ((1 - y) *\<^sub>R x1 + y *\<^sub>R x2)) < e" + unfolding * and scaleR_right_diff_distrib[THEN sym] and norm_scaleR 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" + hence "\d>0. \y. \y - x\ < d \ norm ((1 - x) *\<^sub>R x1 + x *\<^sub>R x2 - ((1 - y) *\<^sub>R x1 + y *\<^sub>R 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" + have "\x\0. x \ 1 \ (1 - x) *\<^sub>R x1 + x *\<^sub>R x2 \ e1 \ (1 - x) *\<^sub>R x1 + x *\<^sub>R x2 \ e2" apply(rule connected_real_lemma) apply (simp add: `x1\e1` `x2\e2` dist_commute)+ using * apply(simp add: dist_norm) using as(1,2)[unfolded open_dist] apply simp using as(1,2)[unfolded open_dist] 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 + then obtain x where "x\0" "x\1" "(1 - x) *\<^sub>R x1 + x *\<^sub>R x2 \ e1" "(1 - x) *\<^sub>R x1 + x *\<^sub>R 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 } @@ -592,7 +601,7 @@ 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)" + (\x\s. \y\s. \u\0. \v\0. u + v = 1 \ f (u *\<^sub>R x + v *\<^sub>R 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 @@ -603,11 +612,11 @@ 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)" + ultimately have "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R 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) } + hence "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R 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 @@ -621,7 +630,7 @@ 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)" + shows "f (u *\<^sub>R x + v *\<^sub>R 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) @@ -642,13 +651,13 @@ 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` + hence "f ((1-u) *\<^sub>R x + u *\<^sub>R 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_norm unfolding * and norm_mul and abs_of_pos[OF `0R x + u *\<^sub>R y) = u *\<^sub>R (x - y)" by (simp add: algebra_simps) + have "(1 - u) *\<^sub>R x + u *\<^sub>R y \ ball x e" unfolding mem_ball dist_norm unfolding * and norm_scaleR and abs_of_pos[OF `0 f ((1 - u) *s x + u *s y)" using assms(4) by auto + hence "f x \ f ((1 - u) *\<^sub>R x + u *\<^sub>R 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 @@ -656,31 +665,32 @@ proof(auto simp add: convex_on_def dist_norm) 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 + have "a = u *\<^sub>R a + v *\<^sub>R a" unfolding scaleR_left_distrib[THEN sym] and `u+v=1` by simp + hence *:"a - (u *\<^sub>R x + v *\<^sub>R y) = (u *\<^sub>R (a - x)) + (v *\<^sub>R (a - y))" + by (auto simp add: algebra_simps) + show "norm (a - (u *\<^sub>R x + v *\<^sub>R y)) \ u * norm (a - x) + v * norm (a - y)" + unfolding * using norm_triangle_ineq[of "u *\<^sub>R (a - x)" "v *\<^sub>R (a - y)"] unfolding norm_scaleR 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)" +lemma convex_scaling: "convex s \ convex ((\x. c *\<^sub>R 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) + apply (rule_tac x="u *\<^sub>R x+v *\<^sub>R y" in bexI) by (auto simp add: algebra_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 + apply (rule_tac x="u *\<^sub>R x+v *\<^sub>R 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) +proof(auto simp add: convex_def image_iff scaleR_right_distrib) 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) + show "\x y. u *\<^sub>R xa + u *\<^sub>R ya + (v *\<^sub>R xb + v *\<^sub>R yb) = x + y \ x \ s \ y \ t" + apply(rule_tac x="u *\<^sub>R xa + v *\<^sub>R xb" in exI) apply(rule_tac x="u *\<^sub>R ya + v *\<^sub>R 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 @@ -700,17 +710,17 @@ 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 +lemma convex_affinity: assumes "convex (s::(real^'n::finite) set)" shows "convex ((\x. a + c *\<^sub>R x) ` s)" +proof- have "(\x. a + c *\<^sub>R x) ` s = op + a ` op *\<^sub>R 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] + show "u *\<^sub>R f x + v *\<^sub>R f y \ f ` s" unfolding image_iff + apply(rule_tac x="u *\<^sub>R x + v *\<^sub>R y" in bexI) + unfolding linear_add[OF l] linear_cmul[OF l, unfolded smult_conv_scaleR] using c[unfolded convex_def] xy uv by auto qed @@ -720,18 +730,18 @@ 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 + have "dist x (u *\<^sub>R y + v *\<^sub>R 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 + thus "dist x (u *\<^sub>R y + v *\<^sub>R 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 + have "dist x (u *\<^sub>R y + v *\<^sub>R 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 + thus "dist x (u *\<^sub>R y + v *\<^sub>R z) \ e" using real_convex_bound_le[OF yz uv] by auto qed lemma connected_ball: "connected(ball (x::real^_) e)" (* FIXME: generalize *) @@ -770,14 +780,14 @@ 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") + b \ (convex hull s) \ (x = u *\<^sub>R a + v *\<^sub>R 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 + then obtain u v b where obt:"u\0" "v\0" "u + v = 1" "b \ convex hull s" "x = u *\<^sub>R a + v *\<^sub>R b" by auto have "a\convex hull insert a s" "b\convex hull insert a s" using hull_mono[of s "insert a s" convex] hull_mono[of "{a}" "insert a s" convex] and obt(4) by auto thus "x\ convex hull insert a s" unfolding obt(5) using convex_convex_hull[of "insert a s", unfolded convex_def] @@ -785,16 +795,16 @@ 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)" + from as(4) obtain u1 v1 b1 where obt1:"u1\0" "v1\0" "u1 + v1 = 1" "b1 \ convex hull s" "x = u1 *\<^sub>R a + v1 *\<^sub>R b1" by auto + from as(5) obtain u2 v2 b2 where obt2:"u2\0" "v2\0" "u2 + v2 = 1" "b2 \ convex hull s" "y = u2 *\<^sub>R a + v2 *\<^sub>R b2" by auto + have *:"\(x::real^_) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" by (auto simp add: algebra_simps) + have "\b \ convex hull s. u *\<^sub>R x + v *\<^sub>R y = (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (b - (u * u1) *\<^sub>R b - (v * u2) *\<^sub>R b)" proof(cases "u * v1 + v * v2 = 0") - have *:"\x s1 s2. x - s1 *s x - s2 *s x = ((1::real) - (s1 + s2)) *s x" by auto + have *:"\(x::real^_) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" by (auto simp add: algebra_simps) 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: **) + thus ?thesis unfolding obt1(5) obt2(5) * using assms hull_subset[of s convex] by(auto simp add: ** scaleR_right_distrib) 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) @@ -803,9 +813,10 @@ 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_tac x="((u * v1) / (u * v1 + v * v2)) *\<^sub>R b1 + ((v * v2) / (u * v1 + v * v2)) *\<^sub>R b2" in bexI) defer apply(rule 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 + unfolding add_divide_distrib[THEN sym] and real_0_le_divide_iff + by (auto simp add: scaleR_left_distrib scaleR_right_distrib) 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 @@ -813,9 +824,9 @@ 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) + show "u *\<^sub>R x + v *\<^sub>R y \ ?hull" unfolding mem_Collect_eq apply(rule_tac x="u * u1 + v * u2" in exI) apply(rule conjI) defer apply(rule_tac x="1 - u * u1 - v * u2" in exI) unfolding Bex_def - using as(1,2) obt1(1,2) obt2(1,2) * by(auto intro!: mult_nonneg_nonneg add_nonneg_nonneg simp add:field_simps) + using as(1,2) obt1(1,2) obt2(1,2) * by(auto intro!: mult_nonneg_nonneg add_nonneg_nonneg simp add: algebra_simps) qed qed @@ -825,7 +836,7 @@ 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") + (setsum (\i. u i *\<^sub>R 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- @@ -834,22 +845,22 @@ 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" + 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 *\<^sub>R 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)" + 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 *\<^sub>R x1 i) = x" by auto + from xy obtain k2 u2 x2 where y:"\i\{1::nat..k2}. 0\u2 i \ x2 i \ s" "setsum u2 {Suc 0..k2} = 1" "(\i = Suc 0..k2. u2 i *\<^sub>R x2 i) = y" by auto + have *:"\P x1 x2 s1 s2 i.(if P i then s1 else s2) *\<^sub>R (if P i then x1 else x2) = (if P i then s1 *\<^sub>R x1 else s2 *\<^sub>R x2)" "{1..k1 + k2} \ {1..k1} = {1..k1}" "{1..k1 + k2} \ - {1..k1} = (\i. i + k1) ` {1..k2}" prefer 3 apply(rule,rule) unfolding image_iff apply(rule_tac x="x - k1" in bexI) 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) + show "u *\<^sub>R x + v *\<^sub>R 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- + unfolding scaleR_scaleR[THEN sym] scaleR_right.setsum [symmetric] 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}") @@ -862,11 +873,11 @@ qed lemma convex_hull_finite: - assumes "finite (s::(real^'n)set)" + assumes "finite (s::(real^'n::finite)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") + setsum u s = 1 \ setsum (\x. u x *\<^sub>R 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" + fix x assume "x\s" thus " \u. (\x\s. 0 \ u x) \ setsum u s = 1 \ (\x\s. u x *\<^sub>R x) = x" apply(rule_tac x="\y. if x=y then 1 else 0" in exI) apply auto unfolding setsum_delta'[OF assms] and setsum_delta''[OF assms] by auto next @@ -878,14 +889,14 @@ 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)" + moreover have "(\x\s. (u * ux x + v * uy x) *\<^sub>R x) = u *\<^sub>R (\x\s. ux x *\<^sub>R x) + v *\<^sub>R (\x\s. uy x *\<^sub>R x)" + unfolding scaleR_left_distrib and setsum_addf and scaleR_scaleR[THEN sym] and scaleR_right.setsum [symmetric] by auto + ultimately show "\uc. (\x\s. 0 \ uc x) \ setsum uc s = 1 \ (\x\s. uc x *\<^sub>R x) = u *\<^sub>R (\x\s. ux x *\<^sub>R x) + v *\<^sub>R (\x\s. uy x *\<^sub>R x)" apply(rule_tac x="\x. u * ux x + v * uy x" in exI) 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]] + thus "(\x\s. u x *\<^sub>R x) \ t" using t(2)[unfolded convex_explicit, THEN spec[where x=s], THEN spec[where x=u]] using assms and t(1) by auto qed @@ -893,10 +904,10 @@ 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") + (\x\s. 0 \ u x) \ setsum u s = 1 \ setsum (\v. u v *\<^sub>R 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" + 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 *\<^sub>R y i) = x" unfolding convex_hull_indexed by auto have fin:"finite {1..k}" by auto @@ -908,16 +919,16 @@ 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" + moreover have "(\v\y ` {1..k}. setsum u {i \ {1..k}. y i = v} *\<^sub>R v) = x" + using setsum_image_gen[OF fin, of "\i. u i *\<^sub>R y i" y, THEN sym] + unfolding scaleR_left.setsum 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 *\<^sub>R 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 + then obtain s u where obt:"finite s" "s \ p" "\x\s. 0 \ u x" "setsum u s = 1" "(\v\s. u v *\<^sub>R v) = y" by auto obtain f where f:"inj_on f {1..card s}" "f ` {1..card s} = s" using ex_bij_betw_nat_finite_1[OF obt(1)] unfolding bij_betw_def by auto @@ -929,14 +940,14 @@ 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"] + 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) *\<^sub>R f x) = u y *\<^sub>R y" by auto } + + hence "(\x = 1..card s. u (f x)) = 1" "(\i = 1..card s. u (f i) *\<^sub>R f i) = y" + unfolding setsum_image_gen[OF *(1), of "\x. u (f x) *\<^sub>R 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) *\<^sub>R f x)" "\v. u v *\<^sub>R 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" + 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 *\<^sub>R 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 @@ -946,24 +957,24 @@ 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") + shows "(\u. (\x\insert a s. 0 \ u x) \ setsum u (insert a s) = w \ setsum (\x. u x *\<^sub>R x) (insert a s) = y) + \ (\v\0. \u. (\x\s. 0 \ u x) \ setsum u s = w - v \ setsum (\x. u x *\<^sub>R x) s = y - v *\<^sub>R 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 ?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 *\<^sub>R 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] + assume ?rhs then obtain v u where uv:"v\0" "\x\s. 0 \ u x" "setsum u s = w - v" "(\x\s. u x *\<^sub>R x) = y - v *\<^sub>R a" by auto + show ?lhs apply(rule_tac x="\x. (if a = x then v else 0) + u x" in exI) unfolding scaleR_left_distrib and 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)" + assume ?rhs then obtain v u where uv:"v\0" "\x\s. 0 \ u x" "setsum u s = w - v" "(\x\s. u x *\<^sub>R x) = y - v *\<^sub>R 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) *\<^sub>R x) = (\x\s. u x *\<^sub>R 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 @@ -971,20 +982,20 @@ 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}" + "convex hull {a,b} = {u *\<^sub>R a + v *\<^sub>R 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}" +lemma convex_hull_2_alt: "convex hull {a,b} = {a + u *\<^sub>R (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 + fix x show "(\v u. x = v *\<^sub>R a + u *\<^sub>R b \ 0 \ v \ 0 \ u \ v + u = 1) = (\u. x = a + u *\<^sub>R (b - a) \ 0 \ u \ u \ 1)" + unfolding * apply auto apply(rule_tac[!] x=u in exI) by (auto simp add: algebra_simps) 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}" + "convex hull {a::real^'n::finite,b,c} = { u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R 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" @@ -995,15 +1006,15 @@ 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}" + "convex hull {a,b,c} = {a + u *\<^sub>R (b - a) + v *\<^sub>R (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 + 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 add: algebra_simps) + apply(rule_tac x=u in exI) apply(rule_tac x=v in exI) by (simp add: algebra_simps) qed subsection {* Relations among closure notions and corresponding hulls. *} lemma subspace_imp_affine: "subspace s \ affine s" - unfolding subspace_def affine_def by auto + unfolding subspace_def affine_def smult_conv_scaleR by auto lemma affine_imp_convex: "affine s \ convex s" unfolding affine_def convex_def by auto @@ -1031,8 +1042,8 @@ 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 + from assms(1)[unfolded dependent_explicit smult_conv_scaleR] obtain S u v + where obt:"finite S" "S \ {x - a |x. x \ s}" "v\S" "u v \ 0" "(\v\S. u v *\<^sub>R v) = 0" by auto def t \ "(\x. x + a) ` S" have inj:"inj_on (\x. x + a) S" unfolding inj_on_def by auto @@ -1046,24 +1057,24 @@ 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)" + moreover have *:"\P Q. (\x\t. (if x = a then P x else Q x) *\<^sub>R x) = (\x\t. Q x *\<^sub>R x)" apply(rule 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" + have "(\x\t. u (x - a)) *\<^sub>R a = (\v\t. u (v - a) *\<^sub>R v)" + unfolding scaleR_left.setsum unfolding t_def and setsum_reindex[OF inj] and o_def + using obt(5) by (auto simp add: setsum_addf scaleR_right_distrib) + hence "(\v\insert a t. (if v = a then - (\x\t. u (x - a)) else u (v - a)) *\<^sub>R 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") + "convex s \ cone s \ (\x\s. \y\s. (x + y) \ s) \ (\x\s. \c\0. (c *\<^sub>R 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 "2 *\<^sub>R x \s" "2 *\<^sub>R 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="2*\<^sub>R x" in ballE) apply(erule_tac x="2*\<^sub>R y" in ballE) apply(erule_tac x="1/2" in allE) apply simp apply(erule_tac x="1/2" in allE) by auto } thus ?thesis unfolding convex_def cone_def by blast qed @@ -1104,20 +1115,20 @@ 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}" + (\x\s. 0 \ u x) \ setsum u s = 1 \ setsum (\v. u v *\<^sub>R 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" + 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 *\<^sub>R v) = y" + assume "\s u. finite s \ s \ p \ (\x\s. 0 \ u x) \ setsum u s = 1 \ (\v\s. u v *\<^sub>R 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 + 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 *\<^sub>R 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" + then obtain w v where wv:"setsum w s = 0" "v\s" "w v \ 0" "(\v\s. w v *\<^sub>R 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) @@ -1147,15 +1158,15 @@ 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) + moreover have "(\v\s. u v *\<^sub>R v + (t * w v) *\<^sub>R v) - (u a *\<^sub>R a + (t * w a) *\<^sub>R a) = y" + unfolding setsum_addf obt(6) scaleR_scaleR[THEN sym] scaleR_right.setsum [symmetric] wv(4) using a(2) [THEN eq_neg_iff_add_eq_0 [THEN iffD2]] by (simp add: vector_smult_lneg) 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: *) + apply(rule_tac x="\v. u v + t * w v" in exI) using obt(1-3) and t and a by (auto simp add: * scaleR_left_distrib) 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 + \ (\x\s. 0 \ u x) \ setsum u s = 1 \ (\v\s. u v *\<^sub>R v) = y" using obt by auto qed auto lemma caratheodory: @@ -1164,7 +1175,7 @@ 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 + "\x\s. 0 \ u x" "setsum u s = 1" "(\v\s. u v *\<^sub>R 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 @@ -1181,14 +1192,14 @@ 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 + assume "\sa u. finite sa \ sa \ s \ (\x\sa. 0 \ u x) \ setsum u sa = 1 \ (\v\sa. u v *\<^sub>R v) = a" + then obtain t u where obt:"finite t" "t\s" "\x\t. 0 \ u x" "setsum u t = 1" "(\v\t. u v *\<^sub>R 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}" + 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 *\<^sub>R v) = y}" apply(rule_tac x="Min i" in exI) unfolding subset_eq apply rule defer apply rule unfolding mem_Collect_eq proof- show "0 < Min i" unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] `b \` t\{}`] @@ -1204,10 +1215,10 @@ 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" + moreover have "(\v\(\v. v + (y - a)) ` t. u (v - (y - a)) *\<^sub>R 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" + by (simp add: setsum_addf setsum_subtractf scaleR_left.setsum[THEN sym] scaleR_right_distrib) + ultimately show "\sa u. finite sa \ (\x\sa. x \ s) \ (\x\sa. 0 \ u x) \ setsum u sa = 1 \ (\v\sa. u v *\<^sub>R v) = y" apply(rule_tac x="(\v. v + (y - a)) ` t" in exI) apply(rule_tac x="\v. u (v - (y - a))" in exI) using obt(1, 3) by auto qed @@ -1246,18 +1257,18 @@ lemma compact_convex_combinations: fixes s t :: "(real ^ 'n::finite) set" 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}" + shows "compact { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ s \ y \ t}" proof- let ?X = "{0..1} \ s \ t" - let ?h = "(\z. (1 - fst z) *s fst (snd z) + fst z *s snd (snd z))" - have *:"{ (1 - u) *s x + u *s y | x y u. 0 \ u \ u \ 1 \ x \ s \ y \ t} = ?h ` ?X" + let ?h = "(\z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" + have *:"{ (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ s \ y \ t} = ?h ` ?X" apply(rule set_ext) unfolding image_iff mem_Collect_eq apply rule apply auto apply (rule_tac x=u in rev_bexI, simp) apply (erule rev_bexI, erule rev_bexI, simp) by auto have "continuous_on ({0..1} \ s \ t) - (\z. (1 - fst z) *s fst (snd z) + fst z *s snd (snd z))" + (\z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" unfolding continuous_on by (rule ballI) (intro tendsto_intros) thus ?thesis unfolding * apply (rule compact_continuous_image) @@ -1297,14 +1308,14 @@ 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. + { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ s \ y \ {x. \t. finite t \ t \ s \ card t \ n \ x \ convex hull t}}" 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 \ + fix x assume "\u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \ 0 \ c \ c \ 1 \ u \ s \ (\t. finite t \ t \ s \ card t \ n \ v \ convex hull t)" - then obtain u v c t where obt:"x = (1 - c) *s u + c *s v" + then obtain u v c t where obt:"x = (1 - c) *\<^sub>R u + c *\<^sub>R v" "0 \ c \ c \ 1" "u \ s" "finite t" "t \ s" "card t \ n" "v \ convex hull t" by auto - moreover have "(1 - c) *s u + c *s v \ convex hull insert u t" + moreover have "(1 - c) *\<^sub>R u + c *\<^sub>R 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" @@ -1312,7 +1323,7 @@ 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 \ + let ?P = "\u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \ 0 \ c \ c \ 1 \ u \ s \ (\t. finite t \ t \ s \ card t \ n \ v \ convex hull t)" show ?P proof(cases "card t = Suc n") case False hence "card t \ n" using t(3) by auto @@ -1325,7 +1336,7 @@ 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" + case False obtain ux vx b where obt:"ux\0" "vx\0" "ux + vx = 1" "b \ convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b" using t(4)[unfolded au convex_hull_insert[OF False]] 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) @@ -1373,7 +1384,7 @@ 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" + obtain u v b where obt:"u\0" "v\0" "u + v = 1" "b \ convex hull s" "y = u *\<^sub>R x + v *\<^sub>R b" using y(1)[unfolded convex_hull_insert[OF False]] by auto show "\z\convex hull insert x s. norm (y - a) < norm (z - a)" proof(cases "y\convex hull s") @@ -1392,24 +1403,24 @@ 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]) + using obt(3) by(auto simp add: scaleR_left_distrib[THEN sym]) thus False using obt(4) and False by simp qed - hence *:"w *s (x - b) \ 0" using w(1) by auto + hence *:"w *\<^sub>R (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_commute[of a] unfolding dist_norm obt(5) by (simp add: ring_simps) - moreover have "(u + w) *s x + (v - w) *s b \ convex hull insert x s" + assume "dist a y < dist a (y + w *\<^sub>R (x - b))" + hence "norm (y - a) < norm ((u + w) *\<^sub>R x + (v - w) *\<^sub>R b - a)" + unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps) + moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \ convex hull insert x s" unfolding convex_hull_insert[OF `s\{}`] and mem_Collect_eq apply(rule_tac x="u + w" in exI) apply rule defer apply(rule_tac x="v - w" in exI) using `u\0` and w and obt(3,4) 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_commute[of a] unfolding dist_norm obt(5) by (simp add: ring_simps) - moreover have "(u - w) *s x + (v + w) *s b \ convex hull insert x s" + assume "dist a y < dist a (y - w *\<^sub>R (x - b))" + hence "norm (y - a) < norm ((u - w) *\<^sub>R x + (v + w) *\<^sub>R b - a)" + unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps) + moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \ convex hull insert x s" unfolding convex_hull_insert[OF `s\{}`] and mem_Collect_eq apply(rule_tac x="u - w" in exI) apply rule defer apply(rule_tac x="v + w" in exI) using `u\0` and w and obt(3,4) by auto @@ -1495,19 +1506,19 @@ 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" + shows "\u>0. \v>0. v \ u \ norm(v *\<^sub>R 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 _ `0R 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)" + shows "\u>0. u \ 1 \ dist (x + u *\<^sub>R (z - x)) y < dist x y" +proof- obtain u where "u>0" and u:"\v>0. v \ u \ norm (v *\<^sub>R (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_norm by(auto simp add: norm_minus_commute field_simps) qed @@ -1516,9 +1527,9 @@ 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_commute field_simps) qed + then obtain u where u:"u>0" "u\1" "dist (x + u *\<^sub>R (y - x)) a < dist x a" using closer_point_lemma[of a x y] by auto + let ?z = "(1 - u) *\<^sub>R x + u *\<^sub>R y" have "?z \ s" 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_commute algebra_simps) qed lemma any_closest_point_unique: assumes "convex s" "closed s" "x \ s" "y \ s" @@ -1580,11 +1591,11 @@ 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)" + fix x assume "x\s" have *:"\u. 0 \ u \ u \ 1 \ dist z y \ dist z ((1 - u) *\<^sub>R y + u *\<^sub>R x)" using assms(1)[unfolded convex_alt] and y and `x\s` and `y\s` by auto assume "\ (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_commute field_simps) + "v>0" "v\1" "dist (y + v *\<^sub>R (x - y)) z < dist y z" using closer_point_lemma[of z y x] by auto + thus False using *[THEN spec[where x=v]] by(auto simp add: dist_commute algebra_simps) qed auto qed @@ -1601,11 +1612,11 @@ 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)"]] + assume "\u>0. u \ 1 \ dist (y + u *\<^sub>R (x - y)) z < dist y z" + then obtain u where "u>0" "u\1" "dist (y + u *\<^sub>R (x - y)) z < dist y z" by auto + thus False using y[THEN bspec[where x="y + u *\<^sub>R (x - y)"]] using assms(1)[unfolded convex_alt, THEN bspec[where x=y]] - using `x\s` `y\s` by (auto simp add: dist_commute field_simps) qed + using `x\s` `y\s` by (auto simp add: dist_commute algebra_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" @@ -1680,10 +1691,10 @@ 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 + hence "\x. norm x = 1 \ (\y\c. 0 \ y \ x)" apply(rule_tac x="inverse(norm a) *\<^sub>R a" in exI) + using hull_subset[of c convex] unfolding subset_eq and dot_rmult [where 'a=real, unfolded smult_conv_scaleR] apply- apply rule defer apply rule apply(rule mult_nonneg_nonneg) - by(auto simp add: dot_sym elim!: ballE) + by(auto simp add: dot_sym norm_scaleR elim!: ballE) thus "frontier (cball 0 1) \ \f \ {}" unfolding c(1) frontier_cball dist_norm by auto qed(insert closed_halfspace_ge, auto) then obtain x where "norm x = 1" "\y\s. x\?k y" unfolding frontier_cball dist_norm by auto @@ -1704,7 +1715,7 @@ 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_tac x="\n. u *\<^sub>R xb n + v *\<^sub>R 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 @@ -1712,13 +1723,13 @@ 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) + show "\e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \ s" apply(rule_tac x="min d e" in exI) apply rule unfolding subset_eq defer apply rule proof- - 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" + fix z assume "z \ ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) (min d e)" + hence "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \ s" apply(rule_tac assms[unfolded convex_alt, rule_format]) - using ed(1,2) and u unfolding subset_eq mem_ball Ball_def dist_norm 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 + using ed(1,2) and u unfolding subset_eq mem_ball Ball_def dist_norm by(auto simp add: algebra_simps) + thus "z \ s" using u by (auto simp add: algebra_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 @@ -1741,17 +1752,17 @@ 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)" + "(convex hull ((\x. c *\<^sub>R x) ` s)) \ (\x. c *\<^sub>R 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)" + "convex hull ((\x. c *\<^sub>R x) ` s) = (\x. c *\<^sub>R 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) + unfolding image_image scaleR_scaleR 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)" + "convex hull ((\x. a + c *\<^sub>R x) ` s) = (\x. a + c *\<^sub>R x) ` (convex hull s)" unfolding image_image[THEN sym] convex_hull_scaling convex_hull_translation .. subsection {* Convex set as intersection of halfspaces. *} @@ -1770,9 +1781,9 @@ 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" + shows "\u. setsum u c = 0 \ (\v\c. u v \ 0) \ setsum (\v. u v *\<^sub>R 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 + thus ?thesis apply(rule_tac x="\v. if v\s then u v else 0" in exI) unfolding if_smult scaleR_zero_left and setsum_restrict_set[OF assms(1), THEN sym] by(auto simp add: Int_absorb1) qed lemma radon_s_lemma: @@ -1793,9 +1804,9 @@ 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 + obtain u v where uv:"setsum u c = 0" "v\c" "u v \ 0" "(\v\c. u v *\<^sub>R 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}" + def z \ "(inverse (setsum u {x\c. u x > 0})) *\<^sub>R setsum (\x. u x *\<^sub>R 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") @@ -1807,23 +1818,23 @@ 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)" + "(\x\{x \ c. 0 < u x} \ {x \ c. u x < 0}. u x *\<^sub>R x) = (\x\c. u x *\<^sub>R x)" using assms(1) apply(rule_tac[!] 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)" + "(\x\{x \ c. 0 < u x}. u x *\<^sub>R x) = - (\x\{x \ c. 0 > u x}. u x *\<^sub>R x)" unfolding eq_neg_iff_add_eq_0 using uv(1,4) by (auto simp add: 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 + using assms(1) unfolding scaleR_scaleR[THEN sym] scaleR_right.setsum [symmetric] 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 * + using assms(1) unfolding scaleR_scaleR[THEN sym] scaleR_right.setsum [symmetric] 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 @@ -1889,10 +1900,10 @@ 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] + unfolding convex_def by(auto simp add: linear_cmul[OF assms, unfolded smult_conv_scaleR] 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]) + unfolding convex_def by (auto simp add: linear_cmul[OF assms, THEN sym, unfolded smult_conv_scaleR] linear_add[OF assms, THEN sym]) qed auto lemma in_convex_hull_linear_image: @@ -1904,51 +1915,51 @@ lemma compact_frontier_line_lemma: fixes s :: "(real ^ _) set" assumes "compact s" "0 \ s" "x \ 0" - obtains u where "0 \ u" "(u *s x) \ frontier s" "\v>u. (v *s x) \ s" + obtains u where "0 \ u" "(u *\<^sub>R x) \ frontier s" "\v>u. (v *\<^sub>R 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] + let ?A = "{y. \u. 0 \ u \ u \ b / norm(x) \ (y = u *\<^sub>R x)}" + have A:"?A = (\u. dest_vec1 u *\<^sub>R x) ` {0 .. vec1 (b / norm x)}" + unfolding image_image[of "\u. u *\<^sub>R 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) apply (rule continuous_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)]) + moreover have "{y. \u\0. u \ b / norm x \ y = u *\<^sub>R 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" + ultimately obtain u y where obt: "u\0" "u \ b / norm x" "y = u *\<^sub>R 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" + { fix v assume as:"v > u" "v *\<^sub>R 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 + using `u\0` unfolding pos_le_divide_eq[OF `norm x > 0`] and norm_scaleR by auto + hence "norm (v *\<^sub>R 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) + hence False unfolding obt(3) unfolding norm_scaleR 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" + have "u *\<^sub>R x \ frontier s" unfolding frontier_straddle apply(rule,rule,rule) apply(rule_tac x="u *\<^sub>R x" in bexI) unfolding obt(3)[THEN sym] + prefer 3 apply(rule_tac x="(u + (e / 2) / norm x) *\<^sub>R x" in exI) apply(rule, rule) proof- + fix e assume "0 < e" and as:"(u + e / 2 / norm x) *\<^sub>R 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_norm obt(3)) + qed(insert `y\s`, auto simp add: dist_norm scaleR_left_distrib obt(3) norm_scaleR) 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 )" + "\x\s. \u. 0 \ u \ u < 1 \ (u *\<^sub>R 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" + def pi \ "\x::real^'n. inverse (norm x) *\<^sub>R 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 injpi:"\x y. pi x = pi y \ norm x = norm y \ x = y" unfolding pi_def by (auto simp add: scaleR_cancel_left) have contpi:"continuous_on (UNIV - {0}) pi" apply(rule continuous_at_imp_continuous_on) apply rule unfolding pi_def @@ -1959,20 +1970,20 @@ apply (rule continuous_at_id) done 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 pi:"\x. x \ 0 \ pi x \ sphere" "\x u. u>0 \ pi (u *\<^sub>R x) = pi x" unfolding pi_def sphere_def by (auto simp add: norm_scaleR scaleR_cancel_right) 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) + have front_smul:"\x\frontier s. \u\0. u *\<^sub>R 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" + obtain v where v:"0 \ v" "v *\<^sub>R x \ frontier s" "\w>v. w *\<^sub>R 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"]] + assume "v>1" thus False using assms(3)[THEN bspec[where x="v *\<^sub>R 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") + show "u *\<^sub>R x \ s \ u \ 1" apply rule using v(3)[unfolded `v=1`, THEN spec[where x=u]] proof- + assume "u\1" thus "u *\<^sub>R 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" @@ -1982,14 +1993,14 @@ 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" + then obtain u where "0 \ u" "u *\<^sub>R x \ frontier s" "\v>u. v *\<^sub>R 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 + thus "x \ pi ` frontier s" unfolding image_iff le_less pi_def apply(rule_tac x="u *\<^sub>R x" in bexI) using `norm x = 1` `0\frontier s` by (auto simp add: norm_scaleR) 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 + from nor have x:"x = norm x *\<^sub>R ((inverse (norm y)) *\<^sub>R y)" unfolding as(3)[unfolded pi_def, THEN sym] by auto + from nor have y:"y = norm y *\<^sub>R ((inverse (norm x)) *\<^sub>R 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 @@ -2005,7 +2016,7 @@ 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") + have "norm x *\<^sub>R 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_norm) thus ?thesis apply(rule_tac assms(3)[rule_format, THEN DiffD1]) apply(rule_tac fs[unfolded subset_eq, rule_format]) @@ -2014,34 +2025,34 @@ 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") + hence "x \ (\x. norm x *\<^sub>R 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 **:"norm x *\<^sub>R (?a *\<^sub>R surf (pi x)) = x" apply(rule_tac scaleR_left_imp_eq[OF invn]) unfolding pi_def using invn 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))" + hence "norm x = norm ((?a * norm x) *\<^sub>R surf (pi x))" + unfolding norm_scaleR 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) *\<^sub>R 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_norm + hence "dist 0 (inverse (norm (surf (pi x))) *\<^sub>R x) \ 1" unfolding dist_norm 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`] + using False `x\s` by(auto simp add:field_simps norm_scaleR) + ultimately show ?thesis unfolding image_iff apply(rule_tac x="inverse (norm (surf(pi x))) *\<^sub>R x" in bexI) + apply(subst injpi[THEN sym]) unfolding norm_scaleR 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)"]) + show ?thesis apply(subst homeomorphic_sym) apply(rule homeomorphic_compact[where f="\x. norm x *\<^sub>R 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") + thus "continuous (at x) (\x. norm x *\<^sub>R surf (pi x))" proof(cases "x=0") case False thus ?thesis apply(rule_tac continuous_mul, rule_tac continuous_at_norm) 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] .. @@ -2050,7 +2061,7 @@ unfolding Ball_def mem_cball dist_norm 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_norm diff_0_right norm_mul abs_norm_cancel proof- + unfolding norm_0 scaleR_zero_left dist_norm diff_0_right norm_scaleR 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 @@ -2065,8 +2076,8 @@ 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)" + show "inj_on (\x. norm x *\<^sub>R 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 *\<^sub>R surf (pi x) = norm y *\<^sub>R 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 @@ -2074,7 +2085,7 @@ 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) + moreover have "norm x = norm y" using as(3)[unfolded *] using False by(auto dest:surf_0 simp add: scaleR_cancel_right) ultimately show ?thesis using injpi by auto qed qed qed auto qed @@ -2083,18 +2094,18 @@ 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) + hence "u *\<^sub>R x \ interior s" unfolding interior_def mem_Collect_eq + apply(rule_tac x="ball (u *\<^sub>R 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" + fix y assume "dist (u *\<^sub>R x) y < 1 - u" + hence "inverse (1 - u) *\<^sub>R (y - u *\<^sub>R x) \ s" using assms(3) apply(erule_tac subsetD) unfolding mem_cball dist_commute dist_norm - unfolding group_add_class.diff_0 group_add_class.diff_0_right norm_minus_cancel norm_mul + unfolding group_add_class.diff_0 group_add_class.diff_0_right norm_minus_cancel norm_scaleR 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 + thus "y \ s" using assms(1)[unfolded convex_def, rule_format, of "inverse(1 - u) *\<^sub>R (y - u *\<^sub>R x)" x "1 - u" u] + using as unfolding scaleR_scaleR by auto qed auto + thus "u *\<^sub>R 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" @@ -2102,16 +2113,16 @@ 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 + have "cball ?n 1 \ (\x. inverse d *\<^sub>R (x - a)) ` s" + apply(rule, rule_tac x="d *\<^sub>R x + a" in image_eqI) defer apply(rule d[unfolded subset_eq, rule_format]) using `d>0` unfolding mem_cball dist_norm - 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) + by(auto simp add: mult_right_le_one_le norm_scaleR) + hence "(\x. inverse d *\<^sub>R (x - a)) ` s homeomorphic cball ?n 1" + using homeomorphic_convex_compact_lemma[of "(\x. ?d *\<^sub>R -a + ?d *\<^sub>R x) ` s", OF convex_affinity compact_affinity] + using assms(1,2) by(auto simp add: uminus_add_conv_diff scaleR_right_diff_distrib) 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 + apply(rule homeomorphic_trans[OF homeomorphic_affinity[of "?d" s "?d *\<^sub>R -a"]]) + using `d>0` `e>0` by(auto simp add: uminus_add_conv_diff scaleR_right_diff_distrib) qed lemma homeomorphic_convex_compact: fixes s::"(real^'n::finite) set" and t::"(real^'n) set" assumes "convex s" "compact s" "interior s \ {}" @@ -2128,8 +2139,8 @@ 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 + unfolding mem_Collect_eq fstcart_pastecart sndcart_pastecart sndcart_add sndcart_cmul [where 'a=real, unfolded smult_conv_scaleR] fstcart_add fstcart_cmul [where 'a=real, unfolded smult_conv_scaleR] + unfolding Ball_def[symmetric] unfolding dest_vec1_add dest_vec1_cmul [where 'a=real, unfolded smult_conv_scaleR] 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" @@ -2158,11 +2169,11 @@ 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} ) " + f (setsum (\i. u i *\<^sub>R 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 + unfolding fstcart_pastecart sndcart_pastecart sndcart_add sndcart_cmul [where 'a=real, unfolded smult_conv_scaleR] fstcart_add fstcart_cmul [where 'a=real, unfolded smult_conv_scaleR] + unfolding dest_vec1_add dest_vec1_cmul [where 'a=real, unfolded smult_conv_scaleR] 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) @@ -2184,7 +2195,7 @@ 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_def, rule_format, OF as(1,2)]) + ultimately show "u *\<^sub>R x + v *\<^sub>R y \ s" apply- apply(rule assms[unfolded is_interval_def, rule_format, OF as(1,2)]) using as(3-) dimindex_ge_1 apply- by(auto simp add: vector_component) qed lemma is_interval_connected: @@ -2259,7 +2270,7 @@ 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" + 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 *\<^sub>R 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) @@ -2297,7 +2308,7 @@ 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 + case False hence *:"x = x$i *\<^sub>R (\ j. if x$j = 0 then 0 else 1) + (1 - x$i) *\<^sub>R (\ 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 @@ -2331,7 +2342,7 @@ 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) + have *:"{x - ?d .. x + ?d} = (\y. x - ?d + (2 * d) *\<^sub>R 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]] @@ -2341,19 +2352,19 @@ 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 + hence "inverse (2 * d) *\<^sub>R (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) + thus "\z\{0..1}. y = x - ?d + (2 * d) *\<^sub>R z" apply- apply(rule_tac x="inverse (2 * d) *\<^sub>R (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" + fix y z assume as:"z\{0..1}" "y = x - ?d + (2*d) *\<^sub>R 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 + thus ?thesis apply(rule_tac that[of "(\y. x - ?d + (2 * d) *\<^sub>R y)` s"]) unfolding * and convex_hull_affinity by auto qed subsection {* Bounded convex function on open set is continuous. *} @@ -2374,12 +2385,12 @@ 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_norm 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)" + { def w \ "x + t *\<^sub>R (y - x)" have "w\s" unfolding w_def apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_norm - unfolding t_def using `k>0` 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 + unfolding t_def using `k>0` by(auto simp add: norm_scaleR simp del: scaleR_right_diff_distrib) + have "(1 / t) *\<^sub>R x + - x + ((t - 1) / t) *\<^sub>R x = (1 / t - 1 + (t - 1) / t) *\<^sub>R x" by (auto simp add: algebra_simps) + also have "\ = 0" using `t>0` by(auto simp add:field_simps simp del:scaleR_left_distrib) + finally have w:"(1 / t) *\<^sub>R w + ((t - 1) / t) *\<^sub>R x = y" unfolding w_def using False and `t>0` by (auto simp add: algebra_simps) have "2 * B < e * t" 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) @@ -2387,12 +2398,12 @@ 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)" + { def w \ "x - t *\<^sub>R (y - x)" have "w\s" unfolding w_def apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_norm - unfolding t_def using `k>0` 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 + unfolding t_def using `k>0` by(auto simp add: norm_scaleR simp del: scaleR_right_diff_distrib) + have "(1 / (1 + t)) *\<^sub>R x + (t / (1 + t)) *\<^sub>R x = (1 / (1 + t) + t / (1 + t)) *\<^sub>R x" by (auto simp add: algebra_simps) + also have "\=x" using `t>0` by (auto simp add:field_simps simp del:scaleR_left_distrib) + finally have w:"(1 / (1+t)) *\<^sub>R w + (t / (1 + t)) *\<^sub>R y = x" unfolding w_def using False and `t>0` by (auto simp add: algebra_simps) 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" @@ -2411,10 +2422,10 @@ assumes "convex_on (cball x e) f" "\y \ 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 + fix y assume y:"y\cball x e" def z \ "2 *\<^sub>R x - y" + have *:"x - (2 *\<^sub>R x - y) = y - x" by vector have z:"z\cball x e" using y unfolding z_def mem_cball dist_norm * by(auto simp add: norm_minus_commute) - have "(1 / 2) *s y + (1 / 2) *s z = x" unfolding z_def by auto + have "(1 / 2) *\<^sub>R y + (1 / 2) *\<^sub>R z = x" unfolding z_def by (auto simp add: algebra_simps) 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" @@ -2463,11 +2474,17 @@ (* 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 + midpoint :: "real ^ 'n::finite \ real ^ 'n \ real ^ 'n" where + "midpoint a b = (inverse (2::real)) *\<^sub>R (a + b)" + +definition + open_segment :: "real ^ 'n::finite \ real ^ 'n \ (real ^ 'n) set" where + "open_segment a b = {(1 - u) *\<^sub>R a + u *\<^sub>R b | u::real. 0 < u \ u < 1}" + +definition + closed_segment :: "real ^ 'n::finite \ real ^ 'n \ (real ^ 'n) set" where + "closed_segment a b = {(1 - u) *\<^sub>R a + u *\<^sub>R b | u::real. 0 \ u \ u \ 1}" definition "between = (\ (a,b). closed_segment a b)" @@ -2476,9 +2493,9 @@ 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 + unfolding midpoint_def unfolding scaleR_right_distrib unfolding scaleR_left_distrib[THEN sym] by auto + +lemma midpoint_sym: "midpoint a b = midpoint b a" unfolding midpoint_def by (auto simp add: scaleR_right_distrib) lemma dist_midpoint: "dist a (midpoint a b) = (dist a b) / 2" (is ?t1) @@ -2486,8 +2503,9 @@ "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 + have *: "\x y::real^'n::finite. 2 *\<^sub>R x = - y \ norm x = (norm y) / 2" unfolding equation_minus_iff by (auto simp add: norm_scaleR) + have **:"\x y::real^'n::finite. 2 *\<^sub>R x = y \ norm x = (norm y) / 2" by (auto simp add: norm_scaleR) + note scaleR_right_distrib [simp] show ?t1 unfolding midpoint_def dist_norm apply (rule **) by(auto,vector) show ?t2 unfolding midpoint_def dist_norm apply (rule *) by(auto,vector) show ?t3 unfolding midpoint_def dist_norm apply (rule *) by(auto,vector) @@ -2534,7 +2552,7 @@ 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 segment_refl:"closed_segment a a = {a}" unfolding segment by (auto simp add: algebra_simps) lemma between_mem_segment: "between (a,b) x \ x \ closed_segment a b" unfolding between_def mem_def by auto @@ -2544,32 +2562,32 @@ case True thus ?thesis unfolding between_def split_conv mem_def[of x, symmetric] by(auto simp add:segment_refl dist_commute) 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 + have *:"\u. a - ((1 - u) *\<^sub>R a + u *\<^sub>R b) = u *\<^sub>R (a - b)" by (auto simp add: algebra_simps) 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) + fix u assume as:"x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \ u" "u \ 1" + hence *:"a - x = u *\<^sub>R (a - b)" "x - b = (1 - u) *\<^sub>R (a - b)" + unfolding as(1) by(auto simp add:algebra_simps) + show "norm (a - x) *\<^sub>R (x - b) = norm (x - b) *\<^sub>R (a - x)" + unfolding norm_minus_commute[of x a] * norm_scaleR 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_norm] 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) + thus "\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 \ u \ u \ 1" apply(rule_tac x="dist a x / dist a b" in exI) unfolding dist_norm 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 = + fix i::'n have "((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R 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_norm] 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 + finally show "x $ i = ((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) $ i" by auto qed(insert Fal2, auto) qed qed lemma between_midpoint: fixes a::"real^'n::finite" shows "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 +proof- have *:"\x y z. x = (1/2::real) *\<^sub>R z \ y = (1/2) *\<^sub>R z \ norm z = norm x + norm y" by (auto simp add: norm_scaleR) show ?t1 ?t2 unfolding between midpoint_def dist_norm apply(rule_tac[!] *) by(auto simp add:field_simps Cart_eq vector_component_simps) qed @@ -2581,16 +2599,16 @@ 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" + shows "x - e *\<^sub>R (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_norm unfolding norm_mul[THEN sym] apply(rule norm_eqI) using `e>0` + fix y assume as:"dist (x - e *\<^sub>R (x - c)) y < e * d" + have *:"y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" using `e>0` by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib) + have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = abs(1/e) * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" + unfolding dist_norm unfolding norm_scaleR[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 "\ = abs(1/e) * norm (x - e *\<^sub>R (x - c) - y)" by(auto intro!:norm_eqI simp add: algebra_simps) also have "\ < d" using as[unfolded dist_norm] and `e>0` by(auto simp add:pos_divide_less_eq[OF `e>0`] real_mult_commute) finally show "y \ s" apply(subst *) apply(rule assms(1)[unfolded convex_alt,rule_format]) @@ -2599,7 +2617,7 @@ 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" + shows "x - e *\<^sub>R (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 @@ -2614,11 +2632,11 @@ 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_norm 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 + def z \ "c + ((1 - e) / e) *\<^sub>R (x - y)" + have *:"x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" unfolding z_def using `e>0` by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) 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_norm using y and assms(4,5) - by(auto simp del:vector_ssub_ldistrib simp add:field_simps norm_minus_commute) + by(auto simp del:scaleR_right_diff_distrib simp add:field_simps norm_minus_commute norm_scaleR) thus ?thesis unfolding * apply - apply(rule mem_interior_convex_shrink) using assms(1,4-5) `y\s` by auto qed @@ -2626,7 +2644,7 @@ 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)}" + shows "convex hull (insert 0 s) = { y. (\u. (\x\s. 0 \ u x) \ setsum u s \ 1 \ setsum (\x. u x *\<^sub>R 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) @@ -2641,16 +2659,16 @@ 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 + 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 *\<^sub>R x) = x" + have *:"\i. u (basis i) = x$i" using as(3) unfolding sumbas and basis_expansion_unique [where 'a=real, unfolded smult_conv_scaleR] 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" + 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 *\<^sub>R 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 + unfolding sumbas using as(2) and basis_expansion_unique [where 'a=real, unfolded smult_conv_scaleR] by(auto simp add:dot_basis) qed qed lemma interior_std_simplex: "interior (convex hull (insert 0 { basis i| i. i\UNIV})) = @@ -2659,14 +2677,14 @@ 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_norm by(auto simp add: norm_basis vector_component_simps basis_component elim:allE[where x=i]) + fix i::'n show "0 < x $ i" using as[THEN spec[where x="x - (e / 2) *\<^sub>R basis i"]] and `e>0` + unfolding dist_norm by(auto simp add: norm_basis vector_component_simps basis_component norm_scaleR 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_norm 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 + have **:"dist x (x + (e / 2) *\<^sub>R basis a) < e" using `e>0` and norm_basis[of a] + unfolding dist_norm by(auto simp add: vector_component_simps basis_component norm_scaleR intro!: mult_strict_left_mono_comm) + have "\i. (x + (e / 2) *\<^sub>R 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) *\<^sub>R 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) *\<^sub>R 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 @@ -2692,8 +2710,8 @@ 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 + let ?D = "UNIV::'n set" let ?a = "setsum (\b::real^'n. inverse (2 * real CARD('n)) *\<^sub>R b) {(basis i) | i. i \ ?D}" + have *:"{basis i :: real ^ 'n | 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) @@ -2718,7 +2736,7 @@ 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))" + where "joinpaths g1 g2 = (\x. if dest_vec1 x \ ((1 / 2)::real) then g1 (2 *\<^sub>R x) else g2(2 *\<^sub>R 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)" @@ -2763,7 +2781,7 @@ 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) + have "2 *\<^sub>R 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 @@ -2786,9 +2804,9 @@ 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}" + have *:"g1 = (\x. g1 (2 *\<^sub>R x)) \ (\x. (1/2) *\<^sub>R x)" + "g2 = (\x. g2 (2 *\<^sub>R x - 1)) \ (\x. (1/2) *\<^sub>R (x + 1))" unfolding o_def by auto + have "op *\<^sub>R (1 / 2) ` {0::real^1..1} \ {0..1}" "(\x. (1 / 2) *\<^sub>R (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) @@ -2796,35 +2814,35 @@ 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}" + fix x assume "x \ op *\<^sub>R (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}" + thus "(g1 +++ g2) x = g1 (2 *\<^sub>R x)" unfolding joinpaths_def by auto next + fix x assume "x \ (\x. (1 / 2) *\<^sub>R (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 "(g1 +++ g2) x = g2 (2 *\<^sub>R x - 1)" proof(cases "dest_vec1 x = 1 / 2") + case True hence "x = (1/2) *\<^sub>R 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}" + have *:"{0 .. 1::real^1} = {0.. (1/2)*\<^sub>R 1} \ {(1/2) *\<^sub>R 1 .. 1}" by(auto simp add: vector_component_simps) + have **:"op *\<^sub>R 2 ` {0..(1 / 2) *\<^sub>R 1} = {0..1::real^1}" apply(rule set_ext, rule) unfolding image_iff + defer apply(rule_tac x="(1/2)*\<^sub>R x" in bexI) by(auto simp add: vector_component_simps) + have ***:"(\x. 2 *\<^sub>R x - 1) ` {(1 / 2) *\<^sub>R 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) + have ****:"\x::real^1. x $ 1 * 2 = 1 \ x = (1/2) *\<^sub>R 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 + show "continuous_on {0..(1 / 2) *\<^sub>R 1} (g1 +++ g2)" apply(rule continuous_on_eq[of _ "\x. g1 (2 *\<^sub>R 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 + show "continuous_on {(1/2)*\<^sub>R1..1} (g1 +++ g2)" apply(rule continuous_on_eq[of _ "g2 \ (\x. 2 *\<^sub>R 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))" + then obtain y where y:"y\{0..1}" "x = (if dest_vec1 y \ 1 / 2 then g1 (2 *\<^sub>R y) else g2 (2 *\<^sub>R 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) @@ -2841,12 +2859,12 @@ 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 + apply(rule_tac x="(1/2) *\<^sub>R 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] + apply(rule_tac x="(1/2) *\<^sub>R (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: @@ -2858,6 +2876,10 @@ 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 dest_vec1_scaleR [simp]: + "dest_vec1 (scaleR a x) = scaleR a (dest_vec1 x)" +unfolding dest_vec1_def by simp + 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}" @@ -2867,40 +2889,40 @@ 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 + hence "g1 (2 *\<^sub>R x) = g1 (2 *\<^sub>R y)" using xy(3) unfolding joinpaths_def dest_vec1_def by auto + moreover have "2 *\<^sub>R x \ {0..1}" "2 *\<^sub>R 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 + ultimately show ?thesis using inj(1)[of "2*\<^sub>R x" "2*\<^sub>R y"] by (auto simp add: scaleR_cancel_left) 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 + hence "g2 (2 *\<^sub>R x - 1) = g2 (2 *\<^sub>R y - 1)" using xy(3) unfolding joinpaths_def dest_vec1_def by auto + moreover have "2 *\<^sub>R x - 1 \ {0..1}" "2 *\<^sub>R 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 + ultimately show ?thesis using inj(2)[of "2*\<^sub>R x - 1" "2*\<^sub>R y - 1"] by (auto simp add: scaleR_cancel_left) 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] + using inj(2)[of "2 *\<^sub>R 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) + using inj(1)[of "2 *\<^sub>R 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) + using inj(2)[of "2 *\<^sub>R 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] + using inj(2)[of "2 *\<^sub>R 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) + using inj(1)[of "2 *\<^sub>R 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) + using inj(2)[of "2 *\<^sub>R x - 1" 1] by(auto simp add:vector_component_simps Cart_eq) ultimately show ?thesis by auto qed qed lemma injective_path_join: @@ -2911,22 +2933,22 @@ 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) + assume "x $ 1 \ 1 / 2" "y $ 1 \ 1 / 2" thus ?thesis using inj(1)[of "2*\<^sub>R x" "2*\<^sub>R y"] and xy + unfolding mem_interval_1 joinpaths_def by(auto simp add:vector_component_simps scaleR_cancel_left) + next assume "x $ 1 > 1 / 2" "y $ 1 > 1 / 2" thus ?thesis using inj(2)[of "2*\<^sub>R x - 1" "2*\<^sub>R y - 1"] and xy + unfolding mem_interval_1 joinpaths_def by(auto simp add:vector_component_simps scaleR_cancel_left) 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) + thus ?thesis using as and inj(1)[of "2 *\<^sub>R x" 1] inj(2)[of "2 *\<^sub>R 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) + thus ?thesis using as and inj(2)[of "2 *\<^sub>R x - 1" 0] inj(1)[of "2 *\<^sub>R 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 @@ -2993,7 +3015,7 @@ definition linepath :: "real ^ 'n::finite \ real ^ 'n \ real ^ 1 \ real ^ 'n" where - "linepath a b = (\x. (1 - dest_vec1 x) *s a + dest_vec1 x *s b)" + "linepath a b = (\x. (1 - dest_vec1 x) *\<^sub>R a + dest_vec1 x *\<^sub>R b)" lemma pathstart_linepath[simp]: "pathstart(linepath a b) = a" unfolding pathstart_def linepath_def by auto @@ -3013,7 +3035,7 @@ 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) + unfolding mem_Collect_eq image_iff apply(erule exE) apply(rule_tac x="u *\<^sub>R 1" in bexI) by(auto simp add:vector_component_simps) lemma reversepath_linepath[simp]: "reversepath(linepath a b) = linepath b a" @@ -3021,10 +3043,10 @@ 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" + fix x y::"real^1" assume "x $ 1 *\<^sub>R b + y $ 1 *\<^sub>R a = x $ 1 *\<^sub>R a + y $ 1 *\<^sub>R 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 + thus ?thesis unfolding injective_path_def linepath_def by(auto simp add:vector_component_simps algebra_simps) qed lemma simple_path_linepath[intro]: "a \ b \ simple_path(linepath a b)" by(auto intro!: injective_imp_simple_path injective_path_linepath) @@ -3164,7 +3186,7 @@ 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) + apply(rule_tac x="linepath a a" in exI) by(auto simp add:segment scaleR_left_diff_distrib) 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) @@ -3225,11 +3247,11 @@ 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 + case False hence *:"{x::real^'n. norm(x - a) = r} = (\x. a + r *\<^sub>R x) ` {x. norm x = 1}" unfolding not_le apply -apply(rule set_ext,rule) + unfolding image_iff apply(rule_tac x="(1/r) *\<^sub>R (x - a)" in bexI) unfolding mem_Collect_eq norm_scaleR by (auto simp add: scaleR_right_diff_distrib norm_scaleR) 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 **:"{x::real^'n. norm x = 1} = (\x. (1/norm x) *\<^sub>R x) ` (UNIV - {0})" apply(rule set_ext,rule) + unfolding image_iff apply(rule_tac x=x in bexI) unfolding mem_Collect_eq norm_scaleR by(auto intro!: *** simp add: norm_scaleR) have "continuous_on (UNIV - {0}) (\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_norm[unfolded o_def]) by auto diff -r 809dfb3d9c76 -r 0e4906ccf280 src/HOL/Library/Topology_Euclidean_Space.thy --- a/src/HOL/Library/Topology_Euclidean_Space.thy Thu Jun 11 20:19:26 2009 -0700 +++ b/src/HOL/Library/Topology_Euclidean_Space.thy Fri Jun 12 11:23:37 2009 -0700 @@ -748,7 +748,7 @@ { fix x have "x \ interior S \ x \ UNIV - (closure (UNIV - S))" unfolding interior_def closure_def islimpt_def - by blast + by blast (* FIXME: VERY slow! *) } thus ?thesis by blast @@ -1253,12 +1253,8 @@ lemma Lim_cmul: fixes f :: "'a \ real ^ 'n::finite" - shows "(f ---> l) net ==> ((\x. c *s f x) ---> c *s l) net" - apply (rule Lim_linear[where f = f]) - apply simp - apply (rule linear_compose_cmul) - apply (rule linear_id[unfolded id_def]) - done + shows "(f ---> l) net ==> ((\x. c *\<^sub>R f x) ---> c *\<^sub>R l) net" + by (intro tendsto_intros) lemma Lim_neg: fixes f :: "'a \ 'b::real_normed_vector" @@ -2099,9 +2095,11 @@ lemma bounded_scaling: fixes S :: "(real ^ 'n::finite) set" - shows "bounded S \ bounded ((\x. c *s x) ` S)" + shows "bounded S \ bounded ((\x. c *\<^sub>R x) ` S)" apply (rule bounded_linear_image, assumption) - by (rule linear_compose_cmul, rule linear_id[unfolded id_def]) + apply (simp only: linear_conv_bounded_linear) + apply (rule scaleR.bounded_linear_right) + done lemma bounded_translation: fixes S :: "'a::real_normed_vector set" @@ -3452,7 +3450,7 @@ lemma continuous_cmul: fixes f :: "'a::metric_space \ real ^ 'n::finite" - shows "continuous net f ==> continuous net (\x. c *s f x)" + shows "continuous net f ==> continuous net (\x. c *\<^sub>R f x)" by (auto simp add: continuous_def Lim_cmul) lemma continuous_neg: @@ -3478,7 +3476,7 @@ lemma continuous_on_cmul: fixes f :: "'a::metric_space \ real ^ _" - shows "continuous_on s f ==> continuous_on s (\x. c *s (f x))" + shows "continuous_on s f ==> continuous_on s (\x. c *\<^sub>R (f x))" unfolding continuous_on_eq_continuous_within using continuous_cmul by blast lemma continuous_on_neg: @@ -3507,12 +3505,12 @@ lemma uniformly_continuous_on_cmul: fixes f :: "'a::real_normed_vector \ real ^ _" assumes "uniformly_continuous_on s f" - shows "uniformly_continuous_on s (\x. c *s f(x))" + shows "uniformly_continuous_on s (\x. c *\<^sub>R f(x))" proof- { fix x y assume "((\n. f (x n) - f (y n)) ---> 0) sequentially" - hence "((\n. c *s f (x n) - c *s f (y n)) ---> 0) sequentially" + hence "((\n. c *\<^sub>R f (x n) - c *\<^sub>R f (y n)) ---> 0) sequentially" using Lim_cmul[of "(\n. f (x n) - f (y n))" 0 sequentially c] - unfolding vector_smult_rzero vector_ssub_ldistrib[of c] by auto + unfolding scaleR_zero_right scaleR_right_diff_distrib by auto } thus ?thesis using assms unfolding uniformly_continuous_on_sequentially by auto qed @@ -3837,20 +3835,20 @@ text{* Some arithmetical combinations (more to prove). *} lemma open_scaling[intro]: - fixes s :: "(real ^ _) set" + fixes s :: "'a::real_normed_vector set" assumes "c \ 0" "open s" - shows "open((\x. c *s x) ` s)" + shows "open((\x. c *\<^sub>R x) ` s)" proof- { fix x assume "x \ s" then obtain e where "e>0" and e:"\x'. dist x' x < e \ x' \ s" using assms(2)[unfolded open_dist, THEN bspec[where x=x]] by auto have "e * abs c > 0" using assms(1)[unfolded zero_less_abs_iff[THEN sym]] using real_mult_order[OF `e>0`] by auto moreover - { fix y assume "dist y (c *s x) < e * \c\" - hence "norm ((1 / c) *s y - x) < e" unfolding dist_norm - using norm_mul[of c "(1 / c) *s y - x", unfolded vector_ssub_ldistrib, unfolded vector_smult_assoc] assms(1) + { fix y assume "dist y (c *\<^sub>R x) < e * \c\" + hence "norm ((1 / c) *\<^sub>R y - x) < e" unfolding dist_norm + using norm_scaleR[of c "(1 / c) *\<^sub>R y - x", unfolded scaleR_right_diff_distrib, unfolded scaleR_scaleR] assms(1) assms(1)[unfolded zero_less_abs_iff[THEN sym]] by (simp del:zero_less_abs_iff) - hence "y \ op *s c ` s" using rev_image_eqI[of "(1 / c) *s y" s y "op *s c"] e[THEN spec[where x="(1 / c) *s y"]] assms(1) unfolding dist_norm vector_smult_assoc by auto } - ultimately have "\e>0. \x'. dist x' (c *s x) < e \ x' \ op *s c ` s" apply(rule_tac x="e * abs c" in exI) by auto } + hence "y \ op *\<^sub>R c ` s" using rev_image_eqI[of "(1 / c) *\<^sub>R y" s y "op *\<^sub>R c"] e[THEN spec[where x="(1 / c) *\<^sub>R y"]] assms(1) unfolding dist_norm scaleR_scaleR by auto } + ultimately have "\e>0. \x'. dist x' (c *\<^sub>R x) < e \ x' \ op *\<^sub>R c ` s" apply(rule_tac x="e * abs c" in exI) by auto } thus ?thesis unfolding open_dist by auto qed @@ -3860,9 +3858,10 @@ by (auto intro!: image_eqI [where f="\x. - x"]) lemma open_negations: - fixes s :: "(real ^ _) set" (* FIXME: generalize *) + fixes s :: "'a::real_normed_vector set" shows "open s ==> open ((\ x. -x) ` s)" - unfolding vector_sneg_minus1 by auto + unfolding scaleR_minus1_left [symmetric] + by (rule open_scaling, auto) lemma open_translation: fixes s :: "'a::real_normed_vector set" @@ -3876,11 +3875,11 @@ lemma open_affinity: fixes s :: "(real ^ _) set" assumes "open s" "c \ 0" - shows "open ((\x. a + c *s x) ` s)" + shows "open ((\x. a + c *\<^sub>R x) ` s)" proof- - have *:"(\x. a + c *s x) = (\x. a + x) \ (\x. c *s x)" unfolding o_def .. - have "op + a ` op *s c ` s = (op + a \ op *s c) ` s" by auto - thus ?thesis using assms open_translation[of "op *s c ` s" a] unfolding * by auto + have *:"(\x. a + c *\<^sub>R x) = (\x. a + x) \ (\x. c *\<^sub>R x)" unfolding o_def .. + have "op + a ` op *\<^sub>R c ` s = (op + a \ op *\<^sub>R c) ` s" by auto + thus ?thesis using assms open_translation[of "op *\<^sub>R c ` s" a] unfolding * by auto qed lemma interior_translation: @@ -4233,37 +4232,39 @@ subsection{* We can now extend limit compositions to consider the scalar multiplier. *} -lemma Lim_mul [tendsto_intros]: +lemma Lim_mul: fixes f :: "'a \ real ^ _" assumes "(c ---> d) net" "(f ---> l) net" - shows "((\x. c(x) *s f x) ---> (d *s l)) net" + shows "((\x. c(x) *\<^sub>R f x) ---> (d *\<^sub>R l)) net" unfolding smult_conv_scaleR using assms by (rule scaleR.tendsto) lemma Lim_vmul: - fixes c :: "'a \ real" - shows "(c ---> d) net ==> ((\x. c(x) *s v) ---> d *s v) net" - using Lim_mul[of c d net "\x. v" v] using Lim_const[of v] by auto + fixes c :: "'a \ real" and v :: "'b::real_normed_vector" + shows "(c ---> d) net ==> ((\x. c(x) *\<^sub>R v) ---> d *\<^sub>R v) net" + by (intro tendsto_intros) lemma continuous_vmul: - fixes c :: "'a::metric_space \ real" - shows "continuous net c ==> continuous net (\x. c(x) *s v)" + fixes c :: "'a::metric_space \ real" and v :: "'b::real_normed_vector" + shows "continuous net c ==> continuous net (\x. c(x) *\<^sub>R v)" unfolding continuous_def using Lim_vmul[of c] by auto lemma continuous_mul: fixes c :: "'a::metric_space \ real" + fixes f :: "'a::metric_space \ 'b::real_normed_vector" shows "continuous net c \ continuous net f - ==> continuous net (\x. c(x) *s f x) " - unfolding continuous_def using Lim_mul[of c] by auto + ==> continuous net (\x. c(x) *\<^sub>R f x) " + unfolding continuous_def by (intro tendsto_intros) lemma continuous_on_vmul: - fixes c :: "'a::metric_space \ real" - shows "continuous_on s c ==> continuous_on s (\x. c(x) *s v)" + fixes c :: "'a::metric_space \ real" and v :: "'b::real_normed_vector" + shows "continuous_on s c ==> continuous_on s (\x. c(x) *\<^sub>R v)" unfolding continuous_on_eq_continuous_within using continuous_vmul[of _ c] by auto lemma continuous_on_mul: fixes c :: "'a::metric_space \ real" + fixes f :: "'a::metric_space \ 'b::real_normed_vector" shows "continuous_on s c \ continuous_on s f - ==> continuous_on s (\x. c(x) *s f x)" + ==> continuous_on s (\x. c(x) *\<^sub>R f x)" unfolding continuous_on_eq_continuous_within using continuous_mul[of _ c] by auto text{* And so we have continuity of inverse. *} @@ -4378,7 +4379,7 @@ lemma compact_scaling: fixes s :: "(real ^ _) set" - assumes "compact s" shows "compact ((\x. c *s x) ` s)" + assumes "compact s" shows "compact ((\x. c *\<^sub>R x) ` s)" using assms unfolding smult_conv_scaleR by (rule compact_scaleR_image) lemma compact_negations: @@ -4416,9 +4417,9 @@ lemma compact_affinity: fixes s :: "(real ^ _) set" - assumes "compact s" shows "compact ((\x. a + c *s x) ` s)" + assumes "compact s" shows "compact ((\x. a + c *\<^sub>R x) ` s)" proof- - have "op + a ` op *s c ` s = (\x. a + c *s x) ` s" by auto + have "op + a ` op *\<^sub>R c ` s = (\x. a + c *\<^sub>R x) ` s" by auto thus ?thesis using compact_translation[OF compact_scaling[OF assms], of a c] by auto qed @@ -4525,7 +4526,7 @@ lemma closed_scaling: fixes s :: "(real ^ _) set" - assumes "closed s" shows "closed ((\x. c *s x) ` s)" + assumes "closed s" shows "closed ((\x. c *\<^sub>R x) ` s)" using assms unfolding smult_conv_scaleR by (rule closed_scaleR_image) lemma closed_negations: @@ -4685,10 +4686,10 @@ hence False using as by auto } moreover { assume as:"\i. \ (b$i \ a$i)" - let ?x = "(1/2) *s (a + b)" + let ?x = "(1/2) *\<^sub>R (a + b)" { fix i have "a$i < b$i" using as[THEN spec[where x=i]] by auto - hence "a$i < ((1/2) *s (a+b)) $ i" "((1/2) *s (a+b)) $ i < b$i" + hence "a$i < ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i < b$i" unfolding vector_smult_component and vector_add_component by (auto simp add: less_divide_eq_number_of1) } hence "{a <..< b} \ {}" using mem_interval(1)[of "?x" a b] by auto } @@ -4700,10 +4701,10 @@ hence False using as by auto } moreover { assume as:"\i. \ (b$i < a$i)" - let ?x = "(1/2) *s (a + b)" + let ?x = "(1/2) *\<^sub>R (a + b)" { fix i have "a$i \ b$i" using as[THEN spec[where x=i]] by auto - hence "a$i \ ((1/2) *s (a+b)) $ i" "((1/2) *s (a+b)) $ i \ b$i" + hence "a$i \ ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i \ b$i" unfolding vector_smult_component and vector_add_component by (auto simp add: less_divide_eq_number_of1) } hence "{a .. b} \ {}" using mem_interval(2)[of "?x" a b] by auto } @@ -4884,14 +4885,14 @@ then obtain s where s:"open s" "x \ s" "s \ {a..b}" by auto then obtain e where "e>0" and e:"\x'. dist x' x < e \ x' \ {a..b}" unfolding open_dist and subset_eq by auto { fix i - have "dist (x - (e / 2) *s basis i) x < e" - "dist (x + (e / 2) *s basis i) x < e" + have "dist (x - (e / 2) *\<^sub>R basis i) x < e" + "dist (x + (e / 2) *\<^sub>R basis i) x < e" unfolding dist_norm apply auto - unfolding norm_minus_cancel and norm_mul using norm_basis[of i] and `e>0` by auto - hence "a $ i \ (x - (e / 2) *s basis i) $ i" - "(x + (e / 2) *s basis i) $ i \ b $ i" - using e[THEN spec[where x="x - (e/2) *s basis i"]] - and e[THEN spec[where x="x + (e/2) *s basis i"]] + unfolding norm_minus_cancel and norm_scaleR using norm_basis[of i] and `e>0` by auto + hence "a $ i \ (x - (e / 2) *\<^sub>R basis i) $ i" + "(x + (e / 2) *\<^sub>R basis i) $ i \ b $ i" + using e[THEN spec[where x="x - (e/2) *\<^sub>R basis i"]] + and e[THEN spec[where x="x + (e/2) *\<^sub>R basis i"]] unfolding mem_interval by (auto elim!: allE[where x=i]) hence "a $ i < x $ i" and "x $ i < b $ i" unfolding vector_minus_component and vector_add_component @@ -4929,10 +4930,10 @@ using bounded_closed_imp_compact using bounded_interval[of a b] using closed_interval[of a b] by auto lemma open_interval_midpoint: fixes a :: "real^'n::finite" - assumes "{a<.. {}" shows "((1/2) *s (a + b)) \ {a<.. {}" shows "((1/2) *\<^sub>R (a + b)) \ {a<.. ((1 / 2) *s (a + b)) $ i < b $ i" + have "a $ i < ((1 / 2) *\<^sub>R (a + b)) $ i \ ((1 / 2) *\<^sub>R (a + b)) $ i < b $ i" using assms[unfolded interval_ne_empty, THEN spec[where x=i]] unfolding vector_smult_component and vector_add_component by(auto simp add: less_divide_eq_number_of1) } @@ -4941,7 +4942,7 @@ lemma open_closed_interval_convex: fixes x :: "real^'n::finite" assumes x:"x \ {a<.. {a .. b}" and e:"0 < e" "e \ 1" - shows "(e *s x + (1 - e) *s y) \ {a<..R x + (1 - e) *\<^sub>R y) \ {a<..R x + (1 - e) *\<^sub>R y) $ i" by auto moreover { have "b $ i = e * b$i + (1 - e) * b$i" unfolding left_diff_distrib by simp also have "\ > e * x $ i + (1 - e) * y $ i" apply(rule add_less_le_mono) @@ -4958,8 +4959,8 @@ using x unfolding mem_interval apply simp using y unfolding mem_interval apply simp done - finally have "(e *s x + (1 - e) *s y) $ i < b $ i" by auto - } ultimately have "a $ i < (e *s x + (1 - e) *s y) $ i \ (e *s x + (1 - e) *s y) $ i < b $ i" by auto } + finally have "(e *\<^sub>R x + (1 - e) *\<^sub>R y) $ i < b $ i" by auto + } ultimately have "a $ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) $ i \ (e *\<^sub>R x + (1 - e) *\<^sub>R y) $ i < b $ i" by auto } thus ?thesis unfolding mem_interval by auto qed @@ -4968,13 +4969,14 @@ shows "closure {a<.. {a .. b}" - def f == "\n::nat. x + (inverse (real n + 1)) *s (?c - x)" + def f == "\n::nat. x + (inverse (real n + 1)) *\<^sub>R (?c - x)" { fix n assume fn:"f n < b \ a < f n \ f n = x" and xc:"x \ ?c" have *:"0 < inverse (real n + 1)" "inverse (real n + 1) \ 1" unfolding inverse_le_1_iff by auto - have "(inverse (real n + 1)) *s ((1 / 2) *s (a + b)) + (1 - inverse (real n + 1)) *s x = - x + (inverse (real n + 1)) *s ((1 / 2 *s (a + b)) - x)" by (auto simp add: vector_ssub_ldistrib vector_add_ldistrib field_simps vector_sadd_rdistrib[THEN sym]) + have "(inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b)) + (1 - inverse (real n + 1)) *\<^sub>R x = + x + (inverse (real n + 1)) *\<^sub>R (((1 / 2) *\<^sub>R (a + b)) - x)" + by (auto simp add: algebra_simps) hence "f n < b" and "a < f n" using open_closed_interval_convex[OF open_interval_midpoint[OF assms] as *] unfolding f_def by auto hence False using fn unfolding f_def using xc by(auto simp add: vector_mul_lcancel vector_ssub_ldistrib) } moreover @@ -4987,8 +4989,8 @@ hence "((\n. inverse (real n + 1)) ---> 0) sequentially" unfolding Lim_sequentially by(auto simp add: dist_norm) hence "(f ---> x) sequentially" unfolding f_def - using Lim_add[OF Lim_const, of "\n::nat. (inverse (real n + 1)) *s ((1 / 2) *s (a + b) - x)" 0 sequentially x] - using Lim_vmul[of "\n::nat. inverse (real n + 1)" 0 sequentially "((1 / 2) *s (a + b) - x)"] by auto } + using Lim_add[OF Lim_const, of "\n::nat. (inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x)" 0 sequentially x] + using Lim_vmul[of "\n::nat. inverse (real n + 1)" 0 sequentially "((1 / 2) *\<^sub>R (a + b) - x)"] by auto } ultimately have "x \ closure {a<.. 0" shows "s homeomorphic ((\x. c *s x) ` s)" + assumes "c \ 0" shows "s homeomorphic ((\x. c *\<^sub>R x) ` s)" unfolding homeomorphic_minimal - apply(rule_tac x="\x. c *s x" in exI) - apply(rule_tac x="\x. (1 / c) *s x" in exI) - apply auto unfolding vector_smult_assoc using assms apply auto + apply(rule_tac x="\x. c *\<^sub>R x" in exI) + apply(rule_tac x="\x. (1 / c) *\<^sub>R x" in exI) + using assms apply auto using continuous_on_cmul[OF continuous_on_id] by auto lemma homeomorphic_translation: @@ -5463,13 +5465,13 @@ using continuous_on_add[OF continuous_on_const continuous_on_id] by auto lemma homeomorphic_affinity: - assumes "c \ 0" shows "s homeomorphic ((\x. a + c *s x) ` s)" + assumes "c \ 0" shows "s homeomorphic ((\x. a + c *\<^sub>R x) ` s)" proof- - have *:"op + a ` op *s c ` s = (\x. a + c *s x) ` s" by auto + have *:"op + a ` op *\<^sub>R c ` s = (\x. a + c *\<^sub>R x) ` s" by auto show ?thesis using homeomorphic_trans using homeomorphic_scaling[OF assms, of s] - using homeomorphic_translation[of "(\x. c *s x) ` s" a] unfolding * by auto + using homeomorphic_translation[of "(\x. c *\<^sub>R x) ` s" a] unfolding * by auto qed lemma homeomorphic_balls: fixes a b ::"real^'a::finite" @@ -5479,27 +5481,23 @@ proof- have *:"\e / d\ > 0" "\d / e\ >0" using assms using divide_pos_pos by auto show ?th unfolding homeomorphic_minimal - apply(rule_tac x="\x. b + (e/d) *s (x - a)" in exI) - apply(rule_tac x="\x. a + (d/e) *s (x - b)" in exI) - apply (auto simp add: dist_commute) unfolding dist_norm and vector_smult_assoc using assms apply auto - unfolding norm_minus_cancel and norm_mul - using continuous_on_add[OF continuous_on_const continuous_on_cmul[OF continuous_on_sub[OF continuous_on_id continuous_on_const]]] - apply (auto simp add: dist_commute) - using pos_less_divide_eq[OF *(1), THEN sym] unfolding real_mult_commute[of _ "\e / d\"] - using pos_less_divide_eq[OF *(2), THEN sym] unfolding real_mult_commute[of _ "\d / e\"] - by (auto simp add: dist_commute) + apply(rule_tac x="\x. b + (e/d) *\<^sub>R (x - a)" in exI) + apply(rule_tac x="\x. a + (d/e) *\<^sub>R (x - b)" in exI) + using assms apply (auto simp add: dist_commute) + unfolding dist_norm + apply (auto simp add: norm_scaleR pos_divide_less_eq mult_strict_left_mono) + unfolding continuous_on + by (intro ballI tendsto_intros, simp, assumption)+ next have *:"\e / d\ > 0" "\d / e\ >0" using assms using divide_pos_pos by auto show ?cth unfolding homeomorphic_minimal - apply(rule_tac x="\x. b + (e/d) *s (x - a)" in exI) - apply(rule_tac x="\x. a + (d/e) *s (x - b)" in exI) - apply (auto simp add: dist_commute) unfolding dist_norm and vector_smult_assoc using assms apply auto - unfolding norm_minus_cancel and norm_mul - using continuous_on_add[OF continuous_on_const continuous_on_cmul[OF continuous_on_sub[OF continuous_on_id continuous_on_const]]] - apply auto - using pos_le_divide_eq[OF *(1), THEN sym] unfolding real_mult_commute[of _ "\e / d\"] - using pos_le_divide_eq[OF *(2), THEN sym] unfolding real_mult_commute[of _ "\d / e\"] - by auto + apply(rule_tac x="\x. b + (e/d) *\<^sub>R (x - a)" in exI) + apply(rule_tac x="\x. a + (d/e) *\<^sub>R (x - b)" in exI) + using assms apply (auto simp add: dist_commute) + unfolding dist_norm + apply (auto simp add: norm_scaleR pos_divide_le_eq) + unfolding continuous_on + by (intro ballI tendsto_intros, simp, assumption)+ qed text{* "Isometry" (up to constant bounds) of injective linear map etc. *} @@ -5587,10 +5585,10 @@ next case False hence *:"0 < norm a / norm x" using `a\0` unfolding zero_less_norm_iff[THEN sym] by(simp only: divide_pos_pos) - have "\c. \x\s. c *s x \ s" using s[unfolded subspace_def] by auto - hence "(norm a / norm x) *s x \ {x \ s. norm x = norm a}" using `x\s` and `x\0` by auto - thus "norm (f b) / norm b * norm x \ norm (f x)" using b[THEN bspec[where x="(norm a / norm x) *s x"]] - unfolding linear_cmul[OF f(1)] and norm_mul and ba using `x\0` `a\0` + have "\c. \x\s. c *\<^sub>R x \ s" using s[unfolded subspace_def smult_conv_scaleR] by auto + hence "(norm a / norm x) *\<^sub>R x \ {x \ s. norm x = norm a}" using `x\s` and `x\0` by (auto simp add: norm_scaleR) + thus "norm (f b) / norm b * norm x \ norm (f x)" using b[THEN bspec[where x="(norm a / norm x) *\<^sub>R x"]] + unfolding linear_cmul[OF f(1), unfolded smult_conv_scaleR] and norm_scaleR and ba using `x\0` `a\0` by (auto simp add: real_mult_commute pos_le_divide_eq pos_divide_le_eq) qed } ultimately @@ -5654,8 +5652,8 @@ case (insert k F) hence *:"\i. i \ insert k F \ x $ i = 0" by auto have **:"F \ insert k F" by auto - def y \ "x - x$k *s basis k" - have y:"x = y + (x$k) *s basis k" unfolding y_def by auto + def y \ "x - x$k *\<^sub>R basis k" + have y:"x = y + (x$k) *\<^sub>R basis k" unfolding y_def by auto { fix i assume i':"i \ F" hence "y $ i = 0" unfolding y_def unfolding vector_minus_component and vector_smult_component and basis_component @@ -5665,9 +5663,10 @@ using image_mono[OF **, of basis] by auto moreover have "basis k \ span (?bas ` (insert k F))" by(rule span_superset, auto) - hence "x$k *s basis k \ span (?bas ` (insert k F))" using span_mul by auto + hence "x$k *\<^sub>R basis k \ span (?bas ` (insert k F))" + using span_mul [where 'a=real, unfolded smult_conv_scaleR] by auto ultimately - have "y + x$k *s basis k \ span (?bas ` (insert k F))" + have "y + x$k *\<^sub>R basis k \ span (?bas ` (insert k F))" using span_add by auto thus ?case using y by auto qed @@ -5788,45 +5787,45 @@ lemma image_affinity_interval: fixes m::real fixes a b c :: "real^'n::finite" - shows "(\x. m *s x + c) ` {a .. b} = + shows "(\x. m *\<^sub>R x + c) ` {a .. b} = (if {a .. b} = {} then {} - else (if 0 \ m then {m *s a + c .. m *s b + c} - else {m *s b + c .. m *s a + c}))" + else (if 0 \ m then {m *\<^sub>R a + c .. m *\<^sub>R b + c} + else {m *\<^sub>R b + c .. m *\<^sub>R a + c}))" proof(cases "m=0") { fix x assume "x \ c" "c \ x" hence "x=c" unfolding vector_less_eq_def and Cart_eq by (auto intro: order_antisym) } moreover case True - moreover have "c \ {m *s a + c..m *s b + c}" unfolding True by(auto simp add: vector_less_eq_def) + moreover have "c \ {m *\<^sub>R a + c..m *\<^sub>R b + c}" unfolding True by(auto simp add: vector_less_eq_def) ultimately show ?thesis by auto next case False { fix y assume "a \ y" "y \ b" "m > 0" - hence "m *s a + c \ m *s y + c" "m *s y + c \ m *s b + c" + hence "m *\<^sub>R a + c \ m *\<^sub>R y + c" "m *\<^sub>R y + c \ m *\<^sub>R b + c" unfolding vector_less_eq_def by(auto simp add: vector_smult_component vector_add_component) } moreover { fix y assume "a \ y" "y \ b" "m < 0" - hence "m *s b + c \ m *s y + c" "m *s y + c \ m *s a + c" + hence "m *\<^sub>R b + c \ m *\<^sub>R y + c" "m *\<^sub>R y + c \ m *\<^sub>R a + c" unfolding vector_less_eq_def by(auto simp add: vector_smult_component vector_add_component mult_left_mono_neg elim!:ballE) } moreover - { fix y assume "m > 0" "m *s a + c \ y" "y \ m *s b + c" - hence "y \ (\x. m *s x + c) ` {a..b}" + { fix y assume "m > 0" "m *\<^sub>R a + c \ y" "y \ m *\<^sub>R b + c" + hence "y \ (\x. m *\<^sub>R x + c) ` {a..b}" unfolding image_iff Bex_def mem_interval vector_less_eq_def apply(auto simp add: vector_smult_component vector_add_component vector_minus_component vector_smult_assoc pth_3[symmetric] - intro!: exI[where x="(1 / m) *s (y - c)"]) + intro!: exI[where x="(1 / m) *\<^sub>R (y - c)"]) by(auto simp add: pos_le_divide_eq pos_divide_le_eq real_mult_commute diff_le_iff) } moreover - { fix y assume "m *s b + c \ y" "y \ m *s a + c" "m < 0" - hence "y \ (\x. m *s x + c) ` {a..b}" + { fix y assume "m *\<^sub>R b + c \ y" "y \ m *\<^sub>R a + c" "m < 0" + hence "y \ (\x. m *\<^sub>R x + c) ` {a..b}" unfolding image_iff Bex_def mem_interval vector_less_eq_def apply(auto simp add: vector_smult_component vector_add_component vector_minus_component vector_smult_assoc pth_3[symmetric] - intro!: exI[where x="(1 / m) *s (y - c)"]) + intro!: exI[where x="(1 / m) *\<^sub>R (y - c)"]) by(auto simp add: neg_le_divide_eq neg_divide_le_eq real_mult_commute diff_le_iff) } ultimately show ?thesis using False by auto qed -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})" +lemma image_smult_interval:"(\x. m *\<^sub>R (x::real^'n::finite)) ` {a..b} = + (if {a..b} = {} then {} else if 0 \ m then {m *\<^sub>R a..m *\<^sub>R b} else {m *\<^sub>R b..m *\<^sub>R a})" using image_affinity_interval[of m 0 a b] by auto subsection{* Banach fixed point theorem (not really topological...) *}