# HG changeset patch # User haftmann # Date 1244903355 -7200 # Node ID ef6d67b1ad1066040095a1d17a1e11218cef273a # Parent dc65b2f786647a68f1dad36adc84c78592197488# Parent 78ac5c304db7023d4cb6ac638ea19b1e9f62ce59 merged diff -r 78ac5c304db7 -r ef6d67b1ad10 Admin/isatest/annomaly.ML --- a/Admin/isatest/annomaly.ML Thu Jun 11 19:49:02 2009 +0200 +++ b/Admin/isatest/annomaly.ML Sat Jun 13 16:29:15 2009 +0200 @@ -1,5 +1,4 @@ (* Title: Admin/isatest/annomaly.ML - ID: $Id$ Author: Martin von Gagern *) diff -r 78ac5c304db7 -r ef6d67b1ad10 Admin/isatest/isatest-annomaly --- a/Admin/isatest/isatest-annomaly Thu Jun 11 19:49:02 2009 +0200 +++ b/Admin/isatest/isatest-annomaly Sat Jun 13 16:29:15 2009 +0200 @@ -1,7 +1,5 @@ #!/usr/bin/env bash # -# $Id$ -# # Create AnnoMaLy documentation for Isabelle # # Based on http://martin.von-gagern.net/projects/annomaly/ diff -r 78ac5c304db7 -r ef6d67b1ad10 Admin/isatest/isatest-check --- a/Admin/isatest/isatest-check Thu Jun 11 19:49:02 2009 +0200 +++ b/Admin/isatest/isatest-check Sat Jun 13 16:29:15 2009 +0200 @@ -1,6 +1,5 @@ #!/usr/bin/env bash # -# $Id$ # Author: Gerwin Klein, TU Muenchen # # DESCRIPTION: sends email for failed tests, checks for error.log, diff -r 78ac5c304db7 -r ef6d67b1ad10 Admin/isatest/isatest-doc --- a/Admin/isatest/isatest-doc Thu Jun 11 19:49:02 2009 +0200 +++ b/Admin/isatest/isatest-doc Sat Jun 13 16:29:15 2009 +0200 @@ -1,6 +1,5 @@ #!/usr/bin/env bash # -# $Id$ # Author: Gerwin Klein, NICTA # # Run IsaMakefile for every Doc/ subdirectory. diff -r 78ac5c304db7 -r ef6d67b1ad10 Admin/isatest/isatest-lint --- a/Admin/isatest/isatest-lint Thu Jun 11 19:49:02 2009 +0200 +++ b/Admin/isatest/isatest-lint Sat Jun 13 16:29:15 2009 +0200 @@ -1,6 +1,5 @@ #!/usr/bin/env perl # -# $Id$ # Author: Florian Haftmann, TUM # # Do consistency and quality checks on the isabelle sources diff -r 78ac5c304db7 -r ef6d67b1ad10 Admin/isatest/isatest-makeall --- a/Admin/isatest/isatest-makeall Thu Jun 11 19:49:02 2009 +0200 +++ b/Admin/isatest/isatest-makeall Sat Jun 13 16:29:15 2009 +0200 @@ -1,6 +1,5 @@ #!/usr/bin/env bash # -# $Id$ # Author: Gerwin Klein, TU Muenchen # # DESCRIPTION: Run isabelle makeall from specified distribution and settings. @@ -71,14 +70,19 @@ NICE="nice" ;; + macbroy2) + MFLAGS="" + NICE="" + ;; + macbroy5) MFLAGS="-j 2" NICE="" ;; macbroy23) - MFLAGS="" - NICE="" + MFLAGS="-j 2" + NICE="nice" ;; macbroy2[0-9]) diff -r 78ac5c304db7 -r ef6d67b1ad10 Admin/isatest/isatest-makedist --- a/Admin/isatest/isatest-makedist Thu Jun 11 19:49:02 2009 +0200 +++ b/Admin/isatest/isatest-makedist Sat Jun 13 16:29:15 2009 +0200 @@ -1,6 +1,5 @@ #!/usr/bin/env bash # -# $Id$ # Author: Gerwin Klein, TU Muenchen # # DESCRIPTION: Build distribution and run isatest-make for lots of platforms. @@ -14,7 +13,6 @@ MAKEDIST=$HOME/bin/makedist MAKEALL=$HOME/bin/isatest-makeall TAR=tar -CVS2CL="$HOME/bin/cvs2cl --follow-only TRUNK" SSH="ssh -f" @@ -46,7 +44,6 @@ rm -f $RUNNING/*.runnning export DISTPREFIX -export CVS2CL DATE=$(date "+%Y-%m-%d") DISTLOG=$LOGPREFIX/isatest-makedist-$DATE.log @@ -109,9 +106,11 @@ sleep 15 $SSH atbroy101 "$MAKEALL $HOME/settings/at64-poly" sleep 15 +$SSH macbroy2 "$MAKEALL $HOME/settings/at-mac-poly-5.1-para; $MAKEALL $HOME/settings/mac-poly-M8" +sleep 15 $SSH macbroy5 "$MAKEALL $HOME/settings/mac-poly" sleep 15 -$SSH macbroy6 "/usr/stud/isatest/bin/isatest-makeall $HOME/settings/at-mac-poly-5.1-para" +$SSH macbroy6 "$MAKEALL $HOME/settings/mac-poly-M4" sleep 15 $SSH atbroy51 "$HOME/admin/isatest/isatest-annomaly" diff -r 78ac5c304db7 -r ef6d67b1ad10 Admin/isatest/isatest-settings --- a/Admin/isatest/isatest-settings Thu Jun 11 19:49:02 2009 +0200 +++ b/Admin/isatest/isatest-settings Sat Jun 13 16:29:15 2009 +0200 @@ -1,5 +1,5 @@ # -*- shell-script -*- :mode=shellscript: -# $Id$ +# # Author: Gerwin Klein, NICTA # # DESCRIPTION: common settings for the isatest-* scripts diff -r 78ac5c304db7 -r ef6d67b1ad10 Admin/isatest/pmail --- a/Admin/isatest/pmail Thu Jun 11 19:49:02 2009 +0200 +++ b/Admin/isatest/pmail Sat Jun 13 16:29:15 2009 +0200 @@ -1,6 +1,5 @@ #!/usr/bin/env bash # -# $Id$ # Author: Gerwin Klein, TU Muenchen # # DESCRIPTION: send email with text attachments. diff -r 78ac5c304db7 -r ef6d67b1ad10 Admin/isatest/settings/mac-poly-M4 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Admin/isatest/settings/mac-poly-M4 Sat Jun 13 16:29:15 2009 +0200 @@ -0,0 +1,28 @@ +# -*- shell-script -*- :mode=shellscript: + + POLYML_HOME="/home/polyml/polyml-svn" + ML_SYSTEM="polyml-experimental" + ML_PLATFORM="x86-darwin" + ML_HOME="$POLYML_HOME/$ML_PLATFORM" + ML_OPTIONS="--mutable 800 --immutable 2000" + + +ISABELLE_HOME_USER=~/isabelle-mac-poly-M4 + +# Where to look for isabelle tools (multiple dirs separated by ':'). +ISABELLE_TOOLS="$ISABELLE_HOME/lib/Tools" + +# Location for temporary files (should be on a local file system). +ISABELLE_TMP_PREFIX="/tmp/isabelle-$USER" + + +# Heap input locations. ML system identifier is included in lookup. +ISABELLE_PATH="$ISABELLE_HOME_USER/heaps:$ISABELLE_HOME/heaps" + +# Heap output location. ML system identifier is appended automatically later on. +ISABELLE_OUTPUT="$ISABELLE_HOME_USER/heaps" +ISABELLE_BROWSER_INFO="$ISABELLE_HOME_USER/browser_info" + +ISABELLE_USEDIR_OPTIONS="-i false -d false -M 4" + +HOL_USEDIR_OPTIONS="-p 2 -Q false" diff -r 78ac5c304db7 -r ef6d67b1ad10 Admin/isatest/settings/mac-poly-M8 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Admin/isatest/settings/mac-poly-M8 Sat Jun 13 16:29:15 2009 +0200 @@ -0,0 +1,28 @@ +# -*- shell-script -*- :mode=shellscript: + + POLYML_HOME="/home/polyml/polyml-svn" + ML_SYSTEM="polyml-experimental" + ML_PLATFORM="x86-darwin" + ML_HOME="$POLYML_HOME/$ML_PLATFORM" + ML_OPTIONS="--mutable 800 --immutable 2000" + + +ISABELLE_HOME_USER=~/isabelle-mac-poly-M8 + +# Where to look for isabelle tools (multiple dirs separated by ':'). +ISABELLE_TOOLS="$ISABELLE_HOME/lib/Tools" + +# Location for temporary files (should be on a local file system). +ISABELLE_TMP_PREFIX="/tmp/isabelle-$USER" + + +# Heap input locations. ML system identifier is included in lookup. +ISABELLE_PATH="$ISABELLE_HOME_USER/heaps:$ISABELLE_HOME/heaps" + +# Heap output location. ML system identifier is appended automatically later on. +ISABELLE_OUTPUT="$ISABELLE_HOME_USER/heaps" +ISABELLE_BROWSER_INFO="$ISABELLE_HOME_USER/browser_info" + +ISABELLE_USEDIR_OPTIONS="-i false -d false -M 8" + +HOL_USEDIR_OPTIONS="-p 2 -Q false" diff -r 78ac5c304db7 -r ef6d67b1ad10 src/HOL/Library/Convex_Euclidean_Space.thy --- a/src/HOL/Library/Convex_Euclidean_Space.thy Thu Jun 11 19:49:02 2009 +0200 +++ b/src/HOL/Library/Convex_Euclidean_Space.thy Sat Jun 13 16:29:15 2009 +0200 @@ -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 @@ -379,28 +384,30 @@ lemma convex_Int: "convex s \ convex t \ convex (s \ t)" unfolding convex_def by auto -lemma convex_halfspace_le: "convex {x. a \ x \ b}" +lemma convex_halfspace_le: "convex {x. inner a x \ b}" unfolding convex_def apply auto - unfolding dot_radd dot_rmult by (metis real_convex_bound_le) - -lemma convex_halfspace_ge: "convex {x. a \ x \ b}" -proof- have *:"{x. a \ x \ b} = {x. -a \ x \ -b}" by auto + unfolding inner_add inner_scaleR + by (metis real_convex_bound_le) + +lemma convex_halfspace_ge: "convex {x. inner a x \ b}" +proof- have *:"{x. inner a x \ b} = {x. inner (-a) x \ -b}" by auto show ?thesis apply(unfold *) using convex_halfspace_le[of "-a" "-b"] by auto qed -lemma convex_hyperplane: "convex {x. a \ x = b}" +lemma convex_hyperplane: "convex {x. inner a x = b}" proof- - have *:"{x. a \ x = b} = {x. a \ x \ b} \ {x. a \ x \ b}" by auto + have *:"{x. inner a x = b} = {x. inner a x \ b} \ {x. inner a x \ b}" by auto show ?thesis unfolding * apply(rule convex_Int) using convex_halfspace_le convex_halfspace_ge by auto qed -lemma convex_halfspace_lt: "convex {x. a \ x < b}" - unfolding convex_def by(auto simp add: real_convex_bound_lt dot_radd dot_rmult) - -lemma convex_halfspace_gt: "convex {x. a \ x > b}" - using convex_halfspace_lt[of "-a" "-b"] by(auto simp add: dot_lneg neg_less_iff_less) - -lemma convex_positive_orthant: "convex {x::real^'n. (\i. 0 \ x$i)}" +lemma convex_halfspace_lt: "convex {x. inner a x < b}" + unfolding convex_def + by(auto simp add: real_convex_bound_lt inner_add) + +lemma convex_halfspace_gt: "convex {x. inner a x > b}" + using convex_halfspace_lt[of "-a" "-b"] by auto + +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] 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)"] 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,41 +1215,65 @@ 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 qed +lemma open_dest_vec1_vimage: "open S \ open (dest_vec1 -` S)" +unfolding open_vector_def all_1 +by (auto simp add: dest_vec1_def) + +lemma tendsto_dest_vec1 [tendsto_intros]: + "(f ---> l) net \ ((\x. dest_vec1 (f x)) ---> dest_vec1 l) net" + unfolding tendsto_def + apply clarify + apply (drule_tac x="dest_vec1 -` S" in spec) + apply (simp add: open_dest_vec1_vimage) + done + +lemma continuous_dest_vec1: "continuous net f \ continuous net (\x. dest_vec1 (f x))" + unfolding continuous_def by (rule tendsto_dest_vec1) + +(* TODO: move *) +lemma compact_real_interval: + fixes a b :: real shows "compact {a..b}" +proof - + have "continuous_on {vec1 a .. vec1 b} dest_vec1" + unfolding continuous_on + by (simp add: tendsto_dest_vec1 Lim_at_within Lim_ident_at) + moreover have "compact {vec1 a .. vec1 b}" by (rule compact_interval) + ultimately have "compact (dest_vec1 ` {vec1 a .. vec1 b})" + by (rule compact_continuous_image) + also have "dest_vec1 ` {vec1 a .. vec1 b} = {a..b}" + by (auto simp add: image_def Bex_def exists_vec1) + finally show ?thesis . +qed lemma compact_convex_combinations: - fixes s t :: "(real ^ _) set" + 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 = "{ pastecart u w | u w. u \ {vec1 0 .. vec1 1} \ w \ { pastecart x y |x y. x \ s \ y \ t} }" - let ?h = "(\z. (1 - dest_vec1(fstcart z)) *s fstcart(sndcart z) + dest_vec1(fstcart z) *s sndcart(sndcart z))" - have *:"{ (1 - u) *s x + u *s y | x y u. 0 \ u \ u \ 1 \ x \ s \ y \ t} = ?h ` ?X" - apply(rule set_ext) unfolding image_iff mem_Collect_eq unfolding mem_interval_1 vec1_dest_vec1 - apply rule apply auto apply(rule_tac x="pastecart (vec1 u) (pastecart xa y)" in exI) apply simp - apply(rule_tac x="vec1 u" in exI) apply(rule_tac x="pastecart xa y" in exI) by auto - { fix u::"real^1" fix x y assume as:"0 \ dest_vec1 u" "dest_vec1 u \ 1" "x \ s" "y \ t" - hence "continuous (at (pastecart u (pastecart x y))) - (\z. fstcart (sndcart z) - dest_vec1 (fstcart z) *s fstcart (sndcart z) + - dest_vec1 (fstcart z) *s sndcart (sndcart z))" - apply (auto intro!: continuous_add continuous_sub continuous_mul simp add: o_def vec1_dest_vec1) - using linear_continuous_at linear_fstcart linear_sndcart linear_sndcart - using linear_compose[unfolded o_def] by auto } - hence "continuous_on {pastecart u w |u w. u \ {vec1 0..vec1 1} \ w \ {pastecart x y |x y. x \ s \ y \ t}} - (\z. (1 - dest_vec1 (fstcart z)) *s fstcart (sndcart z) + dest_vec1 (fstcart z) *s sndcart (sndcart z))" - apply(rule_tac continuous_at_imp_continuous_on) unfolding mem_Collect_eq - unfolding mem_interval_1 vec1_dest_vec1 by auto - thus ?thesis unfolding * apply(rule compact_continuous_image) - defer apply(rule compact_pastecart) defer apply(rule compact_pastecart) - using compact_interval assms by auto + let ?X = "{0..1} \ s \ t" + let ?h = "(\z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" + have *:"{ (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ s \ y \ t} = ?h ` ?X" + apply(rule set_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) *\<^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) + apply (intro compact_Times compact_real_interval assms) + done qed lemma compact_convex_hull: fixes s::"(real^'n::finite) set" @@ -1273,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" @@ -1288,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 @@ -1301,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) @@ -1325,16 +1360,16 @@ fixes a b d :: "real ^ 'n::finite" assumes "d \ 0" shows "dist a (b + d) > dist a b \ dist a (b - d) > dist a b" -proof(cases "a \ d - b \ d > 0") - case True hence "0 < d \ d + (a \ d * 2 - b \ d * 2)" +proof(cases "inner a d - inner b d > 0") + case True hence "0 < inner d d + (inner a d * 2 - inner b d * 2)" apply(rule_tac add_pos_pos) using assms by auto - thus ?thesis apply(rule_tac disjI2) unfolding dist_norm and real_vector_norm_def and real_sqrt_less_iff - by(simp add: dot_rsub dot_radd dot_lsub dot_ladd dot_sym field_simps) + thus ?thesis apply(rule_tac disjI2) unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff + by (simp add: algebra_simps inner_commute) next - case False hence "0 < d \ d + (b \ d * 2 - a \ d * 2)" + case False hence "0 < inner d d + (inner b d * 2 - inner a d * 2)" apply(rule_tac add_pos_nonneg) using assms by auto - thus ?thesis apply(rule_tac disjI1) unfolding dist_norm and real_vector_norm_def and real_sqrt_less_iff - by(simp add: dot_rsub dot_radd dot_lsub dot_ladd dot_sym field_simps) + thus ?thesis apply(rule_tac disjI1) unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff + by (simp add: algebra_simps inner_commute) qed lemma norm_increases_online: @@ -1349,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") @@ -1368,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 @@ -1469,38 +1504,59 @@ "closed s \ s \ {} \ (closest_point s x = x \ x \ s)" using closest_point_in_set[of s x] closest_point_self[of x s] by auto +(* TODO: move *) +lemma norm_lt: "norm x < norm y \ inner x x < inner y y" + unfolding norm_eq_sqrt_inner by simp + +(* TODO: move *) +lemma norm_le: "norm x \ norm y \ inner x x \ inner y y" + unfolding norm_eq_sqrt_inner by simp + lemma closer_points_lemma: fixes y::"real^'n::finite" - assumes "y \ z > 0" - shows "\u>0. \v>0. v \ u \ norm(v *s z - y) < norm y" -proof- have z:"z \ z > 0" unfolding dot_pos_lt using assms by auto - thus ?thesis using assms apply(rule_tac x="(y \ z) / (z \ z)" in exI) apply(rule) defer proof(rule+) - fix v assume "0 y \ z / (z \ z)" - thus "norm (v *s z - y) < norm y" unfolding norm_lt using z and assms - by (simp add: field_simps dot_sym mult_strict_left_mono[OF _ `0 0" + shows "\u>0. \v>0. v \ u \ norm(v *\<^sub>R z - y) < norm y" +proof- have z:"inner z z > 0" unfolding inner_gt_zero_iff using assms by auto + thus ?thesis using assms apply(rule_tac x="inner y z / inner z z" in exI) apply(rule) defer proof(rule+) + fix v assume "0 inner y z / inner z z" + thus "norm (v *\<^sub>R z - y) < norm y" unfolding norm_lt using z and assms + by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ `0 (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)" + assumes "inner (y - x) (z - x) > 0" + 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 lemma any_closest_point_dot: assumes "convex s" "closed s" "x \ s" "y \ s" "\z\s. dist a x \ dist a z" - shows "(a - x) \ (y - x) \ 0" -proof(rule ccontr) assume "\ (a - x) \ (y - x) \ 0" - then obtain u where u:"u>0" "u\1" "dist (x + u *s (y - x)) a < dist x a" using closer_point_lemma[of a x y] by auto - let ?z = "(1 - u) *s x + u *s y" have "?z \ s" using mem_convex[OF assms(1,3,4), of u] using u by auto - thus False using assms(5)[THEN bspec[where x="?z"]] and u(3) by (auto simp add: dist_commute field_simps) qed + shows "inner (a - x) (y - x) \ 0" +proof(rule ccontr) assume "\ inner (a - x) (y - x) \ 0" + 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 + +(* TODO: move *) +lemma norm_le_square: "norm x \ a \ 0 \ a \ inner x x \ a\" +proof - + have "norm x \ a \ 0 \ a \ norm x \ a" + using norm_ge_zero [of x] by arith + also have "\ \ 0 \ a \ (norm x)\ \ a\" + by (auto intro: power_mono dest: power2_le_imp_le) + also have "\ \ 0 \ a \ inner x x \ a\" + unfolding power2_norm_eq_inner .. + finally show ?thesis . +qed lemma any_closest_point_unique: assumes "convex s" "closed s" "x \ s" "y \ s" "\z\s. dist a x \ dist a z" "\z\s. dist a y \ dist a z" shows "x = y" using any_closest_point_dot[OF assms(1-4,5)] and any_closest_point_dot[OF assms(1-2,4,3,6)] - unfolding norm_pths(1) and norm_le_square by auto + unfolding norm_pths(1) and norm_le_square + by (auto simp add: algebra_simps) lemma closest_point_unique: assumes "convex s" "closed s" "x \ s" "\z\s. dist a x \ dist a z" @@ -1510,7 +1566,7 @@ lemma closest_point_dot: assumes "convex s" "closed s" "x \ s" - shows "(a - closest_point s a) \ (x - closest_point s a) \ 0" + shows "inner (a - closest_point s a) (x - closest_point s a) \ 0" apply(rule any_closest_point_dot[OF assms(1,2) _ assms(3)]) using closest_point_exists[OF assms(2)] and assms(3) by auto @@ -1525,13 +1581,13 @@ assumes "convex s" "closed s" "s \ {}" shows "dist (closest_point s x) (closest_point s y) \ dist x y" proof- - have "(x - closest_point s x) \ (closest_point s y - closest_point s x) \ 0" - "(y - closest_point s y) \ (closest_point s x - closest_point s y) \ 0" + have "inner (x - closest_point s x) (closest_point s y - closest_point s x) \ 0" + "inner (y - closest_point s y) (closest_point s x - closest_point s y) \ 0" apply(rule_tac[!] any_closest_point_dot[OF assms(1-2)]) using closest_point_exists[OF assms(2-3)] by auto thus ?thesis unfolding dist_norm and norm_le - using dot_pos_le[of "(x - closest_point s x) - (y - closest_point s y)"] - by (auto simp add: dot_sym dot_ladd dot_radd) qed + using inner_ge_zero[of "(x - closest_point s x) - (y - closest_point s y)"] + by (simp add: inner_add inner_diff inner_commute) qed lemma continuous_at_closest_point: assumes "convex s" "closed s" "s \ {}" @@ -1548,50 +1604,50 @@ lemma supporting_hyperplane_closed_point: assumes "convex s" "closed s" "s \ {}" "z \ s" - shows "\a b. \y\s. a \ z < b \ (a \ y = b) \ (\x\s. a \ x \ b)" + shows "\a b. \y\s. inner a z < b \ (inner a y = b) \ (\x\s. inner a x \ b)" proof- from distance_attains_inf[OF assms(2-3)] obtain y where "y\s" and y:"\x\s. dist z y \ dist z x" by auto - show ?thesis apply(rule_tac x="y - z" in exI, rule_tac x="(y - z) \ y" in exI, rule_tac x=y in bexI) + show ?thesis apply(rule_tac x="y - z" in exI, rule_tac x="inner (y - z) y" in exI, rule_tac x=y in bexI) apply rule defer apply rule defer apply(rule, rule ccontr) using `y\s` proof- - show "(y - z) \ z < (y - z) \ y" apply(subst diff_less_iff(1)[THEN sym]) - unfolding dot_rsub[THEN sym] and dot_pos_lt using `y\s` `z\s` by auto + show "inner (y - z) z < inner (y - z) y" apply(subst diff_less_iff(1)[THEN sym]) + unfolding inner_diff_right[THEN sym] and inner_gt_zero_iff 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) + assume "\ inner (y - z) y \ inner (y - z) x" then obtain v where + "v>0" "v\1" "dist (y + v *\<^sub>R (x - y)) z < dist y z" using closer_point_lemma[of z y x] apply - by (auto simp add: inner_diff) + thus False using *[THEN spec[where x=v]] by(auto simp add: dist_commute algebra_simps) qed auto qed lemma separating_hyperplane_closed_point: assumes "convex s" "closed s" "z \ s" - shows "\a b. a \ z < b \ (\x\s. a \ x > b)" + shows "\a b. inner a z < b \ (\x\s. inner a x > b)" proof(cases "s={}") case True thus ?thesis apply(rule_tac x="-z" in exI, rule_tac x=1 in exI) - using less_le_trans[OF _ dot_pos_le[of z]] by auto + using less_le_trans[OF _ inner_ge_zero[of z]] by auto next case False obtain y where "y\s" and y:"\x\s. dist z y \ dist z x" using distance_attains_inf[OF assms(2) False] by auto - show ?thesis apply(rule_tac x="y - z" in exI, rule_tac x="(y - z) \ z + (norm(y - z))\ / 2" in exI) + show ?thesis apply(rule_tac x="y - z" in exI, rule_tac x="inner (y - z) z + (norm(y - z))\ / 2" in exI) apply rule defer apply rule proof- fix x assume "x\s" - have "\ 0 < (z - y) \ (x - y)" apply(rule_tac notI) proof(drule closer_point_lemma) - assume "\u>0. u \ 1 \ dist (y + u *s (x - y)) z < dist y z" - then obtain u where "u>0" "u\1" "dist (y + u *s (x - y)) z < dist y z" by auto - thus False using y[THEN bspec[where x="y + u *s (x - y)"]] + have "\ 0 < inner (z - y) (x - y)" apply(rule_tac notI) proof(drule closer_point_lemma) + 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" - unfolding norm_pow_2 and dlo_simps(3) by (auto simp add: field_simps dot_sym) + hence "0 < inner (y - z) (y - z)" unfolding power2_norm_eq_inner by simp + ultimately show "inner (y - z) z + (norm (y - z))\ / 2 < inner (y - z) x" + unfolding power2_norm_eq_inner and not_less by (auto simp add: field_simps inner_commute inner_diff) qed(insert `y\s` `z\s`, auto) qed lemma separating_hyperplane_closed_0: assumes "convex (s::(real^'n::finite) set)" "closed s" "0 \ s" - shows "\a b. a \ 0 \ 0 < b \ (\x\s. a \ x > b)" + shows "\a b. a \ 0 \ 0 < b \ (\x\s. inner a x > b)" proof(cases "s={}") guess a using UNIV_witness[where 'a='n] .. case True have "norm ((basis a)::real^'n::finite) = 1" using norm_basis and dimindex_ge_1 by auto @@ -1603,41 +1659,41 @@ lemma separating_hyperplane_closed_compact: assumes "convex (s::(real^'n::finite) set)" "closed s" "convex t" "compact t" "t \ {}" "s \ t = {}" - shows "\a b. (\x\s. a \ x < b) \ (\x\t. a \ x > b)" + shows "\a b. (\x\s. inner a x < b) \ (\x\t. inner a x > b)" proof(cases "s={}") case True obtain b where b:"b>0" "\x\t. norm x \ b" using compact_imp_bounded[OF assms(4)] unfolding bounded_pos by auto obtain z::"real^'n" where z:"norm z = b + 1" using vector_choose_size[of "b + 1"] and b(1) by auto hence "z\t" using b(2)[THEN bspec[where x=z]] by auto - then obtain a b where ab:"a \ z < b" "\x\t. b < a \ x" + then obtain a b where ab:"inner a z < b" "\x\t. b < inner a x" using separating_hyperplane_closed_point[OF assms(3) compact_imp_closed[OF assms(4)], of z] by auto thus ?thesis using True by auto next case False then obtain y where "y\s" by auto - obtain a b where "0 < b" "\x\{x - y |x y. x \ s \ y \ t}. b < a \ x" + obtain a b where "0 < b" "\x\{x - y |x y. x \ s \ y \ t}. b < inner a x" using separating_hyperplane_closed_point[OF convex_differences[OF assms(1,3)], of 0] using closed_compact_differences[OF assms(2,4)] using assms(6) by(auto, blast) - hence ab:"\x\s. \y\t. b + a \ y < a \ x" apply- apply(rule,rule) apply(erule_tac x="x - y" in ballE) by auto - def k \ "rsup ((\x. a \ x) ` t)" + hence ab:"\x\s. \y\t. b + inner a y < inner a x" apply- apply(rule,rule) apply(erule_tac x="x - y" in ballE) by (auto simp add: inner_diff) + def k \ "rsup ((\x. inner a x) ` t)" show ?thesis apply(rule_tac x="-a" in exI, rule_tac x="-(k + b / 2)" in exI) - apply(rule,rule) defer apply(rule) unfolding dot_lneg and neg_less_iff_less proof- - from ab have "((\x. a \ x) ` t) *<= (a \ y - b)" + apply(rule,rule) defer apply(rule) unfolding inner_minus_left and neg_less_iff_less proof- + from ab have "((\x. inner a x) ` t) *<= (inner a y - b)" apply(erule_tac x=y in ballE) apply(rule setleI) using `y\s` by auto - hence k:"isLub UNIV ((\x. a \ x) ` t) k" unfolding k_def apply(rule_tac rsup) using assms(5) by auto - fix x assume "x\t" thus "a \ x < (k + b / 2)" using `0 x"] by auto + hence k:"isLub UNIV ((\x. inner a x) ` t) k" unfolding k_def apply(rule_tac rsup) using assms(5) by auto + fix x assume "x\t" thus "inner a x < (k + b / 2)" using `0s" - hence "k \ a \ x - b" unfolding k_def apply(rule_tac rsup_le) using assms(5) + hence "k \ inner a x - b" unfolding k_def apply(rule_tac rsup_le) using assms(5) unfolding setle_def using ab[THEN bspec[where x=x]] by auto - thus "k + b / 2 < a \ x" using `0 < b` by auto + thus "k + b / 2 < inner a x" using `0 < b` by auto qed qed lemma separating_hyperplane_compact_closed: assumes "convex s" "compact s" "s \ {}" "convex t" "closed t" "s \ t = {}" - shows "\a b. (\x\s. a \ x < b) \ (\x\t. a \ x > b)" -proof- obtain a b where "(\x\t. a \ x < b) \ (\x\s. b < a \ x)" + shows "\a b. (\x\s. inner a x < b) \ (\x\t. inner a x > b)" +proof- obtain a b where "(\x\t. inner a x < b) \ (\x\s. b < inner a x)" using separating_hyperplane_closed_compact[OF assms(4-5,1-2,3)] and assms(6) by auto thus ?thesis apply(rule_tac x="-a" in exI, rule_tac x="-b" in exI) by auto qed @@ -1645,33 +1701,33 @@ lemma separating_hyperplane_set_0: assumes "convex s" "(0::real^'n::finite) \ s" - shows "\a. a \ 0 \ (\x\s. 0 \ a \ x)" -proof- let ?k = "\c. {x::real^'n. 0 \ c \ x}" + shows "\a. a \ 0 \ (\x\s. 0 \ inner a x)" +proof- let ?k = "\c. {x::real^'n. 0 \ inner c x}" have "frontier (cball 0 1) \ (\ (?k ` s)) \ {}" apply(rule compact_imp_fip) apply(rule compact_frontier[OF compact_cball]) defer apply(rule,rule,erule conjE) proof- fix f assume as:"f \ ?k ` s" "finite f" obtain c where c:"f = ?k ` c" "c\s" "finite c" using finite_subset_image[OF as(2,1)] by auto - then obtain a b where ab:"a \ 0" "0 < b" "\x\convex hull c. b < a \ x" + then obtain a b where ab:"a \ 0" "0 < b" "\x\convex hull c. b < inner a x" using separating_hyperplane_closed_0[OF convex_convex_hull, of c] 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 \ inner y x)" apply(rule_tac x="inverse(norm a) *\<^sub>R a" in exI) + using hull_subset[of c convex] unfolding subset_eq and inner_scaleR apply- apply rule defer apply rule apply(rule mult_nonneg_nonneg) - by(auto simp add: dot_sym elim!: ballE) + by(auto simp add: inner_commute 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 - thus ?thesis apply(rule_tac x=x in exI) by(auto simp add: dot_sym) qed + thus ?thesis apply(rule_tac x=x in exI) by(auto simp add: inner_commute) qed lemma separating_hyperplane_sets: assumes "convex s" "convex (t::(real^'n::finite) set)" "s \ {}" "t \ {}" "s \ t = {}" - shows "\a b. a \ 0 \ (\x\s. a \ x \ b) \ (\x\t. a \ x \ b)" + shows "\a b. a \ 0 \ (\x\s. inner a x \ b) \ (\x\t. inner a x \ b)" proof- from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]] - obtain a where "a\0" "\x\{x - y |x y. x \ t \ y \ s}. 0 \ a \ x" using assms(3-5) by auto - hence "\x\t. \y\s. a \ y \ a \ x" apply- apply(rule, rule) apply(erule_tac x="x - y" in ballE) by auto - thus ?thesis apply(rule_tac x=a in exI, rule_tac x="rsup ((\x. a \ x) ` s)" in exI) using `a\0` + obtain a where "a\0" "\x\{x - y |x y. x \ t \ y \ s}. 0 \ inner a x" using assms(3-5) by auto + hence "\x\t. \y\s. inner a y \ inner a x" apply- apply(rule, rule) apply(erule_tac x="x - y" in ballE) by (auto simp add: inner_diff) + thus ?thesis apply(rule_tac x=a in exI, rule_tac x="rsup ((\x. inner a x) ` s)" in exI) using `a\0` apply(rule) apply(rule,rule) apply(rule rsup[THEN isLubD2]) prefer 4 apply(rule,rule rsup_le) unfolding setle_def prefer 4 using assms(3-5) by blast+ qed @@ -1680,7 +1736,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 @@ -1688,13 +1744,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 @@ -1717,27 +1773,27 @@ 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. *} lemma convex_halfspace_intersection: assumes "closed s" "convex s" - shows "s = \ {h. s \ h \ (\a b. h = {x. a \ x \ b})}" + shows "s = \ {h. s \ h \ (\a b. h = {x. inner a x \ b})}" apply(rule set_ext, rule) unfolding Inter_iff Ball_def mem_Collect_eq apply(rule,rule,erule conjE) proof- - fix x assume "\xa. s \ xa \ (\a b. xa = {x. a \ x \ b}) \ x \ xa" - hence "\a b. s \ {x. a \ x \ b} \ x \ {x. a \ x \ b}" by blast + fix x assume "\xa. s \ xa \ (\a b. xa = {x. inner a x \ b}) \ x \ xa" + hence "\a b. s \ {x. inner a x \ b} \ x \ {x. inner a x \ b}" by blast thus "x\s" apply(rule_tac ccontr) apply(drule separating_hyperplane_closed_point[OF assms(2,1)]) apply(erule exE)+ apply(erule_tac x="-a" in allE, erule_tac x="-b" in allE) by auto qed auto @@ -1746,9 +1802,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: @@ -1769,9 +1825,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") @@ -1783,23 +1839,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 @@ -1865,10 +1921,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: @@ -1880,72 +1936,75 @@ 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) unfolding o_def vec1_dest_vec1 apply(rule continuous_at_id) by(rule compact_interval) - moreover have "{y. \u\0. u \ b / norm x \ y = u *s x} \ s \ {}" apply(rule not_disjointI[OF _ assms(2)]) + 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 *\<^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`] 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) 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)) 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 contpi:"continuous_on (UNIV - {0}) pi" apply(rule continuous_at_imp_continuous_on) - apply rule unfolding pi_def apply(rule continuous_mul) unfolding o_def - apply(rule continuous_at_inv[unfolded o_def]) unfolding continuous_at_vec1_range[unfolded o_def] - apply(rule,rule) apply(rule_tac x=e in exI) apply(rule,assumption,rule,rule) - proof- fix e x y assume "0 < e" "norm (y - x::real^'n) < e" - thus "\norm y - norm x\ < e" using norm_triangle_ineq3[of y x] by auto - qed(auto intro!:continuous_at_id) + apply rule unfolding pi_def + apply (rule continuous_mul) + apply (rule continuous_at_inv[unfolded o_def]) + apply (rule continuous_at_norm) + apply simp + 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 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" @@ -1955,14 +2014,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 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 @@ -1978,7 +2037,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]) @@ -1987,35 +2046,35 @@ 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`] + 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 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") - case False thus ?thesis apply(rule_tac continuous_mul, rule_tac continuous_at_vec1_norm) + 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] .. obtain B where B:"\x\s. norm x \ B" using compact_imp_bounded[OF assms(1)] unfolding bounded_iff by auto @@ -2023,7 +2082,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 @@ -2038,8 +2097,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 @@ -2056,18 +2115,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" @@ -2075,16 +2134,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) + 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 \ {}" @@ -2101,8 +2160,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" @@ -2131,11 +2190,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) @@ -2157,7 +2216,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: @@ -2179,10 +2238,10 @@ apply(rule,rule,rule,rule,erule conjE,rule ccontr) proof- fix a b x assume as:"connected s" "a \ s" "b \ s" "dest_vec1 a \ dest_vec1 x" "dest_vec1 x \ dest_vec1 b" "x\s" hence *:"dest_vec1 a < dest_vec1 x" "dest_vec1 x < dest_vec1 b" apply(rule_tac [!] ccontr) unfolding not_less by auto - let ?halfl = "{z. basis 1 \ z < dest_vec1 x} " and ?halfr = "{z. basis 1 \ z > dest_vec1 x} " + let ?halfl = "{z. inner (basis 1) z < dest_vec1 x} " and ?halfr = "{z. inner (basis 1) z > dest_vec1 x} " { fix y assume "y \ s" have "y \ ?halfr \ ?halfl" apply(rule ccontr) - using as(6) `y\s` by (auto simp add: basis_component field_simps dest_vec1_eq[unfolded dest_vec1_def One_nat_def] dest_vec1_def) } - moreover have "a\?halfl" "b\?halfr" using * by (auto simp add: basis_component field_simps dest_vec1_def) + using as(6) `y\s` by (auto simp add: inner_vector_def dest_vec1_eq [unfolded dest_vec1_def] dest_vec1_def) } + moreover have "a\?halfl" "b\?halfr" using * by (auto simp add: inner_vector_def dest_vec1_def) hence "?halfl \ s \ {}" "?halfr \ s \ {}" using as(2-3) by auto ultimately show False apply(rule_tac notE[OF as(1)[unfolded connected_def]]) apply(rule_tac x="?halfl" in exI, rule_tac x="?halfr" in exI) @@ -2232,7 +2291,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) @@ -2270,7 +2329,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 @@ -2304,7 +2363,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]] @@ -2314,26 +2373,26 @@ 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. *} lemma convex_on_bounded_continuous: assumes "open s" "convex_on s f" "\x\s. abs(f x) \ b" - shows "continuous_on s (vec1 o f)" - apply(rule continuous_at_imp_continuous_on) unfolding continuous_at_vec1_range proof(rule,rule,rule) + shows "continuous_on s f" + apply(rule continuous_at_imp_continuous_on) unfolding continuous_at_real_range proof(rule,rule,rule) fix x e assume "x\s" "(0::real) < e" def B \ "abs b + 1" have B:"0 < B" "\x. x\s \ abs (f x) \ B" @@ -2347,12 +2406,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 + 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) + 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) @@ -2360,12 +2419,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 + 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) + 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" @@ -2384,10 +2443,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" @@ -2398,7 +2457,7 @@ lemma convex_on_continuous: assumes "open (s::(real^'n::finite) set)" "convex_on s f" - shows "continuous_on s (vec1 \ f)" + shows "continuous_on s f" unfolding continuous_on_eq_continuous_at[OF assms(1)] proof note dimge1 = dimindex_ge_1[where 'a='n] fix x assume "x\s" @@ -2428,19 +2487,25 @@ using order_trans[OF component_le_norm y[unfolded mem_cball dist_norm], of i] by(auto simp add: vector_component) } thus "f y \ k" apply(rule_tac k[rule_format]) unfolding mem_cball mem_interval dist_norm by(auto simp add: vector_component_simps) qed - hence "continuous_on (ball x d) (vec1 \ f)" apply(rule_tac convex_on_bounded_continuous) + hence "continuous_on (ball x d) f" apply(rule_tac convex_on_bounded_continuous) apply(rule open_ball, rule convex_on_subset[OF conv], rule ball_subset_cball) by auto - thus "continuous (at x) (vec1 \ f)" unfolding continuous_on_eq_continuous_at[OF open_ball] using `d>0` by auto qed + thus "continuous (at x) f" unfolding continuous_on_eq_continuous_at[OF open_ball] using `d>0` by auto qed subsection {* Line segments, starlike sets etc. *) (* Use the same overloading tricks as for intervals, so that *) (* segment[a,b] is closed and segment(a,b) is open relative to affine hull. *} -definition "midpoint a b = (inverse (2::real)) *s (a + b)" - -definition "open_segment a b = {(1 - u) *s a + u *s b | u::real. 0 < u \ u < 1}" - -definition "closed_segment a b = {(1 - u) *s a + u *s b | u::real. 0 \ u \ u \ 1}" +definition + 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)" @@ -2449,9 +2514,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) @@ -2459,8 +2524,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 + have **:"\x y::real^'n::finite. 2 *\<^sub>R x = y \ norm x = (norm y) / 2" by auto + 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) @@ -2507,7 +2573,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 @@ -2517,32 +2583,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] * 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 show ?t1 ?t2 unfolding between midpoint_def dist_norm apply(rule_tac[!] *) by(auto simp add:field_simps Cart_eq vector_component_simps) qed @@ -2554,16 +2620,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]) @@ -2572,7 +2638,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 @@ -2587,11 +2653,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 add:field_simps norm_minus_commute) thus ?thesis unfolding * apply - apply(rule mem_interior_convex_shrink) using assms(1,4-5) `y\s` by auto qed @@ -2599,7 +2665,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) @@ -2614,16 +2680,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" - 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 + 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. inner 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 [where 'a=real, unfolded smult_conv_scaleR] by(auto simp add:inner_basis) qed qed lemma interior_std_simplex: "interior (convex hull (insert 0 { basis i| i. i\UNIV})) = @@ -2632,14 +2698,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` + 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 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] + 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 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 "\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 @@ -2665,8 +2731,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) @@ -2691,7 +2757,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)" @@ -2736,7 +2802,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 @@ -2759,9 +2825,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) @@ -2769,35 +2835,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) @@ -2814,12 +2880,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: @@ -2831,6 +2897,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}" @@ -2840,40 +2910,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 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 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: @@ -2884,22 +2954,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 + 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) - 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 + 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) 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 @@ -2966,7 +3036,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 @@ -2975,7 +3045,8 @@ unfolding pathfinish_def linepath_def by auto lemma continuous_linepath_at[intro]: "continuous (at x) (linepath a b)" - unfolding linepath_def by(auto simp add: vec1_dest_vec1 o_def intro!: continuous_intros) + unfolding linepath_def + by (intro continuous_intros continuous_dest_vec1) lemma continuous_on_linepath[intro]: "continuous_on s (linepath a b)" using continuous_linepath_at by(auto intro!: continuous_at_imp_continuous_on) @@ -2985,7 +3056,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" @@ -2993,10 +3064,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) @@ -3136,7 +3207,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) @@ -3153,18 +3224,18 @@ obtain \ where \:"bij_betw \ {1..CARD('n)} (UNIV::'n set)" using ex_bij_betw_nat_finite_1[OF finite_UNIV] by auto let ?U = "UNIV::(real^'n) set" let ?u = "?U - {0}" let ?basis = "\k. basis (\ k)" - let ?A = "\k. {x::real^'n. \i\{1..k}. (basis (\ i)) \ x \ 0}" + let ?A = "\k. {x::real^'n. \i\{1..k}. inner (basis (\ i)) x \ 0}" have "\k\{2..CARD('n)}. path_connected (?A k)" proof - have *:"\k. ?A (Suc k) = {x. ?basis (Suc k) \ x < 0} \ {x. ?basis (Suc k) \ x > 0} \ ?A k" apply(rule set_ext,rule) defer + have *:"\k. ?A (Suc k) = {x. inner (?basis (Suc k)) x < 0} \ {x. inner (?basis (Suc k)) x > 0} \ ?A k" apply(rule set_ext,rule) defer apply(erule UnE)+ unfolding mem_Collect_eq apply(rule_tac[1-2] x="Suc k" in bexI) by(auto elim!: ballE simp add: not_less le_Suc_eq) fix k assume "k \ {2..CARD('n)}" thus "path_connected (?A k)" proof(induct k) case (Suc k) show ?case proof(cases "k = 1") case False from Suc have d:"k \ {1..CARD('n)}" "Suc k \ {1..CARD('n)}" by auto hence "\ k \ \ (Suc k)" using \[unfolded bij_betw_def inj_on_def, THEN conjunct1, THEN bspec[where x=k]] by auto - hence **:"?basis k + ?basis (Suc k) \ {x. 0 < ?basis (Suc k) \ x} \ (?A k)" - "?basis k - ?basis (Suc k) \ {x. 0 > ?basis (Suc k) \ x} \ ({x. 0 < ?basis (Suc k) \ x} \ (?A k))" using d - by(auto simp add: dot_basis vector_component_simps intro!:bexI[where x=k]) + hence **:"?basis k + ?basis (Suc k) \ {x. 0 < inner (?basis (Suc k)) x} \ (?A k)" + "?basis k - ?basis (Suc k) \ {x. 0 > inner (?basis (Suc k)) x} \ ({x. 0 < inner (?basis (Suc k)) x} \ (?A k))" using d + by(auto simp add: inner_basis vector_component_simps intro!:bexI[where x=k]) show ?thesis unfolding * Un_assoc apply(rule path_connected_Un) defer apply(rule path_connected_Un) prefer 5 apply(rule_tac[1-2] convex_imp_path_connected, rule convex_halfspace_lt, rule convex_halfspace_gt) apply(rule Suc(1)) apply(rule_tac[2-3] ccontr) using d ** False by auto @@ -3177,18 +3248,18 @@ apply(rule_tac[5] x=" ?basis 1 + ?basis 2" in nequals0I) apply(rule_tac[6] x="-?basis 1 + ?basis 2" in nequals0I) apply(rule_tac[7] x="-?basis 1 - ?basis 2" in nequals0I) - using d unfolding *** by(auto intro!: convex_halfspace_gt convex_halfspace_lt, auto simp add:vector_component_simps dot_basis) + using d unfolding *** by(auto intro!: convex_halfspace_gt convex_halfspace_lt, auto simp add:vector_component_simps inner_basis) qed qed auto qed note lem = this - have ***:"\x::real^'n. (\i\{1..CARD('n)}. basis (\ i) \ x \ 0) \ (\i. basis i \ x \ 0)" + have ***:"\x::real^'n. (\i\{1..CARD('n)}. inner (basis (\ i)) x \ 0) \ (\i. inner (basis i) x \ 0)" apply rule apply(erule bexE) apply(rule_tac x="\ i" in exI) defer apply(erule exE) proof- - fix x::"real^'n" and i assume as:"basis i \ x \ 0" + fix x::"real^'n" and i assume as:"inner (basis i) x \ 0" have "i\\ ` {1..CARD('n)}" using \[unfolded bij_betw_def, THEN conjunct2] by auto then obtain j where "j\{1..CARD('n)}" "\ j = i" by auto - thus "\i\{1..CARD('n)}. basis (\ i) \ x \ 0" apply(rule_tac x=j in bexI) using as by auto qed auto + thus "\i\{1..CARD('n)}. inner (basis (\ i)) x \ 0" apply(rule_tac x=j in bexI) using as by auto qed auto have *:"?U - {a} = (\x. x + a) ` {x. x \ 0}" apply(rule set_ext) unfolding image_iff apply rule apply(rule_tac x="x - a" in bexI) by auto - have **:"\x::real^'n. x\0 \ (\i. basis i \ x \ 0)" unfolding Cart_eq by(auto simp add: dot_basis) + have **:"\x::real^'n. x\0 \ (\i. inner (basis i) x \ 0)" unfolding Cart_eq by(auto simp add: inner_basis) show ?thesis unfolding * apply(rule path_connected_continuous_image) apply(rule continuous_on_intros)+ unfolding ** apply(rule lem[THEN bspec[where x="CARD('n)"], unfolded ***]) using assms by auto qed @@ -3197,14 +3268,14 @@ 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) have ***:"\xa. (if xa = 0 then 0 else 1) \ 1 \ xa = 0" apply(rule ccontr) by auto - have **:"{x::real^'n. norm x = 1} = (\x. (1/norm x) *s x) ` (UNIV - {0})" apply(rule set_ext,rule) - unfolding image_iff apply(rule_tac x=x in bexI) unfolding mem_Collect_eq norm_mul by(auto intro!: ***) - have "continuous_on (UNIV - {0}) (vec1 \ (\x::real^'n. 1 / norm x))" unfolding o_def continuous_on_eq_continuous_within + 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 by(auto intro!: ***) + 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_vec1_norm[unfolded o_def]) by auto + apply(rule continuous_at_norm[unfolded o_def]) by auto thus ?thesis unfolding * ** using path_connected_punctured_universe[OF assms] by(auto intro!: path_connected_continuous_image continuous_on_intros continuous_on_mul) qed diff -r 78ac5c304db7 -r ef6d67b1ad10 src/HOL/Library/Euclidean_Space.thy --- a/src/HOL/Library/Euclidean_Space.thy Thu Jun 11 19:49:02 2009 +0200 +++ b/src/HOL/Library/Euclidean_Space.thy Sat Jun 13 16:29:15 2009 +0200 @@ -369,19 +369,33 @@ end -lemma tendsto_Cart_nth: - fixes f :: "'a \ 'b::topological_space ^ 'n::finite" +lemma open_vector_box: "\i. open (S i) \ open {x. \i. x $ i \ S i}" +unfolding open_vector_def by auto + +lemma open_vimage_Cart_nth: "open S \ open ((\x. x $ i) -` S)" +unfolding open_vector_def +apply clarify +apply (rule_tac x="\k. if k = i then S else UNIV" in exI, simp) +done + +lemma closed_vimage_Cart_nth: "closed S \ closed ((\x. x $ i) -` S)" +unfolding closed_open vimage_Compl [symmetric] +by (rule open_vimage_Cart_nth) + +lemma closed_vector_box: "\i. closed (S i) \ closed {x. \i. x $ i \ S i}" +proof - + have "{x. \i. x $ i \ S i} = (\i. (\x. x $ i) -` S i)" by auto + thus "\i. closed (S i) \ closed {x. \i. x $ i \ S i}" + by (simp add: closed_INT closed_vimage_Cart_nth) +qed + +lemma tendsto_Cart_nth [tendsto_intros]: assumes "((\x. f x) ---> a) net" shows "((\x. f x $ i) ---> a $ i) net" proof (rule topological_tendstoI) - fix S :: "'b set" assume "open S" "a $ i \ S" + fix S assume "open S" "a $ i \ S" then have "open ((\y. y $ i) -` S)" "a \ ((\y. y $ i) -` S)" - unfolding open_vector_def - apply simp_all - apply clarify - apply (rule_tac x="\k. if k = i then S else UNIV" in exI) - apply simp - done + by (simp_all add: open_vimage_Cart_nth) with assms have "eventually (\x. f x \ (\y. y $ i) -` S) net" by (rule topological_tendstoD) then show "eventually (\x. f x $ i \ S) net" @@ -736,7 +750,7 @@ instantiation "^" :: (real_normed_vector, finite) real_normed_vector begin -definition vector_norm_def: +definition norm_vector_def: "norm (x::'a^'b) = setL2 (\i. norm (x$i)) UNIV" definition vector_sgn_def: @@ -745,30 +759,30 @@ instance proof fix a :: real and x y :: "'a ^ 'b" show "0 \ norm x" - unfolding vector_norm_def + unfolding norm_vector_def by (rule setL2_nonneg) show "norm x = 0 \ x = 0" - unfolding vector_norm_def + unfolding norm_vector_def by (simp add: setL2_eq_0_iff Cart_eq) show "norm (x + y) \ norm x + norm y" - unfolding vector_norm_def + unfolding norm_vector_def apply (rule order_trans [OF _ setL2_triangle_ineq]) apply (simp add: setL2_mono norm_triangle_ineq) done show "norm (scaleR a x) = \a\ * norm x" - unfolding vector_norm_def - by (simp add: norm_scaleR setL2_right_distrib) + unfolding norm_vector_def + by (simp add: setL2_right_distrib) show "sgn x = scaleR (inverse (norm x)) x" by (rule vector_sgn_def) show "dist x y = norm (x - y)" - unfolding dist_vector_def vector_norm_def + unfolding dist_vector_def norm_vector_def by (simp add: dist_norm) qed end lemma norm_nth_le: "norm (x $ i) \ norm x" -unfolding vector_norm_def +unfolding norm_vector_def by (rule member_le_setL2) simp_all interpretation Cart_nth: bounded_linear "\x. x $ i" @@ -785,28 +799,28 @@ instantiation "^" :: (real_inner, finite) real_inner begin -definition vector_inner_def: +definition inner_vector_def: "inner x y = setsum (\i. inner (x$i) (y$i)) UNIV" instance proof fix r :: real and x y z :: "'a ^ 'b" show "inner x y = inner y x" - unfolding vector_inner_def + unfolding inner_vector_def by (simp add: inner_commute) show "inner (x + y) z = inner x z + inner y z" - unfolding vector_inner_def - by (simp add: inner_left_distrib setsum_addf) + unfolding inner_vector_def + by (simp add: inner_add_left setsum_addf) show "inner (scaleR r x) y = r * inner x y" - unfolding vector_inner_def - by (simp add: inner_scaleR_left setsum_right_distrib) + unfolding inner_vector_def + by (simp add: setsum_right_distrib) show "0 \ inner x x" - unfolding vector_inner_def + unfolding inner_vector_def by (simp add: setsum_nonneg) show "inner x x = 0 \ x = 0" - unfolding vector_inner_def + unfolding inner_vector_def by (simp add: Cart_eq setsum_nonneg_eq_0_iff) show "norm x = sqrt (inner x x)" - unfolding vector_inner_def vector_norm_def setL2_def + unfolding inner_vector_def norm_vector_def setL2_def by (simp add: power2_norm_eq_inner) qed @@ -864,7 +878,7 @@ done lemma norm_vector_1: "norm (x :: _^1) = norm (x$1)" - by (simp add: vector_norm_def UNIV_1) + by (simp add: norm_vector_def UNIV_1) lemma norm_real: "norm(x::real ^ 1) = abs(x$1)" by (simp add: norm_vector_1) @@ -983,12 +997,12 @@ by (rule norm_zero) lemma norm_mul[simp]: "norm(a *s x) = abs(a) * norm x" - by (simp add: vector_norm_def vector_component setL2_right_distrib + by (simp add: norm_vector_def vector_component setL2_right_distrib abs_mult cong: strong_setL2_cong) lemma norm_eq_0_dot: "(norm x = 0) \ (x \ x = (0::real))" - by (simp add: vector_norm_def dot_def setL2_def power2_eq_square) + by (simp add: norm_vector_def dot_def setL2_def power2_eq_square) lemma real_vector_norm_def: "norm x = sqrt (x \ x)" - by (simp add: vector_norm_def setL2_def dot_def power2_eq_square) + by (simp add: norm_vector_def setL2_def dot_def power2_eq_square) lemma norm_pow_2: "norm x ^ 2 = x \ x" by (simp add: real_vector_norm_def) lemma norm_eq_0_imp: "norm x = 0 ==> x = (0::real ^'n::finite)" by (metis norm_eq_zero) @@ -1064,7 +1078,7 @@ qed lemma component_le_norm: "\x$i\ <= norm (x::real ^ 'n::finite)" - apply (simp add: vector_norm_def) + apply (simp add: norm_vector_def) apply (rule member_le_setL2, simp_all) done @@ -1077,7 +1091,7 @@ by (metis component_le_norm basic_trans_rules(21)) lemma norm_le_l1: "norm (x:: real ^'n::finite) <= setsum(\i. \x$i\) UNIV" - by (simp add: vector_norm_def setL2_le_setsum) + by (simp add: norm_vector_def setL2_le_setsum) lemma real_abs_norm: "\norm x\ = norm (x :: real ^ _)" by (rule abs_norm_cancel) @@ -1522,6 +1536,13 @@ shows "basis i \ x = x$i" "x \ (basis i :: 'a^'n::finite) = (x$i :: 'a::semiring_1)" by (auto simp add: dot_def basis_def cond_application_beta cond_value_iff setsum_delta cong del: if_weak_cong) +lemma inner_basis: + fixes x :: "'a::{real_inner, real_algebra_1} ^ 'n::finite" + shows "inner (basis i) x = inner 1 (x $ i)" + and "inner x (basis i) = inner (x $ i) 1" + unfolding inner_vector_def basis_def + by (auto simp add: cond_application_beta cond_value_iff setsum_delta cong del: if_weak_cong) + lemma basis_eq_0: "basis i = (0::'a::semiring_1^'n) \ False" by (auto simp add: Cart_eq) @@ -2917,7 +2938,7 @@ done lemma norm_pastecart: "norm (pastecart x y) <= norm x + norm y" - unfolding vector_norm_def setL2_def setsum_UNIV_sum + unfolding norm_vector_def setL2_def setsum_UNIV_sum by (simp add: sqrt_add_le_add_sqrt setsum_nonneg) subsection {* A generic notion of "hull" (convex, affine, conic hull and closure). *} diff -r 78ac5c304db7 -r ef6d67b1ad10 src/HOL/Library/Inner_Product.thy --- a/src/HOL/Library/Inner_Product.thy Thu Jun 11 19:49:02 2009 +0200 +++ b/src/HOL/Library/Inner_Product.thy Sat Jun 13 16:29:15 2009 +0200 @@ -27,28 +27,28 @@ class real_inner = real_vector + sgn_div_norm + dist_norm + open_dist + fixes inner :: "'a \ 'a \ real" assumes inner_commute: "inner x y = inner y x" - and inner_left_distrib: "inner (x + y) z = inner x z + inner y z" - and inner_scaleR_left: "inner (scaleR r x) y = r * (inner x y)" + and inner_add_left: "inner (x + y) z = inner x z + inner y z" + and inner_scaleR_left [simp]: "inner (scaleR r x) y = r * (inner x y)" and inner_ge_zero [simp]: "0 \ inner x x" and inner_eq_zero_iff [simp]: "inner x x = 0 \ x = 0" and norm_eq_sqrt_inner: "norm x = sqrt (inner x x)" begin lemma inner_zero_left [simp]: "inner 0 x = 0" - using inner_left_distrib [of 0 0 x] by simp + using inner_add_left [of 0 0 x] by simp lemma inner_minus_left [simp]: "inner (- x) y = - inner x y" - using inner_left_distrib [of x "- x" y] by simp + using inner_add_left [of x "- x" y] by simp lemma inner_diff_left: "inner (x - y) z = inner x z - inner y z" - by (simp add: diff_minus inner_left_distrib) + by (simp add: diff_minus inner_add_left) text {* Transfer distributivity rules to right argument. *} -lemma inner_right_distrib: "inner x (y + z) = inner x y + inner x z" - using inner_left_distrib [of y z x] by (simp only: inner_commute) +lemma inner_add_right: "inner x (y + z) = inner x y + inner x z" + using inner_add_left [of y z x] by (simp only: inner_commute) -lemma inner_scaleR_right: "inner x (scaleR r y) = r * (inner x y)" +lemma inner_scaleR_right [simp]: "inner x (scaleR r y) = r * (inner x y)" using inner_scaleR_left [of r y x] by (simp only: inner_commute) lemma inner_zero_right [simp]: "inner x 0 = 0" @@ -60,9 +60,14 @@ lemma inner_diff_right: "inner x (y - z) = inner x y - inner x z" using inner_diff_left [of y z x] by (simp only: inner_commute) +lemmas inner_add [algebra_simps] = inner_add_left inner_add_right +lemmas inner_diff [algebra_simps] = inner_diff_left inner_diff_right +lemmas inner_scaleR = inner_scaleR_left inner_scaleR_right + +text {* Legacy theorem names *} +lemmas inner_left_distrib = inner_add_left +lemmas inner_right_distrib = inner_add_right lemmas inner_distrib = inner_left_distrib inner_right_distrib -lemmas inner_diff = inner_diff_left inner_diff_right -lemmas inner_scaleR = inner_scaleR_left inner_scaleR_right lemma inner_gt_zero_iff [simp]: "0 < inner x x \ x \ 0" by (simp add: order_less_le) @@ -81,7 +86,7 @@ have "0 \ inner (x - scaleR ?r y) (x - scaleR ?r y)" by (rule inner_ge_zero) also have "\ = inner x x - inner y x * ?r" - by (simp add: inner_diff inner_scaleR) + by (simp add: inner_diff) also have "\ = inner x x - (inner x y)\ / inner y y" by (simp add: power2_eq_square inner_commute) finally have "0 \ inner x x - (inner x y)\ / inner y y" . @@ -116,7 +121,7 @@ by (rule order_trans [OF abs_ge_self Cauchy_Schwarz_ineq2]) thus "(norm (x + y))\ \ (norm x + norm y)\" unfolding power2_sum power2_norm_eq_inner - by (simp add: inner_distrib inner_commute) + by (simp add: inner_add inner_commute) show "0 \ norm x + norm y" unfolding norm_eq_sqrt_inner by (simp add: add_nonneg_nonneg) @@ -125,7 +130,7 @@ by (simp add: real_sqrt_mult_distrib) then show "norm (a *\<^sub>R x) = \a\ * norm x" unfolding norm_eq_sqrt_inner - by (simp add: inner_scaleR power2_eq_square mult_assoc) + by (simp add: power2_eq_square mult_assoc) qed end @@ -149,9 +154,9 @@ proof fix x y z :: 'a and r :: real show "inner (x + y) z = inner x z + inner y z" - by (rule inner_left_distrib) + by (rule inner_add_left) show "inner x (y + z) = inner x y + inner x z" - by (rule inner_right_distrib) + by (rule inner_add_right) show "inner (scaleR r x) y = scaleR r (inner x y)" unfolding real_scaleR_def by (rule inner_scaleR_left) show "inner x (scaleR r y) = scaleR r (inner x y)" @@ -244,7 +249,7 @@ \ GDERIV (\x. g (f x)) x :> scaleR dg df" unfolding gderiv_def deriv_fderiv apply (drule (1) FDERIV_compose) - apply (simp add: inner_scaleR_right mult_ac) + apply (simp add: mult_ac) done lemma FDERIV_subst: "\FDERIV f x :> df; df = d\ \ FDERIV f x :> d" @@ -286,7 +291,7 @@ unfolding gderiv_def apply (rule FDERIV_subst) apply (erule (1) FDERIV_mult) - apply (simp add: inner_distrib inner_scaleR mult_ac) + apply (simp add: inner_add mult_ac) done lemma GDERIV_inverse: @@ -302,7 +307,7 @@ have 1: "FDERIV (\x. inner x x) x :> (\h. inner x h + inner h x)" by (intro inner.FDERIV FDERIV_ident) have 2: "(\h. inner x h + inner h x) = (\h. inner h (scaleR 2 x))" - by (simp add: expand_fun_eq inner_scaleR inner_commute) + by (simp add: expand_fun_eq inner_commute) have "0 < inner x x" using `x \ 0` by simp then have 3: "DERIV sqrt (inner x x) :> (inverse (sqrt (inner x x)) / 2)" by (rule DERIV_real_sqrt) diff -r 78ac5c304db7 -r ef6d67b1ad10 src/HOL/Library/Product_Vector.thy --- a/src/HOL/Library/Product_Vector.thy Thu Jun 11 19:49:02 2009 +0200 +++ b/src/HOL/Library/Product_Vector.thy Sat Jun 13 16:29:15 2009 +0200 @@ -72,6 +72,37 @@ end +lemma open_Times: "open S \ open T \ open (S \ T)" +unfolding open_prod_def by auto + +lemma fst_vimage_eq_Times: "fst -` S = S \ UNIV" +by auto + +lemma snd_vimage_eq_Times: "snd -` S = UNIV \ S" +by auto + +lemma open_vimage_fst: "open S \ open (fst -` S)" +by (simp add: fst_vimage_eq_Times open_Times) + +lemma open_vimage_snd: "open S \ open (snd -` S)" +by (simp add: snd_vimage_eq_Times open_Times) + +lemma closed_vimage_fst: "closed S \ closed (fst -` S)" +unfolding closed_open vimage_Compl [symmetric] +by (rule open_vimage_fst) + +lemma closed_vimage_snd: "closed S \ closed (snd -` S)" +unfolding closed_open vimage_Compl [symmetric] +by (rule open_vimage_snd) + +lemma closed_Times: "closed S \ closed T \ closed (S \ T)" +proof - + have "S \ T = (fst -` S) \ (snd -` T)" by auto + thus "closed S \ closed T \ closed (S \ T)" + by (simp add: closed_vimage_fst closed_vimage_snd closed_Int) +qed + + subsection {* Product is a metric space *} instantiation @@ -87,25 +118,21 @@ instance proof fix x y :: "'a \ 'b" show "dist x y = 0 \ x = y" - unfolding dist_prod_def - by (simp add: expand_prod_eq) + unfolding dist_prod_def expand_prod_eq by simp next fix x y z :: "'a \ 'b" show "dist x y \ dist x z + dist y z" unfolding dist_prod_def - apply (rule order_trans [OF _ real_sqrt_sum_squares_triangle_ineq]) - apply (rule real_sqrt_le_mono) - apply (rule order_trans [OF add_mono]) - apply (rule power_mono [OF dist_triangle2 [of _ _ "fst z"] zero_le_dist]) - apply (rule power_mono [OF dist_triangle2 [of _ _ "snd z"] zero_le_dist]) - apply (simp only: real_sum_squared_expand) - done + by (intro order_trans [OF _ real_sqrt_sum_squares_triangle_ineq] + real_sqrt_le_mono add_mono power_mono dist_triangle2 zero_le_dist) next (* FIXME: long proof! *) (* Maybe it would be easier to define topological spaces *) (* in terms of neighborhoods instead of open sets? *) fix S :: "('a \ 'b) set" show "open S \ (\x\S. \e>0. \y. dist y x < e \ y \ S)" + proof + assume "open S" thus "\x\S. \e>0. \y. dist y x < e \ y \ S" unfolding open_prod_def open_dist apply safe apply (drule (1) bspec) @@ -121,7 +148,11 @@ apply (erule le_less_trans [OF real_sqrt_sum_squares_ge1]) apply (drule spec, erule mp) apply (erule le_less_trans [OF real_sqrt_sum_squares_ge2]) - + done + next + assume "\x\S. \e>0. \y. dist y x < e \ y \ S" thus "open S" + unfolding open_prod_def open_dist + apply safe apply (drule (1) bspec) apply clarify apply (subgoal_tac "\r>0. \s>0. e = sqrt (r\ + s\)") @@ -132,14 +163,14 @@ apply clarify apply (rule_tac x="r - dist x a" in exI, rule conjI, simp) apply clarify - apply (rule le_less_trans [OF dist_triangle]) - apply (erule less_le_trans [OF add_strict_right_mono], simp) + apply (simp add: less_diff_eq) + apply (erule le_less_trans [OF dist_triangle]) apply (rule conjI) apply clarify apply (rule_tac x="s - dist x b" in exI, rule conjI, simp) apply clarify - apply (rule le_less_trans [OF dist_triangle]) - apply (erule less_le_trans [OF add_strict_right_mono], simp) + apply (simp add: less_diff_eq) + apply (erule le_less_trans [OF dist_triangle]) apply (rule conjI) apply simp apply (clarify, rename_tac c d) @@ -149,6 +180,7 @@ apply (rule_tac x="e / sqrt 2" in exI, simp add: divide_pos_pos) apply (simp add: power_divide) done + qed qed end @@ -161,7 +193,7 @@ lemma dist_snd_le: "dist (snd x) (snd y) \ dist x y" unfolding dist_prod_def by simp -lemma tendsto_fst: +lemma tendsto_fst [tendsto_intros]: assumes "(f ---> a) net" shows "((\x. fst (f x)) ---> fst a) net" proof (rule topological_tendstoI) @@ -180,7 +212,7 @@ by simp qed -lemma tendsto_snd: +lemma tendsto_snd [tendsto_intros]: assumes "(f ---> a) net" shows "((\x. snd (f x)) ---> snd a) net" proof (rule topological_tendstoI) @@ -199,7 +231,7 @@ by simp qed -lemma tendsto_Pair: +lemma tendsto_Pair [tendsto_intros]: assumes "(f ---> a) net" and "(g ---> b) net" shows "((\x. (f x, g x)) ---> (a, b)) net" proof (rule topological_tendstoI) @@ -315,7 +347,7 @@ done show "norm (scaleR r x) = \r\ * norm x" unfolding norm_prod_def - apply (simp add: norm_scaleR power_mult_distrib) + apply (simp add: power_mult_distrib) apply (simp add: right_distrib [symmetric]) apply (simp add: real_sqrt_mult_distrib) done @@ -349,10 +381,10 @@ by (simp add: inner_commute) show "inner (x + y) z = inner x z + inner y z" unfolding inner_prod_def - by (simp add: inner_left_distrib) + by (simp add: inner_add_left) show "inner (scaleR r x) y = r * inner x y" unfolding inner_prod_def - by (simp add: inner_scaleR_left right_distrib) + by (simp add: right_distrib) show "0 \ inner x x" unfolding inner_prod_def by (intro add_nonneg_nonneg inner_ge_zero) diff -r 78ac5c304db7 -r ef6d67b1ad10 src/HOL/Library/Topology_Euclidean_Space.thy --- a/src/HOL/Library/Topology_Euclidean_Space.thy Thu Jun 11 19:49:02 2009 +0200 +++ b/src/HOL/Library/Topology_Euclidean_Space.thy Sat Jun 13 16:29:15 2009 +0200 @@ -6,7 +6,7 @@ header {* Elementary topology in Euclidean space. *} theory Topology_Euclidean_Space -imports SEQ Euclidean_Space +imports SEQ Euclidean_Space Product_Vector begin declare fstcart_pastecart[simp] sndcart_pastecart[simp] @@ -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 @@ -1031,7 +1031,7 @@ unfolding trivial_limit_def Rep_net_at_infinity apply (clarsimp simp add: expand_set_eq) apply (drule_tac x="scaleR r (sgn 1)" in spec) - apply (simp add: norm_scaleR norm_sgn) + apply (simp add: norm_sgn) done lemma trivial_limit_sequentially: "\ trivial_limit sequentially" @@ -1245,17 +1245,16 @@ unfolding linear_conv_bounded_linear by (rule bounded_linear.tendsto) +lemma Lim_ident_at: "((\x. x) ---> a) (at a)" + unfolding tendsto_def Limits.eventually_at_topological by fast + lemma Lim_const: "((\x. a) ---> a) net" by (rule tendsto_const) 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" @@ -1277,38 +1276,34 @@ lemma Lim_null_norm: fixes f :: "'a \ 'b::real_normed_vector" - shows "(f ---> 0) net \ ((\x. vec1(norm(f x))) ---> 0) net" - by (simp add: Lim dist_norm norm_vec1) + shows "(f ---> 0) net \ ((\x. norm(f x)) ---> 0) net" + by (simp add: Lim dist_norm) lemma Lim_null_comparison: fixes f :: "'a \ 'b::real_normed_vector" - assumes "eventually (\x. norm(f x) <= g x) net" "((\x. vec1(g x)) ---> 0) net" + assumes "eventually (\x. norm (f x) \ g x) net" "(g ---> 0) net" shows "(f ---> 0) net" proof(simp add: tendsto_iff, rule+) fix e::real assume "0 g x" "dist (vec1 (g x)) 0 < e" - hence "dist (f x) 0 < e" unfolding vec_def using dist_vec1[of "g x" "0"] - by (vector dist_norm norm_vec1 real_vector_norm_def dot_def vec1_def) + assume "norm (f x) \ g x" "dist (g x) 0 < e" + hence "dist (f x) 0 < e" by (simp add: dist_norm) } thus "eventually (\x. dist (f x) 0 < e) net" - using eventually_and[of "\x. norm(f x) <= g x" "\x. dist (vec1 (g x)) 0 < e" net] - using eventually_mono[of "(\x. norm (f x) \ g x \ dist (vec1 (g x)) 0 < e)" "(\x. dist (f x) 0 < e)" net] + using eventually_and[of "\x. norm(f x) <= g x" "\x. dist (g x) 0 < e" net] + using eventually_mono[of "(\x. norm (f x) \ g x \ dist (g x) 0 < e)" "(\x. dist (f x) 0 < e)" net] using assms `e>0` unfolding tendsto_iff by auto qed -lemma Lim_component: "(f ---> l) net - ==> ((\a. vec1((f a :: real ^'n::finite)$i)) ---> vec1(l$i)) net" +lemma Lim_component: + fixes f :: "'a \ 'b::metric_space ^ 'n::finite" + shows "(f ---> l) net \ ((\a. f a $i) ---> l$i) net" unfolding tendsto_iff - apply (simp add: dist_norm vec1_sub[symmetric] norm_vec1 vector_minus_component[symmetric] del: vector_minus_component) - apply (auto simp del: vector_minus_component) - apply (erule_tac x=e in allE) - apply clarify - apply (erule eventually_rev_mono) - apply (auto simp del: vector_minus_component) - apply (rule order_le_less_trans) - apply (rule component_le_norm) - by auto + apply (clarify) + apply (drule spec, drule (1) mp) + apply (erule eventually_elim1) + apply (erule le_less_trans [OF dist_nth_le]) + done lemma Lim_transform_bound: fixes f :: "'a \ 'b::real_normed_vector" @@ -1504,12 +1499,6 @@ netlimit :: "'a::metric_space net \ 'a" where "netlimit net = (SOME a. \r>0. eventually (\x. dist x a < r) net)" -lemma dist_triangle3: - fixes x y :: "'a::metric_space" - shows "dist x y \ dist a x + dist a y" -using dist_triangle2 [of x y a] -by (simp add: dist_commute) - lemma netlimit_within: assumes "\ trivial_limit (at a within S)" shows "netlimit (at a within S) = a" @@ -1694,14 +1683,14 @@ apply (subgoal_tac "\N k (n::nat). N + k <= n ==> N <= n - k \ (n - k) + k = n") by metis arith -lemma seq_harmonic: "((\n. vec1(inverse (real n))) ---> 0) sequentially" +lemma seq_harmonic: "((\n. inverse (real n)) ---> 0) sequentially" proof- { fix e::real assume "e>0" hence "\N::nat. \n::nat\N. inverse (real n) < e" using real_arch_inv[of e] apply auto apply(rule_tac x=n in exI) - by (metis dlo_simps(4) le_imp_inverse_le linorder_not_less real_of_nat_gt_zero_cancel_iff real_of_nat_less_iff xt1(7)) + by (metis not_le le_imp_inverse_le not_less real_of_nat_gt_zero_cancel_iff real_of_nat_less_iff xt1(7)) } - thus ?thesis unfolding Lim_sequentially dist_norm apply simp unfolding norm_vec1 by auto + thus ?thesis unfolding Lim_sequentially dist_norm by simp qed text{* More properties of closed balls. *} @@ -1768,7 +1757,7 @@ also have "\ = \- 1 + d / (2 * norm (x - y))\ * norm (x - y)" using scaleR_left_distrib[of "- 1" "d / (2 * norm (y - x))", THEN sym, of "y - x"] unfolding scaleR_minus_left scaleR_one - by (auto simp add: norm_minus_commute norm_scaleR) + by (auto simp add: norm_minus_commute) also have "\ = \- norm (x - y) + d / 2\" unfolding abs_mult_pos[of "norm (x - y)", OF norm_ge_zero[of "x - y"]] unfolding real_add_mult_distrib using `x\y`[unfolded dist_nz, unfolded dist_norm] by auto @@ -1780,7 +1769,7 @@ have "(d / (2*dist y x)) *\<^sub>R (y - x) \ 0" using `x\y`[unfolded dist_nz] `d>0` unfolding scaleR_eq_0_iff by (auto simp add: dist_commute) moreover - have "dist (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) y < d" unfolding dist_norm apply simp unfolding norm_minus_cancel norm_scaleR + have "dist (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) y < d" unfolding dist_norm apply simp unfolding norm_minus_cancel using `d>0` `x\y`[unfolded dist_nz] dist_commute[of x y] unfolding dist_norm by auto ultimately show "\x'\ball x e. x' \ y \ dist x' y < d" by (rule_tac x="y - (d / (2*dist y x)) *\<^sub>R (y - x)" in bexI) auto @@ -1819,11 +1808,11 @@ unfolding z_def by (simp add: algebra_simps) have "dist z y < r" unfolding z_def k_def using `0 < r` - by (simp add: dist_norm norm_scaleR min_def) + by (simp add: dist_norm min_def) hence "z \ T" using `\z. dist z y < r \ z \ T` by simp have "dist x z < dist x y" unfolding z_def2 dist_norm - apply (simp add: norm_scaleR norm_minus_commute) + apply (simp add: norm_minus_commute) apply (simp only: dist_norm [symmetric]) apply (subgoal_tac "\1 - k\ * dist x y < 1 * dist x y", simp) apply (rule mult_strict_right_mono) @@ -1856,9 +1845,10 @@ apply (simp add: zero_less_dist_iff) done +(* In a trivial vector space, this fails for e = 0. *) lemma interior_cball: - fixes x :: "real ^ _" (* FIXME: generalize *) - shows "interior(cball x e) = ball x e" + fixes x :: "'a::{real_normed_vector, perfect_space}" + shows "interior (cball x e) = ball x e" proof(cases "e\0") case False note cs = this from cs have "ball x e = {}" using ball_empty[of e x] by auto moreover @@ -1873,9 +1863,9 @@ { fix S y assume as: "S \ cball x e" "open S" "y\S" then obtain d where "d>0" and d:"\x'. dist x' y < d \ x' \ S" unfolding open_dist by blast - then obtain xa where xa:"dist y xa = d / 2" using vector_choose_dist[of "d/2" y] by auto - hence xa_y:"xa \ y" using dist_nz[of y xa] using `d>0` by auto - have "xa\S" using d[THEN spec[where x=xa]] using xa apply(auto simp add: dist_commute) unfolding dist_nz[THEN sym] using xa_y by auto + then obtain xa where xa_y: "xa \ y" and xa: "dist xa y < d" + using perfect_choose_dist [of d] by auto + have "xa\S" using d[THEN spec[where x=xa]] using xa by(auto simp add: dist_commute) hence xa_cball:"xa \ cball x e" using as(1) by auto hence "y \ ball x e" proof(cases "x = y") @@ -1884,18 +1874,19 @@ thus "y \ ball x e" using `x = y ` by simp next case False - have "dist (y + (d / 2 / dist y x) *s (y - x)) y < d" unfolding dist_norm + have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) y < d" unfolding dist_norm using `d>0` norm_ge_zero[of "y - x"] `x \ y` by auto - hence *:"y + (d / 2 / dist y x) *s (y - x) \ cball x e" using d as(1)[unfolded subset_eq] by blast + hence *:"y + (d / 2 / dist y x) *\<^sub>R (y - x) \ cball x e" using d as(1)[unfolded subset_eq] by blast have "y - x \ 0" using `x \ y` by auto hence **:"d / (2 * norm (y - x)) > 0" unfolding zero_less_norm_iff[THEN sym] using `d>0` divide_pos_pos[of d "2*norm (y - x)"] by auto - have "dist (y + (d / 2 / dist y x) *s (y - x)) x = norm (y + (d / (2 * norm (y - x))) *s y - (d / (2 * norm (y - x))) *s x - x)" - by (auto simp add: dist_norm vector_ssub_ldistrib add_diff_eq) - also have "\ = norm ((1 + d / (2 * norm (y - x))) *s (y - x))" - by (auto simp add: vector_sadd_rdistrib vector_smult_lid ring_simps vector_sadd_rdistrib vector_ssub_ldistrib) - also have "\ = \1 + d / (2 * norm (y - x))\ * norm (y - x)" using ** by auto + have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) x = norm (y + (d / (2 * norm (y - x))) *\<^sub>R y - (d / (2 * norm (y - x))) *\<^sub>R x - x)" + by (auto simp add: dist_norm algebra_simps) + also have "\ = norm ((1 + d / (2 * norm (y - x))) *\<^sub>R (y - x))" + by (auto simp add: algebra_simps) + also have "\ = \1 + d / (2 * norm (y - x))\ * norm (y - x)" + using ** by auto also have "\ = (dist y x) + d/2"using ** by (auto simp add: left_distrib dist_norm) finally have "e \ dist x y +d/2" using *[unfolded mem_cball] by (auto simp add: dist_commute) thus "y \ ball x e" unfolding mem_ball using `d>0` by auto @@ -1905,14 +1896,14 @@ qed lemma frontier_ball: - fixes a :: "real ^ _" (* FIXME: generalize *) + fixes a :: "'a::real_normed_vector" shows "0 < e ==> frontier(ball a e) = {x. dist a x = e}" apply (simp add: frontier_def closure_ball interior_open open_ball order_less_imp_le) apply (simp add: expand_set_eq) by arith lemma frontier_cball: - fixes a :: "real ^ _" (* FIXME: generalize *) + fixes a :: "'a::{real_normed_vector, perfect_space}" shows "frontier(cball a e) = {x. dist a x = e}" apply (simp add: frontier_def interior_cball closed_cball closure_closed order_less_imp_le) apply (simp add: expand_set_eq) @@ -1924,20 +1915,20 @@ lemma cball_empty: "e < 0 ==> cball x e = {}" by (simp add: cball_eq_empty) lemma cball_eq_sing: - fixes x :: "real ^ _" (* FIXME: generalize *) + fixes x :: "'a::perfect_space" shows "(cball x e = {x}) \ e = 0" -proof- - { assume as:"\xa. (dist x xa \ e) = (xa = x)" - hence "e \ 0" apply (erule_tac x=x in allE) by auto - then obtain y where y:"dist x y = e" using vector_choose_dist[of e] by auto - hence "e = 0" using as apply(erule_tac x=y in allE) by auto - } - thus ?thesis unfolding expand_set_eq mem_cball by (auto simp add: dist_nz) -qed +proof (rule linorder_cases) + assume e: "0 < e" + obtain a where "a \ x" "dist a x < e" + using perfect_choose_dist [OF e] by auto + hence "a \ x" "dist x a \ e" by (auto simp add: dist_commute) + with e show ?thesis by (auto simp add: expand_set_eq) +qed auto lemma cball_sing: - fixes x :: "real ^ _" (* FIXME: generalize *) - shows "e = 0 ==> cball x e = {x}" by (simp add: cball_eq_sing) + fixes x :: "'a::metric_space" + shows "e = 0 ==> cball x e = {x}" + by (auto simp add: expand_set_eq) text{* For points in the interior, localization of limits makes no difference. *} @@ -2082,7 +2073,7 @@ fix b::real assume b: "b >0" have b1: "b +1 \ 0" using b by simp with `x \ 0` have "b < norm (scaleR (b + 1) (sgn x))" - by (simp add: norm_scaleR norm_sgn) + by (simp add: norm_sgn) then show "\x::'a. b < norm x" .. qed @@ -2104,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" @@ -2123,26 +2116,26 @@ text{* Some theorems on sups and infs using the notion "bounded". *} -lemma bounded_vec1: +lemma bounded_real: fixes S :: "real set" - shows "bounded(vec1 ` S) \ (\a. \x\S. abs x <= a)" - by (simp add: bounded_iff forall_vec1 norm_vec1 vec1_in_image_vec1) - -lemma bounded_has_rsup: assumes "bounded(vec1 ` S)" "S \ {}" + shows "bounded S \ (\a. \x\S. abs x <= a)" + by (simp add: bounded_iff) + +lemma bounded_has_rsup: assumes "bounded S" "S \ {}" shows "\x\S. x <= rsup S" and "\b. (\x\S. x <= b) \ rsup S <= b" proof fix x assume "x\S" - from assms(1) obtain a where a:"\x\S. \x\ \ a" unfolding bounded_vec1 by auto + from assms(1) obtain a where a:"\x\S. \x\ \ a" unfolding bounded_real by auto hence *:"S *<= a" using setleI[of S a] by (metis abs_le_interval_iff mem_def) - thus "x \ rsup S" using rsup[OF `S\{}`] using assms(1)[unfolded bounded_vec1] using isLubD2[of UNIV S "rsup S" x] using `x\S` by auto + thus "x \ rsup S" using rsup[OF `S\{}`] using assms(1)[unfolded bounded_real] using isLubD2[of UNIV S "rsup S" x] using `x\S` by auto next show "\b. (\x\S. x \ b) \ rsup S \ b" using assms using rsup[of S, unfolded isLub_def isUb_def leastP_def setle_def setge_def] - apply (auto simp add: bounded_vec1) + apply (auto simp add: bounded_real) by (auto simp add: isLub_def isUb_def leastP_def setle_def setge_def) qed -lemma rsup_insert: assumes "bounded (vec1 ` S)" +lemma rsup_insert: assumes "bounded S" shows "rsup(insert x S) = (if S = {} then x else max x (rsup S))" proof(cases "S={}") case True thus ?thesis using rsup_finite_in[of "{x}"] by auto @@ -2168,17 +2161,17 @@ by simp lemma bounded_has_rinf: - assumes "bounded(vec1 ` S)" "S \ {}" + assumes "bounded S" "S \ {}" shows "\x\S. x >= rinf S" and "\b. (\x\S. x >= b) \ rinf S >= b" proof fix x assume "x\S" - from assms(1) obtain a where a:"\x\S. \x\ \ a" unfolding bounded_vec1 by auto + from assms(1) obtain a where a:"\x\S. \x\ \ a" unfolding bounded_real by auto hence *:"- a <=* S" using setgeI[of S "-a"] unfolding abs_le_interval_iff by auto thus "x \ rinf S" using rinf[OF `S\{}`] using isGlbD2[of UNIV S "rinf S" x] using `x\S` by auto next show "\b. (\x\S. x >= b) \ rinf S \ b" using assms using rinf[of S, unfolded isGlb_def isLb_def greatestP_def setle_def setge_def] - apply (auto simp add: bounded_vec1) + apply (auto simp add: bounded_real) by (auto simp add: isGlb_def isLb_def greatestP_def setle_def setge_def) qed @@ -2189,7 +2182,7 @@ apply (blast intro!: order_antisym dest!: isGlb_le_isLb) done -lemma rinf_insert: assumes "bounded (vec1 ` S)" +lemma rinf_insert: assumes "bounded S" shows "rinf(insert x S) = (if S = {} then x else min x (rinf S))" (is "?lhs = ?rhs") proof(cases "S={}") case True thus ?thesis using rinf_finite_in[of "{x}"] by auto @@ -2217,8 +2210,8 @@ definition compact :: "'a::metric_space set \ bool" where (* TODO: generalize *) "compact S \ - (\f. (\n::nat. f n \ S) \ - (\l\S. \r. (\m n. m < n \ r m < r n) \ ((f o r) ---> l) sequentially))" + (\f. (\n. f n \ S) \ + (\l\S. \r. subseq r \ ((f o r) ---> l) sequentially))" text {* A metric space (or topological vector space) is said to have the @@ -2227,44 +2220,43 @@ class heine_borel = assumes bounded_imp_convergent_subsequence: - "bounded s \ \n::nat. f n \ s - \ \l r. (\m n. m < n --> r m < r n) \ ((f \ r) ---> l) sequentially" + "bounded s \ \n. f n \ s + \ \l r. subseq r \ ((f \ r) ---> l) sequentially" lemma bounded_closed_imp_compact: fixes s::"'a::heine_borel set" assumes "bounded s" and "closed s" shows "compact s" proof (unfold compact_def, clarify) fix f :: "nat \ 'a" assume f: "\n. f n \ s" - obtain l r where r: "\m n. m < n \ r m < r n" and l: "((f \ r) ---> l) sequentially" + obtain l r where r: "subseq r" and l: "((f \ r) ---> l) sequentially" using bounded_imp_convergent_subsequence [OF `bounded s` `\n. f n \ s`] by auto from f have fr: "\n. (f \ r) n \ s" by simp have "l \ s" using `closed s` fr l unfolding closed_sequential_limits by blast - show "\l\s. \r. (\m n. m < n \ r m < r n) \ ((f \ r) ---> l) sequentially" + show "\l\s. \r. subseq r \ ((f \ r) ---> l) sequentially" using `l \ s` r l by blast qed -lemma monotone_bigger: fixes r::"nat\nat" - assumes "\m n::nat. m < n --> r m < r n" - shows "n \ r n" +lemma subseq_bigger: assumes "subseq r" shows "n \ r n" proof(induct n) show "0 \ r 0" by auto next fix n assume "n \ r n" - moreover have "r n < r (Suc n)" using assms by auto + moreover have "r n < r (Suc n)" + using assms [unfolded subseq_def] by auto ultimately show "Suc n \ r (Suc n)" by auto qed -lemma eventually_subsequence: - assumes r: "\m n. m < n \ r m < r n" +lemma eventually_subseq: + assumes r: "subseq r" shows "eventually P sequentially \ eventually (\n. P (r n)) sequentially" unfolding eventually_sequentially -by (metis monotone_bigger [OF r] le_trans) - -lemma lim_subsequence: - fixes l :: "'a::metric_space" (* TODO: generalize *) - shows "\m n. m < n \ r m < r n \ (s ---> l) sequentially \ ((s o r) ---> l) sequentially" -unfolding Lim_sequentially by (simp, metis monotone_bigger le_trans) +by (metis subseq_bigger [OF r] le_trans) + +lemma lim_subseq: + "subseq r \ (s ---> l) sequentially \ ((s o r) ---> l) sequentially" +unfolding tendsto_def eventually_sequentially o_def +by (metis subseq_bigger le_trans) lemma num_Axiom: "EX! g. g 0 = e \ (\n. g (Suc n) = f n (g n))" unfolding Ex1_def @@ -2280,9 +2272,8 @@ apply (simp) done - lemma convergent_bounded_increasing: fixes s ::"nat\real" - assumes "\m n. m \ n --> s m \ s n" and "\n. abs(s n) \ b" + assumes "incseq s" and "\n. abs(s n) \ b" shows "\ l. \e::real>0. \ N. \n \ N. abs(s n - l) < e" proof- have "isUb UNIV (range s) b" using assms(2) and abs_le_D1 unfolding isUb_def and setle_def by auto @@ -2292,27 +2283,27 @@ obtain N where "N\n" and n:"\s N - t\ \ e" using as[THEN spec[where x=n]] by auto have "t \ s N" using isLub_isUb[OF t, unfolded isUb_def setle_def] by auto with n have "s N \ t - e" using `e>0` by auto - hence "s n \ t - e" using assms(1)[THEN spec[where x=n], THEN spec[where x=N]] using `n\N` by auto } + hence "s n \ t - e" using assms(1)[unfolded incseq_def, THEN spec[where x=n], THEN spec[where x=N]] using `n\N` by auto } hence "isUb UNIV (range s) (t - e)" unfolding isUb_def and setle_def by auto hence False using isLub_le_isUb[OF t, of "t - e"] and `e>0` by auto } thus ?thesis by blast qed lemma convergent_bounded_monotone: fixes s::"nat \ real" - assumes "\n. abs(s n) \ b" and "(\m n. m \ n --> s m \ s n) \ (\m n. m \ n --> s n \ s m)" + assumes "\n. abs(s n) \ b" and "monoseq s" shows "\l. \e::real>0. \N. \n\N. abs(s n - l) < e" using convergent_bounded_increasing[of s b] assms using convergent_bounded_increasing[of "\n. - s n" b] + unfolding monoseq_def incseq_def apply auto unfolding minus_add_distrib[THEN sym, unfolded diff_minus[THEN sym]] unfolding abs_minus_cancel by(rule_tac x="-l" in exI)auto lemma compact_real_lemma: assumes "\n::nat. abs(s n) \ b" - shows "\(l::real) r. (\m n::nat. m < n --> r m < r n) \ ((s \ r) ---> l) sequentially" + shows "\(l::real) r. subseq r \ ((s \ r) ---> l) sequentially" proof- - obtain r where r:"\m n::nat. m < n \ r m < r n" - "(\m n. m \ n \ s (r m) \ s (r n)) \ (\m n. m \ n \ s (r n) \ s (r m))" - using seq_monosub[of s] by (auto simp add: subseq_def monoseq_def) - thus ?thesis using convergent_bounded_monotone[of "s o r" b] and assms + obtain r where r:"subseq r" "monoseq (\n. s (r n))" + using seq_monosub[of s] by auto + thus ?thesis using convergent_bounded_monotone[of "\n. s (r n)" b] and assms unfolding tendsto_iff dist_norm eventually_sequentially by auto qed @@ -2323,9 +2314,9 @@ then obtain b where b: "\n. abs (f n) \ b" unfolding bounded_iff by auto obtain l :: real and r :: "nat \ nat" where - r: "\m n. m < n \ r m < r n" and l: "((f \ r) ---> l) sequentially" + r: "subseq r" and l: "((f \ r) ---> l) sequentially" using compact_real_lemma [OF b] by auto - thus "\l r. (\m n. m < n \ r m < r n) \ ((f \ r) ---> l) sequentially" + thus "\l r. subseq r \ ((f \ r) ---> l) sequentially" by auto qed @@ -2342,28 +2333,29 @@ fixes f :: "nat \ 'a::heine_borel ^ 'n::finite" assumes "bounded s" and "\n. f n \ s" shows "\d. - \l r. (\n m::nat. m < n --> r m < r n) \ + \l r. subseq r \ (\e>0. eventually (\n. \i\d. dist (f (r n) $ i) (l $ i) < e) sequentially)" proof fix d::"'n set" have "finite d" by simp - thus "\l::'a ^ 'n. \r. (\n m::nat. m < n --> r m < r n) \ + thus "\l::'a ^ 'n. \r. subseq r \ (\e>0. eventually (\n. \i\d. dist (f (r n) $ i) (l $ i) < e) sequentially)" - proof(induct d) case empty thus ?case by auto + proof(induct d) case empty thus ?case unfolding subseq_def by auto next case (insert k d) have s': "bounded ((\x. x $ k) ` s)" using `bounded s` by (rule bounded_component) - obtain l1::"'a^'n" and r1 where r1:"\n m::nat. m < n \ r1 m < r1 n" and lr1:"\e>0. eventually (\n. \i\d. dist (f (r1 n) $ i) (l1 $ i) < e) sequentially" + obtain l1::"'a^'n" and r1 where r1:"subseq r1" and lr1:"\e>0. eventually (\n. \i\d. dist (f (r1 n) $ i) (l1 $ i) < e) sequentially" using insert(3) by auto have f': "\n. f (r1 n) $ k \ (\x. x $ k) ` s" using `\n. f n \ s` by simp - obtain l2 r2 where r2:"\m n::nat. m < n \ r2 m < r2 n" and lr2:"((\i. f (r1 (r2 i)) $ k) ---> l2) sequentially" + obtain l2 r2 where r2:"subseq r2" and lr2:"((\i. f (r1 (r2 i)) $ k) ---> l2) sequentially" using bounded_imp_convergent_subsequence[OF s' f'] unfolding o_def by auto - def r \ "r1 \ r2" have r:"\m n. m < n \ r m < r n" unfolding r_def o_def using r1 and r2 by auto + def r \ "r1 \ r2" have r:"subseq r" + using r1 and r2 unfolding r_def o_def subseq_def by auto moreover def l \ "(\ i. if i = k then l2 else l1$i)::'a^'n" { fix e::real assume "e>0" from lr1 `e>0` have N1:"eventually (\n. \i\d. dist (f (r1 n) $ i) (l1 $ i) < e) sequentially" by blast from lr2 `e>0` have N2:"eventually (\n. dist (f (r1 (r2 n)) $ k) l2 < e) sequentially" by (rule tendstoD) from r2 N1 have N1': "eventually (\n. \i\d. dist (f (r1 (r2 n)) $ i) (l1 $ i) < e) sequentially" - by (rule eventually_subsequence) + by (rule eventually_subseq) have "eventually (\n. \i\(insert k d). dist (f (r n) $ i) (l $ i) < e) sequentially" using N1' N2 by (rule eventually_elim2, simp add: l_def r_def) } @@ -2375,7 +2367,7 @@ proof fix s :: "('a ^ 'b) set" and f :: "nat \ 'a ^ 'b" assume s: "bounded s" and f: "\n. f n \ s" - then obtain l r where r: "\n m::nat. m < n --> r m < r n" + then obtain l r where r: "subseq r" and l: "\e>0. eventually (\n. \i\UNIV. dist (f (r n) $ i) (l $ i) < e) sequentially" using compact_lemma [OF s f] by blast let ?d = "UNIV::'b set" @@ -2396,7 +2388,55 @@ by (rule eventually_elim1) } hence *:"((f \ r) ---> l) sequentially" unfolding o_def tendsto_iff by simp - with r show "\l r. (\m n. m < n \ r m < r n) \ ((f \ r) ---> l) sequentially" by auto + with r show "\l r. subseq r \ ((f \ r) ---> l) sequentially" by auto +qed + +lemma bounded_fst: "bounded s \ bounded (fst ` s)" +unfolding bounded_def +apply clarify +apply (rule_tac x="a" in exI) +apply (rule_tac x="e" in exI) +apply clarsimp +apply (drule (1) bspec) +apply (simp add: dist_Pair_Pair) +apply (erule order_trans [OF real_sqrt_sum_squares_ge1]) +done + +lemma bounded_snd: "bounded s \ bounded (snd ` s)" +unfolding bounded_def +apply clarify +apply (rule_tac x="b" in exI) +apply (rule_tac x="e" in exI) +apply clarsimp +apply (drule (1) bspec) +apply (simp add: dist_Pair_Pair) +apply (erule order_trans [OF real_sqrt_sum_squares_ge2]) +done + +instance "*" :: (heine_borel, heine_borel) heine_borel +proof + fix s :: "('a * 'b) set" and f :: "nat \ 'a * 'b" + assume s: "bounded s" and f: "\n. f n \ s" + from s have s1: "bounded (fst ` s)" by (rule bounded_fst) + from f have f1: "\n. fst (f n) \ fst ` s" by simp + obtain l1 r1 where r1: "subseq r1" + and l1: "((\n. fst (f (r1 n))) ---> l1) sequentially" + using bounded_imp_convergent_subsequence [OF s1 f1] + unfolding o_def by fast + from s have s2: "bounded (snd ` s)" by (rule bounded_snd) + from f have f2: "\n. snd (f (r1 n)) \ snd ` s" by simp + obtain l2 r2 where r2: "subseq r2" + and l2: "((\n. snd (f (r1 (r2 n)))) ---> l2) sequentially" + using bounded_imp_convergent_subsequence [OF s2 f2] + unfolding o_def by fast + have l1': "((\n. fst (f (r1 (r2 n)))) ---> l1) sequentially" + using lim_subseq [OF r2 l1] unfolding o_def . + have l: "((f \ (r1 \ r2)) ---> (l1, l2)) sequentially" + using tendsto_Pair [OF l1' l2] unfolding o_def by simp + have r: "subseq (r1 \ r2)" + using r1 r2 unfolding subseq_def by simp + show "\l r. subseq r \ ((f \ r) ---> l) sequentially" + using l r by fast qed subsection{* Completeness. *} @@ -2461,14 +2501,9 @@ lemma compact_imp_complete: assumes "compact s" shows "complete s" proof- { fix f assume as: "(\n::nat. f n \ s)" "Cauchy f" - from as(1) obtain l r where lr: "l\s" "(\m n. m < n \ r m < r n)" "((f \ r) ---> l) sequentially" using assms unfolding compact_def by blast - - { fix n :: nat have lr':"n \ r n" - proof (induct n) - show "0 \ r 0" using lr(2) by blast - next fix na assume "na \ r na" moreover have "na < Suc na \ r na < r (Suc na)" using lr(2) by blast - ultimately show "Suc na \ r (Suc na)" by auto - qed } note lr' = this + from as(1) obtain l r where lr: "l\s" "subseq r" "((f \ r) ---> l) sequentially" using assms unfolding compact_def by blast + + note lr' = subseq_bigger [OF lr(2)] { fix e::real assume "e>0" from as(2) obtain N where N:"\m n. N \ m \ N \ n \ dist (f m) (f n) < e/2" unfolding cauchy_def using `e>0` apply (erule_tac x="e/2" in allE) by auto @@ -2575,12 +2610,12 @@ thus "x n \ s \ (\m dist (x m) (x n) < e)" unfolding Q_def by auto qed } hence "\n::nat. x n \ s" and x:"\n. \m < n. \ (dist (x m) (x n) < e)" by blast+ - then obtain l r where "l\s" and r:"\m n. m < n \ r m < r n" and "((x \ r) ---> l) sequentially" using assms(1)[unfolded compact_def, THEN spec[where x=x]] by auto + then obtain l r where "l\s" and r:"subseq r" and "((x \ r) ---> l) sequentially" using assms(1)[unfolded compact_def, THEN spec[where x=x]] by auto from this(3) have "Cauchy (x \ r)" using convergent_imp_cauchy by auto then obtain N::nat where N:"\m n. N \ m \ N \ n \ dist ((x \ r) m) ((x \ r) n) < e" unfolding cauchy_def using `e>0` by auto show False using N[THEN spec[where x=N], THEN spec[where x="N+1"]] - using r[THEN spec[where x=N], THEN spec[where x="N+1"]] + using r[unfolded subseq_def, THEN spec[where x=N], THEN spec[where x="N+1"]] using x[THEN spec[where x="r (N+1)"], THEN spec[where x="r (N)"]] by auto qed @@ -2599,7 +2634,7 @@ then obtain f where f:"\n::nat. f n \ s \ (\xa\t. \ ball (f n) (inverse (real (n + 1))) \ xa)" using choice[of "\n::nat. \x. x\s \ (\xa\t. \ ball x (inverse (real (n + 1))) \ xa)"] by auto - then obtain l r where l:"l\s" and r:"\m n. m < n \ r m < r n" and lr:"((f \ r) ---> l) sequentially" + then obtain l r where l:"l\s" and r:"subseq r" and lr:"((f \ r) ---> l) sequentially" using assms(1)[unfolded compact_def, THEN spec[where x=f]] by auto obtain b where "l\b" "b\t" using assms(2) and l by auto @@ -2612,7 +2647,7 @@ obtain N2::nat where N2:"N2>0" "inverse (real N2) < e /2" using real_arch_inv[of "e/2"] and `e>0` by auto have N2':"inverse (real (r (N1 + N2) +1 )) < e/2" apply(rule order_less_trans) apply(rule less_imp_inverse_less) using N2 - using monotone_bigger[OF r, of "N1 + N2"] by auto + using subseq_bigger[OF r, of "N1 + N2"] by auto def x \ "(f (r (N1 + N2)))" have x:"\ ball x (inverse (real (r (N1 + N2) + 1))) \ b" unfolding x_def @@ -2926,7 +2961,7 @@ by blast lemma compact_sing [simp]: "compact {a}" - unfolding compact_def o_def + unfolding compact_def o_def subseq_def by (auto simp add: tendsto_const) lemma compact_cball[simp]: @@ -2987,7 +3022,7 @@ from assms(2) obtain x where x:"\n::nat. x n \ s n" using choice[of "\n x. x\ s n"] by auto from assms(4,1) have *:"compact (s 0)" using bounded_closed_imp_compact[of "s 0"] by auto - then obtain l r where lr:"l\s 0" "\m n. m < n \ r m < r n" "((x \ r) ---> l) sequentially" + then obtain l r where lr:"l\s 0" "subseq r" "((x \ r) ---> l) sequentially" unfolding compact_def apply(erule_tac x=x in allE) using x using assms(3) by blast { fix n::nat @@ -2995,7 +3030,7 @@ with lr(3) obtain N where N:"\m\N. dist ((x \ r) m) l < e" unfolding Lim_sequentially by auto hence "dist ((x \ r) (max N n)) l < e" by auto moreover - have "r (max N n) \ n" using lr(2) using monotone_bigger[of r "max N n"] by auto + have "r (max N n) \ n" using lr(2) using subseq_bigger[of r "max N n"] by auto hence "(x \ r) (max N n) \ s n" using x apply(erule_tac x=n in allE) using x apply(erule_tac x="r (max N n)" in allE) @@ -3415,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: @@ -3441,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: @@ -3470,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 @@ -3701,13 +3736,11 @@ qed lemma continuous_open_preimage_univ: - fixes f :: "real ^ _ \ real ^ _" (* FIXME: generalize *) - shows "\x. continuous (at x) f \ open s \ open {x. f x \ s}" + "\x. continuous (at x) f \ open s \ open {x. f x \ s}" using continuous_open_preimage[of UNIV f s] open_UNIV continuous_at_imp_continuous_on by auto lemma continuous_closed_preimage_univ: - fixes f :: "real ^ _ \ real ^ _" (* FIXME: generalize *) - shows "(\x. continuous (at x) f) \ closed s ==> closed {x. f x \ s}" + "(\x. continuous (at x) f) \ closed s ==> closed {x. f x \ s}" using continuous_closed_preimage[of UNIV f s] closed_UNIV continuous_at_imp_continuous_on by auto text{* Equality of continuous functions on closure and related results. *} @@ -3802,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 @@ -3825,12 +3858,13 @@ 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 :: "(real ^ _) set" (* FIXME: generalize *) + fixes s :: "'a::real_normed_vector set" assumes "open s" shows "open((\x. a + x) ` s)" proof- { fix x have "continuous (at x) (\x. x - a)" using continuous_sub[of "at x" "\x. x" "\x. a"] continuous_at_id[of x] continuous_const[of "at x" a] by auto } @@ -3841,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: @@ -3875,13 +3909,13 @@ proof- { fix x assume x:"\n::nat. x n \ f ` s" then obtain y where y:"\n. y n \ s \ x n = f (y n)" unfolding image_iff Bex_def using choice[of "\n xa. xa \ s \ x n = f xa"] by auto - then obtain l r where "l\s" and r:"\m n. m < n \ r m < r n" and lr:"((y \ r) ---> l) sequentially" using assms(2)[unfolded compact_def, THEN spec[where x=y]] by auto + then obtain l r where "l\s" and r:"subseq r" and lr:"((y \ r) ---> l) sequentially" using assms(2)[unfolded compact_def, THEN spec[where x=y]] by auto { fix e::real assume "e>0" then obtain d where "d>0" and d:"\x'\s. dist x' l < d \ dist (f x') (f l) < e" using assms(1)[unfolded continuous_on_def, THEN bspec[where x=l], OF `l\s`] by auto then obtain N::nat where N:"\n\N. dist ((y \ r) n) l < d" using lr[unfolded Lim_sequentially, THEN spec[where x=d]] by auto { fix n::nat assume "n\N" hence "dist ((x \ r) n) (f l) < e" using N[THEN spec[where x=n]] d[THEN bspec[where x="y (r n)"]] y[THEN spec[where x="r n"]] by auto } hence "\N. \n\N. dist ((x \ r) n) (f l) < e" by auto } - hence "\l\f ` s. \r. (\m n. m < n \ r m < r n) \ ((x \ r) ---> l) sequentially" unfolding Lim_sequentially using r lr `l\s` by auto } + hence "\l\f ` s. \r. subseq r \ ((x \ r) ---> l) sequentially" unfolding Lim_sequentially using r lr `l\s` by auto } thus ?thesis unfolding compact_def by auto qed @@ -3938,7 +3972,8 @@ text{* Continuity of inverse function on compact domain. *} lemma continuous_on_inverse: - fixes f :: "real ^ _ \ real ^ _" + fixes f :: "'a::heine_borel \ 'b::heine_borel" + (* TODO: can this be generalized more? *) assumes "continuous_on s f" "compact s" "\x \ s. g (f x) = x" shows "continuous_on (f ` s) g" proof- @@ -4050,68 +4085,52 @@ subsection{* Topological stuff lifted from and dropped to R *} -lemma open_vec1: - fixes s :: "real set" shows - "open(vec1 ` s) \ - (\x \ s. \e>0. \x'. abs(x' - x) < e --> x' \ s)" (is "?lhs = ?rhs") - unfolding open_dist apply simp unfolding forall_vec1 dist_vec1 vec1_in_image_vec1 by simp - -lemma islimpt_approachable_vec1: +lemma open_real: fixes s :: "real set" shows - "(vec1 x) islimpt (vec1 ` s) \ - (\e>0. \x'\ s. x' \ x \ abs(x' - x) < e)" - by (auto simp add: islimpt_approachable dist_vec1 vec1_eq) - -lemma closed_vec1: - fixes s :: "real set" shows - "closed (vec1 ` s) \ + "open s \ + (\x \ s. \e>0. \x'. abs(x' - x) < e --> x' \ s)" (is "?lhs = ?rhs") + unfolding open_dist dist_norm by simp + +lemma islimpt_approachable_real: + fixes s :: "real set" + shows "x islimpt s \ (\e>0. \x'\ s. x' \ x \ abs(x' - x) < e)" + unfolding islimpt_approachable dist_norm by simp + +lemma closed_real: + fixes s :: "real set" + shows "closed s \ (\x. (\e>0. \x' \ s. x' \ x \ abs(x' - x) < e) --> x \ s)" - unfolding closed_limpt islimpt_approachable forall_vec1 apply simp - unfolding dist_vec1 vec1_in_image_vec1 abs_minus_commute by auto - -lemma continuous_at_vec1_range: - fixes f :: "real ^ _ \ real" - shows "continuous (at x) (vec1 o f) \ (\e>0. \d>0. + unfolding closed_limpt islimpt_approachable dist_norm by simp + +lemma continuous_at_real_range: + fixes f :: "'a::real_normed_vector \ real" + shows "continuous (at x) f \ (\e>0. \d>0. \x'. norm(x' - x) < d --> abs(f x' - f x) < e)" - unfolding continuous_at unfolding Lim_at apply simp unfolding dist_vec1 unfolding dist_nz[THEN sym] unfolding dist_norm apply auto + unfolding continuous_at unfolding Lim_at + unfolding dist_nz[THEN sym] unfolding dist_norm apply auto apply(erule_tac x=e in allE) apply auto apply (rule_tac x=d in exI) apply auto apply (erule_tac x=x' in allE) apply auto apply(erule_tac x=e in allE) by auto -lemma continuous_on_vec1_range: +lemma continuous_on_real_range: fixes f :: "'a::real_normed_vector \ real" - shows "continuous_on s (vec1 o f) \ (\x \ s. \e>0. \d>0. (\x' \ s. norm(x' - x) < d --> abs(f x' - f x) < e))" - unfolding continuous_on_def apply (simp del: dist_commute) unfolding dist_vec1 unfolding dist_norm .. - -lemma continuous_at_vec1_norm: - fixes x :: "real ^ _" - shows "continuous (at x) (vec1 o norm)" - unfolding continuous_at_vec1_range using real_abs_sub_norm order_le_less_trans by blast - -lemma continuous_on_vec1_norm: - fixes s :: "(real ^ _) set" - shows "continuous_on s (vec1 o norm)" -unfolding continuous_on_vec1_range norm_vec1[THEN sym] by (metis norm_vec1 order_le_less_trans real_abs_sub_norm) - -lemma continuous_at_vec1_component: - shows "continuous (at (a::real^'a::finite)) (\ x. vec1(x$i))" -proof- - { fix e::real and x assume "0 < dist x a" "dist x a < e" "e>0" - hence "\x $ i - a $ i\ < e" using component_le_norm[of "x - a" i] unfolding dist_norm by auto } - thus ?thesis unfolding continuous_at tendsto_iff eventually_at dist_vec1 by auto -qed - -lemma continuous_on_vec1_component: - shows "continuous_on s (\ x::real^'a::finite. vec1(x$i))" -proof- - { fix e::real and x xa assume "x\s" "e>0" "xa\s" "0 < norm (xa - x) \ norm (xa - x) < e" - hence "\xa $ i - x $ i\ < e" using component_le_norm[of "xa - x" i] by auto } - thus ?thesis unfolding continuous_on Lim_within dist_vec1 unfolding dist_norm by auto -qed - -lemma continuous_at_vec1_infnorm: - "continuous (at x) (vec1 o infnorm)" - unfolding continuous_at Lim_at o_def unfolding dist_vec1 unfolding dist_norm + shows "continuous_on s f \ (\x \ s. \e>0. \d>0. (\x' \ s. norm(x' - x) < d --> abs(f x' - f x) < e))" + unfolding continuous_on_def dist_norm by simp + +lemma continuous_at_norm: "continuous (at x) norm" + unfolding continuous_at by (intro tendsto_intros) + +lemma continuous_on_norm: "continuous_on s norm" +unfolding continuous_on by (intro ballI tendsto_intros) + +lemma continuous_at_component: "continuous (at a) (\x. x $ i)" +unfolding continuous_at by (intro tendsto_intros) + +lemma continuous_on_component: "continuous_on s (\x. x $ i)" +unfolding continuous_on by (intro ballI tendsto_intros) + +lemma continuous_at_infnorm: "continuous (at x) infnorm" + unfolding continuous_at Lim_at o_def unfolding dist_norm apply auto apply (rule_tac x=e in exI) apply auto using order_trans[OF real_abs_sub_infnorm infnorm_le_norm, of _ x] by (metis xt1(7)) @@ -4119,23 +4138,23 @@ lemma compact_attains_sup: fixes s :: "real set" - assumes "compact (vec1 ` s)" "s \ {}" + assumes "compact s" "s \ {}" shows "\x \ s. \y \ s. y \ x" proof- - from assms(1) have a:"bounded (vec1 ` s)" "closed (vec1 ` s)" unfolding compact_eq_bounded_closed by auto + from assms(1) have a:"bounded s" "closed s" unfolding compact_eq_bounded_closed by auto { fix e::real assume as: "\x\s. x \ rsup s" "rsup s \ s" "0 < e" "\x'\s. x' = rsup s \ \ rsup s - x' < e" have "isLub UNIV s (rsup s)" using rsup[OF assms(2)] unfolding setle_def using as(1) by auto moreover have "isUb UNIV s (rsup s - e)" unfolding isUb_def unfolding setle_def using as(4,2) by auto ultimately have False using isLub_le_isUb[of UNIV s "rsup s" "rsup s - e"] using `e>0` by auto } - thus ?thesis using bounded_has_rsup(1)[OF a(1) assms(2)] using a(2)[unfolded closed_vec1, THEN spec[where x="rsup s"]] + thus ?thesis using bounded_has_rsup(1)[OF a(1) assms(2)] using a(2)[unfolded closed_real, THEN spec[where x="rsup s"]] apply(rule_tac x="rsup s" in bexI) by auto qed lemma compact_attains_inf: fixes s :: "real set" - assumes "compact (vec1 ` s)" "s \ {}" shows "\x \ s. \y \ s. x \ y" + assumes "compact s" "s \ {}" shows "\x \ s. \y \ s. x \ y" proof- - from assms(1) have a:"bounded (vec1 ` s)" "closed (vec1 ` s)" unfolding compact_eq_bounded_closed by auto + from assms(1) have a:"bounded s" "closed s" unfolding compact_eq_bounded_closed by auto { fix e::real assume as: "\x\s. x \ rinf s" "rinf s \ s" "0 < e" "\x'\s. x' = rinf s \ \ abs (x' - rinf s) < e" have "isGlb UNIV s (rinf s)" using rinf[OF assms(2)] unfolding setge_def using as(1) by auto @@ -4145,43 +4164,40 @@ have "rinf s + e \ x" using as(4)[THEN bspec[where x=x]] using as(2) `x\s` unfolding * by auto } hence "isLb UNIV s (rinf s + e)" unfolding isLb_def and setge_def by auto ultimately have False using isGlb_le_isLb[of UNIV s "rinf s" "rinf s + e"] using `e>0` by auto } - thus ?thesis using bounded_has_rinf(1)[OF a(1) assms(2)] using a(2)[unfolded closed_vec1, THEN spec[where x="rinf s"]] + thus ?thesis using bounded_has_rinf(1)[OF a(1) assms(2)] using a(2)[unfolded closed_real, THEN spec[where x="rinf s"]] apply(rule_tac x="rinf s" in bexI) by auto qed lemma continuous_attains_sup: fixes f :: "'a::metric_space \ real" - shows "compact s \ s \ {} \ continuous_on s (vec1 o f) + shows "compact s \ s \ {} \ continuous_on s f ==> (\x \ s. \y \ s. f y \ f x)" using compact_attains_sup[of "f ` s"] - using compact_continuous_image[of s "vec1 \ f"] unfolding image_compose by auto + using compact_continuous_image[of s f] by auto lemma continuous_attains_inf: fixes f :: "'a::metric_space \ real" - shows "compact s \ s \ {} \ continuous_on s (vec1 o f) + shows "compact s \ s \ {} \ continuous_on s f \ (\x \ s. \y \ s. f x \ f y)" using compact_attains_inf[of "f ` s"] - using compact_continuous_image[of s "vec1 \ f"] unfolding image_compose by auto + using compact_continuous_image[of s f] by auto lemma distance_attains_sup: - fixes s :: "(real ^ _) set" assumes "compact s" "s \ {}" shows "\x \ s. \y \ s. dist a y \ dist a x" -proof- - { fix x assume "x\s" fix e::real assume "e>0" - { fix x' assume "x'\s" and as:"norm (x' - x) < e" - hence "\norm (x' - a) - norm (x - a)\ < e" - using real_abs_sub_norm[of "x' - a" "x - a"] by auto } - hence "\d>0. \x'\s. norm (x' - x) < d \ \dist x' a - dist x a\ < e" using `e>0` unfolding dist_norm by auto } - thus ?thesis using assms - using continuous_attains_sup[of s "\x. dist a x"] - unfolding continuous_on_vec1_range by (auto simp add: dist_commute) +proof (rule continuous_attains_sup [OF assms]) + { fix x assume "x\s" + have "(dist a ---> dist a x) (at x within s)" + by (intro tendsto_dist tendsto_const Lim_at_within Lim_ident_at) + } + thus "continuous_on s (dist a)" + unfolding continuous_on .. qed text{* For *minimal* distance, we only need closure, not compactness. *} lemma distance_attains_inf: - fixes a :: "real ^ _" (* FIXME: generalize *) + fixes a :: "'a::heine_borel" assumes "closed s" "s \ {}" shows "\x \ s. \y \ s. dist a x \ dist a y" proof- @@ -4192,14 +4208,25 @@ moreover { fix x assume "x\?B" fix e::real assume "e>0" - { fix x' assume "x'\?B" and as:"norm (x' - x) < e" - hence "\norm (x' - a) - norm (x - a)\ < e" - using real_abs_sub_norm[of "x' - a" "x - a"] by auto } - hence "\d>0. \x'\?B. norm (x' - x) < d \ \dist x' a - dist x a\ < e" using `e>0` unfolding dist_norm by auto } - hence "continuous_on (cball a (dist b a) \ s) (vec1 \ dist a)" unfolding continuous_on_vec1_range - by (auto simp add: dist_commute) - moreover have "compact ?B" using compact_cball[of a "dist b a"] unfolding compact_eq_bounded_closed using bounded_Int and closed_Int and assms(1) by auto - ultimately obtain x where "x\cball a (dist b a) \ s" "\y\cball a (dist b a) \ s. dist a x \ dist a y" using continuous_attains_inf[of ?B "dist a"] by fastsimp + { fix x' assume "x'\?B" and as:"dist x' x < e" + from as have "\dist a x' - dist a x\ < e" + unfolding abs_less_iff minus_diff_eq + using dist_triangle2 [of a x' x] + using dist_triangle [of a x x'] + by arith + } + hence "\d>0. \x'\?B. dist x' x < d \ \dist a x' - dist a x\ < e" + using `e>0` by auto + } + hence "continuous_on (cball a (dist b a) \ s) (dist a)" + unfolding continuous_on Lim_within dist_norm real_norm_def + by fast + moreover have "compact ?B" + using compact_cball[of a "dist b a"] + unfolding compact_eq_bounded_closed + using bounded_Int and closed_Int and assms(1) by auto + ultimately obtain x where "x\cball a (dist b a) \ s" "\y\cball a (dist b a) \ s. dist a x \ dist a y" + using continuous_attains_inf[of ?B "dist a"] by fastsimp thus ?thesis by fastsimp qed @@ -4207,99 +4234,64 @@ lemma Lim_mul: fixes f :: "'a \ real ^ _" - assumes "((vec1 o c) ---> vec1 d) net" "(f ---> l) net" - shows "((\x. c(x) *s f x) ---> (d *s l)) net" -proof- - have "bilinear (\x. op *s (dest_vec1 (x::real^1)))" unfolding bilinear_def linear_def - unfolding dest_vec1_add dest_vec1_cmul - apply vector apply auto unfolding semiring_class.right_distrib semiring_class.left_distrib by auto - thus ?thesis using Lim_bilinear[OF assms, of "\x y. (dest_vec1 x) *s y"] by auto -qed + assumes "(c ---> d) net" "(f ---> 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 "((vec1 o c) ---> vec1 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 (vec1 o 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" - shows "continuous net (vec1 o c) \ continuous net f - ==> continuous net (\x. c(x) *s f x) " - unfolding continuous_def using Lim_mul[of c] by auto + fixes f :: "'a::metric_space \ 'b::real_normed_vector" + shows "continuous net c \ continuous net f + ==> 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 (vec1 o 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" - shows "continuous_on s (vec1 o c) \ continuous_on s f - ==> continuous_on s (\x. c(x) *s f x)" + fixes f :: "'a::metric_space \ 'b::real_normed_vector" + shows "continuous_on s c \ continuous_on s f + ==> 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. *} lemma Lim_inv: fixes f :: "'a \ real" - assumes "((vec1 o f) ---> vec1 l) (net::'a net)" "l \ 0" - shows "((vec1 o inverse o f) ---> vec1(inverse l)) net" -proof - - { fix e::real assume "e>0" - let ?d = "min (\l\ / 2) (l\ * e / 2)" - have "0 < ?d" using `l\0` `e>0` mult_pos_pos[of "l^2" "e/2"] by auto - with assms(1) have "eventually (\x. dist ((vec1 o f) x) (vec1 l) < ?d) net" - by (rule tendstoD) - moreover - { fix x assume "dist ((vec1 o f) x) (vec1 l) < ?d" - hence *:"\f x - l\ < min (\l\ / 2) (l\ * e / 2)" unfolding o_def dist_vec1 by auto - hence fx0:"f x \ 0" using `l \ 0` by auto - hence fxl0: "(f x) * l \ 0" using `l \ 0` by auto - from * have **:"\f x - l\ < l\ * e / 2" by auto - have "\f x\ * 2 \ \l\" using * by (auto simp del: less_divide_eq_number_of1) - hence "\f x\ * 2 * \l\ \ \l\ * \l\" unfolding mult_le_cancel_right by auto - hence "\f x * l\ * 2 \ \l\^2" unfolding real_mult_commute and power2_eq_square by auto - hence ***:"inverse \f x * l\ \ inverse (l\ / 2)" using fxl0 - using le_imp_inverse_le[of "l^2 / 2" "\f x * l\"] by auto - - have "dist ((vec1 \ inverse \ f) x) (vec1 (inverse l)) < e" unfolding o_def unfolding dist_vec1 - unfolding inverse_diff_inverse[OF fx0 `l\0`] apply simp - unfolding mult_commute[of "inverse (f x)"] - unfolding real_divide_def[THEN sym] - unfolding divide_divide_eq_left - unfolding nonzero_abs_divide[OF fxl0] - using mult_less_le_imp_less[OF **, of "inverse \f x * l\", of "inverse (l^2 / 2)"] using *** using fx0 `l\0` - unfolding inverse_eq_divide using `e>0` by auto - } - ultimately - have "eventually (\x. dist ((vec1 o inverse o f) x) (vec1(inverse l)) < e) net" - by (auto elim: eventually_rev_mono) - } - thus ?thesis unfolding tendsto_iff by auto -qed + assumes "(f ---> l) (net::'a net)" "l \ 0" + shows "((inverse o f) ---> inverse l) net" + unfolding o_def using assms by (rule tendsto_inverse) lemma continuous_inv: fixes f :: "'a::metric_space \ real" - shows "continuous net (vec1 o f) \ f(netlimit net) \ 0 - ==> continuous net (vec1 o inverse o f)" + shows "continuous net f \ f(netlimit net) \ 0 + ==> continuous net (inverse o f)" unfolding continuous_def using Lim_inv by auto lemma continuous_at_within_inv: - fixes f :: "real ^ _ \ real" - assumes "continuous (at a within s) (vec1 o f)" "f a \ 0" - shows "continuous (at a within s) (vec1 o inverse o f)" - using assms unfolding continuous_within o_apply - by (rule Lim_inv) + fixes f :: "'a::metric_space \ 'b::real_normed_field" + assumes "continuous (at a within s) f" "f a \ 0" + shows "continuous (at a within s) (inverse o f)" + using assms unfolding continuous_within o_def + by (intro tendsto_intros) lemma continuous_at_inv: - fixes f :: "real ^ _ \ real" - shows "continuous (at a) (vec1 o f) \ f a \ 0 - ==> continuous (at a) (vec1 o inverse o f) " + fixes f :: "'a::metric_space \ 'b::real_normed_field" + shows "continuous (at a) f \ f a \ 0 + ==> continuous (at a) (inverse o f) " using within_UNIV[THEN sym, of "at a"] using continuous_at_within_inv[of a UNIV] by auto subsection{* Preservation properties for pasted sets. *} @@ -4316,6 +4308,16 @@ thus ?thesis unfolding bounded_iff by auto qed +lemma bounded_Times: + assumes "bounded s" "bounded t" shows "bounded (s \ t)" +proof- + obtain x y a b where "\z\s. dist x z \ a" "\z\t. dist y z \ b" + using assms [unfolded bounded_def] by auto + then have "\z\s \ t. dist (x, y) z \ sqrt (a\ + b\)" + by (auto simp add: dist_Pair_Pair real_sqrt_le_mono add_mono power_mono) + thus ?thesis unfolding bounded_any_center [where a="(x, y)"] by auto +qed + lemma closed_pastecart: fixes s :: "(real ^ 'a::finite) set" (* FIXME: generalize *) assumes "closed s" "closed t" @@ -4343,6 +4345,26 @@ shows "compact s \ compact t ==> compact {pastecart x y | x y . x \ s \ y \ t}" unfolding compact_eq_bounded_closed using bounded_pastecart[of s t] closed_pastecart[of s t] by auto +lemma mem_Times_iff: "x \ A \ B \ fst x \ A \ snd x \ B" +by (induct x) simp + +lemma compact_Times: "compact s \ compact t \ compact (s \ t)" +unfolding compact_def +apply clarify +apply (drule_tac x="fst \ f" in spec) +apply (drule mp, simp add: mem_Times_iff) +apply (clarify, rename_tac l1 r1) +apply (drule_tac x="snd \ f \ r1" in spec) +apply (drule mp, simp add: mem_Times_iff) +apply (clarify, rename_tac l2 r2) +apply (rule_tac x="(l1, l2)" in rev_bexI, simp) +apply (rule_tac x="r1 \ r2" in exI) +apply (rule conjI, simp add: subseq_def) +apply (drule_tac r=r2 in lim_subseq [COMP swap_prems_rl], assumption) +apply (drule (1) tendsto_Pair) back +apply (simp add: o_def) +done + text{* Hence some useful properties follow quite easily. *} lemma compact_scaleR_image: @@ -4357,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: @@ -4366,30 +4388,27 @@ using compact_scaleR_image [OF assms, of "- 1"] by auto lemma compact_sums: - fixes s t :: "(real ^ _) set" + fixes s t :: "'a::real_normed_vector set" assumes "compact s" "compact t" shows "compact {x + y | x y. x \ s \ y \ t}" proof- - have *:"{x + y | x y. x \ s \ y \ t} =(\z. fstcart z + sndcart z) ` {pastecart x y | x y. x \ s \ y \ t}" - apply auto unfolding image_iff apply(rule_tac x="pastecart xa y" in bexI) unfolding fstcart_pastecart sndcart_pastecart by auto - have "linear (\z::real^('a + 'a). fstcart z + sndcart z)" unfolding linear_def - unfolding fstcart_add sndcart_add apply auto - unfolding vector_add_ldistrib fstcart_cmul[THEN sym] sndcart_cmul[THEN sym] by auto - hence "continuous_on {pastecart x y |x y. x \ s \ y \ t} (\z. fstcart z + sndcart z)" - using continuous_at_imp_continuous_on linear_continuous_at by auto - thus ?thesis unfolding * using compact_continuous_image compact_pastecart[OF assms] by auto + have *:"{x + y | x y. x \ s \ y \ t} = (\z. fst z + snd z) ` (s \ t)" + apply auto unfolding image_iff apply(rule_tac x="(xa, y)" in bexI) by auto + have "continuous_on (s \ t) (\z. fst z + snd z)" + unfolding continuous_on by (rule ballI) (intro tendsto_intros) + thus ?thesis unfolding * using compact_continuous_image compact_Times [OF assms] by auto qed lemma compact_differences: - fixes s t :: "(real ^ 'a::finite) set" + fixes s t :: "'a::real_normed_vector set" assumes "compact s" "compact t" shows "compact {x - y | x y. x \ s \ y \ t}" proof- - have "{x - y | x y::real^'a. x\s \ y \ t} = {x + y | x y. x \ s \ y \ (uminus ` t)}" + have "{x - y | x y. x\s \ y \ t} = {x + y | x y. x \ s \ y \ (uminus ` t)}" apply auto apply(rule_tac x= xa in exI) apply auto apply(rule_tac x=xa in exI) by auto thus ?thesis using compact_sums[OF assms(1) compact_negations[OF assms(2)]] by auto qed lemma compact_translation: - fixes s :: "(real ^ _) set" + fixes s :: "'a::real_normed_vector set" assumes "compact s" shows "compact ((\x. a + x) ` s)" proof- have "{x + y |x y. x \ s \ y \ {a}} = (\x. a + x) ` s" by auto @@ -4398,23 +4417,23 @@ 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 text{* Hence we get the following. *} lemma compact_sup_maxdistance: - fixes s :: "(real ^ _) set" + fixes s :: "'a::real_normed_vector set" assumes "compact s" "s \ {}" shows "\x\s. \y\s. \u\s. \v\s. norm(u - v) \ norm(x - y)" proof- have "{x - y | x y . x\s \ y\s} \ {}" using `s \ {}` by auto then obtain x where x:"x\{x - y |x y. x \ s \ y \ s}" "\y\{x - y |x y. x \ s \ y \ s}. norm y \ norm x" using compact_differences[OF assms(1) assms(1)] - using distance_attains_sup[unfolded dist_norm, of "{x - y | x y . x\s \ y\s}" 0] by(auto simp add: norm_minus_cancel) + using distance_attains_sup[where 'a="'a", unfolded dist_norm, of "{x - y | x y . x\s \ y\s}" 0] by(auto simp add: norm_minus_cancel) from x(1) obtain a b where "a\s" "b\s" "x = a - b" by auto thus ?thesis using x(2)[unfolded `x = a - b`] by blast qed @@ -4458,11 +4477,11 @@ using diameter_bounded by blast lemma diameter_compact_attained: - fixes s :: "(real ^ _) set" + fixes s :: "'a::real_normed_vector set" assumes "compact s" "s \ {}" shows "\x\s. \y\s. (norm(x - y) = diameter s)" proof- - have b:"bounded s" using assms(1) compact_eq_bounded_closed by auto + have b:"bounded s" using assms(1) by (rule compact_imp_bounded) then obtain x y where xys:"x\s" "y\s" and xy:"\u\s. \v\s. norm (u - v) \ norm (x - y)" using compact_sup_maxdistance[OF assms] by auto hence "diameter s \ norm (x - y)" using rsup_le[of "{norm (x - y) |x y. x \ s \ y \ s}" "norm (x - y)"] unfolding setle_def and diameter_def by auto @@ -4495,7 +4514,7 @@ then obtain N where "\n\N. dist (x n) l < e * \c\" using as(2)[unfolded Lim_sequentially, THEN spec[where x="e * abs c"]] by auto hence "\N. \n\N. dist (scaleR (1 / c) (x n)) (scaleR (1 / c) l) < e" - unfolding dist_norm unfolding scaleR_right_diff_distrib[THEN sym] norm_scaleR + unfolding dist_norm unfolding scaleR_right_diff_distrib[THEN sym] using mult_imp_div_pos_less[of "abs c" _ e] `c\0` by auto } hence "((\n. scaleR (1 / c) (x n)) ---> scaleR (1 / c) l) sequentially" unfolding Lim_sequentially by auto ultimately have "l \ scaleR c ` s" @@ -4507,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: @@ -4523,10 +4542,10 @@ { fix x l assume as:"\n. x n \ ?S" "(x ---> l) sequentially" from as(1) obtain f where f:"\n. x n = fst (f n) + snd (f n)" "\n. fst (f n) \ s" "\n. snd (f n) \ t" using choice[of "\n y. x n = (fst y) + (snd y) \ fst y \ s \ snd y \ t"] by auto - obtain l' r where "l'\s" and r:"\m n. m < n \ r m < r n" and lr:"(((\n. fst (f n)) \ r) ---> l') sequentially" + obtain l' r where "l'\s" and r:"subseq r" and lr:"(((\n. fst (f n)) \ r) ---> l') sequentially" using assms(1)[unfolded compact_def, THEN spec[where x="\ n. fst (f n)"]] using f(2) by auto have "((\n. snd (f (r n))) ---> l - l') sequentially" - using Lim_sub[OF lim_subsequence[OF r as(2)] lr] and f(1) unfolding o_def by auto + using Lim_sub[OF lim_subseq[OF r as(2)] lr] and f(1) unfolding o_def by auto hence "l - l' \ t" using assms(2)[unfolded closed_sequential_limits, THEN spec[where x="\ n. snd (f (r n))"], THEN spec[where x="l - l'"]] using f(3) by auto @@ -4600,7 +4619,7 @@ subsection{* Separation between points and sets. *} lemma separate_point_closed: - fixes s :: "(real ^ _) set" (* FIXME: generalize *) + fixes s :: "'a::heine_borel set" shows "closed s \ a \ s ==> (\d>0. \x\s. d \ dist a x)" proof(cases "s = {}") case True @@ -4613,7 +4632,8 @@ qed lemma separate_compact_closed: - fixes s t :: "(real ^ _) set" + fixes s t :: "'a::{heine_borel, real_normed_vector} set" + (* TODO: does this generalize to heine_borel? *) assumes "compact s" and "closed t" and "s \ t = {}" shows "\d>0. \x\s. \y\t. d \ dist x y" proof- @@ -4629,7 +4649,7 @@ qed lemma separate_closed_compact: - fixes s t :: "(real ^ _) set" + fixes s t :: "'a::{heine_borel, real_normed_vector} set" assumes "closed s" and "compact t" and "s \ t = {}" shows "\d>0. \x\s. \y\t. d \ dist x y" proof- @@ -4666,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 } @@ -4681,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 } @@ -4865,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 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 @@ -4910,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) } @@ -4922,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) @@ -4939,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 @@ -4949,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 @@ -4965,11 +4986,11 @@ then obtain N::nat where "inverse (real (N + 1)) < e" by auto hence "\n\N. inverse (real n + 1) < e" by (auto, metis Suc_le_mono le_SucE less_imp_inverse_less nat_le_real_less order_less_trans real_of_nat_Suc real_of_nat_Suc_gt_zero) hence "\N::nat. \n\N. inverse (real n + 1) < e" by auto } - hence "((vec1 \ (\n. inverse (real n + 1))) ---> vec1 0) sequentially" - unfolding Lim_sequentially by(auto simp add: dist_vec1) + 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<.. real^'n::finite" - assumes "(f ---> l) net" shows "((vec1 o (\y. a \ (f y))) ---> vec1(a \ l)) net" -proof(cases "a = vec 0") - case True thus ?thesis using dot_lzero and Lim_const[of 0 net] unfolding vec1_vec and o_def by auto -next - case False - { fix e::real - assume "0 < e" - with `a \ vec 0` have "0 < e / norm a" by (simp add: divide_pos_pos) - with assms(1) have "eventually (\x. dist (f x) l < e / norm a) net" - by (rule tendstoD) - moreover - { fix z assume "dist (f z) l < e / norm a" - hence "norm a * norm (f z - l) < e" unfolding dist_norm and - pos_less_divide_eq[OF False[unfolded vec_0 zero_less_norm_iff[of a, THEN sym]]] and real_mult_commute by auto - hence "\a \ f z - a \ l\ < e" - using order_le_less_trans[OF norm_cauchy_schwarz_abs[of a "f z - l"], of e] - unfolding dot_rsub[symmetric] by auto } - ultimately have "eventually (\x. \a \ f x - a \ l\ < e) net" - by (auto elim: eventually_rev_mono) - } - thus ?thesis unfolding tendsto_iff - by (auto simp add: dist_vec1) -qed - -lemma continuous_at_vec1_dot: +lemma Lim_dot: fixes f :: "real^'m \ real^'n::finite" + assumes "(f ---> l) net" shows "((\y. a \ (f y)) ---> a \ l) net" + unfolding dot_def by (intro tendsto_intros assms) + +lemma continuous_at_dot: fixes x :: "real ^ _" - shows "continuous (at x) (vec1 o (\y. a \ y))" + shows "continuous (at x) (\y. a \ y)" proof- have "((\x. x) ---> x) (at x)" unfolding Lim_at by auto - thus ?thesis unfolding continuous_at and o_def using Lim_vec1_dot[of "\x. x" x "at x" a] by auto -qed - -lemma continuous_on_vec1_dot: + thus ?thesis unfolding continuous_at and o_def using Lim_dot[of "\x. x" x "at x" a] by auto +qed + +lemma continuous_on_dot: fixes s :: "(real ^ _) set" - shows "continuous_on s (vec1 o (\y. a \ y)) " - using continuous_at_imp_continuous_on[of s "vec1 o (\y. a \ y)"] - using continuous_at_vec1_dot + shows "continuous_on s (\y. a \ y)" + using continuous_at_imp_continuous_on[of s "\y. a \ y"] + using continuous_at_dot by auto -lemma closed_halfspace_le: fixes a::"real^'n::finite" - shows "closed {x. a \ x \ b}" +lemma continuous_on_inner: + fixes s :: "'a::real_inner set" + shows "continuous_on s (inner a)" + unfolding continuous_on by (rule ballI) (intro tendsto_intros) + +lemma closed_halfspace_le: "closed {x. inner a x \ b}" proof- - have *:"{x \ UNIV. (vec1 \ op \ a) x \ vec1 ` {r. \x. a \ x = r \ r \ b}} = {x. a \ x \ b}" by auto - let ?T = "{x::real^1. (\i. x$i \ (vec1 b)$i)}" - have "closed ?T" using closed_interval_left[of "vec1 b"] by simp - moreover have "vec1 ` {r. \x. a \ x = r \ r \ b} = range (vec1 \ op \ a) \ ?T" unfolding all_1 + have *:"{x \ UNIV. inner a x \ {r. \x. inner a x = r \ r \ b}} = {x. inner a x \ b}" by auto + let ?T = "{..b}" + have "closed ?T" by (rule closed_real_atMost) + moreover have "{r. \x. inner a x = r \ r \ b} = range (inner a) \ ?T" unfolding image_def by auto - ultimately have "\T. closed T \ vec1 ` {r. \x. a \ x = r \ r \ b} = range (vec1 \ op \ a) \ T" by auto - hence "closedin euclidean {x \ UNIV. (vec1 \ op \ a) x \ vec1 ` {r. \x. a \ x = r \ r \ b}}" - using continuous_on_vec1_dot[of UNIV a, unfolded continuous_on_closed subtopology_UNIV] unfolding closedin_closed - by (auto elim!: allE[where x="vec1 ` {r. (\x. a \ x = r \ r \ b)}"]) + ultimately have "\T. closed T \ {r. \x. inner a x = r \ r \ b} = range (inner a) \ T" by fast + hence "closedin euclidean {x \ UNIV. inner a x \ {r. \x. inner a x = r \ r \ b}}" + using continuous_on_inner[of UNIV a, unfolded continuous_on_closed subtopology_UNIV] unfolding closedin_closed + by (fast elim!: allE[where x="{r. (\x. inner a x = r \ r \ b)}"]) thus ?thesis unfolding closed_closedin[THEN sym] and * by auto qed -lemma closed_halfspace_ge: "closed {x::real^_. a \ x \ b}" - using closed_halfspace_le[of "-a" "-b"] unfolding dot_lneg by auto - -lemma closed_hyperplane: "closed {x::real^_. a \ x = b}" +lemma closed_halfspace_ge: "closed {x. inner a x \ b}" + using closed_halfspace_le[of "-a" "-b"] unfolding inner_minus_left by auto + +lemma closed_hyperplane: "closed {x. inner a x = b}" proof- - have "{x. a \ x = b} = {x. a \ x \ b} \ {x. a \ x \ b}" by auto + have "{x. inner a x = b} = {x. inner a x \ b} \ {x. inner a x \ b}" by auto thus ?thesis using closed_halfspace_le[of a b] and closed_halfspace_ge[of b a] using closed_Int by auto qed lemma closed_halfspace_component_le: shows "closed {x::real^'n::finite. x$i \ a}" - using closed_halfspace_le[of "(basis i)::real^'n" a] unfolding dot_basis[OF assms] by auto + using closed_halfspace_le[of "(basis i)::real^'n" a] unfolding inner_basis[OF assms] by auto lemma closed_halfspace_component_ge: shows "closed {x::real^'n::finite. x$i \ a}" - using closed_halfspace_ge[of a "(basis i)::real^'n"] unfolding dot_basis[OF assms] by auto + using closed_halfspace_ge[of a "(basis i)::real^'n"] unfolding inner_basis[OF assms] by auto text{* Openness of halfspaces. *} -lemma open_halfspace_lt: "open {x::real^_. a \ x < b}" +lemma open_halfspace_lt: "open {x. inner a x < b}" proof- - have "UNIV - {x. b \ a \ x} = {x. a \ x < b}" by auto + have "UNIV - {x. b \ inner a x} = {x. inner a x < b}" by auto thus ?thesis using closed_halfspace_ge[unfolded closed_def Compl_eq_Diff_UNIV, of b a] by auto qed -lemma open_halfspace_gt: "open {x::real^_. a \ x > b}" +lemma open_halfspace_gt: "open {x. inner a x > b}" proof- - have "UNIV - {x. b \ a \ x} = {x. a \ x > b}" by auto + have "UNIV - {x. b \ inner a x} = {x. inner a x > b}" by auto thus ?thesis using closed_halfspace_le[unfolded closed_def Compl_eq_Diff_UNIV, of a b] by auto qed lemma open_halfspace_component_lt: shows "open {x::real^'n::finite. x$i < a}" - using open_halfspace_lt[of "(basis i)::real^'n" a] unfolding dot_basis[OF assms] by auto + using open_halfspace_lt[of "(basis i)::real^'n" a] unfolding inner_basis[OF assms] by auto lemma open_halfspace_component_gt: shows "open {x::real^'n::finite. x$i > a}" - using open_halfspace_gt[of a "(basis i)::real^'n"] unfolding dot_basis[OF assms] by auto + using open_halfspace_gt[of a "(basis i)::real^'n"] unfolding inner_basis[OF assms] by auto text{* This gives a simple derivation of limit component bounds. *} @@ -5228,8 +5232,8 @@ assumes "(f ---> l) net" "\ (trivial_limit net)" "eventually (\x. f(x)$i \ b) net" shows "l$i \ b" proof- - { fix x have "x \ {x::real^'n. basis i \ x \ b} \ x$i \ b" unfolding dot_basis by auto } note * = this - show ?thesis using Lim_in_closed_set[of "{x. basis i \ x \ b}" f net l] unfolding * + { fix x have "x \ {x::real^'n. inner (basis i) x \ b} \ x$i \ b" unfolding inner_basis by auto } note * = this + show ?thesis using Lim_in_closed_set[of "{x. inner (basis i) x \ b}" f net l] unfolding * using closed_halfspace_le[of "(basis i)::real^'n" b] and assms(1,2,3) by auto qed @@ -5237,8 +5241,8 @@ assumes "(f ---> l) net" "\ (trivial_limit net)" "eventually (\x. b \ (f x)$i) net" shows "b \ l$i" proof- - { fix x have "x \ {x::real^'n. basis i \ x \ b} \ x$i \ b" unfolding dot_basis by auto } note * = this - show ?thesis using Lim_in_closed_set[of "{x. basis i \ x \ b}" f net l] unfolding * + { fix x have "x \ {x::real^'n. inner (basis i) x \ b} \ x$i \ b" unfolding inner_basis by auto } note * = this + show ?thesis using Lim_in_closed_set[of "{x. inner (basis i) x \ b}" f net l] unfolding * using closed_halfspace_ge[of b "(basis i)::real^'n"] and assms(1,2,3) by auto qed @@ -5292,13 +5296,13 @@ text{* Some more convenient intermediate-value theorem formulations. *} -lemma connected_ivt_hyperplane: fixes y :: "real^'n::finite" - assumes "connected s" "x \ s" "y \ s" "a \ x \ b" "b \ a \ y" - shows "\z \ s. a \ z = b" +lemma connected_ivt_hyperplane: + assumes "connected s" "x \ s" "y \ s" "inner a x \ b" "b \ inner a y" + shows "\z \ s. inner a z = b" proof(rule ccontr) - assume as:"\ (\z\s. a \ z = b)" - let ?A = "{x::real^'n. a \ x < b}" - let ?B = "{x::real^'n. a \ x > b}" + assume as:"\ (\z\s. inner a z = b)" + let ?A = "{x. inner a x < b}" + let ?B = "{x. inner a x > b}" have "open ?A" "open ?B" using open_halfspace_lt and open_halfspace_gt by auto moreover have "?A \ ?B = {}" by auto moreover have "s \ ?A \ ?B" using as by auto @@ -5307,7 +5311,7 @@ lemma connected_ivt_component: fixes x::"real^'n::finite" shows "connected s \ x \ s \ y \ s \ x$k \ a \ a \ y$k \ (\z\s. z$k = a)" - using connected_ivt_hyperplane[of s x y "(basis k)::real^'n" a] by (auto simp add: dot_basis) + using connected_ivt_hyperplane[of s x y "(basis k)::real^'n" a] by (auto simp add: inner_basis) text{* Also more convenient formulations of monotone convergence. *} @@ -5320,7 +5324,7 @@ have "\ n. n\m \ dest_vec1 (s m) \ dest_vec1 (s n)" apply(induct_tac n) apply simp using assms(2) apply(erule_tac x="na" in allE) by(auto simp add: not_less_eq_eq) } hence "\m n. m \ n \ dest_vec1 (s m) \ dest_vec1 (s n)" by auto - then obtain l where "\e>0. \N. \n\N. \dest_vec1 (s n) - l\ < e" using convergent_bounded_monotone[OF a] by auto + then obtain l where "\e>0. \N. \n\N. \dest_vec1 (s n) - l\ < e" using convergent_bounded_monotone[OF a] unfolding monoseq_def by auto thus ?thesis unfolding Lim_sequentially apply(rule_tac x="vec1 l" in exI) unfolding dist_norm unfolding abs_dest_vec1 and dest_vec1_sub by auto qed @@ -5429,11 +5433,11 @@ text{* Results on translation, scaling etc. *} lemma homeomorphic_scaling: - assumes "c \ 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: @@ -5444,13 +5448,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" @@ -5460,27 +5464,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: 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: 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. *} @@ -5568,10 +5568,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 + 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 ba using `x\0` `a\0` by (auto simp add: real_mult_commute pos_le_divide_eq pos_divide_le_eq) qed } ultimately @@ -5598,16 +5598,16 @@ "closed {x::real^'n::finite. \i. P i --> x$i = 0}" (is "closed ?A") proof- let ?D = "{i. P i}" - let ?Bs = "{{x::real^'n. basis i \ x = 0}| i. i \ ?D}" + let ?Bs = "{{x::real^'n. inner (basis i) x = 0}| i. i \ ?D}" { fix x { assume "x\?A" hence x:"\i\?D. x $ i = 0" by auto - hence "x\ \ ?Bs" by(auto simp add: dot_basis x) } + hence "x\ \ ?Bs" by(auto simp add: inner_basis x) } moreover { assume x:"x\\?Bs" { fix i assume i:"i \ ?D" - then obtain B where BB:"B \ ?Bs" and B:"B = {x::real^'n. basis i \ x = 0}" by auto - hence "x $ i = 0" unfolding B using x unfolding dot_basis by auto } + then obtain B where BB:"B \ ?Bs" and B:"B = {x::real^'n. inner (basis i) x = 0}" by auto + hence "x $ i = 0" unfolding B using x unfolding inner_basis by auto } hence "x\?A" by auto } ultimately have "x\?A \ x\ \?Bs" by auto } hence "?A = \ ?Bs" by auto @@ -5635,8 +5635,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 @@ -5646,9 +5646,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 @@ -5769,45 +5770,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...) *} @@ -5940,9 +5941,10 @@ subsection{* Edelstein fixed point theorem. *} lemma edelstein_fix: + fixes s :: "'a::real_normed_vector set" assumes s:"compact s" "s \ {}" and gs:"(g ` s) \ s" and dist:"\x\s. \y\s. x \ y \ dist (g x) (g y) < dist x y" - shows "\! x::real^'a::finite\s. g x = x" + shows "\! x\s. g x = x" proof(cases "\x\s. g x \ x") obtain x where "x\s" using s(2) by auto case False hence g:"\x\s. g x = x" by auto @@ -5985,26 +5987,23 @@ qed qed } note distf = this - def h \ "\n. pastecart (f n x) (f n y)" - let ?s2 = "{pastecart x y |x y. x \ s \ y \ s}" - obtain l r where "l\?s2" and r:"\m n. m < n \ r m < r n" and lr:"((h \ r) ---> l) sequentially" - using compact_pastecart[OF s(1) s(1), unfolded compact_def, THEN spec[where x=h]] unfolding h_def + def h \ "\n. (f n x, f n y)" + let ?s2 = "s \ s" + obtain l r where "l\?s2" and r:"subseq r" and lr:"((h \ r) ---> l) sequentially" + using compact_Times [OF s(1) s(1), unfolded compact_def, THEN spec[where x=h]] unfolding h_def using fs[OF `x\s`] and fs[OF `y\s`] by blast - def a \ "fstcart l" def b \ "sndcart l" - have lab:"l = pastecart a b" unfolding a_def b_def and pastecart_fst_snd by simp + def a \ "fst l" def b \ "snd l" + have lab:"l = (a, b)" unfolding a_def b_def by simp have [simp]:"a\s" "b\s" unfolding a_def b_def using `l\?s2` by auto - have "continuous_on (UNIV :: (real ^ _) set) fstcart" - and "continuous_on (UNIV :: (real ^ _) set) sndcart" - using linear_continuous_on using linear_fstcart and linear_sndcart by auto - hence lima:"((fstcart \ (h \ r)) ---> a) sequentially" and limb:"((sndcart \ (h \ r)) ---> b) sequentially" - unfolding atomize_conj unfolding continuous_on_sequentially - apply(erule_tac x="h \ r" in allE) apply(erule_tac x="h \ r" in allE) using lr - unfolding o_def and h_def a_def b_def by auto + have lima:"((fst \ (h \ r)) ---> a) sequentially" + and limb:"((snd \ (h \ r)) ---> b) sequentially" + using lr + unfolding o_def a_def b_def by (simp_all add: tendsto_intros) { fix n::nat - have *:"\fx fy (x::real^_) y. dist fx fy \ dist x y \ \ (dist (fx - fy) (a - b) < dist a b - dist x y)" unfolding dist_norm by norm - { fix x y ::"real^'a" + have *:"\fx fy (x::'a) y. dist fx fy \ dist x y \ \ (dist (fx - fy) (a - b) < dist a b - dist x y)" unfolding dist_norm by norm + { fix x y :: 'a have "dist (-x) (-y) = dist x y" unfolding dist_norm using norm_minus_cancel[of "x - y"] by (auto simp add: uminus_add_conv_diff) } note ** = this @@ -6021,7 +6020,7 @@ moreover have "dist (f (r (Na + Nb + n)) x - f (r (Na + Nb + n)) y) (a - b) \ dist a b - dist (f n x) (f n y)" using distf[of n "r (Na+Nb+n)", OF _ `x\s` `y\s`] - using monotone_bigger[OF r, of "Na+Nb+n"] + using subseq_bigger[OF r, of "Na+Nb+n"] using *[of "f (r (Na + Nb + n)) x" "f (r (Na + Nb + n)) y" "f n x" "f n y"] by auto ultimately have False by simp } @@ -6049,8 +6048,8 @@ have "dist y x < e \ dist (g y) (g x) < e" using dist by fastsimp } hence "continuous_on s g" unfolding continuous_on_def by auto - hence "((sndcart \ h \ r) ---> g a) sequentially" unfolding continuous_on_sequentially - apply (rule allE[where x="\n. (fstcart \ h \ r) n"]) apply (erule ballE[where x=a]) + hence "((snd \ h \ r) ---> g a) sequentially" unfolding continuous_on_sequentially + apply (rule allE[where x="\n. (fst \ h \ r) n"]) apply (erule ballE[where x=a]) using lima unfolding h_def o_def using fs[OF `x\s`] by (auto simp add: y_def) hence "g a = a" using Lim_unique[OF trivial_limit_sequentially limb, of "g a"] unfolding `a=b` and o_assoc by auto diff -r 78ac5c304db7 -r ef6d67b1ad10 src/HOL/Limits.thy --- a/src/HOL/Limits.thy Thu Jun 11 19:49:02 2009 +0200 +++ b/src/HOL/Limits.thy Sat Jun 13 16:29:15 2009 +0200 @@ -358,6 +358,14 @@ where [code del]: "(f ---> l) net \ (\S. open S \ l \ S \ eventually (\x. f x \ S) net)" +ML{* +structure TendstoIntros = + NamedThmsFun(val name = "tendsto_intros" + val description = "introduction rules for tendsto"); +*} + +setup TendstoIntros.setup + lemma topological_tendstoI: "(\S. open S \ l \ S \ eventually (\x. f x \ S) net) \ (f ---> l) net" @@ -395,12 +403,38 @@ lemma tendsto_Zfun_iff: "(f ---> a) net = Zfun (\x. f x - a) net" by (simp only: tendsto_iff Zfun_def dist_norm) -lemma tendsto_const: "((\x. k) ---> k) net" +lemma tendsto_ident_at [tendsto_intros]: "((\x. x) ---> a) (at a)" +unfolding tendsto_def eventually_at_topological by auto + +lemma tendsto_ident_at_within [tendsto_intros]: + "a \ S \ ((\x. x) ---> a) (at a within S)" +unfolding tendsto_def eventually_within eventually_at_topological by auto + +lemma tendsto_const [tendsto_intros]: "((\x. k) ---> k) net" by (simp add: tendsto_def) -lemma tendsto_norm: - fixes a :: "'a::real_normed_vector" - shows "(f ---> a) net \ ((\x. norm (f x)) ---> norm a) net" +lemma tendsto_dist [tendsto_intros]: + assumes f: "(f ---> l) net" and g: "(g ---> m) net" + shows "((\x. dist (f x) (g x)) ---> dist l m) net" +proof (rule tendstoI) + fix e :: real assume "0 < e" + hence e2: "0 < e/2" by simp + from tendstoD [OF f e2] tendstoD [OF g e2] + show "eventually (\x. dist (dist (f x) (g x)) (dist l m) < e) net" + proof (rule eventually_elim2) + fix x assume "dist (f x) l < e/2" "dist (g x) m < e/2" + then show "dist (dist (f x) (g x)) (dist l m) < e" + unfolding dist_real_def + using dist_triangle2 [of "f x" "g x" "l"] + using dist_triangle2 [of "g x" "l" "m"] + using dist_triangle3 [of "l" "m" "f x"] + using dist_triangle [of "f x" "m" "g x"] + by arith + qed +qed + +lemma tendsto_norm [tendsto_intros]: + "(f ---> a) net \ ((\x. norm (f x)) ---> norm a) net" apply (simp add: tendsto_iff dist_norm, safe) apply (drule_tac x="e" in spec, safe) apply (erule eventually_elim1) @@ -417,12 +451,12 @@ shows "(- a) - (- b) = - (a - b)" by simp -lemma tendsto_add: +lemma tendsto_add [tendsto_intros]: fixes a b :: "'a::real_normed_vector" shows "\(f ---> a) net; (g ---> b) net\ \ ((\x. f x + g x) ---> a + b) net" by (simp only: tendsto_Zfun_iff add_diff_add Zfun_add) -lemma tendsto_minus: +lemma tendsto_minus [tendsto_intros]: fixes a :: "'a::real_normed_vector" shows "(f ---> a) net \ ((\x. - f x) ---> - a) net" by (simp only: tendsto_Zfun_iff minus_diff_minus Zfun_minus) @@ -432,16 +466,34 @@ shows "((\x. - f x) ---> - a) net \ (f ---> a) net" by (drule tendsto_minus, simp) -lemma tendsto_diff: +lemma tendsto_diff [tendsto_intros]: fixes a b :: "'a::real_normed_vector" shows "\(f ---> a) net; (g ---> b) net\ \ ((\x. f x - g x) ---> a - b) net" by (simp add: diff_minus tendsto_add tendsto_minus) -lemma (in bounded_linear) tendsto: +lemma tendsto_setsum [tendsto_intros]: + fixes f :: "'a \ 'b \ 'c::real_normed_vector" + assumes "\i. i \ S \ (f i ---> a i) net" + shows "((\x. \i\S. f i x) ---> (\i\S. a i)) net" +proof (cases "finite S") + assume "finite S" thus ?thesis using assms + proof (induct set: finite) + case empty show ?case + by (simp add: tendsto_const) + next + case (insert i F) thus ?case + by (simp add: tendsto_add) + qed +next + assume "\ finite S" thus ?thesis + by (simp add: tendsto_const) +qed + +lemma (in bounded_linear) tendsto [tendsto_intros]: "(g ---> a) net \ ((\x. f (g x)) ---> f a) net" by (simp only: tendsto_Zfun_iff diff [symmetric] Zfun) -lemma (in bounded_bilinear) tendsto: +lemma (in bounded_bilinear) tendsto [tendsto_intros]: "\(f ---> a) net; (g ---> b) net\ \ ((\x. f x ** g x) ---> a ** b) net" by (simp only: tendsto_Zfun_iff prod_diff_prod Zfun_add Zfun Zfun_left Zfun_right) @@ -556,7 +608,7 @@ apply (simp add: tendsto_Zfun_iff) done -lemma tendsto_inverse: +lemma tendsto_inverse [tendsto_intros]: fixes a :: "'a::real_normed_div_algebra" assumes f: "(f ---> a) net" assumes a: "a \ 0" @@ -571,7 +623,7 @@ by (rule tendsto_inverse_lemma) qed -lemma tendsto_divide: +lemma tendsto_divide [tendsto_intros]: fixes a b :: "'a::real_normed_field" shows "\(f ---> a) net; (g ---> b) net; b \ 0\ \ ((\x. f x / g x) ---> a / b) net" diff -r 78ac5c304db7 -r ef6d67b1ad10 src/HOL/List.thy --- a/src/HOL/List.thy Thu Jun 11 19:49:02 2009 +0200 +++ b/src/HOL/List.thy Sat Jun 13 16:29:15 2009 +0200 @@ -2931,7 +2931,7 @@ done -subsubsection {* Infiniteness *} +subsubsection {* (In)finiteness *} lemma finite_maxlen: "finite (M::'a list set) ==> EX n. ALL s:M. size s < n" @@ -2944,6 +2944,27 @@ thus ?case .. qed +lemma finite_lists_length_eq: +assumes "finite A" +shows "finite {xs. set xs \ A \ length xs = n}" (is "finite (?S n)") +proof(induct n) + case 0 show ?case by simp +next + case (Suc n) + have "?S (Suc n) = (\x\A. (\xs. x#xs) ` ?S n)" + by (auto simp:length_Suc_conv) + then show ?case using `finite A` + by (auto intro: finite_imageI Suc) (* FIXME metis? *) +qed + +lemma finite_lists_length_le: + assumes "finite A" shows "finite {xs. set xs \ A \ length xs \ n}" + (is "finite ?S") +proof- + have "?S = (\n\{0..n}. {xs. set xs \ A \ length xs = n})" by auto + thus ?thesis by (auto intro: finite_lists_length_eq[OF `finite A`]) +qed + lemma infinite_UNIV_listI: "~ finite(UNIV::'a list set)" apply(rule notI) apply(drule finite_maxlen) diff -r 78ac5c304db7 -r ef6d67b1ad10 src/HOL/RealVector.thy --- a/src/HOL/RealVector.thy Thu Jun 11 19:49:02 2009 +0200 +++ b/src/HOL/RealVector.thy Sat Jun 13 16:29:15 2009 +0200 @@ -116,11 +116,11 @@ thus "a = b" by (simp only: right_minus_eq) qed -lemma scale_cancel_left: +lemma scale_cancel_left [simp]: "scale a x = scale a y \ x = y \ a = 0" by (auto intro: scale_left_imp_eq) -lemma scale_cancel_right: +lemma scale_cancel_right [simp]: "scale a x = scale b x \ a = b \ x = 0" by (auto intro: scale_right_imp_eq) @@ -530,6 +530,9 @@ lemma dist_triangle: "dist x z \ dist x y + dist y z" using dist_triangle2 [of x z y] by (simp add: dist_commute) +lemma dist_triangle3: "dist x y \ dist a x + dist a y" +using dist_triangle2 [of x y a] by (simp add: dist_commute) + subclass topological_space proof have "\e::real. 0 < e" @@ -568,7 +571,7 @@ assumes norm_ge_zero [simp]: "0 \ norm x" and norm_eq_zero [simp]: "norm x = 0 \ x = 0" and norm_triangle_ineq: "norm (x + y) \ norm x + norm y" - and norm_scaleR: "norm (scaleR a x) = \a\ * norm x" + and norm_scaleR [simp]: "norm (scaleR a x) = \a\ * norm x" class real_normed_algebra = real_algebra + real_normed_vector + assumes norm_mult_ineq: "norm (x * y) \ norm x * norm y" @@ -592,32 +595,6 @@ thus "norm (1::'a) = 1" by simp qed -instantiation real :: real_normed_field -begin - -definition real_norm_def [simp]: - "norm r = \r\" - -definition dist_real_def: - "dist x y = \x - y\" - -definition open_real_def [code del]: - "open (S :: real set) \ (\x\S. \e>0. \y. dist y x < e \ y \ S)" - -instance -apply (intro_classes, unfold real_norm_def real_scaleR_def) -apply (rule dist_real_def) -apply (rule open_real_def) -apply (simp add: real_sgn_def) -apply (rule abs_ge_zero) -apply (rule abs_eq_0) -apply (rule abs_triangle_ineq) -apply (rule abs_mult) -apply (rule abs_mult) -done - -end - lemma norm_zero [simp]: "norm (0::'a::real_normed_vector) = 0" by simp @@ -724,7 +701,7 @@ lemma norm_of_real [simp]: "norm (of_real r :: 'a::real_normed_algebra_1) = \r\" -unfolding of_real_def by (simp add: norm_scaleR) +unfolding of_real_def by simp lemma norm_number_of [simp]: "norm (number_of w::'a::{number_ring,real_normed_algebra_1}) @@ -797,6 +774,76 @@ using norm_triangle_ineq4 [of "x - z" "y - z"] by simp qed + +subsection {* Class instances for real numbers *} + +instantiation real :: real_normed_field +begin + +definition real_norm_def [simp]: + "norm r = \r\" + +definition dist_real_def: + "dist x y = \x - y\" + +definition open_real_def [code del]: + "open (S :: real set) \ (\x\S. \e>0. \y. dist y x < e \ y \ S)" + +instance +apply (intro_classes, unfold real_norm_def real_scaleR_def) +apply (rule dist_real_def) +apply (rule open_real_def) +apply (simp add: real_sgn_def) +apply (rule abs_ge_zero) +apply (rule abs_eq_0) +apply (rule abs_triangle_ineq) +apply (rule abs_mult) +apply (rule abs_mult) +done + +end + +lemma open_real_lessThan [simp]: + fixes a :: real shows "open {.. (\y. \y - x\ < a - x \ y \ {..e>0. \y. \y - x\ < e \ y \ {.. (\y. \y - x\ < x - a \ y \ {a<..})" by auto + thus "\e>0. \y. \y - x\ < e \ y \ {a<..}" .. +qed + +lemma open_real_greaterThanLessThan [simp]: + fixes a b :: real shows "open {a<.. {.. {..b}" by auto + thus "closed {a..b}" by (simp add: closed_Int) +qed + + subsection {* Extra type constraints *} text {* Only allow @{term "open"} in class @{text topological_space}. *} @@ -819,7 +866,7 @@ lemma norm_sgn: "norm (sgn(x::'a::real_normed_vector)) = (if x = 0 then 0 else 1)" -by (simp add: sgn_div_norm norm_scaleR) +by (simp add: sgn_div_norm) lemma sgn_zero [simp]: "sgn(0::'a::real_normed_vector) = 0" by (simp add: sgn_div_norm) @@ -832,7 +879,7 @@ lemma sgn_scaleR: "sgn (scaleR r x) = scaleR (sgn r) (sgn(x::'a::real_normed_vector))" -by (simp add: sgn_div_norm norm_scaleR mult_ac) +by (simp add: sgn_div_norm mult_ac) lemma sgn_one [simp]: "sgn (1::'a::real_normed_algebra_1) = 1" by (simp add: sgn_div_norm) @@ -1000,8 +1047,7 @@ apply (rule scaleR_right_distrib) apply simp apply (rule scaleR_left_commute) -apply (rule_tac x="1" in exI) -apply (simp add: norm_scaleR) +apply (rule_tac x="1" in exI, simp) done interpretation scaleR_left: bounded_linear "\r. scaleR r x" diff -r 78ac5c304db7 -r ef6d67b1ad10 src/HOL/SEQ.thy --- a/src/HOL/SEQ.thy Thu Jun 11 19:49:02 2009 +0200 +++ b/src/HOL/SEQ.thy Sat Jun 13 16:29:15 2009 +0200 @@ -348,23 +348,7 @@ fixes L :: "'a \ 'b::real_normed_vector" assumes n: "\n. n \ S \ X n ----> L n" shows "(\m. \n\S. X n m) ----> (\n\S. L n)" -proof (cases "finite S") - case True - thus ?thesis using n - proof (induct) - case empty - show ?case - by (simp add: LIMSEQ_const) - next - case insert - thus ?case - by (simp add: LIMSEQ_add) - qed -next - case False - thus ?thesis - by (simp add: LIMSEQ_const) -qed +using n unfolding LIMSEQ_conv_tendsto by (rule tendsto_setsum) lemma LIMSEQ_setprod: fixes L :: "'a \ 'b::{real_normed_algebra,comm_ring_1}" diff -r 78ac5c304db7 -r ef6d67b1ad10 src/HOL/Transitive_Closure.thy --- a/src/HOL/Transitive_Closure.thy Thu Jun 11 19:49:02 2009 +0200 +++ b/src/HOL/Transitive_Closure.thy Sat Jun 13 16:29:15 2009 +0200 @@ -486,6 +486,28 @@ lemmas tranclD = tranclpD [to_set] +lemma converse_tranclpE: + assumes major: "tranclp r x z" + assumes base: "r x z ==> P" + assumes step: "\ y. [| r x y; tranclp r y z |] ==> P" + shows P +proof - + from tranclpD[OF major] + obtain y where "r x y" and "rtranclp r y z" by iprover + from this(2) show P + proof (cases rule: rtranclp.cases) + case rtrancl_refl + with `r x y` base show P by iprover + next + case rtrancl_into_rtrancl + from this have "tranclp r y z" + by (iprover intro: rtranclp_into_tranclp1) + with `r x y` step show P by iprover + qed +qed + +lemmas converse_tranclE = converse_tranclpE [to_set] + lemma tranclD2: "(x, y) \ R\<^sup>+ \ \z. (x, z) \ R\<^sup>* \ (z, y) \ R" by (blast elim: tranclE intro: trancl_into_rtrancl) diff -r 78ac5c304db7 -r ef6d67b1ad10 src/HOL/ex/Predicate_Compile_ex.thy --- a/src/HOL/ex/Predicate_Compile_ex.thy Thu Jun 11 19:49:02 2009 +0200 +++ b/src/HOL/ex/Predicate_Compile_ex.thy Sat Jun 13 16:29:15 2009 +0200 @@ -7,8 +7,6 @@ | "even n \ odd (Suc n)" | "odd n \ even (Suc n)" - - code_pred even . thm odd.equation @@ -57,19 +55,23 @@ (*FIXME values 10 "{(ys, zs). partition (\n. n mod 2 = 0) [0, Suc 0, 2, 3, 4, 5, 6, 7] ys zs}" *) -thm tranclp.intros -(* lemma [code_pred_intros]: "r a b ==> tranclp r a b" "r a b ==> tranclp r b c ==> tranclp r a c" by auto -*) + +(* Setup requires quick and dirty proof *) (* -code_pred tranclp . +code_pred tranclp +proof - + case tranclp + from this converse_tranclpE[OF this(1)] show thesis by metis +qed thm tranclp.equation *) + inductive succ :: "nat \ nat \ bool" where "succ 0 1" | "succ m n \ succ (Suc m) (Suc n)" @@ -77,8 +79,7 @@ code_pred succ . thm succ.equation - -(* TODO: requires alternative rules for trancl *) +(* FIXME: why does this not terminate? *) (* values 20 "{n. tranclp succ 10 n}" values "{n. tranclp succ n 10}" diff -r 78ac5c304db7 -r ef6d67b1ad10 src/HOL/ex/predicate_compile.ML --- a/src/HOL/ex/predicate_compile.ML Thu Jun 11 19:49:02 2009 +0200 +++ b/src/HOL/ex/predicate_compile.ML Sat Jun 13 16:29:15 2009 +0200 @@ -135,6 +135,7 @@ cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map string_of_mode ms)) modes)); + datatype predfun_data = PredfunData of { name : string, definition : thm, @@ -180,9 +181,7 @@ (* queries *) fun lookup_pred_data thy name = - case try (Graph.get_node (PredData.get thy)) name of - SOME pred_data => SOME (rep_pred_data pred_data) - | NONE => NONE + Option.map rep_pred_data (try (Graph.get_node (PredData.get thy)) name) fun the_pred_data thy name = case lookup_pred_data thy name of NONE => error ("No such predicate " ^ quote name) @@ -243,16 +242,73 @@ in thy end; *) + +fun imp_prems_conv cv ct = + case Thm.term_of ct of + Const ("==>", _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv) (imp_prems_conv cv) ct + | _ => Conv.all_conv ct + +fun Trueprop_conv cv ct = + case Thm.term_of ct of + Const ("Trueprop", _) $ _ => Conv.arg_conv cv ct + | _ => error "Trueprop_conv" + +fun preprocess_intro thy rule = + Conv.fconv_rule + (imp_prems_conv + (Trueprop_conv (Conv.try_conv (Conv.rewr_conv (Thm.symmetric @{thm Predicate.eq_is_eq}))))) + (Thm.transfer thy rule) + +fun preprocess_elim thy nargs elimrule = let + fun replace_eqs (Const ("Trueprop", _) $ (Const ("op =", T) $ lhs $ rhs)) = + HOLogic.mk_Trueprop (Const (@{const_name Predicate.eq}, T) $ lhs $ rhs) + | replace_eqs t = t + fun preprocess_case t = let + val params = Logic.strip_params t + val (assums1, assums2) = chop nargs (Logic.strip_assums_hyp t) + val assums_hyp' = assums1 @ (map replace_eqs assums2) + in list_all (params, Logic.list_implies (assums_hyp', Logic.strip_assums_concl t)) end + val prems = Thm.prems_of elimrule + val cases' = map preprocess_case (tl prems) + val elimrule' = Logic.list_implies ((hd prems) :: cases', Thm.concl_of elimrule) + in + Thm.equal_elim + (Thm.symmetric (Conv.implies_concl_conv (MetaSimplifier.rewrite true [@{thm eq_is_eq}]) + (cterm_of thy elimrule'))) + elimrule + end; + +fun fetch_pred_data thy name = + case try (InductivePackage.the_inductive (ProofContext.init thy)) name of + SOME (info as (_, result)) => + let + fun is_intro_of intro = + let + val (const, _) = strip_comb (HOLogic.dest_Trueprop (concl_of intro)) + in (fst (dest_Const const) = name) end; + val intros = map (preprocess_intro thy) (filter is_intro_of (#intrs result)) + val elim = nth (#elims result) (find_index (fn s => s = name) (#names (fst info))) + val nparams = length (InductivePackage.params_of (#raw_induct result)) + in (intros, elim, nparams) end + | NONE => error ("No such predicate: " ^ quote name) + (* updaters *) fun add_predfun name mode data = let val add = apsnd (cons (mode, mk_predfun_data data)) in PredData.map (Graph.map_node name (map_pred_data add)) end -fun add_intro thm = let +fun add_intro thm thy = let val (name, _) = dest_Const (fst (strip_intro_concl 0 (prop_of thm))) - fun set (intros, elim, nparams) = (thm::intros, elim, nparams) - in PredData.map (Graph.map_node name (map_pred_data (apfst set))) end + fun cons_intro gr = + case try (Graph.get_node gr) name of + SOME pred_data => Graph.map_node name (map_pred_data + (apfst (fn (intro, elim, nparams) => (thm::intro, elim, nparams)))) gr + | NONE => + let + val nparams = the_default 0 (try (#3 o fetch_pred_data thy) name) + in Graph.new_node (name, mk_pred_data (([thm], NONE, nparams), [])) gr end; + in PredData.map cons_intro thy end fun set_elim thm = let val (name, _) = dest_Const (fst @@ -814,42 +870,6 @@ fun is_Type (Type _) = true | is_Type _ = false -fun imp_prems_conv cv ct = - case Thm.term_of ct of - Const ("==>", _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv) (imp_prems_conv cv) ct - | _ => Conv.all_conv ct - -fun Trueprop_conv cv ct = - case Thm.term_of ct of - Const ("Trueprop", _) $ _ => Conv.arg_conv cv ct - | _ => error "Trueprop_conv" - -fun preprocess_intro thy rule = - Conv.fconv_rule - (imp_prems_conv - (Trueprop_conv (Conv.try_conv (Conv.rewr_conv (Thm.symmetric @{thm Predicate.eq_is_eq}))))) - (Thm.transfer thy rule) - -fun preprocess_elim thy nargs elimrule = let - fun replace_eqs (Const ("Trueprop", _) $ (Const ("op =", T) $ lhs $ rhs)) = - HOLogic.mk_Trueprop (Const (@{const_name Predicate.eq}, T) $ lhs $ rhs) - | replace_eqs t = t - fun preprocess_case t = let - val params = Logic.strip_params t - val (assums1, assums2) = chop nargs (Logic.strip_assums_hyp t) - val assums_hyp' = assums1 @ (map replace_eqs assums2) - in list_all (params, Logic.list_implies (assums_hyp', Logic.strip_assums_concl t)) end - val prems = Thm.prems_of elimrule - val cases' = map preprocess_case (tl prems) - val elimrule' = Logic.list_implies ((hd prems) :: cases', Thm.concl_of elimrule) - in - Thm.equal_elim - (Thm.symmetric (Conv.implies_concl_conv (MetaSimplifier.rewrite true [@{thm eq_is_eq}]) - (cterm_of thy elimrule'))) - elimrule - end; - - (* returns true if t is an application of an datatype constructor *) (* which then consequently would be splitted *) (* else false *) @@ -1369,7 +1389,7 @@ fun mk_casesrule ctxt nparams introrules = let - val intros = map prop_of introrules + val intros = map (Logic.unvarify o prop_of) introrules val (pred, (params, args)) = strip_intro_concl nparams (hd intros) val ([propname], ctxt1) = Variable.variant_fixes ["thesis"] ctxt val prop = HOLogic.mk_Trueprop (Free (propname, HOLogic.boolT)) @@ -1390,26 +1410,13 @@ in Logic.list_implies (assm :: cases, prop) end; (* code dependency graph *) - -fun fetch_pred_data thy name = - case try (InductivePackage.the_inductive (ProofContext.init thy)) name of - SOME (info as (_, result)) => - let - fun is_intro_of intro = - let - val (const, _) = strip_comb (HOLogic.dest_Trueprop (concl_of intro)) - in (fst (dest_Const const) = name) end; - val intros = map (preprocess_intro thy) (filter is_intro_of (#intrs result)) - val elim = nth (#elims result) (find_index (fn s => s = name) (#names (fst info))) - val nparams = length (InductivePackage.params_of (#raw_induct result)) - in mk_pred_data ((intros, SOME elim, nparams), []) end - | NONE => error ("No such predicate: " ^ quote name) -fun dependencies_of (thy : theory) name = +fun dependencies_of thy name = let fun is_inductive_predicate thy name = is_some (try (InductivePackage.the_inductive (ProofContext.init thy)) name) - val data = fetch_pred_data thy name + val (intro, elim, nparams) = fetch_pred_data thy name + val data = mk_pred_data ((intro, SOME elim, nparams), []) val intros = map Thm.prop_of (#intros (rep_pred_data data)) val keys = fold Term.add_consts intros [] |> map fst |> filter (is_inductive_predicate thy) @@ -1458,12 +1465,20 @@ let val nparams = nparams_of thy' const val intros = intros_of thy' const - in mk_casesrule lthy' nparams intros end + in mk_casesrule lthy' nparams intros end val cases_rules = map mk_cases preds + val cases = + map (fn case_rule => RuleCases.Case {fixes = [], + assumes = [("", Logic.strip_imp_prems case_rule)], + binds = [], cases = []}) cases_rules + val case_env = map2 (fn p => fn c => (Long_Name.base_name p, SOME c)) preds cases + val _ = Output.tracing (commas (map fst case_env)) + val lthy'' = ProofContext.add_cases true case_env lthy' + fun after_qed thms = LocalTheory.theory (fold set_elim (map the_single thms) #> add_equations const) in - Proof.theorem_i NONE after_qed (map (single o (rpair [])) cases_rules) lthy' + Proof.theorem_i NONE after_qed (map (single o (rpair [])) cases_rules) lthy'' end; structure P = OuterParse