# HG changeset patch # User himmelma # Date 1256296998 -7200 # Node ID 2083bde13ce101679fbff733a007d35689c62511 # Parent 1fad3160d873412c09582184472b3229dbfd7182 distinguished session for multivariate analysis diff -r 1fad3160d873 -r 2083bde13ce1 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Fri Oct 23 14:33:07 2009 +0200 +++ b/src/HOL/IsaMakefile Fri Oct 23 13:23:18 2009 +0200 @@ -323,15 +323,14 @@ $(LOG)/HOL-Library.gz: $(OUT)/HOL Library/SetsAndFunctions.thy \ Library/Abstract_Rat.thy Library/BigO.thy Library/ContNotDenum.thy \ - Library/Efficient_Nat.thy Library/Euclidean_Space.thy \ + Library/Efficient_Nat.thy \ Library/Sum_Of_Squares.thy Library/Sum_Of_Squares/sos_wrapper.ML \ Library/Sum_Of_Squares/sum_of_squares.ML Library/Fset.thy \ - Library/Convex_Euclidean_Space.thy Library/Glbs.thy \ + Library/Glbs.thy \ Library/normarith.ML Library/Executable_Set.thy \ Library/Infinite_Set.thy Library/FuncSet.thy \ - Library/Permutations.thy Library/Determinants.thy Library/Bit.thy \ - Library/Topology_Euclidean_Space.thy \ - Library/Finite_Cartesian_Product.thy Library/FrechetDeriv.thy \ + Library/Permutations.thy Library/Bit.thy \ + Library/FrechetDeriv.thy \ Library/Fraction_Field.thy Library/Fundamental_Theorem_Algebra.thy \ Library/Inner_Product.thy Library/Kleene_Algebra.thy \ Library/Lattice_Syntax.thy \ @@ -1005,6 +1004,19 @@ @cd TLA; $(ISABELLE_TOOL) usedir $(OUT)/TLA Memory +## HOL-Multivariate_Analysis + +HOL-Multivariate_Analysis: HOL $(OUT)/HOL-Multivariate_Analysis + +$(OUT)/HOL-Multivariate_Analysis: $(OUT)/HOL Multivariate_Analysis/ROOT.ML \ + Multivariate_Analysis/Multivariate_Analysis.thy \ + Multivariate_Analysis/Determinants.thy \ + Multivariate_Analysis/Finite_Cartesian_Product.thy \ + Multivariate_Analysis/Euclidean_Space.thy \ + Multivariate_Analysis/Topology_Euclidean_Space.thy \ + Multivariate_Analysis/Convex_Euclidean_Space.thy + @cd Multivariate_Analysis; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL HOL-Multivariate_Analysis + ## HOL-Nominal HOL-Nominal: HOL $(OUT)/HOL-Nominal diff -r 1fad3160d873 -r 2083bde13ce1 src/HOL/Library/#Topology_Euclidean_Space.thy# --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/#Topology_Euclidean_Space.thy# Fri Oct 23 13:23:18 2009 +0200 @@ -0,0 +1,6029 @@ +(* Title: Topology + Author: Amine Chaieb, University of Cambridge + Author: Robert Himmelmann, TU Muenchen +*) + +header {* Elementary topology in Euclidean space. *} + +theory Topology_Euclidean_Space +imports SEQ Euclidean_Space Product_Vector +begin + +declare fstcart_pastecart[simp] sndcart_pastecart[simp] + +subsection{* General notion of a topology *} + +definition "istopology L \ {} \ L \ (\S \L. \T \L. S \ T \ L) \ (\K. K \L \ \ K \ L)" +typedef (open) 'a topology = "{L::('a set) set. istopology L}" + morphisms "openin" "topology" + unfolding istopology_def by blast + +lemma istopology_open_in[intro]: "istopology(openin U)" + using openin[of U] by blast + +lemma topology_inverse': "istopology U \ openin (topology U) = U" + using topology_inverse[unfolded mem_def Collect_def] . + +lemma topology_inverse_iff: "istopology U \ openin (topology U) = U" + using topology_inverse[of U] istopology_open_in[of "topology U"] by auto + +lemma topology_eq: "T1 = T2 \ (\S. openin T1 S \ openin T2 S)" +proof- + {assume "T1=T2" hence "\S. openin T1 S \ openin T2 S" by simp} + moreover + {assume H: "\S. openin T1 S \ openin T2 S" + hence "openin T1 = openin T2" by (metis mem_def set_ext) + hence "topology (openin T1) = topology (openin T2)" by simp + hence "T1 = T2" unfolding openin_inverse .} + ultimately show ?thesis by blast +qed + +text{* Infer the "universe" from union of all sets in the topology. *} + +definition "topspace T = \{S. openin T S}" + +subsection{* Main properties of open sets *} + +lemma openin_clauses: + fixes U :: "'a topology" + shows "openin U {}" + "\S T. openin U S \ openin U T \ openin U (S\T)" + "\K. (\S \ K. openin U S) \ openin U (\K)" + using openin[of U] unfolding istopology_def Collect_def mem_def + by (metis mem_def subset_eq)+ + +lemma openin_subset[intro]: "openin U S \ S \ topspace U" + unfolding topspace_def by blast +lemma openin_empty[simp]: "openin U {}" by (simp add: openin_clauses) + +lemma openin_Int[intro]: "openin U S \ openin U T \ openin U (S \ T)" + by (simp add: openin_clauses) + +lemma openin_Union[intro]: "(\S \K. openin U S) \ openin U (\ K)" by (simp add: openin_clauses) + +lemma openin_Un[intro]: "openin U S \ openin U T \ openin U (S \ T)" + using openin_Union[of "{S,T}" U] by auto + +lemma openin_topspace[intro, simp]: "openin U (topspace U)" by (simp add: openin_Union topspace_def) + +lemma openin_subopen: "openin U S \ (\x \ S. \T. openin U T \ x \ T \ T \ S)" (is "?lhs \ ?rhs") +proof- + {assume ?lhs then have ?rhs by auto } + moreover + {assume H: ?rhs + then obtain t where t: "\x\S. openin U (t x) \ x \ t x \ t x \ S" + unfolding Ball_def ex_simps(6)[symmetric] choice_iff by blast + from t have th0: "\x\ t`S. openin U x" by auto + have "\ t`S = S" using t by auto + with openin_Union[OF th0] have "openin U S" by simp } + ultimately show ?thesis by blast +qed + +subsection{* Closed sets *} + +definition "closedin U S \ S \ topspace U \ openin U (topspace U - S)" + +lemma closedin_subset: "closedin U S \ S \ topspace U" by (metis closedin_def) +lemma closedin_empty[simp]: "closedin U {}" by (simp add: closedin_def) +lemma closedin_topspace[intro,simp]: + "closedin U (topspace U)" by (simp add: closedin_def) +lemma closedin_Un[intro]: "closedin U S \ closedin U T \ closedin U (S \ T)" + by (auto simp add: Diff_Un closedin_def) + +lemma Diff_Inter[intro]: "A - \S = \ {A - s|s. s\S}" by auto +lemma closedin_Inter[intro]: assumes Ke: "K \ {}" and Kc: "\S \K. closedin U S" + shows "closedin U (\ K)" using Ke Kc unfolding closedin_def Diff_Inter by auto + +lemma closedin_Int[intro]: "closedin U S \ closedin U T \ closedin U (S \ T)" + using closedin_Inter[of "{S,T}" U] by auto + +lemma Diff_Diff_Int: "A - (A - B) = A \ B" by blast +lemma openin_closedin_eq: "openin U S \ S \ topspace U \ closedin U (topspace U - S)" + apply (auto simp add: closedin_def) + apply (metis openin_subset subset_eq) + apply (auto simp add: Diff_Diff_Int) + apply (subgoal_tac "topspace U \ S = S") + by auto + +lemma openin_closedin: "S \ topspace U \ (openin U S \ closedin U (topspace U - S))" + by (simp add: openin_closedin_eq) + +lemma openin_diff[intro]: assumes oS: "openin U S" and cT: "closedin U T" shows "openin U (S - T)" +proof- + have "S - T = S \ (topspace U - T)" using openin_subset[of U S] oS cT + by (auto simp add: topspace_def openin_subset) + then show ?thesis using oS cT by (auto simp add: closedin_def) +qed + +lemma closedin_diff[intro]: assumes oS: "closedin U S" and cT: "openin U T" shows "closedin U (S - T)" +proof- + have "S - T = S \ (topspace U - T)" using closedin_subset[of U S] oS cT + by (auto simp add: topspace_def ) + then show ?thesis using oS cT by (auto simp add: openin_closedin_eq) +qed + +subsection{* Subspace topology. *} + +definition "subtopology U V = topology {S \ V |S. openin U S}" + +lemma istopology_subtopology: "istopology {S \ V |S. openin U S}" (is "istopology ?L") +proof- + have "{} \ ?L" by blast + {fix A B assume A: "A \ ?L" and B: "B \ ?L" + from A B obtain Sa and Sb where Sa: "openin U Sa" "A = Sa \ V" and Sb: "openin U Sb" "B = Sb \ V" by blast + have "A\B = (Sa \ Sb) \ V" "openin U (Sa \ Sb)" using Sa Sb by blast+ + then have "A \ B \ ?L" by blast} + moreover + {fix K assume K: "K \ ?L" + have th0: "?L = (\S. S \ V) ` openin U " + apply (rule set_ext) + apply (simp add: Ball_def image_iff) + by (metis mem_def) + from K[unfolded th0 subset_image_iff] + obtain Sk where Sk: "Sk \ openin U" "K = (\S. S \ V) ` Sk" by blast + have "\K = (\Sk) \ V" using Sk by auto + moreover have "openin U (\ Sk)" using Sk by (auto simp add: subset_eq mem_def) + ultimately have "\K \ ?L" by blast} + ultimately show ?thesis unfolding istopology_def by blast +qed + +lemma openin_subtopology: + "openin (subtopology U V) S \ (\ T. (openin U T) \ (S = T \ V))" + unfolding subtopology_def topology_inverse'[OF istopology_subtopology] + by (auto simp add: Collect_def) + +lemma topspace_subtopology: "topspace(subtopology U V) = topspace U \ V" + by (auto simp add: topspace_def openin_subtopology) + +lemma closedin_subtopology: + "closedin (subtopology U V) S \ (\T. closedin U T \ S = T \ V)" + unfolding closedin_def topspace_subtopology + apply (simp add: openin_subtopology) + apply (rule iffI) + apply clarify + apply (rule_tac x="topspace U - T" in exI) + by auto + +lemma openin_subtopology_refl: "openin (subtopology U V) V \ V \ topspace U" + unfolding openin_subtopology + apply (rule iffI, clarify) + apply (frule openin_subset[of U]) apply blast + apply (rule exI[where x="topspace U"]) + by auto + +lemma subtopology_superset: assumes UV: "topspace U \ V" + shows "subtopology U V = U" +proof- + {fix S + {fix T assume T: "openin U T" "S = T \ V" + from T openin_subset[OF T(1)] UV have eq: "S = T" by blast + have "openin U S" unfolding eq using T by blast} + moreover + {assume S: "openin U S" + hence "\T. openin U T \ S = T \ V" + using openin_subset[OF S] UV by auto} + ultimately have "(\T. openin U T \ S = T \ V) \ openin U S" by blast} + then show ?thesis unfolding topology_eq openin_subtopology by blast +qed + + +lemma subtopology_topspace[simp]: "subtopology U (topspace U) = U" + by (simp add: subtopology_superset) + +lemma subtopology_UNIV[simp]: "subtopology U UNIV = U" + by (simp add: subtopology_superset) + +subsection{* The universal Euclidean versions are what we use most of the time *} + +definition + euclidean :: "'a::topological_space topology" where + "euclidean = topology open" + +lemma open_openin: "open S \ openin euclidean S" + unfolding euclidean_def + apply (rule cong[where x=S and y=S]) + apply (rule topology_inverse[symmetric]) + apply (auto simp add: istopology_def) + by (auto simp add: mem_def subset_eq) + +lemma topspace_euclidean: "topspace euclidean = UNIV" + apply (simp add: topspace_def) + apply (rule set_ext) + by (auto simp add: open_openin[symmetric]) + +lemma topspace_euclidean_subtopology[simp]: "topspace (subtopology euclidean S) = S" + by (simp add: topspace_euclidean topspace_subtopology) + +lemma closed_closedin: "closed S \ closedin euclidean S" + by (simp add: closed_def closedin_def topspace_euclidean open_openin Compl_eq_Diff_UNIV) + +lemma open_subopen: "open S \ (\x\S. \T. open T \ x \ T \ T \ S)" + by (simp add: open_openin openin_subopen[symmetric]) + +subsection{* Open and closed balls. *} + +definition + ball :: "'a::metric_space \ real \ 'a set" where + "ball x e = {y. dist x y < e}" + +definition + cball :: "'a::metric_space \ real \ 'a set" where + "cball x e = {y. dist x y \ e}" + +lemma mem_ball[simp]: "y \ ball x e \ dist x y < e" by (simp add: ball_def) +lemma mem_cball[simp]: "y \ cball x e \ dist x y \ e" by (simp add: cball_def) + +lemma mem_ball_0 [simp]: + fixes x :: "'a::real_normed_vecto" + shows "x \ ball 0 e \ norm x < e" + by (simp add: dist_norm) + +lemma mem_cball_0 [simp]: + fixes x :: "'a::real_normed_vector" + shows "x \ cball 0 e \ norm x \ e" + by (simp add: dist_norm) + +lemma centre_in_cball[simp]: "x \ cball x e \ 0\ e" by simp +lemma ball_subset_cball[simp,intro]: "ball x e \ cball x e" by (simp add: subset_eq) +lemma subset_ball[intro]: "d <= e ==> ball x d \ ball x e" by (simp add: subset_eq) +lemma subset_cball[intro]: "d <= e ==> cball x d \ cball x e" by (simp add: subset_eq) +lemma ball_max_Un: "ball a (max r s) = ball a r \ ball a s" + by (simp add: expand_set_eq) arith + +lemma ball_min_Int: "ball a (min r s) = ball a r \ ball a s" + by (simp add: expand_set_eq) + +subsection{* Topological properties of open balls *} + +lemma diff_less_iff: "(a::real) - b > 0 \ a > b" + "(a::real) - b < 0 \ a < b" + "a - b < c \ a < c +b" "a - b > c \ a > c +b" by arith+ +lemma diff_le_iff: "(a::real) - b \ 0 \ a \ b" "(a::real) - b \ 0 \ a \ b" + "a - b \ c \ a \ c +b" "a - b \ c \ a \ c +b" by arith+ + +lemma open_ball[intro, simp]: "open (ball x e)" + unfolding open_dist ball_def Collect_def Ball_def mem_def + unfolding dist_commute + apply clarify + apply (rule_tac x="e - dist xa x" in exI) + using dist_triangle_alt[where z=x] + apply (clarsimp simp add: diff_less_iff) + apply atomize + apply (erule_tac x="y" in allE) + apply (erule_tac x="xa" in allE) + by arith + +lemma centre_in_ball[simp]: "x \ ball x e \ e > 0" by (metis mem_ball dist_self) +lemma open_contains_ball: "open S \ (\x\S. \e>0. ball x e \ S)" + unfolding open_dist subset_eq mem_ball Ball_def dist_commute .. + +lemma open_contains_ball_eq: "open S \ \x. x\S \ (\e>0. ball x e \ S)" + by (metis open_contains_ball subset_eq centre_in_ball) + +lemma ball_eq_empty[simp]: "ball x e = {} \ e \ 0" + unfolding mem_ball expand_set_eq + apply (simp add: not_less) + by (metis zero_le_dist order_trans dist_self) + +lemma ball_empty[intro]: "e \ 0 ==> ball x e = {}" by simp + +subsection{* Basic "localization" results are handy for connectedness. *} + +lemma openin_open: "openin (subtopology euclidean U) S \ (\T. open T \ (S = U \ T))" + by (auto simp add: openin_subtopology open_openin[symmetric]) + +lemma openin_open_Int[intro]: "open S \ openin (subtopology euclidean U) (U \ S)" + by (auto simp add: openin_open) + +lemma open_openin_trans[trans]: + "open S \ open T \ T \ S \ openin (subtopology euclidean S) T" + by (metis Int_absorb1 openin_open_Int) + +lemma open_subset: "S \ T \ open S \ openin (subtopology euclidean T) S" + by (auto simp add: openin_open) + +lemma closedin_closed: "closedin (subtopology euclidean U) S \ (\T. closed T \ S = U \ T)" + by (simp add: closedin_subtopology closed_closedin Int_ac) + +lemma closedin_closed_Int: "closed S ==> closedin (subtopology euclidean U) (U \ S)" + by (metis closedin_closed) + +lemma closed_closedin_trans: "closed S \ closed T \ T \ S \ closedin (subtopology euclidean S) T" + apply (subgoal_tac "S \ T = T" ) + apply auto + apply (frule closedin_closed_Int[of T S]) + by simp + +lemma closed_subset: "S \ T \ closed S \ closedin (subtopology euclidean T) S" + by (auto simp add: closedin_closed) + +lemma openin_euclidean_subtopology_iff: + fixes S U :: "'a::metric_space set" + shows "openin (subtopology euclidean U) S + \ S \ U \ (\x\S. \e>0. \x'\U. dist x' x < e \ x'\ S)" (is "?lhs \ ?rhs") +proof- + {assume ?lhs hence ?rhs unfolding openin_subtopology open_openin[symmetric] + by (simp add: open_dist) blast} + moreover + {assume SU: "S \ U" and H: "\x. x \ S \ \e>0. \x'\U. dist x' x < e \ x' \ S" + from H obtain d where d: "\x . x\ S \ d x > 0 \ (\x' \ U. dist x' x < d x \ x' \ S)" + by metis + let ?T = "\{B. \x\S. B = ball x (d x)}" + have oT: "open ?T" by auto + { fix x assume "x\S" + hence "x \ \{B. \x\S. B = ball x (d x)}" + apply simp apply(rule_tac x="ball x(d x)" in exI) apply auto + by (rule d [THEN conjunct1]) + hence "x\ ?T \ U" using SU and `x\S` by auto } + moreover + { fix y assume "y\?T" + then obtain B where "y\B" "B\{B. \x\S. B = ball x (d x)}" by auto + then obtain x where "x\S" and x:"y \ ball x (d x)" by auto + assume "y\U" + hence "y\S" using d[OF `x\S`] and x by(auto simp add: dist_commute) } + ultimately have "S = ?T \ U" by blast + with oT have ?lhs unfolding openin_subtopology open_openin[symmetric] by blast} + ultimately show ?thesis by blast +qed + +text{* These "transitivity" results are handy too. *} + +lemma openin_trans[trans]: "openin (subtopology euclidean T) S \ openin (subtopology euclidean U) T + \ openin (subtopology euclidean U) S" + unfolding open_openin openin_open by blast + +lemma openin_open_trans: "openin (subtopology euclidean T) S \ open T \ open S" + by (auto simp add: openin_open intro: openin_trans) + +lemma closedin_trans[trans]: + "closedin (subtopology euclidean T) S \ + closedin (subtopology euclidean U) T + ==> closedin (subtopology euclidean U) S" + by (auto simp add: closedin_closed closed_closedin closed_Inter Int_assoc) + +lemma closedin_closed_trans: "closedin (subtopology euclidean T) S \ closed T \ closed S" + by (auto simp add: closedin_closed intro: closedin_trans) + +subsection{* Connectedness *} + +definition "connected S \ + ~(\e1 e2. open e1 \ open e2 \ S \ (e1 \ e2) \ (e1 \ e2 \ S = {}) + \ ~(e1 \ S = {}) \ ~(e2 \ S = {}))" + +lemma connected_local: + "connected S \ ~(\e1 e2. + openin (subtopology euclidean S) e1 \ + openin (subtopology euclidean S) e2 \ + S \ e1 \ e2 \ + e1 \ e2 = {} \ + ~(e1 = {}) \ + ~(e2 = {}))" +unfolding connected_def openin_open by (safe, blast+) + +lemma exists_diff: "(\S. P(UNIV - S)) \ (\S. P S)" (is "?lhs \ ?rhs") +proof- + + {assume "?lhs" hence ?rhs by blast } + moreover + {fix S assume H: "P S" + have "S = UNIV - (UNIV - S)" by auto + with H have "P (UNIV - (UNIV - S))" by metis } + ultimately show ?thesis by metis +qed + +lemma connected_clopen: "connected S \ + (\T. openin (subtopology euclidean S) T \ + closedin (subtopology euclidean S) T \ T = {} \ T = S)" (is "?lhs \ ?rhs") +proof- + have " \ connected S \ (\e1 e2. open e1 \ open (UNIV - e2) \ S \ e1 \ (UNIV - e2) \ e1 \ (UNIV - e2) \ S = {} \ e1 \ S \ {} \ (UNIV - e2) \ S \ {})" + unfolding connected_def openin_open closedin_closed + apply (subst exists_diff) by blast + hence th0: "connected S \ \ (\e2 e1. closed e2 \ open e1 \ S \ e1 \ (UNIV - e2) \ e1 \ (UNIV - e2) \ S = {} \ e1 \ S \ {} \ (UNIV - e2) \ S \ {})" + (is " _ \ \ (\e2 e1. ?P e2 e1)") apply (simp add: closed_def Compl_eq_Diff_UNIV) by metis + + have th1: "?rhs \ \ (\t' t. closed t'\t = S\t' \ t\{} \ t\S \ (\t'. open t' \ t = S \ t'))" + (is "_ \ \ (\t' t. ?Q t' t)") + unfolding connected_def openin_open closedin_closed by auto + {fix e2 + {fix e1 have "?P e2 e1 \ (\t. closed e2 \ t = S\e2 \ open e1 \ t = S\e1 \ t\{} \ t\S)" + by auto} + then have "(\e1. ?P e2 e1) \ (\t. ?Q e2 t)" by metis} + then have "\e2. (\e1. ?P e2 e1) \ (\t. ?Q e2 t)" by blast + then show ?thesis unfolding th0 th1 by simp +qed + +lemma connected_empty[simp, intro]: "connected {}" + by (simp add: connected_def) + +subsection{* Hausdorff and other separation properties *} + +class t0_space = + assumes t0_space: "x \ y \ \U. open U \ \ (x \ U \ y \ U)" + +class t1_space = + assumes t1_space: "x \ y \ \U V. open U \ open V \ x \ U \ y \ U \ x \ V \ y \ V" +begin + +subclass t0_space +proof +qed (fast dest: t1_space) + +end + +text {* T2 spaces are also known as Hausdorff spaces. *} + +class t2_space = + assumes hausdorff: "x \ y \ \U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" +begin + +subclass t1_space +proof +qed (fast dest: hausdorff) + +end + +instance metric_space \ t2_space +proof + fix x y :: "'a::metric_space" + assume xy: "x \ y" + let ?U = "ball x (dist x y / 2)" + let ?V = "ball y (dist x y / 2)" + have th0: "\d x y z. (d x z :: real) <= d x y + d y z \ d y z = d z y + ==> ~(d x y * 2 < d x z \ d z y * 2 < d x z)" by arith + have "open ?U \ open ?V \ x \ ?U \ y \ ?V \ ?U \ ?V = {}" + using dist_pos_lt[OF xy] th0[of dist,OF dist_triangle dist_commute] + by (auto simp add: expand_set_eq) + then show "\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" + by blast +qed + +lemma separation_t2: + fixes x y :: "'a::t2_space" + shows "x \ y \ (\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {})" + using hausdorff[of x y] by blast + +lemma separation_t1: + fixes x y :: "'a::t1_space" + shows "x \ y \ (\U V. open U \ open V \ x \U \ y\ U \ x\V \ y\V)" + using t1_space[of x y] by blast + +lemma separation_t0: + fixes x y :: "'a::t0_space" + shows "x \ y \ (\U. open U \ ~(x\U \ y\U))" + using t0_space[of x y] by blast + +subsection{* Limit points *} + +definition + islimpt:: "'a::topological_space \ 'a set \ bool" + (infixr "islimpt" 60) where + "x islimpt S \ (\T. x\T \ open T \ (\y\S. y\T \ y\x))" + +lemma islimptI: + assumes "\T. x \ T \ open T \ \y\S. y \ T \ y \ x" + shows "x islimpt S" + using assms unfolding islimpt_def by auto + +lemma islimptE: + assumes "x islimpt S" and "x \ T" and "open T" + obtains y where "y \ S" and "y \ T" and "y \ x" + using assms unfolding islimpt_def by auto + +lemma islimpt_subset: "x islimpt S \ S \ T ==> x islimpt T" by (auto simp add: islimpt_def) + +lemma islimpt_approachable: + fixes x :: "'a::metric_space" + shows "x islimpt S \ (\e>0. \x'\S. x' \ x \ dist x' x < e)" + unfolding islimpt_def + apply auto + apply(erule_tac x="ball x e" in allE) + apply auto + apply(rule_tac x=y in bexI) + apply (auto simp add: dist_commute) + apply (simp add: open_dist, drule (1) bspec) + apply (clarify, drule spec, drule (1) mp, auto) + done + +lemma islimpt_approachable_le: + fixes x :: "'a::metric_space" + shows "x islimpt S \ (\e>0. \x'\ S. x' \ x \ dist x' x <= e)" + unfolding islimpt_approachable + using approachable_lt_le[where f="\x'. dist x' x" and P="\x'. \ (x'\S \ x'\x)"] + by metis (* FIXME: VERY slow! *) + +class perfect_space = + (* FIXME: perfect_space should inherit from topological_space *) + assumes islimpt_UNIV [simp, intro]: "(x::'a::metric_space) islimpt UNIV" + +lemma perfect_choose_dist: + fixes x :: "'a::perfect_space" + shows "0 < r \ \a. a \ x \ dist a x < r" +using islimpt_UNIV [of x] +by (simp add: islimpt_approachable) + +instance real :: perfect_space +apply default +apply (rule islimpt_approachable [THEN iffD2]) +apply (clarify, rule_tac x="x + e/2" in bexI) +apply (auto simp add: dist_norm) +done + +instance "^" :: (perfect_space, finite) perfect_space +proof + fix x :: "'a ^ 'b" + { + fix e :: real assume "0 < e" + def a \ "x $ arbitrary" + have "a islimpt UNIV" by (rule islimpt_UNIV) + with `0 < e` obtain b where "b \ a" and "dist b a < e" + unfolding islimpt_approachable by auto + def y \ "Cart_lambda ((Cart_nth x)(arbitrary := b))" + from `b \ a` have "y \ x" + unfolding a_def y_def by (simp add: Cart_eq) + from `dist b a < e` have "dist y x < e" + unfolding dist_vector_def a_def y_def + apply simp + apply (rule le_less_trans [OF setL2_le_setsum [OF zero_le_dist]]) + apply (subst setsum_diff1' [where a=arbitrary], simp, simp, simp) + done + from `y \ x` and `dist y x < e` + have "\y\UNIV. y \ x \ dist y x < e" by auto + } + then show "x islimpt UNIV" unfolding islimpt_approachable by blast +qed + +lemma closed_limpt: "closed S \ (\x. x islimpt S \ x \ S)" + unfolding closed_def + apply (subst open_subopen) + apply (simp add: islimpt_def subset_eq Compl_eq_Diff_UNIV) + by (metis DiffE DiffI UNIV_I insertCI insert_absorb mem_def) + +lemma islimpt_EMPTY[simp]: "\ x islimpt {}" + unfolding islimpt_def by auto + +lemma closed_positive_orthant: "closed {x::real^'n::finite. \i. 0 \x$i}" +proof- + let ?U = "UNIV :: 'n set" + let ?O = "{x::real^'n. \i. x$i\0}" + {fix x:: "real^'n" and i::'n assume H: "\e>0. \x'\?O. x' \ x \ dist x' x < e" + and xi: "x$i < 0" + from xi have th0: "-x$i > 0" by arith + from H[rule_format, OF th0] obtain x' where x': "x' \?O" "x' \ x" "dist x' x < -x $ i" by blast + have th:" \b a (x::real). abs x <= b \ b <= a ==> ~(a + x < 0)" by arith + have th': "\x (y::real). x < 0 \ 0 <= y ==> abs x <= abs (y - x)" by arith + have th1: "\x$i\ \ \(x' - x)$i\" using x'(1) xi + apply (simp only: vector_component) + by (rule th') auto + have th2: "\dist x x'\ \ \(x' - x)$i\" using component_le_norm[of "x'-x" i] + apply (simp add: dist_norm) by norm + from th[OF th1 th2] x'(3) have False by (simp add: dist_commute) } + then show ?thesis unfolding closed_limpt islimpt_approachable + unfolding not_le[symmetric] by blast +qed + +lemma finite_set_avoid: + fixes a :: "'a::metric_space" + assumes fS: "finite S" shows "\d>0. \x\S. x \ a \ d <= dist a x" +proof(induct rule: finite_induct[OF fS]) + case 1 thus ?case apply auto by ferrack +next + case (2 x F) + from 2 obtain d where d: "d >0" "\x\F. x\a \ d \ dist a x" by blast + {assume "x = a" hence ?case using d by auto } + moreover + {assume xa: "x\a" + let ?d = "min d (dist a x)" + have dp: "?d > 0" using xa d(1) using dist_nz by auto + from d have d': "\x\F. x\a \ ?d \ dist a x" by auto + with dp xa have ?case by(auto intro!: exI[where x="?d"]) } + ultimately show ?case by blast +qed + +lemma islimpt_finite: + fixes S :: "'a::metric_space set" + assumes fS: "finite S" shows "\ a islimpt S" + unfolding islimpt_approachable + using finite_set_avoid[OF fS, of a] by (metis dist_commute not_le) + +lemma islimpt_Un: "x islimpt (S \ T) \ x islimpt S \ x islimpt T" + apply (rule iffI) + defer + apply (metis Un_upper1 Un_upper2 islimpt_subset) + unfolding islimpt_def + apply (rule ccontr, clarsimp, rename_tac A B) + apply (drule_tac x="A \ B" in spec) + apply (auto simp add: open_Int) + done + +lemma discrete_imp_closed: + fixes S :: "'a::metric_space set" + assumes e: "0 < e" and d: "\x \ S. \y \ S. dist y x < e \ y = x" + shows "closed S" +proof- + {fix x assume C: "\e>0. \x'\S. x' \ x \ dist x' x < e" + from e have e2: "e/2 > 0" by arith + from C[rule_format, OF e2] obtain y where y: "y \ S" "y\x" "dist y x < e/2" by blast + let ?m = "min (e/2) (dist x y) " + from e2 y(2) have mp: "?m > 0" by (simp add: dist_nz[THEN sym]) + from C[rule_format, OF mp] obtain z where z: "z \ S" "z\x" "dist z x < ?m" by blast + have th: "dist z y < e" using z y + by (intro dist_triangle_lt [where z=x], simp) + from d[rule_format, OF y(1) z(1) th] y z + have False by (auto simp add: dist_commute)} + then show ?thesis by (metis islimpt_approachable closed_limpt [where 'a='a]) +qed + +subsection{* Interior of a Set *} +definition "interior S = {x. \T. open T \ x \ T \ T \ S}" + +lemma interior_eq: "interior S = S \ open S" + apply (simp add: expand_set_eq interior_def) + apply (subst (2) open_subopen) by (safe, blast+) + +lemma interior_open: "open S ==> (interior S = S)" by (metis interior_eq) + +lemma interior_empty[simp]: "interior {} = {}" by (simp add: interior_def) + +lemma open_interior[simp, intro]: "open(interior S)" + apply (simp add: interior_def) + apply (subst open_subopen) by blast + +lemma interior_interior[simp]: "interior(interior S) = interior S" by (metis interior_eq open_interior) +lemma interior_subset: "interior S \ S" by (auto simp add: interior_def) +lemma subset_interior: "S \ T ==> (interior S) \ (interior T)" by (auto simp add: interior_def) +lemma interior_maximal: "T \ S \ open T ==> T \ (interior S)" by (auto simp add: interior_def) +lemma interior_unique: "T \ S \ open T \ (\T'. T' \ S \ open T' \ T' \ T) \ interior S = T" + by (metis equalityI interior_maximal interior_subset open_interior) +lemma mem_interior: "x \ interior S \ (\e. 0 < e \ ball x e \ S)" + apply (simp add: interior_def) + by (metis open_contains_ball centre_in_ball open_ball subset_trans) + +lemma open_subset_interior: "open S ==> S \ interior T \ S \ T" + by (metis interior_maximal interior_subset subset_trans) + +lemma interior_inter[simp]: "interior(S \ T) = interior S \ interior T" + apply (rule equalityI, simp) + apply (metis Int_lower1 Int_lower2 subset_interior) + by (metis Int_mono interior_subset open_Int open_interior open_subset_interior) + +lemma interior_limit_point [intro]: + fixes x :: "'a::perfect_space" + assumes x: "x \ interior S" shows "x islimpt S" +proof- + from x obtain e where e: "e>0" "\x'. dist x x' < e \ x' \ S" + unfolding mem_interior subset_eq Ball_def mem_ball by blast + { + fix d::real assume d: "d>0" + let ?m = "min d e" + have mde2: "0 < ?m" using e(1) d(1) by simp + from perfect_choose_dist [OF mde2, of x] + obtain y where "y \ x" and "dist y x < ?m" by blast + then have "dist y x < e" "dist y x < d" by simp_all + from `dist y x < e` e(2) have "y \ S" by (simp add: dist_commute) + have "\x'\S. x'\ x \ dist x' x < d" + using `y \ S` `y \ x` `dist y x < d` by fast + } + then show ?thesis unfolding islimpt_approachable by blast +qed + +lemma interior_closed_Un_empty_interior: + assumes cS: "closed S" and iT: "interior T = {}" + shows "interior(S \ T) = interior S" +proof + show "interior S \ interior (S\T)" + by (rule subset_interior, blast) +next + show "interior (S \ T) \ interior S" + proof + fix x assume "x \ interior (S \ T)" + then obtain R where "open R" "x \ R" "R \ S \ T" + unfolding interior_def by fast + show "x \ interior S" + proof (rule ccontr) + assume "x \ interior S" + with `x \ R` `open R` obtain y where "y \ R - S" + unfolding interior_def expand_set_eq by fast + from `open R` `closed S` have "open (R - S)" by (rule open_Diff) + from `R \ S \ T` have "R - S \ T" by fast + from `y \ R - S` `open (R - S)` `R - S \ T` `interior T = {}` + show "False" unfolding interior_def by fast + qed + qed +qed + + +subsection{* Closure of a Set *} + +definition "closure S = S \ {x | x. x islimpt S}" + +lemma closure_interior: "closure S = UNIV - interior (UNIV - S)" +proof- + { fix x + have "x\UNIV - interior (UNIV - S) \ x \ closure S" (is "?lhs = ?rhs") + proof + let ?exT = "\ y. (\T. open T \ y \ T \ T \ UNIV - S)" + assume "?lhs" + hence *:"\ ?exT x" + unfolding interior_def + by simp + { assume "\ ?rhs" + hence False using * + unfolding closure_def islimpt_def + by blast + } + thus "?rhs" + by blast + next + assume "?rhs" thus "?lhs" + unfolding closure_def interior_def islimpt_def + by blast + qed + } + thus ?thesis + by blast +qed + +lemma interior_closure: "interior S = UNIV - (closure (UNIV - S))" +proof- + { fix x + have "x \ interior S \ x \ UNIV - (closure (UNIV - S))" + unfolding interior_def closure_def islimpt_def + by blast (* FIXME: VERY slow! *) + } + thus ?thesis + by blast +qed + +lemma closed_closure[simp, intro]: "closed (closure S)" +proof- + have "closed (UNIV - interior (UNIV -S))" by blast + thus ?thesis using closure_interior[of S] by simp +qed + +lemma closure_hull: "closure S = closed hull S" +proof- + have "S \ closure S" + unfolding closure_def + by blast + moreover + have "closed (closure S)" + using closed_closure[of S] + by assumption + moreover + { fix t + assume *:"S \ t" "closed t" + { fix x + assume "x islimpt S" + hence "x islimpt t" using *(1) + using islimpt_subset[of x, of S, of t] + by blast + } + with * have "closure S \ t" + unfolding closure_def + using closed_limpt[of t] + by auto + } + ultimately show ?thesis + using hull_unique[of S, of "closure S", of closed] + unfolding mem_def + by simp +qed + +lemma closure_eq: "closure S = S \ closed S" + unfolding closure_hull + using hull_eq[of closed, unfolded mem_def, OF closed_Inter, of S] + by (metis mem_def subset_eq) + +lemma closure_closed[simp]: "closed S \ closure S = S" + using closure_eq[of S] + by simp + +lemma closure_closure[simp]: "closure (closure S) = closure S" + unfolding closure_hull + using hull_hull[of closed S] + by assumption + +lemma closure_subset: "S \ closure S" + unfolding closure_hull + using hull_subset[of S closed] + by assumption + +lemma subset_closure: "S \ T \ closure S \ closure T" + unfolding closure_hull + using hull_mono[of S T closed] + by assumption + +lemma closure_minimal: "S \ T \ closed T \ closure S \ T" + using hull_minimal[of S T closed] + unfolding closure_hull mem_def + by simp + +lemma closure_unique: "S \ T \ closed T \ (\ T'. S \ T' \ closed T' \ T \ T') \ closure S = T" + using hull_unique[of S T closed] + unfolding closure_hull mem_def + by simp + +lemma closure_empty[simp]: "closure {} = {}" + using closed_empty closure_closed[of "{}"] + by simp + +lemma closure_univ[simp]: "closure UNIV = UNIV" + using closure_closed[of UNIV] + by simp + +lemma closure_eq_empty: "closure S = {} \ S = {}" + using closure_empty closure_subset[of S] + by blast + +lemma closure_subset_eq: "closure S \ S \ closed S" + using closure_eq[of S] closure_subset[of S] + by simp + +lemma open_inter_closure_eq_empty: + "open S \ (S \ closure T) = {} \ S \ T = {}" + using open_subset_interior[of S "UNIV - T"] + using interior_subset[of "UNIV - T"] + unfolding closure_interior + by auto + +lemma open_inter_closure_subset: + "open S \ (S \ (closure T)) \ closure(S \ T)" +proof + fix x + assume as: "open S" "x \ S \ closure T" + { assume *:"x islimpt T" + have "x islimpt (S \ T)" + proof (rule islimptI) + fix A + assume "x \ A" "open A" + with as have "x \ A \ S" "open (A \ S)" + by (simp_all add: open_Int) + with * obtain y where "y \ T" "y \ A \ S" "y \ x" + by (rule islimptE) + hence "y \ S \ T" "y \ A \ y \ x" + by simp_all + thus "\y\(S \ T). y \ A \ y \ x" .. + qed + } + then show "x \ closure (S \ T)" using as + unfolding closure_def + by blast +qed + +lemma closure_complement: "closure(UNIV - S) = UNIV - interior(S)" +proof- + have "S = UNIV - (UNIV - S)" + by auto + thus ?thesis + unfolding closure_interior + by auto +qed + +lemma interior_complement: "interior(UNIV - S) = UNIV - closure(S)" + unfolding closure_interior + by blast + +subsection{* Frontier (aka boundary) *} + +definition "frontier S = closure S - interior S" + +lemma frontier_closed: "closed(frontier S)" + by (simp add: frontier_def closed_Diff) + +lemma frontier_closures: "frontier S = (closure S) \ (closure(UNIV - S))" + by (auto simp add: frontier_def interior_closure) + +lemma frontier_straddle: + fixes a :: "'a::metric_space" + shows "a \ frontier S \ (\e>0. (\x\S. dist a x < e) \ (\x. x \ S \ dist a x < e))" (is "?lhs \ ?rhs") +proof + assume "?lhs" + { fix e::real + assume "e > 0" + let ?rhse = "(\x\S. dist a x < e) \ (\x. x \ S \ dist a x < e)" + { assume "a\S" + have "\x\S. dist a x < e" using `e>0` `a\S` by(rule_tac x=a in bexI) auto + moreover have "\x. x \ S \ dist a x < e" using `?lhs` `a\S` + unfolding frontier_closures closure_def islimpt_def using `e>0` + by (auto, erule_tac x="ball a e" in allE, auto) + ultimately have ?rhse by auto + } + moreover + { assume "a\S" + hence ?rhse using `?lhs` + unfolding frontier_closures closure_def islimpt_def + using open_ball[of a e] `e > 0` + by (auto, erule_tac x = "ball a e" in allE, auto) (* FIXME: VERY slow! *) + } + ultimately have ?rhse by auto + } + thus ?rhs by auto +next + assume ?rhs + moreover + { fix T assume "a\S" and + as:"\e>0. (\x\S. dist a x < e) \ (\x. x \ S \ dist a x < e)" "a \ S" "a \ T" "open T" + from `open T` `a \ T` have "\e>0. ball a e \ T" unfolding open_contains_ball[of T] by auto + then obtain e where "e>0" "ball a e \ T" by auto + then obtain y where y:"y\S" "dist a y < e" using as(1) by auto + have "\y\S. y \ T \ y \ a" + using `dist a y < e` `ball a e \ T` unfolding ball_def using `y\S` `a\S` by auto + } + hence "a \ closure S" unfolding closure_def islimpt_def using `?rhs` by auto + moreover + { fix T assume "a \ T" "open T" "a\S" + then obtain e where "e>0" and balle: "ball a e \ T" unfolding open_contains_ball using `?rhs` by auto + obtain x where "x \ S" "dist a x < e" using `?rhs` using `e>0` by auto + hence "\y\UNIV - S. y \ T \ y \ a" using balle `a\S` unfolding ball_def by (rule_tac x=x in bexI)auto + } + hence "a islimpt (UNIV - S) \ a\S" unfolding islimpt_def by auto + ultimately show ?lhs unfolding frontier_closures using closure_def[of "UNIV - S"] by auto +qed + +lemma frontier_subset_closed: "closed S \ frontier S \ S" + by (metis frontier_def closure_closed Diff_subset) + +lemma frontier_empty: "frontier {} = {}" + by (simp add: frontier_def closure_empty) + +lemma frontier_subset_eq: "frontier S \ S \ closed S" +proof- + { assume "frontier S \ S" + hence "closure S \ S" using interior_subset unfolding frontier_def by auto + hence "closed S" using closure_subset_eq by auto + } + thus ?thesis using frontier_subset_closed[of S] by auto +qed + +lemma frontier_complement: "frontier(UNIV - S) = frontier S" + by (auto simp add: frontier_def closure_complement interior_complement) + +lemma frontier_disjoint_eq: "frontier S \ S = {} \ open S" + using frontier_complement frontier_subset_eq[of "UNIV - S"] + unfolding open_closed Compl_eq_Diff_UNIV by auto + +subsection{* Common nets and The "within" modifier for nets. *} + +definition + at_infinity :: "'a::real_normed_vector net" where + "at_infinity = Abs_net (range (\r. {x. r \ norm x}))" + +definition + indirection :: "'a::real_normed_vector \ 'a \ 'a net" (infixr "indirection" 70) where + "a indirection v = (at a) within {b. \c\0. b - a = scaleR c v}" + +text{* Prove That They are all nets. *} + +lemma Rep_net_at_infinity: + "Rep_net at_infinity = range (\r. {x. r \ norm x})" +unfolding at_infinity_def +apply (rule Abs_net_inverse') +apply (rule image_nonempty, simp) +apply (clarsimp, rename_tac r s) +apply (rule_tac x="max r s" in exI, auto) +done + +lemma within_UNIV: "net within UNIV = net" + by (simp add: Rep_net_inject [symmetric] Rep_net_within) + +subsection{* Identify Trivial limits, where we can't approach arbitrarily closely. *} + +definition + trivial_limit :: "'a net \ bool" where + "trivial_limit net \ {} \ Rep_net net" + +lemma trivial_limit_within: + shows "trivial_limit (at a within S) \ \ a islimpt S" +proof + assume "trivial_limit (at a within S)" + thus "\ a islimpt S" + unfolding trivial_limit_def + unfolding Rep_net_within Rep_net_at + unfolding islimpt_def + apply (clarsimp simp add: expand_set_eq) + apply (rename_tac T, rule_tac x=T in exI) + apply (clarsimp, drule_tac x=y in spec, simp) + done +next + assume "\ a islimpt S" + thus "trivial_limit (at a within S)" + unfolding trivial_limit_def + unfolding Rep_net_within Rep_net_at + unfolding islimpt_def + apply (clarsimp simp add: image_image) + apply (rule_tac x=T in image_eqI) + apply (auto simp add: expand_set_eq) + done +qed + +lemma trivial_limit_at_iff: "trivial_limit (at a) \ \ a islimpt UNIV" + using trivial_limit_within [of a UNIV] + by (simp add: within_UNIV) + +lemma trivial_limit_at: + fixes a :: "'a::perfect_space" + shows "\ trivial_limit (at a)" + by (simp add: trivial_limit_at_iff) + +lemma trivial_limit_at_infinity: + "\ trivial_limit (at_infinity :: ('a::{real_normed_vector,zero_neq_one}) net)" + (* FIXME: find a more appropriate type class *) + 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_sgn) + done + +lemma trivial_limit_sequentially: "\ trivial_limit sequentially" + by (auto simp add: trivial_limit_def Rep_net_sequentially) + +subsection{* Some property holds "sufficiently close" to the limit point. *} + +lemma eventually_at: (* FIXME: this replaces Limits.eventually_at *) + "eventually P (at a) \ (\d>0. \x. 0 < dist x a \ dist x a < d \ P x)" +unfolding eventually_at dist_nz by auto + +lemma eventually_at_infinity: + "eventually P at_infinity \ (\b. \x. norm x >= b \ P x)" +unfolding eventually_def Rep_net_at_infinity by auto + +lemma eventually_within: "eventually P (at a within S) \ + (\d>0. \x\S. 0 < dist x a \ dist x a < d \ P x)" +unfolding eventually_within eventually_at dist_nz by auto + +lemma eventually_within_le: "eventually P (at a within S) \ + (\d>0. \x\S. 0 < dist x a \ dist x a <= d \ P x)" (is "?lhs = ?rhs") +unfolding eventually_within +apply safe +apply (rule_tac x="d/2" in exI, simp) +apply (rule_tac x="d" in exI, simp) +done + +lemma eventually_happens: "eventually P net ==> trivial_limit net \ (\x. P x)" + unfolding eventually_def trivial_limit_def + using Rep_net_nonempty [of net] by auto + +lemma always_eventually: "(\x. P x) ==> eventually P net" + unfolding eventually_def trivial_limit_def + using Rep_net_nonempty [of net] by auto + +lemma trivial_limit_eventually: "trivial_limit net \ eventually P net" + unfolding trivial_limit_def eventually_def by auto + +lemma eventually_False: "eventually (\x. False) net \ trivial_limit net" + unfolding trivial_limit_def eventually_def by auto + +lemma trivial_limit_eq: "trivial_limit net \ (\P. eventually P net)" + apply (safe elim!: trivial_limit_eventually) + apply (simp add: eventually_False [symmetric]) + done + +text{* Combining theorems for "eventually" *} + +lemma eventually_conjI: + "\eventually (\x. P x) net; eventually (\x. Q x) net\ + \ eventually (\x. P x \ Q x) net" +by (rule eventually_conj) + +lemma eventually_rev_mono: + "eventually P net \ (\x. P x \ Q x) \ eventually Q net" +using eventually_mono [of P Q] by fast + +lemma eventually_and: " eventually (\x. P x \ Q x) net \ eventually P net \ eventually Q net" + by (auto intro!: eventually_conjI elim: eventually_rev_mono) + +lemma eventually_false: "eventually (\x. False) net \ trivial_limit net" + by (auto simp add: eventually_False) + +lemma not_eventually: "(\x. \ P x ) \ ~(trivial_limit net) ==> ~(eventually (\x. P x) net)" + by (simp add: eventually_False) + +subsection{* Limits, defined as vacuously true when the limit is trivial. *} + + text{* Notation Lim to avoid collition with lim defined in analysis *} +definition + Lim :: "'a net \ ('a \ 'b::t2_space) \ 'b" where + "Lim net f = (THE l. (f ---> l) net)" + +lemma Lim: + "(f ---> l) net \ + trivial_limit net \ + (\e>0. eventually (\x. dist (f x) l < e) net)" + unfolding tendsto_iff trivial_limit_eq by auto + + +text{* Show that they yield usual definitions in the various cases. *} + +lemma Lim_within_le: "(f ---> l)(at a within S) \ + (\e>0. \d>0. \x\S. 0 < dist x a \ dist x a <= d \ dist (f x) l < e)" + by (auto simp add: tendsto_iff eventually_within_le) + +lemma Lim_within: "(f ---> l) (at a within S) \ + (\e >0. \d>0. \x \ S. 0 < dist x a \ dist x a < d \ dist (f x) l < e)" + by (auto simp add: tendsto_iff eventually_within) + +lemma Lim_at: "(f ---> l) (at a) \ + (\e >0. \d>0. \x. 0 < dist x a \ dist x a < d \ dist (f x) l < e)" + by (auto simp add: tendsto_iff eventually_at) + +lemma Lim_at_iff_LIM: "(f ---> l) (at a) \ f -- a --> l" + unfolding Lim_at LIM_def by (simp only: zero_less_dist_iff) + +lemma Lim_at_infinity: + "(f ---> l) at_infinity \ (\e>0. \b. \x. norm x >= b \ dist (f x) l < e)" + by (auto simp add: tendsto_iff eventually_at_infinity) + +lemma Lim_sequentially: + "(S ---> l) sequentially \ + (\e>0. \N. \n\N. dist (S n) l < e)" + by (auto simp add: tendsto_iff eventually_sequentially) + +lemma Lim_sequentially_iff_LIMSEQ: "(S ---> l) sequentially \ S ----> l" + unfolding Lim_sequentially LIMSEQ_def .. + +lemma Lim_eventually: "eventually (\x. f x = l) net \ (f ---> l) net" + by (rule topological_tendstoI, auto elim: eventually_rev_mono) + +text{* The expected monotonicity property. *} + +lemma Lim_within_empty: "(f ---> l) (net within {})" + unfolding tendsto_def Limits.eventually_within by simp + +lemma Lim_within_subset: "(f ---> l) (net within S) \ T \ S \ (f ---> l) (net within T)" + unfolding tendsto_def Limits.eventually_within + by (auto elim!: eventually_elim1) + +lemma Lim_Un: assumes "(f ---> l) (net within S)" "(f ---> l) (net within T)" + shows "(f ---> l) (net within (S \ T))" + using assms unfolding tendsto_def Limits.eventually_within + apply clarify + apply (drule spec, drule (1) mp, drule (1) mp) + apply (drule spec, drule (1) mp, drule (1) mp) + apply (auto elim: eventually_elim2) + done + +lemma Lim_Un_univ: + "(f ---> l) (net within S) \ (f ---> l) (net within T) \ S \ T = UNIV + ==> (f ---> l) net" + by (metis Lim_Un within_UNIV) + +text{* Interrelations between restricted and unrestricted limits. *} + +lemma Lim_at_within: "(f ---> l) net ==> (f ---> l)(net within S)" + (* FIXME: rename *) + unfolding tendsto_def Limits.eventually_within + apply (clarify, drule spec, drule (1) mp, drule (1) mp) + by (auto elim!: eventually_elim1) + +lemma Lim_within_open: + fixes f :: "'a::topological_space \ 'b::topological_space" + assumes"a \ S" "open S" + shows "(f ---> l)(at a within S) \ (f ---> l)(at a)" (is "?lhs \ ?rhs") +proof + assume ?lhs + { fix A assume "open A" "l \ A" + with `?lhs` have "eventually (\x. f x \ A) (at a within S)" + by (rule topological_tendstoD) + hence "eventually (\x. x \ S \ f x \ A) (at a)" + unfolding Limits.eventually_within . + then obtain T where "open T" "a \ T" "\x\T. x \ a \ x \ S \ f x \ A" + unfolding eventually_at_topological by fast + hence "open (T \ S)" "a \ T \ S" "\x\(T \ S). x \ a \ f x \ A" + using assms by auto + hence "\T. open T \ a \ T \ (\x\T. x \ a \ f x \ A)" + by fast + hence "eventually (\x. f x \ A) (at a)" + unfolding eventually_at_topological . + } + thus ?rhs by (rule topological_tendstoI) +next + assume ?rhs + thus ?lhs by (rule Lim_at_within) +qed + +text{* Another limit point characterization. *} + +lemma islimpt_sequential: + fixes x :: "'a::metric_space" (* FIXME: generalize to topological_space *) + shows "x islimpt S \ (\f. (\n::nat. f n \ S -{x}) \ (f ---> x) sequentially)" + (is "?lhs = ?rhs") +proof + assume ?lhs + then obtain f where f:"\y. y>0 \ f y \ S \ f y \ x \ dist (f y) x < y" + unfolding islimpt_approachable using choice[of "\e y. e>0 \ y\S \ y\x \ dist y x < e"] by auto + { fix n::nat + have "f (inverse (real n + 1)) \ S - {x}" using f by auto + } + moreover + { fix e::real assume "e>0" + hence "\N::nat. inverse (real (N + 1)) < e" using real_arch_inv[of e] apply (auto simp add: Suc_pred') apply(rule_tac x="n - 1" in exI) by auto + 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) + moreover have "\n\N. dist (f (inverse (real n + 1))) x < (inverse (real n + 1))" using f `e>0` by auto + ultimately have "\N::nat. \n\N. dist (f (inverse (real n + 1))) x < e" apply(rule_tac x=N in exI) apply auto apply(erule_tac x=n in allE)+ by auto + } + hence " ((\n. f (inverse (real n + 1))) ---> x) sequentially" + unfolding Lim_sequentially using f by auto + ultimately show ?rhs apply (rule_tac x="(\n::nat. f (inverse (real n + 1)))" in exI) by auto +next + assume ?rhs + then obtain f::"nat\'a" where f:"(\n. f n \ S - {x})" "(\e>0. \N. \n\N. dist (f n) x < e)" unfolding Lim_sequentially by auto + { fix e::real assume "e>0" + then obtain N where "dist (f N) x < e" using f(2) by auto + moreover have "f N\S" "f N \ x" using f(1) by auto + ultimately have "\x'\S. x' \ x \ dist x' x < e" by auto + } + thus ?lhs unfolding islimpt_approachable by auto +qed + +text{* Basic arithmetical combining theorems for limits. *} + +lemma Lim_linear: + assumes "(f ---> l) net" "bounded_linear h" + shows "((\x. h (f x)) ---> h l) net" +using `bounded_linear h` `(f ---> l) net` +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 \ 'b::real_normed_vector" + 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" + shows "(f ---> l) net ==> ((\x. -(f x)) ---> -l) net" + by (rule tendsto_minus) + +lemma Lim_add: fixes f :: "'a \ 'b::real_normed_vector" shows + "(f ---> l) net \ (g ---> m) net \ ((\x. f(x) + g(x)) ---> l + m) net" + by (rule tendsto_add) + +lemma Lim_sub: + fixes f :: "'a \ 'b::real_normed_vector" + shows "(f ---> l) net \ (g ---> m) net \ ((\x. f(x) - g(x)) ---> l - m) net" + by (rule tendsto_diff) + +lemma Lim_null: + fixes f :: "'a \ 'b::real_normed_vector" + shows "(f ---> l) net \ ((\x. f(x) - l) ---> 0) net" by (simp add: Lim dist_norm) + +lemma Lim_null_norm: + fixes f :: "'a \ 'b::real_normed_vector" + 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" "(g ---> 0) net" + shows "(f ---> 0) net" +proof(simp add: tendsto_iff, rule+) + fix e::real assume "0 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 (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: + fixes f :: "'a \ 'b::metric_space ^ 'n::finite" + shows "(f ---> l) net \ ((\a. f a $i) ---> l$i) net" + unfolding tendsto_iff + 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" + fixes g :: "'a \ 'c::real_normed_vector" + assumes "eventually (\n. norm(f n) <= norm(g n)) net" "(g ---> 0) net" + shows "(f ---> 0) net" +proof (rule tendstoI) + fix e::real assume "e>0" + { fix x + assume "norm (f x) \ norm (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) \ norm (g x)" "\x. dist (g x) 0 < e" net] + using eventually_mono[of "\x. norm (f x) \ norm (g x) \ dist (g x) 0 < e" "\x. dist (f x) 0 < e" net] + using assms `e>0` unfolding tendsto_iff by blast +qed + +text{* Deducing things about the limit from the elements. *} + +lemma Lim_in_closed_set: + assumes "closed S" "eventually (\x. f(x) \ S) net" "\(trivial_limit net)" "(f ---> l) net" + shows "l \ S" +proof (rule ccontr) + assume "l \ S" + with `closed S` have "open (- S)" "l \ - S" + by (simp_all add: open_Compl) + with assms(4) have "eventually (\x. f x \ - S) net" + by (rule topological_tendstoD) + with assms(2) have "eventually (\x. False) net" + by (rule eventually_elim2) simp + with assms(3) show "False" + by (simp add: eventually_False) +qed + +text{* Need to prove closed(cball(x,e)) before deducing this as a corollary. *} + +lemma Lim_dist_ubound: + assumes "\(trivial_limit net)" "(f ---> l) net" "eventually (\x. dist a (f x) <= e) net" + shows "dist a l <= e" +proof (rule ccontr) + assume "\ dist a l \ e" + then have "0 < dist a l - e" by simp + with assms(2) have "eventually (\x. dist (f x) l < dist a l - e) net" + by (rule tendstoD) + with assms(3) have "eventually (\x. dist a (f x) \ e \ dist (f x) l < dist a l - e) net" + by (rule eventually_conjI) + then obtain w where "dist a (f w) \ e" "dist (f w) l < dist a l - e" + using assms(1) eventually_happens by auto + hence "dist a (f w) + dist (f w) l < e + (dist a l - e)" + by (rule add_le_less_mono) + hence "dist a (f w) + dist (f w) l < dist a l" + by simp + also have "\ \ dist a (f w) + dist (f w) l" + by (rule dist_triangle) + finally show False by simp +qed + +lemma Lim_norm_ubound: + fixes f :: "'a \ 'b::real_normed_vector" + assumes "\(trivial_limit net)" "(f ---> l) net" "eventually (\x. norm(f x) <= e) net" + shows "norm(l) <= e" +proof (rule ccontr) + assume "\ norm l \ e" + then have "0 < norm l - e" by simp + with assms(2) have "eventually (\x. dist (f x) l < norm l - e) net" + by (rule tendstoD) + with assms(3) have "eventually (\x. norm (f x) \ e \ dist (f x) l < norm l - e) net" + by (rule eventually_conjI) + then obtain w where "norm (f w) \ e" "dist (f w) l < norm l - e" + using assms(1) eventually_happens by auto + hence "norm (f w - l) < norm l - e" "norm (f w) \ e" by (simp_all add: dist_norm) + hence "norm (f w - l) + norm (f w) < norm l" by simp + hence "norm (f w - l - f w) < norm l" by (rule le_less_trans [OF norm_triangle_ineq4]) + thus False using `\ norm l \ e` by simp +qed + +lemma Lim_norm_lbound: + fixes f :: "'a \ 'b::real_normed_vector" + assumes "\ (trivial_limit net)" "(f ---> l) net" "eventually (\x. e <= norm(f x)) net" + shows "e \ norm l" +proof (rule ccontr) + assume "\ e \ norm l" + then have "0 < e - norm l" by simp + with assms(2) have "eventually (\x. dist (f x) l < e - norm l) net" + by (rule tendstoD) + with assms(3) have "eventually (\x. e \ norm (f x) \ dist (f x) l < e - norm l) net" + by (rule eventually_conjI) + then obtain w where "e \ norm (f w)" "dist (f w) l < e - norm l" + using assms(1) eventually_happens by auto + hence "norm (f w - l) + norm l < e" "e \ norm (f w)" by (simp_all add: dist_norm) + hence "norm (f w - l) + norm l < norm (f w)" by (rule less_le_trans) + hence "norm (f w - l + l) < norm (f w)" by (rule le_less_trans [OF norm_triangle_ineq]) + thus False by simp +qed + +text{* Uniqueness of the limit, when nontrivial. *} + +lemma Lim_unique: + fixes f :: "'a \ 'b::t2_space" + assumes "\ trivial_limit net" "(f ---> l) net" "(f ---> l') net" + shows "l = l'" +proof (rule ccontr) + assume "l \ l'" + obtain U V where "open U" "open V" "l \ U" "l' \ V" "U \ V = {}" + using hausdorff [OF `l \ l'`] by fast + have "eventually (\x. f x \ U) net" + using `(f ---> l) net` `open U` `l \ U` by (rule topological_tendstoD) + moreover + have "eventually (\x. f x \ V) net" + using `(f ---> l') net` `open V` `l' \ V` by (rule topological_tendstoD) + ultimately + have "eventually (\x. False) net" + proof (rule eventually_elim2) + fix x + assume "f x \ U" "f x \ V" + hence "f x \ U \ V" by simp + with `U \ V = {}` show "False" by simp + qed + with `\ trivial_limit net` show "False" + by (simp add: eventually_False) +qed + +lemma tendsto_Lim: + fixes f :: "'a \ 'b::t2_space" + shows "~(trivial_limit net) \ (f ---> l) net ==> Lim net f = l" + unfolding Lim_def using Lim_unique[of net f] by auto + +text{* Limit under bilinear function *} + +lemma Lim_bilinear: + assumes "(f ---> l) net" and "(g ---> m) net" and "bounded_bilinear h" + shows "((\x. h (f x) (g x)) ---> (h l m)) net" +using `bounded_bilinear h` `(f ---> l) net` `(g ---> m) net` +by (rule bounded_bilinear.tendsto) + +text{* These are special for limits out of the same vector space. *} + +lemma Lim_within_id: "(id ---> a) (at a within s)" + unfolding tendsto_def Limits.eventually_within eventually_at_topological + by auto + +lemma Lim_at_id: "(id ---> a) (at a)" +apply (subst within_UNIV[symmetric]) by (simp add: Lim_within_id) + +lemma Lim_at_zero: + fixes a :: "'a::real_normed_vector" + fixes l :: "'b::topological_space" + shows "(f ---> l) (at a) \ ((\x. f(a + x)) ---> l) (at 0)" (is "?lhs = ?rhs") +proof + assume "?lhs" + { fix S assume "open S" "l \ S" + with `?lhs` have "eventually (\x. f x \ S) (at a)" + by (rule topological_tendstoD) + then obtain d where d: "d>0" "\x. x \ a \ dist x a < d \ f x \ S" + unfolding Limits.eventually_at by fast + { fix x::"'a" assume "x \ 0 \ dist x 0 < d" + hence "f (a + x) \ S" using d + apply(erule_tac x="x+a" in allE) + by(auto simp add: comm_monoid_add.mult_commute dist_norm dist_commute) + } + hence "\d>0. \x. x \ 0 \ dist x 0 < d \ f (a + x) \ S" + using d(1) by auto + hence "eventually (\x. f (a + x) \ S) (at 0)" + unfolding Limits.eventually_at . + } + thus "?rhs" by (rule topological_tendstoI) +next + assume "?rhs" + { fix S assume "open S" "l \ S" + with `?rhs` have "eventually (\x. f (a + x) \ S) (at 0)" + by (rule topological_tendstoD) + then obtain d where d: "d>0" "\x. x \ 0 \ dist x 0 < d \ f (a + x) \ S" + unfolding Limits.eventually_at by fast + { fix x::"'a" assume "x \ a \ dist x a < d" + hence "f x \ S" using d apply(erule_tac x="x-a" in allE) + by(auto simp add: comm_monoid_add.mult_commute dist_norm dist_commute) + } + hence "\d>0. \x. x \ a \ dist x a < d \ f x \ S" using d(1) by auto + hence "eventually (\x. f x \ S) (at a)" unfolding Limits.eventually_at . + } + thus "?lhs" by (rule topological_tendstoI) +qed + +text{* It's also sometimes useful to extract the limit point from the net. *} + +definition + netlimit :: "'a::t2_space net \ 'a" where + "netlimit net = (SOME a. ((\x. x) ---> a) net)" + +lemma netlimit_within: + assumes "\ trivial_limit (at a within S)" + shows "netlimit (at a within S) = a" +unfolding netlimit_def +apply (rule some_equality) +apply (rule Lim_at_within) +apply (rule Lim_ident_at) +apply (erule Lim_unique [OF assms]) +apply (rule Lim_at_within) +apply (rule Lim_ident_at) +done + +lemma netlimit_at: + fixes a :: "'a::perfect_space" + shows "netlimit (at a) = a" + apply (subst within_UNIV[symmetric]) + using netlimit_within[of a UNIV] + by (simp add: trivial_limit_at within_UNIV) + +text{* Transformation of limit. *} + +lemma Lim_transform: + fixes f g :: "'a::type \ 'b::real_normed_vector" + assumes "((\x. f x - g x) ---> 0) net" "(f ---> l) net" + shows "(g ---> l) net" +proof- + from assms have "((\x. f x - g x - f x) ---> 0 - l) net" using Lim_sub[of "\x. f x - g x" 0 net f l] by auto + thus "?thesis" using Lim_neg [of "\ x. - g x" "-l" net] by auto +qed + +lemma Lim_transform_eventually: + "eventually (\x. f x = g x) net \ (f ---> l) net ==> (g ---> l) net" + apply (rule topological_tendstoI) + apply (drule (2) topological_tendstoD) + apply (erule (1) eventually_elim2, simp) + done + +lemma Lim_transform_within: + fixes l :: "'b::metric_space" (* TODO: generalize *) + assumes "0 < d" "(\x'\S. 0 < dist x' x \ dist x' x < d \ f x' = g x')" + "(f ---> l) (at x within S)" + shows "(g ---> l) (at x within S)" + using assms(1,3) unfolding Lim_within + apply - + apply (clarify, rename_tac e) + apply (drule_tac x=e in spec, clarsimp, rename_tac r) + apply (rule_tac x="min d r" in exI, clarsimp, rename_tac y) + apply (drule_tac x=y in bspec, assumption, clarsimp) + apply (simp add: assms(2)) + done + +lemma Lim_transform_at: + fixes l :: "'b::metric_space" (* TODO: generalize *) + shows "0 < d \ (\x'. 0 < dist x' x \ dist x' x < d \ f x' = g x') \ + (f ---> l) (at x) ==> (g ---> l) (at x)" + apply (subst within_UNIV[symmetric]) + using Lim_transform_within[of d UNIV x f g l] + by (auto simp add: within_UNIV) + +text{* Common case assuming being away from some crucial point like 0. *} + +lemma Lim_transform_away_within: + fixes a b :: "'a::metric_space" + fixes l :: "'b::metric_space" (* TODO: generalize *) + assumes "a\b" "\x\ S. x \ a \ x \ b \ f x = g x" + and "(f ---> l) (at a within S)" + shows "(g ---> l) (at a within S)" +proof- + have "\x'\S. 0 < dist x' a \ dist x' a < dist a b \ f x' = g x'" using assms(2) + apply auto apply(erule_tac x=x' in ballE) by (auto simp add: dist_commute) + thus ?thesis using Lim_transform_within[of "dist a b" S a f g l] using assms(1,3) unfolding dist_nz by auto +qed + +lemma Lim_transform_away_at: + fixes a b :: "'a::metric_space" + fixes l :: "'b::metric_space" (* TODO: generalize *) + assumes ab: "a\b" and fg: "\x. x \ a \ x \ b \ f x = g x" + and fl: "(f ---> l) (at a)" + shows "(g ---> l) (at a)" + using Lim_transform_away_within[OF ab, of UNIV f g l] fg fl + by (auto simp add: within_UNIV) + +text{* Alternatively, within an open set. *} + +lemma Lim_transform_within_open: + fixes a :: "'a::metric_space" + fixes l :: "'b::metric_space" (* TODO: generalize *) + assumes "open S" "a \ S" "\x\S. x \ a \ f x = g x" "(f ---> l) (at a)" + shows "(g ---> l) (at a)" +proof- + from assms(1,2) obtain e::real where "e>0" and e:"ball a e \ S" unfolding open_contains_ball by auto + hence "\x'. 0 < dist x' a \ dist x' a < e \ f x' = g x'" using assms(3) + unfolding ball_def subset_eq apply auto apply(erule_tac x=x' in allE) apply(erule_tac x=x' in ballE) by(auto simp add: dist_commute) + thus ?thesis using Lim_transform_at[of e a f g l] `e>0` assms(4) by auto +qed + +text{* A congruence rule allowing us to transform limits assuming not at point. *} + +(* FIXME: Only one congruence rule for tendsto can be used at a time! *) + +lemma Lim_cong_within[cong add]: + fixes a :: "'a::metric_space" + fixes l :: "'b::metric_space" (* TODO: generalize *) + shows "(\x. x \ a \ f x = g x) ==> ((\x. f x) ---> l) (at a within S) \ ((g ---> l) (at a within S))" + by (simp add: Lim_within dist_nz[symmetric]) + +lemma Lim_cong_at[cong add]: + fixes a :: "'a::metric_space" + fixes l :: "'b::metric_space" (* TODO: generalize *) + shows "(\x. x \ a ==> f x = g x) ==> (((\x. f x) ---> l) (at a) \ ((g ---> l) (at a)))" + by (simp add: Lim_at dist_nz[symmetric]) + +text{* Useful lemmas on closure and set of possible sequential limits.*} + +lemma closure_sequential: + fixes l :: "'a::metric_space" (* TODO: generalize *) + shows "l \ closure S \ (\x. (\n. x n \ S) \ (x ---> l) sequentially)" (is "?lhs = ?rhs") +proof + assume "?lhs" moreover + { assume "l \ S" + hence "?rhs" using Lim_const[of l sequentially] by auto + } moreover + { assume "l islimpt S" + hence "?rhs" unfolding islimpt_sequential by auto + } ultimately + show "?rhs" unfolding closure_def by auto +next + assume "?rhs" + thus "?lhs" unfolding closure_def unfolding islimpt_sequential by auto +qed + +lemma closed_sequential_limits: + fixes S :: "'a::metric_space set" + shows "closed S \ (\x l. (\n. x n \ S) \ (x ---> l) sequentially \ l \ S)" + unfolding closed_limpt + using closure_sequential [where 'a='a] closure_closed [where 'a='a] closed_limpt [where 'a='a] islimpt_sequential [where 'a='a] mem_delete [where 'a='a] + by metis + +lemma closure_approachable: + fixes S :: "'a::metric_space set" + shows "x \ closure S \ (\e>0. \y\S. dist y x < e)" + apply (auto simp add: closure_def islimpt_approachable) + by (metis dist_self) + +lemma closed_approachable: + fixes S :: "'a::metric_space set" + shows "closed S ==> (\e>0. \y\S. dist y x < e) \ x \ S" + by (metis closure_closed closure_approachable) + +text{* Some other lemmas about sequences. *} + +lemma seq_offset: + fixes l :: "'a::metric_space" (* TODO: generalize *) + shows "(f ---> l) sequentially ==> ((\i. f( i + k)) ---> l) sequentially" + apply (auto simp add: Lim_sequentially) + by (metis trans_le_add1 ) + +lemma seq_offset_neg: + "(f ---> l) sequentially ==> ((\i. f(i - k)) ---> l) sequentially" + apply (rule topological_tendstoI) + apply (drule (2) topological_tendstoD) + apply (simp only: eventually_sequentially) + apply (subgoal_tac "\N k (n::nat). N + k <= n ==> N <= n - k") + apply metis + by arith + +lemma seq_offset_rev: + "((\i. f(i + k)) ---> l) sequentially ==> (f ---> l) sequentially" + apply (rule topological_tendstoI) + apply (drule (2) topological_tendstoD) + apply (simp only: eventually_sequentially) + apply (subgoal_tac "\N k (n::nat). N + k <= n ==> N <= n - k \ (n - k) + k = n") + by metis arith + +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 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 by simp +qed + +text{* More properties of closed balls. *} + +lemma closed_cball: "closed (cball x e)" +unfolding cball_def closed_def +unfolding Collect_neg_eq [symmetric] not_le +apply (clarsimp simp add: open_dist, rename_tac y) +apply (rule_tac x="dist x y - e" in exI, clarsimp) +apply (rename_tac x') +apply (cut_tac x=x and y=x' and z=y in dist_triangle) +apply simp +done + +lemma open_contains_cball: "open S \ (\x\S. \e>0. cball x e \ S)" +proof- + { fix x and e::real assume "x\S" "e>0" "ball x e \ S" + hence "\d>0. cball x d \ S" unfolding subset_eq by (rule_tac x="e/2" in exI, auto) + } moreover + { fix x and e::real assume "x\S" "e>0" "cball x e \ S" + hence "\d>0. ball x d \ S" unfolding subset_eq apply(rule_tac x="e/2" in exI) by auto + } ultimately + show ?thesis unfolding open_contains_ball by auto +qed + +lemma open_contains_cball_eq: "open S ==> (\x. x \ S \ (\e>0. cball x e \ S))" + by (metis open_contains_cball subset_eq order_less_imp_le centre_in_cball mem_def) + +lemma mem_interior_cball: "x \ interior S \ (\e>0. cball x e \ S)" + apply (simp add: interior_def, safe) + apply (force simp add: open_contains_cball) + apply (rule_tac x="ball x e" in exI) + apply (simp add: open_ball centre_in_ball subset_trans [OF ball_subset_cball]) + done + +lemma islimpt_ball: + fixes x y :: "'a::{real_normed_vector,perfect_space}" + shows "y islimpt ball x e \ 0 < e \ y \ cball x e" (is "?lhs = ?rhs") +proof + assume "?lhs" + { assume "e \ 0" + hence *:"ball x e = {}" using ball_eq_empty[of x e] by auto + have False using `?lhs` unfolding * using islimpt_EMPTY[of y] by auto + } + hence "e > 0" by (metis not_less) + moreover + have "y \ cball x e" using closed_cball[of x e] islimpt_subset[of y "ball x e" "cball x e"] ball_subset_cball[of x e] `?lhs` unfolding closed_limpt by auto + ultimately show "?rhs" by auto +next + assume "?rhs" hence "e>0" by auto + { fix d::real assume "d>0" + have "\x'\ball x e. x' \ y \ dist x' y < d" + proof(cases "d \ dist x y") + case True thus "\x'\ball x e. x' \ y \ dist x' y < d" + proof(cases "x=y") + case True hence False using `d \ dist x y` `d>0` by auto + thus "\x'\ball x e. x' \ y \ dist x' y < d" by auto + next + case False + + have "dist x (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) + = norm (x - y + (d / (2 * norm (y - x))) *\<^sub>R (y - x))" + unfolding mem_cball mem_ball dist_norm diff_diff_eq2 diff_add_eq[THEN sym] by auto + 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) + 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 + also have "\ \ e - d/2" using `d \ dist x y` and `d>0` and `?rhs` by(auto simp add: dist_norm) + finally have "y - (d / (2 * dist y x)) *\<^sub>R (y - x) \ ball x e" using `d>0` by auto + + moreover + + 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 + 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 + qed + next + case False hence "d > dist x y" by auto + show "\x'\ball x e. x' \ y \ dist x' y < d" + proof(cases "x=y") + case True + obtain z where **: "z \ y" "dist z y < min e d" + using perfect_choose_dist[of "min e d" y] + using `d > 0` `e>0` by auto + show "\x'\ball x e. x' \ y \ dist x' y < d" + unfolding `x = y` + using `z \ y` ** + by (rule_tac x=z in bexI, auto simp add: dist_commute) + next + case False thus "\x'\ball x e. x' \ y \ dist x' y < d" + using `d>0` `d > dist x y` `?rhs` by(rule_tac x=x in bexI, auto) + qed + qed } + thus "?lhs" unfolding mem_cball islimpt_approachable mem_ball by auto +qed + +lemma closure_ball_lemma: + fixes x y :: "'a::real_normed_vector" + assumes "x \ y" shows "y islimpt ball x (dist x y)" +proof (rule islimptI) + fix T assume "y \ T" "open T" + then obtain r where "0 < r" "\z. dist z y < r \ z \ T" + unfolding open_dist by fast + (* choose point between x and y, within distance r of y. *) + def k \ "min 1 (r / (2 * dist x y))" + def z \ "y + scaleR k (x - y)" + have z_def2: "z = x + scaleR (1 - k) (y - x)" + 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 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_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) + apply (simp add: k_def divide_pos_pos zero_less_dist_iff `0 < r` `x \ y`) + apply (simp add: zero_less_dist_iff `x \ y`) + done + hence "z \ ball x (dist x y)" by simp + have "z \ y" + unfolding z_def k_def using `x \ y` `0 < r` + by (simp add: min_def) + show "\z\ball x (dist x y). z \ T \ z \ y" + using `z \ ball x (dist x y)` `z \ T` `z \ y` + by fast +qed + +lemma closure_ball: + fixes x :: "'a::real_normed_vector" + shows "0 < e \ closure (ball x e) = cball x e" +apply (rule equalityI) +apply (rule closure_minimal) +apply (rule ball_subset_cball) +apply (rule closed_cball) +apply (rule subsetI, rename_tac y) +apply (simp add: le_less [where 'a=real]) +apply (erule disjE) +apply (rule subsetD [OF closure_subset], simp) +apply (simp add: closure_def) +apply clarify +apply (rule closure_ball_lemma) +apply (simp add: zero_less_dist_iff) +done + +(* In a trivial vector space, this fails for e = 0. *) +lemma interior_cball: + 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 + { fix y assume "y \ cball x e" + hence False unfolding mem_cball using dist_nz[of x y] cs by auto } + hence "cball x e = {}" by auto + hence "interior (cball x e) = {}" using interior_empty by auto + ultimately show ?thesis by blast +next + case True note cs = this + have "ball x e \ cball x e" using ball_subset_cball by auto moreover + { 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_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") + case True + hence "e>0" using xa_y[unfolded dist_nz] xa_cball[unfolded mem_cball] by (auto simp add: dist_commute) + thus "y \ ball x e" using `x = y ` by simp + next + case False + 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) *\<^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) *\<^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 + qed } + hence "\S \ cball x e. open S \ S \ ball x e" by auto + ultimately show ?thesis using interior_unique[of "ball x e" "cball x e"] using open_ball[of x e] by auto +qed + +lemma frontier_ball: + 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 :: "'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) + by arith + +lemma cball_eq_empty: "(cball x e = {}) \ e < 0" + apply (simp add: expand_set_eq not_le) + by (metis zero_le_dist dist_self order_less_le_trans) +lemma cball_empty: "e < 0 ==> cball x e = {}" by (simp add: cball_eq_empty) + +lemma cball_eq_sing: + fixes x :: "'a::perfect_space" + shows "(cball x e = {x}) \ e = 0" +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 :: "'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. *} + +lemma eventually_within_interior: + assumes "x \ interior S" + shows "eventually P (at x within S) \ eventually P (at x)" (is "?lhs = ?rhs") +proof- + from assms obtain T where T: "open T" "x \ T" "T \ S" + unfolding interior_def by fast + { assume "?lhs" + then obtain A where "open A" "x \ A" "\y\A. y \ x \ y \ S \ P y" + unfolding Limits.eventually_within Limits.eventually_at_topological + by auto + with T have "open (A \ T)" "x \ A \ T" "\y\(A \ T). y \ x \ P y" + by auto + then have "?rhs" + unfolding Limits.eventually_at_topological by auto + } moreover + { assume "?rhs" hence "?lhs" + unfolding Limits.eventually_within + by (auto elim: eventually_elim1) + } ultimately + show "?thesis" .. +qed + +lemma lim_within_interior: + "x \ interior S \ (f ---> l) (at x within S) \ (f ---> l) (at x)" + unfolding tendsto_def by (simp add: eventually_within_interior) + +lemma netlimit_within_interior: + fixes x :: "'a::{perfect_space, real_normed_vector}" + (* FIXME: generalize to perfect_space *) + assumes "x \ interior S" + shows "netlimit(at x within S) = x" (is "?lhs = ?rhs") +proof- + from assms obtain e::real where e:"e>0" "ball x e \ S" using open_interior[of S] unfolding open_contains_ball using interior_subset[of S] by auto + hence "\ trivial_limit (at x within S)" using islimpt_subset[of x "ball x e" S] unfolding trivial_limit_within islimpt_ball centre_in_cball by auto + thus ?thesis using netlimit_within by auto +qed + +subsection{* Boundedness. *} + + (* FIXME: This has to be unified with BSEQ!! *) +definition + bounded :: "'a::metric_space set \ bool" where + "bounded S \ (\x e. \y\S. dist x y \ e)" + +lemma bounded_any_center: "bounded S \ (\e. \y\S. dist a y \ e)" +unfolding bounded_def +apply safe +apply (rule_tac x="dist a x + e" in exI, clarify) +apply (drule (1) bspec) +apply (erule order_trans [OF dist_triangle add_left_mono]) +apply auto +done + +lemma bounded_iff: "bounded S \ (\a. \x\S. norm x \ a)" +unfolding bounded_any_center [where a=0] +by (simp add: dist_norm) + +lemma bounded_empty[simp]: "bounded {}" by (simp add: bounded_def) +lemma bounded_subset: "bounded T \ S \ T ==> bounded S" + by (metis bounded_def subset_eq) + +lemma bounded_interior[intro]: "bounded S ==> bounded(interior S)" + by (metis bounded_subset interior_subset) + +lemma bounded_closure[intro]: assumes "bounded S" shows "bounded(closure S)" +proof- + from assms obtain x and a where a: "\y\S. dist x y \ a" unfolding bounded_def by auto + { fix y assume "y \ closure S" + then obtain f where f: "\n. f n \ S" "(f ---> y) sequentially" + unfolding closure_sequential by auto + have "\n. f n \ S \ dist x (f n) \ a" using a by simp + hence "eventually (\n. dist x (f n) \ a) sequentially" + by (rule eventually_mono, simp add: f(1)) + have "dist x y \ a" + apply (rule Lim_dist_ubound [of sequentially f]) + apply (rule trivial_limit_sequentially) + apply (rule f(2)) + apply fact + done + } + thus ?thesis unfolding bounded_def by auto +qed + +lemma bounded_cball[simp,intro]: "bounded (cball x e)" + apply (simp add: bounded_def) + apply (rule_tac x=x in exI) + apply (rule_tac x=e in exI) + apply auto + done + +lemma bounded_ball[simp,intro]: "bounded(ball x e)" + by (metis ball_subset_cball bounded_cball bounded_subset) + +lemma finite_imp_bounded[intro]: assumes "finite S" shows "bounded S" +proof- + { fix a F assume as:"bounded F" + then obtain x e where "\y\F. dist x y \ e" unfolding bounded_def by auto + hence "\y\(insert a F). dist x y \ max e (dist x a)" by auto + hence "bounded (insert a F)" unfolding bounded_def by (intro exI) + } + thus ?thesis using finite_induct[of S bounded] using bounded_empty assms by auto +qed + +lemma bounded_Un[simp]: "bounded (S \ T) \ bounded S \ bounded T" + apply (auto simp add: bounded_def) + apply (rename_tac x y r s) + apply (rule_tac x=x in exI) + apply (rule_tac x="max r (dist x y + s)" in exI) + apply (rule ballI, rename_tac z, safe) + apply (drule (1) bspec, simp) + apply (drule (1) bspec) + apply (rule min_max.le_supI2) + apply (erule order_trans [OF dist_triangle add_left_mono]) + done + +lemma bounded_Union[intro]: "finite F \ (\S\F. bounded S) \ bounded(\F)" + by (induct rule: finite_induct[of F], auto) + +lemma bounded_pos: "bounded S \ (\b>0. \x\ S. norm x <= b)" + apply (simp add: bounded_iff) + apply (subgoal_tac "\x (y::real). 0 < 1 + abs y \ (x <= y \ x <= 1 + abs y)") + by metis arith + +lemma bounded_Int[intro]: "bounded S \ bounded T \ bounded (S \ T)" + by (metis Int_lower1 Int_lower2 bounded_subset) + +lemma bounded_diff[intro]: "bounded S ==> bounded (S - T)" +apply (metis Diff_subset bounded_subset) +done + +lemma bounded_insert[intro]:"bounded(insert x S) \ bounded S" + by (metis Diff_cancel Un_empty_right Un_insert_right bounded_Un bounded_subset finite.emptyI finite_imp_bounded infinite_remove subset_insertI) + +lemma not_bounded_UNIV[simp, intro]: + "\ bounded (UNIV :: 'a::{real_normed_vector, perfect_space} set)" +proof(auto simp add: bounded_pos not_le) + obtain x :: 'a where "x \ 0" + using perfect_choose_dist [OF zero_less_one] by fast + 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_sgn) + then show "\x::'a. b < norm x" .. +qed + +lemma bounded_linear_image: + assumes "bounded S" "bounded_linear f" + shows "bounded(f ` S)" +proof- + from assms(1) obtain b where b:"b>0" "\x\S. norm x \ b" unfolding bounded_pos by auto + from assms(2) obtain B where B:"B>0" "\x. norm (f x) \ B * norm x" using bounded_linear.pos_bounded by (auto simp add: mult_ac) + { fix x assume "x\S" + hence "norm x \ b" using b by auto + hence "norm (f x) \ B * b" using B(2) apply(erule_tac x=x in allE) + by (metis B(1) B(2) real_le_trans real_mult_le_cancel_iff2) + } + thus ?thesis unfolding bounded_pos apply(rule_tac x="b*B" in exI) + using b B real_mult_order[of b B] by (auto simp add: real_mult_commute) +qed + +lemma bounded_scaling: + fixes S :: "'a::real_normed_vector set" + shows "bounded S \ bounded ((\x. c *\<^sub>R x) ` S)" + apply (rule bounded_linear_image, assumption) + apply (rule scaleR.bounded_linear_right) + done + +lemma bounded_translation: + fixes S :: "'a::real_normed_vector set" + assumes "bounded S" shows "bounded ((\x. a + x) ` S)" +proof- + from assms obtain b where b:"b>0" "\x\S. norm x \ b" unfolding bounded_pos by auto + { fix x assume "x\S" + hence "norm (a + x) \ b + norm a" using norm_triangle_ineq[of a x] b by auto + } + thus ?thesis unfolding bounded_pos using norm_ge_zero[of a] b(1) using add_strict_increasing[of b 0 "norm a"] + by (auto intro!: add exI[of _ "b + norm a"]) +qed + + +text{* Some theorems on sups and infs using the notion "bounded". *} + +lemma bounded_real: + fixes S :: "real set" + 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_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_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_real) + by (auto simp add: isLub_def isUb_def leastP_def setle_def setge_def) +qed + +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 +next + let ?S = "insert x S" + case False + hence *:"\x\S. x \ rsup S" using bounded_has_rsup(1)[of S] using assms by auto + hence "insert x S *<= max x (rsup S)" unfolding setle_def by auto + hence "isLub UNIV ?S (rsup ?S)" using rsup[of ?S] by auto + moreover + have **:"isUb UNIV ?S (max x (rsup S))" unfolding isUb_def setle_def using * by auto + { fix y assume as:"isUb UNIV (insert x S) y" + hence "max x (rsup S) \ y" unfolding isUb_def using rsup_le[OF `S\{}`] + unfolding setle_def by auto } + hence "max x (rsup S) <=* isUb UNIV (insert x S)" unfolding setge_def Ball_def mem_def by auto + hence "isLub UNIV ?S (max x (rsup S))" using ** isLubI2[of UNIV ?S "max x (rsup S)"] unfolding Collect_def by auto + ultimately show ?thesis using real_isLub_unique[of UNIV ?S] using `S\{}` by auto +qed + +lemma sup_insert_finite: "finite S \ rsup(insert x S) = (if S = {} then x else max x (rsup S))" + apply (rule rsup_insert) + apply (rule finite_imp_bounded) + by simp + +lemma bounded_has_rinf: + 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_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_real) + by (auto simp add: isGlb_def isLb_def greatestP_def setle_def setge_def) +qed + +(* TODO: Move this to RComplete.thy -- would need to include Glb into RComplete *) +lemma real_isGlb_unique: "[| isGlb R S x; isGlb R S y |] ==> x = (y::real)" + apply (frule isGlb_isLb) + apply (frule_tac x = y in isGlb_isLb) + apply (blast intro!: order_antisym dest!: isGlb_le_isLb) + done + +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 +next + let ?S = "insert x S" + case False + hence *:"\x\S. x \ rinf S" using bounded_has_rinf(1)[of S] using assms by auto + hence "min x (rinf S) <=* insert x S" unfolding setge_def by auto + hence "isGlb UNIV ?S (rinf ?S)" using rinf[of ?S] by auto + moreover + have **:"isLb UNIV ?S (min x (rinf S))" unfolding isLb_def setge_def using * by auto + { fix y assume as:"isLb UNIV (insert x S) y" + hence "min x (rinf S) \ y" unfolding isLb_def using rinf_ge[OF `S\{}`] + unfolding setge_def by auto } + hence "isLb UNIV (insert x S) *<= min x (rinf S)" unfolding setle_def Ball_def mem_def by auto + hence "isGlb UNIV ?S (min x (rinf S))" using ** isGlbI2[of UNIV ?S "min x (rinf S)"] unfolding Collect_def by auto + ultimately show ?thesis using real_isGlb_unique[of UNIV ?S] using `S\{}` by auto +qed + +lemma inf_insert_finite: "finite S ==> rinf(insert x S) = (if S = {} then x else min x (rinf S))" + by (rule rinf_insert, rule finite_imp_bounded, simp) + +subsection{* Compactness (the definition is the one based on convegent subsequences). *} + +definition + compact :: "'a::metric_space set \ bool" where (* TODO: generalize *) + "compact S \ + (\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 + Heine-Borel property if every closed and bounded subset is compact. +*} + +class heine_borel = + assumes bounded_imp_convergent_subsequence: + "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: "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. subseq r \ ((f \ r) ---> l) sequentially" + using `l \ s` r l by blast +qed + +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 [unfolded subseq_def] by auto + ultimately show "Suc n \ r (Suc n)" by auto +qed + +lemma eventually_subseq: + assumes r: "subseq r" + shows "eventually P sequentially \ eventually (\n. P (r n)) sequentially" +unfolding eventually_sequentially +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 + apply (rule_tac x="nat_rec e f" in exI) + apply (rule conjI)+ +apply (rule def_nat_rec_0, simp) +apply (rule allI, rule def_nat_rec_Suc, simp) +apply (rule allI, rule impI, rule ext) +apply (erule conjE) +apply (induct_tac x) +apply (simp add: nat_rec_0) +apply (erule_tac x="n" in allE) +apply (simp) +done + +lemma convergent_bounded_increasing: fixes s ::"nat\real" + 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 + then obtain t where t:"isLub UNIV (range s) t" using reals_complete[of "range s" ] by auto + { fix e::real assume "e>0" and as:"\N. \n\N. \ \s n - t\ < e" + { fix n::nat + 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)[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 "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. subseq r \ ((s \ r) ---> l) sequentially" +proof- + 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 + +instance real :: heine_borel +proof + fix s :: "real set" and f :: "nat \ real" + assume s: "bounded s" and f: "\n. f n \ s" + then obtain b where b: "\n. abs (f n) \ b" + unfolding bounded_iff by auto + obtain l :: real and r :: "nat \ nat" where + r: "subseq r" and l: "((f \ r) ---> l) sequentially" + using compact_real_lemma [OF b] by auto + thus "\l r. subseq r \ ((f \ r) ---> l) sequentially" + by auto +qed + +lemma bounded_component: "bounded s \ bounded ((\x. x $ i) ` s)" +unfolding bounded_def +apply clarify +apply (rule_tac x="x $ i" in exI) +apply (rule_tac x="e" in exI) +apply clarify +apply (rule order_trans [OF dist_nth_le], simp) +done + +lemma compact_lemma: + fixes f :: "nat \ 'a::heine_borel ^ 'n::finite" + assumes "bounded s" and "\n. f n \ s" + shows "\d. + \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. subseq r \ + (\e>0. eventually (\n. \i\d. dist (f (r n) $ i) (l $ i) < e) sequentially)" + 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:"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:"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:"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_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) + } + ultimately show ?case by auto + qed +qed + +instance "^" :: (heine_borel, finite) heine_borel +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: "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" + { fix e::real assume "e>0" + hence "0 < e / (real_of_nat (card ?d))" + using zero_less_card_finite using divide_pos_pos[of e, of "real_of_nat (card ?d)"] by auto + with l have "eventually (\n. \i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))) sequentially" + by simp + moreover + { fix n assume n: "\i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))" + have "dist (f (r n)) l \ (\i\?d. dist (f (r n) $ i) (l $ i))" + unfolding dist_vector_def using zero_le_dist by (rule setL2_le_setsum) + also have "\ < (\i\?d. e / (real_of_nat (card ?d)))" + by (rule setsum_strict_mono) (simp_all add: n) + finally have "dist (f (r n)) l < e" by simp + } + ultimately have "eventually (\n. dist (f (r n)) l < e) sequentially" + by (rule eventually_elim1) + } + hence *:"((f \ r) ---> l) sequentially" unfolding o_def tendsto_iff by simp + 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. *} + +lemma cauchy_def: + "Cauchy s \ (\e>0. \N. \m n. m \ N \ n \ N --> dist(s m)(s n) < e)" +unfolding Cauchy_def by blast + +definition + complete :: "'a::metric_space set \ bool" where + "complete s \ (\f. (\n. f n \ s) \ Cauchy f + --> (\l \ s. (f ---> l) sequentially))" + +lemma cauchy: "Cauchy s \ (\e>0.\ N::nat. \n\N. dist(s n)(s N) < e)" (is "?lhs = ?rhs") +proof- + { assume ?rhs + { fix e::real + assume "e>0" + with `?rhs` obtain N where N:"\n\N. dist (s n) (s N) < e/2" + by (erule_tac x="e/2" in allE) auto + { fix n m + assume nm:"N \ m \ N \ n" + hence "dist (s m) (s n) < e" using N + using dist_triangle_half_l[of "s m" "s N" "e" "s n"] + by blast + } + hence "\N. \m n. N \ m \ N \ n \ dist (s m) (s n) < e" + by blast + } + hence ?lhs + unfolding cauchy_def + by blast + } + thus ?thesis + unfolding cauchy_def + using dist_triangle_half_l + by blast +qed + +lemma convergent_imp_cauchy: + "(s ---> l) sequentially ==> Cauchy s" +proof(simp only: cauchy_def, rule, rule) + fix e::real assume "e>0" "(s ---> l) sequentially" + then obtain N::nat where N:"\n\N. dist (s n) l < e/2" unfolding Lim_sequentially by(erule_tac x="e/2" in allE) auto + thus "\N. \m n. N \ m \ N \ n \ dist (s m) (s n) < e" using dist_triangle_half_l[of _ l e _] by (rule_tac x=N in exI) auto +qed + +lemma cauchy_imp_bounded: assumes "Cauchy s" shows "bounded {y. (\n::nat. y = s n)}" +proof- + from assms obtain N::nat where "\m n. N \ m \ N \ n \ dist (s m) (s n) < 1" unfolding cauchy_def apply(erule_tac x= 1 in allE) by auto + hence N:"\n. N \ n \ dist (s N) (s n) < 1" by auto + moreover + have "bounded (s ` {0..N})" using finite_imp_bounded[of "s ` {1..N}"] by auto + then obtain a where a:"\x\s ` {0..N}. dist (s N) x \ a" + unfolding bounded_any_center [where a="s N"] by auto + ultimately show "?thesis" + unfolding bounded_any_center [where a="s N"] + apply(rule_tac x="max a 1" in exI) apply auto + apply(erule_tac x=n in allE) apply(erule_tac x=n in ballE) by auto +qed + +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" "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 + from lr(3)[unfolded Lim_sequentially, THEN spec[where x="e/2"]] obtain M where M:"\n\M. dist ((f \ r) n) l < e/2" using `e>0` by auto + { fix n::nat assume n:"n \ max N M" + have "dist ((f \ r) n) l < e/2" using n M by auto + moreover have "r n \ N" using lr'[of n] n by auto + hence "dist (f n) ((f \ r) n) < e / 2" using N using n by auto + ultimately have "dist (f n) l < e" using dist_triangle_half_r[of "f (r n)" "f n" e l] by (auto simp add: dist_commute) } + hence "\N. \n\N. dist (f n) l < e" by blast } + hence "\l\s. (f ---> l) sequentially" using `l\s` unfolding Lim_sequentially by auto } + thus ?thesis unfolding complete_def by auto +qed + +instance heine_borel < complete_space +proof + fix f :: "nat \ 'a" assume "Cauchy f" + hence "bounded (range f)" unfolding image_def + using cauchy_imp_bounded [of f] by auto + hence "compact (closure (range f))" + using bounded_closed_imp_compact [of "closure (range f)"] by auto + hence "complete (closure (range f))" + using compact_imp_complete by auto + moreover have "\n. f n \ closure (range f)" + using closure_subset [of "range f"] by auto + ultimately have "\l\closure (range f). (f ---> l) sequentially" + using `Cauchy f` unfolding complete_def by auto + then show "convergent f" + unfolding convergent_def LIMSEQ_conv_tendsto [symmetric] by auto +qed + +lemma complete_univ: "complete (UNIV :: 'a::complete_space set)" +proof(simp add: complete_def, rule, rule) + fix f :: "nat \ 'a" assume "Cauchy f" + hence "convergent f" by (rule Cauchy_convergent) + hence "\l. f ----> l" unfolding convergent_def . + thus "\l. (f ---> l) sequentially" unfolding LIMSEQ_conv_tendsto . +qed + +lemma complete_imp_closed: assumes "complete s" shows "closed s" +proof - + { fix x assume "x islimpt s" + then obtain f where f: "\n. f n \ s - {x}" "(f ---> x) sequentially" + unfolding islimpt_sequential by auto + then obtain l where l: "l\s" "(f ---> l) sequentially" + using `complete s`[unfolded complete_def] using convergent_imp_cauchy[of f x] by auto + hence "x \ s" using Lim_unique[of sequentially f l x] trivial_limit_sequentially f(2) by auto + } + thus "closed s" unfolding closed_limpt by auto +qed + +lemma complete_eq_closed: + fixes s :: "'a::complete_space set" + shows "complete s \ closed s" (is "?lhs = ?rhs") +proof + assume ?lhs thus ?rhs by (rule complete_imp_closed) +next + assume ?rhs + { fix f assume as:"\n::nat. f n \ s" "Cauchy f" + then obtain l where "(f ---> l) sequentially" using complete_univ[unfolded complete_def, THEN spec[where x=f]] by auto + hence "\l\s. (f ---> l) sequentially" using `?rhs`[unfolded closed_sequential_limits, THEN spec[where x=f], THEN spec[where x=l]] using as(1) by auto } + thus ?lhs unfolding complete_def by auto +qed + +lemma convergent_eq_cauchy: + fixes s :: "nat \ 'a::complete_space" + shows "(\l. (s ---> l) sequentially) \ Cauchy s" (is "?lhs = ?rhs") +proof + assume ?lhs then obtain l where "(s ---> l) sequentially" by auto + thus ?rhs using convergent_imp_cauchy by auto +next + assume ?rhs thus ?lhs using complete_univ[unfolded complete_def, THEN spec[where x=s]] by auto +qed + +lemma convergent_imp_bounded: + fixes s :: "nat \ 'a::metric_space" + shows "(s ---> l) sequentially ==> bounded (s ` (UNIV::(nat set)))" + using convergent_imp_cauchy[of s] + using cauchy_imp_bounded[of s] + unfolding image_def + by auto + +subsection{* Total boundedness. *} + +fun helper_1::"('a::metric_space set) \ real \ nat \ 'a" where + "helper_1 s e n = (SOME y::'a. y \ s \ (\m (dist (helper_1 s e m) y < e)))" +declare helper_1.simps[simp del] + +lemma compact_imp_totally_bounded: + assumes "compact s" + shows "\e>0. \k. finite k \ k \ s \ s \ (\((\x. ball x e) ` k))" +proof(rule, rule, rule ccontr) + fix e::real assume "e>0" and assm:"\ (\k. finite k \ k \ s \ s \ \(\x. ball x e) ` k)" + def x \ "helper_1 s e" + { fix n + have "x n \ s \ (\m dist (x m) (x n) < e)" + proof(induct_tac rule:nat_less_induct) + fix n def Q \ "(\y. y \ s \ (\m dist (x m) y < e))" + assume as:"\m s \ (\ma dist (x ma) (x m) < e)" + have "\ s \ (\x\x ` {0..s" "z \ (\x\x ` {0.. 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:"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[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 + +subsection{* Heine-Borel theorem (following Burkill \& Burkill vol. 2) *} + +lemma heine_borel_lemma: fixes s::"'a::metric_space set" + assumes "compact s" "s \ (\ t)" "\b \ t. open b" + shows "\e>0. \x \ s. \b \ t. ball x e \ b" +proof(rule ccontr) + assume "\ (\e>0. \x\s. \b\t. ball x e \ b)" + hence cont:"\e>0. \x\s. \xa\t. \ (ball x e \ xa)" by auto + { fix n::nat + have "1 / real (n + 1) > 0" by auto + hence "\x. x\s \ (\xa\t. \ (ball x (inverse (real (n+1))) \ xa))" using cont unfolding Bex_def by auto } + hence "\n::nat. \x. x \ s \ (\xa\t. \ ball x (inverse (real (n + 1))) \ xa)" by auto + 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:"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 + then obtain e where "e>0" and e:"\z. dist z l < e \ z\b" + using assms(3)[THEN bspec[where x=b]] unfolding open_dist by auto + + then obtain N1 where N1:"\n\N1. dist ((f \ r) n) l < e / 2" + using lr[unfolded Lim_sequentially, THEN spec[where x="e/2"]] by auto + + 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 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 + using f[THEN spec[where x="r (N1 + N2)"]] using `b\t` by auto + have "\y\ball x (inverse (real (r (N1 + N2) + 1))). y\b" apply(rule ccontr) using x by auto + then obtain y where y:"y \ ball x (inverse (real (r (N1 + N2) + 1)))" "y \ b" by auto + + have "dist x l < e/2" using N1 unfolding x_def o_def by auto + hence "dist y l < e" using y N2' using dist_triangle[of y l x]by (auto simp add:dist_commute) + + thus False using e and `y\b` by auto +qed + +lemma compact_imp_heine_borel: "compact s ==> (\f. (\t \ f. open t) \ s \ (\ f) + \ (\f'. f' \ f \ finite f' \ s \ (\ f')))" +proof clarify + fix f assume "compact s" " \t\f. open t" "s \ \f" + then obtain e::real where "e>0" and "\x\s. \b\f. ball x e \ b" using heine_borel_lemma[of s f] by auto + hence "\x\s. \b. b\f \ ball x e \ b" by auto + hence "\bb. \x\s. bb x \f \ ball x e \ bb x" using bchoice[of s "\x b. b\f \ ball x e \ b"] by auto + then obtain bb where bb:"\x\s. (bb x) \ f \ ball x e \ (bb x)" by blast + + from `compact s` have "\ k. finite k \ k \ s \ s \ \(\x. ball x e) ` k" using compact_imp_totally_bounded[of s] `e>0` by auto + then obtain k where k:"finite k" "k \ s" "s \ \(\x. ball x e) ` k" by auto + + have "finite (bb ` k)" using k(1) by auto + moreover + { fix x assume "x\s" + hence "x\\(\x. ball x e) ` k" using k(3) unfolding subset_eq by auto + hence "\X\bb ` k. x \ X" using bb k(2) by blast + hence "x \ \(bb ` k)" using Union_iff[of x "bb ` k"] by auto + } + ultimately show "\f'\f. finite f' \ s \ \f'" using bb k(2) by (rule_tac x="bb ` k" in exI) auto +qed + +subsection{* Bolzano-Weierstrass property. *} + +lemma heine_borel_imp_bolzano_weierstrass: + assumes "\f. (\t \ f. open t) \ s \ (\ f) --> (\f'. f' \ f \ finite f' \ s \ (\ f'))" + "infinite t" "t \ s" + shows "\x \ s. x islimpt t" +proof(rule ccontr) + assume "\ (\x \ s. x islimpt t)" + then obtain f where f:"\x\s. x \ f x \ open (f x) \ (\y\t. y \ f x \ y = x)" unfolding islimpt_def + using bchoice[of s "\ x T. x \ T \ open T \ (\y\t. y \ T \ y = x)"] by auto + obtain g where g:"g\{t. \x. x \ s \ t = f x}" "finite g" "s \ \g" + using assms(1)[THEN spec[where x="{t. \x. x\s \ t = f x}"]] using f by auto + from g(1,3) have g':"\x\g. \xa \ s. x = f xa" by auto + { fix x y assume "x\t" "y\t" "f x = f y" + hence "x \ f x" "y \ f x \ y = x" using f[THEN bspec[where x=x]] and `t\s` by auto + hence "x = y" using `f x = f y` and f[THEN bspec[where x=y]] and `y\t` and `t\s` by auto } + hence "infinite (f ` t)" using assms(2) using finite_imageD[unfolded inj_on_def, of f t] by auto + moreover + { fix x assume "x\t" "f x \ g" + from g(3) assms(3) `x\t` obtain h where "h\g" and "x\h" by auto + then obtain y where "y\s" "h = f y" using g'[THEN bspec[where x=h]] by auto + hence "y = x" using f[THEN bspec[where x=y]] and `x\t` and `x\h`[unfolded `h = f y`] by auto + hence False using `f x \ g` `h\g` unfolding `h = f y` by auto } + hence "f ` t \ g" by auto + ultimately show False using g(2) using finite_subset by auto +qed + +subsection{* Complete the chain of compactness variants. *} + +primrec helper_2::"(real \ 'a::metric_space) \ nat \ 'a" where + "helper_2 beyond 0 = beyond 0" | + "helper_2 beyond (Suc n) = beyond (dist arbitrary (helper_2 beyond n) + 1 )" + +lemma bolzano_weierstrass_imp_bounded: fixes s::"'a::metric_space set" + assumes "\t. infinite t \ t \ s --> (\x \ s. x islimpt t)" + shows "bounded s" +proof(rule ccontr) + assume "\ bounded s" + then obtain beyond where "\a. beyond a \s \ \ dist arbitrary (beyond a) \ a" + unfolding bounded_any_center [where a=arbitrary] + apply simp using choice[of "\a x. x\s \ \ dist arbitrary x \ a"] by auto + hence beyond:"\a. beyond a \s" "\a. dist arbitrary (beyond a) > a" + unfolding linorder_not_le by auto + def x \ "helper_2 beyond" + + { fix m n ::nat assume "mn" + have "1 < dist (x m) (x n)" + proof(cases "mn` by auto + hence "1 < dist arbitrary (x m) - dist arbitrary (x n)" using *[of n m] by auto + thus ?thesis using dist_triangle2 [of arbitrary "x m" "x n"] by arith + qed } note ** = this + { fix a b assume "x a = x b" "a \ b" + hence False using **[of a b] by auto } + hence "inj x" unfolding inj_on_def by auto + moreover + { fix n::nat + have "x n \ s" + proof(cases "n = 0") + case True thus ?thesis unfolding x_def using beyond by auto + next + case False then obtain z where "n = Suc z" using not0_implies_Suc by auto + thus ?thesis unfolding x_def using beyond by auto + qed } + ultimately have "infinite (range x) \ range x \ s" unfolding x_def using range_inj_infinite[of "helper_2 beyond"] using beyond(1) by auto + + then obtain l where "l\s" and l:"l islimpt range x" using assms[THEN spec[where x="range x"]] by auto + then obtain y where "x y \ l" and y:"dist (x y) l < 1/2" unfolding islimpt_approachable apply(erule_tac x="1/2" in allE) by auto + then obtain z where "x z \ l" and z:"dist (x z) l < dist (x y) l" using l[unfolded islimpt_approachable, THEN spec[where x="dist (x y) l"]] + unfolding dist_nz by auto + show False using y and z and dist_triangle_half_l[of "x y" l 1 "x z"] and **[of y z] by auto +qed + +lemma sequence_infinite_lemma: + fixes l :: "'a::metric_space" (* TODO: generalize *) + assumes "\n::nat. (f n \ l)" "(f ---> l) sequentially" + shows "infinite {y. (\ n. y = f n)}" +proof(rule ccontr) + let ?A = "(\x. dist x l) ` {y. \n. y = f n}" + assume "\ infinite {y. \n. y = f n}" + hence **:"finite ?A" "?A \ {}" by auto + obtain k where k:"dist (f k) l = Min ?A" using Min_in[OF **] by auto + have "0 < Min ?A" using assms(1) unfolding dist_nz unfolding Min_gr_iff[OF **] by auto + then obtain N where "dist (f N) l < Min ?A" using assms(2)[unfolded Lim_sequentially, THEN spec[where x="Min ?A"]] by auto + moreover have "dist (f N) l \ ?A" by auto + ultimately show False using Min_le[OF **(1), of "dist (f N) l"] by auto +qed + +lemma sequence_unique_limpt: + fixes l :: "'a::metric_space" (* TODO: generalize *) + assumes "\n::nat. (f n \ l)" "(f ---> l) sequentially" "l' islimpt {y. (\n. y = f n)}" + shows "l' = l" +proof(rule ccontr) + def e \ "dist l' l" + assume "l' \ l" hence "e>0" unfolding dist_nz e_def by auto + then obtain N::nat where N:"\n\N. dist (f n) l < e / 2" + using assms(2)[unfolded Lim_sequentially, THEN spec[where x="e/2"]] by auto + def d \ "Min (insert (e/2) ((\n. if dist (f n) l' = 0 then e/2 else dist (f n) l') ` {0 .. N}))" + have "d>0" using `e>0` unfolding d_def e_def using zero_le_dist[of _ l', unfolded order_le_less] by auto + obtain k where k:"f k \ l'" "dist (f k) l' < d" using `d>0` and assms(3)[unfolded islimpt_approachable, THEN spec[where x="d"]] by auto + have "k\N" using k(1)[unfolded dist_nz] using k(2)[unfolded d_def] + by force + hence "dist l' l < e" using N[THEN spec[where x=k]] using k(2)[unfolded d_def] and dist_triangle_half_r[of "f k" l' e l] by auto + thus False unfolding e_def by auto +qed + +lemma bolzano_weierstrass_imp_closed: + fixes s :: "'a::metric_space set" (* TODO: can this be generalized? *) + assumes "\t. infinite t \ t \ s --> (\x \ s. x islimpt t)" + shows "closed s" +proof- + { fix x l assume as: "\n::nat. x n \ s" "(x ---> l) sequentially" + hence "l \ s" + proof(cases "\n. x n \ l") + case False thus "l\s" using as(1) by auto + next + case True note cas = this + with as(2) have "infinite {y. \n. y = x n}" using sequence_infinite_lemma[of x l] by auto + then obtain l' where "l'\s" "l' islimpt {y. \n. y = x n}" using assms[THEN spec[where x="{y. \n. y = x n}"]] as(1) by auto + thus "l\s" using sequence_unique_limpt[of x l l'] using as cas by auto + qed } + thus ?thesis unfolding closed_sequential_limits by fast +qed + +text{* Hence express everything as an equivalence. *} + +lemma compact_eq_heine_borel: + fixes s :: "'a::heine_borel set" + shows "compact s \ + (\f. (\t \ f. open t) \ s \ (\ f) + --> (\f'. f' \ f \ finite f' \ s \ (\ f')))" (is "?lhs = ?rhs") +proof + assume ?lhs thus ?rhs using compact_imp_heine_borel[of s] by blast +next + assume ?rhs + hence "\t. infinite t \ t \ s \ (\x\s. x islimpt t)" + by (blast intro: heine_borel_imp_bolzano_weierstrass[of s]) + thus ?lhs using bolzano_weierstrass_imp_bounded[of s] bolzano_weierstrass_imp_closed[of s] bounded_closed_imp_compact[of s] by blast +qed + +lemma compact_eq_bolzano_weierstrass: + fixes s :: "'a::heine_borel set" + shows "compact s \ (\t. infinite t \ t \ s --> (\x \ s. x islimpt t))" (is "?lhs = ?rhs") +proof + assume ?lhs thus ?rhs unfolding compact_eq_heine_borel using heine_borel_imp_bolzano_weierstrass[of s] by auto +next + assume ?rhs thus ?lhs using bolzano_weierstrass_imp_bounded bolzano_weierstrass_imp_closed bounded_closed_imp_compact by auto +qed + +lemma compact_eq_bounded_closed: + fixes s :: "'a::heine_borel set" + shows "compact s \ bounded s \ closed s" (is "?lhs = ?rhs") +proof + assume ?lhs thus ?rhs unfolding compact_eq_bolzano_weierstrass using bolzano_weierstrass_imp_bounded bolzano_weierstrass_imp_closed by auto +next + assume ?rhs thus ?lhs using bounded_closed_imp_compact by auto +qed + +lemma compact_imp_bounded: + fixes s :: "'a::metric_space set" + shows "compact s ==> bounded s" +proof - + assume "compact s" + hence "\f. (\t\f. open t) \ s \ \f \ (\f'\f. finite f' \ s \ \f')" + by (rule compact_imp_heine_borel) + hence "\t. infinite t \ t \ s \ (\x \ s. x islimpt t)" + using heine_borel_imp_bolzano_weierstrass[of s] by auto + thus "bounded s" + by (rule bolzano_weierstrass_imp_bounded) +qed + +lemma compact_imp_closed: + fixes s :: "'a::metric_space set" + shows "compact s ==> closed s" +proof - + assume "compact s" + hence "\f. (\t\f. open t) \ s \ \f \ (\f'\f. finite f' \ s \ \f')" + by (rule compact_imp_heine_borel) + hence "\t. infinite t \ t \ s \ (\x \ s. x islimpt t)" + using heine_borel_imp_bolzano_weierstrass[of s] by auto + thus "closed s" + by (rule bolzano_weierstrass_imp_closed) +qed + +text{* In particular, some common special cases. *} + +lemma compact_empty[simp]: + "compact {}" + unfolding compact_def + by simp + +(* TODO: can any of the next 3 lemmas be generalized to metric spaces? *) + + (* FIXME : Rename *) +lemma compact_union[intro]: + fixes s t :: "'a::heine_borel set" + shows "compact s \ compact t ==> compact (s \ t)" + unfolding compact_eq_bounded_closed + using bounded_Un[of s t] + using closed_Un[of s t] + by simp + +lemma compact_inter[intro]: + fixes s t :: "'a::heine_borel set" + shows "compact s \ compact t ==> compact (s \ t)" + unfolding compact_eq_bounded_closed + using bounded_Int[of s t] + using closed_Int[of s t] + by simp + +lemma compact_inter_closed[intro]: + fixes s t :: "'a::heine_borel set" + shows "compact s \ closed t ==> compact (s \ t)" + unfolding compact_eq_bounded_closed + using closed_Int[of s t] + using bounded_subset[of "s \ t" s] + by blast + +lemma closed_inter_compact[intro]: + fixes s t :: "'a::heine_borel set" + shows "closed s \ compact t ==> compact (s \ t)" +proof- + assume "closed s" "compact t" + moreover + have "s \ t = t \ s" by auto ultimately + show ?thesis + using compact_inter_closed[of t s] + by auto +qed + +lemma closed_sing [simp]: + fixes a :: "'a::metric_space" + shows "closed {a}" + apply (clarsimp simp add: closed_def open_dist) + apply (rule ccontr) + apply (drule_tac x="dist x a" in spec) + apply (simp add: dist_nz dist_commute) + done + +lemma finite_imp_closed: + fixes s :: "'a::metric_space set" + shows "finite s ==> closed s" +proof (induct set: finite) + case empty show "closed {}" by simp +next + case (insert x F) + hence "closed ({x} \ F)" by (simp only: closed_Un closed_sing) + thus "closed (insert x F)" by simp +qed + +lemma finite_imp_compact: + fixes s :: "'a::heine_borel set" + shows "finite s ==> compact s" + unfolding compact_eq_bounded_closed + using finite_imp_closed finite_imp_bounded + by blast + +lemma compact_sing [simp]: "compact {a}" + unfolding compact_def o_def subseq_def + by (auto simp add: tendsto_const) + +lemma compact_cball[simp]: + fixes x :: "'a::heine_borel" + shows "compact(cball x e)" + using compact_eq_bounded_closed bounded_cball closed_cball + by blast + +lemma compact_frontier_bounded[intro]: + fixes s :: "'a::heine_borel set" + shows "bounded s ==> compact(frontier s)" + unfolding frontier_def + using compact_eq_bounded_closed + by blast + +lemma compact_frontier[intro]: + fixes s :: "'a::heine_borel set" + shows "compact s ==> compact (frontier s)" + using compact_eq_bounded_closed compact_frontier_bounded + by blast + +lemma frontier_subset_compact: + fixes s :: "'a::heine_borel set" + shows "compact s ==> frontier s \ s" + using frontier_subset_closed compact_eq_bounded_closed + by blast + +lemma open_delete: + fixes s :: "'a::metric_space set" + shows "open s ==> open(s - {x})" + using open_Diff[of s "{x}"] closed_sing + by blast + +text{* Finite intersection property. I could make it an equivalence in fact. *} + +lemma compact_imp_fip: + fixes s :: "'a::heine_borel set" + assumes "compact s" "\t \ f. closed t" + "\f'. finite f' \ f' \ f --> (s \ (\ f') \ {})" + shows "s \ (\ f) \ {}" +proof + assume as:"s \ (\ f) = {}" + hence "s \ \op - UNIV ` f" by auto + moreover have "Ball (op - UNIV ` f) open" using open_Diff closed_Diff using assms(2) by auto + ultimately obtain f' where f':"f' \ op - UNIV ` f" "finite f'" "s \ \f'" using assms(1)[unfolded compact_eq_heine_borel, THEN spec[where x="(\t. UNIV - t) ` f"]] by auto + hence "finite (op - UNIV ` f') \ op - UNIV ` f' \ f" by(auto simp add: Diff_Diff_Int) + hence "s \ \op - UNIV ` f' \ {}" using assms(3)[THEN spec[where x="op - UNIV ` f'"]] by auto + thus False using f'(3) unfolding subset_eq and Union_iff by blast +qed + +subsection{* Bounded closed nest property (proof does not use Heine-Borel). *} + +lemma bounded_closed_nest: + assumes "\n. closed(s n)" "\n. (s n \ {})" + "(\m n. m \ n --> s n \ s m)" "bounded(s 0)" + shows "\a::'a::heine_borel. \n::nat. a \ s(n)" +proof- + 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" "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 + { fix e::real assume "e>0" + 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 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) + using assms(3) apply(erule_tac x=n in allE)apply( erule_tac x="r (max N n)" in allE) by auto + ultimately have "\y\s n. dist y l < e" by auto + } + hence "l \ s n" using closed_approachable[of "s n" l] assms(1) by blast + } + thus ?thesis by auto +qed + +text{* Decreasing case does not even need compactness, just completeness. *} + +lemma decreasing_closed_nest: + assumes "\n. closed(s n)" + "\n. (s n \ {})" + "\m n. m \ n --> s n \ s m" + "\e>0. \n. \x \ (s n). \ y \ (s n). dist x y < e" + shows "\a::'a::heine_borel. \n::nat. a \ s n" +proof- + have "\n. \ x. x\s n" using assms(2) by auto + hence "\t. \n. t n \ s n" using choice[of "\ n x. x \ s n"] by auto + then obtain t where t: "\n. t n \ s n" by auto + { fix e::real assume "e>0" + then obtain N where N:"\x\s N. \y\s N. dist x y < e" using assms(4) by auto + { fix m n ::nat assume "N \ m \ N \ n" + hence "t m \ s N" "t n \ s N" using assms(3) t unfolding subset_eq t by blast+ + hence "dist (t m) (t n) < e" using N by auto + } + hence "\N. \m n. N \ m \ N \ n \ dist (t m) (t n) < e" by auto + } + hence "Cauchy t" unfolding cauchy_def by auto + then obtain l where l:"(t ---> l) sequentially" using complete_univ unfolding complete_def by auto + { fix n::nat + { fix e::real assume "e>0" + then obtain N::nat where N:"\n\N. dist (t n) l < e" using l[unfolded Lim_sequentially] by auto + have "t (max n N) \ s n" using assms(3) unfolding subset_eq apply(erule_tac x=n in allE) apply (erule_tac x="max n N" in allE) using t by auto + hence "\y\s n. dist y l < e" apply(rule_tac x="t (max n N)" in bexI) using N by auto + } + hence "l \ s n" using closed_approachable[of "s n" l] assms(1) by auto + } + then show ?thesis by auto +qed + +text{* Strengthen it to the intersection actually being a singleton. *} + +lemma decreasing_closed_nest_sing: + assumes "\n. closed(s n)" + "\n. s n \ {}" + "\m n. m \ n --> s n \ s m" + "\e>0. \n. \x \ (s n). \ y\(s n). dist x y < e" + shows "\a::'a::heine_borel. \ {t. (\n::nat. t = s n)} = {a}" +proof- + obtain a where a:"\n. a \ s n" using decreasing_closed_nest[of s] using assms by auto + { fix b assume b:"b \ \{t. \n. t = s n}" + { fix e::real assume "e>0" + hence "dist a b < e" using assms(4 )using b using a by blast + } + hence "dist a b = 0" by (metis dist_eq_0_iff dist_nz real_less_def) + } + with a have "\{t. \n. t = s n} = {a}" by auto + thus ?thesis by auto +qed + +text{* Cauchy-type criteria for uniform convergence. *} + +lemma uniformly_convergent_eq_cauchy: fixes s::"nat \ 'b \ 'a::heine_borel" shows + "(\l. \e>0. \N. \n x. N \ n \ P x --> dist(s n x)(l x) < e) \ + (\e>0. \N. \m n x. N \ m \ N \ n \ P x --> dist (s m x) (s n x) < e)" (is "?lhs = ?rhs") +proof(rule) + assume ?lhs + then obtain l where l:"\e>0. \N. \n x. N \ n \ P x \ dist (s n x) (l x) < e" by auto + { fix e::real assume "e>0" + then obtain N::nat where N:"\n x. N \ n \ P x \ dist (s n x) (l x) < e / 2" using l[THEN spec[where x="e/2"]] by auto + { fix n m::nat and x::"'b" assume "N \ m \ N \ n \ P x" + hence "dist (s m x) (s n x) < e" + using N[THEN spec[where x=m], THEN spec[where x=x]] + using N[THEN spec[where x=n], THEN spec[where x=x]] + using dist_triangle_half_l[of "s m x" "l x" e "s n x"] by auto } + hence "\N. \m n x. N \ m \ N \ n \ P x --> dist (s m x) (s n x) < e" by auto } + thus ?rhs by auto +next + assume ?rhs + hence "\x. P x \ Cauchy (\n. s n x)" unfolding cauchy_def apply auto by (erule_tac x=e in allE)auto + then obtain l where l:"\x. P x \ ((\n. s n x) ---> l x) sequentially" unfolding convergent_eq_cauchy[THEN sym] + using choice[of "\x l. P x \ ((\n. s n x) ---> l) sequentially"] by auto + { fix e::real assume "e>0" + then obtain N where N:"\m n x. N \ m \ N \ n \ P x \ dist (s m x) (s n x) < e/2" + using `?rhs`[THEN spec[where x="e/2"]] by auto + { fix x assume "P x" + then obtain M where M:"\n\M. dist (s n x) (l x) < e/2" + using l[THEN spec[where x=x], unfolded Lim_sequentially] using `e>0` by(auto elim!: allE[where x="e/2"]) + fix n::nat assume "n\N" + hence "dist(s n x)(l x) < e" using `P x`and N[THEN spec[where x=n], THEN spec[where x="N+M"], THEN spec[where x=x]] + using M[THEN spec[where x="N+M"]] and dist_triangle_half_l[of "s n x" "s (N+M) x" e "l x"] by (auto simp add: dist_commute) } + hence "\N. \n x. N \ n \ P x \ dist(s n x)(l x) < e" by auto } + thus ?lhs by auto +qed + +lemma uniformly_cauchy_imp_uniformly_convergent: + fixes s :: "nat \ 'a \ 'b::heine_borel" + assumes "\e>0.\N. \m (n::nat) x. N \ m \ N \ n \ P x --> dist(s m x)(s n x) < e" + "\x. P x --> (\e>0. \N. \n. N \ n --> dist(s n x)(l x) < e)" + shows "\e>0. \N. \n x. N \ n \ P x --> dist(s n x)(l x) < e" +proof- + obtain l' where l:"\e>0. \N. \n x. N \ n \ P x \ dist (s n x) (l' x) < e" + using assms(1) unfolding uniformly_convergent_eq_cauchy[THEN sym] by auto + moreover + { fix x assume "P x" + hence "l x = l' x" using Lim_unique[OF trivial_limit_sequentially, of "\n. s n x" "l x" "l' x"] + using l and assms(2) unfolding Lim_sequentially by blast } + ultimately show ?thesis by auto +qed + +subsection{* Define continuity over a net to take in restrictions of the set. *} + +definition + continuous :: "'a::t2_space net \ ('a \ 'b::topological_space) \ bool" where + "continuous net f \ (f ---> f(netlimit net)) net" + +lemma continuous_trivial_limit: + "trivial_limit net ==> continuous net f" + unfolding continuous_def tendsto_def trivial_limit_eq by auto + +lemma continuous_within: "continuous (at x within s) f \ (f ---> f(x)) (at x within s)" + unfolding continuous_def + unfolding tendsto_def + using netlimit_within[of x s] + by (cases "trivial_limit (at x within s)") (auto simp add: trivial_limit_eventually) + +lemma continuous_at: "continuous (at x) f \ (f ---> f(x)) (at x)" + using continuous_within [of x UNIV f] by (simp add: within_UNIV) + +lemma continuous_at_within: + assumes "continuous (at x) f" shows "continuous (at x within s) f" + using assms unfolding continuous_at continuous_within + by (rule Lim_at_within) + +text{* Derive the epsilon-delta forms, which we often use as "definitions" *} + +lemma continuous_within_eps_delta: + "continuous (at x within s) f \ (\e>0. \d>0. \x'\ s. dist x' x < d --> dist (f x') (f x) < e)" + unfolding continuous_within and Lim_within + apply auto unfolding dist_nz[THEN sym] apply(auto elim!:allE) apply(rule_tac x=d in exI) by auto + +lemma continuous_at_eps_delta: "continuous (at x) f \ (\e>0. \d>0. + \x'. dist x' x < d --> dist(f x')(f x) < e)" + using continuous_within_eps_delta[of x UNIV f] + unfolding within_UNIV by blast + +text{* Versions in terms of open balls. *} + +lemma continuous_within_ball: + "continuous (at x within s) f \ (\e>0. \d>0. + f ` (ball x d \ s) \ ball (f x) e)" (is "?lhs = ?rhs") +proof + assume ?lhs + { fix e::real assume "e>0" + then obtain d where d: "d>0" "\xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" + using `?lhs`[unfolded continuous_within Lim_within] by auto + { fix y assume "y\f ` (ball x d \ s)" + hence "y \ ball (f x) e" using d(2) unfolding dist_nz[THEN sym] + apply (auto simp add: dist_commute mem_ball) apply(erule_tac x=xa in ballE) apply auto using `e>0` by auto + } + hence "\d>0. f ` (ball x d \ s) \ ball (f x) e" using `d>0` unfolding subset_eq ball_def by (auto simp add: dist_commute) } + thus ?rhs by auto +next + assume ?rhs thus ?lhs unfolding continuous_within Lim_within ball_def subset_eq + apply (auto simp add: dist_commute) apply(erule_tac x=e in allE) by auto +qed + +lemma continuous_at_ball: + "continuous (at x) f \ (\e>0. \d>0. f ` (ball x d) \ ball (f x) e)" (is "?lhs = ?rhs") +proof + assume ?lhs thus ?rhs unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball + apply auto apply(erule_tac x=e in allE) apply auto apply(rule_tac x=d in exI) apply auto apply(erule_tac x=xa in allE) apply (auto simp add: dist_commute dist_nz) + unfolding dist_nz[THEN sym] by auto +next + assume ?rhs thus ?lhs unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball + apply auto apply(erule_tac x=e in allE) apply auto apply(rule_tac x=d in exI) apply auto apply(erule_tac x="f xa" in allE) by (auto simp add: dist_commute dist_nz) +qed + +text{* For setwise continuity, just start from the epsilon-delta definitions. *} + +definition + continuous_on :: "'a::metric_space set \ ('a \ 'b::metric_space) \ bool" where + "continuous_on s f \ (\x \ s. \e>0. \d::real>0. \x' \ s. dist x' x < d --> dist (f x') (f x) < e)" + + +definition + uniformly_continuous_on :: + "'a::metric_space set \ ('a \ 'b::metric_space) \ bool" where + "uniformly_continuous_on s f \ + (\e>0. \d>0. \x\s. \ x'\s. dist x' x < d + --> dist (f x') (f x) < e)" + +text{* Some simple consequential lemmas. *} + +lemma uniformly_continuous_imp_continuous: + " uniformly_continuous_on s f ==> continuous_on s f" + unfolding uniformly_continuous_on_def continuous_on_def by blast + +lemma continuous_at_imp_continuous_within: + "continuous (at x) f ==> continuous (at x within s) f" + unfolding continuous_within continuous_at using Lim_at_within by auto + +lemma continuous_at_imp_continuous_on: assumes "(\x \ s. continuous (at x) f)" + shows "continuous_on s f" +proof(simp add: continuous_at continuous_on_def, rule, rule, rule) + fix x and e::real assume "x\s" "e>0" + hence "eventually (\xa. dist (f xa) (f x) < e) (at x)" using assms unfolding continuous_at tendsto_iff by auto + then obtain d where d:"d>0" "\xa. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" unfolding eventually_at by auto + { fix x' assume "\ 0 < dist x' x" + hence "x=x'" + using dist_nz[of x' x] by auto + hence "dist (f x') (f x) < e" using `e>0` by auto + } + thus "\d>0. \x'\s. dist x' x < d \ dist (f x') (f x) < e" using d by auto +qed + +lemma continuous_on_eq_continuous_within: + "continuous_on s f \ (\x \ s. continuous (at x within s) f)" (is "?lhs = ?rhs") +proof + assume ?rhs + { fix x assume "x\s" + fix e::real assume "e>0" + assume "\d>0. \xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" + then obtain d where "d>0" and d:"\xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" by auto + { fix x' assume as:"x'\s" "dist x' x < d" + hence "dist (f x') (f x) < e" using `e>0` d `x'\s` dist_eq_0_iff[of x' x] zero_le_dist[of x' x] as(2) by (metis dist_eq_0_iff dist_nz) } + hence "\d>0. \x'\s. dist x' x < d \ dist (f x') (f x) < e" using `d>0` by auto + } + thus ?lhs using `?rhs` unfolding continuous_on_def continuous_within Lim_within by auto +next + assume ?lhs + thus ?rhs unfolding continuous_on_def continuous_within Lim_within by blast +qed + +lemma continuous_on: + "continuous_on s f \ (\x \ s. (f ---> f(x)) (at x within s))" + by (auto simp add: continuous_on_eq_continuous_within continuous_within) + +lemma continuous_on_eq_continuous_at: + "open s ==> (continuous_on s f \ (\x \ s. continuous (at x) f))" + by (auto simp add: continuous_on continuous_at Lim_within_open) + +lemma continuous_within_subset: + "continuous (at x within s) f \ t \ s + ==> continuous (at x within t) f" + unfolding continuous_within by(metis Lim_within_subset) + +lemma continuous_on_subset: + "continuous_on s f \ t \ s ==> continuous_on t f" + unfolding continuous_on by (metis subset_eq Lim_within_subset) + +lemma continuous_on_interior: + "continuous_on s f \ x \ interior s ==> continuous (at x) f" +unfolding interior_def +apply simp +by (meson continuous_on_eq_continuous_at continuous_on_subset) + +lemma continuous_on_eq: + "(\x \ s. f x = g x) \ continuous_on s f + ==> continuous_on s g" + by (simp add: continuous_on_def) + +text{* Characterization of various kinds of continuity in terms of sequences. *} + +(* \ could be generalized, but \ requires metric space *) +lemma continuous_within_sequentially: + fixes f :: "'a::metric_space \ 'b::metric_space" + shows "continuous (at a within s) f \ + (\x. (\n::nat. x n \ s) \ (x ---> a) sequentially + --> ((f o x) ---> f a) sequentially)" (is "?lhs = ?rhs") +proof + assume ?lhs + { fix x::"nat \ 'a" assume x:"\n. x n \ s" "\e>0. \N. \n\N. dist (x n) a < e" + fix e::real assume "e>0" + from `?lhs` obtain d where "d>0" and d:"\x\s. 0 < dist x a \ dist x a < d \ dist (f x) (f a) < e" unfolding continuous_within Lim_within using `e>0` by auto + from x(2) `d>0` obtain N where N:"\n\N. dist (x n) a < d" by auto + hence "\N. \n\N. dist ((f \ x) n) (f a) < e" + apply(rule_tac x=N in exI) using N d apply auto using x(1) + apply(erule_tac x=n in allE) apply(erule_tac x=n in allE) + apply(erule_tac x="x n" in ballE) apply auto unfolding dist_nz[THEN sym] apply auto using `e>0` by auto + } + thus ?rhs unfolding continuous_within unfolding Lim_sequentially by simp +next + assume ?rhs + { fix e::real assume "e>0" + assume "\ (\d>0. \x\s. 0 < dist x a \ dist x a < d \ dist (f x) (f a) < e)" + hence "\d. \x. d>0 \ x\s \ (0 < dist x a \ dist x a < d \ \ dist (f x) (f a) < e)" by blast + then obtain x where x:"\d>0. x d \ s \ (0 < dist (x d) a \ dist (x d) a < d \ \ dist (f (x d)) (f a) < e)" + using choice[of "\d x.0 x\s \ (0 < dist x a \ dist x a < d \ \ dist (f x) (f a) < e)"] by auto + { fix d::real assume "d>0" + hence "\N::nat. inverse (real (N + 1)) < d" using real_arch_inv[of d] by (auto, rule_tac x="n - 1" in exI)auto + then obtain N::nat where N:"inverse (real (N + 1)) < d" by auto + { fix n::nat assume n:"n\N" + hence "dist (x (inverse (real (n + 1)))) a < inverse (real (n + 1))" using x[THEN spec[where x="inverse (real (n + 1))"]] by auto + moreover have "inverse (real (n + 1)) < d" using N n 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) + ultimately have "dist (x (inverse (real (n + 1)))) a < d" by auto + } + hence "\N::nat. \n\N. dist (x (inverse (real (n + 1)))) a < d" by auto + } + hence "(\n::nat. x (inverse (real (n + 1))) \ s) \ (\e>0. \N::nat. \n\N. dist (x (inverse (real (n + 1)))) a < e)" using x by auto + hence "\e>0. \N::nat. \n\N. dist (f (x (inverse (real (n + 1))))) (f a) < e" using `?rhs`[THEN spec[where x="\n::nat. x (inverse (real (n+1)))"], unfolded Lim_sequentially] by auto + hence "False" apply(erule_tac x=e in allE) using `e>0` using x by auto + } + thus ?lhs unfolding continuous_within unfolding Lim_within unfolding Lim_sequentially by blast +qed + +lemma continuous_at_sequentially: + fixes f :: "'a::metric_space \ 'b::metric_space" + shows "continuous (at a) f \ (\x. (x ---> a) sequentially + --> ((f o x) ---> f a) sequentially)" + using continuous_within_sequentially[of a UNIV f] unfolding within_UNIV by auto + +lemma continuous_on_sequentially: + "continuous_on s f \ (\x. \a \ s. (\n. x(n) \ s) \ (x ---> a) sequentially + --> ((f o x) ---> f(a)) sequentially)" (is "?lhs = ?rhs") +proof + assume ?rhs thus ?lhs using continuous_within_sequentially[of _ s f] unfolding continuous_on_eq_continuous_within by auto +next + assume ?lhs thus ?rhs unfolding continuous_on_eq_continuous_within using continuous_within_sequentially[of _ s f] by auto +qed + +lemma uniformly_continuous_on_sequentially: + fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" + shows "uniformly_continuous_on s f \ (\x y. (\n. x n \ s) \ (\n. y n \ s) \ + ((\n. x n - y n) ---> 0) sequentially + \ ((\n. f(x n) - f(y n)) ---> 0) sequentially)" (is "?lhs = ?rhs") +proof + assume ?lhs + { fix x y assume x:"\n. x n \ s" and y:"\n. y n \ s" and xy:"((\n. x n - y n) ---> 0) sequentially" + { fix e::real assume "e>0" + then obtain d where "d>0" and d:"\x\s. \x'\s. dist x' x < d \ dist (f x') (f x) < e" + using `?lhs`[unfolded uniformly_continuous_on_def, THEN spec[where x=e]] by auto + obtain N where N:"\n\N. norm (x n - y n - 0) < d" using xy[unfolded Lim_sequentially dist_norm] and `d>0` by auto + { fix n assume "n\N" + hence "norm (f (x n) - f (y n) - 0) < e" + using N[THEN spec[where x=n]] using d[THEN bspec[where x="x n"], THEN bspec[where x="y n"]] using x and y + unfolding dist_commute and dist_norm by simp } + hence "\N. \n\N. norm (f (x n) - f (y n) - 0) < e" by auto } + hence "((\n. f(x n) - f(y n)) ---> 0) sequentially" unfolding Lim_sequentially and dist_norm by auto } + thus ?rhs by auto +next + assume ?rhs + { assume "\ ?lhs" + then obtain e where "e>0" "\d>0. \x\s. \x'\s. dist x' x < d \ \ dist (f x') (f x) < e" unfolding uniformly_continuous_on_def by auto + then obtain fa where fa:"\x. 0 < x \ fst (fa x) \ s \ snd (fa x) \ s \ dist (fst (fa x)) (snd (fa x)) < x \ \ dist (f (fst (fa x))) (f (snd (fa x))) < e" + using choice[of "\d x. d>0 \ fst x \ s \ snd x \ s \ dist (snd x) (fst x) < d \ \ dist (f (snd x)) (f (fst x)) < e"] unfolding Bex_def + by (auto simp add: dist_commute) + def x \ "\n::nat. fst (fa (inverse (real n + 1)))" + def y \ "\n::nat. snd (fa (inverse (real n + 1)))" + have xyn:"\n. x n \ s \ y n \ s" and xy0:"\n. dist (x n) (y n) < inverse (real n + 1)" and fxy:"\n. \ dist (f (x n)) (f (y n)) < e" + unfolding x_def and y_def using fa by auto + have 1:"\(x::'a) y. dist (x - y) 0 = dist x y" unfolding dist_norm by auto + have 2:"\(x::'b) y. dist (x - y) 0 = dist x y" unfolding dist_norm by auto + { fix e::real assume "e>0" + then obtain N::nat where "N \ 0" and N:"0 < inverse (real N) \ inverse (real N) < e" unfolding real_arch_inv[of e] by auto + { fix n::nat assume "n\N" + hence "inverse (real n + 1) < inverse (real N)" using real_of_nat_ge_zero and `N\0` by auto + also have "\ < e" using N by auto + finally have "inverse (real n + 1) < e" by auto + hence "dist (x n - y n) 0 < e" unfolding 1 using xy0[THEN spec[where x=n]] by auto } + hence "\N. \n\N. dist (x n - y n) 0 < e" by auto } + hence "\e>0. \N. \n\N. dist (f (x n) - f (y n)) 0 < e" using `?rhs`[THEN spec[where x=x], THEN spec[where x=y]] and xyn unfolding Lim_sequentially by auto + hence False unfolding 2 using fxy and `e>0` by auto } + thus ?lhs unfolding uniformly_continuous_on_def by blast +qed + +text{* The usual transformation theorems. *} + +lemma continuous_transform_within: + fixes f g :: "'a::metric_space \ 'b::metric_space" + assumes "0 < d" "x \ s" "\x' \ s. dist x' x < d --> f x' = g x'" + "continuous (at x within s) f" + shows "continuous (at x within s) g" +proof- + { fix e::real assume "e>0" + then obtain d' where d':"d'>0" "\xa\s. 0 < dist xa x \ dist xa x < d' \ dist (f xa) (f x) < e" using assms(4) unfolding continuous_within Lim_within by auto + { fix x' assume "x'\s" "0 < dist x' x" "dist x' x < (min d d')" + hence "dist (f x') (g x) < e" using assms(2,3) apply(erule_tac x=x in ballE) using d' by auto } + hence "\xa\s. 0 < dist xa x \ dist xa x < (min d d') \ dist (f xa) (g x) < e" by blast + hence "\d>0. \xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (g x) < e" using `d>0` `d'>0` by(rule_tac x="min d d'" in exI)auto } + hence "(f ---> g x) (at x within s)" unfolding Lim_within using assms(1) by auto + thus ?thesis unfolding continuous_within using Lim_transform_within[of d s x f g "g x"] using assms by blast +qed + +lemma continuous_transform_at: + fixes f g :: "'a::metric_space \ 'b::metric_space" + assumes "0 < d" "\x'. dist x' x < d --> f x' = g x'" + "continuous (at x) f" + shows "continuous (at x) g" +proof- + { fix e::real assume "e>0" + then obtain d' where d':"d'>0" "\xa. 0 < dist xa x \ dist xa x < d' \ dist (f xa) (f x) < e" using assms(3) unfolding continuous_at Lim_at by auto + { fix x' assume "0 < dist x' x" "dist x' x < (min d d')" + hence "dist (f x') (g x) < e" using assms(2) apply(erule_tac x=x in allE) using d' by auto + } + hence "\xa. 0 < dist xa x \ dist xa x < (min d d') \ dist (f xa) (g x) < e" by blast + hence "\d>0. \xa. 0 < dist xa x \ dist xa x < d \ dist (f xa) (g x) < e" using `d>0` `d'>0` by(rule_tac x="min d d'" in exI)auto + } + hence "(f ---> g x) (at x)" unfolding Lim_at using assms(1) by auto + thus ?thesis unfolding continuous_at using Lim_transform_at[of d x f g "g x"] using assms by blast +qed + +text{* Combination results for pointwise continuity. *} + +lemma continuous_const: "continuous net (\x. c)" + by (auto simp add: continuous_def Lim_const) + +lemma continuous_cmul: + fixes f :: "'a::t2_space \ 'b::real_normed_vector" + shows "continuous net f ==> continuous net (\x. c *\<^sub>R f x)" + by (auto simp add: continuous_def Lim_cmul) + +lemma continuous_neg: + fixes f :: "'a::t2_space \ 'b::real_normed_vector" + shows "continuous net f ==> continuous net (\x. -(f x))" + by (auto simp add: continuous_def Lim_neg) + +lemma continuous_add: + fixes f g :: "'a::t2_space \ 'b::real_normed_vector" + shows "continuous net f \ continuous net g \ continuous net (\x. f x + g x)" + by (auto simp add: continuous_def Lim_add) + +lemma continuous_sub: + fixes f g :: "'a::t2_space \ 'b::real_normed_vector" + shows "continuous net f \ continuous net g \ continuous net (\x. f x - g x)" + by (auto simp add: continuous_def Lim_sub) + +text{* Same thing for setwise continuity. *} + +lemma continuous_on_const: + "continuous_on s (\x. c)" + unfolding continuous_on_eq_continuous_within using continuous_const by blast + +lemma continuous_on_cmul: + fixes f :: "'a::metric_space \ 'b::real_normed_vector" + 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: + fixes f :: "'a::metric_space \ 'b::real_normed_vector" + shows "continuous_on s f \ continuous_on s (\x. - f x)" + unfolding continuous_on_eq_continuous_within using continuous_neg by blast + +lemma continuous_on_add: + fixes f g :: "'a::metric_space \ 'b::real_normed_vector" + shows "continuous_on s f \ continuous_on s g + \ continuous_on s (\x. f x + g x)" + unfolding continuous_on_eq_continuous_within using continuous_add by blast + +lemma continuous_on_sub: + fixes f g :: "'a::metric_space \ 'b::real_normed_vector" + shows "continuous_on s f \ continuous_on s g + \ continuous_on s (\x. f x - g x)" + unfolding continuous_on_eq_continuous_within using continuous_sub by blast + +text{* Same thing for uniform continuity, using sequential formulations. *} + +lemma uniformly_continuous_on_const: + "uniformly_continuous_on s (\x. c)" + unfolding uniformly_continuous_on_def by simp + +lemma uniformly_continuous_on_cmul: + fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" + (* FIXME: generalize 'a to metric_space *) + assumes "uniformly_continuous_on s f" + 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 *\<^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 scaleR_zero_right scaleR_right_diff_distrib by auto + } + thus ?thesis using assms unfolding uniformly_continuous_on_sequentially by auto +qed + +lemma dist_minus: + fixes x y :: "'a::real_normed_vector" + shows "dist (- x) (- y) = dist x y" + unfolding dist_norm minus_diff_minus norm_minus_cancel .. + +lemma uniformly_continuous_on_neg: + fixes f :: "'a::metric_space \ 'b::real_normed_vector" + shows "uniformly_continuous_on s f + ==> uniformly_continuous_on s (\x. -(f x))" + unfolding uniformly_continuous_on_def dist_minus . + +lemma uniformly_continuous_on_add: + fixes f g :: "'a::real_normed_vector \ 'b::real_normed_vector" (* FIXME: generalize 'a *) + assumes "uniformly_continuous_on s f" "uniformly_continuous_on s g" + shows "uniformly_continuous_on s (\x. f x + g x)" +proof- + { fix x y assume "((\n. f (x n) - f (y n)) ---> 0) sequentially" + "((\n. g (x n) - g (y n)) ---> 0) sequentially" + hence "((\xa. f (x xa) - f (y xa) + (g (x xa) - g (y xa))) ---> 0 + 0) sequentially" + using Lim_add[of "\ n. f (x n) - f (y n)" 0 sequentially "\ n. g (x n) - g (y n)" 0] by auto + hence "((\n. f (x n) + g (x n) - (f (y n) + g (y n))) ---> 0) sequentially" unfolding Lim_sequentially and add_diff_add [symmetric] by auto } + thus ?thesis using assms unfolding uniformly_continuous_on_sequentially by auto +qed + +lemma uniformly_continuous_on_sub: + fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" (* FIXME: generalize 'a *) + shows "uniformly_continuous_on s f \ uniformly_continuous_on s g + ==> uniformly_continuous_on s (\x. f x - g x)" + unfolding ab_diff_minus + using uniformly_continuous_on_add[of s f "\x. - g x"] + using uniformly_continuous_on_neg[of s g] by auto + +text{* Identity function is continuous in every sense. *} + +lemma continuous_within_id: + "continuous (at a within s) (\x. x)" + unfolding continuous_within by (rule Lim_at_within [OF Lim_ident_at]) + +lemma continuous_at_id: + "continuous (at a) (\x. x)" + unfolding continuous_at by (rule Lim_ident_at) + +lemma continuous_on_id: + "continuous_on s (\x. x)" + unfolding continuous_on Lim_within by auto + +lemma uniformly_continuous_on_id: + "uniformly_continuous_on s (\x. x)" + unfolding uniformly_continuous_on_def by auto + +text{* Continuity of all kinds is preserved under composition. *} + +lemma continuous_within_compose: + fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) + fixes g :: "'b::metric_space \ 'c::metric_space" + assumes "continuous (at x within s) f" "continuous (at (f x) within f ` s) g" + shows "continuous (at x within s) (g o f)" +proof- + { fix e::real assume "e>0" + with assms(2)[unfolded continuous_within Lim_within] obtain d where "d>0" and d:"\xa\f ` s. 0 < dist xa (f x) \ dist xa (f x) < d \ dist (g xa) (g (f x)) < e" by auto + from assms(1)[unfolded continuous_within Lim_within] obtain d' where "d'>0" and d':"\xa\s. 0 < dist xa x \ dist xa x < d' \ dist (f xa) (f x) < d" using `d>0` by auto + { fix y assume as:"y\s" "0 < dist y x" "dist y x < d'" + hence "dist (f y) (f x) < d" using d'[THEN bspec[where x=y]] by (auto simp add:dist_commute) + hence "dist (g (f y)) (g (f x)) < e" using as(1) d[THEN bspec[where x="f y"]] unfolding dist_nz[THEN sym] using `e>0` by auto } + hence "\d>0. \xa\s. 0 < dist xa x \ dist xa x < d \ dist (g (f xa)) (g (f x)) < e" using `d'>0` by auto } + thus ?thesis unfolding continuous_within Lim_within by auto +qed + +lemma continuous_at_compose: + fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) + fixes g :: "'b::metric_space \ 'c::metric_space" + assumes "continuous (at x) f" "continuous (at (f x)) g" + shows "continuous (at x) (g o f)" +proof- + have " continuous (at (f x) within range f) g" using assms(2) using continuous_within_subset[of "f x" UNIV g "range f", unfolded within_UNIV] by auto + thus ?thesis using assms(1) using continuous_within_compose[of x UNIV f g, unfolded within_UNIV] by auto +qed + +lemma continuous_on_compose: + "continuous_on s f \ continuous_on (f ` s) g \ continuous_on s (g o f)" + unfolding continuous_on_eq_continuous_within using continuous_within_compose[of _ s f g] by auto + +lemma uniformly_continuous_on_compose: + assumes "uniformly_continuous_on s f" "uniformly_continuous_on (f ` s) g" + shows "uniformly_continuous_on s (g o f)" +proof- + { fix e::real assume "e>0" + then obtain d where "d>0" and d:"\x\f ` s. \x'\f ` s. dist x' x < d \ dist (g x') (g x) < e" using assms(2) unfolding uniformly_continuous_on_def by auto + obtain d' where "d'>0" "\x\s. \x'\s. dist x' x < d' \ dist (f x') (f x) < d" using `d>0` using assms(1) unfolding uniformly_continuous_on_def by auto + hence "\d>0. \x\s. \x'\s. dist x' x < d \ dist ((g \ f) x') ((g \ f) x) < e" using `d>0` using d by auto } + thus ?thesis using assms unfolding uniformly_continuous_on_def by auto +qed + +text{* Continuity in terms of open preimages. *} + +lemma continuous_at_open: + fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) + shows "continuous (at x) f \ (\t. open t \ f x \ t --> (\s. open s \ x \ s \ (\x' \ s. (f x') \ t)))" (is "?lhs = ?rhs") +proof + assume ?lhs + { fix t assume as: "open t" "f x \ t" + then obtain e where "e>0" and e:"ball (f x) e \ t" unfolding open_contains_ball by auto + + obtain d where "d>0" and d:"\y. 0 < dist y x \ dist y x < d \ dist (f y) (f x) < e" using `e>0` using `?lhs`[unfolded continuous_at Lim_at open_dist] by auto + + have "open (ball x d)" using open_ball by auto + moreover have "x \ ball x d" unfolding centre_in_ball using `d>0` by simp + moreover + { fix x' assume "x'\ball x d" hence "f x' \ t" + using e[unfolded subset_eq Ball_def mem_ball, THEN spec[where x="f x'"]] d[THEN spec[where x=x']] + unfolding mem_ball apply (auto simp add: dist_commute) + unfolding dist_nz[THEN sym] using as(2) by auto } + hence "\x'\ball x d. f x' \ t" by auto + ultimately have "\s. open s \ x \ s \ (\x'\s. f x' \ t)" + apply(rule_tac x="ball x d" in exI) by simp } + thus ?rhs by auto +next + assume ?rhs + { fix e::real assume "e>0" + then obtain s where s: "open s" "x \ s" "\x'\s. f x' \ ball (f x) e" using `?rhs`[unfolded continuous_at Lim_at, THEN spec[where x="ball (f x) e"]] + unfolding centre_in_ball[of "f x" e, THEN sym] by auto + then obtain d where "d>0" and d:"ball x d \ s" unfolding open_contains_ball by auto + { fix y assume "0 < dist y x \ dist y x < d" + hence "dist (f y) (f x) < e" using d[unfolded subset_eq Ball_def mem_ball, THEN spec[where x=y]] + using s(3)[THEN bspec[where x=y], unfolded mem_ball] by (auto simp add: dist_commute) } + hence "\d>0. \xa. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" using `d>0` by auto } + thus ?lhs unfolding continuous_at Lim_at by auto +qed + +lemma continuous_on_open: + "continuous_on s f \ + (\t. openin (subtopology euclidean (f ` s)) t + --> openin (subtopology euclidean s) {x \ s. f x \ t})" (is "?lhs = ?rhs") +proof + assume ?lhs + { fix t assume as:"openin (subtopology euclidean (f ` s)) t" + have "{x \ s. f x \ t} \ s" using as[unfolded openin_euclidean_subtopology_iff] by auto + moreover + { fix x assume as':"x\{x \ s. f x \ t}" + then obtain e where e: "e>0" "\x'\f ` s. dist x' (f x) < e \ x' \ t" using as[unfolded openin_euclidean_subtopology_iff, THEN conjunct2, THEN bspec[where x="f x"]] by auto + from this(1) obtain d where d: "d>0" "\xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" using `?lhs`[unfolded continuous_on Lim_within, THEN bspec[where x=x]] using as' by auto + have "\e>0. \x'\s. dist x' x < e \ x' \ {x \ s. f x \ t}" using d e unfolding dist_nz[THEN sym] by (rule_tac x=d in exI, auto) } + ultimately have "openin (subtopology euclidean s) {x \ s. f x \ t}" unfolding openin_euclidean_subtopology_iff by auto } + thus ?rhs unfolding continuous_on Lim_within using openin by auto +next + assume ?rhs + { fix e::real and x assume "x\s" "e>0" + { fix xa x' assume "dist (f xa) (f x) < e" "xa \ s" "x' \ s" "dist (f xa) (f x') < e - dist (f xa) (f x)" + hence "dist (f x') (f x) < e" using dist_triangle[of "f x'" "f x" "f xa"] + by (auto simp add: dist_commute) } + hence "ball (f x) e \ f ` s \ f ` s \ (\xa\ball (f x) e \ f ` s. \ea>0. \x'\f ` s. dist x' xa < ea \ x' \ ball (f x) e \ f ` s)" apply auto + apply(rule_tac x="e - dist (f xa) (f x)" in exI) using `e>0` by (auto simp add: dist_commute) + hence "\xa\{xa \ s. f xa \ ball (f x) e \ f ` s}. \ea>0. \x'\s. dist x' xa < ea \ x' \ {xa \ s. f xa \ ball (f x) e \ f ` s}" + using `?rhs`[unfolded openin_euclidean_subtopology_iff, THEN spec[where x="ball (f x) e \ f ` s"]] by auto + hence "\d>0. \xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" apply(erule_tac x=x in ballE) apply auto using `e>0` `x\s` by (auto simp add: dist_commute) } + thus ?lhs unfolding continuous_on Lim_within by auto +qed + +(* ------------------------------------------------------------------------- *) +(* Similarly in terms of closed sets. *) +(* ------------------------------------------------------------------------- *) + +lemma continuous_on_closed: + "continuous_on s f \ (\t. closedin (subtopology euclidean (f ` s)) t --> closedin (subtopology euclidean s) {x \ s. f x \ t})" (is "?lhs = ?rhs") +proof + assume ?lhs + { fix t + have *:"s - {x \ s. f x \ f ` s - t} = {x \ s. f x \ t}" by auto + have **:"f ` s - (f ` s - (f ` s - t)) = f ` s - t" by auto + assume as:"closedin (subtopology euclidean (f ` s)) t" + hence "closedin (subtopology euclidean (f ` s)) (f ` s - (f ` s - t))" unfolding closedin_def topspace_euclidean_subtopology unfolding ** by auto + hence "closedin (subtopology euclidean s) {x \ s. f x \ t}" using `?lhs`[unfolded continuous_on_open, THEN spec[where x="(f ` s) - t"]] + unfolding openin_closedin_eq topspace_euclidean_subtopology unfolding * by auto } + thus ?rhs by auto +next + assume ?rhs + { fix t + have *:"s - {x \ s. f x \ f ` s - t} = {x \ s. f x \ t}" by auto + assume as:"openin (subtopology euclidean (f ` s)) t" + hence "openin (subtopology euclidean s) {x \ s. f x \ t}" using `?rhs`[THEN spec[where x="(f ` s) - t"]] + unfolding openin_closedin_eq topspace_euclidean_subtopology *[THEN sym] closedin_subtopology by auto } + thus ?lhs unfolding continuous_on_open by auto +qed + +text{* Half-global and completely global cases. *} + +lemma continuous_open_in_preimage: + assumes "continuous_on s f" "open t" + shows "openin (subtopology euclidean s) {x \ s. f x \ t}" +proof- + have *:"\x. x \ s \ f x \ t \ x \ s \ f x \ (t \ f ` s)" by auto + have "openin (subtopology euclidean (f ` s)) (t \ f ` s)" + using openin_open_Int[of t "f ` s", OF assms(2)] unfolding openin_open by auto + thus ?thesis using assms(1)[unfolded continuous_on_open, THEN spec[where x="t \ f ` s"]] using * by auto +qed + +lemma continuous_closed_in_preimage: + assumes "continuous_on s f" "closed t" + shows "closedin (subtopology euclidean s) {x \ s. f x \ t}" +proof- + have *:"\x. x \ s \ f x \ t \ x \ s \ f x \ (t \ f ` s)" by auto + have "closedin (subtopology euclidean (f ` s)) (t \ f ` s)" + using closedin_closed_Int[of t "f ` s", OF assms(2)] unfolding Int_commute by auto + thus ?thesis + using assms(1)[unfolded continuous_on_closed, THEN spec[where x="t \ f ` s"]] using * by auto +qed + +lemma continuous_open_preimage: + assumes "continuous_on s f" "open s" "open t" + shows "open {x \ s. f x \ t}" +proof- + obtain T where T: "open T" "{x \ s. f x \ t} = s \ T" + using continuous_open_in_preimage[OF assms(1,3)] unfolding openin_open by auto + thus ?thesis using open_Int[of s T, OF assms(2)] by auto +qed + +lemma continuous_closed_preimage: + assumes "continuous_on s f" "closed s" "closed t" + shows "closed {x \ s. f x \ t}" +proof- + obtain T where T: "closed T" "{x \ s. f x \ t} = s \ T" + using continuous_closed_in_preimage[OF assms(1,3)] unfolding closedin_closed by auto + thus ?thesis using closed_Int[of s T, OF assms(2)] by auto +qed + +lemma continuous_open_preimage_univ: + fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) + shows "\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 :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) + shows "(\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 + +lemma continuous_open_vimage: + fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) + shows "\x. continuous (at x) f \ open s \ open (f -` s)" + unfolding vimage_def by (rule continuous_open_preimage_univ) + +lemma continuous_closed_vimage: + fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) + shows "\x. continuous (at x) f \ closed s \ closed (f -` s)" + unfolding vimage_def by (rule continuous_closed_preimage_univ) + +text{* Equality of continuous functions on closure and related results. *} + +lemma continuous_closed_in_preimage_constant: + "continuous_on s f ==> closedin (subtopology euclidean s) {x \ s. f x = a}" + using continuous_closed_in_preimage[of s f "{a}"] closed_sing by auto + +lemma continuous_closed_preimage_constant: + "continuous_on s f \ closed s ==> closed {x \ s. f x = a}" + using continuous_closed_preimage[of s f "{a}"] closed_sing by auto + +lemma continuous_constant_on_closure: + assumes "continuous_on (closure s) f" + "\x \ s. f x = a" + shows "\x \ (closure s). f x = a" + using continuous_closed_preimage_constant[of "closure s" f a] + assms closure_minimal[of s "{x \ closure s. f x = a}"] closure_subset unfolding subset_eq by auto + +lemma image_closure_subset: + assumes "continuous_on (closure s) f" "closed t" "(f ` s) \ t" + shows "f ` (closure s) \ t" +proof- + have "s \ {x \ closure s. f x \ t}" using assms(3) closure_subset by auto + moreover have "closed {x \ closure s. f x \ t}" + using continuous_closed_preimage[OF assms(1)] and assms(2) by auto + ultimately have "closure s = {x \ closure s . f x \ t}" + using closure_minimal[of s "{x \ closure s. f x \ t}"] by auto + thus ?thesis by auto +qed + +lemma continuous_on_closure_norm_le: + fixes f :: "'a::metric_space \ 'b::real_normed_vector" + assumes "continuous_on (closure s) f" "\y \ s. norm(f y) \ b" "x \ (closure s)" + shows "norm(f x) \ b" +proof- + have *:"f ` s \ cball 0 b" using assms(2)[unfolded mem_cball_0[THEN sym]] by auto + show ?thesis + using image_closure_subset[OF assms(1) closed_cball[of 0 b] *] assms(3) + unfolding subset_eq apply(erule_tac x="f x" in ballE) by (auto simp add: dist_norm) +qed + +text{* Making a continuous function avoid some value in a neighbourhood. *} + +lemma continuous_within_avoid: + fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) + assumes "continuous (at x within s) f" "x \ s" "f x \ a" + shows "\e>0. \y \ s. dist x y < e --> f y \ a" +proof- + obtain d where "d>0" and d:"\xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < dist (f x) a" + using assms(1)[unfolded continuous_within Lim_within, THEN spec[where x="dist (f x) a"]] assms(3)[unfolded dist_nz] by auto + { fix y assume " y\s" "dist x y < d" + hence "f y \ a" using d[THEN bspec[where x=y]] assms(3)[unfolded dist_nz] + apply auto unfolding dist_nz[THEN sym] by (auto simp add: dist_commute) } + thus ?thesis using `d>0` by auto +qed + +lemma continuous_at_avoid: + fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) + assumes "continuous (at x) f" "f x \ a" + shows "\e>0. \y. dist x y < e \ f y \ a" +using assms using continuous_within_avoid[of x UNIV f a, unfolded within_UNIV] by auto + +lemma continuous_on_avoid: + assumes "continuous_on s f" "x \ s" "f x \ a" + shows "\e>0. \y \ s. dist x y < e \ f y \ a" +using assms(1)[unfolded continuous_on_eq_continuous_within, THEN bspec[where x=x], OF assms(2)] continuous_within_avoid[of x s f a] assms(2,3) by auto + +lemma continuous_on_open_avoid: + assumes "continuous_on s f" "open s" "x \ s" "f x \ a" + shows "\e>0. \y. dist x y < e \ f y \ a" +using assms(1)[unfolded continuous_on_eq_continuous_at[OF assms(2)], THEN bspec[where x=x], OF assms(3)] continuous_at_avoid[of x f a] assms(3,4) by auto + +text{* Proving a function is constant by proving open-ness of level set. *} + +lemma continuous_levelset_open_in_cases: + "connected s \ continuous_on s f \ + openin (subtopology euclidean s) {x \ s. f x = a} + ==> (\x \ s. f x \ a) \ (\x \ s. f x = a)" +unfolding connected_clopen using continuous_closed_in_preimage_constant by auto + +lemma continuous_levelset_open_in: + "connected s \ continuous_on s f \ + openin (subtopology euclidean s) {x \ s. f x = a} \ + (\x \ s. f x = a) ==> (\x \ s. f x = a)" +using continuous_levelset_open_in_cases[of s f ] +by meson + +lemma continuous_levelset_open: + assumes "connected s" "continuous_on s f" "open {x \ s. f x = a}" "\x \ s. f x = a" + shows "\x \ s. f x = a" +using continuous_levelset_open_in[OF assms(1,2), of a, unfolded openin_open] using assms (3,4) by auto + +text{* Some arithmetical combinations (more to prove). *} + +lemma open_scaling[intro]: + fixes s :: "'a::real_normed_vector set" + assumes "c \ 0" "open 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 *\<^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 *\<^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 + +lemma minus_image_eq_vimage: + fixes A :: "'a::ab_group_add set" + shows "(\x. - x) ` A = (\x. - x) -` A" + by (auto intro!: image_eqI [where f="\x. - x"]) + +lemma open_negations: + fixes s :: "'a::real_normed_vector set" + shows "open s ==> open ((\ x. -x) ` s)" + unfolding scaleR_minus1_left [symmetric] + by (rule open_scaling, auto) + +lemma open_translation: + 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 } + moreover have "{x. x - a \ s} = op + a ` s" apply auto unfolding image_iff apply(rule_tac x="x - a" in bexI) by auto + ultimately show ?thesis using continuous_open_preimage_univ[of "\x. x - a" s] using assms by auto +qed + +lemma open_affinity: + fixes s :: "'a::real_normed_vector set" + assumes "open s" "c \ 0" + shows "open ((\x. a + c *\<^sub>R x) ` s)" +proof- + 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: + fixes s :: "'a::real_normed_vector set" + shows "interior ((\x. a + x) ` s) = (\x. a + x) ` (interior s)" +proof (rule set_ext, rule) + fix x assume "x \ interior (op + a ` s)" + then obtain e where "e>0" and e:"ball x e \ op + a ` s" unfolding mem_interior by auto + hence "ball (x - a) e \ s" unfolding subset_eq Ball_def mem_ball dist_norm apply auto apply(erule_tac x="a + xa" in allE) unfolding ab_group_add_class.diff_diff_eq[THEN sym] by auto + thus "x \ op + a ` interior s" unfolding image_iff apply(rule_tac x="x - a" in bexI) unfolding mem_interior using `e > 0` by auto +next + fix x assume "x \ op + a ` interior s" + then obtain y e where "e>0" and e:"ball y e \ s" and y:"x = a + y" unfolding image_iff Bex_def mem_interior by auto + { fix z have *:"a + y - z = y + a - z" by auto + assume "z\ball x e" + hence "z - a \ s" using e[unfolded subset_eq, THEN bspec[where x="z - a"]] unfolding mem_ball dist_norm y ab_group_add_class.diff_diff_eq2 * by auto + hence "z \ op + a ` s" unfolding image_iff by(auto intro!: bexI[where x="z - a"]) } + hence "ball x e \ op + a ` s" unfolding subset_eq by auto + thus "x \ interior (op + a ` s)" unfolding mem_interior using `e>0` by auto +qed + +subsection {* Preservation of compactness and connectedness under continuous function. *} + +lemma compact_continuous_image: + assumes "continuous_on s f" "compact s" + shows "compact(f ` s)" +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:"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. subseq r \ ((x \ r) ---> l) sequentially" unfolding Lim_sequentially using r lr `l\s` by auto } + thus ?thesis unfolding compact_def by auto +qed + +lemma connected_continuous_image: + assumes "continuous_on s f" "connected s" + shows "connected(f ` s)" +proof- + { fix T assume as: "T \ {}" "T \ f ` s" "openin (subtopology euclidean (f ` s)) T" "closedin (subtopology euclidean (f ` s)) T" + have "{x \ s. f x \ T} = {} \ {x \ s. f x \ T} = s" + using assms(1)[unfolded continuous_on_open, THEN spec[where x=T]] + using assms(1)[unfolded continuous_on_closed, THEN spec[where x=T]] + using assms(2)[unfolded connected_clopen, THEN spec[where x="{x \ s. f x \ T}"]] as(3,4) by auto + hence False using as(1,2) + using as(4)[unfolded closedin_def topspace_euclidean_subtopology] by auto } + thus ?thesis unfolding connected_clopen by auto +qed + +text{* Continuity implies uniform continuity on a compact domain. *} + +lemma compact_uniformly_continuous: + assumes "continuous_on s f" "compact s" + shows "uniformly_continuous_on s f" +proof- + { fix x assume x:"x\s" + hence "\xa. \y. 0 < xa \ (y > 0 \ (\x'\s. dist x' x < y \ dist (f x') (f x) < xa))" using assms(1)[unfolded continuous_on_def, THEN bspec[where x=x]] by auto + hence "\fa. \xa>0. \x'\s. fa xa > 0 \ (dist x' x < fa xa \ dist (f x') (f x) < xa)" using choice[of "\e d. e>0 \ d>0 \(\x'\s. (dist x' x < d \ dist (f x') (f x) < e))"] by auto } + then have "\x\s. \y. \xa. 0 < xa \ (\x'\s. y xa > 0 \ (dist x' x < y xa \ dist (f x') (f x) < xa))" by auto + then obtain d where d:"\e>0. \x\s. \x'\s. d x e > 0 \ (dist x' x < d x e \ dist (f x') (f x) < e)" + using bchoice[of s "\x fa. \xa>0. \x'\s. fa xa > 0 \ (dist x' x < fa xa \ dist (f x') (f x) < xa)"] by blast + + { fix e::real assume "e>0" + + { fix x assume "x\s" hence "x \ ball x (d x (e / 2))" unfolding centre_in_ball using d[THEN spec[where x="e/2"]] using `e>0` by auto } + hence "s \ \{ball x (d x (e / 2)) |x. x \ s}" unfolding subset_eq by auto + moreover + { fix b assume "b\{ball x (d x (e / 2)) |x. x \ s}" hence "open b" by auto } + ultimately obtain ea where "ea>0" and ea:"\x\s. \b\{ball x (d x (e / 2)) |x. x \ s}. ball x ea \ b" using heine_borel_lemma[OF assms(2), of "{ball x (d x (e / 2)) | x. x\s }"] by auto + + { fix x y assume "x\s" "y\s" and as:"dist y x < ea" + obtain z where "z\s" and z:"ball x ea \ ball z (d z (e / 2))" using ea[THEN bspec[where x=x]] and `x\s` by auto + hence "x\ball z (d z (e / 2))" using `ea>0` unfolding subset_eq by auto + hence "dist (f z) (f x) < e / 2" using d[THEN spec[where x="e/2"]] and `e>0` and `x\s` and `z\s` + by (auto simp add: dist_commute) + moreover have "y\ball z (d z (e / 2))" using as and `ea>0` and z[unfolded subset_eq] + by (auto simp add: dist_commute) + hence "dist (f z) (f y) < e / 2" using d[THEN spec[where x="e/2"]] and `e>0` and `y\s` and `z\s` + by (auto simp add: dist_commute) + ultimately have "dist (f y) (f x) < e" using dist_triangle_half_r[of "f z" "f x" e "f y"] + by (auto simp add: dist_commute) } + then have "\d>0. \x\s. \x'\s. dist x' x < d \ dist (f x') (f x) < e" using `ea>0` by auto } + thus ?thesis unfolding uniformly_continuous_on_def by auto +qed + +text{* Continuity of inverse function on compact domain. *} + +lemma continuous_on_inverse: + 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- + have *:"g ` f ` s = s" using assms(3) by (auto simp add: image_iff) + { fix t assume t:"closedin (subtopology euclidean (g ` f ` s)) t" + then obtain T where T: "closed T" "t = s \ T" unfolding closedin_closed unfolding * by auto + have "continuous_on (s \ T) f" using continuous_on_subset[OF assms(1), of "s \ t"] + unfolding T(2) and Int_left_absorb by auto + moreover have "compact (s \ T)" + using assms(2) unfolding compact_eq_bounded_closed + using bounded_subset[of s "s \ T"] and T(1) by auto + ultimately have "closed (f ` t)" using T(1) unfolding T(2) + using compact_continuous_image [of "s \ T" f] unfolding compact_eq_bounded_closed by auto + moreover have "{x \ f ` s. g x \ t} = f ` s \ f ` t" using assms(3) unfolding T(2) by auto + ultimately have "closedin (subtopology euclidean (f ` s)) {x \ f ` s. g x \ t}" + unfolding closedin_closed by auto } + thus ?thesis unfolding continuous_on_closed by auto +qed + +subsection{* A uniformly convergent limit of continuous functions is continuous. *} + +lemma norm_triangle_lt: + fixes x y :: "'a::real_normed_vector" + shows "norm x + norm y < e \ norm (x + y) < e" +by (rule le_less_trans [OF norm_triangle_ineq]) + +lemma continuous_uniform_limit: + fixes f :: "'a \ 'b::metric_space \ 'c::real_normed_vector" + assumes "\ (trivial_limit net)" "eventually (\n. continuous_on s (f n)) net" + "\e>0. eventually (\n. \x \ s. norm(f n x - g x) < e) net" + shows "continuous_on s g" +proof- + { fix x and e::real assume "x\s" "e>0" + have "eventually (\n. \x\s. norm (f n x - g x) < e / 3) net" using `e>0` assms(3)[THEN spec[where x="e/3"]] by auto + then obtain n where n:"\xa\s. norm (f n xa - g xa) < e / 3" "continuous_on s (f n)" + using eventually_and[of "(\n. \x\s. norm (f n x - g x) < e / 3)" "(\n. continuous_on s (f n))" net] assms(1,2) eventually_happens by blast + have "e / 3 > 0" using `e>0` by auto + then obtain d where "d>0" and d:"\x'\s. dist x' x < d \ dist (f n x') (f n x) < e / 3" + using n(2)[unfolded continuous_on_def, THEN bspec[where x=x], OF `x\s`, THEN spec[where x="e/3"]] by blast + { fix y assume "y\s" "dist y x < d" + hence "dist (f n y) (f n x) < e / 3" using d[THEN bspec[where x=y]] by auto + hence "norm (f n y - g x) < 2 * e / 3" using norm_triangle_lt[of "f n y - f n x" "f n x - g x" "2*e/3"] + using n(1)[THEN bspec[where x=x], OF `x\s`] unfolding dist_norm unfolding ab_group_add_class.ab_diff_minus by auto + hence "dist (g y) (g x) < e" unfolding dist_norm using n(1)[THEN bspec[where x=y], OF `y\s`] + unfolding norm_minus_cancel[of "f n y - g y", THEN sym] using norm_triangle_lt[of "f n y - g x" "g y - f n y" e] by (auto simp add: uminus_add_conv_diff) } + hence "\d>0. \x'\s. dist x' x < d \ dist (g x') (g x) < e" using `d>0` by auto } + thus ?thesis unfolding continuous_on_def by auto +qed + +subsection{* Topological properties of linear functions. *} + +lemma linear_lim_0: + assumes "bounded_linear f" shows "(f ---> 0) (at (0))" +proof- + interpret f: bounded_linear f by fact + have "(f ---> f 0) (at 0)" + using tendsto_ident_at by (rule f.tendsto) + thus ?thesis unfolding f.zero . +qed + +lemma linear_continuous_at: + assumes "bounded_linear f" shows "continuous (at a) f" + unfolding continuous_at using assms + apply (rule bounded_linear.tendsto) + apply (rule tendsto_ident_at) + done + +lemma linear_continuous_within: + shows "bounded_linear f ==> continuous (at x within s) f" + using continuous_at_imp_continuous_within[of x f s] using linear_continuous_at[of f] by auto + +lemma linear_continuous_on: + shows "bounded_linear f ==> continuous_on s f" + using continuous_at_imp_continuous_on[of s f] using linear_continuous_at[of f] by auto + +text{* Also bilinear functions, in composition form. *} + +lemma bilinear_continuous_at_compose: + shows "continuous (at x) f \ continuous (at x) g \ bounded_bilinear h + ==> continuous (at x) (\x. h (f x) (g x))" + unfolding continuous_at using Lim_bilinear[of f "f x" "(at x)" g "g x" h] by auto + +lemma bilinear_continuous_within_compose: + shows "continuous (at x within s) f \ continuous (at x within s) g \ bounded_bilinear h + ==> continuous (at x within s) (\x. h (f x) (g x))" + unfolding continuous_within using Lim_bilinear[of f "f x"] by auto + +lemma bilinear_continuous_on_compose: + shows "continuous_on s f \ continuous_on s g \ bounded_bilinear h + ==> continuous_on s (\x. h (f x) (g x))" + unfolding continuous_on_eq_continuous_within apply auto apply(erule_tac x=x in ballE) apply auto apply(erule_tac x=x in ballE) apply auto + using bilinear_continuous_within_compose[of _ s f g h] by auto + +subsection{* Topological stuff lifted from and dropped to R *} + + +lemma open_real: + fixes s :: "real set" shows + "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 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 + 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_real_range: + fixes f :: "'a::real_normed_vector \ real" + 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)) + +text{* Hence some handy theorems on distance, diameter etc. of/from a set. *} + +lemma compact_attains_sup: + fixes s :: "real set" + assumes "compact s" "s \ {}" + shows "\x \ s. \y \ s. y \ x" +proof- + 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_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 s" "s \ {}" shows "\x \ s. \y \ s. x \ y" +proof- + 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 + moreover + { fix x assume "x \ s" + hence *:"abs (x - rinf s) = x - rinf s" using as(1)[THEN bspec[where x=x]] by auto + 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_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 f + ==> (\x \ s. \y \ s. f y \ f x)" + using compact_attains_sup[of "f ` s"] + 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 f + \ (\x \ s. \y \ s. f x \ f y)" + using compact_attains_inf[of "f ` s"] + using compact_continuous_image[of s f] by auto + +lemma distance_attains_sup: + assumes "compact s" "s \ {}" + shows "\x \ s. \y \ s. dist a y \ dist a x" +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 :: "'a::heine_borel" + assumes "closed s" "s \ {}" + shows "\x \ s. \y \ s. dist a x \ dist a y" +proof- + from assms(2) obtain b where "b\s" by auto + let ?B = "cball a (dist b a) \ s" + have "b \ ?B" using `b\s` by (simp add: dist_commute) + hence "?B \ {}" by auto + moreover + { fix x assume "x\?B" + fix e::real assume "e>0" + { 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 + +subsection{* We can now extend limit compositions to consider the scalar multiplier. *} + +lemma Lim_mul: + fixes f :: "'a \ 'b::real_normed_vector" + assumes "(c ---> d) net" "(f ---> l) net" + shows "((\x. c(x) *\<^sub>R f x) ---> (d *\<^sub>R l)) net" + using assms by (rule scaleR.tendsto) + +lemma Lim_vmul: + 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" and v :: "'b::real_normed_vector" + shows "continuous net c ==> continuous net (\x. c(x) *\<^sub>R v)" + unfolding continuous_def using Lim_vmul[of c] by auto + +lemma continuous_mul: + fixes c :: "'a::metric_space \ real" + fixes f :: "'a::metric_space \ 'b::real_normed_vector" + shows "continuous net c \ continuous net f + ==> continuous net (\x. c(x) *\<^sub>R f x) " + unfolding continuous_def by (intro tendsto_intros) + +lemma continuous_on_vmul: + fixes c :: "'a::metric_space \ real" and v :: "'b::real_normed_vector" + shows "continuous_on s c ==> continuous_on s (\x. c(x) *\<^sub>R v)" + unfolding continuous_on_eq_continuous_within using continuous_vmul[of _ c] by auto + +lemma continuous_on_mul: + fixes c :: "'a::metric_space \ real" + fixes f :: "'a::metric_space \ 'b::real_normed_vector" + shows "continuous_on s c \ continuous_on s f + ==> continuous_on s (\x. c(x) *\<^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 "(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 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 :: "'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 :: "'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. *} + +lemma bounded_pastecart: + fixes s :: "('a::real_normed_vector ^ _) set" (* FIXME: generalize to metric_space *) + assumes "bounded s" "bounded t" + shows "bounded { pastecart x y | x y . (x \ s \ y \ t)}" +proof- + obtain a b where ab:"\x\s. norm x \ a" "\x\t. norm x \ b" using assms[unfolded bounded_iff] by auto + { fix x y assume "x\s" "y\t" + hence "norm x \ a" "norm y \ b" using ab by auto + hence "norm (pastecart x y) \ a + b" using norm_pastecart[of x y] by auto } + 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" + shows "closed {pastecart x y | x y . x \ s \ y \ t}" +proof- + { fix x l assume as:"\n::nat. x n \ {pastecart x y |x y. x \ s \ y \ t}" "(x ---> l) sequentially" + { fix n::nat have "fstcart (x n) \ s" "sndcart (x n) \ t" using as(1)[THEN spec[where x=n]] by auto } note * = this + moreover + { fix e::real assume "e>0" + then obtain N::nat where N:"\n\N. dist (x n) l < e" using as(2)[unfolded Lim_sequentially, THEN spec[where x=e]] by auto + { fix n::nat assume "n\N" + hence "dist (fstcart (x n)) (fstcart l) < e" "dist (sndcart (x n)) (sndcart l) < e" + using N[THEN spec[where x=n]] dist_fstcart[of "x n" l] dist_sndcart[of "x n" l] by auto } + hence "\N. \n\N. dist (fstcart (x n)) (fstcart l) < e" "\N. \n\N. dist (sndcart (x n)) (sndcart l) < e" by auto } + ultimately have "fstcart l \ s" "sndcart l \ t" + using assms(1)[unfolded closed_sequential_limits, THEN spec[where x="\n. fstcart (x n)"], THEN spec[where x="fstcart l"]] + using assms(2)[unfolded closed_sequential_limits, THEN spec[where x="\n. sndcart (x n)"], THEN spec[where x="sndcart l"]] + unfolding Lim_sequentially by auto + hence "l \ {pastecart x y |x y. x \ s \ y \ t}" using pastecart_fst_snd[THEN sym, of l] by auto } + thus ?thesis unfolding closed_sequential_limits by auto +qed + +lemma compact_pastecart: + fixes s t :: "(real ^ _) set" + 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_scaling: + fixes s :: "'a::real_normed_vector set" + assumes "compact s" shows "compact ((\x. c *\<^sub>R x) ` s)" +proof- + let ?f = "\x. scaleR c x" + have *:"bounded_linear ?f" by (rule scaleR.bounded_linear_right) + show ?thesis using compact_continuous_image[of s ?f] continuous_at_imp_continuous_on[of s ?f] + using linear_continuous_at[OF *] assms by auto +qed + +lemma compact_negations: + fixes s :: "'a::real_normed_vector set" + assumes "compact s" shows "compact ((\x. -x) ` s)" + using compact_scaling [OF assms, of "- 1"] by auto + +lemma compact_sums: + 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. 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 :: "'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} = {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 :: "'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 + thus ?thesis using compact_sums[OF assms compact_sing[of a]] by auto +qed + +lemma compact_affinity: + fixes s :: "'a::real_normed_vector set" + assumes "compact s" shows "compact ((\x. a + c *\<^sub>R x) ` s)" +proof- + 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 :: "'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[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 + +text{* We can state this in terms of diameter of a set. *} + +definition "diameter s = (if s = {} then 0::real else rsup {norm(x - y) | x y. x \ s \ y \ s})" + (* TODO: generalize to class metric_space *) + +lemma diameter_bounded: + assumes "bounded s" + shows "\x\s. \y\s. norm(x - y) \ diameter s" + "\d>0. d < diameter s --> (\x\s. \y\s. norm(x - y) > d)" +proof- + let ?D = "{norm (x - y) |x y. x \ s \ y \ s}" + obtain a where a:"\x\s. norm x \ a" using assms[unfolded bounded_iff] by auto + { fix x y assume "x \ s" "y \ s" + hence "norm (x - y) \ 2 * a" using norm_triangle_ineq[of x "-y", unfolded norm_minus_cancel] a[THEN bspec[where x=x]] a[THEN bspec[where x=y]] by (auto simp add: ring_simps) } + note * = this + { fix x y assume "x\s" "y\s" hence "s \ {}" by auto + have lub:"isLub UNIV ?D (rsup ?D)" using * rsup[of ?D] using `s\{}` unfolding setle_def by auto + have "norm(x - y) \ diameter s" unfolding diameter_def using `s\{}` *[OF `x\s` `y\s`] `x\s` `y\s` isLubD1[OF lub] unfolding setle_def by auto } + moreover + { fix d::real assume "d>0" "d < diameter s" + hence "s\{}" unfolding diameter_def by auto + hence lub:"isLub UNIV ?D (rsup ?D)" using * rsup[of ?D] unfolding setle_def by auto + have "\d' \ ?D. d' > d" + proof(rule ccontr) + assume "\ (\d'\{norm (x - y) |x y. x \ s \ y \ s}. d < d')" + hence as:"\d'\?D. d' \ d" apply auto apply(erule_tac x="norm (x - y)" in allE) by auto + hence "isUb UNIV ?D d" unfolding isUb_def unfolding setle_def by auto + thus False using `d < diameter s` `s\{}` isLub_le_isUb[OF lub, of d] unfolding diameter_def by auto + qed + hence "\x\s. \y\s. norm(x - y) > d" by auto } + ultimately show "\x\s. \y\s. norm(x - y) \ diameter s" + "\d>0. d < diameter s --> (\x\s. \y\s. norm(x - y) > d)" by auto +qed + +lemma diameter_bounded_bound: + "bounded s \ x \ s \ y \ s ==> norm(x - y) \ diameter s" + using diameter_bounded by blast + +lemma diameter_compact_attained: + 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) 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 + thus ?thesis using diameter_bounded(1)[OF b, THEN bspec[where x=x], THEN bspec[where x=y], OF xys] and xys by auto +qed + +text{* Related results with closure as the conclusion. *} + +lemma closed_scaling: + fixes s :: "'a::real_normed_vector set" + assumes "closed s" shows "closed ((\x. c *\<^sub>R x) ` s)" +proof(cases "s={}") + case True thus ?thesis by auto +next + case False + show ?thesis + proof(cases "c=0") + have *:"(\x. 0) ` s = {0}" using `s\{}` by auto + case True thus ?thesis apply auto unfolding * using closed_sing by auto + next + case False + { fix x l assume as:"\n::nat. x n \ scaleR c ` s" "(x ---> l) sequentially" + { fix n::nat have "scaleR (1 / c) (x n) \ s" + using as(1)[THEN spec[where x=n]] + using `c\0` by (auto simp add: vector_smult_assoc) + } + moreover + { fix e::real assume "e>0" + hence "0 < e *\c\" using `c\0` mult_pos_pos[of e "abs c"] by auto + 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] + 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" + using assms[unfolded closed_sequential_limits, THEN spec[where x="\n. scaleR (1/c) (x n)"], THEN spec[where x="scaleR (1/c) l"]] + unfolding image_iff using `c\0` apply(rule_tac x="scaleR (1 / c) l" in bexI) by auto } + thus ?thesis unfolding closed_sequential_limits by fast + qed +qed + +lemma closed_negations: + fixes s :: "'a::real_normed_vector set" + assumes "closed s" shows "closed ((\x. -x) ` s)" + using closed_scaling[OF assms, of "- 1"] by simp + +lemma compact_closed_sums: + fixes s :: "'a::real_normed_vector set" + assumes "compact s" "closed t" shows "closed {x + y | x y. x \ s \ y \ t}" +proof- + let ?S = "{x + y |x y. x \ s \ y \ t}" + { 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:"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_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 + hence "l \ ?S" using `l' \ s` apply auto apply(rule_tac x=l' in exI) apply(rule_tac x="l - l'" in exI) by auto + } + thus ?thesis unfolding closed_sequential_limits by fast +qed + +lemma closed_compact_sums: + fixes s t :: "'a::real_normed_vector set" + assumes "closed s" "compact t" + shows "closed {x + y | x y. x \ s \ y \ t}" +proof- + have "{x + y |x y. x \ t \ y \ s} = {x + y |x y. x \ s \ y \ t}" apply auto + apply(rule_tac x=y in exI) apply auto apply(rule_tac x=y in exI) by auto + thus ?thesis using compact_closed_sums[OF assms(2,1)] by simp +qed + +lemma compact_closed_differences: + fixes s t :: "'a::real_normed_vector set" + assumes "compact s" "closed t" + shows "closed {x - y | x y. x \ s \ y \ t}" +proof- + have "{x + y |x y. x \ s \ y \ uminus ` t} = {x - y |x y. x \ s \ y \ 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_closed_sums[OF assms(1) closed_negations[OF assms(2)]] by auto +qed + +lemma closed_compact_differences: + fixes s t :: "'a::real_normed_vector set" + assumes "closed s" "compact t" + shows "closed {x - y | x y. x \ s \ y \ t}" +proof- + have "{x + y |x y. x \ s \ y \ uminus ` t} = {x - y |x y. x \ s \ y \ t}" + apply auto apply(rule_tac x=xa in exI) apply auto apply(rule_tac x=xa in exI) by auto + thus ?thesis using closed_compact_sums[OF assms(1) compact_negations[OF assms(2)]] by simp +qed + +lemma closed_translation: + fixes a :: "'a::real_normed_vector" + assumes "closed s" shows "closed ((\x. a + x) ` s)" +proof- + have "{a + y |y. y \ s} = (op + a ` s)" by auto + thus ?thesis using compact_closed_sums[OF compact_sing[of a] assms] by auto +qed + +lemma translation_UNIV: + fixes a :: "'a::ab_group_add" shows "range (\x. a + x) = UNIV" + apply (auto simp add: image_iff) apply(rule_tac x="x - a" in exI) by auto + +lemma translation_diff: + fixes a :: "'a::ab_group_add" + shows "(\x. a + x) ` (s - t) = ((\x. a + x) ` s) - ((\x. a + x) ` t)" + by auto + +lemma closure_translation: + fixes a :: "'a::real_normed_vector" + shows "closure ((\x. a + x) ` s) = (\x. a + x) ` (closure s)" +proof- + have *:"op + a ` (UNIV - s) = UNIV - op + a ` s" + apply auto unfolding image_iff apply(rule_tac x="x - a" in bexI) by auto + show ?thesis unfolding closure_interior translation_diff translation_UNIV + using interior_translation[of a "UNIV - s"] unfolding * by auto +qed + +lemma frontier_translation: + fixes a :: "'a::real_normed_vector" + shows "frontier((\x. a + x) ` s) = (\x. a + x) ` (frontier s)" + unfolding frontier_def translation_diff interior_translation closure_translation by auto + +subsection{* Separation between points and sets. *} + +lemma separate_point_closed: + fixes s :: "'a::heine_borel set" + shows "closed s \ a \ s ==> (\d>0. \x\s. d \ dist a x)" +proof(cases "s = {}") + case True + thus ?thesis by(auto intro!: exI[where x=1]) +next + case False + assume "closed s" "a \ s" + then obtain x where "x\s" "\y\s. dist a x \ dist a y" using `s \ {}` distance_attains_inf [of s a] by blast + with `x\s` show ?thesis using dist_pos_lt[of a x] and`a \ s` by blast +qed + +lemma separate_compact_closed: + 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- + have "0 \ {x - y |x y. x \ s \ y \ t}" using assms(3) by auto + then obtain d where "d>0" and d:"\x\{x - y |x y. x \ s \ y \ t}. d \ dist 0 x" + using separate_point_closed[OF compact_closed_differences[OF assms(1,2)], of 0] by auto + { fix x y assume "x\s" "y\t" + hence "x - y \ {x - y |x y. x \ s \ y \ t}" by auto + hence "d \ dist (x - y) 0" using d[THEN bspec[where x="x - y"]] using dist_commute + by (auto simp add: dist_commute) + hence "d \ dist x y" unfolding dist_norm by auto } + thus ?thesis using `d>0` by auto +qed + +lemma separate_closed_compact: + 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- + have *:"t \ s = {}" using assms(3) by auto + show ?thesis using separate_compact_closed[OF assms(2,1) *] + apply auto apply(rule_tac x=d in exI) apply auto apply (erule_tac x=y in ballE) + by (auto simp add: dist_commute) +qed + +(* A cute way of denoting open and closed intervals using overloading. *) + +lemma interval: fixes a :: "'a::ord^'n::finite" shows + "{a <..< b} = {x::'a^'n. \i. a$i < x$i \ x$i < b$i}" and + "{a .. b} = {x::'a^'n. \i. a$i \ x$i \ x$i \ b$i}" + by (auto simp add: expand_set_eq vector_less_def vector_less_eq_def) + +lemma mem_interval: fixes a :: "'a::ord^'n::finite" shows + "x \ {a<.. (\i. a$i < x$i \ x$i < b$i)" + "x \ {a .. b} \ (\i. a$i \ x$i \ x$i \ b$i)" + using interval[of a b] by(auto simp add: expand_set_eq vector_less_def vector_less_eq_def) + +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 forall_1) + +lemma interval_eq_empty: fixes a :: "real^'n::finite" shows + "({a <..< b} = {} \ (\i. b$i \ a$i))" (is ?th1) and + "({a .. b} = {} \ (\i. b$i < a$i))" (is ?th2) +proof- + { fix i x assume as:"b$i \ a$i" and x:"x\{a <..< b}" + hence "a $ i < x $ i \ x $ i < b $ i" unfolding mem_interval by auto + hence "a$i < b$i" by auto + hence False using as by auto } + moreover + { assume as:"\i. \ (b$i \ a$i)" + 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) *\<^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 } + ultimately show ?th1 by blast + + { fix i x assume as:"b$i < a$i" and x:"x\{a .. b}" + hence "a $ i \ x $ i \ x $ i \ b $ i" unfolding mem_interval by auto + hence "a$i \ b$i" by auto + hence False using as by auto } + moreover + { assume as:"\i. \ (b$i < a$i)" + 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) *\<^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 } + ultimately show ?th2 by blast +qed + +lemma interval_ne_empty: fixes a :: "real^'n::finite" shows + "{a .. b} \ {} \ (\i. a$i \ b$i)" and + "{a <..< b} \ {} \ (\i. a$i < b$i)" + unfolding interval_eq_empty[of a b] by (auto simp add: not_less not_le) (* BH: Why doesn't just "auto" work here? *) + +lemma subset_interval_imp: fixes a :: "real^'n::finite" shows + "(\i. a$i \ c$i \ d$i \ b$i) \ {c .. d} \ {a .. b}" and + "(\i. a$i < c$i \ d$i < b$i) \ {c .. d} \ {a<..i. a$i \ c$i \ d$i \ b$i) \ {c<.. {a .. b}" and + "(\i. a$i \ c$i \ d$i \ b$i) \ {c<.. {a<.. {a<.. {a .. b}" +proof(simp add: subset_eq, rule) + fix x + assume x:"x \{a<.. x $ i" + using x order_less_imp_le[of "a$i" "x$i"] + by(simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq) + } + moreover + { fix i + have "x $ i \ b $ i" + using x order_less_imp_le[of "x$i" "b$i"] + by(simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq) + } + ultimately + show "a \ x \ x \ b" + by(simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq) +qed + +lemma subset_interval: fixes a :: "real^'n::finite" shows + "{c .. d} \ {a .. b} \ (\i. c$i \ d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th1) and + "{c .. d} \ {a<.. (\i. c$i \ d$i) --> (\i. a$i < c$i \ d$i < b$i)" (is ?th2) and + "{c<.. {a .. b} \ (\i. c$i < d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th3) and + "{c<.. {a<.. (\i. c$i < d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th4) +proof- + show ?th1 unfolding subset_eq and Ball_def and mem_interval by (auto intro: order_trans) + show ?th2 unfolding subset_eq and Ball_def and mem_interval by (auto intro: le_less_trans less_le_trans order_trans less_imp_le) + { assume as: "{c<.. {a .. b}" "\i. c$i < d$i" + hence "{c<.. {}" unfolding interval_eq_empty by (auto, drule_tac x=i in spec, simp) (* BH: Why doesn't just "auto" work? *) + fix i + (** TODO combine the following two parts as done in the HOL_light version. **) + { let ?x = "(\ j. (if j=i then ((min (a$j) (d$j))+c$j)/2 else (c$j+d$j)/2))::real^'n" + assume as2: "a$i > c$i" + { fix j + have "c $ j < ?x $ j \ ?x $ j < d $ j" unfolding Cart_lambda_beta + apply(cases "j=i") using as(2)[THEN spec[where x=j]] + by (auto simp add: less_divide_eq_number_of1 as2) } + hence "?x\{c<..{a .. b}" + unfolding mem_interval apply auto apply(rule_tac x=i in exI) + using as(2)[THEN spec[where x=i]] and as2 + by (auto simp add: less_divide_eq_number_of1) + ultimately have False using as by auto } + hence "a$i \ c$i" by(rule ccontr)auto + moreover + { let ?x = "(\ j. (if j=i then ((max (b$j) (c$j))+d$j)/2 else (c$j+d$j)/2))::real^'n" + assume as2: "b$i < d$i" + { fix j + have "d $ j > ?x $ j \ ?x $ j > c $ j" unfolding Cart_lambda_beta + apply(cases "j=i") using as(2)[THEN spec[where x=j]] + by (auto simp add: less_divide_eq_number_of1 as2) } + hence "?x\{c<..{a .. b}" + unfolding mem_interval apply auto apply(rule_tac x=i in exI) + using as(2)[THEN spec[where x=i]] and as2 + by (auto simp add: less_divide_eq_number_of1) + ultimately have False using as by auto } + hence "b$i \ d$i" by(rule ccontr)auto + ultimately + have "a$i \ c$i \ d$i \ b$i" by auto + } note part1 = this + thus ?th3 unfolding subset_eq and Ball_def and mem_interval apply auto apply (erule_tac x=ia in allE, simp)+ by (erule_tac x=i in allE, erule_tac x=i in allE, simp)+ + { assume as:"{c<.. {a<..i. c$i < d$i" + fix i + from as(1) have "{c<.. {a..b}" using interval_open_subset_closed[of a b] by auto + hence "a$i \ c$i \ d$i \ b$i" using part1 and as(2) by auto } note * = this + thus ?th4 unfolding subset_eq and Ball_def and mem_interval apply auto apply (erule_tac x=ia in allE, simp)+ by (erule_tac x=i in allE, erule_tac x=i in allE, simp)+ +qed + +lemma disjoint_interval: fixes a::"real^'n::finite" shows + "{a .. b} \ {c .. d} = {} \ (\i. (b$i < a$i \ d$i < c$i \ b$i < c$i \ d$i < a$i))" (is ?th1) and + "{a .. b} \ {c<.. (\i. (b$i < a$i \ d$i \ c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th2) and + "{a<.. {c .. d} = {} \ (\i. (b$i \ a$i \ d$i < c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th3) and + "{a<.. {c<.. (\i. (b$i \ a$i \ d$i \ c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th4) +proof- + let ?z = "(\ i. ((max (a$i) (c$i)) + (min (b$i) (d$i))) / 2)::real^'n" + show ?th1 ?th2 ?th3 ?th4 + unfolding expand_set_eq and Int_iff and empty_iff and mem_interval and all_conj_distrib[THEN sym] and eq_False + apply (auto elim!: allE[where x="?z"]) + apply ((rule_tac x=x in exI, force) | (rule_tac x=i in exI, force))+ + done +qed + +lemma inter_interval: fixes a :: "'a::linorder^'n::finite" shows + "{a .. b} \ {c .. d} = {(\ i. max (a$i) (c$i)) .. (\ i. min (b$i) (d$i))}" + unfolding expand_set_eq and Int_iff and mem_interval + by (auto simp add: less_divide_eq_number_of1 intro!: bexI) + +(* Moved interval_open_subset_closed a bit upwards *) + +lemma open_interval_lemma: fixes x :: "real" shows + "a < x \ x < b ==> (\d>0. \x'. abs(x' - x) < d --> a < x' \ x' < b)" + by(rule_tac x="min (x - a) (b - x)" in exI, auto) + +lemma open_interval: fixes a :: "real^'n::finite" shows "open {a<..{a<..d>0. \x'. abs (x' - (x$i)) < d \ a$i < x' \ x' < b$i" + using x[unfolded mem_interval, THEN spec[where x=i]] + using open_interval_lemma[of "a$i" "x$i" "b$i"] by auto } + + hence "\i. \d>0. \x'. abs (x' - (x$i)) < d \ a$i < x' \ x' < b$i" by auto + then obtain d where d:"\i. 0 < d i \ (\x'. \x' - x $ i\ < d i \ a $ i < x' \ x' < b $ i)" + using bchoice[of "UNIV" "\i d. d>0 \ (\x'. \x' - x $ i\ < d \ a $ i < x' \ x' < b $ i)"] by auto + + let ?d = "Min (range d)" + have **:"finite (range d)" "range d \ {}" by auto + have "?d>0" unfolding Min_gr_iff[OF **] using d by auto + moreover + { fix x' assume as:"dist x' x < ?d" + { fix i + have "\x'$i - x $ i\ < d i" + using norm_bound_component_lt[OF as[unfolded dist_norm], of i] + unfolding vector_minus_component and Min_gr_iff[OF **] by auto + hence "a $ i < x' $ i" "x' $ i < b $ i" using d[THEN spec[where x=i]] by auto } + hence "a < x' \ x' < b" unfolding vector_less_def by auto } + ultimately have "\e>0. \x'. dist x' x < e \ x' \ {a<..e>0. \x'\{a..b}. x' \ x \ dist x' x < e"(* and xab:"a$i > x$i \ b$i < x$i"*) + { assume xa:"a$i > x$i" + with as obtain y where y:"y\{a..b}" "y \ x" "dist y x < a$i - x$i" by(erule_tac x="a$i - x$i" in allE)auto + hence False unfolding mem_interval and dist_norm + using component_le_norm[of "y-x" i, unfolded vector_minus_component] and xa by(auto elim!: allE[where x=i]) + } hence "a$i \ x$i" by(rule ccontr)auto + moreover + { assume xb:"b$i < x$i" + with as obtain y where y:"y\{a..b}" "y \ x" "dist y x < x$i - b$i" by(erule_tac x="x$i - b$i" in allE)auto + hence False unfolding mem_interval and dist_norm + using component_le_norm[of "y-x" i, unfolded vector_minus_component] and xb by(auto elim!: allE[where x=i]) + } hence "x$i \ b$i" by(rule ccontr)auto + ultimately + have "a $ i \ x $ i \ x $ i \ b $ i" by auto } + thus ?thesis unfolding closed_limpt islimpt_approachable mem_interval by auto +qed + +lemma interior_closed_interval: fixes a :: "real^'n::finite" shows + "interior {a .. b} = {a<.. ?L" using interior_maximal[OF interval_open_subset_closed open_interval] by auto +next + { fix x assume "\T. open T \ x \ T \ T \ {a..b}" + 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) *\<^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 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 + unfolding vector_smult_component and basis_component using `e>0` by auto } + hence "x \ {a<.. ?R" unfolding interior_def and subset_eq by auto +qed + +lemma bounded_closed_interval: fixes a :: "real^'n::finite" shows + "bounded {a .. b}" +proof- + let ?b = "\i\UNIV. \a$i\ + \b$i\" + { fix x::"real^'n" assume x:"\i. a $ i \ x $ i \ x $ i \ b $ i" + { fix i + have "\x$i\ \ \a$i\ + \b$i\" using x[THEN spec[where x=i]] by auto } + hence "(\i\UNIV. \x $ i\) \ ?b" by(rule setsum_mono) + hence "norm x \ ?b" using norm_le_l1[of x] by auto } + thus ?thesis unfolding interval and bounded_iff by auto +qed + +lemma bounded_interval: fixes a :: "real^'n::finite" shows + "bounded {a .. b} \ bounded {a<.. UNIV) \ ({a<.. UNIV)" + using bounded_interval[of a b] + by auto + +lemma compact_interval: fixes a :: "real^'n::finite" shows + "compact {a .. b}" + 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) *\<^sub>R (a + b)) \ {a<..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) } + thus ?thesis unfolding mem_interval by auto +qed + +lemma open_closed_interval_convex: fixes x :: "real^'n::finite" + assumes x:"x \ {a<.. {a .. b}" and e:"0 < e" "e \ 1" + shows "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \ {a<.. < e * x $ i + (1 - e) * y $ i" apply(rule add_less_le_mono) + using e unfolding mult_less_cancel_left and mult_le_cancel_left apply simp_all + using x unfolding mem_interval apply simp + using y unfolding mem_interval apply simp + done + finally have "a $ i < (e *\<^sub>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) + using e unfolding mult_less_cancel_left and mult_le_cancel_left apply simp_all + using x unfolding mem_interval apply simp + using y unfolding mem_interval apply simp + done + 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 + +lemma closure_open_interval: fixes a :: "real^'n::finite" + assumes "{a<.. {}" + shows "closure {a<.. {a .. b}" + 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)) *\<^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 + { assume "\ (f ---> x) sequentially" + { fix e::real assume "e>0" + hence "\N::nat. inverse (real (N + 1)) < e" using real_arch_inv[of e] apply (auto simp add: Suc_pred') apply(rule_tac x="n - 1" in exI) by auto + 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 "((\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)) *\<^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<..a. s \ {-a<..0" and b:"\x\s. norm x \ b" using assms[unfolded bounded_pos] by auto + def a \ "(\ i. b+1)::real^'n" + { fix x assume "x\s" + fix i + have "(-a)$i < x$i" and "x$i < a$i" using b[THEN bspec[where x=x], OF `x\s`] and component_le_norm[of x i] + unfolding vector_uminus_component and a_def and Cart_lambda_beta by auto + } + thus ?thesis by(auto intro: exI[where x=a] simp add: vector_less_def) +qed + +lemma bounded_subset_open_interval: + fixes s :: "(real ^ 'n::finite) set" + shows "bounded s ==> (\a b. s \ {a<..a. s \ {-a .. a}" +proof- + obtain a where "s \ {- a<.. (\a b. s \ {a .. b})" + using bounded_subset_closed_interval_symmetric[of s] by auto + +lemma frontier_closed_interval: + fixes a b :: "real ^ _" + shows "frontier {a .. b} = {a .. b} - {a<.. {}" shows "{a<.. {c .. d} = {} \ {a<.. {c<..x::1. P x) \ P 1" + by (metis num1_eq_iff) + +lemma ex_1: "(\x::1. P x) \ P 1" + by auto (metis num1_eq_iff) + +lemma interval_cases_1: fixes x :: "real^1" shows + "x \ {a .. b} ==> x \ {a<.. (x = a) \ (x = b)" + by(simp add: Cart_eq vector_less_def vector_less_eq_def all_1, auto) + +lemma in_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 add: Cart_eq vector_less_def vector_less_eq_def all_1 dest_vec1_def) + +lemma interval_eq_empty_1: fixes a :: "real^1" shows + "{a .. b} = {} \ dest_vec1 b < dest_vec1 a" + "{a<.. dest_vec1 b \ dest_vec1 a" + unfolding interval_eq_empty and ex_1 and dest_vec1_def by auto + +lemma subset_interval_1: fixes a :: "real^1" shows + "({a .. b} \ {c .. d} \ dest_vec1 b < dest_vec1 a \ + dest_vec1 c \ dest_vec1 a \ dest_vec1 a \ dest_vec1 b \ dest_vec1 b \ dest_vec1 d)" + "({a .. b} \ {c<.. dest_vec1 b < dest_vec1 a \ + dest_vec1 c < dest_vec1 a \ dest_vec1 a \ dest_vec1 b \ dest_vec1 b < dest_vec1 d)" + "({a<.. {c .. d} \ dest_vec1 b \ dest_vec1 a \ + dest_vec1 c \ dest_vec1 a \ dest_vec1 a < dest_vec1 b \ dest_vec1 b \ dest_vec1 d)" + "({a<.. {c<.. dest_vec1 b \ dest_vec1 a \ + dest_vec1 c \ dest_vec1 a \ dest_vec1 a < dest_vec1 b \ dest_vec1 b \ dest_vec1 d)" + unfolding subset_interval[of a b c d] unfolding all_1 and dest_vec1_def by auto + +lemma eq_interval_1: fixes a :: "real^1" shows + "{a .. b} = {c .. d} \ + dest_vec1 b < dest_vec1 a \ dest_vec1 d < dest_vec1 c \ + dest_vec1 a = dest_vec1 c \ dest_vec1 b = dest_vec1 d" +using set_eq_subset[of "{a .. b}" "{c .. d}"] +using subset_interval_1(1)[of a b c d] +using subset_interval_1(1)[of c d a b] +by auto (* FIXME: slow *) + +lemma disjoint_interval_1: fixes a :: "real^1" shows + "{a .. b} \ {c .. d} = {} \ dest_vec1 b < dest_vec1 a \ dest_vec1 d < dest_vec1 c \ dest_vec1 b < dest_vec1 c \ dest_vec1 d < dest_vec1 a" + "{a .. b} \ {c<.. dest_vec1 b < dest_vec1 a \ dest_vec1 d \ dest_vec1 c \ dest_vec1 b \ dest_vec1 c \ dest_vec1 d \ dest_vec1 a" + "{a<.. {c .. d} = {} \ dest_vec1 b \ dest_vec1 a \ dest_vec1 d < dest_vec1 c \ dest_vec1 b \ dest_vec1 c \ dest_vec1 d \ dest_vec1 a" + "{a<.. {c<.. dest_vec1 b \ dest_vec1 a \ dest_vec1 d \ dest_vec1 c \ dest_vec1 b \ dest_vec1 c \ dest_vec1 d \ dest_vec1 a" + unfolding disjoint_interval and dest_vec1_def ex_1 by auto + +lemma open_closed_interval_1: fixes a :: "real^1" shows + "{a<.. dest_vec1 b ==> {a .. b} = {a<.. {a,b}" + unfolding expand_set_eq apply simp unfolding vector_less_def and vector_less_eq_def and all_1 and dest_vec1_eq[THEN sym] and dest_vec1_def by auto + +(* Some stuff for half-infinite intervals too; FIXME: notation? *) + +lemma closed_interval_left: fixes b::"real^'n::finite" + shows "closed {x::real^'n. \i. x$i \ b$i}" +proof- + { fix i + fix x::"real^'n" assume x:"\e>0. \x'\{x. \i. x $ i \ b $ i}. x' \ x \ dist x' x < e" + { assume "x$i > b$i" + then obtain y where "y $ i \ b $ i" "y \ x" "dist y x < x$i - b$i" using x[THEN spec[where x="x$i - b$i"]] by auto + hence False using component_le_norm[of "y - x" i] unfolding dist_norm and vector_minus_component by auto } + hence "x$i \ b$i" by(rule ccontr)auto } + thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast +qed + +lemma closed_interval_right: fixes a::"real^'n::finite" + shows "closed {x::real^'n. \i. a$i \ x$i}" +proof- + { fix i + fix x::"real^'n" assume x:"\e>0. \x'\{x. \i. a $ i \ x $ i}. x' \ x \ dist x' x < e" + { assume "a$i > x$i" + then obtain y where "a $ i \ y $ i" "y \ x" "dist y x < a$i - x$i" using x[THEN spec[where x="a$i - x$i"]] by auto + hence False using component_le_norm[of "y - x" i] unfolding dist_norm and vector_minus_component by auto } + hence "a$i \ x$i" by(rule ccontr)auto } + thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast +qed + +subsection{* Intervals in general, including infinite and mixtures of open and closed. *} + +definition "is_interval s \ (\a\s. \b\s. \x. (\i. ((a$i \ x$i \ x$i \ b$i) \ (b$i \ x$i \ x$i \ a$i))) \ x \ s)" + +lemma is_interval_interval: "is_interval {a .. b::real^'n::finite}" (is ?th1) "is_interval {a<..x y z::real. x < y \ y < z \ x < z" by auto + show ?th1 ?th2 unfolding is_interval_def mem_interval Ball_def atLeastAtMost_iff + by(meson real_le_trans le_less_trans less_le_trans *)+ qed + +lemma is_interval_empty: + "is_interval {}" + unfolding is_interval_def + by simp + +lemma is_interval_univ: + "is_interval UNIV" + unfolding is_interval_def + by simp + +subsection{* Closure of halfspaces and hyperplanes. *} + +lemma Lim_inner: + assumes "(f ---> l) net" shows "((\y. inner a (f y)) ---> inner a l) net" + by (intro tendsto_intros assms) + +lemma continuous_at_inner: "continuous (at x) (inner a)" + unfolding continuous_at by (intro tendsto_intros) + +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. continuous (at x) (inner a)" + unfolding continuous_at by (rule allI) (intro tendsto_intros) + hence "closed (inner a -` {..b})" + using closed_real_atMost by (rule continuous_closed_vimage) + moreover have "{x. inner a x \ b} = inner a -` {..b}" by auto + ultimately show ?thesis by simp +qed + +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. 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 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 inner_basis[OF assms] by auto + +text{* Openness of halfspaces. *} + +lemma open_halfspace_lt: "open {x. inner a x < b}" +proof- + 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. inner a x > b}" +proof- + 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 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 inner_basis[OF assms] by auto + +text{* This gives a simple derivation of limit component bounds. *} + +lemma Lim_component_le: fixes f :: "'a \ real^'n::finite" + 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. 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 + +lemma Lim_component_ge: fixes f :: "'a \ real^'n::finite" + 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. 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 + +lemma Lim_component_eq: fixes f :: "'a \ real^'n::finite" + assumes net:"(f ---> l) net" "~(trivial_limit net)" and ev:"eventually (\x. f(x)$i = b) net" + shows "l$i = b" + using ev[unfolded order_eq_iff eventually_and] using Lim_component_ge[OF net, of b i] and Lim_component_le[OF net, of i b] by auto + +lemma Lim_drop_le: fixes f :: "'a \ real^1" shows + "(f ---> l) net \ ~(trivial_limit net) \ eventually (\x. dest_vec1 (f x) \ b) net ==> dest_vec1 l \ b" + using Lim_component_le[of f l net 1 b] unfolding dest_vec1_def by auto + +lemma Lim_drop_ge: fixes f :: "'a \ real^1" shows + "(f ---> l) net \ ~(trivial_limit net) \ eventually (\x. b \ dest_vec1 (f x)) net ==> b \ dest_vec1 l" + using Lim_component_ge[of f l net b 1] unfolding dest_vec1_def by auto + +text{* Limits relative to a union. *} + +lemma eventually_within_Un: + "eventually P (net within (s \ t)) \ + eventually P (net within s) \ eventually P (net within t)" + unfolding Limits.eventually_within + by (auto elim!: eventually_rev_mp) + +lemma Lim_within_union: + "(f ---> l) (net within (s \ t)) \ + (f ---> l) (net within s) \ (f ---> l) (net within t)" + unfolding tendsto_def + by (auto simp add: eventually_within_Un) + +lemma continuous_on_union: + assumes "closed s" "closed t" "continuous_on s f" "continuous_on t f" + shows "continuous_on (s \ t) f" + using assms unfolding continuous_on unfolding Lim_within_union + unfolding Lim unfolding trivial_limit_within unfolding closed_limpt by auto + +lemma continuous_on_cases: + assumes "closed s" "closed t" "continuous_on s f" "continuous_on t g" + "\x. (x\s \ \ P x) \ (x \ t \ P x) \ f x = g x" + shows "continuous_on (s \ t) (\x. if P x then f x else g x)" +proof- + let ?h = "(\x. if P x then f x else g x)" + have "\x\s. f x = (if P x then f x else g x)" using assms(5) by auto + hence "continuous_on s ?h" using continuous_on_eq[of s f ?h] using assms(3) by auto + moreover + have "\x\t. g x = (if P x then f x else g x)" using assms(5) by auto + hence "continuous_on t ?h" using continuous_on_eq[of t g ?h] using assms(4) by auto + ultimately show ?thesis using continuous_on_union[OF assms(1,2), of ?h] by auto +qed + + +text{* Some more convenient intermediate-value theorem formulations. *} + +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. 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 + ultimately show False using assms(1)[unfolded connected_def not_ex, THEN spec[where x="?A"], THEN spec[where x="?B"]] and assms(2-5) by auto +qed + +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: inner_basis) + +text{* Also more convenient formulations of monotone convergence. *} + +lemma bounded_increasing_convergent: fixes s::"nat \ real^1" + assumes "bounded {s n| n::nat. True}" "\n. dest_vec1(s n) \ dest_vec1(s(Suc n))" + shows "\l. (s ---> l) sequentially" +proof- + obtain a where a:"\n. \dest_vec1 (s n)\ \ a" using assms(1)[unfolded bounded_iff abs_dest_vec1] by auto + { fix m::nat + 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] 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 + +subsection{* Basic homeomorphism definitions. *} + +definition "homeomorphism s t f g \ + (\x\s. (g(f x) = x)) \ (f ` s = t) \ continuous_on s f \ + (\y\t. (f(g y) = y)) \ (g ` t = s) \ continuous_on t g" + +definition + homeomorphic :: "'a::metric_space set \ 'b::metric_space set \ bool" + (infixr "homeomorphic" 60) where + homeomorphic_def: "s homeomorphic t \ (\f g. homeomorphism s t f g)" + +lemma homeomorphic_refl: "s homeomorphic s" + unfolding homeomorphic_def + unfolding homeomorphism_def + using continuous_on_id + apply(rule_tac x = "(\x. x)" in exI) + apply(rule_tac x = "(\x. x)" in exI) + by blast + +lemma homeomorphic_sym: + "s homeomorphic t \ t homeomorphic s" +unfolding homeomorphic_def +unfolding homeomorphism_def +by blast (* FIXME: slow *) + +lemma homeomorphic_trans: + assumes "s homeomorphic t" "t homeomorphic u" shows "s homeomorphic u" +proof- + obtain f1 g1 where fg1:"\x\s. g1 (f1 x) = x" "f1 ` s = t" "continuous_on s f1" "\y\t. f1 (g1 y) = y" "g1 ` t = s" "continuous_on t g1" + using assms(1) unfolding homeomorphic_def homeomorphism_def by auto + obtain f2 g2 where fg2:"\x\t. g2 (f2 x) = x" "f2 ` t = u" "continuous_on t f2" "\y\u. f2 (g2 y) = y" "g2 ` u = t" "continuous_on u g2" + using assms(2) unfolding homeomorphic_def homeomorphism_def by auto + + { fix x assume "x\s" hence "(g1 \ g2) ((f2 \ f1) x) = x" using fg1(1)[THEN bspec[where x=x]] and fg2(1)[THEN bspec[where x="f1 x"]] and fg1(2) by auto } + moreover have "(f2 \ f1) ` s = u" using fg1(2) fg2(2) by auto + moreover have "continuous_on s (f2 \ f1)" using continuous_on_compose[OF fg1(3)] and fg2(3) unfolding fg1(2) by auto + moreover { fix y assume "y\u" hence "(f2 \ f1) ((g1 \ g2) y) = y" using fg2(4)[THEN bspec[where x=y]] and fg1(4)[THEN bspec[where x="g2 y"]] and fg2(5) by auto } + moreover have "(g1 \ g2) ` u = s" using fg1(5) fg2(5) by auto + moreover have "continuous_on u (g1 \ g2)" using continuous_on_compose[OF fg2(6)] and fg1(6) unfolding fg2(5) by auto + ultimately show ?thesis unfolding homeomorphic_def homeomorphism_def apply(rule_tac x="f2 \ f1" in exI) apply(rule_tac x="g1 \ g2" in exI) by auto +qed + +lemma homeomorphic_minimal: + "s homeomorphic t \ + (\f g. (\x\s. f(x) \ t \ (g(f(x)) = x)) \ + (\y\t. g(y) \ s \ (f(g(y)) = y)) \ + continuous_on s f \ continuous_on t g)" +unfolding homeomorphic_def homeomorphism_def +apply auto apply (rule_tac x=f in exI) apply (rule_tac x=g in exI) +apply auto apply (rule_tac x=f in exI) apply (rule_tac x=g in exI) apply auto +unfolding image_iff +apply(erule_tac x="g x" in ballE) apply(erule_tac x="x" in ballE) +apply auto apply(rule_tac x="g x" in bexI) apply auto +apply(erule_tac x="f x" in ballE) apply(erule_tac x="x" in ballE) +apply auto apply(rule_tac x="f x" in bexI) by auto + +subsection{* Relatively weak hypotheses if a set is compact. *} + +definition "inv_on f s = (\x. SOME y. y\s \ f y = x)" + +lemma assumes "inj_on f s" "x\s" + shows "inv_on f s (f x) = x" + using assms unfolding inj_on_def inv_on_def by auto + +lemma homeomorphism_compact: + fixes f :: "'a::heine_borel \ 'b::heine_borel" + (* class constraint due to continuous_on_inverse *) + assumes "compact s" "continuous_on s f" "f ` s = t" "inj_on f s" + shows "\g. homeomorphism s t f g" +proof- + def g \ "\x. SOME y. y\s \ f y = x" + have g:"\x\s. g (f x) = x" using assms(3) assms(4)[unfolded inj_on_def] unfolding g_def by auto + { fix y assume "y\t" + then obtain x where x:"f x = y" "x\s" using assms(3) by auto + hence "g (f x) = x" using g by auto + hence "f (g y) = y" unfolding x(1)[THEN sym] by auto } + hence g':"\x\t. f (g x) = x" by auto + moreover + { fix x + have "x\s \ x \ g ` t" using g[THEN bspec[where x=x]] unfolding image_iff using assms(3) by(auto intro!: bexI[where x="f x"]) + moreover + { assume "x\g ` t" + then obtain y where y:"y\t" "g y = x" by auto + then obtain x' where x':"x'\s" "f x' = y" using assms(3) by auto + hence "x \ s" unfolding g_def using someI2[of "\b. b\s \ f b = y" x' "\x. x\s"] unfolding y(2)[THEN sym] and g_def by auto } + ultimately have "x\s \ x \ g ` t" by auto } + hence "g ` t = s" by auto + ultimately + show ?thesis unfolding homeomorphism_def homeomorphic_def + apply(rule_tac x=g in exI) using g and assms(3) and continuous_on_inverse[OF assms(2,1), of g, unfolded assms(3)] and assms(2) by auto +qed + +lemma homeomorphic_compact: + fixes f :: "'a::heine_borel \ 'b::heine_borel" + (* class constraint due to continuous_on_inverse *) + shows "compact s \ continuous_on s f \ (f ` s = t) \ inj_on f s + \ s homeomorphic t" + unfolding homeomorphic_def by(metis homeomorphism_compact) + +text{* Preservation of topological properties. *} + +lemma homeomorphic_compactness: + "s homeomorphic t ==> (compact s \ compact t)" +unfolding homeomorphic_def homeomorphism_def +by (metis compact_continuous_image) + +text{* Results on translation, scaling etc. *} + +lemma homeomorphic_scaling: + fixes s :: "'a::real_normed_vector set" + assumes "c \ 0" shows "s homeomorphic ((\x. c *\<^sub>R x) ` s)" + unfolding homeomorphic_minimal + 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: + fixes s :: "'a::real_normed_vector set" + shows "s homeomorphic ((\x. a + x) ` s)" + unfolding homeomorphic_minimal + apply(rule_tac x="\x. a + x" in exI) + apply(rule_tac x="\x. -a + x" in exI) + using continuous_on_add[OF continuous_on_const continuous_on_id] by auto + +lemma homeomorphic_affinity: + fixes s :: "'a::real_normed_vector set" + assumes "c \ 0" shows "s homeomorphic ((\x. a + c *\<^sub>R x) ` s)" +proof- + 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 *\<^sub>R x) ` s" a] unfolding * by auto +qed + +lemma homeomorphic_balls: + fixes a b ::"'a::real_normed_vector" (* FIXME: generalize to metric_space *) + assumes "0 < d" "0 < e" + shows "(ball a d) homeomorphic (ball b e)" (is ?th) + "(cball a d) homeomorphic (cball b e)" (is ?cth) +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) *\<^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) *\<^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. *} + +lemma cauchy_isometric: + fixes x :: "nat \ real ^ 'n::finite" + assumes e:"0 < e" and s:"subspace s" and f:"bounded_linear f" and normf:"\x\s. norm(f x) \ e * norm(x)" and xs:"\n::nat. x n \ s" and cf:"Cauchy(f o x)" + shows "Cauchy x" +proof- + interpret f: bounded_linear f by fact + { fix d::real assume "d>0" + then obtain N where N:"\n\N. norm (f (x n) - f (x N)) < e * d" + using cf[unfolded cauchy o_def dist_norm, THEN spec[where x="e*d"]] and e and mult_pos_pos[of e d] by auto + { fix n assume "n\N" + hence "norm (f (x n - x N)) < e * d" using N[THEN spec[where x=n]] unfolding f.diff[THEN sym] by auto + moreover have "e * norm (x n - x N) \ norm (f (x n - x N))" + using subspace_sub[OF s, of "x n" "x N"] using xs[THEN spec[where x=N]] and xs[THEN spec[where x=n]] + using normf[THEN bspec[where x="x n - x N"]] by auto + ultimately have "norm (x n - x N) < d" using `e>0` + using mult_left_less_imp_less[of e "norm (x n - x N)" d] by auto } + hence "\N. \n\N. norm (x n - x N) < d" by auto } + thus ?thesis unfolding cauchy and dist_norm by auto +qed + +lemma complete_isometric_image: + fixes f :: "real ^ _ \ real ^ _" + assumes "0 < e" and s:"subspace s" and f:"bounded_linear f" and normf:"\x\s. norm(f x) \ e * norm(x)" and cs:"complete s" + shows "complete(f ` s)" +proof- + { fix g assume as:"\n::nat. g n \ f ` s" and cfg:"Cauchy g" + then obtain x where "\n. x n \ s \ g n = f (x n)" unfolding image_iff and Bex_def + using choice[of "\ n xa. xa \ s \ g n = f xa"] by auto + hence x:"\n. x n \ s" "\n. g n = f (x n)" by auto + hence "f \ x = g" unfolding expand_fun_eq by auto + then obtain l where "l\s" and l:"(x ---> l) sequentially" + using cs[unfolded complete_def, THEN spec[where x="x"]] + using cauchy_isometric[OF `0l\f ` s. (g ---> l) sequentially" + using linear_continuous_at[OF f, unfolded continuous_at_sequentially, THEN spec[where x=x], of l] + unfolding `f \ x = g` by auto } + thus ?thesis unfolding complete_def by auto +qed + +lemma dist_0_norm: + fixes x :: "'a::real_normed_vector" + shows "dist 0 x = norm x" +unfolding dist_norm by simp + +lemma injective_imp_isometric: fixes f::"real^'m::finite \ real^'n::finite" + assumes s:"closed s" "subspace s" and f:"bounded_linear f" "\x\s. (f x = 0) \ (x = 0)" + shows "\e>0. \x\s. norm (f x) \ e * norm(x)" +proof(cases "s \ {0::real^'m}") + case True + { fix x assume "x \ s" + hence "x = 0" using True by auto + hence "norm x \ norm (f x)" by auto } + thus ?thesis by(auto intro!: exI[where x=1]) +next + interpret f: bounded_linear f by fact + case False + then obtain a where a:"a\0" "a\s" by auto + from False have "s \ {}" by auto + let ?S = "{f x| x. (x \ s \ norm x = norm a)}" + let ?S' = "{x::real^'m. x\s \ norm x = norm a}" + let ?S'' = "{x::real^'m. norm x = norm a}" + + have "?S'' = frontier(cball 0 (norm a))" unfolding frontier_cball and dist_norm by (auto simp add: norm_minus_cancel) + hence "compact ?S''" using compact_frontier[OF compact_cball, of 0 "norm a"] by auto + moreover have "?S' = s \ ?S''" by auto + ultimately have "compact ?S'" using closed_inter_compact[of s ?S''] using s(1) by auto + moreover have *:"f ` ?S' = ?S" by auto + ultimately have "compact ?S" using compact_continuous_image[OF linear_continuous_on[OF f(1)], of ?S'] by auto + hence "closed ?S" using compact_imp_closed by auto + moreover have "?S \ {}" using a by auto + ultimately obtain b' where "b'\?S" "\y\?S. norm b' \ norm y" using distance_attains_inf[of ?S 0] unfolding dist_0_norm by auto + then obtain b where "b\s" and ba:"norm b = norm a" and b:"\x\{x \ s. norm x = norm a}. norm (f b) \ norm (f x)" unfolding *[THEN sym] unfolding image_iff by auto + + let ?e = "norm (f b) / norm b" + have "norm b > 0" using ba and a and norm_ge_zero by auto + moreover have "norm (f b) > 0" using f(2)[THEN bspec[where x=b], OF `b\s`] using `norm b >0` unfolding zero_less_norm_iff by auto + ultimately have "0 < norm (f b) / norm b" by(simp only: divide_pos_pos) + moreover + { fix x assume "x\s" + hence "norm (f b) / norm b * norm x \ norm (f x)" + proof(cases "x=0") + case True thus "norm (f b) / norm b * norm x \ norm (f x)" by auto + 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 *\<^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 f.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 + show ?thesis by auto +qed + +lemma closed_injective_image_subspace: + fixes f :: "real ^ _ \ real ^ _" + assumes "subspace s" "bounded_linear f" "\x\s. f x = 0 --> x = 0" "closed s" + shows "closed(f ` s)" +proof- + obtain e where "e>0" and e:"\x\s. e * norm x \ norm (f x)" using injective_imp_isometric[OF assms(4,1,2,3)] by auto + show ?thesis using complete_isometric_image[OF `e>0` assms(1,2) e] and assms(4) + unfolding complete_eq_closed[THEN sym] by auto +qed + +subsection{* Some properties of a canonical subspace. *} + +lemma subspace_substandard: + "subspace {x::real^'n. (\i. P i \ x$i = 0)}" + unfolding subspace_def by(auto simp add: vector_add_component vector_smult_component elim!: ballE) + +lemma closed_substandard: + "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. 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: 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. 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 + thus ?thesis by(auto simp add: closed_Inter closed_hyperplane) +qed + +lemma dim_substandard: + shows "dim {x::real^'n::finite. \i. i \ d \ x$i = 0} = card d" (is "dim ?A = _") +proof- + let ?D = "UNIV::'n set" + let ?B = "(basis::'n\real^'n) ` d" + + let ?bas = "basis::'n \ real^'n" + + have "?B \ ?A" by auto + + moreover + { fix x::"real^'n" assume "x\?A" + with finite[of d] + have "x\ span ?B" + proof(induct d arbitrary: x) + case empty hence "x=0" unfolding Cart_eq by auto + thus ?case using subspace_0[OF subspace_span[of "{}"]] by auto + next + 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 *\<^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 + using *[THEN spec[where x=i]] by auto } + hence "y \ span (basis ` (insert k F))" using insert(3) + using span_mono[of "?bas ` F" "?bas ` (insert k F)"] + 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 *\<^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 *\<^sub>R basis k \ span (?bas ` (insert k F))" + using span_add by auto + thus ?case using y by auto + qed + } + hence "?A \ span ?B" by auto + + moreover + { fix x assume "x \ ?B" + hence "x\{(basis i)::real^'n |i. i \ ?D}" using assms by auto } + hence "independent ?B" using independent_mono[OF independent_stdbasis, of ?B] and assms by auto + + moreover + have "d \ ?D" unfolding subset_eq using assms by auto + hence *:"inj_on (basis::'n\real^'n) d" using subset_inj_on[OF basis_inj, of "d"] by auto + have "?B hassize (card d)" unfolding hassize_def and card_image[OF *] by auto + + ultimately show ?thesis using dim_unique[of "basis ` d" ?A] by auto +qed + +text{* Hence closure and completeness of all subspaces. *} + +lemma closed_subspace_lemma: "n \ card (UNIV::'n::finite set) \ \A::'n set. card A = n" +apply (induct n) +apply (rule_tac x="{}" in exI, simp) +apply clarsimp +apply (subgoal_tac "\x. x \ A") +apply (erule exE) +apply (rule_tac x="insert x A" in exI, simp) +apply (subgoal_tac "A \ UNIV", auto) +done + +lemma closed_subspace: fixes s::"(real^'n::finite) set" + assumes "subspace s" shows "closed s" +proof- + have "dim s \ card (UNIV :: 'n set)" using dim_subset_univ by auto + then obtain d::"'n set" where t: "card d = dim s" + using closed_subspace_lemma by auto + let ?t = "{x::real^'n. \i. i \ d \ x$i = 0}" + obtain f where f:"bounded_linear f" "f ` ?t = s" "inj_on f ?t" + using subspace_isomorphism[unfolded linear_conv_bounded_linear, OF subspace_substandard[of "\i. i \ d"] assms] + using dim_substandard[of d] and t by auto + interpret f: bounded_linear f by fact + have "\x\?t. f x = 0 \ x = 0" using f.zero using f(3)[unfolded inj_on_def] + by(erule_tac x=0 in ballE) auto + moreover have "closed ?t" using closed_substandard . + moreover have "subspace ?t" using subspace_substandard . + ultimately show ?thesis using closed_injective_image_subspace[of ?t f] + unfolding f(2) using f(1) by auto +qed + +lemma complete_subspace: + fixes s :: "(real ^ _) set" shows "subspace s ==> complete s" + using complete_eq_closed closed_subspace + by auto + +lemma dim_closure: + fixes s :: "(real ^ _) set" + shows "dim(closure s) = dim s" (is "?dc = ?d") +proof- + have "?dc \ ?d" using closure_minimal[OF span_inc, of s] + using closed_subspace[OF subspace_span, of s] + using dim_subset[of "closure s" "span s"] unfolding dim_span by auto + thus ?thesis using dim_subset[OF closure_subset, of s] by auto +qed + +text{* Affine transformations of intervals. *} + +lemma affinity_inverses: + assumes m0: "m \ (0::'a::field)" + shows "(\x. m *s x + c) o (\x. inverse(m) *s x + (-(inverse(m) *s c))) = id" + "(\x. inverse(m) *s x + (-(inverse(m) *s c))) o (\x. m *s x + c) = id" + using m0 +apply (auto simp add: expand_fun_eq vector_add_ldistrib vector_smult_assoc) +by (simp add: vector_smult_lneg[symmetric] vector_smult_assoc vector_sneg_minus1[symmetric]) + +lemma real_affinity_le: + "0 < (m::'a::ordered_field) ==> (m * x + c \ y \ x \ inverse(m) * y + -(c / m))" + by (simp add: field_simps inverse_eq_divide) + +lemma real_le_affinity: + "0 < (m::'a::ordered_field) ==> (y \ m * x + c \ inverse(m) * y + -(c / m) \ x)" + by (simp add: field_simps inverse_eq_divide) + +lemma real_affinity_lt: + "0 < (m::'a::ordered_field) ==> (m * x + c < y \ x < inverse(m) * y + -(c / m))" + by (simp add: field_simps inverse_eq_divide) + +lemma real_lt_affinity: + "0 < (m::'a::ordered_field) ==> (y < m * x + c \ inverse(m) * y + -(c / m) < x)" + by (simp add: field_simps inverse_eq_divide) + +lemma real_affinity_eq: + "(m::'a::ordered_field) \ 0 ==> (m * x + c = y \ x = inverse(m) * y + -(c / m))" + by (simp add: field_simps inverse_eq_divide) + +lemma real_eq_affinity: + "(m::'a::ordered_field) \ 0 ==> (y = m * x + c \ inverse(m) * y + -(c / m) = x)" + by (simp add: field_simps inverse_eq_divide) + +lemma vector_affinity_eq: + assumes m0: "(m::'a::field) \ 0" + shows "m *s x + c = y \ x = inverse m *s y + -(inverse m *s c)" +proof + assume h: "m *s x + c = y" + hence "m *s x = y - c" by (simp add: ring_simps) + hence "inverse m *s (m *s x) = inverse m *s (y - c)" by simp + then show "x = inverse m *s y + - (inverse m *s c)" + using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib) +next + assume h: "x = inverse m *s y + - (inverse m *s c)" + show "m *s x + c = y" unfolding h diff_minus[symmetric] + using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib) +qed + +lemma vector_eq_affinity: + "(m::'a::field) \ 0 ==> (y = m *s x + c \ inverse(m) *s y + -(inverse(m) *s c) = x)" + using vector_affinity_eq[where m=m and x=x and y=y and c=c] + by metis + +lemma image_affinity_interval: fixes m::real + fixes a b c :: "real^'n::finite" + shows "(\x. m *\<^sub>R x + c) ` {a .. b} = + (if {a .. b} = {} then {} + 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 *\<^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 *\<^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 *\<^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 *\<^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) *\<^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 *\<^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) *\<^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 *\<^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...) *} + +lemma banach_fix: + assumes s:"complete s" "s \ {}" and c:"0 \ c" "c < 1" and f:"(f ` s) \ s" and + lipschitz:"\x\s. \y\s. dist (f x) (f y) \ c * dist x y" + shows "\! x\s. (f x = x)" +proof- + have "1 - c > 0" using c by auto + + from s(2) obtain z0 where "z0 \ s" by auto + def z \ "\n. (f ^^ n) z0" + { fix n::nat + have "z n \ s" unfolding z_def + proof(induct n) case 0 thus ?case using `z0 \s` by auto + next case Suc thus ?case using f by auto qed } + note z_in_s = this + + def d \ "dist (z 0) (z 1)" + + have fzn:"\n. f (z n) = z (Suc n)" unfolding z_def by auto + { fix n::nat + have "dist (z n) (z (Suc n)) \ (c ^ n) * d" + proof(induct n) + case 0 thus ?case unfolding d_def by auto + next + case (Suc m) + hence "c * dist (z m) (z (Suc m)) \ c ^ Suc m * d" + using `0 \ c` using mult_mono1_class.mult_mono1[of "dist (z m) (z (Suc m))" "c ^ m * d" c] by auto + thus ?case using lipschitz[THEN bspec[where x="z m"], OF z_in_s, THEN bspec[where x="z (Suc m)"], OF z_in_s] + unfolding fzn and mult_le_cancel_left by auto + qed + } note cf_z = this + + { fix n m::nat + have "(1 - c) * dist (z m) (z (m+n)) \ (c ^ m) * d * (1 - c ^ n)" + proof(induct n) + case 0 show ?case by auto + next + case (Suc k) + have "(1 - c) * dist (z m) (z (m + Suc k)) \ (1 - c) * (dist (z m) (z (m + k)) + dist (z (m + k)) (z (Suc (m + k))))" + using dist_triangle and c by(auto simp add: dist_triangle) + also have "\ \ (1 - c) * (dist (z m) (z (m + k)) + c ^ (m + k) * d)" + using cf_z[of "m + k"] and c by auto + also have "\ \ c ^ m * d * (1 - c ^ k) + (1 - c) * c ^ (m + k) * d" + using Suc by (auto simp add: ring_simps) + also have "\ = (c ^ m) * (d * (1 - c ^ k) + (1 - c) * c ^ k * d)" + unfolding power_add by (auto simp add: ring_simps) + also have "\ \ (c ^ m) * d * (1 - c ^ Suc k)" + using c by (auto simp add: ring_simps) + finally show ?case by auto + qed + } note cf_z2 = this + { fix e::real assume "e>0" + hence "\N. \m n. N \ m \ N \ n \ dist (z m) (z n) < e" + proof(cases "d = 0") + case True + hence "\n. z n = z0" using cf_z2[of 0] and c unfolding z_def by (auto simp add: pos_prod_le[OF `1 - c > 0`]) + thus ?thesis using `e>0` by auto + next + case False hence "d>0" unfolding d_def using zero_le_dist[of "z 0" "z 1"] + by (metis False d_def real_less_def) + hence "0 < e * (1 - c) / d" using `e>0` and `1-c>0` + using divide_pos_pos[of "e * (1 - c)" d] and mult_pos_pos[of e "1 - c"] by auto + then obtain N where N:"c ^ N < e * (1 - c) / d" using real_arch_pow_inv[of "e * (1 - c) / d" c] and c by auto + { fix m n::nat assume "m>n" and as:"m\N" "n\N" + have *:"c ^ n \ c ^ N" using `n\N` and c using power_decreasing[OF `n\N`, of c] by auto + have "1 - c ^ (m - n) > 0" using c and power_strict_mono[of c 1 "m - n"] using `m>n` by auto + hence **:"d * (1 - c ^ (m - n)) / (1 - c) > 0" + using real_mult_order[OF `d>0`, of "1 - c ^ (m - n)"] + using divide_pos_pos[of "d * (1 - c ^ (m - n))" "1 - c"] + using `0 < 1 - c` by auto + + have "dist (z m) (z n) \ c ^ n * d * (1 - c ^ (m - n)) / (1 - c)" + using cf_z2[of n "m - n"] and `m>n` unfolding pos_le_divide_eq[OF `1-c>0`] + by (auto simp add: real_mult_commute dist_commute) + also have "\ \ c ^ N * d * (1 - c ^ (m - n)) / (1 - c)" + using mult_right_mono[OF * order_less_imp_le[OF **]] + unfolding real_mult_assoc by auto + also have "\ < (e * (1 - c) / d) * d * (1 - c ^ (m - n)) / (1 - c)" + using mult_strict_right_mono[OF N **] unfolding real_mult_assoc by auto + also have "\ = e * (1 - c ^ (m - n))" using c and `d>0` and `1 - c > 0` by auto + also have "\ \ e" using c and `1 - c ^ (m - n) > 0` and `e>0` using mult_right_le_one_le[of e "1 - c ^ (m - n)"] by auto + finally have "dist (z m) (z n) < e" by auto + } note * = this + { fix m n::nat assume as:"N\m" "N\n" + hence "dist (z n) (z m) < e" + proof(cases "n = m") + case True thus ?thesis using `e>0` by auto + next + case False thus ?thesis using as and *[of n m] *[of m n] unfolding nat_neq_iff by (auto simp add: dist_commute) + qed } + thus ?thesis by auto + qed + } + hence "Cauchy z" unfolding cauchy_def by auto + then obtain x where "x\s" and x:"(z ---> x) sequentially" using s(1)[unfolded compact_def complete_def, THEN spec[where x=z]] and z_in_s by auto + + def e \ "dist (f x) x" + have "e = 0" proof(rule ccontr) + assume "e \ 0" hence "e>0" unfolding e_def using zero_le_dist[of "f x" x] + by (metis dist_eq_0_iff dist_nz e_def) + then obtain N where N:"\n\N. dist (z n) x < e / 2" + using x[unfolded Lim_sequentially, THEN spec[where x="e/2"]] by auto + hence N':"dist (z N) x < e / 2" by auto + + have *:"c * dist (z N) x \ dist (z N) x" unfolding mult_le_cancel_right2 + using zero_le_dist[of "z N" x] and c + by (metis dist_eq_0_iff dist_nz order_less_asym real_less_def) + have "dist (f (z N)) (f x) \ c * dist (z N) x" using lipschitz[THEN bspec[where x="z N"], THEN bspec[where x=x]] + using z_in_s[of N] `x\s` using c by auto + also have "\ < e / 2" using N' and c using * by auto + finally show False unfolding fzn + using N[THEN spec[where x="Suc N"]] and dist_triangle_half_r[of "z (Suc N)" "f x" e x] + unfolding e_def by auto + qed + hence "f x = x" unfolding e_def by auto + moreover + { fix y assume "f y = y" "y\s" + hence "dist x y \ c * dist x y" using lipschitz[THEN bspec[where x=x], THEN bspec[where x=y]] + using `x\s` and `f x = x` by auto + hence "dist x y = 0" unfolding mult_le_cancel_right1 + using c and zero_le_dist[of x y] by auto + hence "y = x" by auto + } + ultimately show ?thesis unfolding Bex1_def using `x\s` by blast+ +qed + +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\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 + { fix y assume "y\s" + hence "x = y" using `x\s` and dist[THEN bspec[where x=x], THEN bspec[where x=y]] + unfolding g[THEN bspec[where x=x], OF `x\s`] + unfolding g[THEN bspec[where x=y], OF `y\s`] by auto } + thus ?thesis unfolding Bex1_def using `x\s` and g by blast+ +next + case True + then obtain x where [simp]:"x\s" and "g x \ x" by auto + { fix x y assume "x \ s" "y \ s" + hence "dist (g x) (g y) \ dist x y" + using dist[THEN bspec[where x=x], THEN bspec[where x=y]] by auto } note dist' = this + def y \ "g x" + have [simp]:"y\s" unfolding y_def using gs[unfolded image_subset_iff] and `x\s` by blast + def f \ "\n. g ^^ n" + have [simp]:"\n z. g (f n z) = f (Suc n) z" unfolding f_def by auto + have [simp]:"\z. f 0 z = z" unfolding f_def by auto + { fix n::nat and z assume "z\s" + have "f n z \ s" unfolding f_def + proof(induct n) + case 0 thus ?case using `z\s` by simp + next + case (Suc n) thus ?case using gs[unfolded image_subset_iff] by auto + qed } note fs = this + { fix m n ::nat assume "m\n" + fix w z assume "w\s" "z\s" + have "dist (f n w) (f n z) \ dist (f m w) (f m z)" using `m\n` + proof(induct n) + case 0 thus ?case by auto + next + case (Suc n) + thus ?case proof(cases "m\n") + case True thus ?thesis using Suc(1) + using dist'[OF fs fs, OF `w\s` `z\s`, of n n] by auto + next + case False hence mn:"m = Suc n" using Suc(2) by simp + show ?thesis unfolding mn by auto + qed + qed } note distf = this + + 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 \ "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 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::'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 + + { assume as:"dist a b > dist (f n x) (f n y)" + then obtain Na Nb where "\m\Na. dist (f (r m) x) a < (dist a b - dist (f n x) (f n y)) / 2" + and "\m\Nb. dist (f (r m) y) b < (dist a b - dist (f n x) (f n y)) / 2" + using lima limb unfolding h_def Lim_sequentially by (fastsimp simp del: less_divide_eq_number_of1) + hence "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)" + apply(erule_tac x="Na+Nb+n" in allE) + apply(erule_tac x="Na+Nb+n" in allE) apply simp + using dist_triangle_add_half[of a "f (r (Na + Nb + n)) x" "dist a b - dist (f n x) (f n y)" + "-b" "- f (r (Na + Nb + n)) y"] + unfolding ** unfolding group_simps(12) by (auto simp add: dist_commute) + 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 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 + } + hence "dist a b \ dist (f n x) (f n y)" by(rule ccontr)auto } + note ab_fn = this + + have [simp]:"a = b" proof(rule ccontr) + def e \ "dist a b - dist (g a) (g b)" + assume "a\b" hence "e > 0" unfolding e_def using dist by fastsimp + hence "\n. dist (f n x) a < e/2 \ dist (f n y) b < e/2" + using lima limb unfolding Lim_sequentially + apply (auto elim!: allE[where x="e/2"]) apply(rule_tac x="r (max N Na)" in exI) unfolding h_def by fastsimp + then obtain n where n:"dist (f n x) a < e/2 \ dist (f n y) b < e/2" by auto + have "dist (f (Suc n) x) (g a) \ dist (f n x) a" + using dist[THEN bspec[where x="f n x"], THEN bspec[where x="a"]] and fs by auto + moreover have "dist (f (Suc n) y) (g b) \ dist (f n y) b" + using dist[THEN bspec[where x="f n y"], THEN bspec[where x="b"]] and fs by auto + ultimately have "dist (f (Suc n) x) (g a) + dist (f (Suc n) y) (g b) < e" using n by auto + thus False unfolding e_def using ab_fn[of "Suc n"] by norm + qed + + have [simp]:"\n. f (Suc n) x = f n y" unfolding f_def y_def by(induct_tac n)auto + { fix x y assume "x\s" "y\s" moreover + fix e::real assume "e>0" ultimately + 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 "((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 + moreover + { fix x assume "x\s" "g x = x" "x\a" + hence "False" using dist[THEN bspec[where x=a], THEN bspec[where x=x]] + using `g a = a` and `a\s` by auto } + ultimately show "\!x\s. g x = x" unfolding Bex1_def using `a\s` by blast +qed + +end diff -r 1fad3160d873 -r 2083bde13ce1 src/HOL/Library/Convex_Euclidean_Space.thy --- a/src/HOL/Library/Convex_Euclidean_Space.thy Fri Oct 23 14:33:07 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,3371 +0,0 @@ -(* Title: HOL/Library/Convex_Euclidean_Space.thy - Author: Robert Himmelmann, TU Muenchen -*) - -header {* Convex sets, functions and related things. *} - -theory Convex_Euclidean_Space -imports Topology_Euclidean_Space -begin - - -(* ------------------------------------------------------------------------- *) -(* To be moved elsewhere *) -(* ------------------------------------------------------------------------- *) - -declare vector_add_ldistrib[simp] vector_ssub_ldistrib[simp] vector_smult_assoc[simp] vector_smult_rneg[simp] -declare vector_sadd_rdistrib[simp] vector_sub_rdistrib[simp] -declare dot_ladd[simp] dot_radd[simp] dot_lsub[simp] dot_rsub[simp] -declare dot_lmult[simp] dot_rmult[simp] dot_lneg[simp] dot_rneg[simp] -declare UNIV_1[simp] - -term "(x::real^'n \ real) 0" - -lemma dim1in[intro]:"Suc 0 \ {1::nat .. CARD(1)}" by auto - -lemmas vector_component_simps = vector_minus_component vector_smult_component vector_add_component vector_less_eq_def Cart_lambda_beta dest_vec1_def basis_component vector_uminus_component - -lemmas continuous_intros = continuous_add continuous_vmul continuous_cmul continuous_const continuous_sub continuous_at_id continuous_within_id - -lemmas continuous_on_intros = continuous_on_add continuous_on_const continuous_on_id continuous_on_compose continuous_on_cmul continuous_on_neg continuous_on_sub - uniformly_continuous_on_add uniformly_continuous_on_const uniformly_continuous_on_id uniformly_continuous_on_compose uniformly_continuous_on_cmul uniformly_continuous_on_neg uniformly_continuous_on_sub - -lemma dest_vec1_simps[simp]: fixes a::"real^1" - shows "a$1 = 0 \ a = 0" (*"a \ 1 \ dest_vec1 a \ 1" "0 \ a \ 0 \ dest_vec1 a"*) - "a \ b \ dest_vec1 a \ dest_vec1 b" "dest_vec1 (1::real^1) = 1" - by(auto simp add:vector_component_simps all_1 Cart_eq) - -lemma nequals0I:"x\A \ A \ {}" by auto - -lemma norm_not_0:"(x::real^'n::finite)\0 \ norm x \ 0" by auto - -lemma setsum_delta_notmem: assumes "x\s" - shows "setsum (\y. if (y = x) then P x else Q y) s = setsum Q s" - "setsum (\y. if (x = y) then P x else Q y) s = setsum Q s" - "setsum (\y. if (y = x) then P y else Q y) s = setsum Q s" - "setsum (\y. if (x = y) then P y else Q y) s = setsum Q s" - apply(rule_tac [!] setsum_cong2) using assms by auto - -lemma setsum_delta'': - fixes s::"'a::real_vector set" assumes "finite s" - 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)) *\<^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)) *\<^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 *\<^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: - "dest_vec1 ` {a .. b} = {dest_vec1 a .. dest_vec1 b}" - "dest_vec1 ` {a<.. b} = {dest_vec1 a<.. dest_vec1 b}" - "dest_vec1 ` {a ..x. dest_vec1 (f x)) S" - using dest_vec1_sum[OF assms] by auto - -lemma dist_triangle_eq: - 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 smult_conv_scaleR *] - by(auto simp add:norm_minus_commute) qed - -lemma norm_eqI:"x = y \ norm x = norm y" by auto -lemma norm_minus_eqI:"(x::real^'n::finite) = - y \ norm x = norm y" by auto - -lemma Min_grI: assumes "finite A" "A \ {}" "\a\A. x < a" shows "x < Min A" - unfolding Min_gr_iff[OF assms(1,2)] using assms(3) by auto - -lemma dimindex_ge_1:"CARD(_::finite) \ 1" - using one_le_card_finite by auto - -lemma real_dimindex_ge_1:"real (CARD('n::finite)) \ 1" - by(metis dimindex_ge_1 linorder_not_less real_eq_of_nat real_le_trans real_of_nat_1 real_of_nat_le_iff) - -lemma real_dimindex_gt_0:"real (CARD('n::finite)) > 0" apply(rule less_le_trans[OF _ real_dimindex_ge_1]) by auto - -subsection {* Affine set and affine hull.*} - -definition - affine :: "'a::real_vector set \ bool" where - "affine s \ (\x\s. \y\s. \u v. 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 *\<^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 - -lemma affine_empty[intro]: "affine {}" - unfolding affine_def by auto - -lemma affine_sing[intro]: "affine {x}" - unfolding affine_alt by (auto simp add: scaleR_left_distrib [symmetric]) - -lemma affine_UNIV[intro]: "affine UNIV" - unfolding affine_def by auto - -lemma affine_Inter: "(\s\f. affine s) \ affine (\ f)" - unfolding affine_def by auto - -lemma affine_Int: "affine s \ affine t \ affine (s \ t)" - unfolding affine_def by auto - -lemma affine_affine_hull: "affine(affine hull s)" - unfolding hull_def using affine_Inter[of "{t \ affine. s \ t}"] - unfolding mem_def by auto - -lemma affine_hull_eq[simp]: "(affine hull s = s) \ affine s" -proof- - { fix f assume "f \ affine" - hence "affine (\f)" using affine_Inter[of f] unfolding subset_eq mem_def by auto } - thus ?thesis using hull_eq[unfolded mem_def, of affine s] by auto -qed - -lemma setsum_restrict_set'': assumes "finite A" - shows "setsum f {x \ A. P x} = (\x\A. if P x then f x else 0)" - unfolding mem_def[of _ P, symmetric] unfolding setsum_restrict_set'[OF assms] .. - -subsection {* Some explicit formulations (from Lars Schewe). *} - -lemma affine: fixes V::"'a::real_vector 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 *\<^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: scaleR_left_distrib[THEN sym]) -next - fix s u assume as:"\x\V. \y\V. \u v. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ V" - "finite s" "s \ {}" "s \ V" "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 *\<^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::"'a set" and u::"'a \ real" - assume IA:"\u s. \2 < card s; \x\V. \y\V. \u v. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ V; finite s; - s \ {}; s \ V; 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 - thus False using as(7) and `card s > 2` by (metis Numeral1_eq1_nat less_0_number_of less_int_code(15) - less_nat_number_of not_less_iff_gr_or_eq of_nat_1 of_nat_eq_iff pos2 rel_simps(4)) qed - then obtain x where x:"x\s" "u x \ 1" by auto - - have c:"card (s - {x}) = card s - 1" apply(rule card_Diff_singleton) using `x\s` as(4) by auto - have *:"s = insert x (s - {x})" "finite (s - {x})" using `x\s` and as(4) by auto - have **:"setsum u (s - {x}) = 1 - u x" - using setsum_clauses(2)[OF *(2), of u x, unfolded *(1)[THEN sym] as(7)] by auto - have ***:"inverse (1 - u x) * setsum u (s - {x}) = 1" unfolding ** using `u x \ 1` by auto - have "(\xa\s - {x}. (inverse (1 - u x) * u xa) *\<^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 - thus ?thesis apply(rule_tac IA[of "s - {x}" "\y. (inverse (1 - u x) * u y)"]) - unfolding setsum_right_distrib[THEN sym] using as and *** and True by auto - next case False hence "card (s - {x}) = Suc (Suc 0)" using as(2) and c by auto - then obtain a b where "(s - {x}) = {a, b}" "a\b" unfolding card_Suc_eq by auto - thus ?thesis using as(3)[THEN bspec[where x=a], THEN bspec[where x=b]] - using *** *(2) and `s \ V` unfolding setsum_right_distrib by(auto simp add: setsum_clauses(2)) qed - thus "(\x\s. u x *\<^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)) *\<^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) - thus ?thesis using as(4,5) by simp - qed(insert `s\{}` `finite s`, auto) -qed - -lemma affine_hull_explicit: - "affine hull p = {y. \s u. finite s \ s \ {} \ s \ p \ setsum u s = 1 \ setsum (\v. (u v) *\<^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 *\<^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 *\<^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 *\<^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 *\<^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 *\<^sub>R v) = u *\<^sub>R x + v *\<^sub>R y" - apply(rule_tac x="sx \ sy" in exI) - apply(rule_tac x="\a. (if a\sx then u * ux a else 0) + (if a\sy then v * uy a else 0)" in exI) - unfolding scaleR_left_distrib 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 *\<^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 *\<^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 *\<^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. *} - -lemma affine_hull_empty[simp]: "affine hull {} = {}" - apply(rule hull_unique) unfolding mem_def by auto - -lemma affine_hull_finite_step: - fixes y :: "'a::real_vector" - shows "(\u. 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 *\<^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 - next - case False thus ?thesis apply(rule_tac x="u a" in exI) using u and `?as` by auto - qed } moreover - { assume ?rhs - then obtain v u where vu:"setsum u s = w - v" "(\x\s. u x *\<^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 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 *\<^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 *\<^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: - fixes a b :: "'a::real_vector" - shows "affine hull {a,b} = {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::'a)" by auto - 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 *\<^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: - fixes a b c :: "'a::real_vector" - shows "affine hull {a,b,c} = { u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c| u v w. u + v + w = 1}" (is "?lhs = ?rhs") -proof- - have *:"\x y z. z = x - y \ y + z = (x::real)" - "\x y z. z = x - y \ y + z = (x::'a)" by auto - show ?thesis apply(simp add: affine_hull_finite affine_hull_finite_step) - unfolding * apply auto - apply(rule_tac x=v in exI) apply(rule_tac x=va in exI) apply auto - apply(rule_tac x=u in exI) by(auto intro!: exI) -qed - -subsection {* Some relations between affine hull and subspaces. *} - -lemma affine_hull_insert_subset_span: - fixes a :: "real ^ _" - shows "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 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 *\<^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 *\<^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: - fixes a :: "real ^ _" - assumes "a \ s" - shows "affine hull (insert a s) = - {a + v | v . v \ span {x - a | x. x \ s}}" - apply(rule, rule affine_hull_insert_subset_span) unfolding subset_eq Ball_def - unfolding affine_hull_explicit and mem_Collect_eq proof(rule,rule,erule exE,erule conjE) - fix y v assume "y = a + v" "v \ span {x - a |x. x \ s}" - then obtain t u where obt:"finite t" "t \ {x - a |x. x \ s}" "a + (\v\t. u v *\<^sub>R v) = y" unfolding span_explicit smult_conv_scaleR by auto - def f \ "(\x. x + a) ` t" - have f:"finite f" "f \ s" "(\v\f. u (v - a) *\<^sub>R (v - a)) = y - a" unfolding f_def using obt - by(auto simp add: 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 *\<^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 scaleR_left.setsum algebra_simps) qed - -lemma affine_hull_span: - fixes a :: "real ^ _" - assumes "a \ s" - shows "affine hull s = {a + v | v. v \ span {x - a | x. x \ s - {a}}}" - using affine_hull_insert_span[of a "s - {a}", unfolded insert_Diff[OF assms]] by auto - -subsection {* Convexity. *} - -definition - convex :: "'a::real_vector set \ bool" where - "convex s \ (\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) - by (auto simp add: *) qed - -lemma mem_convex: - assumes "convex s" "a \ s" "b \ s" "0 \ u" "u \ 1" - 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:scaleR_left_distrib[THEN sym]) - -lemma convex_UNIV[intro]: "convex UNIV" - unfolding convex_def by auto - -lemma convex_Inter: "(\s\f. convex s) ==> convex(\ f)" - unfolding convex_def by auto - -lemma convex_Int: "convex s \ convex t \ convex (s \ t)" - unfolding convex_def by auto - -lemma convex_halfspace_le: "convex {x. inner a x \ b}" - unfolding convex_def apply 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. inner a x = b}" -proof- - 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. 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) - -subsection {* Explicit expressions for convexity in terms of arbitrary sums. *} - -lemma convex: "convex s \ - (\(k::nat) u x. (\i. 1\i \ i\k \ 0 \ u i \ x i \s) \ (setsum u {1..k} = 1) - \ setsum (\i. u i *\<^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 *\<^sub>R x i) \ s" - "x \ s" "y \ s" "0 \ u" "0 \ v" "u + v = (1::real)" - 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 *\<^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 *\<^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) - hence "setsum u {1 .. k} > 0" using ui apply(rule_tac less_le_trans[of _ "u i"]) using Suc(2)[THEN spec[where x=i]] and i(1) by auto - thus False using Suc(3) unfolding setsum_cl_ivl_Suc and True by simp qed - thus ?thesis unfolding setsum_cl_ivl_Suc using True and Suc(2) by auto - next - have *:"setsum u {1..k} = 1 - u (Suc k)" using Suc(3)[unfolded setsum_cl_ivl_Suc] by auto - have **:"u (Suc k) \ 1" apply(rule ccontr) unfolding not_le using Suc(3) using setsum_nonneg[of "{1..k}" u] using Suc(2) by auto - have ***:"\i k. (u i / (1 - u (Suc k))) *\<^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))) *\<^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)) *\<^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 scaleR_right.setsum [symmetric] using nn by auto qed qed auto qed - - -lemma convex_explicit: - fixes s :: "'a::real_vector set" - shows "convex s \ - (\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 *\<^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 *\<^sub>R x + v *\<^sub>R y \ s" "finite (t::'a 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 *\<^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 *\<^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 *\<^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) - hence "setsum u f > 0" using uy apply(rule_tac less_le_trans[of _ "u y"]) using as(4) and y(1) by auto - thus False using as(2,5) unfolding setsum_clauses(2)[OF as(1)] and True by auto qed - thus ?thesis unfolding setsum_clauses(2)[OF as(1)] using as(2,3) unfolding True by auto - next - have *:"setsum u f = setsum u (insert x f) - u x" using as(2) unfolding setsum_clauses(2)[OF as(1)] by auto - have **:"u x \ 1" apply(rule ccontr) unfolding not_le using as(5)[unfolded setsum_clauses(2)[OF as(1)]] and as(2) - using setsum_nonneg[of f u] and as(4) by auto - case False hence "inverse (1 - u x) *\<^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 *\<^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 *\<^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 *\<^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 *\<^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 *\<^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 :: "'a::real_vector set \ bool" where - "cone s \ (\x\s. \c\0. (c *\<^sub>R x) \ s)" - -lemma cone_empty[intro, simp]: "cone {}" - unfolding cone_def by auto - -lemma cone_univ[intro, simp]: "cone UNIV" - unfolding cone_def by auto - -lemma cone_Inter[intro]: "(\s\f. cone s) \ cone(\ f)" - unfolding cone_def by auto - -subsection {* Conic hull. *} - -lemma cone_cone_hull: "cone (cone hull s)" - unfolding hull_def using cone_Inter[of "{t \ conic. s \ t}"] - by (auto simp add: mem_def) - -lemma cone_hull_eq: "(cone hull s = s) \ cone s" - apply(rule hull_eq[unfolded mem_def]) - using cone_Inter unfolding subset_eq by (auto simp add: mem_def) - -subsection {* Affine dependence and consequential theorems (from Lars Schewe). *} - -definition - affine_dependent :: "'a::real_vector 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 *\<^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 *\<^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 *\<^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 *\<^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 *\<^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 scaleR_scaleR[THEN sym] and scaleR_right.setsum [symmetric] unfolding setsum_right_distrib[THEN sym] and setsum_diff1[OF as(1)] using as by auto -qed - -lemma affine_dependent_explicit_finite: - fixes s :: "'a::real_vector set" assumes "finite s" - 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) *\<^sub>R v = (if vt then (u v) *\<^sub>R v else (0::'a))" by auto - assume ?lhs - then obtain t u v where "finite t" "t \ s" "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_absorb1[OF `t\s`] by auto -next - assume ?rhs - 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 - -subsection {* A general lemma. *} - -lemma convex_connected: - fixes s :: "'a::real_normed_vector set" - assumes "convex s" shows "connected s" -proof- - { fix e1 e2 assume as:"open e1" "open e2" "e1 \ e2 \ s = {}" "s \ e1 \ e2" - assume "e1 \ s \ {}" "e2 \ s \ {}" - then obtain x1 x2 where x1:"x1\e1" "x1\s" and x2:"x2\e2" "x2\s" by auto - hence n:"norm (x1 - x2) > 0" unfolding zero_less_norm_iff using as(3) by auto - - { fix x e::real assume as:"0 \ x" "x \ 1" "0 < e" - { fix y have *:"(1 - x) *\<^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) *\<^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) *\<^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) *\<^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) *\<^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 } - thus ?thesis unfolding connected_def by auto -qed - -subsection {* One rather trivial consequence. *} - -lemma connected_UNIV: "connected (UNIV :: 'a::real_normed_vector set)" - by(simp add: convex_connected convex_UNIV) - -subsection {* Convex functions into the reals. *} - -definition - convex_on :: "'a::real_vector set \ ('a \ real) \ bool" where - "convex_on s f \ - (\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 - -lemma convex_add: - assumes "convex_on s f" "convex_on s g" - shows "convex_on s (\x. f x + g x)" -proof- - { fix x y assume "x\s" "y\s" moreover - fix u v ::real assume "0 \ u" "0 \ v" "u + v = 1" - ultimately have "f (u *\<^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 *\<^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 - -lemma convex_cmul: - assumes "0 \ (c::real)" "convex_on s f" - shows "convex_on s (\x. c * f x)" -proof- - have *:"\u c fx v fy ::real. u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" by (simp add: ring_simps) - show ?thesis using assms(2) and mult_mono1[OF _ assms(1)] unfolding convex_on_def and * by auto -qed - -lemma convex_lower: - assumes "convex_on s f" "x\s" "y \ s" "0 \ u" "0 \ v" "u + v = 1" - shows "f (u *\<^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) - using assms(4,5) by(auto simp add: mult_mono1) - also have "\ = max (f x) (f y)" using assms(6) unfolding distrib[THEN sym] by auto - finally show ?thesis using assms(1)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x=u]] - using assms(2-6) by auto -qed - -lemma convex_local_global_minimum: - fixes s :: "'a::real_normed_vector set" - assumes "0 s" "\y\ball x e. f x \ f y" - shows "\y\s. f x \ f y" -proof(rule ccontr) - have "x\s" using assms(1,3) by auto - assume "\ (\y\s. f x \ f y)" - then obtain y where "y\s" and y:"f x > f y" by auto - hence xy:"0 < dist x y" by (auto simp add: dist_nz[THEN sym]) - - then obtain u where "0 < u" "u \ 1" and u:"u < e / dist x y" - using real_lbound_gt_zero[of 1 "e / dist x y"] using xy `e>0` and divide_pos_pos[of e "dist x y"] by auto - hence "f ((1-u) *\<^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) *\<^sub>R 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) *\<^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 - -lemma convex_distance: - fixes s :: "'a::real_normed_vector set" - shows "convex_on s (\x. dist a x)" -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 *\<^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 *\<^sub>R x) ` s)" - unfolding convex_def and image_iff apply auto - 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 *\<^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 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 *\<^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 -qed - -lemma convex_differences: - assumes "convex s" "convex t" - shows "convex {x - y| x y. x \ s \ y \ t}" -proof- - have "{x - y| x y. x \ s \ y \ t} = {x + y |x y. x \ s \ y \ uminus ` t}" unfolding image_iff apply auto - apply(rule_tac x=xa in exI) apply(rule_tac x="-y" in exI) apply simp - apply(rule_tac x=xa in exI) apply(rule_tac x=xb in exI) by simp - thus ?thesis using convex_sums[OF assms(1) convex_negations[OF assms(2)]] by auto -qed - -lemma convex_translation: assumes "convex s" shows "convex ((\x. a + x) ` s)" -proof- have "{a + y |y. y \ s} = (\x. a + x) ` s" by auto - thus ?thesis using convex_sums[OF convex_singleton[of a] assms] by auto qed - -lemma convex_affinity: assumes "convex s" 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:"bounded_linear f" - shows "convex(f ` s)" -proof(auto simp add: convex_def) - interpret f: bounded_linear f by fact - fix x y assume xy:"x \ s" "y \ s" - fix u v ::real assume uv:"0 \ u" "0 \ v" "u + v = 1" - 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 f.add f.scaleR - using c[unfolded convex_def] xy uv by auto -qed - -subsection {* Balls, being convex, are connected. *} - -lemma convex_ball: - fixes x :: "'a::real_normed_vector" - shows "convex (ball x e)" -proof(auto simp add: convex_def) - fix y z assume yz:"dist x y < e" "dist x z < e" - fix u v ::real assume uv:"0 \ u" "0 \ v" "u + v = 1" - have "dist x (u *\<^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 *\<^sub>R y + v *\<^sub>R z) < e" using real_convex_bound_lt[OF yz uv] by auto -qed - -lemma convex_cball: - fixes x :: "'a::real_normed_vector" - shows "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 *\<^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 *\<^sub>R y + v *\<^sub>R z) \ e" using real_convex_bound_le[OF yz uv] by auto -qed - -lemma connected_ball: - fixes x :: "'a::real_normed_vector" - shows "connected (ball x e)" - using convex_connected convex_ball by auto - -lemma connected_cball: - fixes x :: "'a::real_normed_vector" - shows "connected(cball x e)" - using convex_connected convex_cball by auto - -subsection {* Convex hull. *} - -lemma convex_convex_hull: "convex(convex hull s)" - unfolding hull_def using convex_Inter[of "{t\convex. s\t}"] - unfolding mem_def by auto - -lemma convex_hull_eq: "(convex hull s = s) \ convex s" apply(rule hull_eq[unfolded mem_def]) - using convex_Inter[unfolded Ball_def mem_def] by auto - -lemma bounded_convex_hull: - fixes s :: "'a::real_normed_vector set" - assumes "bounded s" shows "bounded(convex hull s)" -proof- from assms obtain B where B:"\x\s. norm x \ B" unfolding bounded_iff by auto - show ?thesis apply(rule bounded_subset[OF bounded_cball, of _ 0 B]) - unfolding subset_hull[unfolded mem_def, of convex, OF convex_cball] - unfolding subset_eq mem_cball dist_norm using B by auto qed - -lemma finite_imp_bounded_convex_hull: - fixes s :: "'a::real_normed_vector set" - shows "finite s \ bounded(convex hull s)" - using bounded_convex_hull finite_imp_bounded by auto - -subsection {* Stepping theorems for convex hulls of finite sets. *} - -lemma convex_hull_empty[simp]: "convex hull {} = {}" - apply(rule hull_unique) unfolding mem_def by auto - -lemma convex_hull_singleton[simp]: "convex hull {a} = {a}" - apply(rule hull_unique) unfolding mem_def by auto - -lemma convex_hull_insert: - fixes s :: "'a::real_vector set" - assumes "s \ {}" - shows "convex hull (insert a s) = {x. \u\0. \v\0. \b. (u + v = 1) \ - b \ (convex hull s) \ (x = u *\<^sub>R a + v *\<^sub>R b)}" (is "?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 *\<^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] - apply(erule_tac x=a in ballE) apply(erule_tac x=b in ballE) apply(erule_tac x=u in allE) using obt by auto -next - show "convex ?hull" unfolding convex_def apply(rule,rule,rule,rule,rule,rule,rule) proof- - fix x y u v assume as:"(0::real) \ u" "0 \ v" "u + v = 1" "x\?hull" "y\?hull" - from as(4) obtain u1 v1 b1 where obt1:"u1\0" "v1\0" "u1 + v1 = 1" "b1 \ convex hull s" "x = u1 *\<^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::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" by (auto simp add: algebra_simps) - have "\b \ convex hull s. u *\<^sub>R x + v *\<^sub>R y = (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::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" by (auto simp add: algebra_simps) - 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: ** 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) - also have "\ = u * v1 + v * v2" by simp finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" by auto - case False have "0 \ u * v1 + v * v2" "0 \ u * v1" "0 \ u * v1 + v * v2" "0 \ v * v2" apply - - apply(rule add_nonneg_nonneg) prefer 4 apply(rule add_nonneg_nonneg) apply(rule_tac [!] mult_nonneg_nonneg) - using as(1,2) obt1(1,2) obt2(1,2) by auto - thus ?thesis unfolding obt1(5) obt2(5) unfolding * and ** using False - apply(rule_tac x="((u * v1) / (u * v1 + v * v2)) *\<^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 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 - have "u1 * u + u2 * v \ (max u1 u2) * u + (max u1 u2) * v" apply(rule add_mono) - apply(rule_tac [!] mult_right_mono) using as(1,2) obt1(1,2) obt2(1,2) by auto - also have "\ \ 1" unfolding mult.add_right[THEN sym] and as(3) using u1 u2 by auto - finally - show "u *\<^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: algebra_simps) - qed -qed - - -subsection {* Explicit expression for convex hull. *} - -lemma convex_hull_indexed: - fixes s :: "'a::real_vector set" - shows "convex hull s = {y. \k u x. (\i\{1::nat .. k}. 0 \ u i \ x i \ s) \ - (setsum u {1..k} = 1) \ - (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- - fix x assume "x\s" - thus "x \ ?hull" unfolding mem_Collect_eq apply(rule_tac x=1 in exI, rule_tac x="\x. 1" in exI) by auto -next - fix t assume as:"s \ t" "convex t" - show "?hull \ t" apply(rule) unfolding mem_Collect_eq apply(erule exE | erule conjE)+ proof- - fix x k u y assume assm:"\i\{1::nat..k}. 0 \ u i \ y i \ s" "setsum u {1..k} = 1" "(\i = 1..k. u i *\<^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 *\<^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::'a) x2 s1 s2 i.(if P i then s1 else s2) *\<^sub>R (if P i then x1 else x2) = (if P i then s1 *\<^sub>R x1 else s2 *\<^sub>R x2)" - "{1..k1 + k2} \ {1..k1} = {1..k1}" "{1..k1 + k2} \ - {1..k1} = (\i. i + k1) ` {1..k2}" - prefer 3 apply(rule,rule) unfolding image_iff apply(rule_tac x="x - k1" in bexI) by(auto simp add: not_le) - have inj:"inj_on (\i. i + k1) {1..k2}" unfolding inj_on_def by auto - 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 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}") - case True thus ?thesis using mult_nonneg_nonneg[of u "u1 i"] and uv(1) x(1)[THEN bspec[where x=i]] by auto - next def j \ "i - k1" - case False with i have "j \ {1..k2}" unfolding j_def by auto - thus ?thesis unfolding j_def[symmetric] using False - using mult_nonneg_nonneg[of v "u2 j"] and uv(2) y(1)[THEN bspec[where x=j]] by auto qed - qed(auto simp add: not_le x(2,3) y(2,3) uv(3)) -qed - -lemma convex_hull_finite: - fixes s :: "'a::real_vector set" - assumes "finite s" - shows "convex hull s = {y. \u. (\x\s. 0 \ u x) \ - 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 *\<^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 - fix u v ::real assume uv:"0 \ u" "0 \ v" "u + v = 1" - fix ux assume ux:"\x\s. 0 \ ux x" "setsum ux s = (1::real)" - fix uy assume uy:"\x\s. 0 \ uy x" "setsum uy s = (1::real)" - { fix x assume "x\s" - hence "0 \ u * ux x + v * uy x" using ux(1)[THEN bspec[where x=x]] uy(1)[THEN bspec[where x=x]] and uv(1,2) - by (auto, metis add_nonneg_nonneg mult_nonneg_nonneg uv(1) uv(2)) } - moreover have "(\x\s. u * ux x + v * uy x) = 1" - unfolding setsum_addf and setsum_right_distrib[THEN sym] and ux(2) uy(2) using uv(3) by auto - moreover have "(\x\s. (u * ux x + v * uy x) *\<^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 *\<^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 - -subsection {* Another formulation from Lars Schewe. *} - -lemma setsum_constant_scaleR: - fixes y :: "'a::real_vector" - shows "(\x\A. y) = of_nat (card A) *\<^sub>R y" -apply (cases "finite A") -apply (induct set: finite) -apply (simp_all add: algebra_simps) -done - -lemma convex_hull_explicit: - fixes p :: "'a::real_vector set" - shows "convex hull p = {y. \s u. finite s \ s \ p \ - (\x\s. 0 \ u x) \ 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 *\<^sub>R y i) = x" - unfolding convex_hull_indexed by auto - - have fin:"finite {1..k}" by auto - have fin':"\v. finite {i \ {1..k}. y i = v}" by auto - { fix j assume "j\{1..k}" - hence "y j \ p" "0 \ setsum u {i. Suc 0 \ i \ i \ k \ y i = y j}" - using obt(1)[THEN bspec[where x=j]] and obt(2) apply simp - apply(rule setsum_nonneg) using obt(1) by auto } - moreover - have "(\v\y ` {1..k}. setsum u {i \ {1..k}. y i = v}) = 1" - unfolding setsum_image_gen[OF fin, THEN sym] using obt(2) by auto - moreover have "(\v\y ` {1..k}. setsum u {i \ {1..k}. y i = v} *\<^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 *\<^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 - - { fix i::nat assume "i\{1..card s}" - hence "f i \ s" apply(subst f(2)[THEN sym]) by auto - hence "0 \ u (f i)" "f i \ p" using obt(2,3) by auto } - moreover have *:"finite {1..card s}" by auto - { fix y assume "y\s" - then obtain i where "i\{1..card s}" "f i = y" using f using image_iff[of y f "{1..card s}"] by auto - hence "{x. Suc 0 \ x \ x \ card s \ f x = y} = {i}" apply auto using f(1)[unfolded inj_on_def] apply(erule_tac x=x in ballE) by auto - hence "card {x. Suc 0 \ x \ x \ card s \ f x = y} = 1" by auto - hence "(\x\{x \ {1..card s}. f x = y}. u (f x)) = u y" - "(\x\{x \ {1..card s}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y" - by (auto simp add: setsum_constant_scaleR) } - - 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 *\<^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 -qed - -subsection {* A stepping theorem for that expansion. *} - -lemma convex_hull_finite_step: - fixes s :: "'a::real_vector set" assumes "finite s" - 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 *\<^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 *\<^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 *\<^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 - -subsection {* Hence some special cases. *} - -lemma convex_hull_2: - "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 *\<^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 *\<^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,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" - "\x y z ::real^'n. x + y + z = 1 \ x = 1 - y - z" by (auto simp add: ring_simps) - show ?thesis unfolding convex_hull_finite[OF fin(1)] and Collect_def and convex_hull_finite_step[OF fin(2)] and * - unfolding convex_hull_finite_step[OF fin(3)] apply(rule ext) apply simp apply auto - apply(rule_tac x=va in exI) apply (rule_tac x="u c" in exI) apply simp - apply(rule_tac x="1 - v - w" in exI) apply simp apply(rule_tac x=v in exI) apply simp apply(rule_tac x="\x. w" in exI) by simp qed - -lemma convex_hull_3_alt: - "convex hull {a,b,c} = {a + u *\<^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 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. *} - -text {* TODO: Generalize linear algebra concepts defined in @{text -Euclidean_Space.thy} so that we can generalize these lemmas. *} - -lemma subspace_imp_affine: - fixes s :: "(real ^ _) set" shows "subspace s \ affine s" - 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 - -lemma subspace_imp_convex: - fixes s :: "(real ^ _) set" shows "subspace s \ convex s" - using subspace_imp_affine affine_imp_convex by auto - -lemma affine_hull_subset_span: - fixes s :: "(real ^ _) set" shows "(affine hull s) \ (span s)" - unfolding span_def apply(rule hull_antimono) unfolding subset_eq Ball_def mem_def - using subspace_imp_affine by auto - -lemma convex_hull_subset_span: - fixes s :: "(real ^ _) set" shows "(convex hull s) \ (span s)" - unfolding span_def apply(rule hull_antimono) unfolding subset_eq Ball_def mem_def - using subspace_imp_convex by auto - -lemma convex_hull_subset_affine_hull: "(convex hull s) \ (affine hull s)" - unfolding span_def apply(rule hull_antimono) unfolding subset_eq Ball_def mem_def - using affine_imp_convex by auto - -lemma affine_dependent_imp_dependent: - fixes s :: "(real ^ _) set" shows "affine_dependent s \ dependent s" - unfolding affine_dependent_def dependent_def - using affine_hull_subset_span by auto - -lemma dependent_imp_affine_dependent: - fixes s :: "(real ^ _) set" - assumes "dependent {x - a| x . x \ s}" "a \ s" - shows "affine_dependent (insert a s)" -proof- - 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 - have "0\S" using obt(2) assms(2) unfolding subset_eq by auto - have fin:"finite t" and "t\s" unfolding t_def using obt(1,2) by auto - - hence "finite (insert a t)" and "insert a t \ insert a s" by auto - moreover have *:"\P Q. (\x\t. (if x = a then P x else Q x)) = (\x\t. Q x)" - apply(rule setsum_cong2) using `a\s` `t\s` by auto - have "(\x\insert a t. if x = a then - (\x\t. u (x - a)) else u (x - a)) = 0" - unfolding setsum_clauses(2)[OF fin] using `a\s` `t\s` apply auto unfolding * by auto - moreover have "\v\insert a t. (if v = a then - (\x\t. u (x - a)) else u (v - a)) \ 0" - apply(rule_tac x="v + a" in bexI) using obt(3,4) and `0\S` unfolding t_def by auto - moreover have *:"\P Q. (\x\t. (if x = a then P x else Q x) *\<^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)) *\<^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 *\<^sub>R x) \ s)" (is "?lhs = ?rhs") -proof- - { fix x y assume "x\s" "y\s" and ?lhs - 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*\<^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 - -lemma affine_dependent_biggerset: fixes s::"(real^'n::finite) set" - assumes "finite s" "card s \ CARD('n) + 2" - shows "affine_dependent s" -proof- - have "s\{}" using assms by auto then obtain a where "a\s" by auto - have *:"{x - a |x. x \ s - {a}} = (\x. x - a) ` (s - {a})" by auto - have "card {x - a |x. x \ s - {a}} = card (s - {a})" unfolding * - apply(rule card_image) unfolding inj_on_def by auto - also have "\ > CARD('n)" using assms(2) - unfolding card_Diff_singleton[OF assms(1) `a\s`] by auto - finally show ?thesis apply(subst insert_Diff[OF `a\s`, THEN sym]) - apply(rule dependent_imp_affine_dependent) - apply(rule dependent_biggerset) by auto qed - -lemma affine_dependent_biggerset_general: - assumes "finite (s::(real^'n::finite) set)" "card s \ dim s + 2" - shows "affine_dependent s" -proof- - from assms(2) have "s \ {}" by auto - then obtain a where "a\s" by auto - have *:"{x - a |x. x \ s - {a}} = (\x. x - a) ` (s - {a})" by auto - have **:"card {x - a |x. x \ s - {a}} = card (s - {a})" unfolding * - apply(rule card_image) unfolding inj_on_def by auto - have "dim {x - a |x. x \ s - {a}} \ dim s" - apply(rule subset_le_dim) unfolding subset_eq - using `a\s` by (auto simp add:span_superset span_sub) - also have "\ < dim s + 1" by auto - also have "\ \ card (s - {a})" using assms - using card_Diff_singleton[OF assms(1) `a\s`] by auto - finally show ?thesis apply(subst insert_Diff[OF `a\s`, THEN sym]) - apply(rule dependent_imp_affine_dependent) apply(rule dependent_biggerset_general) unfolding ** by auto qed - -subsection {* Caratheodory's theorem. *} - -lemma convex_hull_caratheodory: fixes p::"(real^'n::finite) set" - shows "convex hull p = {y. \s u. finite s \ s \ p \ card s \ CARD('n) + 1 \ - (\x\s. 0 \ u x) \ setsum u s = 1 \ setsum (\v. u v *\<^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 *\<^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 *\<^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 *\<^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) - assume as:"\x\s. 0 \ w x" - hence "setsum w (s - {v}) \ 0" apply(rule_tac setsum_nonneg) by auto - hence "setsum w s > 0" unfolding setsum_diff1'[OF obt(1) `v\s`] - using as[THEN bspec[where x=v]] and `v\s` using `w v \ 0` by auto - thus False using wv(1) by auto - qed hence "i\{}" unfolding i_def by auto - - hence "t \ 0" using Min_ge_iff[of i 0 ] and obt(1) unfolding t_def i_def - using obt(4)[unfolded le_less] apply auto unfolding divide_le_0_iff by auto - have t:"\v\s. u v + t * w v \ 0" proof - fix v assume "v\s" hence v:"0\u v" using obt(4)[THEN bspec[where x=v]] by auto - show"0 \ u v + t * w v" proof(cases "w v < 0") - case False thus ?thesis apply(rule_tac add_nonneg_nonneg) - using v apply simp apply(rule mult_nonneg_nonneg) using `t\0` by auto next - case True hence "t \ u v / (- w v)" using `v\s` - unfolding t_def i_def apply(rule_tac Min_le) using obt(1) by auto - thus ?thesis unfolding real_0_le_add_iff - using pos_le_divide_eq[OF True[unfolded neg_0_less_iff_less[THEN sym]]] by auto - qed qed - - obtain a where "a\s" and "t = (\v. (u v) / (- w v)) a" and "w a < 0" - using Min_in[OF _ `i\{}`] and obt(1) unfolding i_def t_def by auto - hence a:"a\s" "u a + t * w a = 0" by auto - have *:"\f. setsum f (s - {a}) = setsum f s - ((f a)::'a::ring)" unfolding setsum_diff1'[OF obt(1) `a\s`] by auto - have "(\v\s. u v + t * w v) = 1" - unfolding setsum_addf wv(1) setsum_right_distrib[THEN sym] obt(5) by auto - moreover have "(\v\s. u v *\<^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: * 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 *\<^sub>R v) = y" using obt by auto -qed auto - -lemma caratheodory: - "convex hull p = {x::real^'n::finite. \s. finite s \ s \ p \ - card s \ CARD('n) + 1 \ x \ convex hull s}" - unfolding expand_set_eq apply(rule, rule) unfolding mem_Collect_eq proof- - fix x assume "x \ convex hull p" - then obtain s u where "finite s" "s \ p" "card s \ CARD('n) + 1" - "\x\s. 0 \ u x" "setsum u s = 1" "(\v\s. u v *\<^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 -next - fix x assume "\s. finite s \ s \ p \ card s \ CARD('n) + 1 \ x \ convex hull s" - then obtain s where "finite s" "s \ p" "card s \ CARD('n) + 1" "x \ convex hull s" by auto - thus "x \ convex hull p" using hull_mono[OF `s\p`] by auto -qed - -subsection {* Openness and compactness are preserved by convex hull operation. *} - -lemma open_convex_hull: - fixes s :: "'a::real_normed_vector set" - assumes "open s" - shows "open(convex hull s)" - unfolding open_contains_cball convex_hull_explicit unfolding mem_Collect_eq ball_simps(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 *\<^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 *\<^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\{}`] - using b apply simp apply rule apply(erule_tac x=x in ballE) using `t\s` by auto - next fix y assume "y \ cball a (Min i)" - hence y:"norm (a - y) \ Min i" unfolding dist_norm[THEN sym] by auto - { fix x assume "x\t" - hence "Min i \ b x" unfolding i_def apply(rule_tac Min_le) using obt(1) by auto - hence "x + (y - a) \ cball x (b x)" using y unfolding mem_cball dist_norm by auto - moreover from `x\t` have "x\s" using obt(2) by auto - ultimately have "x + (y - a) \ s" using y and b[THEN bspec[where x=x]] unfolding subset_eq by auto } - moreover - have *:"inj_on (\v. v + (y - a)) t" unfolding inj_on_def by auto - have "(\v\(\v. v + (y - a)) ` t. u (v - (y - a))) = 1" - unfolding setsum_reindex[OF *] o_def using obt(4) by auto - moreover have "(\v\(\v. v + (y - a)) ` t. u (v - (y - a)) *\<^sub>R v) = y" - unfolding setsum_reindex[OF *] o_def using obt(4,5) - 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 :: "'a::real_normed_vector set" - assumes "compact s" "compact t" - shows "compact { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ s \ y \ t}" -proof- - let ?X = "{0..1} \ s \ t" - let ?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" - assumes "compact s" shows "compact(convex hull s)" -proof(cases "s={}") - case True thus ?thesis using compact_empty by simp -next - case False then obtain w where "w\s" by auto - show ?thesis unfolding caratheodory[of s] - proof(induct "CARD('n) + 1") - have *:"{x.\sa. finite sa \ sa \ s \ card sa \ 0 \ x \ convex hull sa} = {}" - using compact_empty by (auto simp add: convex_hull_empty) - case 0 thus ?case unfolding * by simp - next - case (Suc n) - show ?case proof(cases "n=0") - case True have "{x. \t. finite t \ t \ s \ card t \ Suc n \ x \ convex hull t} = s" - unfolding expand_set_eq and mem_Collect_eq proof(rule, rule) - fix x assume "\t. finite t \ t \ s \ card t \ Suc n \ x \ convex hull t" - then obtain t where t:"finite t" "t \ s" "card t \ Suc n" "x \ convex hull t" by auto - show "x\s" proof(cases "card t = 0") - case True thus ?thesis using t(4) unfolding card_0_eq[OF t(1)] by(simp add: convex_hull_empty) - next - case False hence "card t = Suc 0" using t(3) `n=0` by auto - then obtain a where "t = {a}" unfolding card_Suc_eq by auto - thus ?thesis using t(2,4) by (simp add: convex_hull_singleton) - qed - next - fix x assume "x\s" - thus "\t. finite t \ t \ s \ card t \ Suc n \ x \ convex hull t" - apply(rule_tac x="{x}" in exI) unfolding convex_hull_singleton by auto - qed thus ?thesis using assms by simp - next - case False have "{x. \t. finite t \ t \ s \ card t \ Suc n \ x \ convex hull t} = - { (1 - u) *\<^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) *\<^sub>R u + c *\<^sub>R v \ - 0 \ c \ c \ 1 \ u \ s \ (\t. finite t \ t \ s \ card t \ n \ v \ convex hull t)" - then obtain u v c t where obt:"x = (1 - c) *\<^sub>R u + c *\<^sub>R v" - "0 \ c \ c \ 1" "u \ s" "finite t" "t \ s" "card t \ n" "v \ convex hull t" by auto - 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" - apply(rule_tac x="insert u t" in exI) by (auto simp add: card_insert_if) - next - fix x assume "\t. finite t \ t \ s \ card t \ Suc n \ x \ convex hull t" - then obtain t where t:"finite t" "t \ s" "card t \ Suc n" "x \ convex hull t" by auto - let ?P = "\u v c. x = (1 - c) *\<^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 - thus ?P apply(rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI) using `w\s` and t - by(auto intro!: exI[where x=t]) - next - case True then obtain a u where au:"t = insert a u" "a\u" apply(drule_tac card_eq_SucD) by auto - show ?P proof(cases "u={}") - case True hence "x=a" using t(4)[unfolded au] by auto - show ?P unfolding `x=a` apply(rule_tac x=a in exI, rule_tac x=a in exI, rule_tac x=1 in exI) - using t and `n\0` unfolding au by(auto intro!: exI[where x="{a}"] simp add: convex_hull_singleton) - next - case False obtain ux vx b where obt:"ux\0" "vx\0" "ux + vx = 1" "b \ convex hull u" "x = ux *\<^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) - using obt and t(1-3) unfolding au and * using card_insert_disjoint[OF _ au(2)] - by(auto intro!: exI[where x=u]) - qed - qed - qed - thus ?thesis using compact_convex_combinations[OF assms Suc] by simp - qed - qed -qed - -lemma finite_imp_compact_convex_hull: - fixes s :: "(real ^ _) set" - shows "finite s \ compact(convex hull s)" - apply(drule finite_imp_compact, drule compact_convex_hull) by assumption - -subsection {* Extremal points of a simplex are some vertices. *} - -lemma dist_increases_online: - fixes a b d :: "'a::real_inner" - assumes "d \ 0" - shows "dist a (b + d) > dist a b \ dist a (b - d) > dist a b" -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 norm_eq_sqrt_inner and real_sqrt_less_iff - by (simp add: algebra_simps inner_commute) -next - 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 norm_eq_sqrt_inner and real_sqrt_less_iff - by (simp add: algebra_simps inner_commute) -qed - -lemma norm_increases_online: - fixes d :: "'a::real_inner" - shows "d \ 0 \ norm(a + d) > norm a \ norm(a - d) > norm a" - using dist_increases_online[of d a 0] unfolding dist_norm by auto - -lemma simplex_furthest_lt: - fixes s::"'a::real_inner set" assumes "finite s" - shows "\x \ (convex hull s). x \ s \ (\y\(convex hull s). norm(x - a) < norm(y - a))" -proof(induct_tac rule: finite_induct[of s]) - fix x s assume as:"finite s" "x\s" "\x\convex hull s. x \ s \ (\y\convex hull s. norm (x - a) < norm (y - a))" - show "\xa\convex hull insert x s. xa \ insert x s \ (\y\convex hull insert x s. norm (xa - a) < norm (y - a))" - proof(rule,rule,cases "s = {}") - case False fix y assume y:"y \ convex hull insert x s" "y \ insert x s" - obtain u v b where obt:"u\0" "v\0" "u + v = 1" "b \ convex hull s" "y = u *\<^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") - case True then obtain z where "z\convex hull s" "norm (y - a) < norm (z - a)" - using as(3)[THEN bspec[where x=y]] and y(2) by auto - thus ?thesis apply(rule_tac x=z in bexI) unfolding convex_hull_insert[OF False] by auto - next - case False show ?thesis using obt(3) proof(cases "u=0", case_tac[!] "v=0") - assume "u=0" "v\0" hence "y = b" using obt by auto - thus ?thesis using False and obt(4) by auto - next - assume "u\0" "v=0" hence "y = x" using obt by auto - thus ?thesis using y(2) by auto - next - assume "u\0" "v\0" - then obtain w where w:"w>0" "wb" proof(rule ccontr) - assume "\ x\b" hence "y=b" unfolding obt(5) - using obt(3) by(auto simp add: scaleR_left_distrib[THEN sym]) - thus False using obt(4) and False by simp qed - 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 *\<^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 *\<^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 - qed - qed auto - qed - qed auto -qed (auto simp add: assms) - -lemma simplex_furthest_le: - fixes s :: "(real ^ _) set" - assumes "finite s" "s \ {}" - shows "\y\s. \x\(convex hull s). norm(x - a) \ norm(y - a)" -proof- - have "convex hull s \ {}" using hull_subset[of s convex] and assms(2) by auto - then obtain x where x:"x\convex hull s" "\y\convex hull s. norm (y - a) \ norm (x - a)" - using distance_attains_sup[OF finite_imp_compact_convex_hull[OF assms(1)], of a] - unfolding dist_commute[of a] unfolding dist_norm by auto - thus ?thesis proof(cases "x\s") - case False then obtain y where "y\convex hull s" "norm (x - a) < norm (y - a)" - using simplex_furthest_lt[OF assms(1), THEN bspec[where x=x]] and x(1) by auto - thus ?thesis using x(2)[THEN bspec[where x=y]] by auto - qed auto -qed - -lemma simplex_furthest_le_exists: - fixes s :: "(real ^ _) set" - shows "finite s \ (\x\(convex hull s). \y\s. norm(x - a) \ norm(y - a))" - using simplex_furthest_le[of s] by (cases "s={}")auto - -lemma simplex_extremal_le: - fixes s :: "(real ^ _) set" - assumes "finite s" "s \ {}" - shows "\u\s. \v\s. \x\convex hull s. \y \ convex hull s. norm(x - y) \ norm(u - v)" -proof- - have "convex hull s \ {}" using hull_subset[of s convex] and assms(2) by auto - then obtain u v where obt:"u\convex hull s" "v\convex hull s" - "\x\convex hull s. \y\convex hull s. norm (x - y) \ norm (u - v)" - using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]] by auto - thus ?thesis proof(cases "u\s \ v\s", erule_tac disjE) - assume "u\s" then obtain y where "y\convex hull s" "norm (u - v) < norm (y - v)" - using simplex_furthest_lt[OF assms(1), THEN bspec[where x=u]] and obt(1) by auto - thus ?thesis using obt(3)[THEN bspec[where x=y], THEN bspec[where x=v]] and obt(2) by auto - next - assume "v\s" then obtain y where "y\convex hull s" "norm (v - u) < norm (y - u)" - using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2) by auto - thus ?thesis using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1) - by (auto simp add: norm_minus_commute) - qed auto -qed - -lemma simplex_extremal_le_exists: - fixes s :: "(real ^ _) set" - shows "finite s \ x \ convex hull s \ y \ convex hull s - \ (\u\s. \v\s. norm(x - y) \ norm(u - v))" - using convex_hull_empty simplex_extremal_le[of s] by(cases "s={}")auto - -subsection {* Closest point of a convex set is unique, with a continuous projection. *} - -definition - closest_point :: "(real ^ 'n::finite) set \ real ^ 'n \ real ^ 'n" where - "closest_point s a = (SOME x. x \ s \ (\y\s. dist a x \ dist a y))" - -lemma closest_point_exists: - assumes "closed s" "s \ {}" - shows "closest_point s a \ s" "\y\s. dist a (closest_point s a) \ dist a y" - unfolding closest_point_def apply(rule_tac[!] someI2_ex) - using distance_attains_inf[OF assms(1,2), of a] by auto - -lemma closest_point_in_set: - "closed s \ s \ {} \ (closest_point s a) \ s" - by(meson closest_point_exists) - -lemma closest_point_le: - "closed s \ x \ s \ dist a (closest_point s a) \ dist a x" - using closest_point_exists[of s] by auto - -lemma closest_point_self: - assumes "x \ s" shows "closest_point s x = x" - unfolding closest_point_def apply(rule some1_equality, rule ex1I[of _ x]) - using assms by auto - -lemma closest_point_refl: - "closed s \ s \ {} \ (closest_point s x = x \ x \ s)" - using closest_point_in_set[of s x] closest_point_self[of x s] by auto - -(* 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 "inner y z > 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 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: - fixes s :: "(real ^ _) set" - assumes "convex s" "closed s" "x \ s" "y \ s" "\z\s. dist a x \ dist a z" - shows "inner (a - x) (y - x) \ 0" -proof(rule ccontr) assume "\ 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: - fixes s :: "(real ^ _) set" - assumes "convex s" "closed s" "x \ s" "y \ s" - "\z\s. dist a x \ dist a z" "\z\s. dist a y \ dist a z" - shows "x = y" using any_closest_point_dot[OF assms(1-4,5)] and any_closest_point_dot[OF assms(1-2,4,3,6)] - unfolding norm_pths(1) and norm_le_square - by (auto simp add: algebra_simps) - -lemma closest_point_unique: - assumes "convex s" "closed s" "x \ s" "\z\s. dist a x \ dist a z" - shows "x = closest_point s a" - using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point s a"] - using closest_point_exists[OF assms(2)] and assms(3) by auto - -lemma closest_point_dot: - assumes "convex s" "closed s" "x \ s" - shows "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 - -lemma closest_point_lt: - assumes "convex s" "closed s" "x \ s" "x \ closest_point s a" - shows "dist a (closest_point s a) < dist a x" - apply(rule ccontr) apply(rule_tac notE[OF assms(4)]) - apply(rule closest_point_unique[OF assms(1-3), of a]) - using closest_point_le[OF assms(2), of _ a] by fastsimp - -lemma closest_point_lipschitz: - assumes "convex s" "closed s" "s \ {}" - shows "dist (closest_point s x) (closest_point s y) \ dist x y" -proof- - have "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 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 \ {}" - shows "continuous (at x) (closest_point s)" - unfolding continuous_at_eps_delta - using le_less_trans[OF closest_point_lipschitz[OF assms]] by auto - -lemma continuous_on_closest_point: - assumes "convex s" "closed s" "s \ {}" - shows "continuous_on t (closest_point s)" - apply(rule continuous_at_imp_continuous_on) using continuous_at_closest_point[OF assms] by auto - -subsection {* Various point-to-set separating/supporting hyperplane theorems. *} - -lemma supporting_hyperplane_closed_point: - fixes s :: "(real ^ _) set" - assumes "convex s" "closed s" "s \ {}" "z \ s" - 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="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 "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) *\<^sub>R y + u *\<^sub>R x)" - using assms(1)[unfolded convex_alt] and y and `x\s` and `y\s` by auto - assume "\ inner (y - z) y \ inner (y - z) x" then obtain v where - "v>0" "v\1" "dist (y + v *\<^sub>R (x - y)) z < dist y z" using closer_point_lemma[of z y x] 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: - fixes s :: "(real ^ _) set" - assumes "convex s" "closed s" "z \ s" - 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 _ 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="inner (y - z) z + (norm(y - z))\ / 2" in exI) - apply rule defer apply rule proof- - fix x assume "x\s" - 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 algebra_simps) qed - moreover have "0 < norm (y - z) ^ 2" using `y\s` `z\s` by auto - 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. 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 - thus ?thesis apply(rule_tac x="basis a" in exI, rule_tac x=1 in exI) using True by auto -next case False thus ?thesis using False using separating_hyperplane_closed_point[OF assms] - apply - apply(erule exE)+ unfolding dot_rzero apply(rule_tac x=a in exI, rule_tac x=b in exI) by auto qed - -subsection {* Now set-to-set for closed/compact sets. *} - -lemma separating_hyperplane_closed_compact: - assumes "convex (s::(real^'n::finite) set)" "closed s" "convex t" "compact t" "t \ {}" "s \ t = {}" - shows "\a b. (\x\s. 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:"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 < 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 + 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 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. 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 \ 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 < inner a x" using `0 < b` by auto - qed -qed - -lemma separating_hyperplane_compact_closed: - fixes s :: "(real ^ _) set" - assumes "convex s" "compact s" "s \ {}" "convex t" "closed t" "s \ t = {}" - 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 - -subsection {* General case without assuming closure and getting non-strict separation. *} - -lemma separating_hyperplane_set_0: - assumes "convex s" "(0::real^'n::finite) \ s" - shows "\a. a \ 0 \ (\x\s. 0 \ 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 < 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 \ 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: 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: 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. inner a x \ b) \ (\x\t. inner a x \ b)" -proof- from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]] - obtain a where "a\0" "\x\{x - y |x y. x \ t \ y \ s}. 0 \ inner a x" 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 - -subsection {* More convexity generalities. *} - -lemma convex_closure: - fixes s :: "'a::real_normed_vector set" - 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 *\<^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 - -lemma convex_interior: - fixes s :: "'a::real_normed_vector set" - assumes "convex s" shows "convex(interior s)" - unfolding convex_alt Ball_def mem_interior apply(rule,rule,rule,rule,rule,rule) apply(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) *\<^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) *\<^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: 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 - -subsection {* Moving and scaling convex hulls. *} - -lemma convex_hull_translation_lemma: - "convex hull ((\x. a + x) ` s) \ (\x. a + x) ` (convex hull s)" - apply(rule hull_minimal, rule image_mono, rule hull_subset) unfolding mem_def - using convex_translation[OF convex_convex_hull, of a s] by assumption - -lemma convex_hull_bilemma: fixes neg - assumes "(\s a. (convex hull (up a s)) \ up a (convex hull s))" - shows "(\s. up a (up (neg a) s) = s) \ (\s. up (neg a) (up a s) = s) \ (\s t a. s \ t \ up a s \ up a t) - \ \s. (convex hull (up a s)) = up a (convex hull s)" - using assms by(metis subset_antisym) - -lemma convex_hull_translation: - "convex hull ((\x. a + x) ` s) = (\x. a + x) ` (convex hull s)" - apply(rule convex_hull_bilemma[rule_format, of _ _ "\a. -a"], rule convex_hull_translation_lemma) unfolding image_image by auto - -lemma convex_hull_scaling_lemma: - "(convex hull ((\x. c *\<^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 *\<^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 scaleR_scaleR by(auto simp add:image_constant_conv convex_hull_eq_empty) - -lemma convex_hull_affinity: - "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: - fixes s :: "(real ^ _) set" - assumes "closed s" "convex s" - 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. 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 - -subsection {* Radon's theorem (from Lars Schewe). *} - -lemma radon_ex_lemma: - assumes "finite c" "affine_dependent c" - shows "\u. setsum u c = 0 \ (\v\c. u v \ 0) \ setsum (\v. u v *\<^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 scaleR_zero_left - and setsum_restrict_set[OF assms(1), THEN sym] by(auto simp add: Int_absorb1) qed - -lemma radon_s_lemma: - assumes "finite s" "setsum f s = (0::real)" - shows "setsum f {x\s. 0 < f x} = - setsum f {x\s. f x < 0}" -proof- have *:"\x. (if f x < 0 then f x else 0) + (if 0 < f x then f x else 0) = f x" by auto - show ?thesis unfolding real_add_eq_0_iff[THEN sym] and setsum_restrict_set''[OF assms(1)] and setsum_addf[THEN sym] and * - using assms(2) by assumption qed - -lemma radon_v_lemma: - assumes "finite s" "setsum f s = 0" "\x. g x = (0::real) \ f x = (0::real^'n)" - shows "(setsum f {x\s. 0 < g x}) = - setsum f {x\s. g x < 0}" -proof- - have *:"\x. (if 0 < g x then f x else 0) + (if g x < 0 then f x else 0) = f x" using assms(3) by auto - show ?thesis unfolding eq_neg_iff_add_eq_0 and setsum_restrict_set''[OF assms(1)] and setsum_addf[THEN sym] and * - using assms(2) by assumption qed - -lemma radon_partition: - assumes "finite c" "affine_dependent c" - shows "\m p. m \ p = {} \ m \ p = c \ (convex hull m) \ (convex hull p) \ {}" proof- - obtain u v where uv:"setsum u c = 0" "v\c" "u v \ 0" "(\v\c. u v *\<^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})) *\<^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") - case True thus ?thesis using setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] by auto - next - case False hence "setsum u c \ setsum (\x. if x=v then u v else 0) c" apply(rule_tac setsum_mono) by auto - thus ?thesis unfolding setsum_delta[OF assms(1)] using uv(2) and `u v < 0` and uv(1) by auto qed - qed (insert setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] uv(2-3), auto) - - hence *:"setsum u {x\c. u x > 0} > 0" unfolding real_less_def apply(rule_tac conjI, rule_tac setsum_nonneg) by auto - moreover have "setsum u ({x \ c. 0 < u x} \ {x \ c. u x < 0}) = setsum u c" - "(\x\{x \ c. 0 < u x} \ {x \ c. u x < 0}. u x *\<^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 *\<^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 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 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 - -lemma radon: assumes "affine_dependent c" - obtains m p where "m\c" "p\c" "m \ p = {}" "(convex hull m) \ (convex hull p) \ {}" -proof- from assms[unfolded affine_dependent_explicit] guess s .. then guess u .. - hence *:"finite s" "affine_dependent s" and s:"s \ c" unfolding affine_dependent_explicit by auto - from radon_partition[OF *] guess m .. then guess p .. - thus ?thesis apply(rule_tac that[of p m]) using s by auto qed - -subsection {* Helly's theorem. *} - -lemma helly_induct: fixes f::"(real^'n::finite) set set" - assumes "f hassize n" "n \ CARD('n) + 1" - "\s\f. convex s" "\t\f. card t = CARD('n) + 1 \ \ t \ {}" - shows "\ f \ {}" - using assms unfolding hassize_def apply(erule_tac conjE) proof(induct n arbitrary: f) -case (Suc n) -show "\ f \ {}" apply(cases "n = CARD('n)") apply(rule Suc(4)[rule_format]) - unfolding card_Diff_singleton_if[OF Suc(5)] and Suc(6) proof- - assume ng:"n \ CARD('n)" hence "\X. \s\f. X s \ \(f - {s})" apply(rule_tac bchoice) unfolding ex_in_conv - apply(rule, rule Suc(1)[rule_format]) unfolding card_Diff_singleton_if[OF Suc(5)] and Suc(6) - defer apply(rule Suc(3)[rule_format]) defer apply(rule Suc(4)[rule_format]) using Suc(2,5) by auto - then obtain X where X:"\s\f. X s \ \(f - {s})" by auto - show ?thesis proof(cases "inj_on X f") - case False then obtain s t where st:"s\t" "s\f" "t\f" "X s = X t" unfolding inj_on_def by auto - hence *:"\ f = \ (f - {s}) \ \ (f - {t})" by auto - show ?thesis unfolding * unfolding ex_in_conv[THEN sym] apply(rule_tac x="X s" in exI) - apply(rule, rule X[rule_format]) using X st by auto - next case True then obtain m p where mp:"m \ p = {}" "m \ p = X ` f" "convex hull m \ convex hull p \ {}" - using radon_partition[of "X ` f"] and affine_dependent_biggerset[of "X ` f"] - unfolding card_image[OF True] and Suc(6) using Suc(2,5) and ng by auto - have "m \ X ` f" "p \ X ` f" using mp(2) by auto - then obtain g h where gh:"m = X ` g" "p = X ` h" "g \ f" "h \ f" unfolding subset_image_iff by auto - hence "f \ (g \ h) = f" by auto - hence f:"f = g \ h" using inj_on_Un_image_eq_iff[of X f "g \ h"] and True - unfolding mp(2)[unfolded image_Un[THEN sym] gh] by auto - have *:"g \ h = {}" using mp(1) unfolding gh using inj_on_image_Int[OF True gh(3,4)] by auto - have "convex hull (X ` h) \ \ g" "convex hull (X ` g) \ \ h" - apply(rule_tac [!] hull_minimal) using Suc(3) gh(3-4) unfolding mem_def unfolding subset_eq - apply(rule_tac [2] convex_Inter, rule_tac [4] convex_Inter) apply rule prefer 3 apply rule proof- - fix x assume "x\X ` g" then guess y unfolding image_iff .. - thus "x\\h" using X[THEN bspec[where x=y]] using * f by auto next - fix x assume "x\X ` h" then guess y unfolding image_iff .. - thus "x\\g" using X[THEN bspec[where x=y]] using * f by auto - qed(auto) - thus ?thesis unfolding f using mp(3)[unfolded gh] by blast qed -qed(insert dimindex_ge_1, auto) qed(auto) - -lemma helly: fixes f::"(real^'n::finite) set set" - assumes "finite f" "card f \ CARD('n) + 1" "\s\f. convex s" - "\t\f. card t = CARD('n) + 1 \ \ t \ {}" - shows "\ f \{}" - apply(rule helly_induct) unfolding hassize_def using assms by auto - -subsection {* Convex hull is "preserved" by a linear function. *} - -lemma convex_hull_linear_image: - assumes "bounded_linear f" - shows "f ` (convex hull s) = convex hull (f ` s)" - apply rule unfolding subset_eq ball_simps apply(rule_tac[!] hull_induct, rule hull_inc) prefer 3 - apply(erule imageE)apply(rule_tac x=xa in image_eqI) apply assumption - apply(rule hull_subset[unfolded subset_eq, rule_format]) apply assumption -proof- - interpret f: bounded_linear f by fact - show "convex {x. f x \ convex hull f ` s}" - unfolding convex_def by(auto simp add: f.scaleR f.add convex_convex_hull[unfolded convex_def, rule_format]) next - interpret f: bounded_linear f by fact - show "convex {x. x \ f ` (convex hull s)}" using convex_convex_hull[unfolded convex_def, of s] - unfolding convex_def by (auto simp add: f.scaleR [symmetric] f.add [symmetric]) -qed auto - -lemma in_convex_hull_linear_image: - assumes "bounded_linear f" "x \ convex hull s" - shows "(f x) \ convex hull (f ` s)" -using convex_hull_linear_image[OF assms(1)] assms(2) by auto - -subsection {* Homeomorphism of all convex compact sets with nonempty interior. *} - -lemma compact_frontier_line_lemma: - fixes s :: "(real ^ _) set" - assumes "compact s" "0 \ s" "x \ 0" - 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 *\<^sub>R x)}" - have A:"?A = (\u. dest_vec1 u *\<^sub>R x) ` {0 .. vec1 (b / norm x)}" - unfolding image_image[of "\u. u *\<^sub>R x" "\x. dest_vec1 x", THEN sym] - unfolding dest_vec1_inverval vec1_dest_vec1 by auto - have "compact ?A" unfolding A apply(rule compact_continuous_image, rule continuous_at_imp_continuous_on) - apply(rule, rule continuous_vmul) - apply (rule continuous_dest_vec1) - apply(rule continuous_at_id) by(rule compact_interval) - moreover have "{y. \u\0. u \ b / norm x \ y = u *\<^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 *\<^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 *\<^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`] 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) using `u\0` `norm x > 0` `v>u` by(auto simp add:field_simps) - } note u_max = this - - 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 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 *\<^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) *\<^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) - 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 *\<^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 *\<^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 *\<^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 *\<^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 *\<^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" - apply(rule homeomorphism_compact) apply(rule compact_frontier[OF assms(1)]) - apply(rule continuous_on_subset[OF contpi]) defer apply(rule set_ext,rule) - unfolding inj_on_def prefer 3 apply(rule,rule,rule) - proof- fix x assume "x\pi ` frontier s" then obtain y where "y\frontier s" "x = pi y" by auto - thus "x \ sphere" using pi(1)[of y] and `0 \ frontier s` by auto - next fix x assume "x\sphere" hence "norm x = 1" "x\0" unfolding sphere_def by auto - then obtain u where "0 \ u" "u *\<^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 *\<^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 *\<^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 - using x y and front_smul[THEN bspec, OF as(1), THEN spec[where x="norm y * (inverse (norm x))"]] - using front_smul[THEN bspec, OF as(2), THEN spec[where x="norm x * (inverse (norm y))"]] - using xys nor by(auto simp add:field_simps divide_le_eq_1 divide_inverse[THEN sym]) - thus "x = y" apply(subst injpi[THEN sym]) using as(3) by auto - qed(insert `0 \ frontier s`, auto) - then obtain surf where surf:"\x\frontier s. surf (pi x) = x" "pi ` frontier s = sphere" "continuous_on (frontier s) pi" - "\y\sphere. pi (surf y) = y" "surf ` sphere = frontier s" "continuous_on sphere surf" unfolding homeomorphism_def by auto - - have cont_surfpi:"continuous_on (UNIV - {0}) (surf \ pi)" apply(rule continuous_on_compose, rule contpi) - apply(rule continuous_on_subset[of sphere], rule surf(6)) using pi(1) by auto - - { fix x assume as:"x \ cball (0::real^'n) 1" - have "norm x *\<^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]) - unfolding surf(5)[THEN sym] by auto - next case True thus ?thesis apply rule defer unfolding pi_def apply(rule fs[unfolded subset_eq, rule_format]) - unfolding surf(5)[unfolded sphere_def, THEN sym] using `0\s` by auto qed } note hom = this - - { fix x assume "x\s" - hence "x \ (\x. norm x *\<^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 *\<^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) *\<^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))) *\<^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))) *\<^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 *\<^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 *\<^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 - hence "B > 0" using assms(2) unfolding subset_eq apply(erule_tac x="basis a" in ballE) defer apply(erule_tac x="basis a" in ballE) - unfolding Ball_def mem_cball dist_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 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 - hence "norm x * norm (surf (pi x)) \ norm x * B" using as(2) by auto - also have "\ < e / B * B" apply(rule mult_strict_right_mono) using as(1) `B>0` by auto - also have "\ = e" using `B>0` by auto - finally show "norm x * norm (surf (pi x)) < e" by assumption - qed(insert `B>0`, auto) qed - next { fix x assume as:"surf (pi x) = 0" - have "x = 0" proof(rule ccontr) - assume "x\0" hence "pi x \ sphere" using pi(1) by auto - hence "surf (pi x) \ frontier s" using surf(5) by auto - thus False using `0\frontier s` unfolding as by simp qed - } note surf_0 = this - show "inj_on (\x. norm x *\<^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 - hence "pi (surf (pi x)) = pi (surf (pi y))" using as(3) - using pi(2)[of "norm x" "surf (pi x)"] pi(2)[of "norm y" "surf (pi y)"] by auto - moreover have "pi x \ sphere" "pi y \ sphere" using pi(1) False by auto - ultimately have *:"pi x = pi y" using surf(4)[THEN bspec[where x="pi x"]] surf(4)[THEN bspec[where x="pi y"]] by auto - moreover have "norm x = norm y" using as(3)[unfolded *] using False by(auto dest:surf_0) - ultimately show ?thesis using injpi by auto qed qed - qed auto qed - -lemma homeomorphic_convex_compact_lemma: fixes s::"(real^'n::finite) set" - assumes "convex s" "compact s" "cball 0 1 \ s" - shows "s homeomorphic (cball (0::real^'n) 1)" - apply(rule starlike_compact_projective[OF assms(2-3)]) proof(rule,rule,rule,erule conjE) - fix x u assume as:"x \ s" "0 \ u" "u < (1::real)" - hence "u *\<^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 *\<^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_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) *\<^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" - shows "s homeomorphic (cball (b::real^'n::finite) e)" -proof- obtain a where "a\interior s" using assms(3) by auto - then obtain d where "d>0" and d:"cball a d \ s" unfolding mem_interior_cball by auto - let ?d = "inverse d" and ?n = "0::real^'n" - have "cball ?n 1 \ (\x. inverse d *\<^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 *\<^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 *\<^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 \ {}" - "convex t" "compact t" "interior t \ {}" - shows "s homeomorphic t" - using assms by(meson zero_less_one homeomorphic_trans homeomorphic_convex_compact_cball homeomorphic_sym) - -subsection {* Epigraphs of convex functions. *} - -definition "epigraph s (f::real^'n \ real) = {xy. fstcart xy \ s \ f(fstcart xy) \ dest_vec1 (sndcart xy)}" - -lemma mem_epigraph: "(pastecart x (vec1 y)) \ epigraph s f \ x \ s \ f x \ y" unfolding epigraph_def by auto - -lemma convex_epigraph: - "convex(epigraph s f) \ convex_on s f \ convex s" - unfolding convex_def convex_on_def unfolding Ball_def forall_pastecart epigraph_def - unfolding mem_Collect_eq fstcart_pastecart sndcart_pastecart sndcart_add sndcart_cmul [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" - shows "convex(epigraph s f)" using assms unfolding convex_epigraph by auto - -lemma convex_epigraph_convex: "convex s \ (convex_on s f \ convex(epigraph s f))" - using convex_epigraph by auto - -subsection {* Use this to derive general bound property of convex function. *} - -lemma forall_of_pastecart: - "(\p. P (\x. fstcart (p x)) (\x. sndcart (p x))) \ (\x y. P x y)" apply meson - apply(erule_tac x="\a. pastecart (x a) (y a)" in allE) unfolding o_def by auto - -lemma forall_of_pastecart': - "(\p. P (fstcart p) (sndcart p)) \ (\x y. P x y)" apply meson - apply(erule_tac x="pastecart x y" in allE) unfolding o_def by auto - -lemma forall_of_dest_vec1: "(\v. P (\x. dest_vec1 (v x))) \ (\x. P x)" - apply rule apply rule apply(erule_tac x="(vec1 \ x)" in allE) unfolding o_def vec1_dest_vec1 by auto - -lemma forall_of_dest_vec1': "(\v. P (dest_vec1 v)) \ (\x. P x)" - apply rule apply rule apply(erule_tac x="(vec1 x)" in allE) defer apply rule - apply(erule_tac x="dest_vec1 v" in allE) unfolding o_def vec1_dest_vec1 by auto - -lemma convex_on: - fixes s :: "(real ^ _) set" - 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 *\<^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 [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) - defer apply(rule setsum_mono) apply(erule conjE)+ apply(erule_tac x=i in allE)apply(rule mult_left_mono) - using assms[unfolded convex] by auto - -subsection {* Convexity of general and special intervals. *} - -lemma is_interval_convex: - fixes s :: "(real ^ _) set" - assumes "is_interval s" shows "convex s" - unfolding convex_def apply(rule,rule,rule,rule,rule,rule,rule) proof- - fix x y u v assume as:"x \ s" "y \ s" "0 \ u" "0 \ v" "u + v = (1::real)" - hence *:"u = 1 - v" "1 - v \ 0" and **:"v = 1 - u" "1 - u \ 0" by auto - { fix a b assume "\ b \ u * a + v * b" - hence "u * a < (1 - v) * b" unfolding not_le using as(4) by(auto simp add: field_simps) - hence "a < b" unfolding * using as(4) *(2) apply(rule_tac mult_left_less_imp_less[of "1 - v"]) by(auto simp add: field_simps) - hence "a \ u * a + v * b" unfolding * using as(4) by (auto simp add: field_simps intro!:mult_right_mono) - } moreover - { fix a b assume "\ u * a + v * b \ a" - hence "v * b > (1 - u) * a" unfolding not_le using as(4) by(auto simp add: field_simps) - hence "a < b" unfolding * using as(4) apply(rule_tac mult_left_less_imp_less) by(auto simp add: ring_simps) - hence "u * a + v * b \ b" unfolding ** using **(2) as(3) by(auto simp add: field_simps intro!:mult_right_mono) } - ultimately show "u *\<^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: - fixes s :: "(real ^ _) set" - shows "is_interval s \ connected s" - using is_interval_convex convex_connected by auto - -lemma convex_interval: "convex {a .. b}" "convex {a<.. (\a\s. \b\s. \ x. dest_vec1 a \ dest_vec1 x \ dest_vec1 x \ dest_vec1 b \ x \ s)" - unfolding is_interval_def dest_vec1_def forall_1 by auto - -lemma is_interval_connected_1: "is_interval s \ connected (s::(real^1) set)" - apply(rule, rule is_interval_connected, assumption) unfolding is_interval_1 - apply(rule,rule,rule,rule,erule conjE,rule ccontr) proof- - fix a b x assume as:"connected s" "a \ s" "b \ s" "dest_vec1 a \ dest_vec1 x" "dest_vec1 x \ dest_vec1 b" "x\s" - hence *:"dest_vec1 a < dest_vec1 x" "dest_vec1 x < dest_vec1 b" apply(rule_tac [!] ccontr) unfolding not_less by auto - let ?halfl = "{z. 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: 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) - apply(rule, rule open_halfspace_lt, rule, rule open_halfspace_gt) apply(rule, rule, rule ccontr) - by(auto simp add: basis_component field_simps) qed - -lemma is_interval_convex_1: - "is_interval s \ convex (s::(real^1) set)" - using is_interval_convex convex_connected is_interval_connected_1 by auto - -lemma convex_connected_1: - "connected s \ convex (s::(real^1) set)" - using is_interval_convex convex_connected is_interval_connected_1 by auto - -subsection {* Another intermediate value theorem formulation. *} - -lemma ivt_increasing_component_on_1: fixes f::"real^1 \ real^'n::finite" - assumes "dest_vec1 a \ dest_vec1 b" "continuous_on {a .. b} f" "(f a)$k \ y" "y \ (f b)$k" - shows "\x\{a..b}. (f x)$k = y" -proof- have "f a \ f ` {a..b}" "f b \ f ` {a..b}" apply(rule_tac[!] imageI) - using assms(1) by(auto simp add: vector_less_eq_def dest_vec1_def) - thus ?thesis using connected_ivt_component[of "f ` {a..b}" "f a" "f b" k y] - using connected_continuous_image[OF assms(2) convex_connected[OF convex_interval(1)]] - using assms by(auto intro!: imageI) qed - -lemma ivt_increasing_component_1: fixes f::"real^1 \ real^'n::finite" - assumes "dest_vec1 a \ dest_vec1 b" - "\x\{a .. b}. continuous (at x) f" "f a$k \ y" "y \ f b$k" - shows "\x\{a..b}. (f x)$k = y" - apply(rule ivt_increasing_component_on_1) using assms using continuous_at_imp_continuous_on by auto - -lemma ivt_decreasing_component_on_1: fixes f::"real^1 \ real^'n::finite" - assumes "dest_vec1 a \ dest_vec1 b" "continuous_on {a .. b} f" "(f b)$k \ y" "y \ (f a)$k" - shows "\x\{a..b}. (f x)$k = y" - apply(subst neg_equal_iff_equal[THEN sym]) unfolding vector_uminus_component[THEN sym] - apply(rule ivt_increasing_component_on_1) using assms using continuous_on_neg - by(auto simp add:vector_uminus_component) - -lemma ivt_decreasing_component_1: fixes f::"real^1 \ real^'n::finite" - assumes "dest_vec1 a \ dest_vec1 b" "\x\{a .. b}. continuous (at x) f" "f b$k \ y" "y \ f a$k" - shows "\x\{a..b}. (f x)$k = y" - apply(rule ivt_decreasing_component_on_1) using assms using continuous_at_imp_continuous_on by auto - -subsection {* A bound within a convex hull, and so an interval. *} - -lemma convex_on_convex_hull_bound: - fixes s :: "(real ^ _) set" - 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 *\<^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) - using assms(2) obt(1) by auto - thus "f x \ b" using assms(1)[unfolded convex_on[OF convex_convex_hull], rule_format, of k u v] - unfolding obt(2-3) using obt(1) and hull_subset[unfolded subset_eq, rule_format, of _ s] by auto qed - -lemma unit_interval_convex_hull: - "{0::real^'n::finite .. 1} = convex hull {x. \i. (x$i = 0) \ (x$i = 1)}" (is "?int = convex hull ?points") -proof- have 01:"{0,1} \ convex hull ?points" apply rule apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) by auto - { fix n x assume "x\{0::real^'n .. 1}" "n \ CARD('n)" "card {i. x$i \ 0} \ n" - hence "x\convex hull ?points" proof(induct n arbitrary: x) - case 0 hence "x = 0" apply(subst Cart_eq) apply rule by auto - thus "x\convex hull ?points" using 01 by auto - next - case (Suc n) show "x\convex hull ?points" proof(cases "{i. x$i \ 0} = {}") - case True hence "x = 0" unfolding Cart_eq by auto - thus "x\convex hull ?points" using 01 by auto - next - case False def xi \ "Min ((\i. x$i) ` {i. x$i \ 0})" - have "xi \ (\i. x$i) ` {i. x$i \ 0}" unfolding xi_def apply(rule Min_in) using False by auto - then obtain i where i':"x$i = xi" "x$i \ 0" by auto - have i:"\j. x$j > 0 \ x$i \ x$j" - unfolding i'(1) xi_def apply(rule_tac Min_le) unfolding image_iff - defer apply(rule_tac x=j in bexI) using i' by auto - have i01:"x$i \ 1" "x$i > 0" using Suc(2)[unfolded mem_interval,rule_format,of i] using i'(2) `x$i \ 0` - by(auto simp add: Cart_lambda_beta) - show ?thesis proof(cases "x$i=1") - case True have "\j\{i. x$i \ 0}. x$j = 1" apply(rule, rule ccontr) unfolding mem_Collect_eq proof- - fix j assume "x $ j \ 0" "x $ j \ 1" - hence j:"x$j \ {0<..<1}" using Suc(2) by(auto simp add: vector_less_eq_def elim!:allE[where x=j]) - hence "x$j \ op $ x ` {i. x $ i \ 0}" by auto - hence "x$j \ x$i" unfolding i'(1) xi_def apply(rule_tac Min_le) by auto - thus False using True Suc(2) j by(auto simp add: vector_less_eq_def elim!:ballE[where x=j]) qed - thus "x\convex hull ?points" apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) - by(auto simp add: Cart_lambda_beta) - next let ?y = "\j. if x$j = 0 then 0 else (x$j - x$i) / (1 - x$i)" - case False hence *:"x = x$i *\<^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 - using Suc(2)[unfolded mem_interval, rule_format, of j] by(auto simp add:field_simps Cart_lambda_beta) - hence "0 \ ?y j \ ?y j \ 1" by auto } - moreover have "i\{j. x$j \ 0} - {j. ((\ j. ?y j)::real^'n) $ j \ 0}" using i01 by(auto simp add: Cart_lambda_beta) - hence "{j. x$j \ 0} \ {j. ((\ j. ?y j)::real^'n::finite) $ j \ 0}" by auto - hence **:"{j. ((\ j. ?y j)::real^'n::finite) $ j \ 0} \ {j. x$j \ 0}" apply - apply rule by(auto simp add: Cart_lambda_beta) - have "card {j. ((\ j. ?y j)::real^'n) $ j \ 0} \ n" using less_le_trans[OF psubset_card_mono[OF _ **] Suc(4)] by auto - ultimately show ?thesis apply(subst *) apply(rule convex_convex_hull[unfolded convex_def, rule_format]) - apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) defer apply(rule Suc(1)) - unfolding mem_interval using i01 Suc(3) by (auto simp add: Cart_lambda_beta) - qed qed qed } note * = this - show ?thesis apply rule defer apply(rule hull_minimal) unfolding subset_eq prefer 3 apply rule - apply(rule_tac n2="CARD('n)" in *) prefer 3 apply(rule card_mono) using 01 and convex_interval(1) prefer 5 apply - apply rule - unfolding mem_interval apply rule unfolding mem_Collect_eq apply(erule_tac x=i in allE) - by(auto simp add: vector_less_eq_def mem_def[of _ convex]) qed - -subsection {* And this is a finite set of vertices. *} - -lemma unit_cube_convex_hull: obtains s where "finite s" "{0 .. 1::real^'n::finite} = convex hull s" - apply(rule that[of "{x::real^'n::finite. \i. x$i=0 \ x$i=1}"]) - apply(rule finite_subset[of _ "(\s. (\ i. if i\s then 1::real else 0)::real^'n::finite) ` UNIV"]) - prefer 3 apply(rule unit_interval_convex_hull) apply rule unfolding mem_Collect_eq proof- - fix x::"real^'n" assume as:"\i. x $ i = 0 \ x $ i = 1" - show "x \ (\s. \ i. if i \ s then 1 else 0) ` UNIV" apply(rule image_eqI[where x="{i. x$i = 1}"]) - unfolding Cart_eq using as by(auto simp add:Cart_lambda_beta) qed auto - -subsection {* Hence any cube (could do any nonempty interval). *} - -lemma cube_convex_hull: - assumes "0 < d" obtains s::"(real^'n::finite) set" where "finite s" "{x - (\ i. d) .. x + (\ i. d)} = convex hull s" proof- - let ?d = "(\ i. d)::real^'n" - have *:"{x - ?d .. x + ?d} = (\y. x - ?d + (2 * d) *\<^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]] - by(auto simp add: vector_component) - hence "1 \ inverse d * (x $ i - y $ i)" "1 \ inverse d * (y $ i - x $ i)" - apply(rule_tac[!] mult_left_le_imp_le[OF _ assms]) unfolding mult_assoc[THEN sym] - using assms by(auto simp add: field_simps right_inverse) - hence "inverse d * (x $ i * 2) \ 2 + inverse d * (y $ i * 2)" - "inverse d * (y $ i * 2) \ 2 + inverse d * (x $ i * 2)" by(auto simp add:field_simps) } - hence "inverse (2 * d) *\<^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) *\<^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) *\<^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) *\<^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: - fixes s :: "(real ^ _) set" - assumes "open s" "convex_on s f" "\x\s. abs(f x) \ b" - shows "continuous_on s f" - apply(rule continuous_at_imp_continuous_on) unfolding continuous_at_real_range proof(rule,rule,rule) - fix x e assume "x\s" "(0::real) < e" - def B \ "abs b + 1" - have B:"0 < B" "\x. x\s \ abs (f x) \ B" - unfolding B_def defer apply(drule assms(3)[rule_format]) by auto - obtain k where "k>0"and k:"cball x k \ s" using assms(1)[unfolded open_contains_cball, THEN bspec[where x=x]] using `x\s` by auto - show "\d>0. \x'. norm (x' - x) < d \ \f x' - f x\ < e" - apply(rule_tac x="min (k / 2) (e / (2 * B) * k)" in exI) apply rule defer proof(rule,rule) - fix y assume as:"norm (y - x) < min (k / 2) (e / (2 * B) * k)" - show "\f y - f x\ < e" proof(cases "y=x") - case False def t \ "k / norm (y - x)" - have "2 < t" "00` by(auto simp add:field_simps) - have "y\s" apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_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 *\<^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 - 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) - hence th1:"f y - f x < e" apply- apply(rule le_less_trans) defer apply assumption - using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w] - using `0s` `w\s` by(auto simp add:field_simps) } - moreover - { def w \ "x - t *\<^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 - 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" - using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w] - using `0s` `w\s` by (auto simp add:field_simps) - also have "\ = (f w + t * f y) / (1 + t)" using `t>0` unfolding real_divide_def by (auto simp add:field_simps) - also have "\ < e + f y" using `t>0` * `e>0` by(auto simp add:field_simps) - finally have "f x - f y < e" by auto } - ultimately show ?thesis by auto - qed(insert `0y \ cball x e. f y \ b" - shows "\y \ cball x e. abs(f y) \ b + 2 * abs(f x)" - apply(rule) proof(cases "0 \ e") case True - fix y assume y:"y\cball x e" def z \ "2 *\<^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) *\<^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" - hence "dist x y < 0" using False unfolding mem_cball not_le by (auto simp del: dist_not_less_zero) - thus "\f y\ \ b + 2 * \f x\" using zero_le_dist[of x y] by auto qed - -subsection {* Hence a convex function on an open set is continuous. *} - -lemma convex_on_continuous: - assumes "open (s::(real^'n::finite) set)" "convex_on s f" - shows "continuous_on s f" - unfolding continuous_on_eq_continuous_at[OF assms(1)] proof - note dimge1 = dimindex_ge_1[where 'a='n] - fix x assume "x\s" - then obtain e where e:"cball x e \ s" "e>0" using assms(1) unfolding open_contains_cball by auto - def d \ "e / real CARD('n)" - have "0 < d" unfolding d_def using `e>0` dimge1 by(rule_tac divide_pos_pos, auto) - let ?d = "(\ i. d)::real^'n" - obtain c where c:"finite c" "{x - ?d..x + ?d} = convex hull c" using cube_convex_hull[OF `d>0`, of x] by auto - have "x\{x - ?d..x + ?d}" using `d>0` unfolding mem_interval by(auto simp add:vector_component_simps) - hence "c\{}" apply(rule_tac ccontr) using c by(auto simp add:convex_hull_empty) - def k \ "Max (f ` c)" - have "convex_on {x - ?d..x + ?d} f" apply(rule convex_on_subset[OF assms(2)]) - apply(rule subset_trans[OF _ e(1)]) unfolding subset_eq mem_cball proof - fix z assume z:"z\{x - ?d..x + ?d}" - have e:"e = setsum (\i. d) (UNIV::'n set)" unfolding setsum_constant d_def using dimge1 - by (metis card_enum field_simps d_def not_one_le_zero of_nat_le_iff real_eq_of_nat real_of_nat_1) - show "dist x z \ e" unfolding dist_norm e apply(rule_tac order_trans[OF norm_le_l1], rule setsum_mono) - using z[unfolded mem_interval] apply(erule_tac x=i in allE) by(auto simp add:field_simps vector_component_simps) qed - hence k:"\y\{x - ?d..x + ?d}. f y \ k" unfolding c(2) apply(rule_tac convex_on_convex_hull_bound) apply assumption - unfolding k_def apply(rule, rule Max_ge) using c(1) by auto - have "d \ e" unfolding d_def apply(rule mult_imp_div_pos_le) using `e>0` dimge1 unfolding mult_le_cancel_left1 using real_dimindex_ge_1 by auto - hence dsube:"cball x d \ cball x e" unfolding subset_eq Ball_def mem_cball by auto - have conv:"convex_on (cball x d) f" apply(rule convex_on_subset, rule convex_on_subset[OF assms(2)]) apply(rule e(1)) using dsube by auto - hence "\y\cball x d. abs (f y) \ k + 2 * abs (f x)" apply(rule_tac convex_bounds_lemma) apply assumption proof - fix y assume y:"y\cball x d" - { fix i::'n have "x $ i - d \ y $ i" "y $ i \ x $ i + d" - using order_trans[OF component_le_norm y[unfolded mem_cball dist_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) 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) 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 :: "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)" - -lemmas segment = open_segment_def closed_segment_def - -definition "starlike s \ (\a\s. \x\s. closed_segment a x \ s)" - -lemma midpoint_refl: "midpoint x x = x" - unfolding midpoint_def unfolding 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) - "dist b (midpoint a b) = (dist a b) / 2" (is ?t2) - "dist (midpoint a b) a = (dist a b) / 2" (is ?t3) - "dist (midpoint a b) b = (dist a b) / 2" (is ?t4) -proof- - have *: "\x y::real^'n::finite. 2 *\<^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) - show ?t4 unfolding midpoint_def dist_norm apply (rule **) by(auto,vector) qed - -lemma midpoint_eq_endpoint: - "midpoint a b = a \ a = (b::real^'n::finite)" - "midpoint a b = b \ a = b" - unfolding dist_eq_0_iff[where 'a="real^'n", THEN sym] dist_midpoint by auto - -lemma convex_contains_segment: - "convex s \ (\a\s. \b\s. closed_segment a b \ s)" - unfolding convex_alt closed_segment_def by auto - -lemma convex_imp_starlike: - "convex s \ s \ {} \ starlike s" - unfolding convex_contains_segment starlike_def by auto - -lemma segment_convex_hull: - "closed_segment a b = convex hull {a,b}" proof- - have *:"\x. {x} \ {}" by auto - have **:"\u v. u + v = 1 \ u = 1 - (v::real)" by auto - show ?thesis unfolding segment convex_hull_insert[OF *] convex_hull_singleton apply(rule set_ext) - unfolding mem_Collect_eq apply(rule,erule exE) - apply(rule_tac x="1 - u" in exI) apply rule defer apply(rule_tac x=u in exI) defer - apply(erule exE, (erule conjE)?)+ apply(rule_tac x="1 - u" in exI) unfolding ** by auto qed - -lemma convex_segment: "convex (closed_segment a b)" - unfolding segment_convex_hull by(rule convex_convex_hull) - -lemma ends_in_segment: "a \ closed_segment a b" "b \ closed_segment a b" - unfolding segment_convex_hull apply(rule_tac[!] hull_subset[unfolded subset_eq, rule_format]) by auto - -lemma segment_furthest_le: - assumes "x \ closed_segment a b" shows "norm(y - x) \ norm(y - a) \ norm(y - x) \ norm(y - b)" proof- - obtain z where "z\{a, b}" "norm (x - y) \ norm (z - y)" using simplex_furthest_le[of "{a, b}" y] - using assms[unfolded segment_convex_hull] by auto - thus ?thesis by(auto simp add:norm_minus_commute) qed - -lemma segment_bound: - assumes "x \ closed_segment a b" - shows "norm(x - a) \ norm(b - a)" "norm(x - b) \ norm(b - a)" - using segment_furthest_le[OF assms, of a] - using segment_furthest_le[OF assms, of b] - by (auto simp add:norm_minus_commute) - -lemma segment_refl:"closed_segment a a = {a}" unfolding segment by (auto simp add: algebra_simps) - -lemma between_mem_segment: "between (a,b) x \ x \ closed_segment a b" - unfolding between_def mem_def by auto - -lemma between:"between (a,b) (x::real^'n::finite) \ dist a b = (dist a x) + (dist x b)" -proof(cases "a = b") - case True thus ?thesis unfolding between_def split_conv mem_def[of x, symmetric] - by(auto simp add:segment_refl dist_commute) next - case False hence Fal:"norm (a - b) \ 0" and Fal2: "norm (a - b) > 0" 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) *\<^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) *\<^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)) *\<^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)) *\<^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) *\<^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 - -lemma between_mem_convex_hull: - "between (a,b) x \ x \ convex hull {a,b}" - unfolding between_mem_segment segment_convex_hull .. - -subsection {* Shrinking towards the interior of a convex set. *} - -lemma mem_interior_convex_shrink: - fixes s :: "(real ^ _) set" - assumes "convex s" "c \ interior s" "x \ s" "0 < e" "e \ 1" - 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 *\<^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 *\<^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]) - apply(rule d[unfolded subset_eq,rule_format]) unfolding mem_ball using assms(3-5) by auto - qed(rule mult_pos_pos, insert `e>0` `d>0`, auto) qed - -lemma mem_interior_closure_convex_shrink: - fixes s :: "(real ^ _) set" - assumes "convex s" "c \ interior s" "x \ closure s" "0 < e" "e \ 1" - 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 - case False hence x:"x islimpt s" using assms(3)[unfolded closure_def] by auto - show ?thesis proof(cases "e=1") - case True obtain y where "y\s" "y \ x" "dist y x < 1" - using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto - thus ?thesis apply(rule_tac x=y in bexI) unfolding True using `d>0` by auto next - case False hence "0 < e * d / (1 - e)" and *:"1 - e > 0" - using `e\1` `e>0` `d>0` by(auto intro!:mult_pos_pos divide_pos_pos) - then obtain y where "y\s" "y \ x" "dist y x < e * d / (1 - e)" - using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto - thus ?thesis apply(rule_tac x=y in bexI) unfolding dist_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) *\<^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 add:field_simps norm_minus_commute) - thus ?thesis unfolding * apply - apply(rule mem_interior_convex_shrink) - using assms(1,4-5) `y\s` by auto qed - -subsection {* Some obvious but surprisingly hard simplex lemmas. *} - -lemma simplex: - assumes "finite s" "0 \ s" - shows "convex hull (insert 0 s) = { y. (\u. (\x\s. 0 \ u x) \ setsum u s \ 1 \ setsum (\x. u x *\<^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) - unfolding if_smult and setsum_delta_notmem[OF assms(2)] by auto - -lemma std_simplex: - "convex hull (insert 0 { basis i | i. i\UNIV}) = - {x::real^'n::finite . (\i. 0 \ x$i) \ setsum (\i. x$i) UNIV \ 1 }" (is "convex hull (insert 0 ?p) = ?s") -proof- let ?D = "UNIV::'n set" - have "0\?p" by(auto simp add: basis_nonzero) - have "{(basis i)::real^'n |i. i \ ?D} = basis ` ?D" by auto - note sumbas = this setsum_reindex[OF basis_inj, unfolded o_def] - show ?thesis unfolding simplex[OF finite_stdbasis `0\?p`] apply(rule set_ext) unfolding mem_Collect_eq apply rule - apply(erule exE, (erule conjE)+) apply(erule_tac[2] conjE)+ proof- - fix x::"real^'n" and u assume as: "\x\{basis i |i. i \?D}. 0 \ u x" "setsum u {basis i |i. i \ ?D} \ 1" "(\x\{basis i |i. i \?D}. u x *\<^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 *\<^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})) = - {x::real^'n::finite. (\i. 0 < x$i) \ setsum (\i. x$i) UNIV < 1 }" - apply(rule set_ext) unfolding mem_interior std_simplex unfolding subset_eq mem_Collect_eq Ball_def mem_ball - unfolding Ball_def[symmetric] apply rule apply(erule exE, (erule conjE)+) defer apply(erule conjE) proof- - fix x::"real^'n" and e assume "0xa. dist x xa < e \ (\x. 0 \ xa $ x) \ setsum (op $ xa) UNIV \ 1" - show "(\xa. 0 < x $ xa) \ setsum (op $ x) UNIV < 1" apply(rule,rule) proof- - fix i::'n show "0 < x $ i" using as[THEN spec[where x="x - (e / 2) *\<^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) *\<^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) *\<^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 -next - fix x::"real^'n::finite" assume as:"\i. 0 < x $ i" "setsum (op $ x) UNIV < 1" - guess a using UNIV_witness[where 'a='b] .. - let ?d = "(1 - setsum (op $ x) UNIV) / real (CARD('n))" - have "Min ((op $ x) ` UNIV) > 0" apply(rule Min_grI) using as(1) dimindex_ge_1 by auto - moreover have"?d > 0" apply(rule divide_pos_pos) using as(2) using dimindex_ge_1 by(auto simp add: Suc_le_eq) - ultimately show "\e>0. \y. dist x y < e \ (\i. 0 \ y $ i) \ setsum (op $ y) UNIV \ 1" - apply(rule_tac x="min (Min ((op $ x) ` UNIV)) ?D" in exI) apply rule defer apply(rule,rule) proof- - fix y assume y:"dist x y < min (Min (op $ x ` UNIV)) ?d" - have "setsum (op $ y) UNIV \ setsum (\i. x$i + ?d) UNIV" proof(rule setsum_mono) - fix i::'n have "abs (y$i - x$i) < ?d" apply(rule le_less_trans) using component_le_norm[of "y - x" i] - using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] by(auto simp add:vector_component_simps norm_minus_commute) - thus "y $ i \ x $ i + ?d" by auto qed - also have "\ \ 1" unfolding setsum_addf setsum_constant card_enum real_eq_of_nat using dimindex_ge_1 by(auto simp add: Suc_le_eq) - finally show "(\i. 0 \ y $ i) \ setsum (op $ y) UNIV \ 1" apply- proof(rule,rule) - fix i::'n have "norm (x - y) < x$i" using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1] - using Min_gr_iff[of "op $ x ` dimset x"] dimindex_ge_1 by auto - thus "0 \ y$i" using component_le_norm[of "x - y" i] and as(1)[rule_format, of i] by(auto simp add: vector_component_simps) - qed auto qed auto qed - -lemma interior_std_simplex_nonempty: obtains a::"real^'n::finite" where - "a \ interior(convex hull (insert 0 {basis i | i . i \ UNIV}))" proof- - let ?D = "UNIV::'n set" let ?a = "setsum (\b::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) - unfolding setsum_delta'[OF finite_UNIV[where 'a='n]] and real_dimindex_ge_1[where 'n='n] by(auto simp add: basis_component[of i]) } - note ** = this - show ?thesis apply(rule that[of ?a]) unfolding interior_std_simplex mem_Collect_eq proof(rule,rule) - fix i::'n show "0 < ?a $ i" unfolding ** using dimindex_ge_1 by(auto simp add: Suc_le_eq) next - have "setsum (op $ ?a) ?D = setsum (\i. inverse (2 * real CARD('n))) ?D" by(rule setsum_cong2, rule **) - also have "\ < 1" unfolding setsum_constant card_enum real_eq_of_nat real_divide_def[THEN sym] by (auto simp add:field_simps) - finally show "setsum (op $ ?a) ?D < 1" by auto qed qed - -subsection {* Paths. *} - -definition "path (g::real^1 \ real^'n::finite) \ continuous_on {0 .. 1} g" - -definition "pathstart (g::real^1 \ real^'n) = g 0" - -definition "pathfinish (g::real^1 \ real^'n) = g 1" - -definition "path_image (g::real^1 \ real^'n) = g ` {0 .. 1}" - -definition "reversepath (g::real^1 \ real^'n) = (\x. g(1 - x))" - -definition joinpaths:: "(real^1 \ real^'n) \ (real^1 \ real^'n) \ (real^1 \ real^'n)" (infixr "+++" 75) - where "joinpaths g1 g2 = (\x. if dest_vec1 x \ ((1 / 2)::real) then g1 (2 *\<^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)" - -definition "injective_path (g::real^1 \ real^'n) \ - (\x\{0..1}. \y\{0..1}. g x = g y \ x = y)" - -subsection {* Some lemmas about these concepts. *} - -lemma injective_imp_simple_path: - "injective_path g \ simple_path g" - unfolding injective_path_def simple_path_def by auto - -lemma path_image_nonempty: "path_image g \ {}" - unfolding path_image_def image_is_empty interval_eq_empty by auto - -lemma pathstart_in_path_image[intro]: "(pathstart g) \ path_image g" - unfolding pathstart_def path_image_def apply(rule imageI) - unfolding mem_interval_1 vec_1[THEN sym] dest_vec1_vec by auto - -lemma pathfinish_in_path_image[intro]: "(pathfinish g) \ path_image g" - unfolding pathfinish_def path_image_def apply(rule imageI) - unfolding mem_interval_1 vec_1[THEN sym] dest_vec1_vec by auto - -lemma connected_path_image[intro]: "path g \ connected(path_image g)" - unfolding path_def path_image_def apply(rule connected_continuous_image, assumption) - by(rule convex_connected, rule convex_interval) - -lemma compact_path_image[intro]: "path g \ compact(path_image g)" - unfolding path_def path_image_def apply(rule compact_continuous_image, assumption) - by(rule compact_interval) - -lemma reversepath_reversepath[simp]: "reversepath(reversepath g) = g" - unfolding reversepath_def by auto - -lemma pathstart_reversepath[simp]: "pathstart(reversepath g) = pathfinish g" - unfolding pathstart_def reversepath_def pathfinish_def by auto - -lemma pathfinish_reversepath[simp]: "pathfinish(reversepath g) = pathstart g" - unfolding pathstart_def reversepath_def pathfinish_def by auto - -lemma pathstart_join[simp]: "pathstart(g1 +++ g2) = pathstart g1" - unfolding pathstart_def joinpaths_def pathfinish_def by auto - -lemma pathfinish_join[simp]:"pathfinish(g1 +++ g2) = pathfinish g2" proof- - have "2 *\<^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 - -lemma path_image_reversepath[simp]: "path_image(reversepath g) = path_image g" proof- - have *:"\g. path_image(reversepath g) \ path_image g" - unfolding path_image_def subset_eq reversepath_def Ball_def image_iff apply(rule,rule,erule bexE) - apply(rule_tac x="1 - xa" in bexI) by(auto simp add:vector_less_eq_def vector_component_simps elim!:ballE) - show ?thesis using *[of g] *[of "reversepath g"] unfolding reversepath_reversepath by auto qed - -lemma path_reversepath[simp]: "path(reversepath g) \ path g" proof- - have *:"\g. path g \ path(reversepath g)" unfolding path_def reversepath_def - apply(rule continuous_on_compose[unfolded o_def, of _ "\x. 1 - x"]) - apply(rule continuous_on_sub, rule continuous_on_const, rule continuous_on_id) - apply(rule continuous_on_subset[of "{0..1}"], assumption) - by (auto, auto simp add:vector_less_eq_def vector_component_simps elim!:ballE) - show ?thesis using *[of g] *[of "reversepath g"] unfolding reversepath_reversepath by auto qed - -lemmas reversepath_simps = path_reversepath path_image_reversepath pathstart_reversepath pathfinish_reversepath - -lemma path_join[simp]: assumes "pathfinish g1 = pathstart g2" shows "path (g1 +++ g2) \ path g1 \ path g2" - unfolding path_def pathfinish_def pathstart_def apply rule defer apply(erule conjE) proof- - assume as:"continuous_on {0..1} (g1 +++ g2)" - have *:"g1 = (\x. g1 (2 *\<^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) - apply (rule continuous_on_cmul, rule continuous_on_add, rule continuous_on_id, rule continuous_on_const) defer - apply (rule continuous_on_cmul, rule continuous_on_id) apply(rule_tac[!] continuous_on_eq[of _ "g1 +++ g2"]) defer prefer 3 - apply(rule_tac[1-2] continuous_on_subset[of "{0 .. 1}"]) apply(rule as, assumption, rule as, assumption) - apply(rule) defer apply rule proof- - fix x assume "x \ op *\<^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 *\<^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 *\<^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)*\<^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) *\<^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) *\<^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)*\<^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 *\<^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) - by(auto intro!: imageI simp add: vector_component_simps) qed - -lemma subset_path_image_join: - assumes "path_image g1 \ s" "path_image g2 \ s" shows "path_image(g1 +++ g2) \ s" - using path_image_join_subset[of g1 g2] and assms by auto - -lemma path_image_join: - assumes "path g1" "path g2" "pathfinish g1 = pathstart g2" - shows "path_image(g1 +++ g2) = (path_image g1) \ (path_image g2)" -apply(rule, rule path_image_join_subset, rule) unfolding Un_iff proof(erule disjE) - fix x assume "x \ path_image g1" - then obtain y where y:"y\{0..1}" "x = g1 y" unfolding path_image_def image_iff by auto - thus "x \ path_image (g1 +++ g2)" unfolding joinpaths_def path_image_def image_iff - apply(rule_tac x="(1/2) *\<^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) *\<^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: - assumes "x \ path_image g1" "x \ path_image g2" shows "x \ path_image(g1 +++ g2)" - using assms and path_image_join_subset[of g1 g2] by auto - -lemma simple_path_reversepath: assumes "simple_path g" shows "simple_path (reversepath g)" - using assms unfolding simple_path_def reversepath_def apply- apply(rule ballI)+ - apply(erule_tac x="1-x" in ballE, erule_tac x="1-y" in ballE) - unfolding mem_interval_1 by(auto simp add:vector_component_simps) - -lemma 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}" - shows "simple_path(g1 +++ g2)" -unfolding simple_path_def proof((rule ballI)+, rule impI) let ?g = "g1 +++ g2" - note inj = assms(1,2)[unfolded injective_path_def, rule_format] - fix x y::"real^1" assume xy:"x \ {0..1}" "y \ {0..1}" "?g x = ?g y" - show "x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" proof(case_tac "x$1 \ 1/2",case_tac[!] "y$1 \ 1/2", unfold not_le) - assume as:"x $ 1 \ 1 / 2" "y $ 1 \ 1 / 2" - hence "g1 (2 *\<^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*\<^sub>R x" "2*\<^sub>R y"] by auto - next assume as:"x $ 1 > 1 / 2" "y $ 1 > 1 / 2" - 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*\<^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 *\<^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 *\<^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 *\<^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 *\<^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 *\<^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 *\<^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: - assumes "injective_path g1" "injective_path g2" "pathfinish g1 = pathstart g2" - "(path_image g1 \ path_image g2) \ {pathstart g2}" - shows "injective_path(g1 +++ g2)" - unfolding injective_path_def proof(rule,rule,rule) let ?g = "g1 +++ g2" - note inj = assms(1,2)[unfolded injective_path_def, rule_format] - fix x y assume xy:"x \ {0..1}" "y \ {0..1}" "(g1 +++ g2) x = (g1 +++ g2) y" - show "x = y" proof(cases "x$1 \ 1/2", case_tac[!] "y$1 \ 1/2", unfold not_le) - assume "x $ 1 \ 1 / 2" "y $ 1 \ 1 / 2" thus ?thesis using inj(1)[of "2*\<^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*\<^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 *\<^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 *\<^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 - -lemmas join_paths_simps = path_join path_image_join pathstart_join pathfinish_join - -subsection {* Reparametrizing a closed curve to start at some chosen point. *} - -definition "shiftpath a (f::real^1 \ real^'n) = - (\x. if dest_vec1 (a + x) \ 1 then f(a + x) else f(a + x - 1))" - -lemma pathstart_shiftpath: "a \ 1 \ pathstart(shiftpath a g) = g a" - unfolding pathstart_def shiftpath_def by auto - -(** move this **) -declare forall_1[simp] ex_1[simp] - -lemma pathfinish_shiftpath: assumes "0 \ a" "pathfinish g = pathstart g" - shows "pathfinish(shiftpath a g) = g a" - using assms unfolding pathstart_def pathfinish_def shiftpath_def - by(auto simp add: vector_component_simps) - -lemma endpoints_shiftpath: - assumes "pathfinish g = pathstart g" "a \ {0 .. 1}" - shows "pathfinish(shiftpath a g) = g a" "pathstart(shiftpath a g) = g a" - using assms by(auto intro!:pathfinish_shiftpath pathstart_shiftpath) - -lemma closed_shiftpath: - assumes "pathfinish g = pathstart g" "a \ {0..1}" - shows "pathfinish(shiftpath a g) = pathstart(shiftpath a g)" - using endpoints_shiftpath[OF assms] by auto - -lemma path_shiftpath: - assumes "path g" "pathfinish g = pathstart g" "a \ {0..1}" - shows "path(shiftpath a g)" proof- - have *:"{0 .. 1} = {0 .. 1-a} \ {1-a .. 1}" using assms(3) by(auto simp add: vector_component_simps) - have **:"\x. x + a = 1 \ g (x + a - 1) = g (x + a)" - using assms(2)[unfolded pathfinish_def pathstart_def] by auto - show ?thesis unfolding path_def shiftpath_def * apply(rule continuous_on_union) - apply(rule closed_interval)+ apply(rule continuous_on_eq[of _ "g \ (\x. a + x)"]) prefer 3 - apply(rule continuous_on_eq[of _ "g \ (\x. a - 1 + x)"]) defer prefer 3 - apply(rule continuous_on_intros)+ prefer 2 apply(rule continuous_on_intros)+ - apply(rule_tac[1-2] continuous_on_subset[OF assms(1)[unfolded path_def]]) - using assms(3) and ** by(auto simp add:vector_component_simps field_simps Cart_eq) qed - -lemma shiftpath_shiftpath: assumes "pathfinish g = pathstart g" "a \ {0..1}" "x \ {0..1}" - shows "shiftpath (1 - a) (shiftpath a g) x = g x" - using assms unfolding pathfinish_def pathstart_def shiftpath_def - by(auto simp add: vector_component_simps) - -lemma path_image_shiftpath: - assumes "a \ {0..1}" "pathfinish g = pathstart g" - shows "path_image(shiftpath a g) = path_image g" proof- - { fix x assume as:"g 1 = g 0" "x \ {0..1::real^1}" " \y\{0..1} \ {x. \ a $ 1 + x $ 1 \ 1}. g x \ g (a + y - 1)" - hence "\y\{0..1} \ {x. a $ 1 + x $ 1 \ 1}. g x = g (a + y)" proof(cases "a \ x") - case False thus ?thesis apply(rule_tac x="1 + x - a" in bexI) - using as(1,2) and as(3)[THEN bspec[where x="1 + x - a"]] and assms(1) - by(auto simp add:vector_component_simps field_simps atomize_not) next - case True thus ?thesis using as(1-2) and assms(1) apply(rule_tac x="x - a" in bexI) - by(auto simp add:vector_component_simps field_simps) qed } - thus ?thesis using assms unfolding shiftpath_def path_image_def pathfinish_def pathstart_def - by(auto simp add:vector_component_simps image_iff) qed - -subsection {* Special case of straight-line paths. *} - -definition - linepath :: "real ^ 'n::finite \ real ^ 'n \ real ^ 1 \ real ^ 'n" where - "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 - -lemma pathfinish_linepath[simp]: "pathfinish(linepath a b) = b" - unfolding pathfinish_def linepath_def by auto - -lemma continuous_linepath_at[intro]: "continuous (at x) (linepath a b)" - unfolding linepath_def - by (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) - -lemma path_linepath[intro]: "path(linepath a b)" - unfolding path_def by(rule continuous_on_linepath) - -lemma path_image_linepath[simp]: "path_image(linepath a b) = (closed_segment a b)" - unfolding path_image_def segment linepath_def apply (rule set_ext, rule) defer - unfolding mem_Collect_eq image_iff apply(erule exE) apply(rule_tac x="u *\<^sub>R 1" in bexI) - by(auto simp add:vector_component_simps) - -lemma reversepath_linepath[simp]: "reversepath(linepath a b) = linepath b a" - unfolding reversepath_def linepath_def by(rule ext, auto simp add:vector_component_simps) - -lemma injective_path_linepath: assumes "a \ b" shows "injective_path(linepath a b)" proof- - { obtain i where i:"a$i \ b$i" using assms[unfolded Cart_eq] by auto - fix x y::"real^1" assume "x $ 1 *\<^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 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) - -subsection {* Bounding a point away from a path. *} - -lemma not_on_path_ball: assumes "path g" "z \ path_image g" - shows "\e>0. ball z e \ (path_image g) = {}" proof- - obtain a where "a\path_image g" "\y\path_image g. dist z a \ dist z y" - using distance_attains_inf[OF _ path_image_nonempty, of g z] - using compact_path_image[THEN compact_imp_closed, OF assms(1)] by auto - thus ?thesis apply(rule_tac x="dist z a" in exI) using assms(2) by(auto intro!: dist_pos_lt) qed - -lemma not_on_path_cball: assumes "path g" "z \ path_image g" - shows "\e>0. cball z e \ (path_image g) = {}" proof- - obtain e where "ball z e \ path_image g = {}" "e>0" using not_on_path_ball[OF assms] by auto - moreover have "cball z (e/2) \ ball z e" using `e>0` by auto - ultimately show ?thesis apply(rule_tac x="e/2" in exI) by auto qed - -subsection {* Path component, considered as a "joinability" relation (from Tom Hales). *} - -definition "path_component s x y \ (\g. path g \ path_image g \ s \ pathstart g = x \ pathfinish g = y)" - -lemmas path_defs = path_def pathstart_def pathfinish_def path_image_def path_component_def - -lemma path_component_mem: assumes "path_component s x y" shows "x \ s" "y \ s" - using assms unfolding path_defs by auto - -lemma path_component_refl: assumes "x \ s" shows "path_component s x x" - unfolding path_defs apply(rule_tac x="\u. x" in exI) using assms - by(auto intro!:continuous_on_intros) - -lemma path_component_refl_eq: "path_component s x x \ x \ s" - by(auto intro!: path_component_mem path_component_refl) - -lemma path_component_sym: "path_component s x y \ path_component s y x" - using assms unfolding path_component_def apply(erule exE) apply(rule_tac x="reversepath g" in exI) - by(auto simp add: reversepath_simps) - -lemma path_component_trans: assumes "path_component s x y" "path_component s y z" shows "path_component s x z" - using assms unfolding path_component_def apply- apply(erule exE)+ apply(rule_tac x="g +++ ga" in exI) by(auto simp add: path_image_join) - -lemma path_component_of_subset: "s \ t \ path_component s x y \ path_component t x y" - unfolding path_component_def by auto - -subsection {* Can also consider it as a set, as the name suggests. *} - -lemma path_component_set: "path_component s x = { y. (\g. path g \ path_image g \ s \ pathstart g = x \ pathfinish g = y )}" - apply(rule set_ext) unfolding mem_Collect_eq unfolding mem_def path_component_def by auto - -lemma mem_path_component_set:"x \ path_component s y \ path_component s y x" unfolding mem_def by auto - -lemma path_component_subset: "(path_component s x) \ s" - apply(rule, rule path_component_mem(2)) by(auto simp add:mem_def) - -lemma path_component_eq_empty: "path_component s x = {} \ x \ s" - apply rule apply(drule equals0D[of _ x]) defer apply(rule equals0I) unfolding mem_path_component_set - apply(drule path_component_mem(1)) using path_component_refl by auto - -subsection {* Path connectedness of a space. *} - -definition "path_connected s \ (\x\s. \y\s. \g. path g \ (path_image g) \ s \ pathstart g = x \ pathfinish g = y)" - -lemma path_connected_component: "path_connected s \ (\x\s. \y\s. path_component s x y)" - unfolding path_connected_def path_component_def by auto - -lemma path_connected_component_set: "path_connected s \ (\x\s. path_component s x = s)" - unfolding path_connected_component apply(rule, rule, rule, rule path_component_subset) - unfolding subset_eq mem_path_component_set Ball_def mem_def by auto - -subsection {* Some useful lemmas about path-connectedness. *} - -lemma convex_imp_path_connected: assumes "convex s" shows "path_connected s" - unfolding path_connected_def apply(rule,rule,rule_tac x="linepath x y" in exI) - unfolding path_image_linepath using assms[unfolded convex_contains_segment] by auto - -lemma path_connected_imp_connected: assumes "path_connected s" shows "connected s" - unfolding connected_def not_ex apply(rule,rule,rule ccontr) unfolding not_not apply(erule conjE)+ proof- - fix e1 e2 assume as:"open e1" "open e2" "s \ e1 \ e2" "e1 \ e2 \ s = {}" "e1 \ s \ {}" "e2 \ s \ {}" - then obtain x1 x2 where obt:"x1\e1\s" "x2\e2\s" by auto - then obtain g where g:"path g" "path_image g \ s" "pathstart g = x1" "pathfinish g = x2" - using assms[unfolded path_connected_def,rule_format,of x1 x2] by auto - have *:"connected {0..1::real^1}" by(auto intro!: convex_connected convex_interval) - have "{0..1} \ {x \ {0..1}. g x \ e1} \ {x \ {0..1}. g x \ e2}" using as(3) g(2)[unfolded path_defs] by blast - moreover have "{x \ {0..1}. g x \ e1} \ {x \ {0..1}. g x \ e2} = {}" using as(4) g(2)[unfolded path_defs] unfolding subset_eq by auto - moreover have "{x \ {0..1}. g x \ e1} \ {} \ {x \ {0..1}. g x \ e2} \ {}" using g(3,4)[unfolded path_defs] using obt by(auto intro!: exI) - ultimately show False using *[unfolded connected_local not_ex,rule_format, of "{x\{0..1}. g x \ e1}" "{x\{0..1}. g x \ e2}"] - using continuous_open_in_preimage[OF g(1)[unfolded path_def] as(1)] - using continuous_open_in_preimage[OF g(1)[unfolded path_def] as(2)] by auto qed - -lemma open_path_component: assumes "open s" shows "open(path_component s x)" - unfolding open_contains_ball proof - fix y assume as:"y \ path_component s x" - hence "y\s" apply- apply(rule path_component_mem(2)) unfolding mem_def by auto - then obtain e where e:"e>0" "ball y e \ s" using assms[unfolded open_contains_ball] by auto - show "\e>0. ball y e \ path_component s x" apply(rule_tac x=e in exI) apply(rule,rule `e>0`,rule) unfolding mem_ball mem_path_component_set proof- - fix z assume "dist y z < e" thus "path_component s x z" apply(rule_tac path_component_trans[of _ _ y]) defer - apply(rule path_component_of_subset[OF e(2)]) apply(rule convex_imp_path_connected[OF convex_ball, unfolded path_connected_component, rule_format]) using `e>0` - using as[unfolded mem_def] by auto qed qed - -lemma open_non_path_component: assumes "open s" shows "open(s - path_component s x)" unfolding open_contains_ball proof - fix y assume as:"y\s - path_component s x" - then obtain e where e:"e>0" "ball y e \ s" using assms[unfolded open_contains_ball] by auto - show "\e>0. ball y e \ s - path_component s x" apply(rule_tac x=e in exI) apply(rule,rule `e>0`,rule,rule) defer proof(rule ccontr) - fix z assume "z\ball y e" "\ z \ path_component s x" - hence "y \ path_component s x" unfolding not_not mem_path_component_set using `e>0` - apply- apply(rule path_component_trans,assumption) apply(rule path_component_of_subset[OF e(2)]) - apply(rule convex_imp_path_connected[OF convex_ball, unfolded path_connected_component, rule_format]) by auto - thus False using as by auto qed(insert e(2), auto) qed - -lemma connected_open_path_connected: assumes "open s" "connected s" shows "path_connected s" - unfolding path_connected_component_set proof(rule,rule,rule path_component_subset, rule) - fix x y assume "x \ s" "y \ s" show "y \ path_component s x" proof(rule ccontr) - assume "y \ path_component s x" moreover - have "path_component s x \ s \ {}" using `x\s` path_component_eq_empty path_component_subset[of s x] by auto - ultimately show False using `y\s` open_non_path_component[OF assms(1)] open_path_component[OF assms(1)] - using assms(2)[unfolded connected_def not_ex, rule_format, of"path_component s x" "s - path_component s x"] by auto -qed qed - -lemma path_connected_continuous_image: - assumes "continuous_on s f" "path_connected s" shows "path_connected (f ` s)" - unfolding path_connected_def proof(rule,rule) - fix x' y' assume "x' \ f ` s" "y' \ f ` s" - then obtain x y where xy:"x\s" "y\s" "x' = f x" "y' = f y" by auto - guess g using assms(2)[unfolded path_connected_def,rule_format,OF xy(1,2)] .. - thus "\g. path g \ path_image g \ f ` s \ pathstart g = x' \ pathfinish g = y'" - unfolding xy apply(rule_tac x="f \ g" in exI) unfolding path_defs - using assms(1) by(auto intro!: continuous_on_compose continuous_on_subset[of _ _ "g ` {0..1}"]) qed - -lemma homeomorphic_path_connectedness: - "s homeomorphic t \ (path_connected s \ path_connected t)" - unfolding homeomorphic_def homeomorphism_def apply(erule exE|erule conjE)+ apply rule - apply(drule_tac f=f in path_connected_continuous_image) prefer 3 - apply(drule_tac f=g in path_connected_continuous_image) by auto - -lemma path_connected_empty: "path_connected {}" - unfolding path_connected_def by auto - -lemma path_connected_singleton: "path_connected {a}" - unfolding path_connected_def apply(rule,rule) - apply(rule_tac x="linepath a a" in exI) by(auto simp add:segment 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) - fix x y assume as:"x \ s \ t" "y \ s \ t" - from assms(3) obtain z where "z \ s \ t" by auto - thus "path_component (s \ t) x y" using as using assms(1-2)[unfolded path_connected_component] apply- - apply(erule_tac[!] UnE)+ apply(rule_tac[2-3] path_component_trans[of _ _ z]) - by(auto simp add:path_component_of_subset [OF Un_upper1] path_component_of_subset[OF Un_upper2]) qed - -subsection {* sphere is path-connected. *} - -lemma path_connected_punctured_universe: - assumes "2 \ CARD('n::finite)" shows "path_connected((UNIV::(real^'n::finite) set) - {a})" proof- - obtain \ where \:"bij_betw \ {1..CARD('n)} (UNIV::'n set)" using ex_bij_betw_nat_finite_1[OF finite_UNIV] by auto - let ?U = "UNIV::(real^'n) set" let ?u = "?U - {0}" - let ?basis = "\k. basis (\ k)" - let ?A = "\k. {x::real^'n. \i\{1..k}. inner (basis (\ i)) x \ 0}" - have "\k\{2..CARD('n)}. path_connected (?A k)" proof - 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 < 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 - next case True hence d:"1\{1..CARD('n)}" "2\{1..CARD('n)}" using Suc(2) by auto - have ***:"Suc 1 = 2" by auto - have **:"\s t P Q. s \ t \ {x. P x \ Q x} = (s \ {x. P x}) \ (t \ {x. Q x})" by auto - have "\ 2 \ \ (Suc 0)" apply(rule ccontr) using \[unfolded bij_betw_def inj_on_def, THEN conjunct1, THEN bspec[where x=2]] using assms by auto - thus ?thesis unfolding * True unfolding ** neq_iff bex_disj_distrib apply - - apply(rule path_connected_Un, rule_tac[1-2] path_connected_Un) defer 3 apply(rule_tac[1-4] convex_imp_path_connected) - apply(rule_tac[5] x=" ?basis 1 + ?basis 2" in nequals0I) - apply(rule_tac[6] x="-?basis 1 + ?basis 2" in nequals0I) - apply(rule_tac[7] x="-?basis 1 - ?basis 2" in nequals0I) - using d unfolding *** by(auto intro!: convex_halfspace_gt convex_halfspace_lt, auto simp add:vector_component_simps inner_basis) - qed qed auto qed note lem = this - - 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:"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)}. 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. 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 - -lemma path_connected_sphere: assumes "2 \ CARD('n::finite)" shows "path_connected {x::real^'n::finite. norm(x - a) = r}" proof(cases "r\0") - case True thus ?thesis proof(cases "r=0") - case False hence "{x::real^'n. norm(x - a) = r} = {}" using True by auto - thus ?thesis using path_connected_empty by auto - qed(auto intro!:path_connected_singleton) next - case False hence *:"{x::real^'n. norm(x - a) = r} = (\x. a + r *\<^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 **:"{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 split:split_if_asm) - have "continuous_on (UNIV - {0}) (\x::real^'n. 1 / norm x)" unfolding o_def continuous_on_eq_continuous_within - apply(rule, rule continuous_at_within_inv[unfolded o_def inverse_eq_divide]) apply(rule continuous_at_within) - apply(rule continuous_at_norm[unfolded o_def]) by auto - thus ?thesis unfolding * ** using path_connected_punctured_universe[OF assms] - by(auto intro!: path_connected_continuous_image continuous_on_intros continuous_on_mul) qed - -lemma connected_sphere: "2 \ CARD('n) \ connected {x::real^'n::finite. norm(x - a) = r}" - using path_connected_sphere path_connected_imp_connected by auto - -(** In continuous_at_vec1_norm : Use \ instead of \. **) - -end diff -r 1fad3160d873 -r 2083bde13ce1 src/HOL/Library/Determinants.thy --- a/src/HOL/Library/Determinants.thy Fri Oct 23 14:33:07 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1087 +0,0 @@ -(* Title: Determinants - Author: Amine Chaieb, University of Cambridge -*) - -header {* Traces, Determinant of square matrices and some properties *} - -theory Determinants -imports Euclidean_Space Permutations -begin - -subsection{* First some facts about products*} -lemma setprod_insert_eq: "finite A \ setprod f (insert a A) = (if a \ A then setprod f A else f a * setprod f A)" -apply clarsimp -by(subgoal_tac "insert a A = A", auto) - -lemma setprod_add_split: - assumes mn: "(m::nat) <= n + 1" - shows "setprod f {m.. n+p} = setprod f {m .. n} * setprod f {n+1..n+p}" -proof- - let ?A = "{m .. n+p}" - let ?B = "{m .. n}" - let ?C = "{n+1..n+p}" - from mn have un: "?B \ ?C = ?A" by auto - from mn have dj: "?B \ ?C = {}" by auto - have f: "finite ?B" "finite ?C" by simp_all - from setprod_Un_disjoint[OF f dj, of f, unfolded un] show ?thesis . -qed - - -lemma setprod_offset: "setprod f {(m::nat) + p .. n + p} = setprod (\i. f (i + p)) {m..n}" -apply (rule setprod_reindex_cong[where f="op + p"]) -apply (auto simp add: image_iff Bex_def inj_on_def) -apply arith -apply (rule ext) -apply (simp add: add_commute) -done - -lemma setprod_singleton: "setprod f {x} = f x" by simp - -lemma setprod_singleton_nat_seg: "setprod f {n..n} = f (n::'a::order)" by simp - -lemma setprod_numseg: "setprod f {m..0} = (if m=0 then f 0 else 1)" - "setprod f {m .. Suc n} = (if m \ Suc n then f (Suc n) * setprod f {m..n} - else setprod f {m..n})" - by (auto simp add: atLeastAtMostSuc_conv) - -lemma setprod_le: assumes fS: "finite S" and fg: "\x\S. f x \ 0 \ f x \ (g x :: 'a::ordered_idom)" - shows "setprod f S \ setprod g S" -using fS fg -apply(induct S) -apply simp -apply auto -apply (rule mult_mono) -apply (auto intro: setprod_nonneg) -done - - (* FIXME: In Finite_Set there is a useless further assumption *) -lemma setprod_inversef: "finite A ==> setprod (inverse \ f) A = (inverse (setprod f A) :: 'a:: {division_by_zero, field})" - apply (erule finite_induct) - apply (simp) - apply simp - done - -lemma setprod_le_1: assumes fS: "finite S" and f: "\x\S. f x \ 0 \ f x \ (1::'a::ordered_idom)" - shows "setprod f S \ 1" -using setprod_le[OF fS f] unfolding setprod_1 . - -subsection{* Trace *} - -definition trace :: "'a::semiring_1^'n^'n \ 'a" where - "trace A = setsum (\i. ((A$i)$i)) (UNIV::'n set)" - -lemma trace_0: "trace(mat 0) = 0" - by (simp add: trace_def mat_def) - -lemma trace_I: "trace(mat 1 :: 'a::semiring_1^'n^'n) = of_nat(CARD('n))" - by (simp add: trace_def mat_def) - -lemma trace_add: "trace ((A::'a::comm_semiring_1^'n^'n) + B) = trace A + trace B" - by (simp add: trace_def setsum_addf) - -lemma trace_sub: "trace ((A::'a::comm_ring_1^'n^'n) - B) = trace A - trace B" - by (simp add: trace_def setsum_subtractf) - -lemma trace_mul_sym:"trace ((A::'a::comm_semiring_1^'n^'n) ** B) = trace (B**A)" - apply (simp add: trace_def matrix_matrix_mult_def) - apply (subst setsum_commute) - by (simp add: mult_commute) - -(* ------------------------------------------------------------------------- *) -(* Definition of determinant. *) -(* ------------------------------------------------------------------------- *) - -definition det:: "'a::comm_ring_1^'n^'n \ 'a" where - "det A = setsum (\p. of_int (sign p) * setprod (\i. A$i$p i) (UNIV :: 'n set)) {p. p permutes (UNIV :: 'n set)}" - -(* ------------------------------------------------------------------------- *) -(* A few general lemmas we need below. *) -(* ------------------------------------------------------------------------- *) - -lemma setprod_permute: - assumes p: "p permutes S" - shows "setprod f S = setprod (f o p) S" -proof- - {assume "\ finite S" hence ?thesis by simp} - moreover - {assume fS: "finite S" - then have ?thesis - apply (simp add: setprod_def cong del:strong_setprod_cong) - apply (rule ab_semigroup_mult.fold_image_permute) - apply (auto simp add: p) - apply unfold_locales - done} - ultimately show ?thesis by blast -qed - -lemma setproduct_permute_nat_interval: "p permutes {m::nat .. n} ==> setprod f {m..n} = setprod (f o p) {m..n}" - by (blast intro!: setprod_permute) - -(* ------------------------------------------------------------------------- *) -(* Basic determinant properties. *) -(* ------------------------------------------------------------------------- *) - -lemma det_transp: "det (transp A) = det (A::'a::comm_ring_1 ^'n^'n::finite)" -proof- - let ?di = "\A i j. A$i$j" - let ?U = "(UNIV :: 'n set)" - have fU: "finite ?U" by simp - {fix p assume p: "p \ {p. p permutes ?U}" - from p have pU: "p permutes ?U" by blast - have sth: "sign (inv p) = sign p" - by (metis sign_inverse fU p mem_def Collect_def permutation_permutes) - from permutes_inj[OF pU] - have pi: "inj_on p ?U" by (blast intro: subset_inj_on) - from permutes_image[OF pU] - have "setprod (\i. ?di (transp A) i (inv p i)) ?U = setprod (\i. ?di (transp A) i (inv p i)) (p ` ?U)" by simp - also have "\ = setprod ((\i. ?di (transp A) i (inv p i)) o p) ?U" - unfolding setprod_reindex[OF pi] .. - also have "\ = setprod (\i. ?di A i (p i)) ?U" - proof- - {fix i assume i: "i \ ?U" - from i permutes_inv_o[OF pU] permutes_in_image[OF pU] - have "((\i. ?di (transp A) i (inv p i)) o p) i = ?di A i (p i)" - unfolding transp_def by (simp add: expand_fun_eq)} - then show "setprod ((\i. ?di (transp A) i (inv p i)) o p) ?U = setprod (\i. ?di A i (p i)) ?U" by (auto intro: setprod_cong) - qed - finally have "of_int (sign (inv p)) * (setprod (\i. ?di (transp A) i (inv p i)) ?U) = of_int (sign p) * (setprod (\i. ?di A i (p i)) ?U)" using sth - by simp} - then show ?thesis unfolding det_def apply (subst setsum_permutations_inverse) - apply (rule setsum_cong2) by blast -qed - -lemma det_lowerdiagonal: - fixes A :: "'a::comm_ring_1^'n^'n::{finite,wellorder}" - assumes ld: "\i j. i < j \ A$i$j = 0" - shows "det A = setprod (\i. A$i$i) (UNIV:: 'n set)" -proof- - let ?U = "UNIV:: 'n set" - let ?PU = "{p. p permutes ?U}" - let ?pp = "\p. of_int (sign p) * setprod (\i. A$i$p i) (UNIV :: 'n set)" - have fU: "finite ?U" by simp - from finite_permutations[OF fU] have fPU: "finite ?PU" . - have id0: "{id} \ ?PU" by (auto simp add: permutes_id) - {fix p assume p: "p \ ?PU -{id}" - from p have pU: "p permutes ?U" and pid: "p \ id" by blast+ - from permutes_natset_le[OF pU] pid obtain i where - i: "p i > i" by (metis not_le) - from ld[OF i] have ex:"\i \ ?U. A$i$p i = 0" by blast - from setprod_zero[OF fU ex] have "?pp p = 0" by simp} - then have p0: "\p \ ?PU -{id}. ?pp p = 0" by blast - from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis - unfolding det_def by (simp add: sign_id) -qed - -lemma det_upperdiagonal: - fixes A :: "'a::comm_ring_1^'n^'n::{finite,wellorder}" - assumes ld: "\i j. i > j \ A$i$j = 0" - shows "det A = setprod (\i. A$i$i) (UNIV:: 'n set)" -proof- - let ?U = "UNIV:: 'n set" - let ?PU = "{p. p permutes ?U}" - let ?pp = "(\p. of_int (sign p) * setprod (\i. A$i$p i) (UNIV :: 'n set))" - have fU: "finite ?U" by simp - from finite_permutations[OF fU] have fPU: "finite ?PU" . - have id0: "{id} \ ?PU" by (auto simp add: permutes_id) - {fix p assume p: "p \ ?PU -{id}" - from p have pU: "p permutes ?U" and pid: "p \ id" by blast+ - from permutes_natset_ge[OF pU] pid obtain i where - i: "p i < i" by (metis not_le) - from ld[OF i] have ex:"\i \ ?U. A$i$p i = 0" by blast - from setprod_zero[OF fU ex] have "?pp p = 0" by simp} - then have p0: "\p \ ?PU -{id}. ?pp p = 0" by blast - from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis - unfolding det_def by (simp add: sign_id) -qed - -lemma det_diagonal: - fixes A :: "'a::comm_ring_1^'n^'n::finite" - assumes ld: "\i j. i \ j \ A$i$j = 0" - shows "det A = setprod (\i. A$i$i) (UNIV::'n set)" -proof- - let ?U = "UNIV:: 'n set" - let ?PU = "{p. p permutes ?U}" - let ?pp = "\p. of_int (sign p) * setprod (\i. A$i$p i) (UNIV :: 'n set)" - have fU: "finite ?U" by simp - from finite_permutations[OF fU] have fPU: "finite ?PU" . - have id0: "{id} \ ?PU" by (auto simp add: permutes_id) - {fix p assume p: "p \ ?PU - {id}" - then have "p \ id" by simp - then obtain i where i: "p i \ i" unfolding expand_fun_eq by auto - from ld [OF i [symmetric]] have ex:"\i \ ?U. A$i$p i = 0" by blast - from setprod_zero [OF fU ex] have "?pp p = 0" by simp} - then have p0: "\p \ ?PU - {id}. ?pp p = 0" by blast - from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis - unfolding det_def by (simp add: sign_id) -qed - -lemma det_I: "det (mat 1 :: 'a::comm_ring_1^'n^'n::finite) = 1" -proof- - let ?A = "mat 1 :: 'a::comm_ring_1^'n^'n" - let ?U = "UNIV :: 'n set" - let ?f = "\i j. ?A$i$j" - {fix i assume i: "i \ ?U" - have "?f i i = 1" using i by (vector mat_def)} - hence th: "setprod (\i. ?f i i) ?U = setprod (\x. 1) ?U" - by (auto intro: setprod_cong) - {fix i j assume i: "i \ ?U" and j: "j \ ?U" and ij: "i \ j" - have "?f i j = 0" using i j ij by (vector mat_def) } - then have "det ?A = setprod (\i. ?f i i) ?U" using det_diagonal - by blast - also have "\ = 1" unfolding th setprod_1 .. - finally show ?thesis . -qed - -lemma det_0: "det (mat 0 :: 'a::comm_ring_1^'n^'n::finite) = 0" - by (simp add: det_def setprod_zero) - -lemma det_permute_rows: - fixes A :: "'a::comm_ring_1^'n^'n::finite" - assumes p: "p permutes (UNIV :: 'n::finite set)" - shows "det(\ i. A$p i :: 'a^'n^'n) = of_int (sign p) * det A" - apply (simp add: det_def setsum_right_distrib mult_assoc[symmetric]) - apply (subst sum_permutations_compose_right[OF p]) -proof(rule setsum_cong2) - let ?U = "UNIV :: 'n set" - let ?PU = "{p. p permutes ?U}" - fix q assume qPU: "q \ ?PU" - have fU: "finite ?U" by simp - from qPU have q: "q permutes ?U" by blast - from p q have pp: "permutation p" and qp: "permutation q" - by (metis fU permutation_permutes)+ - from permutes_inv[OF p] have ip: "inv p permutes ?U" . - have "setprod (\i. A$p i$ (q o p) i) ?U = setprod ((\i. A$p i$(q o p) i) o inv p) ?U" - by (simp only: setprod_permute[OF ip, symmetric]) - also have "\ = setprod (\i. A $ (p o inv p) i $ (q o (p o inv p)) i) ?U" - by (simp only: o_def) - also have "\ = setprod (\i. A$i$q i) ?U" by (simp only: o_def permutes_inverses[OF p]) - finally have thp: "setprod (\i. A$p i$ (q o p) i) ?U = setprod (\i. A$i$q i) ?U" - by blast - show "of_int (sign (q o p)) * setprod (\i. A$ p i$ (q o p) i) ?U = of_int (sign p) * of_int (sign q) * setprod (\i. A$i$q i) ?U" - by (simp only: thp sign_compose[OF qp pp] mult_commute of_int_mult) -qed - -lemma det_permute_columns: - fixes A :: "'a::comm_ring_1^'n^'n::finite" - assumes p: "p permutes (UNIV :: 'n set)" - shows "det(\ i j. A$i$ p j :: 'a^'n^'n) = of_int (sign p) * det A" -proof- - let ?Ap = "\ i j. A$i$ p j :: 'a^'n^'n" - let ?At = "transp A" - have "of_int (sign p) * det A = det (transp (\ i. transp A $ p i))" - unfolding det_permute_rows[OF p, of ?At] det_transp .. - moreover - have "?Ap = transp (\ i. transp A $ p i)" - by (simp add: transp_def Cart_eq) - ultimately show ?thesis by simp -qed - -lemma det_identical_rows: - fixes A :: "'a::ordered_idom^'n^'n::finite" - assumes ij: "i \ j" - and r: "row i A = row j A" - shows "det A = 0" -proof- - have tha: "\(a::'a) b. a = b ==> b = - a ==> a = 0" - by simp - have th1: "of_int (-1) = - 1" by (metis of_int_1 of_int_minus number_of_Min) - let ?p = "Fun.swap i j id" - let ?A = "\ i. A $ ?p i" - from r have "A = ?A" by (simp add: Cart_eq row_def swap_def) - hence "det A = det ?A" by simp - moreover have "det A = - det ?A" - by (simp add: det_permute_rows[OF permutes_swap_id] sign_swap_id ij th1) - ultimately show "det A = 0" by (metis tha) -qed - -lemma det_identical_columns: - fixes A :: "'a::ordered_idom^'n^'n::finite" - assumes ij: "i \ j" - and r: "column i A = column j A" - shows "det A = 0" -apply (subst det_transp[symmetric]) -apply (rule det_identical_rows[OF ij]) -by (metis row_transp r) - -lemma det_zero_row: - fixes A :: "'a::{idom, ring_char_0}^'n^'n::finite" - assumes r: "row i A = 0" - shows "det A = 0" -using r -apply (simp add: row_def det_def Cart_eq) -apply (rule setsum_0') -apply (auto simp: sign_nz) -done - -lemma det_zero_column: - fixes A :: "'a::{idom,ring_char_0}^'n^'n::finite" - assumes r: "column i A = 0" - shows "det A = 0" - apply (subst det_transp[symmetric]) - apply (rule det_zero_row [of i]) - by (metis row_transp r) - -lemma det_row_add: - fixes a b c :: "'n::finite \ _ ^ 'n" - shows "det((\ i. if i = k then a i + b i else c i)::'a::comm_ring_1^'n^'n) = - det((\ i. if i = k then a i else c i)::'a::comm_ring_1^'n^'n) + - det((\ i. if i = k then b i else c i)::'a::comm_ring_1^'n^'n)" -unfolding det_def Cart_lambda_beta setsum_addf[symmetric] -proof (rule setsum_cong2) - let ?U = "UNIV :: 'n set" - let ?pU = "{p. p permutes ?U}" - let ?f = "(\i. if i = k then a i + b i else c i)::'n \ 'a::comm_ring_1^'n" - let ?g = "(\ i. if i = k then a i else c i)::'n \ 'a::comm_ring_1^'n" - let ?h = "(\ i. if i = k then b i else c i)::'n \ 'a::comm_ring_1^'n" - fix p assume p: "p \ ?pU" - let ?Uk = "?U - {k}" - from p have pU: "p permutes ?U" by blast - have kU: "?U = insert k ?Uk" by blast - {fix j assume j: "j \ ?Uk" - from j have "?f j $ p j = ?g j $ p j" and "?f j $ p j= ?h j $ p j" - by simp_all} - then have th1: "setprod (\i. ?f i $ p i) ?Uk = setprod (\i. ?g i $ p i) ?Uk" - and th2: "setprod (\i. ?f i $ p i) ?Uk = setprod (\i. ?h i $ p i) ?Uk" - apply - - apply (rule setprod_cong, simp_all)+ - done - have th3: "finite ?Uk" "k \ ?Uk" by auto - have "setprod (\i. ?f i $ p i) ?U = setprod (\i. ?f i $ p i) (insert k ?Uk)" - unfolding kU[symmetric] .. - also have "\ = ?f k $ p k * setprod (\i. ?f i $ p i) ?Uk" - apply (rule setprod_insert) - apply simp - by blast - also have "\ = (a k $ p k * setprod (\i. ?f i $ p i) ?Uk) + (b k$ p k * setprod (\i. ?f i $ p i) ?Uk)" by (simp add: ring_simps) - also have "\ = (a k $ p k * setprod (\i. ?g i $ p i) ?Uk) + (b k$ p k * setprod (\i. ?h i $ p i) ?Uk)" by (metis th1 th2) - also have "\ = setprod (\i. ?g i $ p i) (insert k ?Uk) + setprod (\i. ?h i $ p i) (insert k ?Uk)" - unfolding setprod_insert[OF th3] by simp - finally have "setprod (\i. ?f i $ p i) ?U = setprod (\i. ?g i $ p i) ?U + setprod (\i. ?h i $ p i) ?U" unfolding kU[symmetric] . - then show "of_int (sign p) * setprod (\i. ?f i $ p i) ?U = of_int (sign p) * setprod (\i. ?g i $ p i) ?U + of_int (sign p) * setprod (\i. ?h i $ p i) ?U" - by (simp add: ring_simps) -qed - -lemma det_row_mul: - fixes a b :: "'n::finite \ _ ^ 'n" - shows "det((\ i. if i = k then c *s a i else b i)::'a::comm_ring_1^'n^'n) = - c* det((\ i. if i = k then a i else b i)::'a::comm_ring_1^'n^'n)" - -unfolding det_def Cart_lambda_beta setsum_right_distrib -proof (rule setsum_cong2) - let ?U = "UNIV :: 'n set" - let ?pU = "{p. p permutes ?U}" - let ?f = "(\i. if i = k then c*s a i else b i)::'n \ 'a::comm_ring_1^'n" - let ?g = "(\ i. if i = k then a i else b i)::'n \ 'a::comm_ring_1^'n" - fix p assume p: "p \ ?pU" - let ?Uk = "?U - {k}" - from p have pU: "p permutes ?U" by blast - have kU: "?U = insert k ?Uk" by blast - {fix j assume j: "j \ ?Uk" - from j have "?f j $ p j = ?g j $ p j" by simp} - then have th1: "setprod (\i. ?f i $ p i) ?Uk = setprod (\i. ?g i $ p i) ?Uk" - apply - - apply (rule setprod_cong, simp_all) - done - have th3: "finite ?Uk" "k \ ?Uk" by auto - have "setprod (\i. ?f i $ p i) ?U = setprod (\i. ?f i $ p i) (insert k ?Uk)" - unfolding kU[symmetric] .. - also have "\ = ?f k $ p k * setprod (\i. ?f i $ p i) ?Uk" - apply (rule setprod_insert) - apply simp - by blast - also have "\ = (c*s a k) $ p k * setprod (\i. ?f i $ p i) ?Uk" by (simp add: ring_simps) - also have "\ = c* (a k $ p k * setprod (\i. ?g i $ p i) ?Uk)" - unfolding th1 by (simp add: mult_ac) - also have "\ = c* (setprod (\i. ?g i $ p i) (insert k ?Uk))" - unfolding setprod_insert[OF th3] by simp - finally have "setprod (\i. ?f i $ p i) ?U = c* (setprod (\i. ?g i $ p i) ?U)" unfolding kU[symmetric] . - then show "of_int (sign p) * setprod (\i. ?f i $ p i) ?U = c * (of_int (sign p) * setprod (\i. ?g i $ p i) ?U)" - by (simp add: ring_simps) -qed - -lemma det_row_0: - fixes b :: "'n::finite \ _ ^ 'n" - shows "det((\ i. if i = k then 0 else b i)::'a::comm_ring_1^'n^'n) = 0" -using det_row_mul[of k 0 "\i. 1" b] -apply (simp) - unfolding vector_smult_lzero . - -lemma det_row_operation: - fixes A :: "'a::ordered_idom^'n^'n::finite" - assumes ij: "i \ j" - shows "det (\ k. if k = i then row i A + c *s row j A else row k A) = det A" -proof- - let ?Z = "(\ k. if k = i then row j A else row k A) :: 'a ^'n^'n" - have th: "row i ?Z = row j ?Z" by (vector row_def) - have th2: "((\ k. if k = i then row i A else row k A) :: 'a^'n^'n) = A" - by (vector row_def) - show ?thesis - unfolding det_row_add [of i] det_row_mul[of i] det_identical_rows[OF ij th] th2 - by simp -qed - -lemma det_row_span: - fixes A :: "'a:: ordered_idom^'n^'n::finite" - assumes x: "x \ span {row j A |j. j \ i}" - shows "det (\ k. if k = i then row i A + x else row k A) = det A" -proof- - let ?U = "UNIV :: 'n set" - let ?S = "{row j A |j. j \ i}" - let ?d = "\x. det (\ k. if k = i then x else row k A)" - let ?P = "\x. ?d (row i A + x) = det A" - {fix k - - have "(if k = i then row i A + 0 else row k A) = row k A" by simp} - then have P0: "?P 0" - apply - - apply (rule cong[of det, OF refl]) - by (vector row_def) - moreover - {fix c z y assume zS: "z \ ?S" and Py: "?P y" - from zS obtain j where j: "z = row j A" "i \ j" by blast - let ?w = "row i A + y" - have th0: "row i A + (c*s z + y) = ?w + c*s z" by vector - have thz: "?d z = 0" - apply (rule det_identical_rows[OF j(2)]) - using j by (vector row_def) - have "?d (row i A + (c*s z + y)) = ?d (?w + c*s z)" unfolding th0 .. - then have "?P (c*s z + y)" unfolding thz Py det_row_mul[of i] det_row_add[of i] - by simp } - - ultimately show ?thesis - apply - - apply (rule span_induct_alt[of ?P ?S, OF P0]) - apply blast - apply (rule x) - done -qed - -(* ------------------------------------------------------------------------- *) -(* May as well do this, though it's a bit unsatisfactory since it ignores *) -(* exact duplicates by considering the rows/columns as a set. *) -(* ------------------------------------------------------------------------- *) - -lemma det_dependent_rows: - fixes A:: "'a::ordered_idom^'n^'n::finite" - assumes d: "dependent (rows A)" - shows "det A = 0" -proof- - let ?U = "UNIV :: 'n set" - from d obtain i where i: "row i A \ span (rows A - {row i A})" - unfolding dependent_def rows_def by blast - {fix j k assume jk: "j \ k" - and c: "row j A = row k A" - from det_identical_rows[OF jk c] have ?thesis .} - moreover - {assume H: "\ i j. i \ j \ row i A \ row j A" - have th0: "- row i A \ span {row j A|j. j \ i}" - apply (rule span_neg) - apply (rule set_rev_mp) - apply (rule i) - apply (rule span_mono) - using H i by (auto simp add: rows_def) - from det_row_span[OF th0] - have "det A = det (\ k. if k = i then 0 *s 1 else row k A)" - unfolding right_minus vector_smult_lzero .. - with det_row_mul[of i "0::'a" "\i. 1"] - have "det A = 0" by simp} - ultimately show ?thesis by blast -qed - -lemma det_dependent_columns: assumes d: "dependent(columns (A::'a::ordered_idom^'n^'n::finite))" shows "det A = 0" -by (metis d det_dependent_rows rows_transp det_transp) - -(* ------------------------------------------------------------------------- *) -(* Multilinearity and the multiplication formula. *) -(* ------------------------------------------------------------------------- *) - -lemma Cart_lambda_cong: "(\x. f x = g x) \ (Cart_lambda f::'a^'n) = (Cart_lambda g :: 'a^'n)" - apply (rule iffD1[OF Cart_lambda_unique]) by vector - -lemma det_linear_row_setsum: - assumes fS: "finite S" - shows "det ((\ i. if i = k then setsum (a i) S else c i)::'a::comm_ring_1^'n^'n::finite) = setsum (\j. det ((\ i. if i = k then a i j else c i)::'a^'n^'n)) S" -proof(induct rule: finite_induct[OF fS]) - case 1 thus ?case apply simp unfolding setsum_empty det_row_0[of k] .. -next - case (2 x F) - then show ?case by (simp add: det_row_add cong del: if_weak_cong) -qed - -lemma finite_bounded_functions: - assumes fS: "finite S" - shows "finite {f. (\i \ {1.. (k::nat)}. f i \ S) \ (\i. i \ {1 .. k} \ f i = i)}" -proof(induct k) - case 0 - have th: "{f. \i. f i = i} = {id}" by (auto intro: ext) - show ?case by (auto simp add: th) -next - case (Suc k) - let ?f = "\(y::nat,g) i. if i = Suc k then y else g i" - let ?S = "?f ` (S \ {f. (\i\{1..k}. f i \ S) \ (\i. i \ {1..k} \ f i = i)})" - have "?S = {f. (\i\{1.. Suc k}. f i \ S) \ (\i. i \ {1.. Suc k} \ f i = i)}" - apply (auto simp add: image_iff) - apply (rule_tac x="x (Suc k)" in bexI) - apply (rule_tac x = "\i. if i = Suc k then i else x i" in exI) - apply (auto intro: ext) - done - with finite_imageI[OF finite_cartesian_product[OF fS Suc.hyps(1)], of ?f] - show ?case by metis -qed - - -lemma eq_id_iff[simp]: "(\x. f x = x) = (f = id)" by (auto intro: ext) - -lemma det_linear_rows_setsum_lemma: - assumes fS: "finite S" and fT: "finite T" - shows "det((\ i. if i \ T then setsum (a i) S else c i):: 'a::comm_ring_1^'n^'n::finite) = - setsum (\f. det((\ i. if i \ T then a i (f i) else c i)::'a^'n^'n)) - {f. (\i \ T. f i \ S) \ (\i. i \ T \ f i = i)}" -using fT -proof(induct T arbitrary: a c set: finite) - case empty - have th0: "\x y. (\ i. if i \ {} then x i else y i) = (\ i. y i)" by vector - from "empty.prems" show ?case unfolding th0 by simp -next - case (insert z T a c) - let ?F = "\T. {f. (\i \ T. f i \ S) \ (\i. i \ T \ f i = i)}" - let ?h = "\(y,g) i. if i = z then y else g i" - let ?k = "\h. (h(z),(\i. if i = z then i else h i))" - let ?s = "\ k a c f. det((\ i. if i \ T then a i (f i) else c i)::'a^'n^'n)" - let ?c = "\i. if i = z then a i j else c i" - have thif: "\a b c d. (if a \ b then c else d) = (if a then c else if b then c else d)" by simp - have thif2: "\a b c d e. (if a then b else if c then d else e) = - (if c then (if a then b else d) else (if a then b else e))" by simp - from `z \ T` have nz: "\i. i \ T \ i = z \ False" by auto - have "det (\ i. if i \ insert z T then setsum (a i) S else c i) = - det (\ i. if i = z then setsum (a i) S - else if i \ T then setsum (a i) S else c i)" - unfolding insert_iff thif .. - also have "\ = (\j\S. det (\ i. if i \ T then setsum (a i) S - else if i = z then a i j else c i))" - unfolding det_linear_row_setsum[OF fS] - apply (subst thif2) - using nz by (simp cong del: if_weak_cong cong add: if_cong) - finally have tha: - "det (\ i. if i \ insert z T then setsum (a i) S else c i) = - (\(j, f)\S \ ?F T. det (\ i. if i \ T then a i (f i) - else if i = z then a i j - else c i))" - unfolding insert.hyps unfolding setsum_cartesian_product by blast - show ?case unfolding tha - apply(rule setsum_eq_general_reverses[where h= "?h" and k= "?k"], - blast intro: finite_cartesian_product fS finite, - blast intro: finite_cartesian_product fS finite) - using `z \ T` - apply (auto intro: ext) - apply (rule cong[OF refl[of det]]) - by vector -qed - -lemma det_linear_rows_setsum: - assumes fS: "finite (S::'n::finite set)" - shows "det (\ i. setsum (a i) S) = setsum (\f. det (\ i. a i (f i) :: 'a::comm_ring_1 ^ 'n^'n::finite)) {f. \i. f i \ S}" -proof- - have th0: "\x y. ((\ i. if i \ (UNIV:: 'n set) then x i else y i) :: 'a^'n^'n) = (\ i. x i)" by vector - - from det_linear_rows_setsum_lemma[OF fS, of "UNIV :: 'n set" a, unfolded th0, OF finite] show ?thesis by simp -qed - -lemma matrix_mul_setsum_alt: - fixes A B :: "'a::comm_ring_1^'n^'n::finite" - shows "A ** B = (\ i. setsum (\k. A$i$k *s B $ k) (UNIV :: 'n set))" - by (vector matrix_matrix_mult_def setsum_component) - -lemma det_rows_mul: - "det((\ i. c i *s a i)::'a::comm_ring_1^'n^'n::finite) = - setprod (\i. c i) (UNIV:: 'n set) * det((\ i. a i)::'a^'n^'n)" -proof (simp add: det_def setsum_right_distrib cong add: setprod_cong, rule setsum_cong2) - let ?U = "UNIV :: 'n set" - let ?PU = "{p. p permutes ?U}" - fix p assume pU: "p \ ?PU" - let ?s = "of_int (sign p)" - from pU have p: "p permutes ?U" by blast - have "setprod (\i. c i * a i $ p i) ?U = setprod c ?U * setprod (\i. a i $ p i) ?U" - unfolding setprod_timesf .. - then show "?s * (\xa\?U. c xa * a xa $ p xa) = - setprod c ?U * (?s* (\xa\?U. a xa $ p xa))" by (simp add: ring_simps) -qed - -lemma det_mul: - fixes A B :: "'a::ordered_idom^'n^'n::finite" - shows "det (A ** B) = det A * det B" -proof- - let ?U = "UNIV :: 'n set" - let ?F = "{f. (\i\ ?U. f i \ ?U) \ (\i. i \ ?U \ f i = i)}" - let ?PU = "{p. p permutes ?U}" - have fU: "finite ?U" by simp - have fF: "finite ?F" by (rule finite) - {fix p assume p: "p permutes ?U" - - have "p \ ?F" unfolding mem_Collect_eq permutes_in_image[OF p] - using p[unfolded permutes_def] by simp} - then have PUF: "?PU \ ?F" by blast - {fix f assume fPU: "f \ ?F - ?PU" - have fUU: "f ` ?U \ ?U" using fPU by auto - from fPU have f: "\i \ ?U. f i \ ?U" - "\i. i \ ?U \ f i = i" "\(\y. \!x. f x = y)" unfolding permutes_def - by auto - - let ?A = "(\ i. A$i$f i *s B$f i) :: 'a^'n^'n" - let ?B = "(\ i. B$f i) :: 'a^'n^'n" - {assume fni: "\ inj_on f ?U" - then obtain i j where ij: "f i = f j" "i \ j" - unfolding inj_on_def by blast - from ij - have rth: "row i ?B = row j ?B" by (vector row_def) - from det_identical_rows[OF ij(2) rth] - have "det (\ i. A$i$f i *s B$f i) = 0" - unfolding det_rows_mul by simp} - moreover - {assume fi: "inj_on f ?U" - from f fi have fith: "\i j. f i = f j \ i = j" - unfolding inj_on_def by metis - note fs = fi[unfolded surjective_iff_injective_gen[OF fU fU refl fUU, symmetric]] - - {fix y - from fs f have "\x. f x = y" by blast - then obtain x where x: "f x = y" by blast - {fix z assume z: "f z = y" from fith x z have "z = x" by metis} - with x have "\!x. f x = y" by blast} - with f(3) have "det (\ i. A$i$f i *s B$f i) = 0" by blast} - ultimately have "det (\ i. A$i$f i *s B$f i) = 0" by blast} - hence zth: "\ f\ ?F - ?PU. det (\ i. A$i$f i *s B$f i) = 0" by simp - {fix p assume pU: "p \ ?PU" - from pU have p: "p permutes ?U" by blast - let ?s = "\p. of_int (sign p)" - let ?f = "\q. ?s p * (\i\ ?U. A $ i $ p i) * - (?s q * (\i\ ?U. B $ i $ q i))" - have "(setsum (\q. ?s q * - (\i\ ?U. (\ i. A $ i $ p i *s B $ p i :: 'a^'n^'n) $ i $ q i)) ?PU) = - (setsum (\q. ?s p * (\i\ ?U. A $ i $ p i) * - (?s q * (\i\ ?U. B $ i $ q i))) ?PU)" - unfolding sum_permutations_compose_right[OF permutes_inv[OF p], of ?f] - proof(rule setsum_cong2) - fix q assume qU: "q \ ?PU" - hence q: "q permutes ?U" by blast - from p q have pp: "permutation p" and pq: "permutation q" - unfolding permutation_permutes by auto - have th00: "of_int (sign p) * of_int (sign p) = (1::'a)" - "\a. of_int (sign p) * (of_int (sign p) * a) = a" - unfolding mult_assoc[symmetric] unfolding of_int_mult[symmetric] - by (simp_all add: sign_idempotent) - have ths: "?s q = ?s p * ?s (q o inv p)" - using pp pq permutation_inverse[OF pp] sign_inverse[OF pp] - by (simp add: th00 mult_ac sign_idempotent sign_compose) - have th001: "setprod (\i. B$i$ q (inv p i)) ?U = setprod ((\i. B$i$ q (inv p i)) o p) ?U" - by (rule setprod_permute[OF p]) - have thp: "setprod (\i. (\ i. A$i$p i *s B$p i :: 'a^'n^'n) $i $ q i) ?U = setprod (\i. A$i$p i) ?U * setprod (\i. B$i$ q (inv p i)) ?U" - unfolding th001 setprod_timesf[symmetric] o_def permutes_inverses[OF p] - apply (rule setprod_cong[OF refl]) - using permutes_in_image[OF q] by vector - show "?s q * setprod (\i. (((\ i. A$i$p i *s B$p i) :: 'a^'n^'n)$i$q i)) ?U = ?s p * (setprod (\i. A$i$p i) ?U) * (?s (q o inv p) * setprod (\i. B$i$(q o inv p) i) ?U)" - using ths thp pp pq permutation_inverse[OF pp] sign_inverse[OF pp] - by (simp add: sign_nz th00 ring_simps sign_idempotent sign_compose) - qed - } - then have th2: "setsum (\f. det (\ i. A$i$f i *s B$f i)) ?PU = det A * det B" - unfolding det_def setsum_product - by (rule setsum_cong2) - have "det (A**B) = setsum (\f. det (\ i. A $ i $ f i *s B $ f i)) ?F" - unfolding matrix_mul_setsum_alt det_linear_rows_setsum[OF fU] by simp - also have "\ = setsum (\f. det (\ i. A$i$f i *s B$f i)) ?PU" - using setsum_mono_zero_cong_left[OF fF PUF zth, symmetric] - unfolding det_rows_mul by auto - finally show ?thesis unfolding th2 . -qed - -(* ------------------------------------------------------------------------- *) -(* Relation to invertibility. *) -(* ------------------------------------------------------------------------- *) - -lemma invertible_left_inverse: - fixes A :: "real^'n^'n::finite" - shows "invertible A \ (\(B::real^'n^'n). B** A = mat 1)" - by (metis invertible_def matrix_left_right_inverse) - -lemma invertible_righ_inverse: - fixes A :: "real^'n^'n::finite" - shows "invertible A \ (\(B::real^'n^'n). A** B = mat 1)" - by (metis invertible_def matrix_left_right_inverse) - -lemma invertible_det_nz: - fixes A::"real ^'n^'n::finite" - shows "invertible A \ det A \ 0" -proof- - {assume "invertible A" - then obtain B :: "real ^'n^'n" where B: "A ** B = mat 1" - unfolding invertible_righ_inverse by blast - hence "det (A ** B) = det (mat 1 :: real ^'n^'n)" by simp - hence "det A \ 0" - apply (simp add: det_mul det_I) by algebra } - moreover - {assume H: "\ invertible A" - let ?U = "UNIV :: 'n set" - have fU: "finite ?U" by simp - from H obtain c i where c: "setsum (\i. c i *s row i A) ?U = 0" - and iU: "i \ ?U" and ci: "c i \ 0" - unfolding invertible_righ_inverse - unfolding matrix_right_invertible_independent_rows by blast - have stupid: "\(a::real^'n) b. a + b = 0 \ -a = b" - apply (drule_tac f="op + (- a)" in cong[OF refl]) - apply (simp only: ab_left_minus add_assoc[symmetric]) - apply simp - done - from c ci - have thr0: "- row i A = setsum (\j. (1/ c i) *s (c j *s row j A)) (?U - {i})" - unfolding setsum_diff1'[OF fU iU] setsum_cmul - apply - - apply (rule vector_mul_lcancel_imp[OF ci]) - apply (auto simp add: vector_smult_assoc vector_smult_rneg field_simps) - unfolding stupid .. - have thr: "- row i A \ span {row j A| j. j \ i}" - unfolding thr0 - apply (rule span_setsum) - apply simp - apply (rule ballI) - apply (rule span_mul)+ - apply (rule span_superset) - apply auto - done - let ?B = "(\ k. if k = i then 0 else row k A) :: real ^'n^'n" - have thrb: "row i ?B = 0" using iU by (vector row_def) - have "det A = 0" - unfolding det_row_span[OF thr, symmetric] right_minus - unfolding det_zero_row[OF thrb] ..} - ultimately show ?thesis by blast -qed - -(* ------------------------------------------------------------------------- *) -(* Cramer's rule. *) -(* ------------------------------------------------------------------------- *) - -lemma cramer_lemma_transp: - fixes A:: "'a::ordered_idom^'n^'n::finite" and x :: "'a ^'n::finite" - shows "det ((\ i. if i = k then setsum (\i. x$i *s row i A) (UNIV::'n set) - else row i A)::'a^'n^'n) = x$k * det A" - (is "?lhs = ?rhs") -proof- - let ?U = "UNIV :: 'n set" - let ?Uk = "?U - {k}" - have U: "?U = insert k ?Uk" by blast - have fUk: "finite ?Uk" by simp - have kUk: "k \ ?Uk" by simp - have th00: "\k s. x$k *s row k A + s = (x$k - 1) *s row k A + row k A + s" - by (vector ring_simps) - have th001: "\f k . (\x. if x = k then f k else f x) = f" by (auto intro: ext) - have "(\ i. row i A) = A" by (vector row_def) - then have thd1: "det (\ i. row i A) = det A" by simp - have thd0: "det (\ i. if i = k then row k A + (\i \ ?Uk. x $ i *s row i A) else row i A) = det A" - apply (rule det_row_span) - apply (rule span_setsum[OF fUk]) - apply (rule ballI) - apply (rule span_mul) - apply (rule span_superset) - apply auto - done - show "?lhs = x$k * det A" - apply (subst U) - unfolding setsum_insert[OF fUk kUk] - apply (subst th00) - unfolding add_assoc - apply (subst det_row_add) - unfolding thd0 - unfolding det_row_mul - unfolding th001[of k "\i. row i A"] - unfolding thd1 by (simp add: ring_simps) -qed - -lemma cramer_lemma: - fixes A :: "'a::ordered_idom ^'n^'n::finite" - shows "det((\ i j. if j = k then (A *v x)$i else A$i$j):: 'a^'n^'n) = x$k * det A" -proof- - let ?U = "UNIV :: 'n set" - have stupid: "\c. setsum (\i. c i *s row i (transp A)) ?U = setsum (\i. c i *s column i A) ?U" - by (auto simp add: row_transp intro: setsum_cong2) - show ?thesis unfolding matrix_mult_vsum - unfolding cramer_lemma_transp[of k x "transp A", unfolded det_transp, symmetric] - unfolding stupid[of "\i. x$i"] - apply (subst det_transp[symmetric]) - apply (rule cong[OF refl[of det]]) by (vector transp_def column_def row_def) -qed - -lemma cramer: - fixes A ::"real^'n^'n::finite" - assumes d0: "det A \ 0" - shows "A *v x = b \ x = (\ k. det(\ i j. if j=k then b$i else A$i$j :: real^'n^'n) / det A)" -proof- - from d0 obtain B where B: "A ** B = mat 1" "B ** A = mat 1" - unfolding invertible_det_nz[symmetric] invertible_def by blast - have "(A ** B) *v b = b" by (simp add: B matrix_vector_mul_lid) - hence "A *v (B *v b) = b" by (simp add: matrix_vector_mul_assoc) - then have xe: "\x. A*v x = b" by blast - {fix x assume x: "A *v x = b" - have "x = (\ k. det(\ i j. if j=k then b$i else A$i$j :: real^'n^'n) / det A)" - unfolding x[symmetric] - using d0 by (simp add: Cart_eq cramer_lemma field_simps)} - with xe show ?thesis by auto -qed - -(* ------------------------------------------------------------------------- *) -(* Orthogonality of a transformation and matrix. *) -(* ------------------------------------------------------------------------- *) - -definition "orthogonal_transformation f \ linear f \ (\v w. f v \ f w = v \ w)" - -lemma orthogonal_transformation: "orthogonal_transformation f \ linear f \ (\(v::real ^_). norm (f v) = norm v)" - unfolding orthogonal_transformation_def - apply auto - apply (erule_tac x=v in allE)+ - apply (simp add: real_vector_norm_def) - by (simp add: dot_norm linear_add[symmetric]) - -definition "orthogonal_matrix (Q::'a::semiring_1^'n^'n) \ transp Q ** Q = mat 1 \ Q ** transp Q = mat 1" - -lemma orthogonal_matrix: "orthogonal_matrix (Q:: real ^'n^'n::finite) \ transp Q ** Q = mat 1" - by (metis matrix_left_right_inverse orthogonal_matrix_def) - -lemma orthogonal_matrix_id: "orthogonal_matrix (mat 1 :: _^'n^'n::finite)" - by (simp add: orthogonal_matrix_def transp_mat matrix_mul_lid) - -lemma orthogonal_matrix_mul: - fixes A :: "real ^'n^'n::finite" - assumes oA : "orthogonal_matrix A" - and oB: "orthogonal_matrix B" - shows "orthogonal_matrix(A ** B)" - using oA oB - unfolding orthogonal_matrix matrix_transp_mul - apply (subst matrix_mul_assoc) - apply (subst matrix_mul_assoc[symmetric]) - by (simp add: matrix_mul_rid) - -lemma orthogonal_transformation_matrix: - fixes f:: "real^'n \ real^'n::finite" - shows "orthogonal_transformation f \ linear f \ orthogonal_matrix(matrix f)" - (is "?lhs \ ?rhs") -proof- - let ?mf = "matrix f" - let ?ot = "orthogonal_transformation f" - let ?U = "UNIV :: 'n set" - have fU: "finite ?U" by simp - let ?m1 = "mat 1 :: real ^'n^'n" - {assume ot: ?ot - from ot have lf: "linear f" and fd: "\v w. f v \ f w = v \ w" - unfolding orthogonal_transformation_def orthogonal_matrix by blast+ - {fix i j - let ?A = "transp ?mf ** ?mf" - have th0: "\b (x::'a::comm_ring_1). (if b then 1 else 0)*x = (if b then x else 0)" - "\b (x::'a::comm_ring_1). x*(if b then 1 else 0) = (if b then x else 0)" - by simp_all - from fd[rule_format, of "basis i" "basis j", unfolded matrix_works[OF lf, symmetric] dot_matrix_vector_mul] - have "?A$i$j = ?m1 $ i $ j" - by (simp add: dot_def matrix_matrix_mult_def columnvector_def rowvector_def basis_def th0 setsum_delta[OF fU] mat_def)} - hence "orthogonal_matrix ?mf" unfolding orthogonal_matrix by vector - with lf have ?rhs by blast} - moreover - {assume lf: "linear f" and om: "orthogonal_matrix ?mf" - from lf om have ?lhs - unfolding orthogonal_matrix_def norm_eq orthogonal_transformation - unfolding matrix_works[OF lf, symmetric] - apply (subst dot_matrix_vector_mul) - by (simp add: dot_matrix_product matrix_mul_lid)} - ultimately show ?thesis by blast -qed - -lemma det_orthogonal_matrix: - fixes Q:: "'a::ordered_idom^'n^'n::finite" - assumes oQ: "orthogonal_matrix Q" - shows "det Q = 1 \ det Q = - 1" -proof- - - have th: "\x::'a. x = 1 \ x = - 1 \ x*x = 1" (is "\x::'a. ?ths x") - proof- - fix x:: 'a - have th0: "x*x - 1 = (x - 1)*(x + 1)" by (simp add: ring_simps) - have th1: "\(x::'a) y. x = - y \ x + y = 0" - apply (subst eq_iff_diff_eq_0) by simp - have "x*x = 1 \ x*x - 1 = 0" by simp - also have "\ \ x = 1 \ x = - 1" unfolding th0 th1 by simp - finally show "?ths x" .. - qed - from oQ have "Q ** transp Q = mat 1" by (metis orthogonal_matrix_def) - hence "det (Q ** transp Q) = det (mat 1:: 'a^'n^'n)" by simp - hence "det Q * det Q = 1" by (simp add: det_mul det_I det_transp) - then show ?thesis unfolding th . -qed - -(* ------------------------------------------------------------------------- *) -(* Linearity of scaling, and hence isometry, that preserves origin. *) -(* ------------------------------------------------------------------------- *) -lemma scaling_linear: - fixes f :: "real ^'n \ real ^'n::finite" - assumes f0: "f 0 = 0" and fd: "\x y. dist (f x) (f y) = c * dist x y" - shows "linear f" -proof- - {fix v w - {fix x note fd[rule_format, of x 0, unfolded dist_norm f0 diff_0_right] } - note th0 = this - have "f v \ f w = c^2 * (v \ w)" - unfolding dot_norm_neg dist_norm[symmetric] - unfolding th0 fd[rule_format] by (simp add: power2_eq_square field_simps)} - note fc = this - show ?thesis unfolding linear_def vector_eq - by (simp add: dot_lmult dot_ladd dot_rmult dot_radd fc ring_simps) -qed - -lemma isometry_linear: - "f (0:: real^'n) = (0:: real^'n::finite) \ \x y. dist(f x) (f y) = dist x y - \ linear f" -by (rule scaling_linear[where c=1]) simp_all - -(* ------------------------------------------------------------------------- *) -(* Hence another formulation of orthogonal transformation. *) -(* ------------------------------------------------------------------------- *) - -lemma orthogonal_transformation_isometry: - "orthogonal_transformation f \ f(0::real^'n) = (0::real^'n::finite) \ (\x y. dist(f x) (f y) = dist x y)" - unfolding orthogonal_transformation - apply (rule iffI) - apply clarify - apply (clarsimp simp add: linear_0 linear_sub[symmetric] dist_norm) - apply (rule conjI) - apply (rule isometry_linear) - apply simp - apply simp - apply clarify - apply (erule_tac x=v in allE) - apply (erule_tac x=0 in allE) - by (simp add: dist_norm) - -(* ------------------------------------------------------------------------- *) -(* Can extend an isometry from unit sphere. *) -(* ------------------------------------------------------------------------- *) - -lemma isometry_sphere_extend: - fixes f:: "real ^'n \ real ^'n::finite" - assumes f1: "\x. norm x = 1 \ norm (f x) = 1" - and fd1: "\ x y. norm x = 1 \ norm y = 1 \ dist (f x) (f y) = dist x y" - shows "\g. orthogonal_transformation g \ (\x. norm x = 1 \ g x = f x)" -proof- - {fix x y x' y' x0 y0 x0' y0' :: "real ^'n" - assume H: "x = norm x *s x0" "y = norm y *s y0" - "x' = norm x *s x0'" "y' = norm y *s y0'" - "norm x0 = 1" "norm x0' = 1" "norm y0 = 1" "norm y0' = 1" - "norm(x0' - y0') = norm(x0 - y0)" - - have "norm(x' - y') = norm(x - y)" - apply (subst H(1)) - apply (subst H(2)) - apply (subst H(3)) - apply (subst H(4)) - using H(5-9) - apply (simp add: norm_eq norm_eq_1) - apply (simp add: dot_lsub dot_rsub dot_lmult dot_rmult) - apply (simp add: ring_simps) - by (simp only: right_distrib[symmetric])} - note th0 = this - let ?g = "\x. if x = 0 then 0 else norm x *s f (inverse (norm x) *s x)" - {fix x:: "real ^'n" assume nx: "norm x = 1" - have "?g x = f x" using nx by auto} - hence thfg: "\x. norm x = 1 \ ?g x = f x" by blast - have g0: "?g 0 = 0" by simp - {fix x y :: "real ^'n" - {assume "x = 0" "y = 0" - then have "dist (?g x) (?g y) = dist x y" by simp } - moreover - {assume "x = 0" "y \ 0" - then have "dist (?g x) (?g y) = dist x y" - apply (simp add: dist_norm norm_mul) - apply (rule f1[rule_format]) - by(simp add: norm_mul field_simps)} - moreover - {assume "x \ 0" "y = 0" - then have "dist (?g x) (?g y) = dist x y" - apply (simp add: dist_norm norm_mul) - apply (rule f1[rule_format]) - by(simp add: norm_mul field_simps)} - moreover - {assume z: "x \ 0" "y \ 0" - have th00: "x = norm x *s (inverse (norm x) *s x)" "y = norm y *s (inverse (norm y) *s y)" "norm x *s f ((inverse (norm x) *s x)) = norm x *s f (inverse (norm x) *s x)" - "norm y *s f (inverse (norm y) *s y) = norm y *s f (inverse (norm y) *s y)" - "norm (inverse (norm x) *s x) = 1" - "norm (f (inverse (norm x) *s x)) = 1" - "norm (inverse (norm y) *s y) = 1" - "norm (f (inverse (norm y) *s y)) = 1" - "norm (f (inverse (norm x) *s x) - f (inverse (norm y) *s y)) = - norm (inverse (norm x) *s x - inverse (norm y) *s y)" - using z - by (auto simp add: vector_smult_assoc field_simps norm_mul intro: f1[rule_format] fd1[rule_format, unfolded dist_norm]) - from z th0[OF th00] have "dist (?g x) (?g y) = dist x y" - by (simp add: dist_norm)} - ultimately have "dist (?g x) (?g y) = dist x y" by blast} - note thd = this - show ?thesis - apply (rule exI[where x= ?g]) - unfolding orthogonal_transformation_isometry - using g0 thfg thd by metis -qed - -(* ------------------------------------------------------------------------- *) -(* Rotation, reflection, rotoinversion. *) -(* ------------------------------------------------------------------------- *) - -definition "rotation_matrix Q \ orthogonal_matrix Q \ det Q = 1" -definition "rotoinversion_matrix Q \ orthogonal_matrix Q \ det Q = - 1" - -lemma orthogonal_rotation_or_rotoinversion: - fixes Q :: "'a::ordered_idom^'n^'n::finite" - shows " orthogonal_matrix Q \ rotation_matrix Q \ rotoinversion_matrix Q" - by (metis rotoinversion_matrix_def rotation_matrix_def det_orthogonal_matrix) -(* ------------------------------------------------------------------------- *) -(* Explicit formulas for low dimensions. *) -(* ------------------------------------------------------------------------- *) - -lemma setprod_1: "setprod f {(1::nat)..1} = f 1" by simp - -lemma setprod_2: "setprod f {(1::nat)..2} = f 1 * f 2" - by (simp add: nat_number setprod_numseg mult_commute) -lemma setprod_3: "setprod f {(1::nat)..3} = f 1 * f 2 * f 3" - by (simp add: nat_number setprod_numseg mult_commute) - -lemma det_1: "det (A::'a::comm_ring_1^1^1) = A$1$1" - by (simp add: det_def permutes_sing sign_id UNIV_1) - -lemma det_2: "det (A::'a::comm_ring_1^2^2) = A$1$1 * A$2$2 - A$1$2 * A$2$1" -proof- - have f12: "finite {2::2}" "1 \ {2::2}" by auto - show ?thesis - unfolding det_def UNIV_2 - unfolding setsum_over_permutations_insert[OF f12] - unfolding permutes_sing - apply (simp add: sign_swap_id sign_id swap_id_eq) - by (simp add: arith_simps(31)[symmetric] of_int_minus of_int_1 del: arith_simps(31)) -qed - -lemma det_3: "det (A::'a::comm_ring_1^3^3) = - A$1$1 * A$2$2 * A$3$3 + - A$1$2 * A$2$3 * A$3$1 + - A$1$3 * A$2$1 * A$3$2 - - A$1$1 * A$2$3 * A$3$2 - - A$1$2 * A$2$1 * A$3$3 - - A$1$3 * A$2$2 * A$3$1" -proof- - have f123: "finite {2::3, 3}" "1 \ {2::3, 3}" by auto - have f23: "finite {3::3}" "2 \ {3::3}" by auto - - show ?thesis - unfolding det_def UNIV_3 - unfolding setsum_over_permutations_insert[OF f123] - unfolding setsum_over_permutations_insert[OF f23] - - unfolding permutes_sing - apply (simp add: sign_swap_id permutation_swap_id sign_compose sign_id swap_id_eq) - apply (simp add: arith_simps(31)[symmetric] of_int_minus of_int_1 del: arith_simps(31)) - by (simp add: ring_simps) -qed - -end diff -r 1fad3160d873 -r 2083bde13ce1 src/HOL/Library/Euclidean_Space.thy --- a/src/HOL/Library/Euclidean_Space.thy Fri Oct 23 14:33:07 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,5372 +0,0 @@ -(* Title: Library/Euclidean_Space - Author: Amine Chaieb, University of Cambridge -*) - -header {* (Real) Vectors in Euclidean space, and elementary linear algebra.*} - -theory Euclidean_Space -imports - Complex_Main "~~/src/HOL/Decision_Procs/Dense_Linear_Order" - Finite_Cartesian_Product Glbs Infinite_Set Numeral_Type - Inner_Product -uses "positivstellensatz.ML" ("normarith.ML") -begin - -text{* Some common special cases.*} - -lemma forall_1: "(\i::1. P i) \ P 1" - by (metis num1_eq_iff) - -lemma exhaust_2: - fixes x :: 2 shows "x = 1 \ x = 2" -proof (induct x) - case (of_int z) - then have "0 <= z" and "z < 2" by simp_all - then have "z = 0 | z = 1" by arith - then show ?case by auto -qed - -lemma forall_2: "(\i::2. P i) \ P 1 \ P 2" - by (metis exhaust_2) - -lemma exhaust_3: - fixes x :: 3 shows "x = 1 \ x = 2 \ x = 3" -proof (induct x) - case (of_int z) - then have "0 <= z" and "z < 3" by simp_all - then have "z = 0 \ z = 1 \ z = 2" by arith - then show ?case by auto -qed - -lemma forall_3: "(\i::3. P i) \ P 1 \ P 2 \ P 3" - by (metis exhaust_3) - -lemma UNIV_1: "UNIV = {1::1}" - by (auto simp add: num1_eq_iff) - -lemma UNIV_2: "UNIV = {1::2, 2::2}" - using exhaust_2 by auto - -lemma UNIV_3: "UNIV = {1::3, 2::3, 3::3}" - using exhaust_3 by auto - -lemma setsum_1: "setsum f (UNIV::1 set) = f 1" - unfolding UNIV_1 by simp - -lemma setsum_2: "setsum f (UNIV::2 set) = f 1 + f 2" - unfolding UNIV_2 by simp - -lemma setsum_3: "setsum f (UNIV::3 set) = f 1 + f 2 + f 3" - unfolding UNIV_3 by (simp add: add_ac) - -subsection{* Basic componentwise operations on vectors. *} - -instantiation "^" :: (plus,type) plus -begin -definition vector_add_def : "op + \ (\ x y. (\ i. (x$i) + (y$i)))" -instance .. -end - -instantiation "^" :: (times,type) times -begin - definition vector_mult_def : "op * \ (\ x y. (\ i. (x$i) * (y$i)))" - instance .. -end - -instantiation "^" :: (minus,type) minus begin - definition vector_minus_def : "op - \ (\ x y. (\ i. (x$i) - (y$i)))" -instance .. -end - -instantiation "^" :: (uminus,type) uminus begin - definition vector_uminus_def : "uminus \ (\ x. (\ i. - (x$i)))" -instance .. -end -instantiation "^" :: (zero,type) zero begin - definition vector_zero_def : "0 \ (\ i. 0)" -instance .. -end - -instantiation "^" :: (one,type) one begin - definition vector_one_def : "1 \ (\ i. 1)" -instance .. -end - -instantiation "^" :: (ord,type) ord - begin -definition vector_less_eq_def: - "less_eq (x :: 'a ^'b) y = (ALL i. x$i <= y$i)" -definition vector_less_def: "less (x :: 'a ^'b) y = (ALL i. x$i < y$i)" - -instance by (intro_classes) -end - -instantiation "^" :: (scaleR, type) scaleR -begin -definition vector_scaleR_def: "scaleR = (\ r x. (\ i. scaleR r (x$i)))" -instance .. -end - -text{* Also the scalar-vector multiplication. *} - -definition vector_scalar_mult:: "'a::times \ 'a ^'n \ 'a ^ 'n" (infixl "*s" 70) - where "c *s x = (\ i. c * (x$i))" - -text{* Constant Vectors *} - -definition "vec x = (\ i. x)" - -text{* Dot products. *} - -definition dot :: "'a::{comm_monoid_add, times} ^ 'n \ 'a ^ 'n \ 'a" (infix "\" 70) where - "x \ y = setsum (\i. x$i * y$i) UNIV" - -lemma dot_1[simp]: "(x::'a::{comm_monoid_add, times}^1) \ y = (x$1) * (y$1)" - by (simp add: dot_def setsum_1) - -lemma dot_2[simp]: "(x::'a::{comm_monoid_add, times}^2) \ y = (x$1) * (y$1) + (x$2) * (y$2)" - by (simp add: dot_def setsum_2) - -lemma dot_3[simp]: "(x::'a::{comm_monoid_add, times}^3) \ y = (x$1) * (y$1) + (x$2) * (y$2) + (x$3) * (y$3)" - by (simp add: dot_def setsum_3) - -subsection {* A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space. *} - -method_setup vector = {* -let - val ss1 = HOL_basic_ss addsimps [@{thm dot_def}, @{thm setsum_addf} RS sym, - @{thm setsum_subtractf} RS sym, @{thm setsum_right_distrib}, - @{thm setsum_left_distrib}, @{thm setsum_negf} RS sym] - val ss2 = @{simpset} addsimps - [@{thm vector_add_def}, @{thm vector_mult_def}, - @{thm vector_minus_def}, @{thm vector_uminus_def}, - @{thm vector_one_def}, @{thm vector_zero_def}, @{thm vec_def}, - @{thm vector_scaleR_def}, - @{thm Cart_lambda_beta}, @{thm vector_scalar_mult_def}] - fun vector_arith_tac ths = - simp_tac ss1 - THEN' (fn i => rtac @{thm setsum_cong2} i - ORELSE rtac @{thm setsum_0'} i - ORELSE simp_tac (HOL_basic_ss addsimps [@{thm "Cart_eq"}]) i) - (* THEN' TRY o clarify_tac HOL_cs THEN' (TRY o rtac @{thm iffI}) *) - THEN' asm_full_simp_tac (ss2 addsimps ths) - in - Attrib.thms >> (fn ths => K (SIMPLE_METHOD' (vector_arith_tac ths))) - end -*} "Lifts trivial vector statements to real arith statements" - -lemma vec_0[simp]: "vec 0 = 0" by (vector vector_zero_def) -lemma vec_1[simp]: "vec 1 = 1" by (vector vector_one_def) - - - -text{* Obvious "component-pushing". *} - -lemma vec_component [simp]: "(vec x :: 'a ^ 'n)$i = x" - by (vector vec_def) - -lemma vector_add_component [simp]: - fixes x y :: "'a::{plus} ^ 'n" - shows "(x + y)$i = x$i + y$i" - by vector - -lemma vector_minus_component [simp]: - fixes x y :: "'a::{minus} ^ 'n" - shows "(x - y)$i = x$i - y$i" - by vector - -lemma vector_mult_component [simp]: - fixes x y :: "'a::{times} ^ 'n" - shows "(x * y)$i = x$i * y$i" - by vector - -lemma vector_smult_component [simp]: - fixes y :: "'a::{times} ^ 'n" - shows "(c *s y)$i = c * (y$i)" - by vector - -lemma vector_uminus_component [simp]: - fixes x :: "'a::{uminus} ^ 'n" - shows "(- x)$i = - (x$i)" - by vector - -lemma vector_scaleR_component [simp]: - fixes x :: "'a::scaleR ^ 'n" - shows "(scaleR r x)$i = scaleR r (x$i)" - by vector - -lemma cond_component: "(if b then x else y)$i = (if b then x$i else y$i)" by vector - -lemmas vector_component = - vec_component vector_add_component vector_mult_component - vector_smult_component vector_minus_component vector_uminus_component - vector_scaleR_component cond_component - -subsection {* Some frequently useful arithmetic lemmas over vectors. *} - -instance "^" :: (semigroup_add,type) semigroup_add - apply (intro_classes) by (vector add_assoc) - - -instance "^" :: (monoid_add,type) monoid_add - apply (intro_classes) by vector+ - -instance "^" :: (group_add,type) group_add - apply (intro_classes) by (vector algebra_simps)+ - -instance "^" :: (ab_semigroup_add,type) ab_semigroup_add - apply (intro_classes) by (vector add_commute) - -instance "^" :: (comm_monoid_add,type) comm_monoid_add - apply (intro_classes) by vector - -instance "^" :: (ab_group_add,type) ab_group_add - apply (intro_classes) by vector+ - -instance "^" :: (cancel_semigroup_add,type) cancel_semigroup_add - apply (intro_classes) - by (vector Cart_eq)+ - -instance "^" :: (cancel_ab_semigroup_add,type) cancel_ab_semigroup_add - apply (intro_classes) - by (vector Cart_eq) - -instance "^" :: (real_vector, type) real_vector - by default (vector scaleR_left_distrib scaleR_right_distrib)+ - -instance "^" :: (semigroup_mult,type) semigroup_mult - apply (intro_classes) by (vector mult_assoc) - -instance "^" :: (monoid_mult,type) monoid_mult - apply (intro_classes) by vector+ - -instance "^" :: (ab_semigroup_mult,type) ab_semigroup_mult - apply (intro_classes) by (vector mult_commute) - -instance "^" :: (ab_semigroup_idem_mult,type) ab_semigroup_idem_mult - apply (intro_classes) by (vector mult_idem) - -instance "^" :: (comm_monoid_mult,type) comm_monoid_mult - apply (intro_classes) by vector - -fun vector_power :: "('a::{one,times} ^'n) \ nat \ 'a^'n" where - "vector_power x 0 = 1" - | "vector_power x (Suc n) = x * vector_power x n" - -instance "^" :: (semiring,type) semiring - apply (intro_classes) by (vector ring_simps)+ - -instance "^" :: (semiring_0,type) semiring_0 - apply (intro_classes) by (vector ring_simps)+ -instance "^" :: (semiring_1,type) semiring_1 - apply (intro_classes) by vector -instance "^" :: (comm_semiring,type) comm_semiring - apply (intro_classes) by (vector ring_simps)+ - -instance "^" :: (comm_semiring_0,type) comm_semiring_0 by (intro_classes) -instance "^" :: (cancel_comm_monoid_add, type) cancel_comm_monoid_add .. -instance "^" :: (semiring_0_cancel,type) semiring_0_cancel by (intro_classes) -instance "^" :: (comm_semiring_0_cancel,type) comm_semiring_0_cancel by (intro_classes) -instance "^" :: (ring,type) ring by (intro_classes) -instance "^" :: (semiring_1_cancel,type) semiring_1_cancel by (intro_classes) -instance "^" :: (comm_semiring_1,type) comm_semiring_1 by (intro_classes) - -instance "^" :: (ring_1,type) ring_1 .. - -instance "^" :: (real_algebra,type) real_algebra - apply intro_classes - apply (simp_all add: vector_scaleR_def ring_simps) - apply vector - apply vector - done - -instance "^" :: (real_algebra_1,type) real_algebra_1 .. - -lemma of_nat_index: - "(of_nat n :: 'a::semiring_1 ^'n)$i = of_nat n" - apply (induct n) - apply vector - apply vector - done -lemma zero_index[simp]: - "(0 :: 'a::zero ^'n)$i = 0" by vector - -lemma one_index[simp]: - "(1 :: 'a::one ^'n)$i = 1" by vector - -lemma one_plus_of_nat_neq_0: "(1::'a::semiring_char_0) + of_nat n \ 0" -proof- - have "(1::'a) + of_nat n = 0 \ of_nat 1 + of_nat n = (of_nat 0 :: 'a)" by simp - also have "\ \ 1 + n = 0" by (simp only: of_nat_add[symmetric] of_nat_eq_iff) - finally show ?thesis by simp -qed - -instance "^" :: (semiring_char_0,type) semiring_char_0 -proof (intro_classes) - fix m n ::nat - show "(of_nat m :: 'a^'b) = of_nat n \ m = n" - by (simp add: Cart_eq of_nat_index) -qed - -instance "^" :: (comm_ring_1,type) comm_ring_1 by intro_classes -instance "^" :: (ring_char_0,type) ring_char_0 by intro_classes - -lemma vector_smult_assoc: "a *s (b *s x) = ((a::'a::semigroup_mult) * b) *s x" - by (vector mult_assoc) -lemma vector_sadd_rdistrib: "((a::'a::semiring) + b) *s x = a *s x + b *s x" - by (vector ring_simps) -lemma vector_add_ldistrib: "(c::'a::semiring) *s (x + y) = c *s x + c *s y" - by (vector ring_simps) -lemma vector_smult_lzero[simp]: "(0::'a::mult_zero) *s x = 0" by vector -lemma vector_smult_lid[simp]: "(1::'a::monoid_mult) *s x = x" by vector -lemma vector_ssub_ldistrib: "(c::'a::ring) *s (x - y) = c *s x - c *s y" - by (vector ring_simps) -lemma vector_smult_rneg: "(c::'a::ring) *s -x = -(c *s x)" by vector -lemma vector_smult_lneg: "- (c::'a::ring) *s x = -(c *s x)" by vector -lemma vector_sneg_minus1: "-x = (- (1::'a::ring_1)) *s x" by vector -lemma vector_smult_rzero[simp]: "c *s 0 = (0::'a::mult_zero ^ 'n)" by vector -lemma vector_sub_rdistrib: "((a::'a::ring) - b) *s x = a *s x - b *s x" - by (vector ring_simps) - -lemma vec_eq[simp]: "(vec m = vec n) \ (m = n)" - by (simp add: Cart_eq) - -subsection {* Topological space *} - -instantiation "^" :: (topological_space, finite) topological_space -begin - -definition open_vector_def: - "open (S :: ('a ^ 'b) set) \ - (\x\S. \A. (\i. open (A i) \ x$i \ A i) \ - (\y. (\i. y$i \ A i) \ y \ S))" - -instance proof - show "open (UNIV :: ('a ^ 'b) set)" - unfolding open_vector_def by auto -next - fix S T :: "('a ^ 'b) set" - assume "open S" "open T" thus "open (S \ T)" - unfolding open_vector_def - apply clarify - apply (drule (1) bspec)+ - apply (clarify, rename_tac Sa Ta) - apply (rule_tac x="\i. Sa i \ Ta i" in exI) - apply (simp add: open_Int) - done -next - fix K :: "('a ^ 'b) set set" - assume "\S\K. open S" thus "open (\K)" - unfolding open_vector_def - apply clarify - apply (drule (1) bspec) - apply (drule (1) bspec) - apply clarify - apply (rule_tac x=A in exI) - apply fast - done -qed - -end - -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 assume "open S" "a $ i \ S" - then have "open ((\y. y $ i) -` S)" "a \ ((\y. y $ i) -` S)" - 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" - by simp -qed - -subsection {* Square root of sum of squares *} - -definition - "setL2 f A = sqrt (\i\A. (f i)\)" - -lemma setL2_cong: - "\A = B; \x. x \ B \ f x = g x\ \ setL2 f A = setL2 g B" - unfolding setL2_def by simp - -lemma strong_setL2_cong: - "\A = B; \x. x \ B =simp=> f x = g x\ \ setL2 f A = setL2 g B" - unfolding setL2_def simp_implies_def by simp - -lemma setL2_infinite [simp]: "\ finite A \ setL2 f A = 0" - unfolding setL2_def by simp - -lemma setL2_empty [simp]: "setL2 f {} = 0" - unfolding setL2_def by simp - -lemma setL2_insert [simp]: - "\finite F; a \ F\ \ - setL2 f (insert a F) = sqrt ((f a)\ + (setL2 f F)\)" - unfolding setL2_def by (simp add: setsum_nonneg) - -lemma setL2_nonneg [simp]: "0 \ setL2 f A" - unfolding setL2_def by (simp add: setsum_nonneg) - -lemma setL2_0': "\a\A. f a = 0 \ setL2 f A = 0" - unfolding setL2_def by simp - -lemma setL2_constant: "setL2 (\x. y) A = sqrt (of_nat (card A)) * \y\" - unfolding setL2_def by (simp add: real_sqrt_mult) - -lemma setL2_mono: - assumes "\i. i \ K \ f i \ g i" - assumes "\i. i \ K \ 0 \ f i" - shows "setL2 f K \ setL2 g K" - unfolding setL2_def - by (simp add: setsum_nonneg setsum_mono power_mono prems) - -lemma setL2_strict_mono: - assumes "finite K" and "K \ {}" - assumes "\i. i \ K \ f i < g i" - assumes "\i. i \ K \ 0 \ f i" - shows "setL2 f K < setL2 g K" - unfolding setL2_def - by (simp add: setsum_strict_mono power_strict_mono assms) - -lemma setL2_right_distrib: - "0 \ r \ r * setL2 f A = setL2 (\x. r * f x) A" - unfolding setL2_def - apply (simp add: power_mult_distrib) - apply (simp add: setsum_right_distrib [symmetric]) - apply (simp add: real_sqrt_mult setsum_nonneg) - done - -lemma setL2_left_distrib: - "0 \ r \ setL2 f A * r = setL2 (\x. f x * r) A" - unfolding setL2_def - apply (simp add: power_mult_distrib) - apply (simp add: setsum_left_distrib [symmetric]) - apply (simp add: real_sqrt_mult setsum_nonneg) - done - -lemma setsum_nonneg_eq_0_iff: - fixes f :: "'a \ 'b::pordered_ab_group_add" - shows "\finite A; \x\A. 0 \ f x\ \ setsum f A = 0 \ (\x\A. f x = 0)" - apply (induct set: finite, simp) - apply (simp add: add_nonneg_eq_0_iff setsum_nonneg) - done - -lemma setL2_eq_0_iff: "finite A \ setL2 f A = 0 \ (\x\A. f x = 0)" - unfolding setL2_def - by (simp add: setsum_nonneg setsum_nonneg_eq_0_iff) - -lemma setL2_triangle_ineq: - shows "setL2 (\i. f i + g i) A \ setL2 f A + setL2 g A" -proof (cases "finite A") - case False - thus ?thesis by simp -next - case True - thus ?thesis - proof (induct set: finite) - case empty - show ?case by simp - next - case (insert x F) - hence "sqrt ((f x + g x)\ + (setL2 (\i. f i + g i) F)\) \ - sqrt ((f x + g x)\ + (setL2 f F + setL2 g F)\)" - by (intro real_sqrt_le_mono add_left_mono power_mono insert - setL2_nonneg add_increasing zero_le_power2) - also have - "\ \ sqrt ((f x)\ + (setL2 f F)\) + sqrt ((g x)\ + (setL2 g F)\)" - by (rule real_sqrt_sum_squares_triangle_ineq) - finally show ?case - using insert by simp - qed -qed - -lemma sqrt_sum_squares_le_sum: - "\0 \ x; 0 \ y\ \ sqrt (x\ + y\) \ x + y" - apply (rule power2_le_imp_le) - apply (simp add: power2_sum) - apply (simp add: mult_nonneg_nonneg) - apply (simp add: add_nonneg_nonneg) - done - -lemma setL2_le_setsum [rule_format]: - "(\i\A. 0 \ f i) \ setL2 f A \ setsum f A" - apply (cases "finite A") - apply (induct set: finite) - apply simp - apply clarsimp - apply (erule order_trans [OF sqrt_sum_squares_le_sum]) - apply simp - apply simp - apply simp - done - -lemma sqrt_sum_squares_le_sum_abs: "sqrt (x\ + y\) \ \x\ + \y\" - apply (rule power2_le_imp_le) - apply (simp add: power2_sum) - apply (simp add: mult_nonneg_nonneg) - apply (simp add: add_nonneg_nonneg) - done - -lemma setL2_le_setsum_abs: "setL2 f A \ (\i\A. \f i\)" - apply (cases "finite A") - apply (induct set: finite) - apply simp - apply simp - apply (rule order_trans [OF sqrt_sum_squares_le_sum_abs]) - apply simp - apply simp - done - -lemma setL2_mult_ineq_lemma: - fixes a b c d :: real - shows "2 * (a * c) * (b * d) \ a\ * d\ + b\ * c\" -proof - - have "0 \ (a * d - b * c)\" by simp - also have "\ = a\ * d\ + b\ * c\ - 2 * (a * d) * (b * c)" - by (simp only: power2_diff power_mult_distrib) - also have "\ = a\ * d\ + b\ * c\ - 2 * (a * c) * (b * d)" - by simp - finally show "2 * (a * c) * (b * d) \ a\ * d\ + b\ * c\" - by simp -qed - -lemma setL2_mult_ineq: "(\i\A. \f i\ * \g i\) \ setL2 f A * setL2 g A" - apply (cases "finite A") - apply (induct set: finite) - apply simp - apply (rule power2_le_imp_le, simp) - apply (rule order_trans) - apply (rule power_mono) - apply (erule add_left_mono) - apply (simp add: add_nonneg_nonneg mult_nonneg_nonneg setsum_nonneg) - apply (simp add: power2_sum) - apply (simp add: power_mult_distrib) - apply (simp add: right_distrib left_distrib) - apply (rule ord_le_eq_trans) - apply (rule setL2_mult_ineq_lemma) - apply simp - apply (intro mult_nonneg_nonneg setL2_nonneg) - apply simp - done - -lemma member_le_setL2: "\finite A; i \ A\ \ f i \ setL2 f A" - apply (rule_tac s="insert i (A - {i})" and t="A" in subst) - apply fast - apply (subst setL2_insert) - apply simp - apply simp - apply simp - done - -subsection {* Metric *} - -(* TODO: move somewhere else *) -lemma finite_choice: "finite A \ \x\A. \y. P x y \ \f. \x\A. P x (f x)" -apply (induct set: finite, simp_all) -apply (clarify, rename_tac y) -apply (rule_tac x="f(x:=y)" in exI, simp) -done - -instantiation "^" :: (metric_space, finite) metric_space -begin - -definition dist_vector_def: - "dist (x::'a^'b) (y::'a^'b) = setL2 (\i. dist (x$i) (y$i)) UNIV" - -lemma dist_nth_le: "dist (x $ i) (y $ i) \ dist x y" -unfolding dist_vector_def -by (rule member_le_setL2) simp_all - -instance proof - fix x y :: "'a ^ 'b" - show "dist x y = 0 \ x = y" - unfolding dist_vector_def - by (simp add: setL2_eq_0_iff Cart_eq) -next - fix x y z :: "'a ^ 'b" - show "dist x y \ dist x z + dist y z" - unfolding dist_vector_def - apply (rule order_trans [OF _ setL2_triangle_ineq]) - apply (simp add: setL2_mono dist_triangle2) - done -next - (* FIXME: long proof! *) - fix S :: "('a ^ 'b) set" - show "open S \ (\x\S. \e>0. \y. dist y x < e \ y \ S)" - unfolding open_vector_def open_dist - apply safe - apply (drule (1) bspec) - apply clarify - apply (subgoal_tac "\e>0. \i y. dist y (x$i) < e \ y \ A i") - apply clarify - apply (rule_tac x=e in exI, clarify) - apply (drule spec, erule mp, clarify) - apply (drule spec, drule spec, erule mp) - apply (erule le_less_trans [OF dist_nth_le]) - apply (subgoal_tac "\i\UNIV. \e>0. \y. dist y (x$i) < e \ y \ A i") - apply (drule finite_choice [OF finite], clarify) - apply (rule_tac x="Min (range f)" in exI, simp) - apply clarify - apply (drule_tac x=i in spec, clarify) - apply (erule (1) bspec) - apply (drule (1) bspec, clarify) - apply (subgoal_tac "\r. (\i::'b. 0 < r i) \ e = setL2 r UNIV") - apply clarify - apply (rule_tac x="\i. {y. dist y (x$i) < r i}" in exI) - apply (rule conjI) - apply clarify - apply (rule conjI) - apply (clarify, rename_tac y) - apply (rule_tac x="r i - dist y (x$i)" in exI, rule conjI, simp) - apply clarify - apply (simp only: less_diff_eq) - apply (erule le_less_trans [OF dist_triangle]) - apply simp - apply clarify - apply (drule spec, erule mp) - apply (simp add: dist_vector_def setL2_strict_mono) - apply (rule_tac x="\i. e / sqrt (of_nat CARD('b))" in exI) - apply (simp add: divide_pos_pos setL2_constant) - done -qed - -end - -lemma LIMSEQ_Cart_nth: - "(X ----> a) \ (\n. X n $ i) ----> a $ i" -unfolding LIMSEQ_conv_tendsto by (rule tendsto_Cart_nth) - -lemma LIM_Cart_nth: - "(f -- x --> y) \ (\x. f x $ i) -- x --> y $ i" -unfolding LIM_conv_tendsto by (rule tendsto_Cart_nth) - -lemma Cauchy_Cart_nth: - "Cauchy (\n. X n) \ Cauchy (\n. X n $ i)" -unfolding Cauchy_def by (fast intro: le_less_trans [OF dist_nth_le]) - -lemma LIMSEQ_vector: - fixes X :: "nat \ 'a::metric_space ^ 'n::finite" - assumes X: "\i. (\n. X n $ i) ----> (a $ i)" - shows "X ----> a" -proof (rule metric_LIMSEQ_I) - fix r :: real assume "0 < r" - then have "0 < r / of_nat CARD('n)" (is "0 < ?s") - by (simp add: divide_pos_pos) - def N \ "\i. LEAST N. \n\N. dist (X n $ i) (a $ i) < ?s" - def M \ "Max (range N)" - have "\i. \N. \n\N. dist (X n $ i) (a $ i) < ?s" - using X `0 < ?s` by (rule metric_LIMSEQ_D) - hence "\i. \n\N i. dist (X n $ i) (a $ i) < ?s" - unfolding N_def by (rule LeastI_ex) - hence M: "\i. \n\M. dist (X n $ i) (a $ i) < ?s" - unfolding M_def by simp - { - fix n :: nat assume "M \ n" - have "dist (X n) a = setL2 (\i. dist (X n $ i) (a $ i)) UNIV" - unfolding dist_vector_def .. - also have "\ \ setsum (\i. dist (X n $ i) (a $ i)) UNIV" - by (rule setL2_le_setsum [OF zero_le_dist]) - also have "\ < setsum (\i::'n. ?s) UNIV" - by (rule setsum_strict_mono, simp_all add: M `M \ n`) - also have "\ = r" - by simp - finally have "dist (X n) a < r" . - } - hence "\n\M. dist (X n) a < r" - by simp - then show "\M. \n\M. dist (X n) a < r" .. -qed - -lemma Cauchy_vector: - fixes X :: "nat \ 'a::metric_space ^ 'n::finite" - assumes X: "\i. Cauchy (\n. X n $ i)" - shows "Cauchy (\n. X n)" -proof (rule metric_CauchyI) - fix r :: real assume "0 < r" - then have "0 < r / of_nat CARD('n)" (is "0 < ?s") - by (simp add: divide_pos_pos) - def N \ "\i. LEAST N. \m\N. \n\N. dist (X m $ i) (X n $ i) < ?s" - def M \ "Max (range N)" - have "\i. \N. \m\N. \n\N. dist (X m $ i) (X n $ i) < ?s" - using X `0 < ?s` by (rule metric_CauchyD) - hence "\i. \m\N i. \n\N i. dist (X m $ i) (X n $ i) < ?s" - unfolding N_def by (rule LeastI_ex) - hence M: "\i. \m\M. \n\M. dist (X m $ i) (X n $ i) < ?s" - unfolding M_def by simp - { - fix m n :: nat - assume "M \ m" "M \ n" - have "dist (X m) (X n) = setL2 (\i. dist (X m $ i) (X n $ i)) UNIV" - unfolding dist_vector_def .. - also have "\ \ setsum (\i. dist (X m $ i) (X n $ i)) UNIV" - by (rule setL2_le_setsum [OF zero_le_dist]) - also have "\ < setsum (\i::'n. ?s) UNIV" - by (rule setsum_strict_mono, simp_all add: M `M \ m` `M \ n`) - also have "\ = r" - by simp - finally have "dist (X m) (X n) < r" . - } - hence "\m\M. \n\M. dist (X m) (X n) < r" - by simp - then show "\M. \m\M. \n\M. dist (X m) (X n) < r" .. -qed - -instance "^" :: (complete_space, finite) complete_space -proof - fix X :: "nat \ 'a ^ 'b" assume "Cauchy X" - have "\i. (\n. X n $ i) ----> lim (\n. X n $ i)" - using Cauchy_Cart_nth [OF `Cauchy X`] - by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) - hence "X ----> Cart_lambda (\i. lim (\n. X n $ i))" - by (simp add: LIMSEQ_vector) - then show "convergent X" - by (rule convergentI) -qed - -subsection {* Norms *} - -instantiation "^" :: (real_normed_vector, finite) real_normed_vector -begin - -definition norm_vector_def: - "norm (x::'a^'b) = setL2 (\i. norm (x$i)) UNIV" - -definition vector_sgn_def: - "sgn (x::'a^'b) = scaleR (inverse (norm x)) x" - -instance proof - fix a :: real and x y :: "'a ^ 'b" - show "0 \ norm x" - unfolding norm_vector_def - by (rule setL2_nonneg) - show "norm x = 0 \ x = 0" - unfolding norm_vector_def - by (simp add: setL2_eq_0_iff Cart_eq) - show "norm (x + y) \ norm x + norm y" - 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 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 norm_vector_def - by (simp add: dist_norm) -qed - -end - -lemma norm_nth_le: "norm (x $ i) \ norm x" -unfolding norm_vector_def -by (rule member_le_setL2) simp_all - -interpretation Cart_nth: bounded_linear "\x. x $ i" -apply default -apply (rule vector_add_component) -apply (rule vector_scaleR_component) -apply (rule_tac x="1" in exI, simp add: norm_nth_le) -done - -instance "^" :: (banach, finite) banach .. - -subsection {* Inner products *} - -instantiation "^" :: (real_inner, finite) real_inner -begin - -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 inner_vector_def - by (simp add: inner_commute) - show "inner (x + y) z = inner x z + inner y z" - unfolding inner_vector_def - by (simp add: inner_add_left setsum_addf) - show "inner (scaleR r x) y = r * inner x y" - unfolding inner_vector_def - by (simp add: setsum_right_distrib) - show "0 \ inner x x" - unfolding inner_vector_def - by (simp add: setsum_nonneg) - show "inner x x = 0 \ x = 0" - unfolding inner_vector_def - by (simp add: Cart_eq setsum_nonneg_eq_0_iff) - show "norm x = sqrt (inner x x)" - unfolding inner_vector_def norm_vector_def setL2_def - by (simp add: power2_norm_eq_inner) -qed - -end - -subsection{* Properties of the dot product. *} - -lemma dot_sym: "(x::'a:: {comm_monoid_add, ab_semigroup_mult} ^ 'n) \ y = y \ x" - by (vector mult_commute) -lemma dot_ladd: "((x::'a::ring ^ 'n) + y) \ z = (x \ z) + (y \ z)" - by (vector ring_simps) -lemma dot_radd: "x \ (y + (z::'a::ring ^ 'n)) = (x \ y) + (x \ z)" - by (vector ring_simps) -lemma dot_lsub: "((x::'a::ring ^ 'n) - y) \ z = (x \ z) - (y \ z)" - by (vector ring_simps) -lemma dot_rsub: "(x::'a::ring ^ 'n) \ (y - z) = (x \ y) - (x \ z)" - by (vector ring_simps) -lemma dot_lmult: "(c *s x) \ y = (c::'a::ring) * (x \ y)" by (vector ring_simps) -lemma dot_rmult: "x \ (c *s y) = (c::'a::comm_ring) * (x \ y)" by (vector ring_simps) -lemma dot_lneg: "(-x) \ (y::'a::ring ^ 'n) = -(x \ y)" by vector -lemma dot_rneg: "(x::'a::ring ^ 'n) \ (-y) = -(x \ y)" by vector -lemma dot_lzero[simp]: "0 \ x = (0::'a::{comm_monoid_add, mult_zero})" by vector -lemma dot_rzero[simp]: "x \ 0 = (0::'a::{comm_monoid_add, mult_zero})" by vector -lemma dot_pos_le[simp]: "(0::'a\ordered_ring_strict) <= x \ x" - by (simp add: dot_def setsum_nonneg) - -lemma setsum_squares_eq_0_iff: assumes fS: "finite F" and fp: "\x \ F. f x \ (0 ::'a::pordered_ab_group_add)" shows "setsum f F = 0 \ (ALL x:F. f x = 0)" -using fS fp setsum_nonneg[OF fp] -proof (induct set: finite) - case empty thus ?case by simp -next - case (insert x F) - from insert.prems have Fx: "f x \ 0" and Fp: "\ a \ F. f a \ 0" by simp_all - from insert.hyps Fp setsum_nonneg[OF Fp] - have h: "setsum f F = 0 \ (\a \F. f a = 0)" by metis - from add_nonneg_eq_0_iff[OF Fx setsum_nonneg[OF Fp]] insert.hyps(1,2) - show ?case by (simp add: h) -qed - -lemma dot_eq_0: "x \ x = 0 \ (x::'a::{ordered_ring_strict,ring_no_zero_divisors} ^ 'n::finite) = 0" - by (simp add: dot_def setsum_squares_eq_0_iff Cart_eq) - -lemma dot_pos_lt[simp]: "(0 < x \ x) \ (x::'a::{ordered_ring_strict,ring_no_zero_divisors} ^ 'n::finite) \ 0" using dot_eq_0[of x] dot_pos_le[of x] - by (auto simp add: le_less) - -subsection{* The collapse of the general concepts to dimension one. *} - -lemma vector_one: "(x::'a ^1) = (\ i. (x$1))" - by (simp add: Cart_eq forall_1) - -lemma forall_one: "(\(x::'a ^1). P x) \ (\x. P(\ i. x))" - apply auto - apply (erule_tac x= "x$1" in allE) - apply (simp only: vector_one[symmetric]) - done - -lemma norm_vector_1: "norm (x :: _^1) = norm (x$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) - -lemma dist_real: "dist(x::real ^ 1) y = abs((x$1) - (y$1))" - by (auto simp add: norm_real dist_norm) - -subsection {* A connectedness or intermediate value lemma with several applications. *} - -lemma connected_real_lemma: - fixes f :: "real \ 'a::metric_space" - assumes ab: "a \ b" and fa: "f a \ e1" and fb: "f b \ e2" - and dst: "\e x. a <= x \ x <= b \ 0 < e ==> \d > 0. \y. abs(y - x) < d \ dist(f y) (f x) < e" - and e1: "\y \ e1. \e > 0. \y'. dist y' y < e \ y' \ e1" - and e2: "\y \ e2. \e > 0. \y'. dist y' y < e \ y' \ e2" - and e12: "~(\x \ a. x <= b \ f x \ e1 \ f x \ e2)" - shows "\x \ a. x <= b \ f x \ e1 \ f x \ e2" (is "\ x. ?P x") -proof- - let ?S = "{c. \x \ a. x <= c \ f x \ e1}" - have Se: " \x. x \ ?S" apply (rule exI[where x=a]) by (auto simp add: fa) - have Sub: "\y. isUb UNIV ?S y" - apply (rule exI[where x= b]) - using ab fb e12 by (auto simp add: isUb_def setle_def) - from reals_complete[OF Se Sub] obtain l where - l: "isLub UNIV ?S l"by blast - have alb: "a \ l" "l \ b" using l ab fa fb e12 - apply (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def) - by (metis linorder_linear) - have ale1: "\z \ a. z < l \ f z \ e1" using l - apply (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def) - by (metis linorder_linear not_le) - have th1: "\z x e d :: real. z <= x + e \ e < d ==> z < x \ abs(z - x) < d" by arith - have th2: "\e x:: real. 0 < e ==> ~(x + e <= x)" by arith - have th3: "\d::real. d > 0 \ \e > 0. e < d" by dlo - {assume le2: "f l \ e2" - from le2 fa fb e12 alb have la: "l \ a" by metis - hence lap: "l - a > 0" using alb by arith - from e2[rule_format, OF le2] obtain e where - e: "e > 0" "\y. dist y (f l) < e \ y \ e2" by metis - from dst[OF alb e(1)] obtain d where - d: "d > 0" "\y. \y - l\ < d \ dist (f y) (f l) < e" by metis - have "\d'. d' < d \ d' >0 \ l - d' > a" using lap d(1) - apply ferrack by arith - then obtain d' where d': "d' > 0" "d' < d" "l - d' > a" by metis - from d e have th0: "\y. \y - l\ < d \ f y \ e2" by metis - from th0[rule_format, of "l - d'"] d' have "f (l - d') \ e2" by auto - moreover - have "f (l - d') \ e1" using ale1[rule_format, of "l -d'"] d' by auto - ultimately have False using e12 alb d' by auto} - moreover - {assume le1: "f l \ e1" - from le1 fa fb e12 alb have lb: "l \ b" by metis - hence blp: "b - l > 0" using alb by arith - from e1[rule_format, OF le1] obtain e where - e: "e > 0" "\y. dist y (f l) < e \ y \ e1" by metis - from dst[OF alb e(1)] obtain d where - d: "d > 0" "\y. \y - l\ < d \ dist (f y) (f l) < e" by metis - have "\d'. d' < d \ d' >0" using d(1) by dlo - then obtain d' where d': "d' > 0" "d' < d" by metis - from d e have th0: "\y. \y - l\ < d \ f y \ e1" by auto - hence "\y. l \ y \ y \ l + d' \ f y \ e1" using d' by auto - with ale1 have "\y. a \ y \ y \ l + d' \ f y \ e1" by auto - with l d' have False - by (auto simp add: isLub_def isUb_def setle_def setge_def leastP_def) } - ultimately show ?thesis using alb by metis -qed - -text{* One immediately useful corollary is the existence of square roots! --- Should help to get rid of all the development of square-root for reals as a special case @{typ "real^1"} *} - -lemma square_bound_lemma: "(x::real) < (1 + x) * (1 + x)" -proof- - have "(x + 1/2)^2 + 3/4 > 0" using zero_le_power2[of "x+1/2"] by arith - thus ?thesis by (simp add: ring_simps power2_eq_square) -qed - -lemma square_continuous: "0 < (e::real) ==> \d. 0 < d \ (\y. abs(y - x) < d \ abs(y * y - x * x) < e)" - using isCont_power[OF isCont_ident, of 2, unfolded isCont_def LIM_eq, rule_format, of e x] apply (auto simp add: power2_eq_square) - apply (rule_tac x="s" in exI) - apply auto - apply (erule_tac x=y in allE) - apply auto - done - -lemma real_le_lsqrt: "0 <= x \ 0 <= y \ x <= y^2 ==> sqrt x <= y" - using real_sqrt_le_iff[of x "y^2"] by simp - -lemma real_le_rsqrt: "x^2 \ y \ x \ sqrt y" - using real_sqrt_le_mono[of "x^2" y] by simp - -lemma real_less_rsqrt: "x^2 < y \ x < sqrt y" - using real_sqrt_less_mono[of "x^2" y] by simp - -lemma sqrt_even_pow2: assumes n: "even n" - shows "sqrt(2 ^ n) = 2 ^ (n div 2)" -proof- - from n obtain m where m: "n = 2*m" unfolding even_nat_equiv_def2 - by (auto simp add: nat_number) - from m have "sqrt(2 ^ n) = sqrt ((2 ^ m) ^ 2)" - by (simp only: power_mult[symmetric] mult_commute) - then show ?thesis using m by simp -qed - -lemma real_div_sqrt: "0 <= x ==> x / sqrt(x) = sqrt(x)" - apply (cases "x = 0", simp_all) - using sqrt_divide_self_eq[of x] - apply (simp add: inverse_eq_divide real_sqrt_ge_0_iff field_simps) - done - -text{* Hence derive more interesting properties of the norm. *} - -text {* - This type-specific version is only here - to make @{text normarith.ML} happy. -*} -lemma norm_0: "norm (0::real ^ _) = 0" - by (rule norm_zero) - -lemma norm_mul[simp]: "norm(a *s x) = abs(a) * norm x" - 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: norm_vector_def dot_def setL2_def power2_eq_square) -lemma real_vector_norm_def: "norm x = sqrt (x \ x)" - 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) -lemma vector_mul_eq_0[simp]: "(a *s x = 0) \ a = (0::'a::idom) \ x = 0" - by vector -lemma vector_mul_lcancel[simp]: "a *s x = a *s y \ a = (0::real) \ x = y" - by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_ssub_ldistrib) -lemma vector_mul_rcancel[simp]: "a *s x = b *s x \ (a::real) = b \ x = 0" - by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_sub_rdistrib) -lemma vector_mul_lcancel_imp: "a \ (0::real) ==> a *s x = a *s y ==> (x = y)" - by (metis vector_mul_lcancel) -lemma vector_mul_rcancel_imp: "x \ 0 \ (a::real) *s x = b *s x ==> a = b" - by (metis vector_mul_rcancel) -lemma norm_cauchy_schwarz: - fixes x y :: "real ^ 'n::finite" - shows "x \ y <= norm x * norm y" -proof- - {assume "norm x = 0" - hence ?thesis by (simp add: dot_lzero dot_rzero)} - moreover - {assume "norm y = 0" - hence ?thesis by (simp add: dot_lzero dot_rzero)} - moreover - {assume h: "norm x \ 0" "norm y \ 0" - let ?z = "norm y *s x - norm x *s y" - from h have p: "norm x * norm y > 0" by (metis norm_ge_zero le_less zero_compare_simps) - from dot_pos_le[of ?z] - have "(norm x * norm y) * (x \ y) \ norm x ^2 * norm y ^2" - apply (simp add: dot_rsub dot_lsub dot_lmult dot_rmult ring_simps) - by (simp add: norm_pow_2[symmetric] power2_eq_square dot_sym) - hence "x\y \ (norm x ^2 * norm y ^2) / (norm x * norm y)" using p - by (simp add: field_simps) - hence ?thesis using h by (simp add: power2_eq_square)} - ultimately show ?thesis by metis -qed - -lemma norm_cauchy_schwarz_abs: - fixes x y :: "real ^ 'n::finite" - shows "\x \ y\ \ norm x * norm y" - using norm_cauchy_schwarz[of x y] norm_cauchy_schwarz[of x "-y"] - by (simp add: real_abs_def dot_rneg) - -lemma norm_triangle_sub: - fixes x y :: "'a::real_normed_vector" - shows "norm x \ norm y + norm (x - y)" - using norm_triangle_ineq[of "y" "x - y"] by (simp add: ring_simps) - -lemma norm_triangle_le: "norm(x::real ^'n::finite) + norm y <= e ==> norm(x + y) <= e" - by (metis order_trans norm_triangle_ineq) -lemma norm_triangle_lt: "norm(x::real ^'n::finite) + norm(y) < e ==> norm(x + y) < e" - by (metis basic_trans_rules(21) norm_triangle_ineq) - -lemma component_le_norm: "\x$i\ <= norm (x::real ^ 'n::finite)" - apply (simp add: norm_vector_def) - apply (rule member_le_setL2, simp_all) - done - -lemma norm_bound_component_le: "norm(x::real ^ 'n::finite) <= e - ==> \x$i\ <= e" - by (metis component_le_norm order_trans) - -lemma norm_bound_component_lt: "norm(x::real ^ 'n::finite) < e - ==> \x$i\ < e" - 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: norm_vector_def setL2_le_setsum) - -lemma real_abs_norm: "\norm x\ = norm (x :: real ^ _)" - by (rule abs_norm_cancel) -lemma real_abs_sub_norm: "\norm(x::real ^'n::finite) - norm y\ <= norm(x - y)" - by (rule norm_triangle_ineq3) -lemma norm_le: "norm(x::real ^ _) <= norm(y) \ x \ x <= y \ y" - by (simp add: real_vector_norm_def) -lemma norm_lt: "norm(x::real ^ _) < norm(y) \ x \ x < y \ y" - by (simp add: real_vector_norm_def) -lemma norm_eq: "norm (x::real ^ _) = norm y \ x \ x = y \ y" - by (simp add: order_eq_iff norm_le) -lemma norm_eq_1: "norm(x::real ^ _) = 1 \ x \ x = 1" - by (simp add: real_vector_norm_def) - -text{* Squaring equations and inequalities involving norms. *} - -lemma dot_square_norm: "x \ x = norm(x)^2" - by (simp add: real_vector_norm_def) - -lemma norm_eq_square: "norm(x) = a \ 0 <= a \ x \ x = a^2" - by (auto simp add: real_vector_norm_def) - -lemma real_abs_le_square_iff: "\x\ \ \y\ \ (x::real)^2 \ y^2" -proof- - have "x^2 \ y^2 \ (x -y) * (y + x) \ 0" by (simp add: ring_simps power2_eq_square) - also have "\ \ \x\ \ \y\" apply (simp add: zero_compare_simps real_abs_def not_less) by arith -finally show ?thesis .. -qed - -lemma norm_le_square: "norm(x) <= a \ 0 <= a \ x \ x <= a^2" - apply (simp add: dot_square_norm real_abs_le_square_iff[symmetric]) - using norm_ge_zero[of x] - apply arith - done - -lemma norm_ge_square: "norm(x) >= a \ a <= 0 \ x \ x >= a ^ 2" - apply (simp add: dot_square_norm real_abs_le_square_iff[symmetric]) - using norm_ge_zero[of x] - apply arith - done - -lemma norm_lt_square: "norm(x) < a \ 0 < a \ x \ x < a^2" - by (metis not_le norm_ge_square) -lemma norm_gt_square: "norm(x) > a \ a < 0 \ x \ x > a^2" - by (metis norm_le_square not_less) - -text{* Dot product in terms of the norm rather than conversely. *} - -lemma dot_norm: "x \ y = (norm(x + y) ^2 - norm x ^ 2 - norm y ^ 2) / 2" - by (simp add: norm_pow_2 dot_ladd dot_radd dot_sym) - -lemma dot_norm_neg: "x \ y = ((norm x ^ 2 + norm y ^ 2) - norm(x - y) ^ 2) / 2" - by (simp add: norm_pow_2 dot_ladd dot_radd dot_lsub dot_rsub dot_sym) - - -text{* Equality of vectors in terms of @{term "op \"} products. *} - -lemma vector_eq: "(x:: real ^ 'n::finite) = y \ x \ x = x \ y\ y \ y = x \ x" (is "?lhs \ ?rhs") -proof - assume "?lhs" then show ?rhs by simp -next - assume ?rhs - then have "x \ x - x \ y = 0 \ x \ y - y\ y = 0" by simp - hence "x \ (x - y) = 0 \ y \ (x - y) = 0" - by (simp add: dot_rsub dot_lsub dot_sym) - then have "(x - y) \ (x - y) = 0" by (simp add: ring_simps dot_lsub dot_rsub) - then show "x = y" by (simp add: dot_eq_0) -qed - - -subsection{* General linear decision procedure for normed spaces. *} - -lemma norm_cmul_rule_thm: - fixes x :: "'a::real_normed_vector" - shows "b >= norm(x) ==> \c\ * b >= norm(scaleR c x)" - unfolding norm_scaleR - apply (erule mult_mono1) - apply simp - done - - (* FIXME: Move all these theorems into the ML code using lemma antiquotation *) -lemma norm_add_rule_thm: - fixes x1 x2 :: "'a::real_normed_vector" - shows "norm x1 \ b1 \ norm x2 \ b2 \ norm (x1 + x2) \ b1 + b2" - by (rule order_trans [OF norm_triangle_ineq add_mono]) - -lemma ge_iff_diff_ge_0: "(a::'a::ordered_ring) \ b == a - b \ 0" - by (simp add: ring_simps) - -lemma pth_1: - fixes x :: "'a::real_normed_vector" - shows "x == scaleR 1 x" by simp - -lemma pth_2: - fixes x :: "'a::real_normed_vector" - shows "x - y == x + -y" by (atomize (full)) simp - -lemma pth_3: - fixes x :: "'a::real_normed_vector" - shows "- x == scaleR (-1) x" by simp - -lemma pth_4: - fixes x :: "'a::real_normed_vector" - shows "scaleR 0 x == 0" and "scaleR c 0 = (0::'a)" by simp_all - -lemma pth_5: - fixes x :: "'a::real_normed_vector" - shows "scaleR c (scaleR d x) == scaleR (c * d) x" by simp - -lemma pth_6: - fixes x :: "'a::real_normed_vector" - shows "scaleR c (x + y) == scaleR c x + scaleR c y" - by (simp add: scaleR_right_distrib) - -lemma pth_7: - fixes x :: "'a::real_normed_vector" - shows "0 + x == x" and "x + 0 == x" by simp_all - -lemma pth_8: - fixes x :: "'a::real_normed_vector" - shows "scaleR c x + scaleR d x == scaleR (c + d) x" - by (simp add: scaleR_left_distrib) - -lemma pth_9: - fixes x :: "'a::real_normed_vector" shows - "(scaleR c x + z) + scaleR d x == scaleR (c + d) x + z" - "scaleR c x + (scaleR d x + z) == scaleR (c + d) x + z" - "(scaleR c x + w) + (scaleR d x + z) == scaleR (c + d) x + (w + z)" - by (simp_all add: algebra_simps) - -lemma pth_a: - fixes x :: "'a::real_normed_vector" - shows "scaleR 0 x + y == y" by simp - -lemma pth_b: - fixes x :: "'a::real_normed_vector" shows - "scaleR c x + scaleR d y == scaleR c x + scaleR d y" - "(scaleR c x + z) + scaleR d y == scaleR c x + (z + scaleR d y)" - "scaleR c x + (scaleR d y + z) == scaleR c x + (scaleR d y + z)" - "(scaleR c x + w) + (scaleR d y + z) == scaleR c x + (w + (scaleR d y + z))" - by (simp_all add: algebra_simps) - -lemma pth_c: - fixes x :: "'a::real_normed_vector" shows - "scaleR c x + scaleR d y == scaleR d y + scaleR c x" - "(scaleR c x + z) + scaleR d y == scaleR d y + (scaleR c x + z)" - "scaleR c x + (scaleR d y + z) == scaleR d y + (scaleR c x + z)" - "(scaleR c x + w) + (scaleR d y + z) == scaleR d y + ((scaleR c x + w) + z)" - by (simp_all add: algebra_simps) - -lemma pth_d: - fixes x :: "'a::real_normed_vector" - shows "x + 0 == x" by simp - -lemma norm_imp_pos_and_ge: - fixes x :: "'a::real_normed_vector" - shows "norm x == n \ norm x \ 0 \ n \ norm x" - by atomize auto - -lemma real_eq_0_iff_le_ge_0: "(x::real) = 0 == x \ 0 \ -x \ 0" by arith - -lemma norm_pths: - fixes x :: "'a::real_normed_vector" shows - "x = y \ norm (x - y) \ 0" - "x \ y \ \ (norm (x - y) \ 0)" - using norm_ge_zero[of "x - y"] by auto - -lemma vector_dist_norm: - fixes x :: "'a::real_normed_vector" - shows "dist x y = norm (x - y)" - by (rule dist_norm) - -use "normarith.ML" - -method_setup norm = {* Scan.succeed (SIMPLE_METHOD' o NormArith.norm_arith_tac) -*} "Proves simple linear statements about vector norms" - - -text{* Hence more metric properties. *} - -lemma dist_triangle_alt: - fixes x y z :: "'a::metric_space" - shows "dist y z <= dist x y + dist x z" -using dist_triangle [of y z x] by (simp add: dist_commute) - -lemma dist_pos_lt: - fixes x y :: "'a::metric_space" - shows "x \ y ==> 0 < dist x y" -by (simp add: zero_less_dist_iff) - -lemma dist_nz: - fixes x y :: "'a::metric_space" - shows "x \ y \ 0 < dist x y" -by (simp add: zero_less_dist_iff) - -lemma dist_triangle_le: - fixes x y z :: "'a::metric_space" - shows "dist x z + dist y z <= e \ dist x y <= e" -by (rule order_trans [OF dist_triangle2]) - -lemma dist_triangle_lt: - fixes x y z :: "'a::metric_space" - shows "dist x z + dist y z < e ==> dist x y < e" -by (rule le_less_trans [OF dist_triangle2]) - -lemma dist_triangle_half_l: - fixes x1 x2 y :: "'a::metric_space" - shows "dist x1 y < e / 2 \ dist x2 y < e / 2 \ dist x1 x2 < e" -by (rule dist_triangle_lt [where z=y], simp) - -lemma dist_triangle_half_r: - fixes x1 x2 y :: "'a::metric_space" - shows "dist y x1 < e / 2 \ dist y x2 < e / 2 \ dist x1 x2 < e" -by (rule dist_triangle_half_l, simp_all add: dist_commute) - -lemma dist_triangle_add: - fixes x y x' y' :: "'a::real_normed_vector" - shows "dist (x + y) (x' + y') <= dist x x' + dist y y'" - by norm - -lemma dist_mul[simp]: "dist (c *s x) (c *s y) = \c\ * dist x y" - unfolding dist_norm vector_ssub_ldistrib[symmetric] norm_mul .. - -lemma dist_triangle_add_half: - fixes x x' y y' :: "'a::real_normed_vector" - shows "dist x x' < e / 2 \ dist y y' < e / 2 \ dist(x + y) (x' + y') < e" - by norm - -lemma setsum_component [simp]: - fixes f:: " 'a \ ('b::comm_monoid_add) ^'n" - shows "(setsum f S)$i = setsum (\x. (f x)$i) S" - by (cases "finite S", induct S set: finite, simp_all) - -lemma setsum_eq: "setsum f S = (\ i. setsum (\x. (f x)$i ) S)" - by (simp add: Cart_eq) - -lemma setsum_clauses: - shows "setsum f {} = 0" - and "finite S \ setsum f (insert x S) = - (if x \ S then setsum f S else f x + setsum f S)" - by (auto simp add: insert_absorb) - -lemma setsum_cmul: - fixes f:: "'c \ ('a::semiring_1)^'n" - shows "setsum (\x. c *s f x) S = c *s setsum f S" - by (simp add: Cart_eq setsum_right_distrib) - -lemma setsum_norm: - fixes f :: "'a \ 'b::real_normed_vector" - assumes fS: "finite S" - shows "norm (setsum f S) <= setsum (\x. norm(f x)) S" -proof(induct rule: finite_induct[OF fS]) - case 1 thus ?case by simp -next - case (2 x S) - from "2.hyps" have "norm (setsum f (insert x S)) \ norm (f x) + norm (setsum f S)" by (simp add: norm_triangle_ineq) - also have "\ \ norm (f x) + setsum (\x. norm(f x)) S" - using "2.hyps" by simp - finally show ?case using "2.hyps" by simp -qed - -lemma real_setsum_norm: - fixes f :: "'a \ real ^'n::finite" - assumes fS: "finite S" - shows "norm (setsum f S) <= setsum (\x. norm(f x)) S" -proof(induct rule: finite_induct[OF fS]) - case 1 thus ?case by simp -next - case (2 x S) - from "2.hyps" have "norm (setsum f (insert x S)) \ norm (f x) + norm (setsum f S)" by (simp add: norm_triangle_ineq) - also have "\ \ norm (f x) + setsum (\x. norm(f x)) S" - using "2.hyps" by simp - finally show ?case using "2.hyps" by simp -qed - -lemma setsum_norm_le: - fixes f :: "'a \ 'b::real_normed_vector" - assumes fS: "finite S" - and fg: "\x \ S. norm (f x) \ g x" - shows "norm (setsum f S) \ setsum g S" -proof- - from fg have "setsum (\x. norm(f x)) S <= setsum g S" - by - (rule setsum_mono, simp) - then show ?thesis using setsum_norm[OF fS, of f] fg - by arith -qed - -lemma real_setsum_norm_le: - fixes f :: "'a \ real ^ 'n::finite" - assumes fS: "finite S" - and fg: "\x \ S. norm (f x) \ g x" - shows "norm (setsum f S) \ setsum g S" -proof- - from fg have "setsum (\x. norm(f x)) S <= setsum g S" - by - (rule setsum_mono, simp) - then show ?thesis using real_setsum_norm[OF fS, of f] fg - by arith -qed - -lemma setsum_norm_bound: - fixes f :: "'a \ 'b::real_normed_vector" - assumes fS: "finite S" - and K: "\x \ S. norm (f x) \ K" - shows "norm (setsum f S) \ of_nat (card S) * K" - using setsum_norm_le[OF fS K] setsum_constant[symmetric] - by simp - -lemma real_setsum_norm_bound: - fixes f :: "'a \ real ^ 'n::finite" - assumes fS: "finite S" - and K: "\x \ S. norm (f x) \ K" - shows "norm (setsum f S) \ of_nat (card S) * K" - using real_setsum_norm_le[OF fS K] setsum_constant[symmetric] - by simp - -lemma setsum_vmul: - fixes f :: "'a \ 'b::{real_normed_vector,semiring, mult_zero}" - assumes fS: "finite S" - shows "setsum f S *s v = setsum (\x. f x *s v) S" -proof(induct rule: finite_induct[OF fS]) - case 1 then show ?case by (simp add: vector_smult_lzero) -next - case (2 x F) - from "2.hyps" have "setsum f (insert x F) *s v = (f x + setsum f F) *s v" - by simp - also have "\ = f x *s v + setsum f F *s v" - by (simp add: vector_sadd_rdistrib) - also have "\ = setsum (\x. f x *s v) (insert x F)" using "2.hyps" by simp - finally show ?case . -qed - -(* FIXME : Problem thm setsum_vmul[of _ "f:: 'a \ real ^'n"] --- - Get rid of *s and use real_vector instead! Also prove that ^ creates a real_vector !! *) - - (* FIXME: Here too need stupid finiteness assumption on T!!! *) -lemma setsum_group: - assumes fS: "finite S" and fT: "finite T" and fST: "f ` S \ T" - shows "setsum (\y. setsum g {x. x\ S \ f x = y}) T = setsum g S" - -apply (subst setsum_image_gen[OF fS, of g f]) -apply (rule setsum_mono_zero_right[OF fT fST]) -by (auto intro: setsum_0') - -lemma vsum_norm_allsubsets_bound: - fixes f:: "'a \ real ^'n::finite" - assumes fP: "finite P" and fPs: "\Q. Q \ P \ norm (setsum f Q) \ e" - shows "setsum (\x. norm (f x)) P \ 2 * real CARD('n) * e" -proof- - let ?d = "real CARD('n)" - let ?nf = "\x. norm (f x)" - let ?U = "UNIV :: 'n set" - have th0: "setsum (\x. setsum (\i. \f x $ i\) ?U) P = setsum (\i. setsum (\x. \f x $ i\) P) ?U" - by (rule setsum_commute) - have th1: "2 * ?d * e = of_nat (card ?U) * (2 * e)" by (simp add: real_of_nat_def) - have "setsum ?nf P \ setsum (\x. setsum (\i. \f x $ i\) ?U) P" - apply (rule setsum_mono) - by (rule norm_le_l1) - also have "\ \ 2 * ?d * e" - unfolding th0 th1 - proof(rule setsum_bounded) - fix i assume i: "i \ ?U" - let ?Pp = "{x. x\ P \ f x $ i \ 0}" - let ?Pn = "{x. x \ P \ f x $ i < 0}" - have thp: "P = ?Pp \ ?Pn" by auto - have thp0: "?Pp \ ?Pn ={}" by auto - have PpP: "?Pp \ P" and PnP: "?Pn \ P" by blast+ - have Ppe:"setsum (\x. \f x $ i\) ?Pp \ e" - using component_le_norm[of "setsum (\x. f x) ?Pp" i] fPs[OF PpP] - by (auto intro: abs_le_D1) - have Pne: "setsum (\x. \f x $ i\) ?Pn \ e" - using component_le_norm[of "setsum (\x. - f x) ?Pn" i] fPs[OF PnP] - by (auto simp add: setsum_negf intro: abs_le_D1) - have "setsum (\x. \f x $ i\) P = setsum (\x. \f x $ i\) ?Pp + setsum (\x. \f x $ i\) ?Pn" - apply (subst thp) - apply (rule setsum_Un_zero) - using fP thp0 by auto - also have "\ \ 2*e" using Pne Ppe by arith - finally show "setsum (\x. \f x $ i\) P \ 2*e" . - qed - finally show ?thesis . -qed - -lemma dot_lsum: "finite S \ setsum f S \ (y::'a::{comm_ring}^'n) = setsum (\x. f x \ y) S " - by (induct rule: finite_induct, auto simp add: dot_lzero dot_ladd dot_radd) - -lemma dot_rsum: "finite S \ (y::'a::{comm_ring}^'n) \ setsum f S = setsum (\x. y \ f x) S " - by (induct rule: finite_induct, auto simp add: dot_rzero dot_radd) - -subsection{* Basis vectors in coordinate directions. *} - - -definition "basis k = (\ i. if i = k then 1 else 0)" - -lemma basis_component [simp]: "basis k $ i = (if k=i then 1 else 0)" - unfolding basis_def by simp - -lemma delta_mult_idempotent: - "(if k=a then 1 else (0::'a::semiring_1)) * (if k=a then 1 else 0) = (if k=a then 1 else 0)" by (cases "k=a", auto) - -lemma norm_basis: - shows "norm (basis k :: real ^'n::finite) = 1" - apply (simp add: basis_def real_vector_norm_def dot_def) - apply (vector delta_mult_idempotent) - using setsum_delta[of "UNIV :: 'n set" "k" "\k. 1::real"] - apply auto - done - -lemma norm_basis_1: "norm(basis 1 :: real ^'n::{finite,one}) = 1" - by (rule norm_basis) - -lemma vector_choose_size: "0 <= c ==> \(x::real^'n::finite). norm x = c" - apply (rule exI[where x="c *s basis arbitrary"]) - by (simp only: norm_mul norm_basis) - -lemma vector_choose_dist: assumes e: "0 <= e" - shows "\(y::real^'n::finite). dist x y = e" -proof- - from vector_choose_size[OF e] obtain c:: "real ^'n" where "norm c = e" - by blast - then have "dist x (x - c) = e" by (simp add: dist_norm) - then show ?thesis by blast -qed - -lemma basis_inj: "inj (basis :: 'n \ real ^'n::finite)" - by (simp add: inj_on_def Cart_eq) - -lemma cond_value_iff: "f (if b then x else y) = (if b then f x else f y)" - by auto - -lemma basis_expansion: - "setsum (\i. (x$i) *s basis i) UNIV = (x::('a::ring_1) ^'n::finite)" (is "?lhs = ?rhs" is "setsum ?f ?S = _") - by (auto simp add: Cart_eq cond_value_iff setsum_delta[of "?S", where ?'b = "'a", simplified] cong del: if_weak_cong) - -lemma basis_expansion_unique: - "setsum (\i. f i *s basis i) UNIV = (x::('a::comm_ring_1) ^'n::finite) \ (\i. f i = x$i)" - by (simp add: Cart_eq setsum_delta cond_value_iff cong del: if_weak_cong) - -lemma cond_application_beta: "(if b then f else g) x = (if b then f x else g x)" - by auto - -lemma dot_basis: - 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) - -lemma basis_nonzero: - shows "basis k \ (0:: 'a::semiring_1 ^'n)" - by (simp add: basis_eq_0) - -lemma vector_eq_ldot: "(\x. x \ y = x \ z) \ y = (z::'a::semiring_1^'n::finite)" - apply (auto simp add: Cart_eq dot_basis) - apply (erule_tac x="basis i" in allE) - apply (simp add: dot_basis) - apply (subgoal_tac "y = z") - apply simp - apply (simp add: Cart_eq) - done - -lemma vector_eq_rdot: "(\z. x \ z = y \ z) \ x = (y::'a::semiring_1^'n::finite)" - apply (auto simp add: Cart_eq dot_basis) - apply (erule_tac x="basis i" in allE) - apply (simp add: dot_basis) - apply (subgoal_tac "x = y") - apply simp - apply (simp add: Cart_eq) - done - -subsection{* Orthogonality. *} - -definition "orthogonal x y \ (x \ y = 0)" - -lemma orthogonal_basis: - shows "orthogonal (basis i :: 'a^'n::finite) x \ x$i = (0::'a::ring_1)" - by (auto simp add: orthogonal_def dot_def basis_def cond_value_iff cond_application_beta setsum_delta cong del: if_weak_cong) - -lemma orthogonal_basis_basis: - shows "orthogonal (basis i :: 'a::ring_1^'n::finite) (basis j) \ i \ j" - unfolding orthogonal_basis[of i] basis_component[of j] by simp - - (* FIXME : Maybe some of these require less than comm_ring, but not all*) -lemma orthogonal_clauses: - "orthogonal a (0::'a::comm_ring ^'n)" - "orthogonal a x ==> orthogonal a (c *s x)" - "orthogonal a x ==> orthogonal a (-x)" - "orthogonal a x \ orthogonal a y ==> orthogonal a (x + y)" - "orthogonal a x \ orthogonal a y ==> orthogonal a (x - y)" - "orthogonal 0 a" - "orthogonal x a ==> orthogonal (c *s x) a" - "orthogonal x a ==> orthogonal (-x) a" - "orthogonal x a \ orthogonal y a ==> orthogonal (x + y) a" - "orthogonal x a \ orthogonal y a ==> orthogonal (x - y) a" - unfolding orthogonal_def dot_rneg dot_rmult dot_radd dot_rsub - dot_lzero dot_rzero dot_lneg dot_lmult dot_ladd dot_lsub - by simp_all - -lemma orthogonal_commute: "orthogonal (x::'a::{ab_semigroup_mult,comm_monoid_add} ^'n)y \ orthogonal y x" - by (simp add: orthogonal_def dot_sym) - -subsection{* Explicit vector construction from lists. *} - -primrec from_nat :: "nat \ 'a::{monoid_add,one}" -where "from_nat 0 = 0" | "from_nat (Suc n) = 1 + from_nat n" - -lemma from_nat [simp]: "from_nat = of_nat" -by (rule ext, induct_tac x, simp_all) - -primrec - list_fun :: "nat \ _ list \ _ \ _" -where - "list_fun n [] = (\x. 0)" -| "list_fun n (x # xs) = fun_upd (list_fun (Suc n) xs) (from_nat n) x" - -definition "vector l = (\ i. list_fun 1 l i)" -(*definition "vector l = (\ i. if i <= length l then l ! (i - 1) else 0)"*) - -lemma vector_1: "(vector[x]) $1 = x" - unfolding vector_def by simp - -lemma vector_2: - "(vector[x,y]) $1 = x" - "(vector[x,y] :: 'a^2)$2 = (y::'a::zero)" - unfolding vector_def by simp_all - -lemma vector_3: - "(vector [x,y,z] ::('a::zero)^3)$1 = x" - "(vector [x,y,z] ::('a::zero)^3)$2 = y" - "(vector [x,y,z] ::('a::zero)^3)$3 = z" - unfolding vector_def by simp_all - -lemma forall_vector_1: "(\v::'a::zero^1. P v) \ (\x. P(vector[x]))" - apply auto - apply (erule_tac x="v$1" in allE) - apply (subgoal_tac "vector [v$1] = v") - apply simp - apply (vector vector_def) - apply (simp add: forall_1) - done - -lemma forall_vector_2: "(\v::'a::zero^2. P v) \ (\x y. P(vector[x, y]))" - apply auto - apply (erule_tac x="v$1" in allE) - apply (erule_tac x="v$2" in allE) - apply (subgoal_tac "vector [v$1, v$2] = v") - apply simp - apply (vector vector_def) - apply (simp add: forall_2) - done - -lemma forall_vector_3: "(\v::'a::zero^3. P v) \ (\x y z. P(vector[x, y, z]))" - apply auto - apply (erule_tac x="v$1" in allE) - apply (erule_tac x="v$2" in allE) - apply (erule_tac x="v$3" in allE) - apply (subgoal_tac "vector [v$1, v$2, v$3] = v") - apply simp - apply (vector vector_def) - apply (simp add: forall_3) - done - -subsection{* Linear functions. *} - -definition "linear f \ (\x y. f(x + y) = f x + f y) \ (\c x. f(c *s x) = c *s f x)" - -lemma linear_compose_cmul: "linear f ==> linear (\x. (c::'a::comm_semiring) *s f x)" - by (vector linear_def Cart_eq ring_simps) - -lemma linear_compose_neg: "linear (f :: 'a ^'n \ 'a::comm_ring ^'m) ==> linear (\x. -(f(x)))" by (vector linear_def Cart_eq) - -lemma linear_compose_add: "linear (f :: 'a ^'n \ 'a::semiring_1 ^'m) \ linear g ==> linear (\x. f(x) + g(x))" - by (vector linear_def Cart_eq ring_simps) - -lemma linear_compose_sub: "linear (f :: 'a ^'n \ 'a::ring_1 ^'m) \ linear g ==> linear (\x. f x - g x)" - by (vector linear_def Cart_eq ring_simps) - -lemma linear_compose: "linear f \ linear g ==> linear (g o f)" - by (simp add: linear_def) - -lemma linear_id: "linear id" by (simp add: linear_def id_def) - -lemma linear_zero: "linear (\x. 0::'a::semiring_1 ^ 'n)" by (simp add: linear_def) - -lemma linear_compose_setsum: - assumes fS: "finite S" and lS: "\a \ S. linear (f a :: 'a::semiring_1 ^ 'n \ 'a ^ 'm)" - shows "linear(\x. setsum (\a. f a x :: 'a::semiring_1 ^'m) S)" - using lS - apply (induct rule: finite_induct[OF fS]) - by (auto simp add: linear_zero intro: linear_compose_add) - -lemma linear_vmul_component: - fixes f:: "'a::semiring_1^'m \ 'a^'n" - assumes lf: "linear f" - shows "linear (\x. f x $ k *s v)" - using lf - apply (auto simp add: linear_def ) - by (vector ring_simps)+ - -lemma linear_0: "linear f ==> f 0 = (0::'a::semiring_1 ^'n)" - unfolding linear_def - apply clarsimp - apply (erule allE[where x="0::'a"]) - apply simp - done - -lemma linear_cmul: "linear f ==> f(c*s x) = c *s f x" by (simp add: linear_def) - -lemma linear_neg: "linear (f :: 'a::ring_1 ^'n \ _) ==> f (-x) = - f x" - unfolding vector_sneg_minus1 - using linear_cmul[of f] by auto - -lemma linear_add: "linear f ==> f(x + y) = f x + f y" by (metis linear_def) - -lemma linear_sub: "linear (f::'a::ring_1 ^'n \ _) ==> f(x - y) = f x - f y" - by (simp add: diff_def linear_add linear_neg) - -lemma linear_setsum: - fixes f:: "'a::semiring_1^'n \ _" - assumes lf: "linear f" and fS: "finite S" - shows "f (setsum g S) = setsum (f o g) S" -proof (induct rule: finite_induct[OF fS]) - case 1 thus ?case by (simp add: linear_0[OF lf]) -next - case (2 x F) - have "f (setsum g (insert x F)) = f (g x + setsum g F)" using "2.hyps" - by simp - also have "\ = f (g x) + f (setsum g F)" using linear_add[OF lf] by simp - also have "\ = setsum (f o g) (insert x F)" using "2.hyps" by simp - finally show ?case . -qed - -lemma linear_setsum_mul: - fixes f:: "'a ^'n \ 'a::semiring_1^'m" - assumes lf: "linear f" and fS: "finite S" - shows "f (setsum (\i. c i *s v i) S) = setsum (\i. c i *s f (v i)) S" - using linear_setsum[OF lf fS, of "\i. c i *s v i" , unfolded o_def] - linear_cmul[OF lf] by simp - -lemma linear_injective_0: - assumes lf: "linear (f:: 'a::ring_1 ^ 'n \ _)" - shows "inj f \ (\x. f x = 0 \ x = 0)" -proof- - have "inj f \ (\ x y. f x = f y \ x = y)" by (simp add: inj_on_def) - also have "\ \ (\ x y. f x - f y = 0 \ x - y = 0)" by simp - also have "\ \ (\ x y. f (x - y) = 0 \ x - y = 0)" - by (simp add: linear_sub[OF lf]) - also have "\ \ (\ x. f x = 0 \ x = 0)" by auto - finally show ?thesis . -qed - -lemma linear_bounded: - fixes f:: "real ^'m::finite \ real ^'n::finite" - assumes lf: "linear f" - shows "\B. \x. norm (f x) \ B * norm x" -proof- - let ?S = "UNIV:: 'm set" - let ?B = "setsum (\i. norm(f(basis i))) ?S" - have fS: "finite ?S" by simp - {fix x:: "real ^ 'm" - let ?g = "(\i. (x$i) *s (basis i) :: real ^ 'm)" - have "norm (f x) = norm (f (setsum (\i. (x$i) *s (basis i)) ?S))" - by (simp only: basis_expansion) - also have "\ = norm (setsum (\i. (x$i) *s f (basis i))?S)" - using linear_setsum[OF lf fS, of ?g, unfolded o_def] linear_cmul[OF lf] - by auto - finally have th0: "norm (f x) = norm (setsum (\i. (x$i) *s f (basis i))?S)" . - {fix i assume i: "i \ ?S" - from component_le_norm[of x i] - have "norm ((x$i) *s f (basis i :: real ^'m)) \ norm (f (basis i)) * norm x" - unfolding norm_mul - apply (simp only: mult_commute) - apply (rule mult_mono) - by (auto simp add: ring_simps norm_ge_zero) } - then have th: "\i\ ?S. norm ((x$i) *s f (basis i :: real ^'m)) \ norm (f (basis i)) * norm x" by metis - from real_setsum_norm_le[OF fS, of "\i. (x$i) *s (f (basis i))", OF th] - have "norm (f x) \ ?B * norm x" unfolding th0 setsum_left_distrib by metis} - then show ?thesis by blast -qed - -lemma linear_bounded_pos: - fixes f:: "real ^'n::finite \ real ^ 'm::finite" - assumes lf: "linear f" - shows "\B > 0. \x. norm (f x) \ B * norm x" -proof- - from linear_bounded[OF lf] obtain B where - B: "\x. norm (f x) \ B * norm x" by blast - let ?K = "\B\ + 1" - have Kp: "?K > 0" by arith - {assume C: "B < 0" - have "norm (1::real ^ 'n) > 0" by (simp add: zero_less_norm_iff) - with C have "B * norm (1:: real ^ 'n) < 0" - by (simp add: zero_compare_simps) - with B[rule_format, of 1] norm_ge_zero[of "f 1"] have False by simp - } - then have Bp: "B \ 0" by ferrack - {fix x::"real ^ 'n" - have "norm (f x) \ ?K * norm x" - using B[rule_format, of x] norm_ge_zero[of x] norm_ge_zero[of "f x"] Bp - apply (auto simp add: ring_simps split add: abs_split) - apply (erule order_trans, simp) - done - } - then show ?thesis using Kp by blast -qed - -lemma smult_conv_scaleR: "c *s x = scaleR c x" - unfolding vector_scalar_mult_def vector_scaleR_def by simp - -lemma linear_conv_bounded_linear: - fixes f :: "real ^ _ \ real ^ _" - shows "linear f \ bounded_linear f" -proof - assume "linear f" - show "bounded_linear f" - proof - fix x y show "f (x + y) = f x + f y" - using `linear f` unfolding linear_def by simp - next - fix r x show "f (scaleR r x) = scaleR r (f x)" - using `linear f` unfolding linear_def - by (simp add: smult_conv_scaleR) - next - have "\B. \x. norm (f x) \ B * norm x" - using `linear f` by (rule linear_bounded) - thus "\K. \x. norm (f x) \ norm x * K" - by (simp add: mult_commute) - qed -next - assume "bounded_linear f" - then interpret f: bounded_linear f . - show "linear f" - unfolding linear_def smult_conv_scaleR - by (simp add: f.add f.scaleR) -qed - -subsection{* Bilinear functions. *} - -definition "bilinear f \ (\x. linear(\y. f x y)) \ (\y. linear(\x. f x y))" - -lemma bilinear_ladd: "bilinear h ==> h (x + y) z = (h x z) + (h y z)" - by (simp add: bilinear_def linear_def) -lemma bilinear_radd: "bilinear h ==> h x (y + z) = (h x y) + (h x z)" - by (simp add: bilinear_def linear_def) - -lemma bilinear_lmul: "bilinear h ==> h (c *s x) y = c *s (h x y)" - by (simp add: bilinear_def linear_def) - -lemma bilinear_rmul: "bilinear h ==> h x (c *s y) = c *s (h x y)" - by (simp add: bilinear_def linear_def) - -lemma bilinear_lneg: "bilinear h ==> h (- (x:: 'a::ring_1 ^ 'n)) y = -(h x y)" - by (simp only: vector_sneg_minus1 bilinear_lmul) - -lemma bilinear_rneg: "bilinear h ==> h x (- (y:: 'a::ring_1 ^ 'n)) = - h x y" - by (simp only: vector_sneg_minus1 bilinear_rmul) - -lemma (in ab_group_add) eq_add_iff: "x = x + y \ y = 0" - using add_imp_eq[of x y 0] by auto - -lemma bilinear_lzero: - fixes h :: "'a::ring^'n \ _" assumes bh: "bilinear h" shows "h 0 x = 0" - using bilinear_ladd[OF bh, of 0 0 x] - by (simp add: eq_add_iff ring_simps) - -lemma bilinear_rzero: - fixes h :: "'a::ring^'n \ _" assumes bh: "bilinear h" shows "h x 0 = 0" - using bilinear_radd[OF bh, of x 0 0 ] - by (simp add: eq_add_iff ring_simps) - -lemma bilinear_lsub: "bilinear h ==> h (x - (y:: 'a::ring_1 ^ 'n)) z = h x z - h y z" - by (simp add: diff_def bilinear_ladd bilinear_lneg) - -lemma bilinear_rsub: "bilinear h ==> h z (x - (y:: 'a::ring_1 ^ 'n)) = h z x - h z y" - by (simp add: diff_def bilinear_radd bilinear_rneg) - -lemma bilinear_setsum: - fixes h:: "'a ^'n \ 'a::semiring_1^'m \ 'a ^ 'k" - assumes bh: "bilinear h" and fS: "finite S" and fT: "finite T" - shows "h (setsum f S) (setsum g T) = setsum (\(i,j). h (f i) (g j)) (S \ T) " -proof- - have "h (setsum f S) (setsum g T) = setsum (\x. h (f x) (setsum g T)) S" - apply (rule linear_setsum[unfolded o_def]) - using bh fS by (auto simp add: bilinear_def) - also have "\ = setsum (\x. setsum (\y. h (f x) (g y)) T) S" - apply (rule setsum_cong, simp) - apply (rule linear_setsum[unfolded o_def]) - using bh fT by (auto simp add: bilinear_def) - finally show ?thesis unfolding setsum_cartesian_product . -qed - -lemma bilinear_bounded: - fixes h:: "real ^'m::finite \ real^'n::finite \ real ^ 'k::finite" - assumes bh: "bilinear h" - shows "\B. \x y. norm (h x y) \ B * norm x * norm y" -proof- - let ?M = "UNIV :: 'm set" - let ?N = "UNIV :: 'n set" - let ?B = "setsum (\(i,j). norm (h (basis i) (basis j))) (?M \ ?N)" - have fM: "finite ?M" and fN: "finite ?N" by simp_all - {fix x:: "real ^ 'm" and y :: "real^'n" - have "norm (h x y) = norm (h (setsum (\i. (x$i) *s basis i) ?M) (setsum (\i. (y$i) *s basis i) ?N))" unfolding basis_expansion .. - also have "\ = norm (setsum (\ (i,j). h ((x$i) *s basis i) ((y$j) *s basis j)) (?M \ ?N))" unfolding bilinear_setsum[OF bh fM fN] .. - finally have th: "norm (h x y) = \" . - have "norm (h x y) \ ?B * norm x * norm y" - apply (simp add: setsum_left_distrib th) - apply (rule real_setsum_norm_le) - using fN fM - apply simp - apply (auto simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh] norm_mul ring_simps) - apply (rule mult_mono) - apply (auto simp add: norm_ge_zero zero_le_mult_iff component_le_norm) - apply (rule mult_mono) - apply (auto simp add: norm_ge_zero zero_le_mult_iff component_le_norm) - done} - then show ?thesis by metis -qed - -lemma bilinear_bounded_pos: - fixes h:: "real ^'m::finite \ real^'n::finite \ real ^ 'k::finite" - assumes bh: "bilinear h" - shows "\B > 0. \x y. norm (h x y) \ B * norm x * norm y" -proof- - from bilinear_bounded[OF bh] obtain B where - B: "\x y. norm (h x y) \ B * norm x * norm y" by blast - let ?K = "\B\ + 1" - have Kp: "?K > 0" by arith - have KB: "B < ?K" by arith - {fix x::"real ^'m" and y :: "real ^'n" - from KB Kp - have "B * norm x * norm y \ ?K * norm x * norm y" - apply - - apply (rule mult_right_mono, rule mult_right_mono) - by (auto simp add: norm_ge_zero) - then have "norm (h x y) \ ?K * norm x * norm y" - using B[rule_format, of x y] by simp} - with Kp show ?thesis by blast -qed - -lemma bilinear_conv_bounded_bilinear: - fixes h :: "real ^ _ \ real ^ _ \ real ^ _" - shows "bilinear h \ bounded_bilinear h" -proof - assume "bilinear h" - show "bounded_bilinear h" - proof - fix x y z show "h (x + y) z = h x z + h y z" - using `bilinear h` unfolding bilinear_def linear_def by simp - next - fix x y z show "h x (y + z) = h x y + h x z" - using `bilinear h` unfolding bilinear_def linear_def by simp - next - fix r x y show "h (scaleR r x) y = scaleR r (h x y)" - using `bilinear h` unfolding bilinear_def linear_def - by (simp add: smult_conv_scaleR) - next - fix r x y show "h x (scaleR r y) = scaleR r (h x y)" - using `bilinear h` unfolding bilinear_def linear_def - by (simp add: smult_conv_scaleR) - next - have "\B. \x y. norm (h x y) \ B * norm x * norm y" - using `bilinear h` by (rule bilinear_bounded) - thus "\K. \x y. norm (h x y) \ norm x * norm y * K" - by (simp add: mult_ac) - qed -next - assume "bounded_bilinear h" - then interpret h: bounded_bilinear h . - show "bilinear h" - unfolding bilinear_def linear_conv_bounded_linear - using h.bounded_linear_left h.bounded_linear_right - by simp -qed - -subsection{* Adjoints. *} - -definition "adjoint f = (SOME f'. \x y. f x \ y = x \ f' y)" - -lemma choice_iff: "(\x. \y. P x y) \ (\f. \x. P x (f x))" by metis - -lemma adjoint_works_lemma: - fixes f:: "'a::ring_1 ^'n::finite \ 'a ^ 'm::finite" - assumes lf: "linear f" - shows "\x y. f x \ y = x \ adjoint f y" -proof- - let ?N = "UNIV :: 'n set" - let ?M = "UNIV :: 'm set" - have fN: "finite ?N" by simp - have fM: "finite ?M" by simp - {fix y:: "'a ^ 'm" - let ?w = "(\ i. (f (basis i) \ y)) :: 'a ^ 'n" - {fix x - have "f x \ y = f (setsum (\i. (x$i) *s basis i) ?N) \ y" - by (simp only: basis_expansion) - also have "\ = (setsum (\i. (x$i) *s f (basis i)) ?N) \ y" - unfolding linear_setsum[OF lf fN] - by (simp add: linear_cmul[OF lf]) - finally have "f x \ y = x \ ?w" - apply (simp only: ) - apply (simp add: dot_def setsum_left_distrib setsum_right_distrib setsum_commute[of _ ?M ?N] ring_simps) - done} - } - then show ?thesis unfolding adjoint_def - some_eq_ex[of "\f'. \x y. f x \ y = x \ f' y"] - using choice_iff[of "\a b. \x. f x \ a = x \ b "] - by metis -qed - -lemma adjoint_works: - fixes f:: "'a::ring_1 ^'n::finite \ 'a ^ 'm::finite" - assumes lf: "linear f" - shows "x \ adjoint f y = f x \ y" - using adjoint_works_lemma[OF lf] by metis - - -lemma adjoint_linear: - fixes f :: "'a::comm_ring_1 ^'n::finite \ 'a ^ 'm::finite" - assumes lf: "linear f" - shows "linear (adjoint f)" - by (simp add: linear_def vector_eq_ldot[symmetric] dot_radd dot_rmult adjoint_works[OF lf]) - -lemma adjoint_clauses: - fixes f:: "'a::comm_ring_1 ^'n::finite \ 'a ^ 'm::finite" - assumes lf: "linear f" - shows "x \ adjoint f y = f x \ y" - and "adjoint f y \ x = y \ f x" - by (simp_all add: adjoint_works[OF lf] dot_sym ) - -lemma adjoint_adjoint: - fixes f:: "'a::comm_ring_1 ^ 'n::finite \ 'a ^ 'm::finite" - assumes lf: "linear f" - shows "adjoint (adjoint f) = f" - apply (rule ext) - by (simp add: vector_eq_ldot[symmetric] adjoint_clauses[OF adjoint_linear[OF lf]] adjoint_clauses[OF lf]) - -lemma adjoint_unique: - fixes f:: "'a::comm_ring_1 ^ 'n::finite \ 'a ^ 'm::finite" - assumes lf: "linear f" and u: "\x y. f' x \ y = x \ f y" - shows "f' = adjoint f" - apply (rule ext) - using u - by (simp add: vector_eq_rdot[symmetric] adjoint_clauses[OF lf]) - -text{* Matrix notation. NB: an MxN matrix is of type @{typ "'a^'n^'m"}, not @{typ "'a^'m^'n"} *} - -consts generic_mult :: "'a \ 'b \ 'c" (infixr "\" 75) - -defs (overloaded) -matrix_matrix_mult_def: "(m:: ('a::semiring_1) ^'n^'m) \ (m' :: 'a ^'p^'n) \ (\ i j. setsum (\k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m" - -abbreviation - matrix_matrix_mult' :: "('a::semiring_1) ^'n^'m \ 'a ^'p^'n \ 'a ^ 'p ^'m" (infixl "**" 70) - where "m ** m' == m\ m'" - -defs (overloaded) - matrix_vector_mult_def: "(m::('a::semiring_1) ^'n^'m) \ (x::'a ^'n) \ (\ i. setsum (\j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m" - -abbreviation - matrix_vector_mult' :: "('a::semiring_1) ^'n^'m \ 'a ^'n \ 'a ^ 'm" (infixl "*v" 70) - where - "m *v v == m \ v" - -defs (overloaded) - vector_matrix_mult_def: "(x::'a^'m) \ (m::('a::semiring_1) ^'n^'m) \ (\ j. setsum (\i. ((m$i)$j) * (x$i)) (UNIV :: 'm set)) :: 'a^'n" - -abbreviation - vactor_matrix_mult' :: "'a ^ 'm \ ('a::semiring_1) ^'n^'m \ 'a ^'n " (infixl "v*" 70) - where - "v v* m == v \ m" - -definition "(mat::'a::zero => 'a ^'n^'n) k = (\ i j. if i = j then k else 0)" -definition "(transp::'a^'n^'m \ 'a^'m^'n) A = (\ i j. ((A$j)$i))" -definition "(row::'m => 'a ^'n^'m \ 'a ^'n) i A = (\ j. ((A$i)$j))" -definition "(column::'n =>'a^'n^'m =>'a^'m) j A = (\ i. ((A$i)$j))" -definition "rows(A::'a^'n^'m) = { row i A | i. i \ (UNIV :: 'm set)}" -definition "columns(A::'a^'n^'m) = { column i A | i. i \ (UNIV :: 'n set)}" - -lemma mat_0[simp]: "mat 0 = 0" by (vector mat_def) -lemma matrix_add_ldistrib: "(A ** (B + C)) = (A \ B) + (A \ C)" - by (vector matrix_matrix_mult_def setsum_addf[symmetric] ring_simps) - -lemma matrix_mul_lid: - fixes A :: "'a::semiring_1 ^ 'm ^ 'n::finite" - shows "mat 1 ** A = A" - apply (simp add: matrix_matrix_mult_def mat_def) - apply vector - by (auto simp only: cond_value_iff cond_application_beta setsum_delta'[OF finite] mult_1_left mult_zero_left if_True UNIV_I) - - -lemma matrix_mul_rid: - fixes A :: "'a::semiring_1 ^ 'm::finite ^ 'n" - shows "A ** mat 1 = A" - apply (simp add: matrix_matrix_mult_def mat_def) - apply vector - by (auto simp only: cond_value_iff cond_application_beta setsum_delta[OF finite] mult_1_right mult_zero_right if_True UNIV_I cong: if_cong) - -lemma matrix_mul_assoc: "A ** (B ** C) = (A ** B) ** C" - apply (vector matrix_matrix_mult_def setsum_right_distrib setsum_left_distrib mult_assoc) - apply (subst setsum_commute) - apply simp - done - -lemma matrix_vector_mul_assoc: "A *v (B *v x) = (A ** B) *v x" - apply (vector matrix_matrix_mult_def matrix_vector_mult_def setsum_right_distrib setsum_left_distrib mult_assoc) - apply (subst setsum_commute) - apply simp - done - -lemma matrix_vector_mul_lid: "mat 1 *v x = (x::'a::semiring_1 ^ 'n::finite)" - apply (vector matrix_vector_mult_def mat_def) - by (simp add: cond_value_iff cond_application_beta - setsum_delta' cong del: if_weak_cong) - -lemma matrix_transp_mul: "transp(A ** B) = transp B ** transp (A::'a::comm_semiring_1^'m^'n)" - by (simp add: matrix_matrix_mult_def transp_def Cart_eq mult_commute) - -lemma matrix_eq: - fixes A B :: "'a::semiring_1 ^ 'n::finite ^ 'm" - shows "A = B \ (\x. A *v x = B *v x)" (is "?lhs \ ?rhs") - apply auto - apply (subst Cart_eq) - apply clarify - apply (clarsimp simp add: matrix_vector_mult_def basis_def cond_value_iff cond_application_beta Cart_eq cong del: if_weak_cong) - apply (erule_tac x="basis ia" in allE) - apply (erule_tac x="i" in allE) - by (auto simp add: basis_def cond_value_iff cond_application_beta setsum_delta[OF finite] cong del: if_weak_cong) - -lemma matrix_vector_mul_component: - shows "((A::'a::semiring_1^'n'^'m) *v x)$k = (A$k) \ x" - by (simp add: matrix_vector_mult_def dot_def) - -lemma dot_lmul_matrix: "((x::'a::comm_semiring_1 ^'n) v* A) \ y = x \ (A *v y)" - apply (simp add: dot_def matrix_vector_mult_def vector_matrix_mult_def setsum_left_distrib setsum_right_distrib mult_ac) - apply (subst setsum_commute) - by simp - -lemma transp_mat: "transp (mat n) = mat n" - by (vector transp_def mat_def) - -lemma transp_transp: "transp(transp A) = A" - by (vector transp_def) - -lemma row_transp: - fixes A:: "'a::semiring_1^'n^'m" - shows "row i (transp A) = column i A" - by (simp add: row_def column_def transp_def Cart_eq) - -lemma column_transp: - fixes A:: "'a::semiring_1^'n^'m" - shows "column i (transp A) = row i A" - by (simp add: row_def column_def transp_def Cart_eq) - -lemma rows_transp: "rows(transp (A::'a::semiring_1^'n^'m)) = columns A" -by (auto simp add: rows_def columns_def row_transp intro: set_ext) - -lemma columns_transp: "columns(transp (A::'a::semiring_1^'n^'m)) = rows A" by (metis transp_transp rows_transp) - -text{* Two sometimes fruitful ways of looking at matrix-vector multiplication. *} - -lemma matrix_mult_dot: "A *v x = (\ i. A$i \ x)" - by (simp add: matrix_vector_mult_def dot_def) - -lemma matrix_mult_vsum: "(A::'a::comm_semiring_1^'n^'m) *v x = setsum (\i. (x$i) *s column i A) (UNIV:: 'n set)" - by (simp add: matrix_vector_mult_def Cart_eq column_def mult_commute) - -lemma vector_componentwise: - "(x::'a::ring_1^'n::finite) = (\ j. setsum (\i. (x$i) * (basis i :: 'a^'n)$j) (UNIV :: 'n set))" - apply (subst basis_expansion[symmetric]) - by (vector Cart_eq setsum_component) - -lemma linear_componentwise: - fixes f:: "'a::ring_1 ^ 'm::finite \ 'a ^ 'n" - assumes lf: "linear f" - shows "(f x)$j = setsum (\i. (x$i) * (f (basis i)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs") -proof- - let ?M = "(UNIV :: 'm set)" - let ?N = "(UNIV :: 'n set)" - have fM: "finite ?M" by simp - have "?rhs = (setsum (\i.(x$i) *s f (basis i) ) ?M)$j" - unfolding vector_smult_component[symmetric] - unfolding setsum_component[of "(\i.(x$i) *s f (basis i :: 'a^'m))" ?M] - .. - then show ?thesis unfolding linear_setsum_mul[OF lf fM, symmetric] basis_expansion .. -qed - -text{* Inverse matrices (not necessarily square) *} - -definition "invertible(A::'a::semiring_1^'n^'m) \ (\A'::'a^'m^'n. A ** A' = mat 1 \ A' ** A = mat 1)" - -definition "matrix_inv(A:: 'a::semiring_1^'n^'m) = - (SOME A'::'a^'m^'n. A ** A' = mat 1 \ A' ** A = mat 1)" - -text{* Correspondence between matrices and linear operators. *} - -definition matrix:: "('a::{plus,times, one, zero}^'m \ 'a ^ 'n) \ 'a^'m^'n" -where "matrix f = (\ i j. (f(basis j))$i)" - -lemma matrix_vector_mul_linear: "linear(\x. A *v (x::'a::comm_semiring_1 ^ 'n))" - by (simp add: linear_def matrix_vector_mult_def Cart_eq ring_simps setsum_right_distrib setsum_addf) - -lemma matrix_works: assumes lf: "linear f" shows "matrix f *v x = f (x::'a::comm_ring_1 ^ 'n::finite)" -apply (simp add: matrix_def matrix_vector_mult_def Cart_eq mult_commute) -apply clarify -apply (rule linear_componentwise[OF lf, symmetric]) -done - -lemma matrix_vector_mul: "linear f ==> f = (\x. matrix f *v (x::'a::comm_ring_1 ^ 'n::finite))" by (simp add: ext matrix_works) - -lemma matrix_of_matrix_vector_mul: "matrix(\x. A *v (x :: 'a:: comm_ring_1 ^ 'n::finite)) = A" - by (simp add: matrix_eq matrix_vector_mul_linear matrix_works) - -lemma matrix_compose: - assumes lf: "linear (f::'a::comm_ring_1^'n::finite \ 'a^'m::finite)" - and lg: "linear (g::'a::comm_ring_1^'m::finite \ 'a^'k)" - shows "matrix (g o f) = matrix g ** matrix f" - using lf lg linear_compose[OF lf lg] matrix_works[OF linear_compose[OF lf lg]] - by (simp add: matrix_eq matrix_works matrix_vector_mul_assoc[symmetric] o_def) - -lemma matrix_vector_column:"(A::'a::comm_semiring_1^'n^'m) *v x = setsum (\i. (x$i) *s ((transp A)$i)) (UNIV:: 'n set)" - by (simp add: matrix_vector_mult_def transp_def Cart_eq mult_commute) - -lemma adjoint_matrix: "adjoint(\x. (A::'a::comm_ring_1^'n::finite^'m::finite) *v x) = (\x. transp A *v x)" - apply (rule adjoint_unique[symmetric]) - apply (rule matrix_vector_mul_linear) - apply (simp add: transp_def dot_def matrix_vector_mult_def setsum_left_distrib setsum_right_distrib) - apply (subst setsum_commute) - apply (auto simp add: mult_ac) - done - -lemma matrix_adjoint: assumes lf: "linear (f :: 'a::comm_ring_1^'n::finite \ 'a ^ 'm::finite)" - shows "matrix(adjoint f) = transp(matrix f)" - apply (subst matrix_vector_mul[OF lf]) - unfolding adjoint_matrix matrix_of_matrix_vector_mul .. - -subsection{* Interlude: Some properties of real sets *} - -lemma seq_mono_lemma: assumes "\(n::nat) \ m. (d n :: real) < e n" and "\n \ m. e n <= e m" - shows "\n \ m. d n < e m" - using prems apply auto - apply (erule_tac x="n" in allE) - apply (erule_tac x="n" in allE) - apply auto - done - - -lemma real_convex_bound_lt: - assumes xa: "(x::real) < a" and ya: "y < a" and u: "0 <= u" and v: "0 <= v" - and uv: "u + v = 1" - shows "u * x + v * y < a" -proof- - have uv': "u = 0 \ v \ 0" using u v uv by arith - have "a = a * (u + v)" unfolding uv by simp - hence th: "u * a + v * a = a" by (simp add: ring_simps) - from xa u have "u \ 0 \ u*x < u*a" by (simp add: mult_compare_simps) - from ya v have "v \ 0 \ v * y < v * a" by (simp add: mult_compare_simps) - from xa ya u v have "u * x + v * y < u * a + v * a" - apply (cases "u = 0", simp_all add: uv') - apply(rule mult_strict_left_mono) - using uv' apply simp_all - - apply (rule add_less_le_mono) - apply(rule mult_strict_left_mono) - apply simp_all - apply (rule mult_left_mono) - apply simp_all - done - thus ?thesis unfolding th . -qed - -lemma real_convex_bound_le: - assumes xa: "(x::real) \ a" and ya: "y \ a" and u: "0 <= u" and v: "0 <= v" - and uv: "u + v = 1" - shows "u * x + v * y \ a" -proof- - from xa ya u v have "u * x + v * y \ u * a + v * a" by (simp add: add_mono mult_left_mono) - also have "\ \ (u + v) * a" by (simp add: ring_simps) - finally show ?thesis unfolding uv by simp -qed - -lemma infinite_enumerate: assumes fS: "infinite S" - shows "\r. subseq r \ (\n. r n \ S)" -unfolding subseq_def -using enumerate_in_set[OF fS] enumerate_mono[of _ _ S] fS by auto - -lemma approachable_lt_le: "(\(d::real)>0. \x. f x < d \ P x) \ (\d>0. \x. f x \ d \ P x)" -apply auto -apply (rule_tac x="d/2" in exI) -apply auto -done - - -lemma triangle_lemma: - assumes x: "0 <= (x::real)" and y:"0 <= y" and z: "0 <= z" and xy: "x^2 <= y^2 + z^2" - shows "x <= y + z" -proof- - have "y^2 + z^2 \ y^2 + 2*y*z + z^2" using z y by (simp add: zero_compare_simps) - with xy have th: "x ^2 \ (y+z)^2" by (simp add: power2_eq_square ring_simps) - from y z have yz: "y + z \ 0" by arith - from power2_le_imp_le[OF th yz] show ?thesis . -qed - - -lemma lambda_skolem: "(\i. \x. P i x) \ - (\x::'a ^ 'n. \i. P i (x$i))" (is "?lhs \ ?rhs") -proof- - let ?S = "(UNIV :: 'n set)" - {assume H: "?rhs" - then have ?lhs by auto} - moreover - {assume H: "?lhs" - then obtain f where f:"\i. P i (f i)" unfolding choice_iff by metis - let ?x = "(\ i. (f i)) :: 'a ^ 'n" - {fix i - from f have "P i (f i)" by metis - then have "P i (?x$i)" by auto - } - hence "\i. P i (?x$i)" by metis - hence ?rhs by metis } - ultimately show ?thesis by metis -qed - -(* Supremum and infimum of real sets *) - - -definition rsup:: "real set \ real" where - "rsup S = (SOME a. isLub UNIV S a)" - -lemma rsup_alt: "rsup S = (SOME a. (\x \ S. x \ a) \ (\b. (\x \ S. x \ b) \ a \ b))" by (auto simp add: isLub_def rsup_def leastP_def isUb_def setle_def setge_def) - -lemma rsup: assumes Se: "S \ {}" and b: "\b. S *<= b" - shows "isLub UNIV S (rsup S)" -using Se b -unfolding rsup_def -apply clarify -apply (rule someI_ex) -apply (rule reals_complete) -by (auto simp add: isUb_def setle_def) - -lemma rsup_le: assumes Se: "S \ {}" and Sb: "S *<= b" shows "rsup S \ b" -proof- - from Sb have bu: "isUb UNIV S b" by (simp add: isUb_def setle_def) - from rsup[OF Se] Sb have "isLub UNIV S (rsup S)" by blast - then show ?thesis using bu by (auto simp add: isLub_def leastP_def setle_def setge_def) -qed - -lemma rsup_finite_Max: assumes fS: "finite S" and Se: "S \ {}" - shows "rsup S = Max S" -using fS Se -proof- - let ?m = "Max S" - from Max_ge[OF fS] have Sm: "\ x\ S. x \ ?m" by metis - with rsup[OF Se] have lub: "isLub UNIV S (rsup S)" by (metis setle_def) - from Max_in[OF fS Se] lub have mrS: "?m \ rsup S" - by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def) - moreover - have "rsup S \ ?m" using Sm lub - by (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def) - ultimately show ?thesis by arith -qed - -lemma rsup_finite_in: assumes fS: "finite S" and Se: "S \ {}" - shows "rsup S \ S" - using rsup_finite_Max[OF fS Se] Max_in[OF fS Se] by metis - -lemma rsup_finite_Ub: assumes fS: "finite S" and Se: "S \ {}" - shows "isUb S S (rsup S)" - using rsup_finite_Max[OF fS Se] rsup_finite_in[OF fS Se] Max_ge[OF fS] - unfolding isUb_def setle_def by metis - -lemma rsup_finite_ge_iff: assumes fS: "finite S" and Se: "S \ {}" - shows "a \ rsup S \ (\ x \ S. a \ x)" -using rsup_finite_Ub[OF fS Se] by (auto simp add: isUb_def setle_def) - -lemma rsup_finite_le_iff: assumes fS: "finite S" and Se: "S \ {}" - shows "a \ rsup S \ (\ x \ S. a \ x)" -using rsup_finite_Ub[OF fS Se] by (auto simp add: isUb_def setle_def) - -lemma rsup_finite_gt_iff: assumes fS: "finite S" and Se: "S \ {}" - shows "a < rsup S \ (\ x \ S. a < x)" -using rsup_finite_Ub[OF fS Se] by (auto simp add: isUb_def setle_def) - -lemma rsup_finite_lt_iff: assumes fS: "finite S" and Se: "S \ {}" - shows "a > rsup S \ (\ x \ S. a > x)" -using rsup_finite_Ub[OF fS Se] by (auto simp add: isUb_def setle_def) - -lemma rsup_unique: assumes b: "S *<= b" and S: "\b' < b. \x \ S. b' < x" - shows "rsup S = b" -using b S -unfolding setle_def rsup_alt -apply - -apply (rule some_equality) -apply (metis linorder_not_le order_eq_iff[symmetric])+ -done - -lemma rsup_le_subset: "S\{} \ S \ T \ (\b. T *<= b) \ rsup S \ rsup T" - apply (rule rsup_le) - apply simp - using rsup[of T] by (auto simp add: isLub_def leastP_def setge_def setle_def isUb_def) - -lemma isUb_def': "isUb R S = (\x. S *<= x \ x \ R)" - apply (rule ext) - by (metis isUb_def) - -lemma UNIV_trivial: "UNIV x" using UNIV_I[of x] by (metis mem_def) -lemma rsup_bounds: assumes Se: "S \ {}" and l: "a <=* S" and u: "S *<= b" - shows "a \ rsup S \ rsup S \ b" -proof- - from rsup[OF Se] u have lub: "isLub UNIV S (rsup S)" by blast - hence b: "rsup S \ b" using u by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def') - from Se obtain y where y: "y \ S" by blast - from lub l have "a \ rsup S" apply (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def') - apply (erule ballE[where x=y]) - apply (erule ballE[where x=y]) - apply arith - using y apply auto - done - with b show ?thesis by blast -qed - -lemma rsup_abs_le: "S \ {} \ (\x\S. \x\ \ a) \ \rsup S\ \ a" - unfolding abs_le_interval_iff using rsup_bounds[of S "-a" a] - by (auto simp add: setge_def setle_def) - -lemma rsup_asclose: assumes S:"S \ {}" and b: "\x\S. \x - l\ \ e" shows "\rsup S - l\ \ e" -proof- - have th: "\(x::real) l e. \x - l\ \ e \ l - e \ x \ x \ l + e" by arith - show ?thesis using S b rsup_bounds[of S "l - e" "l+e"] unfolding th - by (auto simp add: setge_def setle_def) -qed - -definition rinf:: "real set \ real" where - "rinf S = (SOME a. isGlb UNIV S a)" - -lemma rinf_alt: "rinf S = (SOME a. (\x \ S. x \ a) \ (\b. (\x \ S. x \ b) \ a \ b))" by (auto simp add: isGlb_def rinf_def greatestP_def isLb_def setle_def setge_def) - -lemma reals_complete_Glb: assumes Se: "\x. x \ S" and lb: "\ y. isLb UNIV S y" - shows "\(t::real). isGlb UNIV S t" -proof- - let ?M = "uminus ` S" - from lb have th: "\y. isUb UNIV ?M y" apply (auto simp add: isUb_def isLb_def setle_def setge_def) - by (rule_tac x="-y" in exI, auto) - from Se have Me: "\x. x \ ?M" by blast - from reals_complete[OF Me th] obtain t where t: "isLub UNIV ?M t" by blast - have "isGlb UNIV S (- t)" using t - apply (auto simp add: isLub_def isGlb_def leastP_def greatestP_def setle_def setge_def isUb_def isLb_def) - apply (erule_tac x="-y" in allE) - apply auto - done - then show ?thesis by metis -qed - -lemma rinf: assumes Se: "S \ {}" and b: "\b. b <=* S" - shows "isGlb UNIV S (rinf S)" -using Se b -unfolding rinf_def -apply clarify -apply (rule someI_ex) -apply (rule reals_complete_Glb) -apply (auto simp add: isLb_def setle_def setge_def) -done - -lemma rinf_ge: assumes Se: "S \ {}" and Sb: "b <=* S" shows "rinf S \ b" -proof- - from Sb have bu: "isLb UNIV S b" by (simp add: isLb_def setge_def) - from rinf[OF Se] Sb have "isGlb UNIV S (rinf S)" by blast - then show ?thesis using bu by (auto simp add: isGlb_def greatestP_def setle_def setge_def) -qed - -lemma rinf_finite_Min: assumes fS: "finite S" and Se: "S \ {}" - shows "rinf S = Min S" -using fS Se -proof- - let ?m = "Min S" - from Min_le[OF fS] have Sm: "\ x\ S. x \ ?m" by metis - with rinf[OF Se] have glb: "isGlb UNIV S (rinf S)" by (metis setge_def) - from Min_in[OF fS Se] glb have mrS: "?m \ rinf S" - by (auto simp add: isGlb_def greatestP_def setle_def setge_def isLb_def) - moreover - have "rinf S \ ?m" using Sm glb - by (auto simp add: isGlb_def greatestP_def isLb_def setle_def setge_def) - ultimately show ?thesis by arith -qed - -lemma rinf_finite_in: assumes fS: "finite S" and Se: "S \ {}" - shows "rinf S \ S" - using rinf_finite_Min[OF fS Se] Min_in[OF fS Se] by metis - -lemma rinf_finite_Lb: assumes fS: "finite S" and Se: "S \ {}" - shows "isLb S S (rinf S)" - using rinf_finite_Min[OF fS Se] rinf_finite_in[OF fS Se] Min_le[OF fS] - unfolding isLb_def setge_def by metis - -lemma rinf_finite_ge_iff: assumes fS: "finite S" and Se: "S \ {}" - shows "a \ rinf S \ (\ x \ S. a \ x)" -using rinf_finite_Lb[OF fS Se] by (auto simp add: isLb_def setge_def) - -lemma rinf_finite_le_iff: assumes fS: "finite S" and Se: "S \ {}" - shows "a \ rinf S \ (\ x \ S. a \ x)" -using rinf_finite_Lb[OF fS Se] by (auto simp add: isLb_def setge_def) - -lemma rinf_finite_gt_iff: assumes fS: "finite S" and Se: "S \ {}" - shows "a < rinf S \ (\ x \ S. a < x)" -using rinf_finite_Lb[OF fS Se] by (auto simp add: isLb_def setge_def) - -lemma rinf_finite_lt_iff: assumes fS: "finite S" and Se: "S \ {}" - shows "a > rinf S \ (\ x \ S. a > x)" -using rinf_finite_Lb[OF fS Se] by (auto simp add: isLb_def setge_def) - -lemma rinf_unique: assumes b: "b <=* S" and S: "\b' > b. \x \ S. b' > x" - shows "rinf S = b" -using b S -unfolding setge_def rinf_alt -apply - -apply (rule some_equality) -apply (metis linorder_not_le order_eq_iff[symmetric])+ -done - -lemma rinf_ge_subset: "S\{} \ S \ T \ (\b. b <=* T) \ rinf S >= rinf T" - apply (rule rinf_ge) - apply simp - using rinf[of T] by (auto simp add: isGlb_def greatestP_def setge_def setle_def isLb_def) - -lemma isLb_def': "isLb R S = (\x. x <=* S \ x \ R)" - apply (rule ext) - by (metis isLb_def) - -lemma rinf_bounds: assumes Se: "S \ {}" and l: "a <=* S" and u: "S *<= b" - shows "a \ rinf S \ rinf S \ b" -proof- - from rinf[OF Se] l have lub: "isGlb UNIV S (rinf S)" by blast - hence b: "a \ rinf S" using l by (auto simp add: isGlb_def greatestP_def setle_def setge_def isLb_def') - from Se obtain y where y: "y \ S" by blast - from lub u have "b \ rinf S" apply (auto simp add: isGlb_def greatestP_def setle_def setge_def isLb_def') - apply (erule ballE[where x=y]) - apply (erule ballE[where x=y]) - apply arith - using y apply auto - done - with b show ?thesis by blast -qed - -lemma rinf_abs_ge: "S \ {} \ (\x\S. \x\ \ a) \ \rinf S\ \ a" - unfolding abs_le_interval_iff using rinf_bounds[of S "-a" a] - by (auto simp add: setge_def setle_def) - -lemma rinf_asclose: assumes S:"S \ {}" and b: "\x\S. \x - l\ \ e" shows "\rinf S - l\ \ e" -proof- - have th: "\(x::real) l e. \x - l\ \ e \ l - e \ x \ x \ l + e" by arith - show ?thesis using S b rinf_bounds[of S "l - e" "l+e"] unfolding th - by (auto simp add: setge_def setle_def) -qed - - - -subsection{* Operator norm. *} - -definition "onorm f = rsup {norm (f x)| x. norm x = 1}" - -lemma norm_bound_generalize: - fixes f:: "real ^'n::finite \ real^'m::finite" - assumes lf: "linear f" - shows "(\x. norm x = 1 \ norm (f x) \ b) \ (\x. norm (f x) \ b * norm x)" (is "?lhs \ ?rhs") -proof- - {assume H: ?rhs - {fix x :: "real^'n" assume x: "norm x = 1" - from H[rule_format, of x] x have "norm (f x) \ b" by simp} - then have ?lhs by blast } - - moreover - {assume H: ?lhs - from H[rule_format, of "basis arbitrary"] - have bp: "b \ 0" using norm_ge_zero[of "f (basis arbitrary)"] - by (auto simp add: norm_basis elim: order_trans [OF norm_ge_zero]) - {fix x :: "real ^'n" - {assume "x = 0" - then have "norm (f x) \ b * norm x" by (simp add: linear_0[OF lf] bp)} - moreover - {assume x0: "x \ 0" - hence n0: "norm x \ 0" by (metis norm_eq_zero) - let ?c = "1/ norm x" - have "norm (?c*s x) = 1" using x0 by (simp add: n0 norm_mul) - with H have "norm (f(?c*s x)) \ b" by blast - hence "?c * norm (f x) \ b" - by (simp add: linear_cmul[OF lf] norm_mul) - hence "norm (f x) \ b * norm x" - using n0 norm_ge_zero[of x] by (auto simp add: field_simps)} - ultimately have "norm (f x) \ b * norm x" by blast} - then have ?rhs by blast} - ultimately show ?thesis by blast -qed - -lemma onorm: - fixes f:: "real ^'n::finite \ real ^'m::finite" - assumes lf: "linear f" - shows "norm (f x) <= onorm f * norm x" - and "\x. norm (f x) <= b * norm x \ onorm f <= b" -proof- - { - let ?S = "{norm (f x) |x. norm x = 1}" - have Se: "?S \ {}" using norm_basis by auto - from linear_bounded[OF lf] have b: "\ b. ?S *<= b" - unfolding norm_bound_generalize[OF lf, symmetric] by (auto simp add: setle_def) - {from rsup[OF Se b, unfolded onorm_def[symmetric]] - show "norm (f x) <= onorm f * norm x" - apply - - apply (rule spec[where x = x]) - unfolding norm_bound_generalize[OF lf, symmetric] - by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)} - { - show "\x. norm (f x) <= b * norm x \ onorm f <= b" - using rsup[OF Se b, unfolded onorm_def[symmetric]] - unfolding norm_bound_generalize[OF lf, symmetric] - by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)} - } -qed - -lemma onorm_pos_le: assumes lf: "linear (f::real ^'n::finite \ real ^'m::finite)" shows "0 <= onorm f" - using order_trans[OF norm_ge_zero onorm(1)[OF lf, of "basis arbitrary"], unfolded norm_basis] by simp - -lemma onorm_eq_0: assumes lf: "linear (f::real ^'n::finite \ real ^'m::finite)" - shows "onorm f = 0 \ (\x. f x = 0)" - using onorm[OF lf] - apply (auto simp add: onorm_pos_le) - apply atomize - apply (erule allE[where x="0::real"]) - using onorm_pos_le[OF lf] - apply arith - done - -lemma onorm_const: "onorm(\x::real^'n::finite. (y::real ^ 'm::finite)) = norm y" -proof- - let ?f = "\x::real^'n. (y::real ^ 'm)" - have th: "{norm (?f x)| x. norm x = 1} = {norm y}" - by(auto intro: vector_choose_size set_ext) - show ?thesis - unfolding onorm_def th - apply (rule rsup_unique) by (simp_all add: setle_def) -qed - -lemma onorm_pos_lt: assumes lf: "linear (f::real ^ 'n::finite \ real ^'m::finite)" - shows "0 < onorm f \ ~(\x. f x = 0)" - unfolding onorm_eq_0[OF lf, symmetric] - using onorm_pos_le[OF lf] by arith - -lemma onorm_compose: - assumes lf: "linear (f::real ^'n::finite \ real ^'m::finite)" - and lg: "linear (g::real^'k::finite \ real^'n::finite)" - shows "onorm (f o g) <= onorm f * onorm g" - apply (rule onorm(2)[OF linear_compose[OF lg lf], rule_format]) - unfolding o_def - apply (subst mult_assoc) - apply (rule order_trans) - apply (rule onorm(1)[OF lf]) - apply (rule mult_mono1) - apply (rule onorm(1)[OF lg]) - apply (rule onorm_pos_le[OF lf]) - done - -lemma onorm_neg_lemma: assumes lf: "linear (f::real ^'n::finite \ real^'m::finite)" - shows "onorm (\x. - f x) \ onorm f" - using onorm[OF linear_compose_neg[OF lf]] onorm[OF lf] - unfolding norm_minus_cancel by metis - -lemma onorm_neg: assumes lf: "linear (f::real ^'n::finite \ real^'m::finite)" - shows "onorm (\x. - f x) = onorm f" - using onorm_neg_lemma[OF lf] onorm_neg_lemma[OF linear_compose_neg[OF lf]] - by simp - -lemma onorm_triangle: - assumes lf: "linear (f::real ^'n::finite \ real ^'m::finite)" and lg: "linear g" - shows "onorm (\x. f x + g x) <= onorm f + onorm g" - apply(rule onorm(2)[OF linear_compose_add[OF lf lg], rule_format]) - apply (rule order_trans) - apply (rule norm_triangle_ineq) - apply (simp add: distrib) - apply (rule add_mono) - apply (rule onorm(1)[OF lf]) - apply (rule onorm(1)[OF lg]) - done - -lemma onorm_triangle_le: "linear (f::real ^'n::finite \ real ^'m::finite) \ linear g \ onorm(f) + onorm(g) <= e - \ onorm(\x. f x + g x) <= e" - apply (rule order_trans) - apply (rule onorm_triangle) - apply assumption+ - done - -lemma onorm_triangle_lt: "linear (f::real ^'n::finite \ real ^'m::finite) \ linear g \ onorm(f) + onorm(g) < e - ==> onorm(\x. f x + g x) < e" - apply (rule order_le_less_trans) - apply (rule onorm_triangle) - by assumption+ - -(* "lift" from 'a to 'a^1 and "drop" from 'a^1 to 'a -- FIXME: potential use of transfer *) - -definition vec1:: "'a \ 'a ^ 1" where "vec1 x = (\ i. x)" - -definition dest_vec1:: "'a ^1 \ 'a" - where "dest_vec1 x = (x$1)" - -lemma vec1_component[simp]: "(vec1 x)$1 = x" - by (simp add: vec1_def) - -lemma vec1_dest_vec1[simp]: "vec1(dest_vec1 x) = x" "dest_vec1(vec1 y) = y" - by (simp_all add: vec1_def dest_vec1_def Cart_eq forall_1) - -lemma forall_vec1: "(\x. P x) \ (\x. P (vec1 x))" by (metis vec1_dest_vec1) - -lemma exists_vec1: "(\x. P x) \ (\x. P(vec1 x))" by (metis vec1_dest_vec1) - -lemma forall_dest_vec1: "(\x. P x) \ (\x. P(dest_vec1 x))" by (metis vec1_dest_vec1) - -lemma exists_dest_vec1: "(\x. P x) \ (\x. P(dest_vec1 x))"by (metis vec1_dest_vec1) - -lemma vec1_eq[simp]: "vec1 x = vec1 y \ x = y" by (metis vec1_dest_vec1) - -lemma dest_vec1_eq[simp]: "dest_vec1 x = dest_vec1 y \ x = y" by (metis vec1_dest_vec1) - -lemma vec1_in_image_vec1: "vec1 x \ (vec1 ` S) \ x \ S" by auto - -lemma vec1_vec: "vec1 x = vec x" by (vector vec1_def) - -lemma vec1_add: "vec1(x + y) = vec1 x + vec1 y" by (vector vec1_def) -lemma vec1_sub: "vec1(x - y) = vec1 x - vec1 y" by (vector vec1_def) -lemma vec1_cmul: "vec1(c* x) = c *s vec1 x " by (vector vec1_def) -lemma vec1_neg: "vec1(- x) = - vec1 x " by (vector vec1_def) - -lemma vec1_setsum: assumes fS: "finite S" - shows "vec1(setsum f S) = setsum (vec1 o f) S" - apply (induct rule: finite_induct[OF fS]) - apply (simp add: vec1_vec) - apply (auto simp add: vec1_add) - done - -lemma dest_vec1_lambda: "dest_vec1(\ i. x i) = x 1" - by (simp add: dest_vec1_def) - -lemma dest_vec1_vec: "dest_vec1(vec x) = x" - by (simp add: vec1_vec[symmetric]) - -lemma dest_vec1_add: "dest_vec1(x + y) = dest_vec1 x + dest_vec1 y" - by (metis vec1_dest_vec1 vec1_add) - -lemma dest_vec1_sub: "dest_vec1(x - y) = dest_vec1 x - dest_vec1 y" - by (metis vec1_dest_vec1 vec1_sub) - -lemma dest_vec1_cmul: "dest_vec1(c*sx) = c * dest_vec1 x" - by (metis vec1_dest_vec1 vec1_cmul) - -lemma dest_vec1_neg: "dest_vec1(- x) = - dest_vec1 x" - by (metis vec1_dest_vec1 vec1_neg) - -lemma dest_vec1_0[simp]: "dest_vec1 0 = 0" by (metis vec_0 dest_vec1_vec) - -lemma dest_vec1_sum: assumes fS: "finite S" - shows "dest_vec1(setsum f S) = setsum (dest_vec1 o f) S" - apply (induct rule: finite_induct[OF fS]) - apply (simp add: dest_vec1_vec) - apply (auto simp add: dest_vec1_add) - done - -lemma norm_vec1: "norm(vec1 x) = abs(x)" - by (simp add: vec1_def norm_real) - -lemma dist_vec1: "dist(vec1 x) (vec1 y) = abs(x - y)" - by (simp only: dist_real vec1_component) -lemma abs_dest_vec1: "norm x = \dest_vec1 x\" - by (metis vec1_dest_vec1 norm_vec1) - -lemma linear_vmul_dest_vec1: - fixes f:: "'a::semiring_1^'n \ 'a^1" - shows "linear f \ linear (\x. dest_vec1(f x) *s v)" - unfolding dest_vec1_def - apply (rule linear_vmul_component) - by auto - -lemma linear_from_scalars: - assumes lf: "linear (f::'a::comm_ring_1 ^1 \ 'a^'n)" - shows "f = (\x. dest_vec1 x *s column 1 (matrix f))" - apply (rule ext) - apply (subst matrix_works[OF lf, symmetric]) - apply (auto simp add: Cart_eq matrix_vector_mult_def dest_vec1_def column_def mult_commute UNIV_1) - done - -lemma linear_to_scalars: assumes lf: "linear (f::'a::comm_ring_1 ^'n::finite \ 'a^1)" - shows "f = (\x. vec1(row 1 (matrix f) \ x))" - apply (rule ext) - apply (subst matrix_works[OF lf, symmetric]) - apply (simp add: Cart_eq matrix_vector_mult_def vec1_def row_def dot_def mult_commute forall_1) - done - -lemma dest_vec1_eq_0: "dest_vec1 x = 0 \ x = 0" - by (simp add: dest_vec1_eq[symmetric]) - -lemma setsum_scalars: assumes fS: "finite S" - shows "setsum f S = vec1 (setsum (dest_vec1 o f) S)" - unfolding vec1_setsum[OF fS] by simp - -lemma dest_vec1_wlog_le: "(\(x::'a::linorder ^ 1) y. P x y \ P y x) \ (\x y. dest_vec1 x <= dest_vec1 y ==> P x y) \ P x y" - apply (cases "dest_vec1 x \ dest_vec1 y") - apply simp - apply (subgoal_tac "dest_vec1 y \ dest_vec1 x") - apply (auto) - done - -text{* Pasting vectors. *} - -lemma linear_fstcart: "linear fstcart" - by (auto simp add: linear_def Cart_eq) - -lemma linear_sndcart: "linear sndcart" - by (auto simp add: linear_def Cart_eq) - -lemma fstcart_vec[simp]: "fstcart(vec x) = vec x" - by (simp add: Cart_eq) - -lemma fstcart_add[simp]:"fstcart(x + y) = fstcart (x::'a::{plus,times}^('b + 'c)) + fstcart y" - by (simp add: Cart_eq) - -lemma fstcart_cmul[simp]:"fstcart(c*s x) = c*s fstcart (x::'a::{plus,times}^('b + 'c))" - by (simp add: Cart_eq) - -lemma fstcart_neg[simp]:"fstcart(- x) = - fstcart (x::'a::ring_1^('b + 'c))" - by (simp add: Cart_eq) - -lemma fstcart_sub[simp]:"fstcart(x - y) = fstcart (x::'a::ring_1^('b + 'c)) - fstcart y" - by (simp add: Cart_eq) - -lemma fstcart_setsum: - fixes f:: "'d \ 'a::semiring_1^_" - assumes fS: "finite S" - shows "fstcart (setsum f S) = setsum (\i. fstcart (f i)) S" - by (induct rule: finite_induct[OF fS], simp_all add: vec_0[symmetric] del: vec_0) - -lemma sndcart_vec[simp]: "sndcart(vec x) = vec x" - by (simp add: Cart_eq) - -lemma sndcart_add[simp]:"sndcart(x + y) = sndcart (x::'a::{plus,times}^('b + 'c)) + sndcart y" - by (simp add: Cart_eq) - -lemma sndcart_cmul[simp]:"sndcart(c*s x) = c*s sndcart (x::'a::{plus,times}^('b + 'c))" - by (simp add: Cart_eq) - -lemma sndcart_neg[simp]:"sndcart(- x) = - sndcart (x::'a::ring_1^('b + 'c))" - by (simp add: Cart_eq) - -lemma sndcart_sub[simp]:"sndcart(x - y) = sndcart (x::'a::ring_1^('b + 'c)) - sndcart y" - by (simp add: Cart_eq) - -lemma sndcart_setsum: - fixes f:: "'d \ 'a::semiring_1^_" - assumes fS: "finite S" - shows "sndcart (setsum f S) = setsum (\i. sndcart (f i)) S" - by (induct rule: finite_induct[OF fS], simp_all add: vec_0[symmetric] del: vec_0) - -lemma pastecart_vec[simp]: "pastecart (vec x) (vec x) = vec x" - by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart) - -lemma pastecart_add[simp]:"pastecart (x1::'a::{plus,times}^_) y1 + pastecart x2 y2 = pastecart (x1 + x2) (y1 + y2)" - by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart) - -lemma pastecart_cmul[simp]: "pastecart (c *s (x1::'a::{plus,times}^_)) (c *s y1) = c *s pastecart x1 y1" - by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart) - -lemma pastecart_neg[simp]: "pastecart (- (x::'a::ring_1^_)) (- y) = - pastecart x y" - unfolding vector_sneg_minus1 pastecart_cmul .. - -lemma pastecart_sub: "pastecart (x1::'a::ring_1^_) y1 - pastecart x2 y2 = pastecart (x1 - x2) (y1 - y2)" - by (simp add: diff_def pastecart_neg[symmetric] del: pastecart_neg) - -lemma pastecart_setsum: - fixes f:: "'d \ 'a::semiring_1^_" - assumes fS: "finite S" - shows "pastecart (setsum f S) (setsum g S) = setsum (\i. pastecart (f i) (g i)) S" - by (simp add: pastecart_eq fstcart_setsum[OF fS] sndcart_setsum[OF fS] fstcart_pastecart sndcart_pastecart) - -lemma setsum_Plus: - "\finite A; finite B\ \ - (\x\A <+> B. g x) = (\x\A. g (Inl x)) + (\x\B. g (Inr x))" - unfolding Plus_def - by (subst setsum_Un_disjoint, auto simp add: setsum_reindex) - -lemma setsum_UNIV_sum: - fixes g :: "'a::finite + 'b::finite \ _" - shows "(\x\UNIV. g x) = (\x\UNIV. g (Inl x)) + (\x\UNIV. g (Inr x))" - apply (subst UNIV_Plus_UNIV [symmetric]) - apply (rule setsum_Plus [OF finite finite]) - done - -lemma norm_fstcart: "norm(fstcart x) <= norm (x::real ^('n::finite + 'm::finite))" -proof- - have th0: "norm x = norm (pastecart (fstcart x) (sndcart x))" - by (simp add: pastecart_fst_snd) - have th1: "fstcart x \ fstcart x \ pastecart (fstcart x) (sndcart x) \ pastecart (fstcart x) (sndcart x)" - by (simp add: dot_def setsum_UNIV_sum pastecart_def setsum_nonneg) - then show ?thesis - unfolding th0 - unfolding real_vector_norm_def real_sqrt_le_iff id_def - by (simp add: dot_def) -qed - -lemma dist_fstcart: "dist(fstcart (x::real^_)) (fstcart y) <= dist x y" - unfolding dist_norm by (metis fstcart_sub[symmetric] norm_fstcart) - -lemma norm_sndcart: "norm(sndcart x) <= norm (x::real ^('n::finite + 'm::finite))" -proof- - have th0: "norm x = norm (pastecart (fstcart x) (sndcart x))" - by (simp add: pastecart_fst_snd) - have th1: "sndcart x \ sndcart x \ pastecart (fstcart x) (sndcart x) \ pastecart (fstcart x) (sndcart x)" - by (simp add: dot_def setsum_UNIV_sum pastecart_def setsum_nonneg) - then show ?thesis - unfolding th0 - unfolding real_vector_norm_def real_sqrt_le_iff id_def - by (simp add: dot_def) -qed - -lemma dist_sndcart: "dist(sndcart (x::real^_)) (sndcart y) <= dist x y" - unfolding dist_norm by (metis sndcart_sub[symmetric] norm_sndcart) - -lemma dot_pastecart: "(pastecart (x1::'a::{times,comm_monoid_add}^'n::finite) (x2::'a::{times,comm_monoid_add}^'m::finite)) \ (pastecart y1 y2) = x1 \ y1 + x2 \ y2" - by (simp add: dot_def setsum_UNIV_sum pastecart_def) - -text {* TODO: move to NthRoot *} -lemma sqrt_add_le_add_sqrt: - assumes x: "0 \ x" and y: "0 \ y" - shows "sqrt (x + y) \ sqrt x + sqrt y" -apply (rule power2_le_imp_le) -apply (simp add: real_sum_squared_expand add_nonneg_nonneg x y) -apply (simp add: mult_nonneg_nonneg x y) -apply (simp add: add_nonneg_nonneg x y) -done - -lemma norm_pastecart: "norm (pastecart x y) <= norm x + norm y" - 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). *} - -definition hull :: "'a set set \ 'a set \ 'a set" (infixl "hull" 75) where - "S hull s = Inter {t. t \ S \ s \ t}" - -lemma hull_same: "s \ S \ S hull s = s" - unfolding hull_def by auto - -lemma hull_in: "(\T. T \ S ==> Inter T \ S) ==> (S hull s) \ S" -unfolding hull_def subset_iff by auto - -lemma hull_eq: "(\T. T \ S ==> Inter T \ S) ==> (S hull s) = s \ s \ S" -using hull_same[of s S] hull_in[of S s] by metis - - -lemma hull_hull: "S hull (S hull s) = S hull s" - unfolding hull_def by blast - -lemma hull_subset: "s \ (S hull s)" - unfolding hull_def by blast - -lemma hull_mono: " s \ t ==> (S hull s) \ (S hull t)" - unfolding hull_def by blast - -lemma hull_antimono: "S \ T ==> (T hull s) \ (S hull s)" - unfolding hull_def by blast - -lemma hull_minimal: "s \ t \ t \ S ==> (S hull s) \ t" - unfolding hull_def by blast - -lemma subset_hull: "t \ S ==> S hull s \ t \ s \ t" - unfolding hull_def by blast - -lemma hull_unique: "s \ t \ t \ S \ (\t'. s \ t' \ t' \ S ==> t \ t') - ==> (S hull s = t)" -unfolding hull_def by auto - -lemma hull_induct: "(\x. x\ S \ P x) \ Q {x. P x} \ \x\ Q hull S. P x" - using hull_minimal[of S "{x. P x}" Q] - by (auto simp add: subset_eq Collect_def mem_def) - -lemma hull_inc: "x \ S \ x \ P hull S" by (metis hull_subset subset_eq) - -lemma hull_union_subset: "(S hull s) \ (S hull t) \ (S hull (s \ t))" -unfolding Un_subset_iff by (metis hull_mono Un_upper1 Un_upper2) - -lemma hull_union: assumes T: "\T. T \ S ==> Inter T \ S" - shows "S hull (s \ t) = S hull (S hull s \ S hull t)" -apply rule -apply (rule hull_mono) -unfolding Un_subset_iff -apply (metis hull_subset Un_upper1 Un_upper2 subset_trans) -apply (rule hull_minimal) -apply (metis hull_union_subset) -apply (metis hull_in T) -done - -lemma hull_redundant_eq: "a \ (S hull s) \ (S hull (insert a s) = S hull s)" - unfolding hull_def by blast - -lemma hull_redundant: "a \ (S hull s) ==> (S hull (insert a s) = S hull s)" -by (metis hull_redundant_eq) - -text{* Archimedian properties and useful consequences. *} - -lemma real_arch_simple: "\n. x <= real (n::nat)" - using reals_Archimedean2[of x] apply auto by (rule_tac x="Suc n" in exI, auto) -lemmas real_arch_lt = reals_Archimedean2 - -lemmas real_arch = reals_Archimedean3 - -lemma real_arch_inv: "0 < e \ (\n::nat. n \ 0 \ 0 < inverse (real n) \ inverse (real n) < e)" - using reals_Archimedean - apply (auto simp add: field_simps inverse_positive_iff_positive) - apply (subgoal_tac "inverse (real n) > 0") - apply arith - apply simp - done - -lemma real_pow_lbound: "0 <= x ==> 1 + real n * x <= (1 + x) ^ n" -proof(induct n) - case 0 thus ?case by simp -next - case (Suc n) - hence h: "1 + real n * x \ (1 + x) ^ n" by simp - from h have p: "1 \ (1 + x) ^ n" using Suc.prems by simp - from h have "1 + real n * x + x \ (1 + x) ^ n + x" by simp - also have "\ \ (1 + x) ^ Suc n" apply (subst diff_le_0_iff_le[symmetric]) - apply (simp add: ring_simps) - using mult_left_mono[OF p Suc.prems] by simp - finally show ?case by (simp add: real_of_nat_Suc ring_simps) -qed - -lemma real_arch_pow: assumes x: "1 < (x::real)" shows "\n. y < x^n" -proof- - from x have x0: "x - 1 > 0" by arith - from real_arch[OF x0, rule_format, of y] - obtain n::nat where n:"y < real n * (x - 1)" by metis - from x0 have x00: "x- 1 \ 0" by arith - from real_pow_lbound[OF x00, of n] n - have "y < x^n" by auto - then show ?thesis by metis -qed - -lemma real_arch_pow2: "\n. (x::real) < 2^ n" - using real_arch_pow[of 2 x] by simp - -lemma real_arch_pow_inv: assumes y: "(y::real) > 0" and x1: "x < 1" - shows "\n. x^n < y" -proof- - {assume x0: "x > 0" - from x0 x1 have ix: "1 < 1/x" by (simp add: field_simps) - from real_arch_pow[OF ix, of "1/y"] - obtain n where n: "1/y < (1/x)^n" by blast - then - have ?thesis using y x0 by (auto simp add: field_simps power_divide) } - moreover - {assume "\ x > 0" with y x1 have ?thesis apply auto by (rule exI[where x=1], auto)} - ultimately show ?thesis by metis -qed - -lemma forall_pos_mono: "(\d e::real. d < e \ P d ==> P e) \ (\n::nat. n \ 0 ==> P(inverse(real n))) \ (\e. 0 < e ==> P e)" - by (metis real_arch_inv) - -lemma forall_pos_mono_1: "(\d e::real. d < e \ P d ==> P e) \ (\n. P(inverse(real (Suc n)))) ==> 0 < e ==> P e" - apply (rule forall_pos_mono) - apply auto - apply (atomize) - apply (erule_tac x="n - 1" in allE) - apply auto - done - -lemma real_archimedian_rdiv_eq_0: assumes x0: "x \ 0" and c: "c \ 0" and xc: "\(m::nat)>0. real m * x \ c" - shows "x = 0" -proof- - {assume "x \ 0" with x0 have xp: "x > 0" by arith - from real_arch[OF xp, rule_format, of c] obtain n::nat where n: "c < real n * x" by blast - with xc[rule_format, of n] have "n = 0" by arith - with n c have False by simp} - then show ?thesis by blast -qed - -(* ------------------------------------------------------------------------- *) -(* Relate max and min to sup and inf. *) -(* ------------------------------------------------------------------------- *) - -lemma real_max_rsup: "max x y = rsup {x,y}" -proof- - have f: "finite {x, y}" "{x,y} \ {}" by simp_all - from rsup_finite_le_iff[OF f, of "max x y"] have "rsup {x,y} \ max x y" by simp - moreover - have "max x y \ rsup {x,y}" using rsup_finite_ge_iff[OF f, of "max x y"] - by (simp add: linorder_linear) - ultimately show ?thesis by arith -qed - -lemma real_min_rinf: "min x y = rinf {x,y}" -proof- - have f: "finite {x, y}" "{x,y} \ {}" by simp_all - from rinf_finite_le_iff[OF f, of "min x y"] have "rinf {x,y} \ min x y" - by (simp add: linorder_linear) - moreover - have "min x y \ rinf {x,y}" using rinf_finite_ge_iff[OF f, of "min x y"] - by simp - ultimately show ?thesis by arith -qed - -(* ------------------------------------------------------------------------- *) -(* Geometric progression. *) -(* ------------------------------------------------------------------------- *) - -lemma sum_gp_basic: "((1::'a::{field}) - x) * setsum (\i. x^i) {0 .. n} = (1 - x^(Suc n))" - (is "?lhs = ?rhs") -proof- - {assume x1: "x = 1" hence ?thesis by simp} - moreover - {assume x1: "x\1" - hence x1': "x - 1 \ 0" "1 - x \ 0" "x - 1 = - (1 - x)" "- (1 - x) \ 0" by auto - from geometric_sum[OF x1, of "Suc n", unfolded x1'] - have "(- (1 - x)) * setsum (\i. x^i) {0 .. n} = - (1 - x^(Suc n))" - unfolding atLeastLessThanSuc_atLeastAtMost - using x1' apply (auto simp only: field_simps) - apply (simp add: ring_simps) - done - then have ?thesis by (simp add: ring_simps) } - ultimately show ?thesis by metis -qed - -lemma sum_gp_multiplied: assumes mn: "m <= n" - shows "((1::'a::{field}) - x) * setsum (op ^ x) {m..n} = x^m - x^ Suc n" - (is "?lhs = ?rhs") -proof- - let ?S = "{0..(n - m)}" - from mn have mn': "n - m \ 0" by arith - let ?f = "op + m" - have i: "inj_on ?f ?S" unfolding inj_on_def by auto - have f: "?f ` ?S = {m..n}" - using mn apply (auto simp add: image_iff Bex_def) by arith - have th: "op ^ x o op + m = (\i. x^m * x^i)" - by (rule ext, simp add: power_add power_mult) - from setsum_reindex[OF i, of "op ^ x", unfolded f th setsum_right_distrib[symmetric]] - have "?lhs = x^m * ((1 - x) * setsum (op ^ x) {0..n - m})" by simp - then show ?thesis unfolding sum_gp_basic using mn - by (simp add: ring_simps power_add[symmetric]) -qed - -lemma sum_gp: "setsum (op ^ (x::'a::{field})) {m .. n} = - (if n < m then 0 else if x = 1 then of_nat ((n + 1) - m) - else (x^ m - x^ (Suc n)) / (1 - x))" -proof- - {assume nm: "n < m" hence ?thesis by simp} - moreover - {assume "\ n < m" hence nm: "m \ n" by arith - {assume x: "x = 1" hence ?thesis by simp} - moreover - {assume x: "x \ 1" hence nz: "1 - x \ 0" by simp - from sum_gp_multiplied[OF nm, of x] nz have ?thesis by (simp add: field_simps)} - ultimately have ?thesis by metis - } - ultimately show ?thesis by metis -qed - -lemma sum_gp_offset: "setsum (op ^ (x::'a::{field})) {m .. m+n} = - (if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))" - unfolding sum_gp[of x m "m + n"] power_Suc - by (simp add: ring_simps power_add) - - -subsection{* A bit of linear algebra. *} - -definition "subspace S \ 0 \ S \ (\x\ S. \y \S. x + y \ S) \ (\c. \x \S. c *s x \S )" -definition "span S = (subspace hull S)" -definition "dependent S \ (\a \ S. a \ span(S - {a}))" -abbreviation "independent s == ~(dependent s)" - -(* Closure properties of subspaces. *) - -lemma subspace_UNIV[simp]: "subspace(UNIV)" by (simp add: subspace_def) - -lemma subspace_0: "subspace S ==> 0 \ S" by (metis subspace_def) - -lemma subspace_add: "subspace S \ x \ S \ y \ S ==> x + y \ S" - by (metis subspace_def) - -lemma subspace_mul: "subspace S \ x \ S \ c *s x \ S" - by (metis subspace_def) - -lemma subspace_neg: "subspace S \ (x::'a::ring_1^'n) \ S \ - x \ S" - by (metis vector_sneg_minus1 subspace_mul) - -lemma subspace_sub: "subspace S \ (x::'a::ring_1^'n) \ S \ y \ S \ x - y \ S" - by (metis diff_def subspace_add subspace_neg) - -lemma subspace_setsum: - assumes sA: "subspace A" and fB: "finite B" - and f: "\x\ B. f x \ A" - shows "setsum f B \ A" - using fB f sA - apply(induct rule: finite_induct[OF fB]) - by (simp add: subspace_def sA, auto simp add: sA subspace_add) - -lemma subspace_linear_image: - assumes lf: "linear (f::'a::semiring_1^'n \ _)" and sS: "subspace S" - shows "subspace(f ` S)" - using lf sS linear_0[OF lf] - unfolding linear_def subspace_def - apply (auto simp add: image_iff) - apply (rule_tac x="x + y" in bexI, auto) - apply (rule_tac x="c*s x" in bexI, auto) - done - -lemma subspace_linear_preimage: "linear (f::'a::semiring_1^'n \ _) ==> subspace S ==> subspace {x. f x \ S}" - by (auto simp add: subspace_def linear_def linear_0[of f]) - -lemma subspace_trivial: "subspace {0::'a::semiring_1 ^_}" - by (simp add: subspace_def) - -lemma subspace_inter: "subspace A \ subspace B ==> subspace (A \ B)" - by (simp add: subspace_def) - - -lemma span_mono: "A \ B ==> span A \ span B" - by (metis span_def hull_mono) - -lemma subspace_span: "subspace(span S)" - unfolding span_def - apply (rule hull_in[unfolded mem_def]) - apply (simp only: subspace_def Inter_iff Int_iff subset_eq) - apply auto - apply (erule_tac x="X" in ballE) - apply (simp add: mem_def) - apply blast - apply (erule_tac x="X" in ballE) - apply (erule_tac x="X" in ballE) - apply (erule_tac x="X" in ballE) - apply (clarsimp simp add: mem_def) - apply simp - apply simp - apply simp - apply (erule_tac x="X" in ballE) - apply (erule_tac x="X" in ballE) - apply (simp add: mem_def) - apply simp - apply simp - done - -lemma span_clauses: - "a \ S ==> a \ span S" - "0 \ span S" - "x\ span S \ y \ span S ==> x + y \ span S" - "x \ span S \ c *s x \ span S" - by (metis span_def hull_subset subset_eq subspace_span subspace_def)+ - -lemma span_induct: assumes SP: "\x. x \ S ==> P x" - and P: "subspace P" and x: "x \ span S" shows "P x" -proof- - from SP have SP': "S \ P" by (simp add: mem_def subset_eq) - from P have P': "P \ subspace" by (simp add: mem_def) - from x hull_minimal[OF SP' P', unfolded span_def[symmetric]] - show "P x" by (metis mem_def subset_eq) -qed - -lemma span_empty: "span {} = {(0::'a::semiring_0 ^ 'n)}" - apply (simp add: span_def) - apply (rule hull_unique) - apply (auto simp add: mem_def subspace_def) - unfolding mem_def[of "0::'a^'n", symmetric] - apply simp - done - -lemma independent_empty: "independent {}" - by (simp add: dependent_def) - -lemma independent_mono: "independent A \ B \ A ==> independent B" - apply (clarsimp simp add: dependent_def span_mono) - apply (subgoal_tac "span (B - {a}) \ span (A - {a})") - apply force - apply (rule span_mono) - apply auto - done - -lemma span_subspace: "A \ B \ B \ span A \ subspace B \ span A = B" - by (metis order_antisym span_def hull_minimal mem_def) - -lemma span_induct': assumes SP: "\x \ S. P x" - and P: "subspace P" shows "\x \ span S. P x" - using span_induct SP P by blast - -inductive span_induct_alt_help for S:: "'a::semiring_1^'n \ bool" - where - span_induct_alt_help_0: "span_induct_alt_help S 0" - | span_induct_alt_help_S: "x \ S \ span_induct_alt_help S z \ span_induct_alt_help S (c *s x + z)" - -lemma span_induct_alt': - assumes h0: "h (0::'a::semiring_1^'n)" and hS: "\c x y. x \ S \ h y \ h (c*s x + y)" shows "\x \ span S. h x" -proof- - {fix x:: "'a^'n" assume x: "span_induct_alt_help S x" - have "h x" - apply (rule span_induct_alt_help.induct[OF x]) - apply (rule h0) - apply (rule hS, assumption, assumption) - done} - note th0 = this - {fix x assume x: "x \ span S" - - have "span_induct_alt_help S x" - proof(rule span_induct[where x=x and S=S]) - show "x \ span S" using x . - next - fix x assume xS : "x \ S" - from span_induct_alt_help_S[OF xS span_induct_alt_help_0, of 1] - show "span_induct_alt_help S x" by simp - next - have "span_induct_alt_help S 0" by (rule span_induct_alt_help_0) - moreover - {fix x y assume h: "span_induct_alt_help S x" "span_induct_alt_help S y" - from h - have "span_induct_alt_help S (x + y)" - apply (induct rule: span_induct_alt_help.induct) - apply simp - unfolding add_assoc - apply (rule span_induct_alt_help_S) - apply assumption - apply simp - done} - moreover - {fix c x assume xt: "span_induct_alt_help S x" - then have "span_induct_alt_help S (c*s x)" - apply (induct rule: span_induct_alt_help.induct) - apply (simp add: span_induct_alt_help_0) - apply (simp add: vector_smult_assoc vector_add_ldistrib) - apply (rule span_induct_alt_help_S) - apply assumption - apply simp - done - } - ultimately show "subspace (span_induct_alt_help S)" - unfolding subspace_def mem_def Ball_def by blast - qed} - with th0 show ?thesis by blast -qed - -lemma span_induct_alt: - assumes h0: "h (0::'a::semiring_1^'n)" and hS: "\c x y. x \ S \ h y \ h (c*s x + y)" and x: "x \ span S" - shows "h x" -using span_induct_alt'[of h S] h0 hS x by blast - -(* Individual closure properties. *) - -lemma span_superset: "x \ S ==> x \ span S" by (metis span_clauses) - -lemma span_0: "0 \ span S" by (metis subspace_span subspace_0) - -lemma span_add: "x \ span S \ y \ span S ==> x + y \ span S" - by (metis subspace_add subspace_span) - -lemma span_mul: "x \ span S ==> (c *s x) \ span S" - by (metis subspace_span subspace_mul) - -lemma span_neg: "x \ span S ==> - (x::'a::ring_1^'n) \ span S" - by (metis subspace_neg subspace_span) - -lemma span_sub: "(x::'a::ring_1^'n) \ span S \ y \ span S ==> x - y \ span S" - by (metis subspace_span subspace_sub) - -lemma span_setsum: "finite A \ \x \ A. f x \ span S ==> setsum f A \ span S" - apply (rule subspace_setsum) - by (metis subspace_span subspace_setsum)+ - -lemma span_add_eq: "(x::'a::ring_1^'n) \ span S \ x + y \ span S \ y \ span S" - apply (auto simp only: span_add span_sub) - apply (subgoal_tac "(x + y) - x \ span S", simp) - by (simp only: span_add span_sub) - -(* Mapping under linear image. *) - -lemma span_linear_image: assumes lf: "linear (f::'a::semiring_1 ^ 'n => _)" - shows "span (f ` S) = f ` (span S)" -proof- - {fix x - assume x: "x \ span (f ` S)" - have "x \ f ` span S" - apply (rule span_induct[where x=x and S = "f ` S"]) - apply (clarsimp simp add: image_iff) - apply (frule span_superset) - apply blast - apply (simp only: mem_def) - apply (rule subspace_linear_image[OF lf]) - apply (rule subspace_span) - apply (rule x) - done} - moreover - {fix x assume x: "x \ span S" - have th0:"(\a. f a \ span (f ` S)) = {x. f x \ span (f ` S)}" apply (rule set_ext) - unfolding mem_def Collect_def .. - have "f x \ span (f ` S)" - apply (rule span_induct[where S=S]) - apply (rule span_superset) - apply simp - apply (subst th0) - apply (rule subspace_linear_preimage[OF lf subspace_span, of "f ` S"]) - apply (rule x) - done} - ultimately show ?thesis by blast -qed - -(* The key breakdown property. *) - -lemma span_breakdown: - assumes bS: "(b::'a::ring_1 ^ 'n) \ S" and aS: "a \ span S" - shows "\k. a - k*s b \ span (S - {b})" (is "?P a") -proof- - {fix x assume xS: "x \ S" - {assume ab: "x = b" - then have "?P x" - apply simp - apply (rule exI[where x="1"], simp) - by (rule span_0)} - moreover - {assume ab: "x \ b" - then have "?P x" using xS - apply - - apply (rule exI[where x=0]) - apply (rule span_superset) - by simp} - ultimately have "?P x" by blast} - moreover have "subspace ?P" - unfolding subspace_def - apply auto - apply (simp add: mem_def) - apply (rule exI[where x=0]) - using span_0[of "S - {b}"] - apply (simp add: mem_def) - apply (clarsimp simp add: mem_def) - apply (rule_tac x="k + ka" in exI) - apply (subgoal_tac "x + y - (k + ka) *s b = (x - k*s b) + (y - ka *s b)") - apply (simp only: ) - apply (rule span_add[unfolded mem_def]) - apply assumption+ - apply (vector ring_simps) - apply (clarsimp simp add: mem_def) - apply (rule_tac x= "c*k" in exI) - apply (subgoal_tac "c *s x - (c * k) *s b = c*s (x - k*s b)") - apply (simp only: ) - apply (rule span_mul[unfolded mem_def]) - apply assumption - by (vector ring_simps) - ultimately show "?P a" using aS span_induct[where S=S and P= "?P"] by metis -qed - -lemma span_breakdown_eq: - "(x::'a::ring_1^'n) \ span (insert a S) \ (\k. (x - k *s a) \ span S)" (is "?lhs \ ?rhs") -proof- - {assume x: "x \ span (insert a S)" - from x span_breakdown[of "a" "insert a S" "x"] - have ?rhs apply clarsimp - apply (rule_tac x= "k" in exI) - apply (rule set_rev_mp[of _ "span (S - {a})" _]) - apply assumption - apply (rule span_mono) - apply blast - done} - moreover - { fix k assume k: "x - k *s a \ span S" - have eq: "x = (x - k *s a) + k *s a" by vector - have "(x - k *s a) + k *s a \ span (insert a S)" - apply (rule span_add) - apply (rule set_rev_mp[of _ "span S" _]) - apply (rule k) - apply (rule span_mono) - apply blast - apply (rule span_mul) - apply (rule span_superset) - apply blast - done - then have ?lhs using eq by metis} - ultimately show ?thesis by blast -qed - -(* Hence some "reversal" results.*) - -lemma in_span_insert: - assumes a: "(a::'a::field^'n) \ span (insert b S)" and na: "a \ span S" - shows "b \ span (insert a S)" -proof- - from span_breakdown[of b "insert b S" a, OF insertI1 a] - obtain k where k: "a - k*s b \ span (S - {b})" by auto - {assume k0: "k = 0" - with k have "a \ span S" - apply (simp) - apply (rule set_rev_mp) - apply assumption - apply (rule span_mono) - apply blast - done - with na have ?thesis by blast} - moreover - {assume k0: "k \ 0" - have eq: "b = (1/k) *s a - ((1/k) *s a - b)" by vector - from k0 have eq': "(1/k) *s (a - k*s b) = (1/k) *s a - b" - by (vector field_simps) - from k have "(1/k) *s (a - k*s b) \ span (S - {b})" - by (rule span_mul) - hence th: "(1/k) *s a - b \ span (S - {b})" - unfolding eq' . - - from k - have ?thesis - apply (subst eq) - apply (rule span_sub) - apply (rule span_mul) - apply (rule span_superset) - apply blast - apply (rule set_rev_mp) - apply (rule th) - apply (rule span_mono) - using na by blast} - ultimately show ?thesis by blast -qed - -lemma in_span_delete: - assumes a: "(a::'a::field^'n) \ span S" - and na: "a \ span (S-{b})" - shows "b \ span (insert a (S - {b}))" - apply (rule in_span_insert) - apply (rule set_rev_mp) - apply (rule a) - apply (rule span_mono) - apply blast - apply (rule na) - done - -(* Transitivity property. *) - -lemma span_trans: - assumes x: "(x::'a::ring_1^'n) \ span S" and y: "y \ span (insert x S)" - shows "y \ span S" -proof- - from span_breakdown[of x "insert x S" y, OF insertI1 y] - obtain k where k: "y -k*s x \ span (S - {x})" by auto - have eq: "y = (y - k *s x) + k *s x" by vector - show ?thesis - apply (subst eq) - apply (rule span_add) - apply (rule set_rev_mp) - apply (rule k) - apply (rule span_mono) - apply blast - apply (rule span_mul) - by (rule x) -qed - -(* ------------------------------------------------------------------------- *) -(* An explicit expansion is sometimes needed. *) -(* ------------------------------------------------------------------------- *) - -lemma span_explicit: - "span P = {y::'a::semiring_1^'n. \S u. finite S \ S \ P \ setsum (\v. u v *s v) S = y}" - (is "_ = ?E" is "_ = {y. ?h y}" is "_ = {y. \S u. ?Q S u y}") -proof- - {fix x assume x: "x \ ?E" - then obtain S u where fS: "finite S" and SP: "S\P" and u: "setsum (\v. u v *s v) S = x" - by blast - have "x \ span P" - unfolding u[symmetric] - apply (rule span_setsum[OF fS]) - using span_mono[OF SP] - by (auto intro: span_superset span_mul)} - moreover - have "\x \ span P. x \ ?E" - unfolding mem_def Collect_def - proof(rule span_induct_alt') - show "?h 0" - apply (rule exI[where x="{}"]) by simp - next - fix c x y - assume x: "x \ P" and hy: "?h y" - from hy obtain S u where fS: "finite S" and SP: "S\P" - and u: "setsum (\v. u v *s v) S = y" by blast - let ?S = "insert x S" - let ?u = "\y. if y = x then (if x \ S then u y + c else c) - else u y" - from fS SP x have th0: "finite (insert x S)" "insert x S \ P" by blast+ - {assume xS: "x \ S" - have S1: "S = (S - {x}) \ {x}" - and Sss:"finite (S - {x})" "finite {x}" "(S -{x}) \ {x} = {}" using xS fS by auto - have "setsum (\v. ?u v *s v) ?S =(\v\S - {x}. u v *s v) + (u x + c) *s x" - using xS - by (simp add: setsum_Un_disjoint[OF Sss, unfolded S1[symmetric]] - setsum_clauses(2)[OF fS] cong del: if_weak_cong) - also have "\ = (\v\S. u v *s v) + c *s x" - apply (simp add: setsum_Un_disjoint[OF Sss, unfolded S1[symmetric]]) - by (vector ring_simps) - also have "\ = c*s x + y" - by (simp add: add_commute u) - finally have "setsum (\v. ?u v *s v) ?S = c*s x + y" . - then have "?Q ?S ?u (c*s x + y)" using th0 by blast} - moreover - {assume xS: "x \ S" - have th00: "(\v\S. (if v = x then c else u v) *s v) = y" - unfolding u[symmetric] - apply (rule setsum_cong2) - using xS by auto - have "?Q ?S ?u (c*s x + y)" using fS xS th0 - by (simp add: th00 setsum_clauses add_commute cong del: if_weak_cong)} - ultimately have "?Q ?S ?u (c*s x + y)" - by (cases "x \ S", simp, simp) - then show "?h (c*s x + y)" - apply - - apply (rule exI[where x="?S"]) - apply (rule exI[where x="?u"]) by metis - qed - ultimately show ?thesis by blast -qed - -lemma dependent_explicit: - "dependent P \ (\S u. finite S \ S \ P \ (\(v::'a::{idom,field}^'n) \S. u v \ 0 \ setsum (\v. u v *s v) S = 0))" (is "?lhs = ?rhs") -proof- - {assume dP: "dependent P" - then obtain a S u where aP: "a \ P" and fS: "finite S" - and SP: "S \ P - {a}" and ua: "setsum (\v. u v *s v) S = a" - unfolding dependent_def span_explicit by blast - let ?S = "insert a S" - let ?u = "\y. if y = a then - 1 else u y" - let ?v = a - from aP SP have aS: "a \ S" by blast - from fS SP aP have th0: "finite ?S" "?S \ P" "?v \ ?S" "?u ?v \ 0" by auto - have s0: "setsum (\v. ?u v *s v) ?S = 0" - using fS aS - apply (simp add: vector_smult_lneg vector_smult_lid setsum_clauses ring_simps ) - apply (subst (2) ua[symmetric]) - apply (rule setsum_cong2) - by auto - with th0 have ?rhs - apply - - apply (rule exI[where x= "?S"]) - apply (rule exI[where x= "?u"]) - by clarsimp} - moreover - {fix S u v assume fS: "finite S" - and SP: "S \ P" and vS: "v \ S" and uv: "u v \ 0" - and u: "setsum (\v. u v *s v) S = 0" - let ?a = v - let ?S = "S - {v}" - let ?u = "\i. (- u i) / u v" - have th0: "?a \ P" "finite ?S" "?S \ P" using fS SP vS by auto - have "setsum (\v. ?u v *s v) ?S = setsum (\v. (- (inverse (u ?a))) *s (u v *s v)) S - ?u v *s v" - using fS vS uv - by (simp add: setsum_diff1 vector_smult_lneg divide_inverse - vector_smult_assoc field_simps) - also have "\ = ?a" - unfolding setsum_cmul u - using uv by (simp add: vector_smult_lneg) - finally have "setsum (\v. ?u v *s v) ?S = ?a" . - with th0 have ?lhs - unfolding dependent_def span_explicit - apply - - apply (rule bexI[where x= "?a"]) - apply simp_all - apply (rule exI[where x= "?S"]) - by auto} - ultimately show ?thesis by blast -qed - - -lemma span_finite: - assumes fS: "finite S" - shows "span S = {(y::'a::semiring_1^'n). \u. setsum (\v. u v *s v) S = y}" - (is "_ = ?rhs") -proof- - {fix y assume y: "y \ span S" - from y obtain S' u where fS': "finite S'" and SS': "S' \ S" and - u: "setsum (\v. u v *s v) S' = y" unfolding span_explicit by blast - let ?u = "\x. if x \ S' then u x else 0" - from setsum_restrict_set[OF fS, of "\v. u v *s v" S', symmetric] SS' - have "setsum (\v. ?u v *s v) S = setsum (\v. u v *s v) S'" - unfolding cond_value_iff cond_application_beta - by (simp add: cond_value_iff inf_absorb2 cong del: if_weak_cong) - hence "setsum (\v. ?u v *s v) S = y" by (metis u) - hence "y \ ?rhs" by auto} - moreover - {fix y u assume u: "setsum (\v. u v *s v) S = y" - then have "y \ span S" using fS unfolding span_explicit by auto} - ultimately show ?thesis by blast -qed - - -(* Standard bases are a spanning set, and obviously finite. *) - -lemma span_stdbasis:"span {basis i :: 'a::ring_1^'n::finite | i. i \ (UNIV :: 'n set)} = UNIV" -apply (rule set_ext) -apply auto -apply (subst basis_expansion[symmetric]) -apply (rule span_setsum) -apply simp -apply auto -apply (rule span_mul) -apply (rule span_superset) -apply (auto simp add: Collect_def mem_def) -done - -lemma has_size_stdbasis: "{basis i ::real ^'n::finite | i. i \ (UNIV :: 'n set)} hassize CARD('n)" (is "?S hassize ?n") -proof- - have eq: "?S = basis ` UNIV" by blast - show ?thesis unfolding eq - apply (rule hassize_image_inj[OF basis_inj]) - by (simp add: hassize_def) -qed - -lemma finite_stdbasis: "finite {basis i ::real^'n::finite |i. i\ (UNIV:: 'n set)}" - using has_size_stdbasis[unfolded hassize_def] - .. - -lemma card_stdbasis: "card {basis i ::real^'n::finite |i. i\ (UNIV :: 'n set)} = CARD('n)" - using has_size_stdbasis[unfolded hassize_def] - .. - -lemma independent_stdbasis_lemma: - assumes x: "(x::'a::semiring_1 ^ 'n) \ span (basis ` S)" - and iS: "i \ S" - shows "(x$i) = 0" -proof- - let ?U = "UNIV :: 'n set" - let ?B = "basis ` S" - let ?P = "\(x::'a^'n). \i\ ?U. i \ S \ x$i =0" - {fix x::"'a^'n" assume xS: "x\ ?B" - from xS have "?P x" by auto} - moreover - have "subspace ?P" - by (auto simp add: subspace_def Collect_def mem_def) - ultimately show ?thesis - using x span_induct[of ?B ?P x] iS by blast -qed - -lemma independent_stdbasis: "independent {basis i ::real^'n::finite |i. i\ (UNIV :: 'n set)}" -proof- - let ?I = "UNIV :: 'n set" - let ?b = "basis :: _ \ real ^'n" - let ?B = "?b ` ?I" - have eq: "{?b i|i. i \ ?I} = ?B" - by auto - {assume d: "dependent ?B" - then obtain k where k: "k \ ?I" "?b k \ span (?B - {?b k})" - unfolding dependent_def by auto - have eq1: "?B - {?b k} = ?B - ?b ` {k}" by simp - have eq2: "?B - {?b k} = ?b ` (?I - {k})" - unfolding eq1 - apply (rule inj_on_image_set_diff[symmetric]) - apply (rule basis_inj) using k(1) by auto - from k(2) have th0: "?b k \ span (?b ` (?I - {k}))" unfolding eq2 . - from independent_stdbasis_lemma[OF th0, of k, simplified] - have False by simp} - then show ?thesis unfolding eq dependent_def .. -qed - -(* This is useful for building a basis step-by-step. *) - -lemma independent_insert: - "independent(insert (a::'a::field ^'n) S) \ - (if a \ S then independent S - else independent S \ a \ span S)" (is "?lhs \ ?rhs") -proof- - {assume aS: "a \ S" - hence ?thesis using insert_absorb[OF aS] by simp} - moreover - {assume aS: "a \ S" - {assume i: ?lhs - then have ?rhs using aS - apply simp - apply (rule conjI) - apply (rule independent_mono) - apply assumption - apply blast - by (simp add: dependent_def)} - moreover - {assume i: ?rhs - have ?lhs using i aS - apply simp - apply (auto simp add: dependent_def) - apply (case_tac "aa = a", auto) - apply (subgoal_tac "insert a S - {aa} = insert a (S - {aa})") - apply simp - apply (subgoal_tac "a \ span (insert aa (S - {aa}))") - apply (subgoal_tac "insert aa (S - {aa}) = S") - apply simp - apply blast - apply (rule in_span_insert) - apply assumption - apply blast - apply blast - done} - ultimately have ?thesis by blast} - ultimately show ?thesis by blast -qed - -(* The degenerate case of the Exchange Lemma. *) - -lemma mem_delete: "x \ (A - {a}) \ x \ a \ x \ A" - by blast - -lemma span_span: "span (span A) = span A" - unfolding span_def hull_hull .. - -lemma span_inc: "S \ span S" - by (metis subset_eq span_superset) - -lemma spanning_subset_independent: - assumes BA: "B \ A" and iA: "independent (A::('a::field ^'n) set)" - and AsB: "A \ span B" - shows "A = B" -proof - from BA show "B \ A" . -next - from span_mono[OF BA] span_mono[OF AsB] - have sAB: "span A = span B" unfolding span_span by blast - - {fix x assume x: "x \ A" - from iA have th0: "x \ span (A - {x})" - unfolding dependent_def using x by blast - from x have xsA: "x \ span A" by (blast intro: span_superset) - have "A - {x} \ A" by blast - hence th1:"span (A - {x}) \ span A" by (metis span_mono) - {assume xB: "x \ B" - from xB BA have "B \ A -{x}" by blast - hence "span B \ span (A - {x})" by (metis span_mono) - with th1 th0 sAB have "x \ span A" by blast - with x have False by (metis span_superset)} - then have "x \ B" by blast} - then show "A \ B" by blast -qed - -(* The general case of the Exchange Lemma, the key to what follows. *) - -lemma exchange_lemma: - assumes f:"finite (t:: ('a::field^'n) set)" and i: "independent s" - and sp:"s \ span t" - shows "\t'. (t' hassize card t) \ s \ t' \ t' \ s \ t \ s \ span t'" -using f i sp -proof(induct c\"card(t - s)" arbitrary: s t rule: nat_less_induct) - fix n:: nat and s t :: "('a ^'n) set" - assume H: " \m(x:: ('a ^'n) set) xa. - finite xa \ - independent x \ - x \ span xa \ - m = card (xa - x) \ - (\t'. (t' hassize card xa) \ - x \ t' \ t' \ x \ xa \ x \ span t')" - and ft: "finite t" and s: "independent s" and sp: "s \ span t" - and n: "n = card (t - s)" - let ?P = "\t'. (t' hassize card t) \ s \ t' \ t' \ s \ t \ s \ span t'" - let ?ths = "\t'. ?P t'" - {assume st: "s \ t" - from st ft span_mono[OF st] have ?ths apply - apply (rule exI[where x=t]) - by (auto simp add: hassize_def intro: span_superset)} - moreover - {assume st: "t \ s" - - from spanning_subset_independent[OF st s sp] - st ft span_mono[OF st] have ?ths apply - apply (rule exI[where x=t]) - by (auto simp add: hassize_def intro: span_superset)} - moreover - {assume st: "\ s \ t" "\ t \ s" - from st(2) obtain b where b: "b \ t" "b \ s" by blast - from b have "t - {b} - s \ t - s" by blast - then have cardlt: "card (t - {b} - s) < n" using n ft - by (auto intro: psubset_card_mono) - from b ft have ct0: "card t \ 0" by auto - {assume stb: "s \ span(t -{b})" - from ft have ftb: "finite (t -{b})" by auto - from H[rule_format, OF cardlt ftb s stb] - obtain u where u: "u hassize card (t-{b})" "s \ u" "u \ s \ (t - {b})" "s \ span u" by blast - let ?w = "insert b u" - have th0: "s \ insert b u" using u by blast - from u(3) b have "u \ s \ t" by blast - then have th1: "insert b u \ s \ t" using u b by blast - have bu: "b \ u" using b u by blast - from u(1) have fu: "finite u" by (simp add: hassize_def) - from u(1) ft b have "u hassize (card t - 1)" by auto - then - have th2: "insert b u hassize card t" - using card_insert_disjoint[OF fu bu] ct0 by (auto simp add: hassize_def) - from u(4) have "s \ span u" . - also have "\ \ span (insert b u)" apply (rule span_mono) by blast - finally have th3: "s \ span (insert b u)" . from th0 th1 th2 th3 have th: "?P ?w" by blast - from th have ?ths by blast} - moreover - {assume stb: "\ s \ span(t -{b})" - from stb obtain a where a: "a \ s" "a \ span (t - {b})" by blast - have ab: "a \ b" using a b by blast - have at: "a \ t" using a ab span_superset[of a "t- {b}"] by auto - have mlt: "card ((insert a (t - {b})) - s) < n" - using cardlt ft n a b by auto - have ft': "finite (insert a (t - {b}))" using ft by auto - {fix x assume xs: "x \ s" - have t: "t \ (insert b (insert a (t -{b})))" using b by auto - from b(1) have "b \ span t" by (simp add: span_superset) - have bs: "b \ span (insert a (t - {b}))" - by (metis in_span_delete a sp mem_def subset_eq) - from xs sp have "x \ span t" by blast - with span_mono[OF t] - have x: "x \ span (insert b (insert a (t - {b})))" .. - from span_trans[OF bs x] have "x \ span (insert a (t - {b}))" .} - then have sp': "s \ span (insert a (t - {b}))" by blast - - from H[rule_format, OF mlt ft' s sp' refl] obtain u where - u: "u hassize card (insert a (t -{b}))" "s \ u" "u \ s \ insert a (t -{b})" - "s \ span u" by blast - from u a b ft at ct0 have "?P u" by (auto simp add: hassize_def) - then have ?ths by blast } - ultimately have ?ths by blast - } - ultimately - show ?ths by blast -qed - -(* This implies corresponding size bounds. *) - -lemma independent_span_bound: - assumes f: "finite t" and i: "independent (s::('a::field^'n) set)" and sp:"s \ span t" - shows "finite s \ card s \ card t" - by (metis exchange_lemma[OF f i sp] hassize_def finite_subset card_mono) - - -lemma finite_Atleast_Atmost_nat[simp]: "finite {f x |x. x\ (UNIV::'a::finite set)}" -proof- - have eq: "{f x |x. x\ UNIV} = f ` UNIV" by auto - show ?thesis unfolding eq - apply (rule finite_imageI) - apply (rule finite) - done -qed - - -lemma independent_bound: - fixes S:: "(real^'n::finite) set" - shows "independent S \ finite S \ card S <= CARD('n)" - apply (subst card_stdbasis[symmetric]) - apply (rule independent_span_bound) - apply (rule finite_Atleast_Atmost_nat) - apply assumption - unfolding span_stdbasis - apply (rule subset_UNIV) - done - -lemma dependent_biggerset: "(finite (S::(real ^'n::finite) set) ==> card S > CARD('n)) ==> dependent S" - by (metis independent_bound not_less) - -(* Hence we can create a maximal independent subset. *) - -lemma maximal_independent_subset_extend: - assumes sv: "(S::(real^'n::finite) set) \ V" and iS: "independent S" - shows "\B. S \ B \ B \ V \ independent B \ V \ span B" - using sv iS -proof(induct d\ "CARD('n) - card S" arbitrary: S rule: nat_less_induct) - fix n and S:: "(real^'n) set" - assume H: "\mS \ V. independent S \ m = CARD('n) - card S \ - (\B. S \ B \ B \ V \ independent B \ V \ span B)" - and sv: "S \ V" and i: "independent S" and n: "n = CARD('n) - card S" - let ?P = "\B. S \ B \ B \ V \ independent B \ V \ span B" - let ?ths = "\x. ?P x" - let ?d = "CARD('n)" - {assume "V \ span S" - then have ?ths using sv i by blast } - moreover - {assume VS: "\ V \ span S" - from VS obtain a where a: "a \ V" "a \ span S" by blast - from a have aS: "a \ S" by (auto simp add: span_superset) - have th0: "insert a S \ V" using a sv by blast - from independent_insert[of a S] i a - have th1: "independent (insert a S)" by auto - have mlt: "?d - card (insert a S) < n" - using aS a n independent_bound[OF th1] - by auto - - from H[rule_format, OF mlt th0 th1 refl] - obtain B where B: "insert a S \ B" "B \ V" "independent B" " V \ span B" - by blast - from B have "?P B" by auto - then have ?ths by blast} - ultimately show ?ths by blast -qed - -lemma maximal_independent_subset: - "\(B:: (real ^'n::finite) set). B\ V \ independent B \ V \ span B" - by (metis maximal_independent_subset_extend[of "{}:: (real ^'n) set"] empty_subsetI independent_empty) - -(* Notion of dimension. *) - -definition "dim V = (SOME n. \B. B \ V \ independent B \ V \ span B \ (B hassize n))" - -lemma basis_exists: "\B. (B :: (real ^'n::finite) set) \ V \ independent B \ V \ span B \ (B hassize dim V)" -unfolding dim_def some_eq_ex[of "\n. \B. B \ V \ independent B \ V \ span B \ (B hassize n)"] -unfolding hassize_def -using maximal_independent_subset[of V] independent_bound -by auto - -(* Consequences of independence or spanning for cardinality. *) - -lemma independent_card_le_dim: "(B::(real ^'n::finite) set) \ V \ independent B \ finite B \ card B \ dim V" -by (metis basis_exists[of V] independent_span_bound[where ?'a=real] hassize_def subset_trans) - -lemma span_card_ge_dim: "(B::(real ^'n::finite) set) \ V \ V \ span B \ finite B \ dim V \ card B" - by (metis basis_exists[of V] independent_span_bound hassize_def subset_trans) - -lemma basis_card_eq_dim: - "B \ (V:: (real ^'n::finite) set) \ V \ span B \ independent B \ finite B \ card B = dim V" - by (metis order_eq_iff independent_card_le_dim span_card_ge_dim independent_mono) - -lemma dim_unique: "(B::(real ^'n::finite) set) \ V \ V \ span B \ independent B \ B hassize n \ dim V = n" - by (metis basis_card_eq_dim hassize_def) - -(* More lemmas about dimension. *) - -lemma dim_univ: "dim (UNIV :: (real^'n::finite) set) = CARD('n)" - apply (rule dim_unique[of "{basis i |i. i\ (UNIV :: 'n set)}"]) - by (auto simp only: span_stdbasis has_size_stdbasis independent_stdbasis) - -lemma dim_subset: - "(S:: (real ^'n::finite) set) \ T \ dim S \ dim T" - using basis_exists[of T] basis_exists[of S] - by (metis independent_span_bound[where ?'a = real and ?'n = 'n] subset_eq hassize_def) - -lemma dim_subset_univ: "dim (S:: (real^'n::finite) set) \ CARD('n)" - by (metis dim_subset subset_UNIV dim_univ) - -(* Converses to those. *) - -lemma card_ge_dim_independent: - assumes BV:"(B::(real ^'n::finite) set) \ V" and iB:"independent B" and dVB:"dim V \ card B" - shows "V \ span B" -proof- - {fix a assume aV: "a \ V" - {assume aB: "a \ span B" - then have iaB: "independent (insert a B)" using iB aV BV by (simp add: independent_insert) - from aV BV have th0: "insert a B \ V" by blast - from aB have "a \B" by (auto simp add: span_superset) - with independent_card_le_dim[OF th0 iaB] dVB have False by auto} - then have "a \ span B" by blast} - then show ?thesis by blast -qed - -lemma card_le_dim_spanning: - assumes BV: "(B:: (real ^'n::finite) set) \ V" and VB: "V \ span B" - and fB: "finite B" and dVB: "dim V \ card B" - shows "independent B" -proof- - {fix a assume a: "a \ B" "a \ span (B -{a})" - from a fB have c0: "card B \ 0" by auto - from a fB have cb: "card (B -{a}) = card B - 1" by auto - from BV a have th0: "B -{a} \ V" by blast - {fix x assume x: "x \ V" - from a have eq: "insert a (B -{a}) = B" by blast - from x VB have x': "x \ span B" by blast - from span_trans[OF a(2), unfolded eq, OF x'] - have "x \ span (B -{a})" . } - then have th1: "V \ span (B -{a})" by blast - have th2: "finite (B -{a})" using fB by auto - from span_card_ge_dim[OF th0 th1 th2] - have c: "dim V \ card (B -{a})" . - from c c0 dVB cb have False by simp} - then show ?thesis unfolding dependent_def by blast -qed - -lemma card_eq_dim: "(B:: (real ^'n::finite) set) \ V \ B hassize dim V \ independent B \ V \ span B" - by (metis hassize_def order_eq_iff card_le_dim_spanning - card_ge_dim_independent) - -(* ------------------------------------------------------------------------- *) -(* More general size bound lemmas. *) -(* ------------------------------------------------------------------------- *) - -lemma independent_bound_general: - "independent (S:: (real^'n::finite) set) \ finite S \ card S \ dim S" - by (metis independent_card_le_dim independent_bound subset_refl) - -lemma dependent_biggerset_general: "(finite (S:: (real^'n::finite) set) \ card S > dim S) \ dependent S" - using independent_bound_general[of S] by (metis linorder_not_le) - -lemma dim_span: "dim (span (S:: (real ^'n::finite) set)) = dim S" -proof- - have th0: "dim S \ dim (span S)" - by (auto simp add: subset_eq intro: dim_subset span_superset) - from basis_exists[of S] - obtain B where B: "B \ S" "independent B" "S \ span B" "B hassize dim S" by blast - from B have fB: "finite B" "card B = dim S" unfolding hassize_def by blast+ - have bSS: "B \ span S" using B(1) by (metis subset_eq span_inc) - have sssB: "span S \ span B" using span_mono[OF B(3)] by (simp add: span_span) - from span_card_ge_dim[OF bSS sssB fB(1)] th0 show ?thesis - using fB(2) by arith -qed - -lemma subset_le_dim: "(S:: (real ^'n::finite) set) \ span T \ dim S \ dim T" - by (metis dim_span dim_subset) - -lemma span_eq_dim: "span (S:: (real ^'n::finite) set) = span T ==> dim S = dim T" - by (metis dim_span) - -lemma spans_image: - assumes lf: "linear (f::'a::semiring_1^'n \ _)" and VB: "V \ span B" - shows "f ` V \ span (f ` B)" - unfolding span_linear_image[OF lf] - by (metis VB image_mono) - -lemma dim_image_le: - fixes f :: "real^'n::finite \ real^'m::finite" - assumes lf: "linear f" shows "dim (f ` S) \ dim (S:: (real ^'n::finite) set)" -proof- - from basis_exists[of S] obtain B where - B: "B \ S" "independent B" "S \ span B" "B hassize dim S" by blast - from B have fB: "finite B" "card B = dim S" unfolding hassize_def by blast+ - have "dim (f ` S) \ card (f ` B)" - apply (rule span_card_ge_dim) - using lf B fB by (auto simp add: span_linear_image spans_image subset_image_iff) - also have "\ \ dim S" using card_image_le[OF fB(1)] fB by simp - finally show ?thesis . -qed - -(* Relation between bases and injectivity/surjectivity of map. *) - -lemma spanning_surjective_image: - assumes us: "UNIV \ span (S:: ('a::semiring_1 ^'n) set)" - and lf: "linear f" and sf: "surj f" - shows "UNIV \ span (f ` S)" -proof- - have "UNIV \ f ` UNIV" using sf by (auto simp add: surj_def) - also have " \ \ span (f ` S)" using spans_image[OF lf us] . -finally show ?thesis . -qed - -lemma independent_injective_image: - assumes iS: "independent (S::('a::semiring_1^'n) set)" and lf: "linear f" and fi: "inj f" - shows "independent (f ` S)" -proof- - {fix a assume a: "a \ S" "f a \ span (f ` S - {f a})" - have eq: "f ` S - {f a} = f ` (S - {a})" using fi - by (auto simp add: inj_on_def) - from a have "f a \ f ` span (S -{a})" - unfolding eq span_linear_image[OF lf, of "S - {a}"] by blast - hence "a \ span (S -{a})" using fi by (auto simp add: inj_on_def) - with a(1) iS have False by (simp add: dependent_def) } - then show ?thesis unfolding dependent_def by blast -qed - -(* ------------------------------------------------------------------------- *) -(* Picking an orthogonal replacement for a spanning set. *) -(* ------------------------------------------------------------------------- *) - (* FIXME : Move to some general theory ?*) -definition "pairwise R S \ (\x \ S. \y\ S. x\y \ R x y)" - -lemma vector_sub_project_orthogonal: "(b::'a::ordered_field^'n::finite) \ (x - ((b \ x) / (b\b)) *s b) = 0" - apply (cases "b = 0", simp) - apply (simp add: dot_rsub dot_rmult) - unfolding times_divide_eq_right[symmetric] - by (simp add: field_simps dot_eq_0) - -lemma basis_orthogonal: - fixes B :: "(real ^'n::finite) set" - assumes fB: "finite B" - shows "\C. finite C \ card C \ card B \ span C = span B \ pairwise orthogonal C" - (is " \C. ?P B C") -proof(induct rule: finite_induct[OF fB]) - case 1 thus ?case apply (rule exI[where x="{}"]) by (auto simp add: pairwise_def) -next - case (2 a B) - note fB = `finite B` and aB = `a \ B` - from `\C. finite C \ card C \ card B \ span C = span B \ pairwise orthogonal C` - obtain C where C: "finite C" "card C \ card B" - "span C = span B" "pairwise orthogonal C" by blast - let ?a = "a - setsum (\x. (x\a / (x\x)) *s x) C" - let ?C = "insert ?a C" - from C(1) have fC: "finite ?C" by simp - from fB aB C(1,2) have cC: "card ?C \ card (insert a B)" by (simp add: card_insert_if) - {fix x k - have th0: "\(a::'b::comm_ring) b c. a - (b - c) = c + (a - b)" by (simp add: ring_simps) - have "x - k *s (a - (\x\C. (x \ a / (x \ x)) *s x)) \ span C \ x - k *s a \ span C" - apply (simp only: vector_ssub_ldistrib th0) - apply (rule span_add_eq) - apply (rule span_mul) - apply (rule span_setsum[OF C(1)]) - apply clarify - apply (rule span_mul) - by (rule span_superset)} - then have SC: "span ?C = span (insert a B)" - unfolding expand_set_eq span_breakdown_eq C(3)[symmetric] by auto - thm pairwise_def - {fix x y assume xC: "x \ ?C" and yC: "y \ ?C" and xy: "x \ y" - {assume xa: "x = ?a" and ya: "y = ?a" - have "orthogonal x y" using xa ya xy by blast} - moreover - {assume xa: "x = ?a" and ya: "y \ ?a" "y \ C" - from ya have Cy: "C = insert y (C - {y})" by blast - have fth: "finite (C - {y})" using C by simp - have "orthogonal x y" - using xa ya - unfolding orthogonal_def xa dot_lsub dot_rsub diff_eq_0_iff_eq - apply simp - apply (subst Cy) - using C(1) fth - apply (simp only: setsum_clauses) - thm dot_ladd - apply (auto simp add: dot_ladd dot_radd dot_lmult dot_rmult dot_eq_0 dot_sym[of y a] dot_lsum[OF fth]) - apply (rule setsum_0') - apply clarsimp - apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format]) - by auto} - moreover - {assume xa: "x \ ?a" "x \ C" and ya: "y = ?a" - from xa have Cx: "C = insert x (C - {x})" by blast - have fth: "finite (C - {x})" using C by simp - have "orthogonal x y" - using xa ya - unfolding orthogonal_def ya dot_rsub dot_lsub diff_eq_0_iff_eq - apply simp - apply (subst Cx) - using C(1) fth - apply (simp only: setsum_clauses) - apply (subst dot_sym[of x]) - apply (auto simp add: dot_radd dot_rmult dot_eq_0 dot_sym[of x a] dot_rsum[OF fth]) - apply (rule setsum_0') - apply clarsimp - apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format]) - by auto} - moreover - {assume xa: "x \ C" and ya: "y \ C" - have "orthogonal x y" using xa ya xy C(4) unfolding pairwise_def by blast} - ultimately have "orthogonal x y" using xC yC by blast} - then have CPO: "pairwise orthogonal ?C" unfolding pairwise_def by blast - from fC cC SC CPO have "?P (insert a B) ?C" by blast - then show ?case by blast -qed - -lemma orthogonal_basis_exists: - fixes V :: "(real ^'n::finite) set" - shows "\B. independent B \ B \ span V \ V \ span B \ (B hassize dim V) \ pairwise orthogonal B" -proof- - from basis_exists[of V] obtain B where B: "B \ V" "independent B" "V \ span B" "B hassize dim V" by blast - from B have fB: "finite B" "card B = dim V" by (simp_all add: hassize_def) - from basis_orthogonal[OF fB(1)] obtain C where - C: "finite C" "card C \ card B" "span C = span B" "pairwise orthogonal C" by blast - from C B - have CSV: "C \ span V" by (metis span_inc span_mono subset_trans) - from span_mono[OF B(3)] C have SVC: "span V \ span C" by (simp add: span_span) - from card_le_dim_spanning[OF CSV SVC C(1)] C(2,3) fB - have iC: "independent C" by (simp add: dim_span) - from C fB have "card C \ dim V" by simp - moreover have "dim V \ card C" using span_card_ge_dim[OF CSV SVC C(1)] - by (simp add: dim_span) - ultimately have CdV: "C hassize dim V" unfolding hassize_def using C(1) by simp - from C B CSV CdV iC show ?thesis by auto -qed - -lemma span_eq: "span S = span T \ S \ span T \ T \ span S" - by (metis set_eq_subset span_mono span_span span_inc) (* FIXME: slow *) - -(* ------------------------------------------------------------------------- *) -(* Low-dimensional subset is in a hyperplane (weak orthogonal complement). *) -(* ------------------------------------------------------------------------- *) - -lemma span_not_univ_orthogonal: - assumes sU: "span S \ UNIV" - shows "\(a:: real ^'n::finite). a \0 \ (\x \ span S. a \ x = 0)" -proof- - from sU obtain a where a: "a \ span S" by blast - from orthogonal_basis_exists obtain B where - B: "independent B" "B \ span S" "S \ span B" "B hassize dim S" "pairwise orthogonal B" - by blast - from B have fB: "finite B" "card B = dim S" by (simp_all add: hassize_def) - from span_mono[OF B(2)] span_mono[OF B(3)] - have sSB: "span S = span B" by (simp add: span_span) - let ?a = "a - setsum (\b. (a\b / (b\b)) *s b) B" - have "setsum (\b. (a\b / (b\b)) *s b) B \ span S" - unfolding sSB - apply (rule span_setsum[OF fB(1)]) - apply clarsimp - apply (rule span_mul) - by (rule span_superset) - with a have a0:"?a \ 0" by auto - have "\x\span B. ?a \ x = 0" - proof(rule span_induct') - show "subspace (\x. ?a \ x = 0)" - by (auto simp add: subspace_def mem_def dot_radd dot_rmult) - next - {fix x assume x: "x \ B" - from x have B': "B = insert x (B - {x})" by blast - have fth: "finite (B - {x})" using fB by simp - have "?a \ x = 0" - apply (subst B') using fB fth - unfolding setsum_clauses(2)[OF fth] - apply simp - apply (clarsimp simp add: dot_lsub dot_ladd dot_lmult dot_lsum dot_eq_0) - apply (rule setsum_0', rule ballI) - unfolding dot_sym - by (auto simp add: x field_simps dot_eq_0 intro: B(5)[unfolded pairwise_def orthogonal_def, rule_format])} - then show "\x \ B. ?a \ x = 0" by blast - qed - with a0 show ?thesis unfolding sSB by (auto intro: exI[where x="?a"]) -qed - -lemma span_not_univ_subset_hyperplane: - assumes SU: "span S \ (UNIV ::(real^'n::finite) set)" - shows "\ a. a \0 \ span S \ {x. a \ x = 0}" - using span_not_univ_orthogonal[OF SU] by auto - -lemma lowdim_subset_hyperplane: - assumes d: "dim S < CARD('n::finite)" - shows "\(a::real ^'n::finite). a \ 0 \ span S \ {x. a \ x = 0}" -proof- - {assume "span S = UNIV" - hence "dim (span S) = dim (UNIV :: (real ^'n) set)" by simp - hence "dim S = CARD('n)" by (simp add: dim_span dim_univ) - with d have False by arith} - hence th: "span S \ UNIV" by blast - from span_not_univ_subset_hyperplane[OF th] show ?thesis . -qed - -(* We can extend a linear basis-basis injection to the whole set. *) - -lemma linear_indep_image_lemma: - assumes lf: "linear f" and fB: "finite B" - and ifB: "independent (f ` B)" - and fi: "inj_on f B" and xsB: "x \ span B" - and fx: "f (x::'a::field^'n) = 0" - shows "x = 0" - using fB ifB fi xsB fx -proof(induct arbitrary: x rule: finite_induct[OF fB]) - case 1 thus ?case by (auto simp add: span_empty) -next - case (2 a b x) - have fb: "finite b" using "2.prems" by simp - have th0: "f ` b \ f ` (insert a b)" - apply (rule image_mono) by blast - from independent_mono[ OF "2.prems"(2) th0] - have ifb: "independent (f ` b)" . - have fib: "inj_on f b" - apply (rule subset_inj_on [OF "2.prems"(3)]) - by blast - from span_breakdown[of a "insert a b", simplified, OF "2.prems"(4)] - obtain k where k: "x - k*s a \ span (b -{a})" by blast - have "f (x - k*s a) \ span (f ` b)" - unfolding span_linear_image[OF lf] - apply (rule imageI) - using k span_mono[of "b-{a}" b] by blast - hence "f x - k*s f a \ span (f ` b)" - by (simp add: linear_sub[OF lf] linear_cmul[OF lf]) - hence th: "-k *s f a \ span (f ` b)" - using "2.prems"(5) by (simp add: vector_smult_lneg) - {assume k0: "k = 0" - from k0 k have "x \ span (b -{a})" by simp - then have "x \ span b" using span_mono[of "b-{a}" b] - by blast} - moreover - {assume k0: "k \ 0" - from span_mul[OF th, of "- 1/ k"] k0 - have th1: "f a \ span (f ` b)" - by (auto simp add: vector_smult_assoc) - from inj_on_image_set_diff[OF "2.prems"(3), of "insert a b " "{a}", symmetric] - have tha: "f ` insert a b - f ` {a} = f ` (insert a b - {a})" by blast - from "2.prems"(2)[unfolded dependent_def bex_simps(10), rule_format, of "f a"] - have "f a \ span (f ` b)" using tha - using "2.hyps"(2) - "2.prems"(3) by auto - with th1 have False by blast - then have "x \ span b" by blast} - ultimately have xsb: "x \ span b" by blast - from "2.hyps"(3)[OF fb ifb fib xsb "2.prems"(5)] - show "x = 0" . -qed - -(* We can extend a linear mapping from basis. *) - -lemma linear_independent_extend_lemma: - assumes fi: "finite B" and ib: "independent B" - shows "\g. (\x\ span B. \y\ span B. g ((x::'a::field^'n) + y) = g x + g y) - \ (\x\ span B. \c. g (c*s x) = c *s g x) - \ (\x\ B. g x = f x)" -using ib fi -proof(induct rule: finite_induct[OF fi]) - case 1 thus ?case by (auto simp add: span_empty) -next - case (2 a b) - from "2.prems" "2.hyps" have ibf: "independent b" "finite b" - by (simp_all add: independent_insert) - from "2.hyps"(3)[OF ibf] obtain g where - g: "\x\span b. \y\span b. g (x + y) = g x + g y" - "\x\span b. \c. g (c *s x) = c *s g x" "\x\b. g x = f x" by blast - let ?h = "\z. SOME k. (z - k *s a) \ span b" - {fix z assume z: "z \ span (insert a b)" - have th0: "z - ?h z *s a \ span b" - apply (rule someI_ex) - unfolding span_breakdown_eq[symmetric] - using z . - {fix k assume k: "z - k *s a \ span b" - have eq: "z - ?h z *s a - (z - k*s a) = (k - ?h z) *s a" - by (simp add: ring_simps vector_sadd_rdistrib[symmetric]) - from span_sub[OF th0 k] - have khz: "(k - ?h z) *s a \ span b" by (simp add: eq) - {assume "k \ ?h z" hence k0: "k - ?h z \ 0" by simp - from k0 span_mul[OF khz, of "1 /(k - ?h z)"] - have "a \ span b" by (simp add: vector_smult_assoc) - with "2.prems"(1) "2.hyps"(2) have False - by (auto simp add: dependent_def)} - then have "k = ?h z" by blast} - with th0 have "z - ?h z *s a \ span b \ (\k. z - k *s a \ span b \ k = ?h z)" by blast} - note h = this - let ?g = "\z. ?h z *s f a + g (z - ?h z *s a)" - {fix x y assume x: "x \ span (insert a b)" and y: "y \ span (insert a b)" - have tha: "\(x::'a^'n) y a k l. (x + y) - (k + l) *s a = (x - k *s a) + (y - l *s a)" - by (vector ring_simps) - have addh: "?h (x + y) = ?h x + ?h y" - apply (rule conjunct2[OF h, rule_format, symmetric]) - apply (rule span_add[OF x y]) - unfolding tha - by (metis span_add x y conjunct1[OF h, rule_format]) - have "?g (x + y) = ?g x + ?g y" - unfolding addh tha - g(1)[rule_format,OF conjunct1[OF h, OF x] conjunct1[OF h, OF y]] - by (simp add: vector_sadd_rdistrib)} - moreover - {fix x:: "'a^'n" and c:: 'a assume x: "x \ span (insert a b)" - have tha: "\(x::'a^'n) c k a. c *s x - (c * k) *s a = c *s (x - k *s a)" - by (vector ring_simps) - have hc: "?h (c *s x) = c * ?h x" - apply (rule conjunct2[OF h, rule_format, symmetric]) - apply (metis span_mul x) - by (metis tha span_mul x conjunct1[OF h]) - have "?g (c *s x) = c*s ?g x" - unfolding hc tha g(2)[rule_format, OF conjunct1[OF h, OF x]] - by (vector ring_simps)} - moreover - {fix x assume x: "x \ (insert a b)" - {assume xa: "x = a" - have ha1: "1 = ?h a" - apply (rule conjunct2[OF h, rule_format]) - apply (metis span_superset insertI1) - using conjunct1[OF h, OF span_superset, OF insertI1] - by (auto simp add: span_0) - - from xa ha1[symmetric] have "?g x = f x" - apply simp - using g(2)[rule_format, OF span_0, of 0] - by simp} - moreover - {assume xb: "x \ b" - have h0: "0 = ?h x" - apply (rule conjunct2[OF h, rule_format]) - apply (metis span_superset insertI1 xb x) - apply simp - apply (metis span_superset xb) - done - have "?g x = f x" - by (simp add: h0[symmetric] g(3)[rule_format, OF xb])} - ultimately have "?g x = f x" using x by blast } - ultimately show ?case apply - apply (rule exI[where x="?g"]) by blast -qed - -lemma linear_independent_extend: - assumes iB: "independent (B:: (real ^'n::finite) set)" - shows "\g. linear g \ (\x\B. g x = f x)" -proof- - from maximal_independent_subset_extend[of B UNIV] iB - obtain C where C: "B \ C" "independent C" "\x. x \ span C" by auto - - from C(2) independent_bound[of C] linear_independent_extend_lemma[of C f] - obtain g where g: "(\x\ span C. \y\ span C. g (x + y) = g x + g y) - \ (\x\ span C. \c. g (c*s x) = c *s g x) - \ (\x\ C. g x = f x)" by blast - from g show ?thesis unfolding linear_def using C - apply clarsimp by blast -qed - -(* Can construct an isomorphism between spaces of same dimension. *) - -lemma card_le_inj: assumes fA: "finite A" and fB: "finite B" - and c: "card A \ card B" shows "(\f. f ` A \ B \ inj_on f A)" -using fB c -proof(induct arbitrary: B rule: finite_induct[OF fA]) - case 1 thus ?case by simp -next - case (2 x s t) - thus ?case - proof(induct rule: finite_induct[OF "2.prems"(1)]) - case 1 then show ?case by simp - next - case (2 y t) - from "2.prems"(1,2,5) "2.hyps"(1,2) have cst:"card s \ card t" by simp - from "2.prems"(3) [OF "2.hyps"(1) cst] obtain f where - f: "f ` s \ t \ inj_on f s" by blast - from f "2.prems"(2) "2.hyps"(2) show ?case - apply - - apply (rule exI[where x = "\z. if z = x then y else f z"]) - by (auto simp add: inj_on_def) - qed -qed - -lemma card_subset_eq: assumes fB: "finite B" and AB: "A \ B" and - c: "card A = card B" - shows "A = B" -proof- - from fB AB have fA: "finite A" by (auto intro: finite_subset) - from fA fB have fBA: "finite (B - A)" by auto - have e: "A \ (B - A) = {}" by blast - have eq: "A \ (B - A) = B" using AB by blast - from card_Un_disjoint[OF fA fBA e, unfolded eq c] - have "card (B - A) = 0" by arith - hence "B - A = {}" unfolding card_eq_0_iff using fA fB by simp - with AB show "A = B" by blast -qed - -lemma subspace_isomorphism: - assumes s: "subspace (S:: (real ^'n::finite) set)" - and t: "subspace (T :: (real ^ 'm::finite) set)" - and d: "dim S = dim T" - shows "\f. linear f \ f ` S = T \ inj_on f S" -proof- - from basis_exists[of S] obtain B where - B: "B \ S" "independent B" "S \ span B" "B hassize dim S" by blast - from basis_exists[of T] obtain C where - C: "C \ T" "independent C" "T \ span C" "C hassize dim T" by blast - from B(4) C(4) card_le_inj[of B C] d obtain f where - f: "f ` B \ C" "inj_on f B" unfolding hassize_def by auto - from linear_independent_extend[OF B(2)] obtain g where - g: "linear g" "\x\ B. g x = f x" by blast - from B(4) have fB: "finite B" by (simp add: hassize_def) - from C(4) have fC: "finite C" by (simp add: hassize_def) - from inj_on_iff_eq_card[OF fB, of f] f(2) - have "card (f ` B) = card B" by simp - with B(4) C(4) have ceq: "card (f ` B) = card C" using d - by (simp add: hassize_def) - have "g ` B = f ` B" using g(2) - by (auto simp add: image_iff) - also have "\ = C" using card_subset_eq[OF fC f(1) ceq] . - finally have gBC: "g ` B = C" . - have gi: "inj_on g B" using f(2) g(2) - by (auto simp add: inj_on_def) - note g0 = linear_indep_image_lemma[OF g(1) fB, unfolded gBC, OF C(2) gi] - {fix x y assume x: "x \ S" and y: "y \ S" and gxy:"g x = g y" - from B(3) x y have x': "x \ span B" and y': "y \ span B" by blast+ - from gxy have th0: "g (x - y) = 0" by (simp add: linear_sub[OF g(1)]) - have th1: "x - y \ span B" using x' y' by (metis span_sub) - have "x=y" using g0[OF th1 th0] by simp } - then have giS: "inj_on g S" - unfolding inj_on_def by blast - from span_subspace[OF B(1,3) s] - have "g ` S = span (g ` B)" by (simp add: span_linear_image[OF g(1)]) - also have "\ = span C" unfolding gBC .. - also have "\ = T" using span_subspace[OF C(1,3) t] . - finally have gS: "g ` S = T" . - from g(1) gS giS show ?thesis by blast -qed - -(* linear functions are equal on a subspace if they are on a spanning set. *) - -lemma subspace_kernel: - assumes lf: "linear (f::'a::semiring_1 ^'n \ _)" - shows "subspace {x. f x = 0}" -apply (simp add: subspace_def) -by (simp add: linear_add[OF lf] linear_cmul[OF lf] linear_0[OF lf]) - -lemma linear_eq_0_span: - assumes lf: "linear f" and f0: "\x\B. f x = 0" - shows "\x \ span B. f x = (0::'a::semiring_1 ^'n)" -proof - fix x assume x: "x \ span B" - let ?P = "\x. f x = 0" - from subspace_kernel[OF lf] have "subspace ?P" unfolding Collect_def . - with x f0 span_induct[of B "?P" x] show "f x = 0" by blast -qed - -lemma linear_eq_0: - assumes lf: "linear f" and SB: "S \ span B" and f0: "\x\B. f x = 0" - shows "\x \ S. f x = (0::'a::semiring_1^'n)" - by (metis linear_eq_0_span[OF lf] subset_eq SB f0) - -lemma linear_eq: - assumes lf: "linear (f::'a::ring_1^'n \ _)" and lg: "linear g" and S: "S \ span B" - and fg: "\ x\ B. f x = g x" - shows "\x\ S. f x = g x" -proof- - let ?h = "\x. f x - g x" - from fg have fg': "\x\ B. ?h x = 0" by simp - from linear_eq_0[OF linear_compose_sub[OF lf lg] S fg'] - show ?thesis by simp -qed - -lemma linear_eq_stdbasis: - assumes lf: "linear (f::'a::ring_1^'m::finite \ 'a^'n::finite)" and lg: "linear g" - and fg: "\i. f (basis i) = g(basis i)" - shows "f = g" -proof- - let ?U = "UNIV :: 'm set" - let ?I = "{basis i:: 'a^'m|i. i \ ?U}" - {fix x assume x: "x \ (UNIV :: ('a^'m) set)" - from equalityD2[OF span_stdbasis] - have IU: " (UNIV :: ('a^'m) set) \ span ?I" by blast - from linear_eq[OF lf lg IU] fg x - have "f x = g x" unfolding Collect_def Ball_def mem_def by metis} - then show ?thesis by (auto intro: ext) -qed - -(* Similar results for bilinear functions. *) - -lemma bilinear_eq: - assumes bf: "bilinear (f:: 'a::ring^'m \ 'a^'n \ 'a^'p)" - and bg: "bilinear g" - and SB: "S \ span B" and TC: "T \ span C" - and fg: "\x\ B. \y\ C. f x y = g x y" - shows "\x\S. \y\T. f x y = g x y " -proof- - let ?P = "\x. \y\ span C. f x y = g x y" - from bf bg have sp: "subspace ?P" - unfolding bilinear_def linear_def subspace_def bf bg - by(auto simp add: span_0 mem_def bilinear_lzero[OF bf] bilinear_lzero[OF bg] span_add Ball_def intro: bilinear_ladd[OF bf]) - - have "\x \ span B. \y\ span C. f x y = g x y" - apply - - apply (rule ballI) - apply (rule span_induct[of B ?P]) - defer - apply (rule sp) - apply assumption - apply (clarsimp simp add: Ball_def) - apply (rule_tac P="\y. f xa y = g xa y" and S=C in span_induct) - using fg - apply (auto simp add: subspace_def) - using bf bg unfolding bilinear_def linear_def - by(auto simp add: span_0 mem_def bilinear_rzero[OF bf] bilinear_rzero[OF bg] span_add Ball_def intro: bilinear_ladd[OF bf]) - then show ?thesis using SB TC by (auto intro: ext) -qed - -lemma bilinear_eq_stdbasis: - assumes bf: "bilinear (f:: 'a::ring_1^'m::finite \ 'a^'n::finite \ 'a^'p)" - and bg: "bilinear g" - and fg: "\i j. f (basis i) (basis j) = g (basis i) (basis j)" - shows "f = g" -proof- - from fg have th: "\x \ {basis i| i. i\ (UNIV :: 'm set)}. \y\ {basis j |j. j \ (UNIV :: 'n set)}. f x y = g x y" by blast - from bilinear_eq[OF bf bg equalityD2[OF span_stdbasis] equalityD2[OF span_stdbasis] th] show ?thesis by (blast intro: ext) -qed - -(* Detailed theorems about left and right invertibility in general case. *) - -lemma left_invertible_transp: - "(\(B::'a^'n^'m). B ** transp (A::'a^'n^'m) = mat (1::'a::comm_semiring_1)) \ (\(B::'a^'m^'n). A ** B = mat 1)" - by (metis matrix_transp_mul transp_mat transp_transp) - -lemma right_invertible_transp: - "(\(B::'a^'n^'m). transp (A::'a^'n^'m) ** B = mat (1::'a::comm_semiring_1)) \ (\(B::'a^'m^'n). B ** A = mat 1)" - by (metis matrix_transp_mul transp_mat transp_transp) - -lemma linear_injective_left_inverse: - assumes lf: "linear (f::real ^'n::finite \ real ^'m::finite)" and fi: "inj f" - shows "\g. linear g \ g o f = id" -proof- - from linear_independent_extend[OF independent_injective_image, OF independent_stdbasis, OF lf fi] - obtain h:: "real ^'m \ real ^'n" where h: "linear h" " \x \ f ` {basis i|i. i \ (UNIV::'n set)}. h x = inv f x" by blast - from h(2) - have th: "\i. (h \ f) (basis i) = id (basis i)" - using inv_o_cancel[OF fi, unfolded stupid_ext[symmetric] id_def o_def] - by auto - - from linear_eq_stdbasis[OF linear_compose[OF lf h(1)] linear_id th] - have "h o f = id" . - then show ?thesis using h(1) by blast -qed - -lemma linear_surjective_right_inverse: - assumes lf: "linear (f:: real ^'m::finite \ real ^'n::finite)" and sf: "surj f" - shows "\g. linear g \ f o g = id" -proof- - from linear_independent_extend[OF independent_stdbasis] - obtain h:: "real ^'n \ real ^'m" where - h: "linear h" "\ x\ {basis i| i. i\ (UNIV :: 'n set)}. h x = inv f x" by blast - from h(2) - have th: "\i. (f o h) (basis i) = id (basis i)" - using sf - apply (auto simp add: surj_iff o_def stupid_ext[symmetric]) - apply (erule_tac x="basis i" in allE) - by auto - - from linear_eq_stdbasis[OF linear_compose[OF h(1) lf] linear_id th] - have "f o h = id" . - then show ?thesis using h(1) by blast -qed - -lemma matrix_left_invertible_injective: -"(\B. (B::real^'m^'n) ** (A::real^'n::finite^'m::finite) = mat 1) \ (\x y. A *v x = A *v y \ x = y)" -proof- - {fix B:: "real^'m^'n" and x y assume B: "B ** A = mat 1" and xy: "A *v x = A*v y" - from xy have "B*v (A *v x) = B *v (A*v y)" by simp - hence "x = y" - unfolding matrix_vector_mul_assoc B matrix_vector_mul_lid .} - moreover - {assume A: "\x y. A *v x = A *v y \ x = y" - hence i: "inj (op *v A)" unfolding inj_on_def by auto - from linear_injective_left_inverse[OF matrix_vector_mul_linear i] - obtain g where g: "linear g" "g o op *v A = id" by blast - have "matrix g ** A = mat 1" - unfolding matrix_eq matrix_vector_mul_lid matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)] - using g(2) by (simp add: o_def id_def stupid_ext) - then have "\B. (B::real ^'m^'n) ** A = mat 1" by blast} - ultimately show ?thesis by blast -qed - -lemma matrix_left_invertible_ker: - "(\B. (B::real ^'m::finite^'n::finite) ** (A::real^'n^'m) = mat 1) \ (\x. A *v x = 0 \ x = 0)" - unfolding matrix_left_invertible_injective - using linear_injective_0[OF matrix_vector_mul_linear, of A] - by (simp add: inj_on_def) - -lemma matrix_right_invertible_surjective: -"(\B. (A::real^'n::finite^'m::finite) ** (B::real^'m^'n) = mat 1) \ surj (\x. A *v x)" -proof- - {fix B :: "real ^'m^'n" assume AB: "A ** B = mat 1" - {fix x :: "real ^ 'm" - have "A *v (B *v x) = x" - by (simp add: matrix_vector_mul_lid matrix_vector_mul_assoc AB)} - hence "surj (op *v A)" unfolding surj_def by metis } - moreover - {assume sf: "surj (op *v A)" - from linear_surjective_right_inverse[OF matrix_vector_mul_linear sf] - obtain g:: "real ^'m \ real ^'n" where g: "linear g" "op *v A o g = id" - by blast - - have "A ** (matrix g) = mat 1" - unfolding matrix_eq matrix_vector_mul_lid - matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)] - using g(2) unfolding o_def stupid_ext[symmetric] id_def - . - hence "\B. A ** (B::real^'m^'n) = mat 1" by blast - } - ultimately show ?thesis unfolding surj_def by blast -qed - -lemma matrix_left_invertible_independent_columns: - fixes A :: "real^'n::finite^'m::finite" - shows "(\(B::real ^'m^'n). B ** A = mat 1) \ (\c. setsum (\i. c i *s column i A) (UNIV :: 'n set) = 0 \ (\i. c i = 0))" - (is "?lhs \ ?rhs") -proof- - let ?U = "UNIV :: 'n set" - {assume k: "\x. A *v x = 0 \ x = 0" - {fix c i assume c: "setsum (\i. c i *s column i A) ?U = 0" - and i: "i \ ?U" - let ?x = "\ i. c i" - have th0:"A *v ?x = 0" - using c - unfolding matrix_mult_vsum Cart_eq - by auto - from k[rule_format, OF th0] i - have "c i = 0" by (vector Cart_eq)} - hence ?rhs by blast} - moreover - {assume H: ?rhs - {fix x assume x: "A *v x = 0" - let ?c = "\i. ((x$i ):: real)" - from H[rule_format, of ?c, unfolded matrix_mult_vsum[symmetric], OF x] - have "x = 0" by vector}} - ultimately show ?thesis unfolding matrix_left_invertible_ker by blast -qed - -lemma matrix_right_invertible_independent_rows: - fixes A :: "real^'n::finite^'m::finite" - shows "(\(B::real^'m^'n). A ** B = mat 1) \ (\c. setsum (\i. c i *s row i A) (UNIV :: 'm set) = 0 \ (\i. c i = 0))" - unfolding left_invertible_transp[symmetric] - matrix_left_invertible_independent_columns - by (simp add: column_transp) - -lemma matrix_right_invertible_span_columns: - "(\(B::real ^'n::finite^'m::finite). (A::real ^'m^'n) ** B = mat 1) \ span (columns A) = UNIV" (is "?lhs = ?rhs") -proof- - let ?U = "UNIV :: 'm set" - have fU: "finite ?U" by simp - have lhseq: "?lhs \ (\y. \(x::real^'m). setsum (\i. (x$i) *s column i A) ?U = y)" - unfolding matrix_right_invertible_surjective matrix_mult_vsum surj_def - apply (subst eq_commute) .. - have rhseq: "?rhs \ (\x. x \ span (columns A))" by blast - {assume h: ?lhs - {fix x:: "real ^'n" - from h[unfolded lhseq, rule_format, of x] obtain y:: "real ^'m" - where y: "setsum (\i. (y$i) *s column i A) ?U = x" by blast - have "x \ span (columns A)" - unfolding y[symmetric] - apply (rule span_setsum[OF fU]) - apply clarify - apply (rule span_mul) - apply (rule span_superset) - unfolding columns_def - by blast} - then have ?rhs unfolding rhseq by blast} - moreover - {assume h:?rhs - let ?P = "\(y::real ^'n). \(x::real^'m). setsum (\i. (x$i) *s column i A) ?U = y" - {fix y have "?P y" - proof(rule span_induct_alt[of ?P "columns A"]) - show "\x\real ^ 'm. setsum (\i. (x$i) *s column i A) ?U = 0" - apply (rule exI[where x=0]) - by (simp add: zero_index vector_smult_lzero) - next - fix c y1 y2 assume y1: "y1 \ columns A" and y2: "?P y2" - from y1 obtain i where i: "i \ ?U" "y1 = column i A" - unfolding columns_def by blast - from y2 obtain x:: "real ^'m" where - x: "setsum (\i. (x$i) *s column i A) ?U = y2" by blast - let ?x = "(\ j. if j = i then c + (x$i) else (x$j))::real^'m" - show "?P (c*s y1 + y2)" - proof(rule exI[where x= "?x"], vector, auto simp add: i x[symmetric] cond_value_iff right_distrib cond_application_beta cong del: if_weak_cong) - fix j - have th: "\xa \ ?U. (if xa = i then (c + (x$i)) * ((column xa A)$j) - else (x$xa) * ((column xa A$j))) = (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))" using i(1) - by (simp add: ring_simps) - have "setsum (\xa. if xa = i then (c + (x$i)) * ((column xa A)$j) - else (x$xa) * ((column xa A$j))) ?U = setsum (\xa. (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))) ?U" - apply (rule setsum_cong[OF refl]) - using th by blast - also have "\ = setsum (\xa. if xa = i then c * ((column i A)$j) else 0) ?U + setsum (\xa. ((x$xa) * ((column xa A)$j))) ?U" - by (simp add: setsum_addf) - also have "\ = c * ((column i A)$j) + setsum (\xa. ((x$xa) * ((column xa A)$j))) ?U" - unfolding setsum_delta[OF fU] - using i(1) by simp - finally show "setsum (\xa. if xa = i then (c + (x$i)) * ((column xa A)$j) - else (x$xa) * ((column xa A$j))) ?U = c * ((column i A)$j) + setsum (\xa. ((x$xa) * ((column xa A)$j))) ?U" . - qed - next - show "y \ span (columns A)" unfolding h by blast - qed} - then have ?lhs unfolding lhseq ..} - ultimately show ?thesis by blast -qed - -lemma matrix_left_invertible_span_rows: - "(\(B::real^'m::finite^'n::finite). B ** (A::real^'n^'m) = mat 1) \ span (rows A) = UNIV" - unfolding right_invertible_transp[symmetric] - unfolding columns_transp[symmetric] - unfolding matrix_right_invertible_span_columns - .. - -(* An injective map real^'n->real^'n is also surjective. *) - -lemma linear_injective_imp_surjective: - assumes lf: "linear (f:: real ^'n::finite \ real ^'n)" and fi: "inj f" - shows "surj f" -proof- - let ?U = "UNIV :: (real ^'n) set" - from basis_exists[of ?U] obtain B - where B: "B \ ?U" "independent B" "?U \ span B" "B hassize dim ?U" - by blast - from B(4) have d: "dim ?U = card B" by (simp add: hassize_def) - have th: "?U \ span (f ` B)" - apply (rule card_ge_dim_independent) - apply blast - apply (rule independent_injective_image[OF B(2) lf fi]) - apply (rule order_eq_refl) - apply (rule sym) - unfolding d - apply (rule card_image) - apply (rule subset_inj_on[OF fi]) - by blast - from th show ?thesis - unfolding span_linear_image[OF lf] surj_def - using B(3) by blast -qed - -(* And vice versa. *) - -lemma surjective_iff_injective_gen: - assumes fS: "finite S" and fT: "finite T" and c: "card S = card T" - and ST: "f ` S \ T" - shows "(\y \ T. \x \ S. f x = y) \ inj_on f S" (is "?lhs \ ?rhs") -proof- - {assume h: "?lhs" - {fix x y assume x: "x \ S" and y: "y \ S" and f: "f x = f y" - from x fS have S0: "card S \ 0" by auto - {assume xy: "x \ y" - have th: "card S \ card (f ` (S - {y}))" - unfolding c - apply (rule card_mono) - apply (rule finite_imageI) - using fS apply simp - using h xy x y f unfolding subset_eq image_iff - apply auto - apply (case_tac "xa = f x") - apply (rule bexI[where x=x]) - apply auto - done - also have " \ \ card (S -{y})" - apply (rule card_image_le) - using fS by simp - also have "\ \ card S - 1" using y fS by simp - finally have False using S0 by arith } - then have "x = y" by blast} - then have ?rhs unfolding inj_on_def by blast} - moreover - {assume h: ?rhs - have "f ` S = T" - apply (rule card_subset_eq[OF fT ST]) - unfolding card_image[OF h] using c . - then have ?lhs by blast} - ultimately show ?thesis by blast -qed - -lemma linear_surjective_imp_injective: - assumes lf: "linear (f::real ^'n::finite => real ^'n)" and sf: "surj f" - shows "inj f" -proof- - let ?U = "UNIV :: (real ^'n) set" - from basis_exists[of ?U] obtain B - where B: "B \ ?U" "independent B" "?U \ span B" "B hassize dim ?U" - by blast - {fix x assume x: "x \ span B" and fx: "f x = 0" - from B(4) have fB: "finite B" by (simp add: hassize_def) - from B(4) have d: "dim ?U = card B" by (simp add: hassize_def) - have fBi: "independent (f ` B)" - apply (rule card_le_dim_spanning[of "f ` B" ?U]) - apply blast - using sf B(3) - unfolding span_linear_image[OF lf] surj_def subset_eq image_iff - apply blast - using fB apply (blast intro: finite_imageI) - unfolding d - apply (rule card_image_le) - apply (rule fB) - done - have th0: "dim ?U \ card (f ` B)" - apply (rule span_card_ge_dim) - apply blast - unfolding span_linear_image[OF lf] - apply (rule subset_trans[where B = "f ` UNIV"]) - using sf unfolding surj_def apply blast - apply (rule image_mono) - apply (rule B(3)) - apply (metis finite_imageI fB) - done - - moreover have "card (f ` B) \ card B" - by (rule card_image_le, rule fB) - ultimately have th1: "card B = card (f ` B)" unfolding d by arith - have fiB: "inj_on f B" - unfolding surjective_iff_injective_gen[OF fB finite_imageI[OF fB] th1 subset_refl, symmetric] by blast - from linear_indep_image_lemma[OF lf fB fBi fiB x] fx - have "x = 0" by blast} - note th = this - from th show ?thesis unfolding linear_injective_0[OF lf] - using B(3) by blast -qed - -(* Hence either is enough for isomorphism. *) - -lemma left_right_inverse_eq: - assumes fg: "f o g = id" and gh: "g o h = id" - shows "f = h" -proof- - have "f = f o (g o h)" unfolding gh by simp - also have "\ = (f o g) o h" by (simp add: o_assoc) - finally show "f = h" unfolding fg by simp -qed - -lemma isomorphism_expand: - "f o g = id \ g o f = id \ (\x. f(g x) = x) \ (\x. g(f x) = x)" - by (simp add: expand_fun_eq o_def id_def) - -lemma linear_injective_isomorphism: - assumes lf: "linear (f :: real^'n::finite \ real ^'n)" and fi: "inj f" - shows "\f'. linear f' \ (\x. f' (f x) = x) \ (\x. f (f' x) = x)" -unfolding isomorphism_expand[symmetric] -using linear_surjective_right_inverse[OF lf linear_injective_imp_surjective[OF lf fi]] linear_injective_left_inverse[OF lf fi] -by (metis left_right_inverse_eq) - -lemma linear_surjective_isomorphism: - assumes lf: "linear (f::real ^'n::finite \ real ^'n)" and sf: "surj f" - shows "\f'. linear f' \ (\x. f' (f x) = x) \ (\x. f (f' x) = x)" -unfolding isomorphism_expand[symmetric] -using linear_surjective_right_inverse[OF lf sf] linear_injective_left_inverse[OF lf linear_surjective_imp_injective[OF lf sf]] -by (metis left_right_inverse_eq) - -(* Left and right inverses are the same for R^N->R^N. *) - -lemma linear_inverse_left: - assumes lf: "linear (f::real ^'n::finite \ real ^'n)" and lf': "linear f'" - shows "f o f' = id \ f' o f = id" -proof- - {fix f f':: "real ^'n \ real ^'n" - assume lf: "linear f" "linear f'" and f: "f o f' = id" - from f have sf: "surj f" - - apply (auto simp add: o_def stupid_ext[symmetric] id_def surj_def) - by metis - from linear_surjective_isomorphism[OF lf(1) sf] lf f - have "f' o f = id" unfolding stupid_ext[symmetric] o_def id_def - by metis} - then show ?thesis using lf lf' by metis -qed - -(* Moreover, a one-sided inverse is automatically linear. *) - -lemma left_inverse_linear: - assumes lf: "linear (f::real ^'n::finite \ real ^'n)" and gf: "g o f = id" - shows "linear g" -proof- - from gf have fi: "inj f" apply (auto simp add: inj_on_def o_def id_def stupid_ext[symmetric]) - by metis - from linear_injective_isomorphism[OF lf fi] - obtain h:: "real ^'n \ real ^'n" where - h: "linear h" "\x. h (f x) = x" "\x. f (h x) = x" by blast - have "h = g" apply (rule ext) using gf h(2,3) - apply (simp add: o_def id_def stupid_ext[symmetric]) - by metis - with h(1) show ?thesis by blast -qed - -lemma right_inverse_linear: - assumes lf: "linear (f:: real ^'n::finite \ real ^'n)" and gf: "f o g = id" - shows "linear g" -proof- - from gf have fi: "surj f" apply (auto simp add: surj_def o_def id_def stupid_ext[symmetric]) - by metis - from linear_surjective_isomorphism[OF lf fi] - obtain h:: "real ^'n \ real ^'n" where - h: "linear h" "\x. h (f x) = x" "\x. f (h x) = x" by blast - have "h = g" apply (rule ext) using gf h(2,3) - apply (simp add: o_def id_def stupid_ext[symmetric]) - by metis - with h(1) show ?thesis by blast -qed - -(* The same result in terms of square matrices. *) - -lemma matrix_left_right_inverse: - fixes A A' :: "real ^'n::finite^'n" - shows "A ** A' = mat 1 \ A' ** A = mat 1" -proof- - {fix A A' :: "real ^'n^'n" assume AA': "A ** A' = mat 1" - have sA: "surj (op *v A)" - unfolding surj_def - apply clarify - apply (rule_tac x="(A' *v y)" in exI) - by (simp add: matrix_vector_mul_assoc AA' matrix_vector_mul_lid) - from linear_surjective_isomorphism[OF matrix_vector_mul_linear sA] - obtain f' :: "real ^'n \ real ^'n" - where f': "linear f'" "\x. f' (A *v x) = x" "\x. A *v f' x = x" by blast - have th: "matrix f' ** A = mat 1" - by (simp add: matrix_eq matrix_works[OF f'(1)] matrix_vector_mul_assoc[symmetric] matrix_vector_mul_lid f'(2)[rule_format]) - hence "(matrix f' ** A) ** A' = mat 1 ** A'" by simp - hence "matrix f' = A'" by (simp add: matrix_mul_assoc[symmetric] AA' matrix_mul_rid matrix_mul_lid) - hence "matrix f' ** A = A' ** A" by simp - hence "A' ** A = mat 1" by (simp add: th)} - then show ?thesis by blast -qed - -(* Considering an n-element vector as an n-by-1 or 1-by-n matrix. *) - -definition "rowvector v = (\ i j. (v$j))" - -definition "columnvector v = (\ i j. (v$i))" - -lemma transp_columnvector: - "transp(columnvector v) = rowvector v" - by (simp add: transp_def rowvector_def columnvector_def Cart_eq) - -lemma transp_rowvector: "transp(rowvector v) = columnvector v" - by (simp add: transp_def columnvector_def rowvector_def Cart_eq) - -lemma dot_rowvector_columnvector: - "columnvector (A *v v) = A ** columnvector v" - by (vector columnvector_def matrix_matrix_mult_def matrix_vector_mult_def) - -lemma dot_matrix_product: "(x::'a::semiring_1^'n::finite) \ y = (((rowvector x ::'a^'n^1) ** (columnvector y :: 'a^1^'n))$1)$1" - by (vector matrix_matrix_mult_def rowvector_def columnvector_def dot_def) - -lemma dot_matrix_vector_mul: - fixes A B :: "real ^'n::finite ^'n" and x y :: "real ^'n" - shows "(A *v x) \ (B *v y) = - (((rowvector x :: real^'n^1) ** ((transp A ** B) ** (columnvector y :: real ^1^'n)))$1)$1" -unfolding dot_matrix_product transp_columnvector[symmetric] - dot_rowvector_columnvector matrix_transp_mul matrix_mul_assoc .. - -(* Infinity norm. *) - -definition "infnorm (x::real^'n::finite) = rsup {abs(x$i) |i. i\ (UNIV :: 'n set)}" - -lemma numseg_dimindex_nonempty: "\i. i \ (UNIV :: 'n set)" - by auto - -lemma infnorm_set_image: - "{abs(x$i) |i. i\ (UNIV :: 'n set)} = - (\i. abs(x$i)) ` (UNIV :: 'n set)" by blast - -lemma infnorm_set_lemma: - shows "finite {abs((x::'a::abs ^'n::finite)$i) |i. i\ (UNIV :: 'n set)}" - and "{abs(x$i) |i. i\ (UNIV :: 'n::finite set)} \ {}" - unfolding infnorm_set_image - by (auto intro: finite_imageI) - -lemma infnorm_pos_le: "0 \ infnorm (x::real^'n::finite)" - unfolding infnorm_def - unfolding rsup_finite_ge_iff[ OF infnorm_set_lemma] - unfolding infnorm_set_image - by auto - -lemma infnorm_triangle: "infnorm ((x::real^'n::finite) + y) \ infnorm x + infnorm y" -proof- - have th: "\x y (z::real). x - y <= z \ x - z <= y" by arith - have th1: "\S f. f ` S = { f i| i. i \ S}" by blast - have th2: "\x (y::real). abs(x + y) - abs(x) <= abs(y)" by arith - show ?thesis - unfolding infnorm_def - unfolding rsup_finite_le_iff[ OF infnorm_set_lemma] - apply (subst diff_le_eq[symmetric]) - unfolding rsup_finite_ge_iff[ OF infnorm_set_lemma] - unfolding infnorm_set_image bex_simps - apply (subst th) - unfolding th1 - unfolding rsup_finite_ge_iff[ OF infnorm_set_lemma] - - unfolding infnorm_set_image ball_simps bex_simps - apply simp - apply (metis th2) - done -qed - -lemma infnorm_eq_0: "infnorm x = 0 \ (x::real ^'n::finite) = 0" -proof- - have "infnorm x <= 0 \ x = 0" - unfolding infnorm_def - unfolding rsup_finite_le_iff[OF infnorm_set_lemma] - unfolding infnorm_set_image ball_simps - by vector - then show ?thesis using infnorm_pos_le[of x] by simp -qed - -lemma infnorm_0: "infnorm 0 = 0" - by (simp add: infnorm_eq_0) - -lemma infnorm_neg: "infnorm (- x) = infnorm x" - unfolding infnorm_def - apply (rule cong[of "rsup" "rsup"]) - apply blast - apply (rule set_ext) - apply auto - done - -lemma infnorm_sub: "infnorm (x - y) = infnorm (y - x)" -proof- - have "y - x = - (x - y)" by simp - then show ?thesis by (metis infnorm_neg) -qed - -lemma real_abs_sub_infnorm: "\ infnorm x - infnorm y\ \ infnorm (x - y)" -proof- - have th: "\(nx::real) n ny. nx <= n + ny \ ny <= n + nx ==> \nx - ny\ <= n" - by arith - from infnorm_triangle[of "x - y" " y"] infnorm_triangle[of "x - y" "-x"] - have ths: "infnorm x \ infnorm (x - y) + infnorm y" - "infnorm y \ infnorm (x - y) + infnorm x" - by (simp_all add: ring_simps infnorm_neg diff_def[symmetric]) - from th[OF ths] show ?thesis . -qed - -lemma real_abs_infnorm: " \infnorm x\ = infnorm x" - using infnorm_pos_le[of x] by arith - -lemma component_le_infnorm: - shows "\x$i\ \ infnorm (x::real^'n::finite)" -proof- - let ?U = "UNIV :: 'n set" - let ?S = "{\x$i\ |i. i\ ?U}" - have fS: "finite ?S" unfolding image_Collect[symmetric] - apply (rule finite_imageI) unfolding Collect_def mem_def by simp - have S0: "?S \ {}" by blast - have th1: "\S f. f ` S = { f i| i. i \ S}" by blast - from rsup_finite_in[OF fS S0] rsup_finite_Ub[OF fS S0] - show ?thesis unfolding infnorm_def isUb_def setle_def - unfolding infnorm_set_image ball_simps by auto -qed - -lemma infnorm_mul_lemma: "infnorm(a *s x) <= \a\ * infnorm x" - apply (subst infnorm_def) - unfolding rsup_finite_le_iff[OF infnorm_set_lemma] - unfolding infnorm_set_image ball_simps - apply (simp add: abs_mult) - apply (rule allI) - apply (cut_tac component_le_infnorm[of x]) - apply (rule mult_mono) - apply auto - done - -lemma infnorm_mul: "infnorm(a *s x) = abs a * infnorm x" -proof- - {assume a0: "a = 0" hence ?thesis by (simp add: infnorm_0) } - moreover - {assume a0: "a \ 0" - from a0 have th: "(1/a) *s (a *s x) = x" - by (simp add: vector_smult_assoc) - from a0 have ap: "\a\ > 0" by arith - from infnorm_mul_lemma[of "1/a" "a *s x"] - have "infnorm x \ 1/\a\ * infnorm (a*s x)" - unfolding th by simp - with ap have "\a\ * infnorm x \ \a\ * (1/\a\ * infnorm (a *s x))" by (simp add: field_simps) - then have "\a\ * infnorm x \ infnorm (a*s x)" - using ap by (simp add: field_simps) - with infnorm_mul_lemma[of a x] have ?thesis by arith } - ultimately show ?thesis by blast -qed - -lemma infnorm_pos_lt: "infnorm x > 0 \ x \ 0" - using infnorm_pos_le[of x] infnorm_eq_0[of x] by arith - -(* Prove that it differs only up to a bound from Euclidean norm. *) - -lemma infnorm_le_norm: "infnorm x \ norm x" - unfolding infnorm_def rsup_finite_le_iff[OF infnorm_set_lemma] - unfolding infnorm_set_image ball_simps - by (metis component_le_norm) -lemma card_enum: "card {1 .. n} = n" by auto -lemma norm_le_infnorm: "norm(x) <= sqrt(real CARD('n)) * infnorm(x::real ^'n::finite)" -proof- - let ?d = "CARD('n)" - have "real ?d \ 0" by simp - hence d2: "(sqrt (real ?d))^2 = real ?d" - by (auto intro: real_sqrt_pow2) - have th: "sqrt (real ?d) * infnorm x \ 0" - by (simp add: zero_le_mult_iff real_sqrt_ge_0_iff infnorm_pos_le) - have th1: "x\x \ (sqrt (real ?d) * infnorm x)^2" - unfolding power_mult_distrib d2 - apply (subst power2_abs[symmetric]) - unfolding real_of_nat_def dot_def power2_eq_square[symmetric] - apply (subst power2_abs[symmetric]) - apply (rule setsum_bounded) - apply (rule power_mono) - unfolding abs_of_nonneg[OF infnorm_pos_le] - unfolding infnorm_def rsup_finite_ge_iff[OF infnorm_set_lemma] - unfolding infnorm_set_image bex_simps - apply blast - by (rule abs_ge_zero) - from real_le_lsqrt[OF dot_pos_le th th1] - show ?thesis unfolding real_vector_norm_def id_def . -qed - -(* Equality in Cauchy-Schwarz and triangle inequalities. *) - -lemma norm_cauchy_schwarz_eq: "(x::real ^'n::finite) \ y = norm x * norm y \ norm x *s y = norm y *s x" (is "?lhs \ ?rhs") -proof- - {assume h: "x = 0" - hence ?thesis by simp} - moreover - {assume h: "y = 0" - hence ?thesis by simp} - moreover - {assume x: "x \ 0" and y: "y \ 0" - from dot_eq_0[of "norm y *s x - norm x *s y"] - have "?rhs \ (norm y * (norm y * norm x * norm x - norm x * (x \ y)) - norm x * (norm y * (y \ x) - norm x * norm y * norm y) = 0)" - using x y - unfolding dot_rsub dot_lsub dot_lmult dot_rmult - unfolding norm_pow_2[symmetric] power2_eq_square diff_eq_0_iff_eq apply (simp add: dot_sym) - apply (simp add: ring_simps) - apply metis - done - also have "\ \ (2 * norm x * norm y * (norm x * norm y - x \ y) = 0)" using x y - by (simp add: ring_simps dot_sym) - also have "\ \ ?lhs" using x y - apply simp - by metis - finally have ?thesis by blast} - ultimately show ?thesis by blast -qed - -lemma norm_cauchy_schwarz_abs_eq: - fixes x y :: "real ^ 'n::finite" - shows "abs(x \ y) = norm x * norm y \ - norm x *s y = norm y *s x \ norm(x) *s y = - norm y *s x" (is "?lhs \ ?rhs") -proof- - have th: "\(x::real) a. a \ 0 \ abs x = a \ x = a \ x = - a" by arith - have "?rhs \ norm x *s y = norm y *s x \ norm (- x) *s y = norm y *s (- x)" - apply simp by vector - also have "\ \(x \ y = norm x * norm y \ - (-x) \ y = norm x * norm y)" - unfolding norm_cauchy_schwarz_eq[symmetric] - unfolding norm_minus_cancel - norm_mul by blast - also have "\ \ ?lhs" - unfolding th[OF mult_nonneg_nonneg, OF norm_ge_zero[of x] norm_ge_zero[of y]] dot_lneg - by arith - finally show ?thesis .. -qed - -lemma norm_triangle_eq: - fixes x y :: "real ^ 'n::finite" - shows "norm(x + y) = norm x + norm y \ norm x *s y = norm y *s x" -proof- - {assume x: "x =0 \ y =0" - hence ?thesis by (cases "x=0", simp_all)} - moreover - {assume x: "x \ 0" and y: "y \ 0" - hence "norm x \ 0" "norm y \ 0" - by simp_all - hence n: "norm x > 0" "norm y > 0" - using norm_ge_zero[of x] norm_ge_zero[of y] - by arith+ - have th: "\(a::real) b c. a + b + c \ 0 ==> (a = b + c \ a^2 = (b + c)^2)" by algebra - have "norm(x + y) = norm x + norm y \ norm(x + y)^ 2 = (norm x + norm y) ^2" - apply (rule th) using n norm_ge_zero[of "x + y"] - by arith - also have "\ \ norm x *s y = norm y *s x" - unfolding norm_cauchy_schwarz_eq[symmetric] - unfolding norm_pow_2 dot_ladd dot_radd - by (simp add: norm_pow_2[symmetric] power2_eq_square dot_sym ring_simps) - finally have ?thesis .} - ultimately show ?thesis by blast -qed - -(* Collinearity.*) - -definition "collinear S \ (\u. \x \ S. \ y \ S. \c. x - y = c *s u)" - -lemma collinear_empty: "collinear {}" by (simp add: collinear_def) - -lemma collinear_sing: "collinear {(x::'a::ring_1^'n)}" - apply (simp add: collinear_def) - apply (rule exI[where x=0]) - by simp - -lemma collinear_2: "collinear {(x::'a::ring_1^'n),y}" - apply (simp add: collinear_def) - apply (rule exI[where x="x - y"]) - apply auto - apply (rule exI[where x=0], simp) - apply (rule exI[where x=1], simp) - apply (rule exI[where x="- 1"], simp add: vector_sneg_minus1[symmetric]) - apply (rule exI[where x=0], simp) - done - -lemma collinear_lemma: "collinear {(0::real^'n),x,y} \ x = 0 \ y = 0 \ (\c. y = c *s x)" (is "?lhs \ ?rhs") -proof- - {assume "x=0 \ y = 0" hence ?thesis - by (cases "x = 0", simp_all add: collinear_2 insert_commute)} - moreover - {assume x: "x \ 0" and y: "y \ 0" - {assume h: "?lhs" - then obtain u where u: "\ x\ {0,x,y}. \y\ {0,x,y}. \c. x - y = c *s u" unfolding collinear_def by blast - from u[rule_format, of x 0] u[rule_format, of y 0] - obtain cx and cy where - cx: "x = cx*s u" and cy: "y = cy*s u" - by auto - from cx x have cx0: "cx \ 0" by auto - from cy y have cy0: "cy \ 0" by auto - let ?d = "cy / cx" - from cx cy cx0 have "y = ?d *s x" - by (simp add: vector_smult_assoc) - hence ?rhs using x y by blast} - moreover - {assume h: "?rhs" - then obtain c where c: "y = c*s x" using x y by blast - have ?lhs unfolding collinear_def c - apply (rule exI[where x=x]) - apply auto - apply (rule exI[where x="- 1"], simp only: vector_smult_lneg vector_smult_lid) - apply (rule exI[where x= "-c"], simp only: vector_smult_lneg) - apply (rule exI[where x=1], simp) - apply (rule exI[where x="1 - c"], simp add: vector_smult_lneg vector_sub_rdistrib) - apply (rule exI[where x="c - 1"], simp add: vector_smult_lneg vector_sub_rdistrib) - done} - ultimately have ?thesis by blast} - ultimately show ?thesis by blast -qed - -lemma norm_cauchy_schwarz_equal: - fixes x y :: "real ^ 'n::finite" - shows "abs(x \ y) = norm x * norm y \ collinear {(0::real^'n),x,y}" -unfolding norm_cauchy_schwarz_abs_eq -apply (cases "x=0", simp_all add: collinear_2) -apply (cases "y=0", simp_all add: collinear_2 insert_commute) -unfolding collinear_lemma -apply simp -apply (subgoal_tac "norm x \ 0") -apply (subgoal_tac "norm y \ 0") -apply (rule iffI) -apply (cases "norm x *s y = norm y *s x") -apply (rule exI[where x="(1/norm x) * norm y"]) -apply (drule sym) -unfolding vector_smult_assoc[symmetric] -apply (simp add: vector_smult_assoc field_simps) -apply (rule exI[where x="(1/norm x) * - norm y"]) -apply clarify -apply (drule sym) -unfolding vector_smult_assoc[symmetric] -apply (simp add: vector_smult_assoc field_simps) -apply (erule exE) -apply (erule ssubst) -unfolding vector_smult_assoc -unfolding norm_mul -apply (subgoal_tac "norm x * c = \c\ * norm x \ norm x * c = - \c\ * norm x") -apply (case_tac "c <= 0", simp add: ring_simps) -apply (simp add: ring_simps) -apply (case_tac "c <= 0", simp add: ring_simps) -apply (simp add: ring_simps) -apply simp -apply simp -done - -end diff -r 1fad3160d873 -r 2083bde13ce1 src/HOL/Library/Finite_Cartesian_Product.thy --- a/src/HOL/Library/Finite_Cartesian_Product.thy Fri Oct 23 14:33:07 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,95 +0,0 @@ -(* Title: HOL/Library/Finite_Cartesian_Product - Author: Amine Chaieb, University of Cambridge -*) - -header {* Definition of finite Cartesian product types. *} - -theory Finite_Cartesian_Product -imports Main (*FIXME: ATP_Linkup is only needed for metis at a few places. We could dispense of that by changing the proofs.*) -begin - -definition hassize (infixr "hassize" 12) where - "(S hassize n) = (finite S \ card S = n)" - -lemma hassize_image_inj: assumes f: "inj_on f S" and S: "S hassize n" - shows "f ` S hassize n" - using f S card_image[OF f] - by (simp add: hassize_def inj_on_def) - - -subsection {* Finite Cartesian products, with indexing and lambdas. *} - -typedef (open Cart) - ('a, 'b) "^" (infixl "^" 15) - = "UNIV :: ('b \ 'a) set" - morphisms Cart_nth Cart_lambda .. - -notation Cart_nth (infixl "$" 90) - -notation (xsymbols) Cart_lambda (binder "\" 10) - -lemma stupid_ext: "(\x. f x = g x) \ (f = g)" - apply auto - apply (rule ext) - apply auto - done - -lemma Cart_eq: "((x:: 'a ^ 'b) = y) \ (\i. x$i = y$i)" - by (simp add: Cart_nth_inject [symmetric] expand_fun_eq) - -lemma Cart_lambda_beta [simp]: "Cart_lambda g $ i = g i" - by (simp add: Cart_lambda_inverse) - -lemma Cart_lambda_unique: - fixes f :: "'a ^ 'b" - shows "(\i. f$i = g i) \ Cart_lambda g = f" - by (auto simp add: Cart_eq) - -lemma Cart_lambda_eta: "(\ i. (g$i)) = g" - by (simp add: Cart_eq) - -text{* A non-standard sum to "paste" Cartesian products. *} - -definition pastecart :: "'a ^ 'm \ 'a ^ 'n \ 'a ^ ('m + 'n)" where - "pastecart f g = (\ i. case i of Inl a \ f$a | Inr b \ g$b)" - -definition fstcart:: "'a ^('m + 'n) \ 'a ^ 'm" where - "fstcart f = (\ i. (f$(Inl i)))" - -definition sndcart:: "'a ^('m + 'n) \ 'a ^ 'n" where - "sndcart f = (\ i. (f$(Inr i)))" - -lemma nth_pastecart_Inl [simp]: "pastecart f g $ Inl a = f$a" - unfolding pastecart_def by simp - -lemma nth_pastecart_Inr [simp]: "pastecart f g $ Inr b = g$b" - unfolding pastecart_def by simp - -lemma nth_fstcart [simp]: "fstcart f $ i = f $ Inl i" - unfolding fstcart_def by simp - -lemma nth_sndtcart [simp]: "sndcart f $ i = f $ Inr i" - unfolding sndcart_def by simp - -lemma finite_sum_image: "(UNIV::('a + 'b) set) = range Inl \ range Inr" -by (auto, case_tac x, auto) - -lemma fstcart_pastecart: "fstcart (pastecart (x::'a ^'m ) (y:: 'a ^ 'n)) = x" - by (simp add: Cart_eq) - -lemma sndcart_pastecart: "sndcart (pastecart (x::'a ^'m ) (y:: 'a ^ 'n)) = y" - by (simp add: Cart_eq) - -lemma pastecart_fst_snd: "pastecart (fstcart z) (sndcart z) = z" - by (simp add: Cart_eq pastecart_def fstcart_def sndcart_def split: sum.split) - -lemma pastecart_eq: "(x = y) \ (fstcart x = fstcart y) \ (sndcart x = sndcart y)" - using pastecart_fst_snd[of x] pastecart_fst_snd[of y] by metis - -lemma forall_pastecart: "(\p. P p) \ (\x y. P (pastecart x y))" - by (metis pastecart_fst_snd fstcart_pastecart sndcart_pastecart) - -lemma exists_pastecart: "(\p. P p) \ (\x y. P (pastecart x y))" - by (metis pastecart_fst_snd fstcart_pastecart sndcart_pastecart) - -end diff -r 1fad3160d873 -r 2083bde13ce1 src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy Fri Oct 23 14:33:07 2009 +0200 +++ b/src/HOL/Library/Library.thy Fri Oct 23 13:23:18 2009 +0200 @@ -14,9 +14,7 @@ Commutative_Ring Continuity ContNotDenum - Convex_Euclidean_Space Countable - Determinants Diagonalize Efficient_Nat Enum @@ -54,7 +52,6 @@ RBT State_Monad Sum_Of_Squares - Topology_Euclidean_Space Univ_Poly While_Combinator Word diff -r 1fad3160d873 -r 2083bde13ce1 src/HOL/Library/Topology_Euclidean_Space.thy --- a/src/HOL/Library/Topology_Euclidean_Space.thy Fri Oct 23 14:33:07 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,6027 +0,0 @@ -(* Title: HOL/Library/Topology_Euclidian_Space.thy - Author: Amine Chaieb, University of Cambridge - Author: Robert Himmelmann, TU Muenchen -*) - -header {* Elementary topology in Euclidean space. *} - -theory Topology_Euclidean_Space -imports SEQ Euclidean_Space Product_Vector -begin - -declare fstcart_pastecart[simp] sndcart_pastecart[simp] - -subsection{* General notion of a topology *} - -definition "istopology L \ {} \ L \ (\S \L. \T \L. S \ T \ L) \ (\K. K \L \ \ K \ L)" -typedef (open) 'a topology = "{L::('a set) set. istopology L}" - morphisms "openin" "topology" - unfolding istopology_def by blast - -lemma istopology_open_in[intro]: "istopology(openin U)" - using openin[of U] by blast - -lemma topology_inverse': "istopology U \ openin (topology U) = U" - using topology_inverse[unfolded mem_def Collect_def] . - -lemma topology_inverse_iff: "istopology U \ openin (topology U) = U" - using topology_inverse[of U] istopology_open_in[of "topology U"] by auto - -lemma topology_eq: "T1 = T2 \ (\S. openin T1 S \ openin T2 S)" -proof- - {assume "T1=T2" hence "\S. openin T1 S \ openin T2 S" by simp} - moreover - {assume H: "\S. openin T1 S \ openin T2 S" - hence "openin T1 = openin T2" by (metis mem_def set_ext) - hence "topology (openin T1) = topology (openin T2)" by simp - hence "T1 = T2" unfolding openin_inverse .} - ultimately show ?thesis by blast -qed - -text{* Infer the "universe" from union of all sets in the topology. *} - -definition "topspace T = \{S. openin T S}" - -subsection{* Main properties of open sets *} - -lemma openin_clauses: - fixes U :: "'a topology" - shows "openin U {}" - "\S T. openin U S \ openin U T \ openin U (S\T)" - "\K. (\S \ K. openin U S) \ openin U (\K)" - using openin[of U] unfolding istopology_def Collect_def mem_def - by (metis mem_def subset_eq)+ - -lemma openin_subset[intro]: "openin U S \ S \ topspace U" - unfolding topspace_def by blast -lemma openin_empty[simp]: "openin U {}" by (simp add: openin_clauses) - -lemma openin_Int[intro]: "openin U S \ openin U T \ openin U (S \ T)" - by (simp add: openin_clauses) - -lemma openin_Union[intro]: "(\S \K. openin U S) \ openin U (\ K)" by (simp add: openin_clauses) - -lemma openin_Un[intro]: "openin U S \ openin U T \ openin U (S \ T)" - using openin_Union[of "{S,T}" U] by auto - -lemma openin_topspace[intro, simp]: "openin U (topspace U)" by (simp add: openin_Union topspace_def) - -lemma openin_subopen: "openin U S \ (\x \ S. \T. openin U T \ x \ T \ T \ S)" (is "?lhs \ ?rhs") -proof- - {assume ?lhs then have ?rhs by auto } - moreover - {assume H: ?rhs - then obtain t where t: "\x\S. openin U (t x) \ x \ t x \ t x \ S" - unfolding Ball_def ex_simps(6)[symmetric] choice_iff by blast - from t have th0: "\x\ t`S. openin U x" by auto - have "\ t`S = S" using t by auto - with openin_Union[OF th0] have "openin U S" by simp } - ultimately show ?thesis by blast -qed - -subsection{* Closed sets *} - -definition "closedin U S \ S \ topspace U \ openin U (topspace U - S)" - -lemma closedin_subset: "closedin U S \ S \ topspace U" by (metis closedin_def) -lemma closedin_empty[simp]: "closedin U {}" by (simp add: closedin_def) -lemma closedin_topspace[intro,simp]: - "closedin U (topspace U)" by (simp add: closedin_def) -lemma closedin_Un[intro]: "closedin U S \ closedin U T \ closedin U (S \ T)" - by (auto simp add: Diff_Un closedin_def) - -lemma Diff_Inter[intro]: "A - \S = \ {A - s|s. s\S}" by auto -lemma closedin_Inter[intro]: assumes Ke: "K \ {}" and Kc: "\S \K. closedin U S" - shows "closedin U (\ K)" using Ke Kc unfolding closedin_def Diff_Inter by auto - -lemma closedin_Int[intro]: "closedin U S \ closedin U T \ closedin U (S \ T)" - using closedin_Inter[of "{S,T}" U] by auto - -lemma Diff_Diff_Int: "A - (A - B) = A \ B" by blast -lemma openin_closedin_eq: "openin U S \ S \ topspace U \ closedin U (topspace U - S)" - apply (auto simp add: closedin_def Diff_Diff_Int inf_absorb2) - apply (metis openin_subset subset_eq) - done - -lemma openin_closedin: "S \ topspace U \ (openin U S \ closedin U (topspace U - S))" - by (simp add: openin_closedin_eq) - -lemma openin_diff[intro]: assumes oS: "openin U S" and cT: "closedin U T" shows "openin U (S - T)" -proof- - have "S - T = S \ (topspace U - T)" using openin_subset[of U S] oS cT - by (auto simp add: topspace_def openin_subset) - then show ?thesis using oS cT by (auto simp add: closedin_def) -qed - -lemma closedin_diff[intro]: assumes oS: "closedin U S" and cT: "openin U T" shows "closedin U (S - T)" -proof- - have "S - T = S \ (topspace U - T)" using closedin_subset[of U S] oS cT - by (auto simp add: topspace_def ) - then show ?thesis using oS cT by (auto simp add: openin_closedin_eq) -qed - -subsection{* Subspace topology. *} - -definition "subtopology U V = topology {S \ V |S. openin U S}" - -lemma istopology_subtopology: "istopology {S \ V |S. openin U S}" (is "istopology ?L") -proof- - have "{} \ ?L" by blast - {fix A B assume A: "A \ ?L" and B: "B \ ?L" - from A B obtain Sa and Sb where Sa: "openin U Sa" "A = Sa \ V" and Sb: "openin U Sb" "B = Sb \ V" by blast - have "A\B = (Sa \ Sb) \ V" "openin U (Sa \ Sb)" using Sa Sb by blast+ - then have "A \ B \ ?L" by blast} - moreover - {fix K assume K: "K \ ?L" - have th0: "?L = (\S. S \ V) ` openin U " - apply (rule set_ext) - apply (simp add: Ball_def image_iff) - by (metis mem_def) - from K[unfolded th0 subset_image_iff] - obtain Sk where Sk: "Sk \ openin U" "K = (\S. S \ V) ` Sk" by blast - have "\K = (\Sk) \ V" using Sk by auto - moreover have "openin U (\ Sk)" using Sk by (auto simp add: subset_eq mem_def) - ultimately have "\K \ ?L" by blast} - ultimately show ?thesis unfolding istopology_def by blast -qed - -lemma openin_subtopology: - "openin (subtopology U V) S \ (\ T. (openin U T) \ (S = T \ V))" - unfolding subtopology_def topology_inverse'[OF istopology_subtopology] - by (auto simp add: Collect_def) - -lemma topspace_subtopology: "topspace(subtopology U V) = topspace U \ V" - by (auto simp add: topspace_def openin_subtopology) - -lemma closedin_subtopology: - "closedin (subtopology U V) S \ (\T. closedin U T \ S = T \ V)" - unfolding closedin_def topspace_subtopology - apply (simp add: openin_subtopology) - apply (rule iffI) - apply clarify - apply (rule_tac x="topspace U - T" in exI) - by auto - -lemma openin_subtopology_refl: "openin (subtopology U V) V \ V \ topspace U" - unfolding openin_subtopology - apply (rule iffI, clarify) - apply (frule openin_subset[of U]) apply blast - apply (rule exI[where x="topspace U"]) - by auto - -lemma subtopology_superset: assumes UV: "topspace U \ V" - shows "subtopology U V = U" -proof- - {fix S - {fix T assume T: "openin U T" "S = T \ V" - from T openin_subset[OF T(1)] UV have eq: "S = T" by blast - have "openin U S" unfolding eq using T by blast} - moreover - {assume S: "openin U S" - hence "\T. openin U T \ S = T \ V" - using openin_subset[OF S] UV by auto} - ultimately have "(\T. openin U T \ S = T \ V) \ openin U S" by blast} - then show ?thesis unfolding topology_eq openin_subtopology by blast -qed - - -lemma subtopology_topspace[simp]: "subtopology U (topspace U) = U" - by (simp add: subtopology_superset) - -lemma subtopology_UNIV[simp]: "subtopology U UNIV = U" - by (simp add: subtopology_superset) - -subsection{* The universal Euclidean versions are what we use most of the time *} - -definition - euclidean :: "'a::topological_space topology" where - "euclidean = topology open" - -lemma open_openin: "open S \ openin euclidean S" - unfolding euclidean_def - apply (rule cong[where x=S and y=S]) - apply (rule topology_inverse[symmetric]) - apply (auto simp add: istopology_def) - by (auto simp add: mem_def subset_eq) - -lemma topspace_euclidean: "topspace euclidean = UNIV" - apply (simp add: topspace_def) - apply (rule set_ext) - by (auto simp add: open_openin[symmetric]) - -lemma topspace_euclidean_subtopology[simp]: "topspace (subtopology euclidean S) = S" - by (simp add: topspace_euclidean topspace_subtopology) - -lemma closed_closedin: "closed S \ closedin euclidean S" - by (simp add: closed_def closedin_def topspace_euclidean open_openin Compl_eq_Diff_UNIV) - -lemma open_subopen: "open S \ (\x\S. \T. open T \ x \ T \ T \ S)" - by (simp add: open_openin openin_subopen[symmetric]) - -subsection{* Open and closed balls. *} - -definition - ball :: "'a::metric_space \ real \ 'a set" where - "ball x e = {y. dist x y < e}" - -definition - cball :: "'a::metric_space \ real \ 'a set" where - "cball x e = {y. dist x y \ e}" - -lemma mem_ball[simp]: "y \ ball x e \ dist x y < e" by (simp add: ball_def) -lemma mem_cball[simp]: "y \ cball x e \ dist x y \ e" by (simp add: cball_def) - -lemma mem_ball_0 [simp]: - fixes x :: "'a::real_normed_vector" - shows "x \ ball 0 e \ norm x < e" - by (simp add: dist_norm) - -lemma mem_cball_0 [simp]: - fixes x :: "'a::real_normed_vector" - shows "x \ cball 0 e \ norm x \ e" - by (simp add: dist_norm) - -lemma centre_in_cball[simp]: "x \ cball x e \ 0\ e" by simp -lemma ball_subset_cball[simp,intro]: "ball x e \ cball x e" by (simp add: subset_eq) -lemma subset_ball[intro]: "d <= e ==> ball x d \ ball x e" by (simp add: subset_eq) -lemma subset_cball[intro]: "d <= e ==> cball x d \ cball x e" by (simp add: subset_eq) -lemma ball_max_Un: "ball a (max r s) = ball a r \ ball a s" - by (simp add: expand_set_eq) arith - -lemma ball_min_Int: "ball a (min r s) = ball a r \ ball a s" - by (simp add: expand_set_eq) - -subsection{* Topological properties of open balls *} - -lemma diff_less_iff: "(a::real) - b > 0 \ a > b" - "(a::real) - b < 0 \ a < b" - "a - b < c \ a < c +b" "a - b > c \ a > c +b" by arith+ -lemma diff_le_iff: "(a::real) - b \ 0 \ a \ b" "(a::real) - b \ 0 \ a \ b" - "a - b \ c \ a \ c +b" "a - b \ c \ a \ c +b" by arith+ - -lemma open_ball[intro, simp]: "open (ball x e)" - unfolding open_dist ball_def Collect_def Ball_def mem_def - unfolding dist_commute - apply clarify - apply (rule_tac x="e - dist xa x" in exI) - using dist_triangle_alt[where z=x] - apply (clarsimp simp add: diff_less_iff) - apply atomize - apply (erule_tac x="y" in allE) - apply (erule_tac x="xa" in allE) - by arith - -lemma centre_in_ball[simp]: "x \ ball x e \ e > 0" by (metis mem_ball dist_self) -lemma open_contains_ball: "open S \ (\x\S. \e>0. ball x e \ S)" - unfolding open_dist subset_eq mem_ball Ball_def dist_commute .. - -lemma open_contains_ball_eq: "open S \ \x. x\S \ (\e>0. ball x e \ S)" - by (metis open_contains_ball subset_eq centre_in_ball) - -lemma ball_eq_empty[simp]: "ball x e = {} \ e \ 0" - unfolding mem_ball expand_set_eq - apply (simp add: not_less) - by (metis zero_le_dist order_trans dist_self) - -lemma ball_empty[intro]: "e \ 0 ==> ball x e = {}" by simp - -subsection{* Basic "localization" results are handy for connectedness. *} - -lemma openin_open: "openin (subtopology euclidean U) S \ (\T. open T \ (S = U \ T))" - by (auto simp add: openin_subtopology open_openin[symmetric]) - -lemma openin_open_Int[intro]: "open S \ openin (subtopology euclidean U) (U \ S)" - by (auto simp add: openin_open) - -lemma open_openin_trans[trans]: - "open S \ open T \ T \ S \ openin (subtopology euclidean S) T" - by (metis Int_absorb1 openin_open_Int) - -lemma open_subset: "S \ T \ open S \ openin (subtopology euclidean T) S" - by (auto simp add: openin_open) - -lemma closedin_closed: "closedin (subtopology euclidean U) S \ (\T. closed T \ S = U \ T)" - by (simp add: closedin_subtopology closed_closedin Int_ac) - -lemma closedin_closed_Int: "closed S ==> closedin (subtopology euclidean U) (U \ S)" - by (metis closedin_closed) - -lemma closed_closedin_trans: "closed S \ closed T \ T \ S \ closedin (subtopology euclidean S) T" - apply (subgoal_tac "S \ T = T" ) - apply auto - apply (frule closedin_closed_Int[of T S]) - by simp - -lemma closed_subset: "S \ T \ closed S \ closedin (subtopology euclidean T) S" - by (auto simp add: closedin_closed) - -lemma openin_euclidean_subtopology_iff: - fixes S U :: "'a::metric_space set" - shows "openin (subtopology euclidean U) S - \ S \ U \ (\x\S. \e>0. \x'\U. dist x' x < e \ x'\ S)" (is "?lhs \ ?rhs") -proof- - {assume ?lhs hence ?rhs unfolding openin_subtopology open_openin[symmetric] - by (simp add: open_dist) blast} - moreover - {assume SU: "S \ U" and H: "\x. x \ S \ \e>0. \x'\U. dist x' x < e \ x' \ S" - from H obtain d where d: "\x . x\ S \ d x > 0 \ (\x' \ U. dist x' x < d x \ x' \ S)" - by metis - let ?T = "\{B. \x\S. B = ball x (d x)}" - have oT: "open ?T" by auto - { fix x assume "x\S" - hence "x \ \{B. \x\S. B = ball x (d x)}" - apply simp apply(rule_tac x="ball x(d x)" in exI) apply auto - by (rule d [THEN conjunct1]) - hence "x\ ?T \ U" using SU and `x\S` by auto } - moreover - { fix y assume "y\?T" - then obtain B where "y\B" "B\{B. \x\S. B = ball x (d x)}" by auto - then obtain x where "x\S" and x:"y \ ball x (d x)" by auto - assume "y\U" - hence "y\S" using d[OF `x\S`] and x by(auto simp add: dist_commute) } - ultimately have "S = ?T \ U" by blast - with oT have ?lhs unfolding openin_subtopology open_openin[symmetric] by blast} - ultimately show ?thesis by blast -qed - -text{* These "transitivity" results are handy too. *} - -lemma openin_trans[trans]: "openin (subtopology euclidean T) S \ openin (subtopology euclidean U) T - \ openin (subtopology euclidean U) S" - unfolding open_openin openin_open by blast - -lemma openin_open_trans: "openin (subtopology euclidean T) S \ open T \ open S" - by (auto simp add: openin_open intro: openin_trans) - -lemma closedin_trans[trans]: - "closedin (subtopology euclidean T) S \ - closedin (subtopology euclidean U) T - ==> closedin (subtopology euclidean U) S" - by (auto simp add: closedin_closed closed_closedin closed_Inter Int_assoc) - -lemma closedin_closed_trans: "closedin (subtopology euclidean T) S \ closed T \ closed S" - by (auto simp add: closedin_closed intro: closedin_trans) - -subsection{* Connectedness *} - -definition "connected S \ - ~(\e1 e2. open e1 \ open e2 \ S \ (e1 \ e2) \ (e1 \ e2 \ S = {}) - \ ~(e1 \ S = {}) \ ~(e2 \ S = {}))" - -lemma connected_local: - "connected S \ ~(\e1 e2. - openin (subtopology euclidean S) e1 \ - openin (subtopology euclidean S) e2 \ - S \ e1 \ e2 \ - e1 \ e2 = {} \ - ~(e1 = {}) \ - ~(e2 = {}))" -unfolding connected_def openin_open by (safe, blast+) - -lemma exists_diff: "(\S. P(UNIV - S)) \ (\S. P S)" (is "?lhs \ ?rhs") -proof- - - {assume "?lhs" hence ?rhs by blast } - moreover - {fix S assume H: "P S" - have "S = UNIV - (UNIV - S)" by auto - with H have "P (UNIV - (UNIV - S))" by metis } - ultimately show ?thesis by metis -qed - -lemma connected_clopen: "connected S \ - (\T. openin (subtopology euclidean S) T \ - closedin (subtopology euclidean S) T \ T = {} \ T = S)" (is "?lhs \ ?rhs") -proof- - have " \ connected S \ (\e1 e2. open e1 \ open (UNIV - e2) \ S \ e1 \ (UNIV - e2) \ e1 \ (UNIV - e2) \ S = {} \ e1 \ S \ {} \ (UNIV - e2) \ S \ {})" - unfolding connected_def openin_open closedin_closed - apply (subst exists_diff) by blast - hence th0: "connected S \ \ (\e2 e1. closed e2 \ open e1 \ S \ e1 \ (UNIV - e2) \ e1 \ (UNIV - e2) \ S = {} \ e1 \ S \ {} \ (UNIV - e2) \ S \ {})" - (is " _ \ \ (\e2 e1. ?P e2 e1)") apply (simp add: closed_def Compl_eq_Diff_UNIV) by metis - - have th1: "?rhs \ \ (\t' t. closed t'\t = S\t' \ t\{} \ t\S \ (\t'. open t' \ t = S \ t'))" - (is "_ \ \ (\t' t. ?Q t' t)") - unfolding connected_def openin_open closedin_closed by auto - {fix e2 - {fix e1 have "?P e2 e1 \ (\t. closed e2 \ t = S\e2 \ open e1 \ t = S\e1 \ t\{} \ t\S)" - by auto} - then have "(\e1. ?P e2 e1) \ (\t. ?Q e2 t)" by metis} - then have "\e2. (\e1. ?P e2 e1) \ (\t. ?Q e2 t)" by blast - then show ?thesis unfolding th0 th1 by simp -qed - -lemma connected_empty[simp, intro]: "connected {}" - by (simp add: connected_def) - -subsection{* Hausdorff and other separation properties *} - -class t0_space = - assumes t0_space: "x \ y \ \U. open U \ \ (x \ U \ y \ U)" - -class t1_space = - assumes t1_space: "x \ y \ \U V. open U \ open V \ x \ U \ y \ U \ x \ V \ y \ V" -begin - -subclass t0_space -proof -qed (fast dest: t1_space) - -end - -text {* T2 spaces are also known as Hausdorff spaces. *} - -class t2_space = - assumes hausdorff: "x \ y \ \U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" -begin - -subclass t1_space -proof -qed (fast dest: hausdorff) - -end - -instance metric_space \ t2_space -proof - fix x y :: "'a::metric_space" - assume xy: "x \ y" - let ?U = "ball x (dist x y / 2)" - let ?V = "ball y (dist x y / 2)" - have th0: "\d x y z. (d x z :: real) <= d x y + d y z \ d y z = d z y - ==> ~(d x y * 2 < d x z \ d z y * 2 < d x z)" by arith - have "open ?U \ open ?V \ x \ ?U \ y \ ?V \ ?U \ ?V = {}" - using dist_pos_lt[OF xy] th0[of dist,OF dist_triangle dist_commute] - by (auto simp add: expand_set_eq) - then show "\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" - by blast -qed - -lemma separation_t2: - fixes x y :: "'a::t2_space" - shows "x \ y \ (\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {})" - using hausdorff[of x y] by blast - -lemma separation_t1: - fixes x y :: "'a::t1_space" - shows "x \ y \ (\U V. open U \ open V \ x \U \ y\ U \ x\V \ y\V)" - using t1_space[of x y] by blast - -lemma separation_t0: - fixes x y :: "'a::t0_space" - shows "x \ y \ (\U. open U \ ~(x\U \ y\U))" - using t0_space[of x y] by blast - -subsection{* Limit points *} - -definition - islimpt:: "'a::topological_space \ 'a set \ bool" - (infixr "islimpt" 60) where - "x islimpt S \ (\T. x\T \ open T \ (\y\S. y\T \ y\x))" - -lemma islimptI: - assumes "\T. x \ T \ open T \ \y\S. y \ T \ y \ x" - shows "x islimpt S" - using assms unfolding islimpt_def by auto - -lemma islimptE: - assumes "x islimpt S" and "x \ T" and "open T" - obtains y where "y \ S" and "y \ T" and "y \ x" - using assms unfolding islimpt_def by auto - -lemma islimpt_subset: "x islimpt S \ S \ T ==> x islimpt T" by (auto simp add: islimpt_def) - -lemma islimpt_approachable: - fixes x :: "'a::metric_space" - shows "x islimpt S \ (\e>0. \x'\S. x' \ x \ dist x' x < e)" - unfolding islimpt_def - apply auto - apply(erule_tac x="ball x e" in allE) - apply auto - apply(rule_tac x=y in bexI) - apply (auto simp add: dist_commute) - apply (simp add: open_dist, drule (1) bspec) - apply (clarify, drule spec, drule (1) mp, auto) - done - -lemma islimpt_approachable_le: - fixes x :: "'a::metric_space" - shows "x islimpt S \ (\e>0. \x'\ S. x' \ x \ dist x' x <= e)" - unfolding islimpt_approachable - using approachable_lt_le[where f="\x'. dist x' x" and P="\x'. \ (x'\S \ x'\x)"] - by metis (* FIXME: VERY slow! *) - -class perfect_space = - (* FIXME: perfect_space should inherit from topological_space *) - assumes islimpt_UNIV [simp, intro]: "(x::'a::metric_space) islimpt UNIV" - -lemma perfect_choose_dist: - fixes x :: "'a::perfect_space" - shows "0 < r \ \a. a \ x \ dist a x < r" -using islimpt_UNIV [of x] -by (simp add: islimpt_approachable) - -instance real :: perfect_space -apply default -apply (rule islimpt_approachable [THEN iffD2]) -apply (clarify, rule_tac x="x + e/2" in bexI) -apply (auto simp add: dist_norm) -done - -instance "^" :: (perfect_space, finite) perfect_space -proof - fix x :: "'a ^ 'b" - { - fix e :: real assume "0 < e" - def a \ "x $ undefined" - have "a islimpt UNIV" by (rule islimpt_UNIV) - with `0 < e` obtain b where "b \ a" and "dist b a < e" - unfolding islimpt_approachable by auto - def y \ "Cart_lambda ((Cart_nth x)(undefined := b))" - from `b \ a` have "y \ x" - unfolding a_def y_def by (simp add: Cart_eq) - from `dist b a < e` have "dist y x < e" - unfolding dist_vector_def a_def y_def - apply simp - apply (rule le_less_trans [OF setL2_le_setsum [OF zero_le_dist]]) - apply (subst setsum_diff1' [where a=undefined], simp, simp, simp) - done - from `y \ x` and `dist y x < e` - have "\y\UNIV. y \ x \ dist y x < e" by auto - } - then show "x islimpt UNIV" unfolding islimpt_approachable by blast -qed - -lemma closed_limpt: "closed S \ (\x. x islimpt S \ x \ S)" - unfolding closed_def - apply (subst open_subopen) - apply (simp add: islimpt_def subset_eq Compl_eq_Diff_UNIV) - by (metis DiffE DiffI UNIV_I insertCI insert_absorb mem_def) - -lemma islimpt_EMPTY[simp]: "\ x islimpt {}" - unfolding islimpt_def by auto - -lemma closed_positive_orthant: "closed {x::real^'n::finite. \i. 0 \x$i}" -proof- - let ?U = "UNIV :: 'n set" - let ?O = "{x::real^'n. \i. x$i\0}" - {fix x:: "real^'n" and i::'n assume H: "\e>0. \x'\?O. x' \ x \ dist x' x < e" - and xi: "x$i < 0" - from xi have th0: "-x$i > 0" by arith - from H[rule_format, OF th0] obtain x' where x': "x' \?O" "x' \ x" "dist x' x < -x $ i" by blast - have th:" \b a (x::real). abs x <= b \ b <= a ==> ~(a + x < 0)" by arith - have th': "\x (y::real). x < 0 \ 0 <= y ==> abs x <= abs (y - x)" by arith - have th1: "\x$i\ \ \(x' - x)$i\" using x'(1) xi - apply (simp only: vector_component) - by (rule th') auto - have th2: "\dist x x'\ \ \(x' - x)$i\" using component_le_norm[of "x'-x" i] - apply (simp add: dist_norm) by norm - from th[OF th1 th2] x'(3) have False by (simp add: dist_commute) } - then show ?thesis unfolding closed_limpt islimpt_approachable - unfolding not_le[symmetric] by blast -qed - -lemma finite_set_avoid: - fixes a :: "'a::metric_space" - assumes fS: "finite S" shows "\d>0. \x\S. x \ a \ d <= dist a x" -proof(induct rule: finite_induct[OF fS]) - case 1 thus ?case apply auto by ferrack -next - case (2 x F) - from 2 obtain d where d: "d >0" "\x\F. x\a \ d \ dist a x" by blast - {assume "x = a" hence ?case using d by auto } - moreover - {assume xa: "x\a" - let ?d = "min d (dist a x)" - have dp: "?d > 0" using xa d(1) using dist_nz by auto - from d have d': "\x\F. x\a \ ?d \ dist a x" by auto - with dp xa have ?case by(auto intro!: exI[where x="?d"]) } - ultimately show ?case by blast -qed - -lemma islimpt_finite: - fixes S :: "'a::metric_space set" - assumes fS: "finite S" shows "\ a islimpt S" - unfolding islimpt_approachable - using finite_set_avoid[OF fS, of a] by (metis dist_commute not_le) - -lemma islimpt_Un: "x islimpt (S \ T) \ x islimpt S \ x islimpt T" - apply (rule iffI) - defer - apply (metis Un_upper1 Un_upper2 islimpt_subset) - unfolding islimpt_def - apply (rule ccontr, clarsimp, rename_tac A B) - apply (drule_tac x="A \ B" in spec) - apply (auto simp add: open_Int) - done - -lemma discrete_imp_closed: - fixes S :: "'a::metric_space set" - assumes e: "0 < e" and d: "\x \ S. \y \ S. dist y x < e \ y = x" - shows "closed S" -proof- - {fix x assume C: "\e>0. \x'\S. x' \ x \ dist x' x < e" - from e have e2: "e/2 > 0" by arith - from C[rule_format, OF e2] obtain y where y: "y \ S" "y\x" "dist y x < e/2" by blast - let ?m = "min (e/2) (dist x y) " - from e2 y(2) have mp: "?m > 0" by (simp add: dist_nz[THEN sym]) - from C[rule_format, OF mp] obtain z where z: "z \ S" "z\x" "dist z x < ?m" by blast - have th: "dist z y < e" using z y - by (intro dist_triangle_lt [where z=x], simp) - from d[rule_format, OF y(1) z(1) th] y z - have False by (auto simp add: dist_commute)} - then show ?thesis by (metis islimpt_approachable closed_limpt [where 'a='a]) -qed - -subsection{* Interior of a Set *} -definition "interior S = {x. \T. open T \ x \ T \ T \ S}" - -lemma interior_eq: "interior S = S \ open S" - apply (simp add: expand_set_eq interior_def) - apply (subst (2) open_subopen) by (safe, blast+) - -lemma interior_open: "open S ==> (interior S = S)" by (metis interior_eq) - -lemma interior_empty[simp]: "interior {} = {}" by (simp add: interior_def) - -lemma open_interior[simp, intro]: "open(interior S)" - apply (simp add: interior_def) - apply (subst open_subopen) by blast - -lemma interior_interior[simp]: "interior(interior S) = interior S" by (metis interior_eq open_interior) -lemma interior_subset: "interior S \ S" by (auto simp add: interior_def) -lemma subset_interior: "S \ T ==> (interior S) \ (interior T)" by (auto simp add: interior_def) -lemma interior_maximal: "T \ S \ open T ==> T \ (interior S)" by (auto simp add: interior_def) -lemma interior_unique: "T \ S \ open T \ (\T'. T' \ S \ open T' \ T' \ T) \ interior S = T" - by (metis equalityI interior_maximal interior_subset open_interior) -lemma mem_interior: "x \ interior S \ (\e. 0 < e \ ball x e \ S)" - apply (simp add: interior_def) - by (metis open_contains_ball centre_in_ball open_ball subset_trans) - -lemma open_subset_interior: "open S ==> S \ interior T \ S \ T" - by (metis interior_maximal interior_subset subset_trans) - -lemma interior_inter[simp]: "interior(S \ T) = interior S \ interior T" - apply (rule equalityI, simp) - apply (metis Int_lower1 Int_lower2 subset_interior) - by (metis Int_mono interior_subset open_Int open_interior open_subset_interior) - -lemma interior_limit_point [intro]: - fixes x :: "'a::perfect_space" - assumes x: "x \ interior S" shows "x islimpt S" -proof- - from x obtain e where e: "e>0" "\x'. dist x x' < e \ x' \ S" - unfolding mem_interior subset_eq Ball_def mem_ball by blast - { - fix d::real assume d: "d>0" - let ?m = "min d e" - have mde2: "0 < ?m" using e(1) d(1) by simp - from perfect_choose_dist [OF mde2, of x] - obtain y where "y \ x" and "dist y x < ?m" by blast - then have "dist y x < e" "dist y x < d" by simp_all - from `dist y x < e` e(2) have "y \ S" by (simp add: dist_commute) - have "\x'\S. x'\ x \ dist x' x < d" - using `y \ S` `y \ x` `dist y x < d` by fast - } - then show ?thesis unfolding islimpt_approachable by blast -qed - -lemma interior_closed_Un_empty_interior: - assumes cS: "closed S" and iT: "interior T = {}" - shows "interior(S \ T) = interior S" -proof - show "interior S \ interior (S\T)" - by (rule subset_interior, blast) -next - show "interior (S \ T) \ interior S" - proof - fix x assume "x \ interior (S \ T)" - then obtain R where "open R" "x \ R" "R \ S \ T" - unfolding interior_def by fast - show "x \ interior S" - proof (rule ccontr) - assume "x \ interior S" - with `x \ R` `open R` obtain y where "y \ R - S" - unfolding interior_def expand_set_eq by fast - from `open R` `closed S` have "open (R - S)" by (rule open_Diff) - from `R \ S \ T` have "R - S \ T" by fast - from `y \ R - S` `open (R - S)` `R - S \ T` `interior T = {}` - show "False" unfolding interior_def by fast - qed - qed -qed - - -subsection{* Closure of a Set *} - -definition "closure S = S \ {x | x. x islimpt S}" - -lemma closure_interior: "closure S = UNIV - interior (UNIV - S)" -proof- - { fix x - have "x\UNIV - interior (UNIV - S) \ x \ closure S" (is "?lhs = ?rhs") - proof - let ?exT = "\ y. (\T. open T \ y \ T \ T \ UNIV - S)" - assume "?lhs" - hence *:"\ ?exT x" - unfolding interior_def - by simp - { assume "\ ?rhs" - hence False using * - unfolding closure_def islimpt_def - by blast - } - thus "?rhs" - by blast - next - assume "?rhs" thus "?lhs" - unfolding closure_def interior_def islimpt_def - by blast - qed - } - thus ?thesis - by blast -qed - -lemma interior_closure: "interior S = UNIV - (closure (UNIV - S))" -proof- - { fix x - have "x \ interior S \ x \ UNIV - (closure (UNIV - S))" - unfolding interior_def closure_def islimpt_def - by blast (* FIXME: VERY slow! *) - } - thus ?thesis - by blast -qed - -lemma closed_closure[simp, intro]: "closed (closure S)" -proof- - have "closed (UNIV - interior (UNIV -S))" by blast - thus ?thesis using closure_interior[of S] by simp -qed - -lemma closure_hull: "closure S = closed hull S" -proof- - have "S \ closure S" - unfolding closure_def - by blast - moreover - have "closed (closure S)" - using closed_closure[of S] - by assumption - moreover - { fix t - assume *:"S \ t" "closed t" - { fix x - assume "x islimpt S" - hence "x islimpt t" using *(1) - using islimpt_subset[of x, of S, of t] - by blast - } - with * have "closure S \ t" - unfolding closure_def - using closed_limpt[of t] - by auto - } - ultimately show ?thesis - using hull_unique[of S, of "closure S", of closed] - unfolding mem_def - by simp -qed - -lemma closure_eq: "closure S = S \ closed S" - unfolding closure_hull - using hull_eq[of closed, unfolded mem_def, OF closed_Inter, of S] - by (metis mem_def subset_eq) - -lemma closure_closed[simp]: "closed S \ closure S = S" - using closure_eq[of S] - by simp - -lemma closure_closure[simp]: "closure (closure S) = closure S" - unfolding closure_hull - using hull_hull[of closed S] - by assumption - -lemma closure_subset: "S \ closure S" - unfolding closure_hull - using hull_subset[of S closed] - by assumption - -lemma subset_closure: "S \ T \ closure S \ closure T" - unfolding closure_hull - using hull_mono[of S T closed] - by assumption - -lemma closure_minimal: "S \ T \ closed T \ closure S \ T" - using hull_minimal[of S T closed] - unfolding closure_hull mem_def - by simp - -lemma closure_unique: "S \ T \ closed T \ (\ T'. S \ T' \ closed T' \ T \ T') \ closure S = T" - using hull_unique[of S T closed] - unfolding closure_hull mem_def - by simp - -lemma closure_empty[simp]: "closure {} = {}" - using closed_empty closure_closed[of "{}"] - by simp - -lemma closure_univ[simp]: "closure UNIV = UNIV" - using closure_closed[of UNIV] - by simp - -lemma closure_eq_empty: "closure S = {} \ S = {}" - using closure_empty closure_subset[of S] - by blast - -lemma closure_subset_eq: "closure S \ S \ closed S" - using closure_eq[of S] closure_subset[of S] - by simp - -lemma open_inter_closure_eq_empty: - "open S \ (S \ closure T) = {} \ S \ T = {}" - using open_subset_interior[of S "UNIV - T"] - using interior_subset[of "UNIV - T"] - unfolding closure_interior - by auto - -lemma open_inter_closure_subset: - "open S \ (S \ (closure T)) \ closure(S \ T)" -proof - fix x - assume as: "open S" "x \ S \ closure T" - { assume *:"x islimpt T" - have "x islimpt (S \ T)" - proof (rule islimptI) - fix A - assume "x \ A" "open A" - with as have "x \ A \ S" "open (A \ S)" - by (simp_all add: open_Int) - with * obtain y where "y \ T" "y \ A \ S" "y \ x" - by (rule islimptE) - hence "y \ S \ T" "y \ A \ y \ x" - by simp_all - thus "\y\(S \ T). y \ A \ y \ x" .. - qed - } - then show "x \ closure (S \ T)" using as - unfolding closure_def - by blast -qed - -lemma closure_complement: "closure(UNIV - S) = UNIV - interior(S)" -proof- - have "S = UNIV - (UNIV - S)" - by auto - thus ?thesis - unfolding closure_interior - by auto -qed - -lemma interior_complement: "interior(UNIV - S) = UNIV - closure(S)" - unfolding closure_interior - by blast - -subsection{* Frontier (aka boundary) *} - -definition "frontier S = closure S - interior S" - -lemma frontier_closed: "closed(frontier S)" - by (simp add: frontier_def closed_Diff) - -lemma frontier_closures: "frontier S = (closure S) \ (closure(UNIV - S))" - by (auto simp add: frontier_def interior_closure) - -lemma frontier_straddle: - fixes a :: "'a::metric_space" - shows "a \ frontier S \ (\e>0. (\x\S. dist a x < e) \ (\x. x \ S \ dist a x < e))" (is "?lhs \ ?rhs") -proof - assume "?lhs" - { fix e::real - assume "e > 0" - let ?rhse = "(\x\S. dist a x < e) \ (\x. x \ S \ dist a x < e)" - { assume "a\S" - have "\x\S. dist a x < e" using `e>0` `a\S` by(rule_tac x=a in bexI) auto - moreover have "\x. x \ S \ dist a x < e" using `?lhs` `a\S` - unfolding frontier_closures closure_def islimpt_def using `e>0` - by (auto, erule_tac x="ball a e" in allE, auto) - ultimately have ?rhse by auto - } - moreover - { assume "a\S" - hence ?rhse using `?lhs` - unfolding frontier_closures closure_def islimpt_def - using open_ball[of a e] `e > 0` - by (auto, erule_tac x = "ball a e" in allE, auto) (* FIXME: VERY slow! *) - } - ultimately have ?rhse by auto - } - thus ?rhs by auto -next - assume ?rhs - moreover - { fix T assume "a\S" and - as:"\e>0. (\x\S. dist a x < e) \ (\x. x \ S \ dist a x < e)" "a \ S" "a \ T" "open T" - from `open T` `a \ T` have "\e>0. ball a e \ T" unfolding open_contains_ball[of T] by auto - then obtain e where "e>0" "ball a e \ T" by auto - then obtain y where y:"y\S" "dist a y < e" using as(1) by auto - have "\y\S. y \ T \ y \ a" - using `dist a y < e` `ball a e \ T` unfolding ball_def using `y\S` `a\S` by auto - } - hence "a \ closure S" unfolding closure_def islimpt_def using `?rhs` by auto - moreover - { fix T assume "a \ T" "open T" "a\S" - then obtain e where "e>0" and balle: "ball a e \ T" unfolding open_contains_ball using `?rhs` by auto - obtain x where "x \ S" "dist a x < e" using `?rhs` using `e>0` by auto - hence "\y\UNIV - S. y \ T \ y \ a" using balle `a\S` unfolding ball_def by (rule_tac x=x in bexI)auto - } - hence "a islimpt (UNIV - S) \ a\S" unfolding islimpt_def by auto - ultimately show ?lhs unfolding frontier_closures using closure_def[of "UNIV - S"] by auto -qed - -lemma frontier_subset_closed: "closed S \ frontier S \ S" - by (metis frontier_def closure_closed Diff_subset) - -lemma frontier_empty: "frontier {} = {}" - by (simp add: frontier_def closure_empty) - -lemma frontier_subset_eq: "frontier S \ S \ closed S" -proof- - { assume "frontier S \ S" - hence "closure S \ S" using interior_subset unfolding frontier_def by auto - hence "closed S" using closure_subset_eq by auto - } - thus ?thesis using frontier_subset_closed[of S] by auto -qed - -lemma frontier_complement: "frontier(UNIV - S) = frontier S" - by (auto simp add: frontier_def closure_complement interior_complement) - -lemma frontier_disjoint_eq: "frontier S \ S = {} \ open S" - using frontier_complement frontier_subset_eq[of "UNIV - S"] - unfolding open_closed Compl_eq_Diff_UNIV by auto - -subsection{* Common nets and The "within" modifier for nets. *} - -definition - at_infinity :: "'a::real_normed_vector net" where - "at_infinity = Abs_net (range (\r. {x. r \ norm x}))" - -definition - indirection :: "'a::real_normed_vector \ 'a \ 'a net" (infixr "indirection" 70) where - "a indirection v = (at a) within {b. \c\0. b - a = scaleR c v}" - -text{* Prove That They are all nets. *} - -lemma Rep_net_at_infinity: - "Rep_net at_infinity = range (\r. {x. r \ norm x})" -unfolding at_infinity_def -apply (rule Abs_net_inverse') -apply (rule image_nonempty, simp) -apply (clarsimp, rename_tac r s) -apply (rule_tac x="max r s" in exI, auto) -done - -lemma within_UNIV: "net within UNIV = net" - by (simp add: Rep_net_inject [symmetric] Rep_net_within) - -subsection{* Identify Trivial limits, where we can't approach arbitrarily closely. *} - -definition - trivial_limit :: "'a net \ bool" where - "trivial_limit net \ {} \ Rep_net net" - -lemma trivial_limit_within: - shows "trivial_limit (at a within S) \ \ a islimpt S" -proof - assume "trivial_limit (at a within S)" - thus "\ a islimpt S" - unfolding trivial_limit_def - unfolding Rep_net_within Rep_net_at - unfolding islimpt_def - apply (clarsimp simp add: expand_set_eq) - apply (rename_tac T, rule_tac x=T in exI) - apply (clarsimp, drule_tac x=y in spec, simp) - done -next - assume "\ a islimpt S" - thus "trivial_limit (at a within S)" - unfolding trivial_limit_def - unfolding Rep_net_within Rep_net_at - unfolding islimpt_def - apply (clarsimp simp add: image_image) - apply (rule_tac x=T in image_eqI) - apply (auto simp add: expand_set_eq) - done -qed - -lemma trivial_limit_at_iff: "trivial_limit (at a) \ \ a islimpt UNIV" - using trivial_limit_within [of a UNIV] - by (simp add: within_UNIV) - -lemma trivial_limit_at: - fixes a :: "'a::perfect_space" - shows "\ trivial_limit (at a)" - by (simp add: trivial_limit_at_iff) - -lemma trivial_limit_at_infinity: - "\ trivial_limit (at_infinity :: ('a::{real_normed_vector,zero_neq_one}) net)" - (* FIXME: find a more appropriate type class *) - 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_sgn) - done - -lemma trivial_limit_sequentially: "\ trivial_limit sequentially" - by (auto simp add: trivial_limit_def Rep_net_sequentially) - -subsection{* Some property holds "sufficiently close" to the limit point. *} - -lemma eventually_at: (* FIXME: this replaces Limits.eventually_at *) - "eventually P (at a) \ (\d>0. \x. 0 < dist x a \ dist x a < d \ P x)" -unfolding eventually_at dist_nz by auto - -lemma eventually_at_infinity: - "eventually P at_infinity \ (\b. \x. norm x >= b \ P x)" -unfolding eventually_def Rep_net_at_infinity by auto - -lemma eventually_within: "eventually P (at a within S) \ - (\d>0. \x\S. 0 < dist x a \ dist x a < d \ P x)" -unfolding eventually_within eventually_at dist_nz by auto - -lemma eventually_within_le: "eventually P (at a within S) \ - (\d>0. \x\S. 0 < dist x a \ dist x a <= d \ P x)" (is "?lhs = ?rhs") -unfolding eventually_within -apply safe -apply (rule_tac x="d/2" in exI, simp) -apply (rule_tac x="d" in exI, simp) -done - -lemma eventually_happens: "eventually P net ==> trivial_limit net \ (\x. P x)" - unfolding eventually_def trivial_limit_def - using Rep_net_nonempty [of net] by auto - -lemma always_eventually: "(\x. P x) ==> eventually P net" - unfolding eventually_def trivial_limit_def - using Rep_net_nonempty [of net] by auto - -lemma trivial_limit_eventually: "trivial_limit net \ eventually P net" - unfolding trivial_limit_def eventually_def by auto - -lemma eventually_False: "eventually (\x. False) net \ trivial_limit net" - unfolding trivial_limit_def eventually_def by auto - -lemma trivial_limit_eq: "trivial_limit net \ (\P. eventually P net)" - apply (safe elim!: trivial_limit_eventually) - apply (simp add: eventually_False [symmetric]) - done - -text{* Combining theorems for "eventually" *} - -lemma eventually_conjI: - "\eventually (\x. P x) net; eventually (\x. Q x) net\ - \ eventually (\x. P x \ Q x) net" -by (rule eventually_conj) - -lemma eventually_rev_mono: - "eventually P net \ (\x. P x \ Q x) \ eventually Q net" -using eventually_mono [of P Q] by fast - -lemma eventually_and: " eventually (\x. P x \ Q x) net \ eventually P net \ eventually Q net" - by (auto intro!: eventually_conjI elim: eventually_rev_mono) - -lemma eventually_false: "eventually (\x. False) net \ trivial_limit net" - by (auto simp add: eventually_False) - -lemma not_eventually: "(\x. \ P x ) \ ~(trivial_limit net) ==> ~(eventually (\x. P x) net)" - by (simp add: eventually_False) - -subsection{* Limits, defined as vacuously true when the limit is trivial. *} - - text{* Notation Lim to avoid collition with lim defined in analysis *} -definition - Lim :: "'a net \ ('a \ 'b::t2_space) \ 'b" where - "Lim net f = (THE l. (f ---> l) net)" - -lemma Lim: - "(f ---> l) net \ - trivial_limit net \ - (\e>0. eventually (\x. dist (f x) l < e) net)" - unfolding tendsto_iff trivial_limit_eq by auto - - -text{* Show that they yield usual definitions in the various cases. *} - -lemma Lim_within_le: "(f ---> l)(at a within S) \ - (\e>0. \d>0. \x\S. 0 < dist x a \ dist x a <= d \ dist (f x) l < e)" - by (auto simp add: tendsto_iff eventually_within_le) - -lemma Lim_within: "(f ---> l) (at a within S) \ - (\e >0. \d>0. \x \ S. 0 < dist x a \ dist x a < d \ dist (f x) l < e)" - by (auto simp add: tendsto_iff eventually_within) - -lemma Lim_at: "(f ---> l) (at a) \ - (\e >0. \d>0. \x. 0 < dist x a \ dist x a < d \ dist (f x) l < e)" - by (auto simp add: tendsto_iff eventually_at) - -lemma Lim_at_iff_LIM: "(f ---> l) (at a) \ f -- a --> l" - unfolding Lim_at LIM_def by (simp only: zero_less_dist_iff) - -lemma Lim_at_infinity: - "(f ---> l) at_infinity \ (\e>0. \b. \x. norm x >= b \ dist (f x) l < e)" - by (auto simp add: tendsto_iff eventually_at_infinity) - -lemma Lim_sequentially: - "(S ---> l) sequentially \ - (\e>0. \N. \n\N. dist (S n) l < e)" - by (auto simp add: tendsto_iff eventually_sequentially) - -lemma Lim_sequentially_iff_LIMSEQ: "(S ---> l) sequentially \ S ----> l" - unfolding Lim_sequentially LIMSEQ_def .. - -lemma Lim_eventually: "eventually (\x. f x = l) net \ (f ---> l) net" - by (rule topological_tendstoI, auto elim: eventually_rev_mono) - -text{* The expected monotonicity property. *} - -lemma Lim_within_empty: "(f ---> l) (net within {})" - unfolding tendsto_def Limits.eventually_within by simp - -lemma Lim_within_subset: "(f ---> l) (net within S) \ T \ S \ (f ---> l) (net within T)" - unfolding tendsto_def Limits.eventually_within - by (auto elim!: eventually_elim1) - -lemma Lim_Un: assumes "(f ---> l) (net within S)" "(f ---> l) (net within T)" - shows "(f ---> l) (net within (S \ T))" - using assms unfolding tendsto_def Limits.eventually_within - apply clarify - apply (drule spec, drule (1) mp, drule (1) mp) - apply (drule spec, drule (1) mp, drule (1) mp) - apply (auto elim: eventually_elim2) - done - -lemma Lim_Un_univ: - "(f ---> l) (net within S) \ (f ---> l) (net within T) \ S \ T = UNIV - ==> (f ---> l) net" - by (metis Lim_Un within_UNIV) - -text{* Interrelations between restricted and unrestricted limits. *} - -lemma Lim_at_within: "(f ---> l) net ==> (f ---> l)(net within S)" - (* FIXME: rename *) - unfolding tendsto_def Limits.eventually_within - apply (clarify, drule spec, drule (1) mp, drule (1) mp) - by (auto elim!: eventually_elim1) - -lemma Lim_within_open: - fixes f :: "'a::topological_space \ 'b::topological_space" - assumes"a \ S" "open S" - shows "(f ---> l)(at a within S) \ (f ---> l)(at a)" (is "?lhs \ ?rhs") -proof - assume ?lhs - { fix A assume "open A" "l \ A" - with `?lhs` have "eventually (\x. f x \ A) (at a within S)" - by (rule topological_tendstoD) - hence "eventually (\x. x \ S \ f x \ A) (at a)" - unfolding Limits.eventually_within . - then obtain T where "open T" "a \ T" "\x\T. x \ a \ x \ S \ f x \ A" - unfolding eventually_at_topological by fast - hence "open (T \ S)" "a \ T \ S" "\x\(T \ S). x \ a \ f x \ A" - using assms by auto - hence "\T. open T \ a \ T \ (\x\T. x \ a \ f x \ A)" - by fast - hence "eventually (\x. f x \ A) (at a)" - unfolding eventually_at_topological . - } - thus ?rhs by (rule topological_tendstoI) -next - assume ?rhs - thus ?lhs by (rule Lim_at_within) -qed - -text{* Another limit point characterization. *} - -lemma islimpt_sequential: - fixes x :: "'a::metric_space" (* FIXME: generalize to topological_space *) - shows "x islimpt S \ (\f. (\n::nat. f n \ S -{x}) \ (f ---> x) sequentially)" - (is "?lhs = ?rhs") -proof - assume ?lhs - then obtain f where f:"\y. y>0 \ f y \ S \ f y \ x \ dist (f y) x < y" - unfolding islimpt_approachable using choice[of "\e y. e>0 \ y\S \ y\x \ dist y x < e"] by auto - { fix n::nat - have "f (inverse (real n + 1)) \ S - {x}" using f by auto - } - moreover - { fix e::real assume "e>0" - hence "\N::nat. inverse (real (N + 1)) < e" using real_arch_inv[of e] apply (auto simp add: Suc_pred') apply(rule_tac x="n - 1" in exI) by auto - 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) - moreover have "\n\N. dist (f (inverse (real n + 1))) x < (inverse (real n + 1))" using f `e>0` by auto - ultimately have "\N::nat. \n\N. dist (f (inverse (real n + 1))) x < e" apply(rule_tac x=N in exI) apply auto apply(erule_tac x=n in allE)+ by auto - } - hence " ((\n. f (inverse (real n + 1))) ---> x) sequentially" - unfolding Lim_sequentially using f by auto - ultimately show ?rhs apply (rule_tac x="(\n::nat. f (inverse (real n + 1)))" in exI) by auto -next - assume ?rhs - then obtain f::"nat\'a" where f:"(\n. f n \ S - {x})" "(\e>0. \N. \n\N. dist (f n) x < e)" unfolding Lim_sequentially by auto - { fix e::real assume "e>0" - then obtain N where "dist (f N) x < e" using f(2) by auto - moreover have "f N\S" "f N \ x" using f(1) by auto - ultimately have "\x'\S. x' \ x \ dist x' x < e" by auto - } - thus ?lhs unfolding islimpt_approachable by auto -qed - -text{* Basic arithmetical combining theorems for limits. *} - -lemma Lim_linear: - assumes "(f ---> l) net" "bounded_linear h" - shows "((\x. h (f x)) ---> h l) net" -using `bounded_linear h` `(f ---> l) net` -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 \ 'b::real_normed_vector" - 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" - shows "(f ---> l) net ==> ((\x. -(f x)) ---> -l) net" - by (rule tendsto_minus) - -lemma Lim_add: fixes f :: "'a \ 'b::real_normed_vector" shows - "(f ---> l) net \ (g ---> m) net \ ((\x. f(x) + g(x)) ---> l + m) net" - by (rule tendsto_add) - -lemma Lim_sub: - fixes f :: "'a \ 'b::real_normed_vector" - shows "(f ---> l) net \ (g ---> m) net \ ((\x. f(x) - g(x)) ---> l - m) net" - by (rule tendsto_diff) - -lemma Lim_null: - fixes f :: "'a \ 'b::real_normed_vector" - shows "(f ---> l) net \ ((\x. f(x) - l) ---> 0) net" by (simp add: Lim dist_norm) - -lemma Lim_null_norm: - fixes f :: "'a \ 'b::real_normed_vector" - 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" "(g ---> 0) net" - shows "(f ---> 0) net" -proof(simp add: tendsto_iff, rule+) - fix e::real assume "0 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 (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: - fixes f :: "'a \ 'b::metric_space ^ 'n::finite" - shows "(f ---> l) net \ ((\a. f a $i) ---> l$i) net" - unfolding tendsto_iff - 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" - fixes g :: "'a \ 'c::real_normed_vector" - assumes "eventually (\n. norm(f n) <= norm(g n)) net" "(g ---> 0) net" - shows "(f ---> 0) net" -proof (rule tendstoI) - fix e::real assume "e>0" - { fix x - assume "norm (f x) \ norm (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) \ norm (g x)" "\x. dist (g x) 0 < e" net] - using eventually_mono[of "\x. norm (f x) \ norm (g x) \ dist (g x) 0 < e" "\x. dist (f x) 0 < e" net] - using assms `e>0` unfolding tendsto_iff by blast -qed - -text{* Deducing things about the limit from the elements. *} - -lemma Lim_in_closed_set: - assumes "closed S" "eventually (\x. f(x) \ S) net" "\(trivial_limit net)" "(f ---> l) net" - shows "l \ S" -proof (rule ccontr) - assume "l \ S" - with `closed S` have "open (- S)" "l \ - S" - by (simp_all add: open_Compl) - with assms(4) have "eventually (\x. f x \ - S) net" - by (rule topological_tendstoD) - with assms(2) have "eventually (\x. False) net" - by (rule eventually_elim2) simp - with assms(3) show "False" - by (simp add: eventually_False) -qed - -text{* Need to prove closed(cball(x,e)) before deducing this as a corollary. *} - -lemma Lim_dist_ubound: - assumes "\(trivial_limit net)" "(f ---> l) net" "eventually (\x. dist a (f x) <= e) net" - shows "dist a l <= e" -proof (rule ccontr) - assume "\ dist a l \ e" - then have "0 < dist a l - e" by simp - with assms(2) have "eventually (\x. dist (f x) l < dist a l - e) net" - by (rule tendstoD) - with assms(3) have "eventually (\x. dist a (f x) \ e \ dist (f x) l < dist a l - e) net" - by (rule eventually_conjI) - then obtain w where "dist a (f w) \ e" "dist (f w) l < dist a l - e" - using assms(1) eventually_happens by auto - hence "dist a (f w) + dist (f w) l < e + (dist a l - e)" - by (rule add_le_less_mono) - hence "dist a (f w) + dist (f w) l < dist a l" - by simp - also have "\ \ dist a (f w) + dist (f w) l" - by (rule dist_triangle) - finally show False by simp -qed - -lemma Lim_norm_ubound: - fixes f :: "'a \ 'b::real_normed_vector" - assumes "\(trivial_limit net)" "(f ---> l) net" "eventually (\x. norm(f x) <= e) net" - shows "norm(l) <= e" -proof (rule ccontr) - assume "\ norm l \ e" - then have "0 < norm l - e" by simp - with assms(2) have "eventually (\x. dist (f x) l < norm l - e) net" - by (rule tendstoD) - with assms(3) have "eventually (\x. norm (f x) \ e \ dist (f x) l < norm l - e) net" - by (rule eventually_conjI) - then obtain w where "norm (f w) \ e" "dist (f w) l < norm l - e" - using assms(1) eventually_happens by auto - hence "norm (f w - l) < norm l - e" "norm (f w) \ e" by (simp_all add: dist_norm) - hence "norm (f w - l) + norm (f w) < norm l" by simp - hence "norm (f w - l - f w) < norm l" by (rule le_less_trans [OF norm_triangle_ineq4]) - thus False using `\ norm l \ e` by simp -qed - -lemma Lim_norm_lbound: - fixes f :: "'a \ 'b::real_normed_vector" - assumes "\ (trivial_limit net)" "(f ---> l) net" "eventually (\x. e <= norm(f x)) net" - shows "e \ norm l" -proof (rule ccontr) - assume "\ e \ norm l" - then have "0 < e - norm l" by simp - with assms(2) have "eventually (\x. dist (f x) l < e - norm l) net" - by (rule tendstoD) - with assms(3) have "eventually (\x. e \ norm (f x) \ dist (f x) l < e - norm l) net" - by (rule eventually_conjI) - then obtain w where "e \ norm (f w)" "dist (f w) l < e - norm l" - using assms(1) eventually_happens by auto - hence "norm (f w - l) + norm l < e" "e \ norm (f w)" by (simp_all add: dist_norm) - hence "norm (f w - l) + norm l < norm (f w)" by (rule less_le_trans) - hence "norm (f w - l + l) < norm (f w)" by (rule le_less_trans [OF norm_triangle_ineq]) - thus False by simp -qed - -text{* Uniqueness of the limit, when nontrivial. *} - -lemma Lim_unique: - fixes f :: "'a \ 'b::t2_space" - assumes "\ trivial_limit net" "(f ---> l) net" "(f ---> l') net" - shows "l = l'" -proof (rule ccontr) - assume "l \ l'" - obtain U V where "open U" "open V" "l \ U" "l' \ V" "U \ V = {}" - using hausdorff [OF `l \ l'`] by fast - have "eventually (\x. f x \ U) net" - using `(f ---> l) net` `open U` `l \ U` by (rule topological_tendstoD) - moreover - have "eventually (\x. f x \ V) net" - using `(f ---> l') net` `open V` `l' \ V` by (rule topological_tendstoD) - ultimately - have "eventually (\x. False) net" - proof (rule eventually_elim2) - fix x - assume "f x \ U" "f x \ V" - hence "f x \ U \ V" by simp - with `U \ V = {}` show "False" by simp - qed - with `\ trivial_limit net` show "False" - by (simp add: eventually_False) -qed - -lemma tendsto_Lim: - fixes f :: "'a \ 'b::t2_space" - shows "~(trivial_limit net) \ (f ---> l) net ==> Lim net f = l" - unfolding Lim_def using Lim_unique[of net f] by auto - -text{* Limit under bilinear function *} - -lemma Lim_bilinear: - assumes "(f ---> l) net" and "(g ---> m) net" and "bounded_bilinear h" - shows "((\x. h (f x) (g x)) ---> (h l m)) net" -using `bounded_bilinear h` `(f ---> l) net` `(g ---> m) net` -by (rule bounded_bilinear.tendsto) - -text{* These are special for limits out of the same vector space. *} - -lemma Lim_within_id: "(id ---> a) (at a within s)" - unfolding tendsto_def Limits.eventually_within eventually_at_topological - by auto - -lemma Lim_at_id: "(id ---> a) (at a)" -apply (subst within_UNIV[symmetric]) by (simp add: Lim_within_id) - -lemma Lim_at_zero: - fixes a :: "'a::real_normed_vector" - fixes l :: "'b::topological_space" - shows "(f ---> l) (at a) \ ((\x. f(a + x)) ---> l) (at 0)" (is "?lhs = ?rhs") -proof - assume "?lhs" - { fix S assume "open S" "l \ S" - with `?lhs` have "eventually (\x. f x \ S) (at a)" - by (rule topological_tendstoD) - then obtain d where d: "d>0" "\x. x \ a \ dist x a < d \ f x \ S" - unfolding Limits.eventually_at by fast - { fix x::"'a" assume "x \ 0 \ dist x 0 < d" - hence "f (a + x) \ S" using d - apply(erule_tac x="x+a" in allE) - by(auto simp add: comm_monoid_add.mult_commute dist_norm dist_commute) - } - hence "\d>0. \x. x \ 0 \ dist x 0 < d \ f (a + x) \ S" - using d(1) by auto - hence "eventually (\x. f (a + x) \ S) (at 0)" - unfolding Limits.eventually_at . - } - thus "?rhs" by (rule topological_tendstoI) -next - assume "?rhs" - { fix S assume "open S" "l \ S" - with `?rhs` have "eventually (\x. f (a + x) \ S) (at 0)" - by (rule topological_tendstoD) - then obtain d where d: "d>0" "\x. x \ 0 \ dist x 0 < d \ f (a + x) \ S" - unfolding Limits.eventually_at by fast - { fix x::"'a" assume "x \ a \ dist x a < d" - hence "f x \ S" using d apply(erule_tac x="x-a" in allE) - by(auto simp add: comm_monoid_add.mult_commute dist_norm dist_commute) - } - hence "\d>0. \x. x \ a \ dist x a < d \ f x \ S" using d(1) by auto - hence "eventually (\x. f x \ S) (at a)" unfolding Limits.eventually_at . - } - thus "?lhs" by (rule topological_tendstoI) -qed - -text{* It's also sometimes useful to extract the limit point from the net. *} - -definition - netlimit :: "'a::t2_space net \ 'a" where - "netlimit net = (SOME a. ((\x. x) ---> a) net)" - -lemma netlimit_within: - assumes "\ trivial_limit (at a within S)" - shows "netlimit (at a within S) = a" -unfolding netlimit_def -apply (rule some_equality) -apply (rule Lim_at_within) -apply (rule Lim_ident_at) -apply (erule Lim_unique [OF assms]) -apply (rule Lim_at_within) -apply (rule Lim_ident_at) -done - -lemma netlimit_at: - fixes a :: "'a::perfect_space" - shows "netlimit (at a) = a" - apply (subst within_UNIV[symmetric]) - using netlimit_within[of a UNIV] - by (simp add: trivial_limit_at within_UNIV) - -text{* Transformation of limit. *} - -lemma Lim_transform: - fixes f g :: "'a::type \ 'b::real_normed_vector" - assumes "((\x. f x - g x) ---> 0) net" "(f ---> l) net" - shows "(g ---> l) net" -proof- - from assms have "((\x. f x - g x - f x) ---> 0 - l) net" using Lim_sub[of "\x. f x - g x" 0 net f l] by auto - thus "?thesis" using Lim_neg [of "\ x. - g x" "-l" net] by auto -qed - -lemma Lim_transform_eventually: - "eventually (\x. f x = g x) net \ (f ---> l) net ==> (g ---> l) net" - apply (rule topological_tendstoI) - apply (drule (2) topological_tendstoD) - apply (erule (1) eventually_elim2, simp) - done - -lemma Lim_transform_within: - fixes l :: "'b::metric_space" (* TODO: generalize *) - assumes "0 < d" "(\x'\S. 0 < dist x' x \ dist x' x < d \ f x' = g x')" - "(f ---> l) (at x within S)" - shows "(g ---> l) (at x within S)" - using assms(1,3) unfolding Lim_within - apply - - apply (clarify, rename_tac e) - apply (drule_tac x=e in spec, clarsimp, rename_tac r) - apply (rule_tac x="min d r" in exI, clarsimp, rename_tac y) - apply (drule_tac x=y in bspec, assumption, clarsimp) - apply (simp add: assms(2)) - done - -lemma Lim_transform_at: - fixes l :: "'b::metric_space" (* TODO: generalize *) - shows "0 < d \ (\x'. 0 < dist x' x \ dist x' x < d \ f x' = g x') \ - (f ---> l) (at x) ==> (g ---> l) (at x)" - apply (subst within_UNIV[symmetric]) - using Lim_transform_within[of d UNIV x f g l] - by (auto simp add: within_UNIV) - -text{* Common case assuming being away from some crucial point like 0. *} - -lemma Lim_transform_away_within: - fixes a b :: "'a::metric_space" - fixes l :: "'b::metric_space" (* TODO: generalize *) - assumes "a\b" "\x\ S. x \ a \ x \ b \ f x = g x" - and "(f ---> l) (at a within S)" - shows "(g ---> l) (at a within S)" -proof- - have "\x'\S. 0 < dist x' a \ dist x' a < dist a b \ f x' = g x'" using assms(2) - apply auto apply(erule_tac x=x' in ballE) by (auto simp add: dist_commute) - thus ?thesis using Lim_transform_within[of "dist a b" S a f g l] using assms(1,3) unfolding dist_nz by auto -qed - -lemma Lim_transform_away_at: - fixes a b :: "'a::metric_space" - fixes l :: "'b::metric_space" (* TODO: generalize *) - assumes ab: "a\b" and fg: "\x. x \ a \ x \ b \ f x = g x" - and fl: "(f ---> l) (at a)" - shows "(g ---> l) (at a)" - using Lim_transform_away_within[OF ab, of UNIV f g l] fg fl - by (auto simp add: within_UNIV) - -text{* Alternatively, within an open set. *} - -lemma Lim_transform_within_open: - fixes a :: "'a::metric_space" - fixes l :: "'b::metric_space" (* TODO: generalize *) - assumes "open S" "a \ S" "\x\S. x \ a \ f x = g x" "(f ---> l) (at a)" - shows "(g ---> l) (at a)" -proof- - from assms(1,2) obtain e::real where "e>0" and e:"ball a e \ S" unfolding open_contains_ball by auto - hence "\x'. 0 < dist x' a \ dist x' a < e \ f x' = g x'" using assms(3) - unfolding ball_def subset_eq apply auto apply(erule_tac x=x' in allE) apply(erule_tac x=x' in ballE) by(auto simp add: dist_commute) - thus ?thesis using Lim_transform_at[of e a f g l] `e>0` assms(4) by auto -qed - -text{* A congruence rule allowing us to transform limits assuming not at point. *} - -(* FIXME: Only one congruence rule for tendsto can be used at a time! *) - -lemma Lim_cong_within[cong add]: - fixes a :: "'a::metric_space" - fixes l :: "'b::metric_space" (* TODO: generalize *) - shows "(\x. x \ a \ f x = g x) ==> ((\x. f x) ---> l) (at a within S) \ ((g ---> l) (at a within S))" - by (simp add: Lim_within dist_nz[symmetric]) - -lemma Lim_cong_at[cong add]: - fixes a :: "'a::metric_space" - fixes l :: "'b::metric_space" (* TODO: generalize *) - shows "(\x. x \ a ==> f x = g x) ==> (((\x. f x) ---> l) (at a) \ ((g ---> l) (at a)))" - by (simp add: Lim_at dist_nz[symmetric]) - -text{* Useful lemmas on closure and set of possible sequential limits.*} - -lemma closure_sequential: - fixes l :: "'a::metric_space" (* TODO: generalize *) - shows "l \ closure S \ (\x. (\n. x n \ S) \ (x ---> l) sequentially)" (is "?lhs = ?rhs") -proof - assume "?lhs" moreover - { assume "l \ S" - hence "?rhs" using Lim_const[of l sequentially] by auto - } moreover - { assume "l islimpt S" - hence "?rhs" unfolding islimpt_sequential by auto - } ultimately - show "?rhs" unfolding closure_def by auto -next - assume "?rhs" - thus "?lhs" unfolding closure_def unfolding islimpt_sequential by auto -qed - -lemma closed_sequential_limits: - fixes S :: "'a::metric_space set" - shows "closed S \ (\x l. (\n. x n \ S) \ (x ---> l) sequentially \ l \ S)" - unfolding closed_limpt - using closure_sequential [where 'a='a] closure_closed [where 'a='a] closed_limpt [where 'a='a] islimpt_sequential [where 'a='a] mem_delete [where 'a='a] - by metis - -lemma closure_approachable: - fixes S :: "'a::metric_space set" - shows "x \ closure S \ (\e>0. \y\S. dist y x < e)" - apply (auto simp add: closure_def islimpt_approachable) - by (metis dist_self) - -lemma closed_approachable: - fixes S :: "'a::metric_space set" - shows "closed S ==> (\e>0. \y\S. dist y x < e) \ x \ S" - by (metis closure_closed closure_approachable) - -text{* Some other lemmas about sequences. *} - -lemma seq_offset: - fixes l :: "'a::metric_space" (* TODO: generalize *) - shows "(f ---> l) sequentially ==> ((\i. f( i + k)) ---> l) sequentially" - apply (auto simp add: Lim_sequentially) - by (metis trans_le_add1 ) - -lemma seq_offset_neg: - "(f ---> l) sequentially ==> ((\i. f(i - k)) ---> l) sequentially" - apply (rule topological_tendstoI) - apply (drule (2) topological_tendstoD) - apply (simp only: eventually_sequentially) - apply (subgoal_tac "\N k (n::nat). N + k <= n ==> N <= n - k") - apply metis - by arith - -lemma seq_offset_rev: - "((\i. f(i + k)) ---> l) sequentially ==> (f ---> l) sequentially" - apply (rule topological_tendstoI) - apply (drule (2) topological_tendstoD) - apply (simp only: eventually_sequentially) - apply (subgoal_tac "\N k (n::nat). N + k <= n ==> N <= n - k \ (n - k) + k = n") - by metis arith - -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 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 by simp -qed - -text{* More properties of closed balls. *} - -lemma closed_cball: "closed (cball x e)" -unfolding cball_def closed_def -unfolding Collect_neg_eq [symmetric] not_le -apply (clarsimp simp add: open_dist, rename_tac y) -apply (rule_tac x="dist x y - e" in exI, clarsimp) -apply (rename_tac x') -apply (cut_tac x=x and y=x' and z=y in dist_triangle) -apply simp -done - -lemma open_contains_cball: "open S \ (\x\S. \e>0. cball x e \ S)" -proof- - { fix x and e::real assume "x\S" "e>0" "ball x e \ S" - hence "\d>0. cball x d \ S" unfolding subset_eq by (rule_tac x="e/2" in exI, auto) - } moreover - { fix x and e::real assume "x\S" "e>0" "cball x e \ S" - hence "\d>0. ball x d \ S" unfolding subset_eq apply(rule_tac x="e/2" in exI) by auto - } ultimately - show ?thesis unfolding open_contains_ball by auto -qed - -lemma open_contains_cball_eq: "open S ==> (\x. x \ S \ (\e>0. cball x e \ S))" - by (metis open_contains_cball subset_eq order_less_imp_le centre_in_cball mem_def) - -lemma mem_interior_cball: "x \ interior S \ (\e>0. cball x e \ S)" - apply (simp add: interior_def, safe) - apply (force simp add: open_contains_cball) - apply (rule_tac x="ball x e" in exI) - apply (simp add: open_ball centre_in_ball subset_trans [OF ball_subset_cball]) - done - -lemma islimpt_ball: - fixes x y :: "'a::{real_normed_vector,perfect_space}" - shows "y islimpt ball x e \ 0 < e \ y \ cball x e" (is "?lhs = ?rhs") -proof - assume "?lhs" - { assume "e \ 0" - hence *:"ball x e = {}" using ball_eq_empty[of x e] by auto - have False using `?lhs` unfolding * using islimpt_EMPTY[of y] by auto - } - hence "e > 0" by (metis not_less) - moreover - have "y \ cball x e" using closed_cball[of x e] islimpt_subset[of y "ball x e" "cball x e"] ball_subset_cball[of x e] `?lhs` unfolding closed_limpt by auto - ultimately show "?rhs" by auto -next - assume "?rhs" hence "e>0" by auto - { fix d::real assume "d>0" - have "\x'\ball x e. x' \ y \ dist x' y < d" - proof(cases "d \ dist x y") - case True thus "\x'\ball x e. x' \ y \ dist x' y < d" - proof(cases "x=y") - case True hence False using `d \ dist x y` `d>0` by auto - thus "\x'\ball x e. x' \ y \ dist x' y < d" by auto - next - case False - - have "dist x (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) - = norm (x - y + (d / (2 * norm (y - x))) *\<^sub>R (y - x))" - unfolding mem_cball mem_ball dist_norm diff_diff_eq2 diff_add_eq[THEN sym] by auto - 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) - 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 - also have "\ \ e - d/2" using `d \ dist x y` and `d>0` and `?rhs` by(auto simp add: dist_norm) - finally have "y - (d / (2 * dist y x)) *\<^sub>R (y - x) \ ball x e" using `d>0` by auto - - moreover - - 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 - 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 - qed - next - case False hence "d > dist x y" by auto - show "\x'\ball x e. x' \ y \ dist x' y < d" - proof(cases "x=y") - case True - obtain z where **: "z \ y" "dist z y < min e d" - using perfect_choose_dist[of "min e d" y] - using `d > 0` `e>0` by auto - show "\x'\ball x e. x' \ y \ dist x' y < d" - unfolding `x = y` - using `z \ y` ** - by (rule_tac x=z in bexI, auto simp add: dist_commute) - next - case False thus "\x'\ball x e. x' \ y \ dist x' y < d" - using `d>0` `d > dist x y` `?rhs` by(rule_tac x=x in bexI, auto) - qed - qed } - thus "?lhs" unfolding mem_cball islimpt_approachable mem_ball by auto -qed - -lemma closure_ball_lemma: - fixes x y :: "'a::real_normed_vector" - assumes "x \ y" shows "y islimpt ball x (dist x y)" -proof (rule islimptI) - fix T assume "y \ T" "open T" - then obtain r where "0 < r" "\z. dist z y < r \ z \ T" - unfolding open_dist by fast - (* choose point between x and y, within distance r of y. *) - def k \ "min 1 (r / (2 * dist x y))" - def z \ "y + scaleR k (x - y)" - have z_def2: "z = x + scaleR (1 - k) (y - x)" - 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 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_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) - apply (simp add: k_def divide_pos_pos zero_less_dist_iff `0 < r` `x \ y`) - apply (simp add: zero_less_dist_iff `x \ y`) - done - hence "z \ ball x (dist x y)" by simp - have "z \ y" - unfolding z_def k_def using `x \ y` `0 < r` - by (simp add: min_def) - show "\z\ball x (dist x y). z \ T \ z \ y" - using `z \ ball x (dist x y)` `z \ T` `z \ y` - by fast -qed - -lemma closure_ball: - fixes x :: "'a::real_normed_vector" - shows "0 < e \ closure (ball x e) = cball x e" -apply (rule equalityI) -apply (rule closure_minimal) -apply (rule ball_subset_cball) -apply (rule closed_cball) -apply (rule subsetI, rename_tac y) -apply (simp add: le_less [where 'a=real]) -apply (erule disjE) -apply (rule subsetD [OF closure_subset], simp) -apply (simp add: closure_def) -apply clarify -apply (rule closure_ball_lemma) -apply (simp add: zero_less_dist_iff) -done - -(* In a trivial vector space, this fails for e = 0. *) -lemma interior_cball: - 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 - { fix y assume "y \ cball x e" - hence False unfolding mem_cball using dist_nz[of x y] cs by auto } - hence "cball x e = {}" by auto - hence "interior (cball x e) = {}" using interior_empty by auto - ultimately show ?thesis by blast -next - case True note cs = this - have "ball x e \ cball x e" using ball_subset_cball by auto moreover - { 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_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") - case True - hence "e>0" using xa_y[unfolded dist_nz] xa_cball[unfolded mem_cball] by (auto simp add: dist_commute) - thus "y \ ball x e" using `x = y ` by simp - next - case False - 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) *\<^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) *\<^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 - qed } - hence "\S \ cball x e. open S \ S \ ball x e" by auto - ultimately show ?thesis using interior_unique[of "ball x e" "cball x e"] using open_ball[of x e] by auto -qed - -lemma frontier_ball: - 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 :: "'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) - by arith - -lemma cball_eq_empty: "(cball x e = {}) \ e < 0" - apply (simp add: expand_set_eq not_le) - by (metis zero_le_dist dist_self order_less_le_trans) -lemma cball_empty: "e < 0 ==> cball x e = {}" by (simp add: cball_eq_empty) - -lemma cball_eq_sing: - fixes x :: "'a::perfect_space" - shows "(cball x e = {x}) \ e = 0" -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 :: "'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. *} - -lemma eventually_within_interior: - assumes "x \ interior S" - shows "eventually P (at x within S) \ eventually P (at x)" (is "?lhs = ?rhs") -proof- - from assms obtain T where T: "open T" "x \ T" "T \ S" - unfolding interior_def by fast - { assume "?lhs" - then obtain A where "open A" "x \ A" "\y\A. y \ x \ y \ S \ P y" - unfolding Limits.eventually_within Limits.eventually_at_topological - by auto - with T have "open (A \ T)" "x \ A \ T" "\y\(A \ T). y \ x \ P y" - by auto - then have "?rhs" - unfolding Limits.eventually_at_topological by auto - } moreover - { assume "?rhs" hence "?lhs" - unfolding Limits.eventually_within - by (auto elim: eventually_elim1) - } ultimately - show "?thesis" .. -qed - -lemma lim_within_interior: - "x \ interior S \ (f ---> l) (at x within S) \ (f ---> l) (at x)" - unfolding tendsto_def by (simp add: eventually_within_interior) - -lemma netlimit_within_interior: - fixes x :: "'a::{perfect_space, real_normed_vector}" - (* FIXME: generalize to perfect_space *) - assumes "x \ interior S" - shows "netlimit(at x within S) = x" (is "?lhs = ?rhs") -proof- - from assms obtain e::real where e:"e>0" "ball x e \ S" using open_interior[of S] unfolding open_contains_ball using interior_subset[of S] by auto - hence "\ trivial_limit (at x within S)" using islimpt_subset[of x "ball x e" S] unfolding trivial_limit_within islimpt_ball centre_in_cball by auto - thus ?thesis using netlimit_within by auto -qed - -subsection{* Boundedness. *} - - (* FIXME: This has to be unified with BSEQ!! *) -definition - bounded :: "'a::metric_space set \ bool" where - "bounded S \ (\x e. \y\S. dist x y \ e)" - -lemma bounded_any_center: "bounded S \ (\e. \y\S. dist a y \ e)" -unfolding bounded_def -apply safe -apply (rule_tac x="dist a x + e" in exI, clarify) -apply (drule (1) bspec) -apply (erule order_trans [OF dist_triangle add_left_mono]) -apply auto -done - -lemma bounded_iff: "bounded S \ (\a. \x\S. norm x \ a)" -unfolding bounded_any_center [where a=0] -by (simp add: dist_norm) - -lemma bounded_empty[simp]: "bounded {}" by (simp add: bounded_def) -lemma bounded_subset: "bounded T \ S \ T ==> bounded S" - by (metis bounded_def subset_eq) - -lemma bounded_interior[intro]: "bounded S ==> bounded(interior S)" - by (metis bounded_subset interior_subset) - -lemma bounded_closure[intro]: assumes "bounded S" shows "bounded(closure S)" -proof- - from assms obtain x and a where a: "\y\S. dist x y \ a" unfolding bounded_def by auto - { fix y assume "y \ closure S" - then obtain f where f: "\n. f n \ S" "(f ---> y) sequentially" - unfolding closure_sequential by auto - have "\n. f n \ S \ dist x (f n) \ a" using a by simp - hence "eventually (\n. dist x (f n) \ a) sequentially" - by (rule eventually_mono, simp add: f(1)) - have "dist x y \ a" - apply (rule Lim_dist_ubound [of sequentially f]) - apply (rule trivial_limit_sequentially) - apply (rule f(2)) - apply fact - done - } - thus ?thesis unfolding bounded_def by auto -qed - -lemma bounded_cball[simp,intro]: "bounded (cball x e)" - apply (simp add: bounded_def) - apply (rule_tac x=x in exI) - apply (rule_tac x=e in exI) - apply auto - done - -lemma bounded_ball[simp,intro]: "bounded(ball x e)" - by (metis ball_subset_cball bounded_cball bounded_subset) - -lemma finite_imp_bounded[intro]: assumes "finite S" shows "bounded S" -proof- - { fix a F assume as:"bounded F" - then obtain x e where "\y\F. dist x y \ e" unfolding bounded_def by auto - hence "\y\(insert a F). dist x y \ max e (dist x a)" by auto - hence "bounded (insert a F)" unfolding bounded_def by (intro exI) - } - thus ?thesis using finite_induct[of S bounded] using bounded_empty assms by auto -qed - -lemma bounded_Un[simp]: "bounded (S \ T) \ bounded S \ bounded T" - apply (auto simp add: bounded_def) - apply (rename_tac x y r s) - apply (rule_tac x=x in exI) - apply (rule_tac x="max r (dist x y + s)" in exI) - apply (rule ballI, rename_tac z, safe) - apply (drule (1) bspec, simp) - apply (drule (1) bspec) - apply (rule min_max.le_supI2) - apply (erule order_trans [OF dist_triangle add_left_mono]) - done - -lemma bounded_Union[intro]: "finite F \ (\S\F. bounded S) \ bounded(\F)" - by (induct rule: finite_induct[of F], auto) - -lemma bounded_pos: "bounded S \ (\b>0. \x\ S. norm x <= b)" - apply (simp add: bounded_iff) - apply (subgoal_tac "\x (y::real). 0 < 1 + abs y \ (x <= y \ x <= 1 + abs y)") - by metis arith - -lemma bounded_Int[intro]: "bounded S \ bounded T \ bounded (S \ T)" - by (metis Int_lower1 Int_lower2 bounded_subset) - -lemma bounded_diff[intro]: "bounded S ==> bounded (S - T)" -apply (metis Diff_subset bounded_subset) -done - -lemma bounded_insert[intro]:"bounded(insert x S) \ bounded S" - by (metis Diff_cancel Un_empty_right Un_insert_right bounded_Un bounded_subset finite.emptyI finite_imp_bounded infinite_remove subset_insertI) - -lemma not_bounded_UNIV[simp, intro]: - "\ bounded (UNIV :: 'a::{real_normed_vector, perfect_space} set)" -proof(auto simp add: bounded_pos not_le) - obtain x :: 'a where "x \ 0" - using perfect_choose_dist [OF zero_less_one] by fast - 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_sgn) - then show "\x::'a. b < norm x" .. -qed - -lemma bounded_linear_image: - assumes "bounded S" "bounded_linear f" - shows "bounded(f ` S)" -proof- - from assms(1) obtain b where b:"b>0" "\x\S. norm x \ b" unfolding bounded_pos by auto - from assms(2) obtain B where B:"B>0" "\x. norm (f x) \ B * norm x" using bounded_linear.pos_bounded by (auto simp add: mult_ac) - { fix x assume "x\S" - hence "norm x \ b" using b by auto - hence "norm (f x) \ B * b" using B(2) apply(erule_tac x=x in allE) - by (metis B(1) B(2) real_le_trans real_mult_le_cancel_iff2) - } - thus ?thesis unfolding bounded_pos apply(rule_tac x="b*B" in exI) - using b B real_mult_order[of b B] by (auto simp add: real_mult_commute) -qed - -lemma bounded_scaling: - fixes S :: "'a::real_normed_vector set" - shows "bounded S \ bounded ((\x. c *\<^sub>R x) ` S)" - apply (rule bounded_linear_image, assumption) - apply (rule scaleR.bounded_linear_right) - done - -lemma bounded_translation: - fixes S :: "'a::real_normed_vector set" - assumes "bounded S" shows "bounded ((\x. a + x) ` S)" -proof- - from assms obtain b where b:"b>0" "\x\S. norm x \ b" unfolding bounded_pos by auto - { fix x assume "x\S" - hence "norm (a + x) \ b + norm a" using norm_triangle_ineq[of a x] b by auto - } - thus ?thesis unfolding bounded_pos using norm_ge_zero[of a] b(1) using add_strict_increasing[of b 0 "norm a"] - by (auto intro!: add exI[of _ "b + norm a"]) -qed - - -text{* Some theorems on sups and infs using the notion "bounded". *} - -lemma bounded_real: - fixes S :: "real set" - 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_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_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_real) - by (auto simp add: isLub_def isUb_def leastP_def setle_def setge_def) -qed - -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 -next - let ?S = "insert x S" - case False - hence *:"\x\S. x \ rsup S" using bounded_has_rsup(1)[of S] using assms by auto - hence "insert x S *<= max x (rsup S)" unfolding setle_def by auto - hence "isLub UNIV ?S (rsup ?S)" using rsup[of ?S] by auto - moreover - have **:"isUb UNIV ?S (max x (rsup S))" unfolding isUb_def setle_def using * by auto - { fix y assume as:"isUb UNIV (insert x S) y" - hence "max x (rsup S) \ y" unfolding isUb_def using rsup_le[OF `S\{}`] - unfolding setle_def by auto } - hence "max x (rsup S) <=* isUb UNIV (insert x S)" unfolding setge_def Ball_def mem_def by auto - hence "isLub UNIV ?S (max x (rsup S))" using ** isLubI2[of UNIV ?S "max x (rsup S)"] unfolding Collect_def by auto - ultimately show ?thesis using real_isLub_unique[of UNIV ?S] using `S\{}` by auto -qed - -lemma sup_insert_finite: "finite S \ rsup(insert x S) = (if S = {} then x else max x (rsup S))" - apply (rule rsup_insert) - apply (rule finite_imp_bounded) - by simp - -lemma bounded_has_rinf: - 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_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_real) - by (auto simp add: isGlb_def isLb_def greatestP_def setle_def setge_def) -qed - -(* TODO: Move this to RComplete.thy -- would need to include Glb into RComplete *) -lemma real_isGlb_unique: "[| isGlb R S x; isGlb R S y |] ==> x = (y::real)" - apply (frule isGlb_isLb) - apply (frule_tac x = y in isGlb_isLb) - apply (blast intro!: order_antisym dest!: isGlb_le_isLb) - done - -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 -next - let ?S = "insert x S" - case False - hence *:"\x\S. x \ rinf S" using bounded_has_rinf(1)[of S] using assms by auto - hence "min x (rinf S) <=* insert x S" unfolding setge_def by auto - hence "isGlb UNIV ?S (rinf ?S)" using rinf[of ?S] by auto - moreover - have **:"isLb UNIV ?S (min x (rinf S))" unfolding isLb_def setge_def using * by auto - { fix y assume as:"isLb UNIV (insert x S) y" - hence "min x (rinf S) \ y" unfolding isLb_def using rinf_ge[OF `S\{}`] - unfolding setge_def by auto } - hence "isLb UNIV (insert x S) *<= min x (rinf S)" unfolding setle_def Ball_def mem_def by auto - hence "isGlb UNIV ?S (min x (rinf S))" using ** isGlbI2[of UNIV ?S "min x (rinf S)"] unfolding Collect_def by auto - ultimately show ?thesis using real_isGlb_unique[of UNIV ?S] using `S\{}` by auto -qed - -lemma inf_insert_finite: "finite S ==> rinf(insert x S) = (if S = {} then x else min x (rinf S))" - by (rule rinf_insert, rule finite_imp_bounded, simp) - -subsection{* Compactness (the definition is the one based on convegent subsequences). *} - -definition - compact :: "'a::metric_space set \ bool" where (* TODO: generalize *) - "compact S \ - (\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 - Heine-Borel property if every closed and bounded subset is compact. -*} - -class heine_borel = - assumes bounded_imp_convergent_subsequence: - "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: "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. subseq r \ ((f \ r) ---> l) sequentially" - using `l \ s` r l by blast -qed - -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 [unfolded subseq_def] by auto - ultimately show "Suc n \ r (Suc n)" by auto -qed - -lemma eventually_subseq: - assumes r: "subseq r" - shows "eventually P sequentially \ eventually (\n. P (r n)) sequentially" -unfolding eventually_sequentially -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 - apply (rule_tac x="nat_rec e f" in exI) - apply (rule conjI)+ -apply (rule def_nat_rec_0, simp) -apply (rule allI, rule def_nat_rec_Suc, simp) -apply (rule allI, rule impI, rule ext) -apply (erule conjE) -apply (induct_tac x) -apply (simp add: nat_rec_0) -apply (erule_tac x="n" in allE) -apply (simp) -done - -lemma convergent_bounded_increasing: fixes s ::"nat\real" - 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 - then obtain t where t:"isLub UNIV (range s) t" using reals_complete[of "range s" ] by auto - { fix e::real assume "e>0" and as:"\N. \n\N. \ \s n - t\ < e" - { fix n::nat - 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)[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 "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. subseq r \ ((s \ r) ---> l) sequentially" -proof- - 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 - -instance real :: heine_borel -proof - fix s :: "real set" and f :: "nat \ real" - assume s: "bounded s" and f: "\n. f n \ s" - then obtain b where b: "\n. abs (f n) \ b" - unfolding bounded_iff by auto - obtain l :: real and r :: "nat \ nat" where - r: "subseq r" and l: "((f \ r) ---> l) sequentially" - using compact_real_lemma [OF b] by auto - thus "\l r. subseq r \ ((f \ r) ---> l) sequentially" - by auto -qed - -lemma bounded_component: "bounded s \ bounded ((\x. x $ i) ` s)" -unfolding bounded_def -apply clarify -apply (rule_tac x="x $ i" in exI) -apply (rule_tac x="e" in exI) -apply clarify -apply (rule order_trans [OF dist_nth_le], simp) -done - -lemma compact_lemma: - fixes f :: "nat \ 'a::heine_borel ^ 'n::finite" - assumes "bounded s" and "\n. f n \ s" - shows "\d. - \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. subseq r \ - (\e>0. eventually (\n. \i\d. dist (f (r n) $ i) (l $ i) < e) sequentially)" - 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:"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:"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:"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_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) - } - ultimately show ?case by auto - qed -qed - -instance "^" :: (heine_borel, finite) heine_borel -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: "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" - { fix e::real assume "e>0" - hence "0 < e / (real_of_nat (card ?d))" - using zero_less_card_finite using divide_pos_pos[of e, of "real_of_nat (card ?d)"] by auto - with l have "eventually (\n. \i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))) sequentially" - by simp - moreover - { fix n assume n: "\i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))" - have "dist (f (r n)) l \ (\i\?d. dist (f (r n) $ i) (l $ i))" - unfolding dist_vector_def using zero_le_dist by (rule setL2_le_setsum) - also have "\ < (\i\?d. e / (real_of_nat (card ?d)))" - by (rule setsum_strict_mono) (simp_all add: n) - finally have "dist (f (r n)) l < e" by simp - } - ultimately have "eventually (\n. dist (f (r n)) l < e) sequentially" - by (rule eventually_elim1) - } - hence *:"((f \ r) ---> l) sequentially" unfolding o_def tendsto_iff by simp - 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. *} - -lemma cauchy_def: - "Cauchy s \ (\e>0. \N. \m n. m \ N \ n \ N --> dist(s m)(s n) < e)" -unfolding Cauchy_def by blast - -definition - complete :: "'a::metric_space set \ bool" where - "complete s \ (\f. (\n. f n \ s) \ Cauchy f - --> (\l \ s. (f ---> l) sequentially))" - -lemma cauchy: "Cauchy s \ (\e>0.\ N::nat. \n\N. dist(s n)(s N) < e)" (is "?lhs = ?rhs") -proof- - { assume ?rhs - { fix e::real - assume "e>0" - with `?rhs` obtain N where N:"\n\N. dist (s n) (s N) < e/2" - by (erule_tac x="e/2" in allE) auto - { fix n m - assume nm:"N \ m \ N \ n" - hence "dist (s m) (s n) < e" using N - using dist_triangle_half_l[of "s m" "s N" "e" "s n"] - by blast - } - hence "\N. \m n. N \ m \ N \ n \ dist (s m) (s n) < e" - by blast - } - hence ?lhs - unfolding cauchy_def - by blast - } - thus ?thesis - unfolding cauchy_def - using dist_triangle_half_l - by blast -qed - -lemma convergent_imp_cauchy: - "(s ---> l) sequentially ==> Cauchy s" -proof(simp only: cauchy_def, rule, rule) - fix e::real assume "e>0" "(s ---> l) sequentially" - then obtain N::nat where N:"\n\N. dist (s n) l < e/2" unfolding Lim_sequentially by(erule_tac x="e/2" in allE) auto - thus "\N. \m n. N \ m \ N \ n \ dist (s m) (s n) < e" using dist_triangle_half_l[of _ l e _] by (rule_tac x=N in exI) auto -qed - -lemma cauchy_imp_bounded: assumes "Cauchy s" shows "bounded {y. (\n::nat. y = s n)}" -proof- - from assms obtain N::nat where "\m n. N \ m \ N \ n \ dist (s m) (s n) < 1" unfolding cauchy_def apply(erule_tac x= 1 in allE) by auto - hence N:"\n. N \ n \ dist (s N) (s n) < 1" by auto - moreover - have "bounded (s ` {0..N})" using finite_imp_bounded[of "s ` {1..N}"] by auto - then obtain a where a:"\x\s ` {0..N}. dist (s N) x \ a" - unfolding bounded_any_center [where a="s N"] by auto - ultimately show "?thesis" - unfolding bounded_any_center [where a="s N"] - apply(rule_tac x="max a 1" in exI) apply auto - apply(erule_tac x=n in allE) apply(erule_tac x=n in ballE) by auto -qed - -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" "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 - from lr(3)[unfolded Lim_sequentially, THEN spec[where x="e/2"]] obtain M where M:"\n\M. dist ((f \ r) n) l < e/2" using `e>0` by auto - { fix n::nat assume n:"n \ max N M" - have "dist ((f \ r) n) l < e/2" using n M by auto - moreover have "r n \ N" using lr'[of n] n by auto - hence "dist (f n) ((f \ r) n) < e / 2" using N using n by auto - ultimately have "dist (f n) l < e" using dist_triangle_half_r[of "f (r n)" "f n" e l] by (auto simp add: dist_commute) } - hence "\N. \n\N. dist (f n) l < e" by blast } - hence "\l\s. (f ---> l) sequentially" using `l\s` unfolding Lim_sequentially by auto } - thus ?thesis unfolding complete_def by auto -qed - -instance heine_borel < complete_space -proof - fix f :: "nat \ 'a" assume "Cauchy f" - hence "bounded (range f)" unfolding image_def - using cauchy_imp_bounded [of f] by auto - hence "compact (closure (range f))" - using bounded_closed_imp_compact [of "closure (range f)"] by auto - hence "complete (closure (range f))" - using compact_imp_complete by auto - moreover have "\n. f n \ closure (range f)" - using closure_subset [of "range f"] by auto - ultimately have "\l\closure (range f). (f ---> l) sequentially" - using `Cauchy f` unfolding complete_def by auto - then show "convergent f" - unfolding convergent_def LIMSEQ_conv_tendsto [symmetric] by auto -qed - -lemma complete_univ: "complete (UNIV :: 'a::complete_space set)" -proof(simp add: complete_def, rule, rule) - fix f :: "nat \ 'a" assume "Cauchy f" - hence "convergent f" by (rule Cauchy_convergent) - hence "\l. f ----> l" unfolding convergent_def . - thus "\l. (f ---> l) sequentially" unfolding LIMSEQ_conv_tendsto . -qed - -lemma complete_imp_closed: assumes "complete s" shows "closed s" -proof - - { fix x assume "x islimpt s" - then obtain f where f: "\n. f n \ s - {x}" "(f ---> x) sequentially" - unfolding islimpt_sequential by auto - then obtain l where l: "l\s" "(f ---> l) sequentially" - using `complete s`[unfolded complete_def] using convergent_imp_cauchy[of f x] by auto - hence "x \ s" using Lim_unique[of sequentially f l x] trivial_limit_sequentially f(2) by auto - } - thus "closed s" unfolding closed_limpt by auto -qed - -lemma complete_eq_closed: - fixes s :: "'a::complete_space set" - shows "complete s \ closed s" (is "?lhs = ?rhs") -proof - assume ?lhs thus ?rhs by (rule complete_imp_closed) -next - assume ?rhs - { fix f assume as:"\n::nat. f n \ s" "Cauchy f" - then obtain l where "(f ---> l) sequentially" using complete_univ[unfolded complete_def, THEN spec[where x=f]] by auto - hence "\l\s. (f ---> l) sequentially" using `?rhs`[unfolded closed_sequential_limits, THEN spec[where x=f], THEN spec[where x=l]] using as(1) by auto } - thus ?lhs unfolding complete_def by auto -qed - -lemma convergent_eq_cauchy: - fixes s :: "nat \ 'a::complete_space" - shows "(\l. (s ---> l) sequentially) \ Cauchy s" (is "?lhs = ?rhs") -proof - assume ?lhs then obtain l where "(s ---> l) sequentially" by auto - thus ?rhs using convergent_imp_cauchy by auto -next - assume ?rhs thus ?lhs using complete_univ[unfolded complete_def, THEN spec[where x=s]] by auto -qed - -lemma convergent_imp_bounded: - fixes s :: "nat \ 'a::metric_space" - shows "(s ---> l) sequentially ==> bounded (s ` (UNIV::(nat set)))" - using convergent_imp_cauchy[of s] - using cauchy_imp_bounded[of s] - unfolding image_def - by auto - -subsection{* Total boundedness. *} - -fun helper_1::"('a::metric_space set) \ real \ nat \ 'a" where - "helper_1 s e n = (SOME y::'a. y \ s \ (\m (dist (helper_1 s e m) y < e)))" -declare helper_1.simps[simp del] - -lemma compact_imp_totally_bounded: - assumes "compact s" - shows "\e>0. \k. finite k \ k \ s \ s \ (\((\x. ball x e) ` k))" -proof(rule, rule, rule ccontr) - fix e::real assume "e>0" and assm:"\ (\k. finite k \ k \ s \ s \ \(\x. ball x e) ` k)" - def x \ "helper_1 s e" - { fix n - have "x n \ s \ (\m dist (x m) (x n) < e)" - proof(induct_tac rule:nat_less_induct) - fix n def Q \ "(\y. y \ s \ (\m dist (x m) y < e))" - assume as:"\m s \ (\ma dist (x ma) (x m) < e)" - have "\ s \ (\x\x ` {0..s" "z \ (\x\x ` {0.. 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:"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[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 - -subsection{* Heine-Borel theorem (following Burkill \& Burkill vol. 2) *} - -lemma heine_borel_lemma: fixes s::"'a::metric_space set" - assumes "compact s" "s \ (\ t)" "\b \ t. open b" - shows "\e>0. \x \ s. \b \ t. ball x e \ b" -proof(rule ccontr) - assume "\ (\e>0. \x\s. \b\t. ball x e \ b)" - hence cont:"\e>0. \x\s. \xa\t. \ (ball x e \ xa)" by auto - { fix n::nat - have "1 / real (n + 1) > 0" by auto - hence "\x. x\s \ (\xa\t. \ (ball x (inverse (real (n+1))) \ xa))" using cont unfolding Bex_def by auto } - hence "\n::nat. \x. x \ s \ (\xa\t. \ ball x (inverse (real (n + 1))) \ xa)" by auto - 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:"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 - then obtain e where "e>0" and e:"\z. dist z l < e \ z\b" - using assms(3)[THEN bspec[where x=b]] unfolding open_dist by auto - - then obtain N1 where N1:"\n\N1. dist ((f \ r) n) l < e / 2" - using lr[unfolded Lim_sequentially, THEN spec[where x="e/2"]] by auto - - 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 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 - using f[THEN spec[where x="r (N1 + N2)"]] using `b\t` by auto - have "\y\ball x (inverse (real (r (N1 + N2) + 1))). y\b" apply(rule ccontr) using x by auto - then obtain y where y:"y \ ball x (inverse (real (r (N1 + N2) + 1)))" "y \ b" by auto - - have "dist x l < e/2" using N1 unfolding x_def o_def by auto - hence "dist y l < e" using y N2' using dist_triangle[of y l x]by (auto simp add:dist_commute) - - thus False using e and `y\b` by auto -qed - -lemma compact_imp_heine_borel: "compact s ==> (\f. (\t \ f. open t) \ s \ (\ f) - \ (\f'. f' \ f \ finite f' \ s \ (\ f')))" -proof clarify - fix f assume "compact s" " \t\f. open t" "s \ \f" - then obtain e::real where "e>0" and "\x\s. \b\f. ball x e \ b" using heine_borel_lemma[of s f] by auto - hence "\x\s. \b. b\f \ ball x e \ b" by auto - hence "\bb. \x\s. bb x \f \ ball x e \ bb x" using bchoice[of s "\x b. b\f \ ball x e \ b"] by auto - then obtain bb where bb:"\x\s. (bb x) \ f \ ball x e \ (bb x)" by blast - - from `compact s` have "\ k. finite k \ k \ s \ s \ \(\x. ball x e) ` k" using compact_imp_totally_bounded[of s] `e>0` by auto - then obtain k where k:"finite k" "k \ s" "s \ \(\x. ball x e) ` k" by auto - - have "finite (bb ` k)" using k(1) by auto - moreover - { fix x assume "x\s" - hence "x\\(\x. ball x e) ` k" using k(3) unfolding subset_eq by auto - hence "\X\bb ` k. x \ X" using bb k(2) by blast - hence "x \ \(bb ` k)" using Union_iff[of x "bb ` k"] by auto - } - ultimately show "\f'\f. finite f' \ s \ \f'" using bb k(2) by (rule_tac x="bb ` k" in exI) auto -qed - -subsection{* Bolzano-Weierstrass property. *} - -lemma heine_borel_imp_bolzano_weierstrass: - assumes "\f. (\t \ f. open t) \ s \ (\ f) --> (\f'. f' \ f \ finite f' \ s \ (\ f'))" - "infinite t" "t \ s" - shows "\x \ s. x islimpt t" -proof(rule ccontr) - assume "\ (\x \ s. x islimpt t)" - then obtain f where f:"\x\s. x \ f x \ open (f x) \ (\y\t. y \ f x \ y = x)" unfolding islimpt_def - using bchoice[of s "\ x T. x \ T \ open T \ (\y\t. y \ T \ y = x)"] by auto - obtain g where g:"g\{t. \x. x \ s \ t = f x}" "finite g" "s \ \g" - using assms(1)[THEN spec[where x="{t. \x. x\s \ t = f x}"]] using f by auto - from g(1,3) have g':"\x\g. \xa \ s. x = f xa" by auto - { fix x y assume "x\t" "y\t" "f x = f y" - hence "x \ f x" "y \ f x \ y = x" using f[THEN bspec[where x=x]] and `t\s` by auto - hence "x = y" using `f x = f y` and f[THEN bspec[where x=y]] and `y\t` and `t\s` by auto } - hence "infinite (f ` t)" using assms(2) using finite_imageD[unfolded inj_on_def, of f t] by auto - moreover - { fix x assume "x\t" "f x \ g" - from g(3) assms(3) `x\t` obtain h where "h\g" and "x\h" by auto - then obtain y where "y\s" "h = f y" using g'[THEN bspec[where x=h]] by auto - hence "y = x" using f[THEN bspec[where x=y]] and `x\t` and `x\h`[unfolded `h = f y`] by auto - hence False using `f x \ g` `h\g` unfolding `h = f y` by auto } - hence "f ` t \ g" by auto - ultimately show False using g(2) using finite_subset by auto -qed - -subsection{* Complete the chain of compactness variants. *} - -primrec helper_2::"(real \ 'a::metric_space) \ nat \ 'a" where - "helper_2 beyond 0 = beyond 0" | - "helper_2 beyond (Suc n) = beyond (dist undefined (helper_2 beyond n) + 1 )" - -lemma bolzano_weierstrass_imp_bounded: fixes s::"'a::metric_space set" - assumes "\t. infinite t \ t \ s --> (\x \ s. x islimpt t)" - shows "bounded s" -proof(rule ccontr) - assume "\ bounded s" - then obtain beyond where "\a. beyond a \s \ \ dist undefined (beyond a) \ a" - unfolding bounded_any_center [where a=undefined] - apply simp using choice[of "\a x. x\s \ \ dist undefined x \ a"] by auto - hence beyond:"\a. beyond a \s" "\a. dist undefined (beyond a) > a" - unfolding linorder_not_le by auto - def x \ "helper_2 beyond" - - { fix m n ::nat assume "mn" - have "1 < dist (x m) (x n)" - proof(cases "mn` by auto - hence "1 < dist undefined (x m) - dist undefined (x n)" using *[of n m] by auto - thus ?thesis using dist_triangle2 [of undefined "x m" "x n"] by arith - qed } note ** = this - { fix a b assume "x a = x b" "a \ b" - hence False using **[of a b] by auto } - hence "inj x" unfolding inj_on_def by auto - moreover - { fix n::nat - have "x n \ s" - proof(cases "n = 0") - case True thus ?thesis unfolding x_def using beyond by auto - next - case False then obtain z where "n = Suc z" using not0_implies_Suc by auto - thus ?thesis unfolding x_def using beyond by auto - qed } - ultimately have "infinite (range x) \ range x \ s" unfolding x_def using range_inj_infinite[of "helper_2 beyond"] using beyond(1) by auto - - then obtain l where "l\s" and l:"l islimpt range x" using assms[THEN spec[where x="range x"]] by auto - then obtain y where "x y \ l" and y:"dist (x y) l < 1/2" unfolding islimpt_approachable apply(erule_tac x="1/2" in allE) by auto - then obtain z where "x z \ l" and z:"dist (x z) l < dist (x y) l" using l[unfolded islimpt_approachable, THEN spec[where x="dist (x y) l"]] - unfolding dist_nz by auto - show False using y and z and dist_triangle_half_l[of "x y" l 1 "x z"] and **[of y z] by auto -qed - -lemma sequence_infinite_lemma: - fixes l :: "'a::metric_space" (* TODO: generalize *) - assumes "\n::nat. (f n \ l)" "(f ---> l) sequentially" - shows "infinite {y. (\ n. y = f n)}" -proof(rule ccontr) - let ?A = "(\x. dist x l) ` {y. \n. y = f n}" - assume "\ infinite {y. \n. y = f n}" - hence **:"finite ?A" "?A \ {}" by auto - obtain k where k:"dist (f k) l = Min ?A" using Min_in[OF **] by auto - have "0 < Min ?A" using assms(1) unfolding dist_nz unfolding Min_gr_iff[OF **] by auto - then obtain N where "dist (f N) l < Min ?A" using assms(2)[unfolded Lim_sequentially, THEN spec[where x="Min ?A"]] by auto - moreover have "dist (f N) l \ ?A" by auto - ultimately show False using Min_le[OF **(1), of "dist (f N) l"] by auto -qed - -lemma sequence_unique_limpt: - fixes l :: "'a::metric_space" (* TODO: generalize *) - assumes "\n::nat. (f n \ l)" "(f ---> l) sequentially" "l' islimpt {y. (\n. y = f n)}" - shows "l' = l" -proof(rule ccontr) - def e \ "dist l' l" - assume "l' \ l" hence "e>0" unfolding dist_nz e_def by auto - then obtain N::nat where N:"\n\N. dist (f n) l < e / 2" - using assms(2)[unfolded Lim_sequentially, THEN spec[where x="e/2"]] by auto - def d \ "Min (insert (e/2) ((\n. if dist (f n) l' = 0 then e/2 else dist (f n) l') ` {0 .. N}))" - have "d>0" using `e>0` unfolding d_def e_def using zero_le_dist[of _ l', unfolded order_le_less] by auto - obtain k where k:"f k \ l'" "dist (f k) l' < d" using `d>0` and assms(3)[unfolded islimpt_approachable, THEN spec[where x="d"]] by auto - have "k\N" using k(1)[unfolded dist_nz] using k(2)[unfolded d_def] - by force - hence "dist l' l < e" using N[THEN spec[where x=k]] using k(2)[unfolded d_def] and dist_triangle_half_r[of "f k" l' e l] by auto - thus False unfolding e_def by auto -qed - -lemma bolzano_weierstrass_imp_closed: - fixes s :: "'a::metric_space set" (* TODO: can this be generalized? *) - assumes "\t. infinite t \ t \ s --> (\x \ s. x islimpt t)" - shows "closed s" -proof- - { fix x l assume as: "\n::nat. x n \ s" "(x ---> l) sequentially" - hence "l \ s" - proof(cases "\n. x n \ l") - case False thus "l\s" using as(1) by auto - next - case True note cas = this - with as(2) have "infinite {y. \n. y = x n}" using sequence_infinite_lemma[of x l] by auto - then obtain l' where "l'\s" "l' islimpt {y. \n. y = x n}" using assms[THEN spec[where x="{y. \n. y = x n}"]] as(1) by auto - thus "l\s" using sequence_unique_limpt[of x l l'] using as cas by auto - qed } - thus ?thesis unfolding closed_sequential_limits by fast -qed - -text{* Hence express everything as an equivalence. *} - -lemma compact_eq_heine_borel: - fixes s :: "'a::heine_borel set" - shows "compact s \ - (\f. (\t \ f. open t) \ s \ (\ f) - --> (\f'. f' \ f \ finite f' \ s \ (\ f')))" (is "?lhs = ?rhs") -proof - assume ?lhs thus ?rhs using compact_imp_heine_borel[of s] by blast -next - assume ?rhs - hence "\t. infinite t \ t \ s \ (\x\s. x islimpt t)" - by (blast intro: heine_borel_imp_bolzano_weierstrass[of s]) - thus ?lhs using bolzano_weierstrass_imp_bounded[of s] bolzano_weierstrass_imp_closed[of s] bounded_closed_imp_compact[of s] by blast -qed - -lemma compact_eq_bolzano_weierstrass: - fixes s :: "'a::heine_borel set" - shows "compact s \ (\t. infinite t \ t \ s --> (\x \ s. x islimpt t))" (is "?lhs = ?rhs") -proof - assume ?lhs thus ?rhs unfolding compact_eq_heine_borel using heine_borel_imp_bolzano_weierstrass[of s] by auto -next - assume ?rhs thus ?lhs using bolzano_weierstrass_imp_bounded bolzano_weierstrass_imp_closed bounded_closed_imp_compact by auto -qed - -lemma compact_eq_bounded_closed: - fixes s :: "'a::heine_borel set" - shows "compact s \ bounded s \ closed s" (is "?lhs = ?rhs") -proof - assume ?lhs thus ?rhs unfolding compact_eq_bolzano_weierstrass using bolzano_weierstrass_imp_bounded bolzano_weierstrass_imp_closed by auto -next - assume ?rhs thus ?lhs using bounded_closed_imp_compact by auto -qed - -lemma compact_imp_bounded: - fixes s :: "'a::metric_space set" - shows "compact s ==> bounded s" -proof - - assume "compact s" - hence "\f. (\t\f. open t) \ s \ \f \ (\f'\f. finite f' \ s \ \f')" - by (rule compact_imp_heine_borel) - hence "\t. infinite t \ t \ s \ (\x \ s. x islimpt t)" - using heine_borel_imp_bolzano_weierstrass[of s] by auto - thus "bounded s" - by (rule bolzano_weierstrass_imp_bounded) -qed - -lemma compact_imp_closed: - fixes s :: "'a::metric_space set" - shows "compact s ==> closed s" -proof - - assume "compact s" - hence "\f. (\t\f. open t) \ s \ \f \ (\f'\f. finite f' \ s \ \f')" - by (rule compact_imp_heine_borel) - hence "\t. infinite t \ t \ s \ (\x \ s. x islimpt t)" - using heine_borel_imp_bolzano_weierstrass[of s] by auto - thus "closed s" - by (rule bolzano_weierstrass_imp_closed) -qed - -text{* In particular, some common special cases. *} - -lemma compact_empty[simp]: - "compact {}" - unfolding compact_def - by simp - -(* TODO: can any of the next 3 lemmas be generalized to metric spaces? *) - - (* FIXME : Rename *) -lemma compact_union[intro]: - fixes s t :: "'a::heine_borel set" - shows "compact s \ compact t ==> compact (s \ t)" - unfolding compact_eq_bounded_closed - using bounded_Un[of s t] - using closed_Un[of s t] - by simp - -lemma compact_inter[intro]: - fixes s t :: "'a::heine_borel set" - shows "compact s \ compact t ==> compact (s \ t)" - unfolding compact_eq_bounded_closed - using bounded_Int[of s t] - using closed_Int[of s t] - by simp - -lemma compact_inter_closed[intro]: - fixes s t :: "'a::heine_borel set" - shows "compact s \ closed t ==> compact (s \ t)" - unfolding compact_eq_bounded_closed - using closed_Int[of s t] - using bounded_subset[of "s \ t" s] - by blast - -lemma closed_inter_compact[intro]: - fixes s t :: "'a::heine_borel set" - shows "closed s \ compact t ==> compact (s \ t)" -proof- - assume "closed s" "compact t" - moreover - have "s \ t = t \ s" by auto ultimately - show ?thesis - using compact_inter_closed[of t s] - by auto -qed - -lemma closed_sing [simp]: - fixes a :: "'a::metric_space" - shows "closed {a}" - apply (clarsimp simp add: closed_def open_dist) - apply (rule ccontr) - apply (drule_tac x="dist x a" in spec) - apply (simp add: dist_nz dist_commute) - done - -lemma finite_imp_closed: - fixes s :: "'a::metric_space set" - shows "finite s ==> closed s" -proof (induct set: finite) - case empty show "closed {}" by simp -next - case (insert x F) - hence "closed ({x} \ F)" by (simp only: closed_Un closed_sing) - thus "closed (insert x F)" by simp -qed - -lemma finite_imp_compact: - fixes s :: "'a::heine_borel set" - shows "finite s ==> compact s" - unfolding compact_eq_bounded_closed - using finite_imp_closed finite_imp_bounded - by blast - -lemma compact_sing [simp]: "compact {a}" - unfolding compact_def o_def subseq_def - by (auto simp add: tendsto_const) - -lemma compact_cball[simp]: - fixes x :: "'a::heine_borel" - shows "compact(cball x e)" - using compact_eq_bounded_closed bounded_cball closed_cball - by blast - -lemma compact_frontier_bounded[intro]: - fixes s :: "'a::heine_borel set" - shows "bounded s ==> compact(frontier s)" - unfolding frontier_def - using compact_eq_bounded_closed - by blast - -lemma compact_frontier[intro]: - fixes s :: "'a::heine_borel set" - shows "compact s ==> compact (frontier s)" - using compact_eq_bounded_closed compact_frontier_bounded - by blast - -lemma frontier_subset_compact: - fixes s :: "'a::heine_borel set" - shows "compact s ==> frontier s \ s" - using frontier_subset_closed compact_eq_bounded_closed - by blast - -lemma open_delete: - fixes s :: "'a::metric_space set" - shows "open s ==> open(s - {x})" - using open_Diff[of s "{x}"] closed_sing - by blast - -text{* Finite intersection property. I could make it an equivalence in fact. *} - -lemma compact_imp_fip: - fixes s :: "'a::heine_borel set" - assumes "compact s" "\t \ f. closed t" - "\f'. finite f' \ f' \ f --> (s \ (\ f') \ {})" - shows "s \ (\ f) \ {}" -proof - assume as:"s \ (\ f) = {}" - hence "s \ \op - UNIV ` f" by auto - moreover have "Ball (op - UNIV ` f) open" using open_Diff closed_Diff using assms(2) by auto - ultimately obtain f' where f':"f' \ op - UNIV ` f" "finite f'" "s \ \f'" using assms(1)[unfolded compact_eq_heine_borel, THEN spec[where x="(\t. UNIV - t) ` f"]] by auto - hence "finite (op - UNIV ` f') \ op - UNIV ` f' \ f" by(auto simp add: Diff_Diff_Int) - hence "s \ \op - UNIV ` f' \ {}" using assms(3)[THEN spec[where x="op - UNIV ` f'"]] by auto - thus False using f'(3) unfolding subset_eq and Union_iff by blast -qed - -subsection{* Bounded closed nest property (proof does not use Heine-Borel). *} - -lemma bounded_closed_nest: - assumes "\n. closed(s n)" "\n. (s n \ {})" - "(\m n. m \ n --> s n \ s m)" "bounded(s 0)" - shows "\a::'a::heine_borel. \n::nat. a \ s(n)" -proof- - 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" "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 - { fix e::real assume "e>0" - 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 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) - using assms(3) apply(erule_tac x=n in allE)apply( erule_tac x="r (max N n)" in allE) by auto - ultimately have "\y\s n. dist y l < e" by auto - } - hence "l \ s n" using closed_approachable[of "s n" l] assms(1) by blast - } - thus ?thesis by auto -qed - -text{* Decreasing case does not even need compactness, just completeness. *} - -lemma decreasing_closed_nest: - assumes "\n. closed(s n)" - "\n. (s n \ {})" - "\m n. m \ n --> s n \ s m" - "\e>0. \n. \x \ (s n). \ y \ (s n). dist x y < e" - shows "\a::'a::heine_borel. \n::nat. a \ s n" -proof- - have "\n. \ x. x\s n" using assms(2) by auto - hence "\t. \n. t n \ s n" using choice[of "\ n x. x \ s n"] by auto - then obtain t where t: "\n. t n \ s n" by auto - { fix e::real assume "e>0" - then obtain N where N:"\x\s N. \y\s N. dist x y < e" using assms(4) by auto - { fix m n ::nat assume "N \ m \ N \ n" - hence "t m \ s N" "t n \ s N" using assms(3) t unfolding subset_eq t by blast+ - hence "dist (t m) (t n) < e" using N by auto - } - hence "\N. \m n. N \ m \ N \ n \ dist (t m) (t n) < e" by auto - } - hence "Cauchy t" unfolding cauchy_def by auto - then obtain l where l:"(t ---> l) sequentially" using complete_univ unfolding complete_def by auto - { fix n::nat - { fix e::real assume "e>0" - then obtain N::nat where N:"\n\N. dist (t n) l < e" using l[unfolded Lim_sequentially] by auto - have "t (max n N) \ s n" using assms(3) unfolding subset_eq apply(erule_tac x=n in allE) apply (erule_tac x="max n N" in allE) using t by auto - hence "\y\s n. dist y l < e" apply(rule_tac x="t (max n N)" in bexI) using N by auto - } - hence "l \ s n" using closed_approachable[of "s n" l] assms(1) by auto - } - then show ?thesis by auto -qed - -text{* Strengthen it to the intersection actually being a singleton. *} - -lemma decreasing_closed_nest_sing: - assumes "\n. closed(s n)" - "\n. s n \ {}" - "\m n. m \ n --> s n \ s m" - "\e>0. \n. \x \ (s n). \ y\(s n). dist x y < e" - shows "\a::'a::heine_borel. \ {t. (\n::nat. t = s n)} = {a}" -proof- - obtain a where a:"\n. a \ s n" using decreasing_closed_nest[of s] using assms by auto - { fix b assume b:"b \ \{t. \n. t = s n}" - { fix e::real assume "e>0" - hence "dist a b < e" using assms(4 )using b using a by blast - } - hence "dist a b = 0" by (metis dist_eq_0_iff dist_nz real_less_def) - } - with a have "\{t. \n. t = s n} = {a}" by auto - thus ?thesis by auto -qed - -text{* Cauchy-type criteria for uniform convergence. *} - -lemma uniformly_convergent_eq_cauchy: fixes s::"nat \ 'b \ 'a::heine_borel" shows - "(\l. \e>0. \N. \n x. N \ n \ P x --> dist(s n x)(l x) < e) \ - (\e>0. \N. \m n x. N \ m \ N \ n \ P x --> dist (s m x) (s n x) < e)" (is "?lhs = ?rhs") -proof(rule) - assume ?lhs - then obtain l where l:"\e>0. \N. \n x. N \ n \ P x \ dist (s n x) (l x) < e" by auto - { fix e::real assume "e>0" - then obtain N::nat where N:"\n x. N \ n \ P x \ dist (s n x) (l x) < e / 2" using l[THEN spec[where x="e/2"]] by auto - { fix n m::nat and x::"'b" assume "N \ m \ N \ n \ P x" - hence "dist (s m x) (s n x) < e" - using N[THEN spec[where x=m], THEN spec[where x=x]] - using N[THEN spec[where x=n], THEN spec[where x=x]] - using dist_triangle_half_l[of "s m x" "l x" e "s n x"] by auto } - hence "\N. \m n x. N \ m \ N \ n \ P x --> dist (s m x) (s n x) < e" by auto } - thus ?rhs by auto -next - assume ?rhs - hence "\x. P x \ Cauchy (\n. s n x)" unfolding cauchy_def apply auto by (erule_tac x=e in allE)auto - then obtain l where l:"\x. P x \ ((\n. s n x) ---> l x) sequentially" unfolding convergent_eq_cauchy[THEN sym] - using choice[of "\x l. P x \ ((\n. s n x) ---> l) sequentially"] by auto - { fix e::real assume "e>0" - then obtain N where N:"\m n x. N \ m \ N \ n \ P x \ dist (s m x) (s n x) < e/2" - using `?rhs`[THEN spec[where x="e/2"]] by auto - { fix x assume "P x" - then obtain M where M:"\n\M. dist (s n x) (l x) < e/2" - using l[THEN spec[where x=x], unfolded Lim_sequentially] using `e>0` by(auto elim!: allE[where x="e/2"]) - fix n::nat assume "n\N" - hence "dist(s n x)(l x) < e" using `P x`and N[THEN spec[where x=n], THEN spec[where x="N+M"], THEN spec[where x=x]] - using M[THEN spec[where x="N+M"]] and dist_triangle_half_l[of "s n x" "s (N+M) x" e "l x"] by (auto simp add: dist_commute) } - hence "\N. \n x. N \ n \ P x \ dist(s n x)(l x) < e" by auto } - thus ?lhs by auto -qed - -lemma uniformly_cauchy_imp_uniformly_convergent: - fixes s :: "nat \ 'a \ 'b::heine_borel" - assumes "\e>0.\N. \m (n::nat) x. N \ m \ N \ n \ P x --> dist(s m x)(s n x) < e" - "\x. P x --> (\e>0. \N. \n. N \ n --> dist(s n x)(l x) < e)" - shows "\e>0. \N. \n x. N \ n \ P x --> dist(s n x)(l x) < e" -proof- - obtain l' where l:"\e>0. \N. \n x. N \ n \ P x \ dist (s n x) (l' x) < e" - using assms(1) unfolding uniformly_convergent_eq_cauchy[THEN sym] by auto - moreover - { fix x assume "P x" - hence "l x = l' x" using Lim_unique[OF trivial_limit_sequentially, of "\n. s n x" "l x" "l' x"] - using l and assms(2) unfolding Lim_sequentially by blast } - ultimately show ?thesis by auto -qed - -subsection{* Define continuity over a net to take in restrictions of the set. *} - -definition - continuous :: "'a::t2_space net \ ('a \ 'b::topological_space) \ bool" where - "continuous net f \ (f ---> f(netlimit net)) net" - -lemma continuous_trivial_limit: - "trivial_limit net ==> continuous net f" - unfolding continuous_def tendsto_def trivial_limit_eq by auto - -lemma continuous_within: "continuous (at x within s) f \ (f ---> f(x)) (at x within s)" - unfolding continuous_def - unfolding tendsto_def - using netlimit_within[of x s] - by (cases "trivial_limit (at x within s)") (auto simp add: trivial_limit_eventually) - -lemma continuous_at: "continuous (at x) f \ (f ---> f(x)) (at x)" - using continuous_within [of x UNIV f] by (simp add: within_UNIV) - -lemma continuous_at_within: - assumes "continuous (at x) f" shows "continuous (at x within s) f" - using assms unfolding continuous_at continuous_within - by (rule Lim_at_within) - -text{* Derive the epsilon-delta forms, which we often use as "definitions" *} - -lemma continuous_within_eps_delta: - "continuous (at x within s) f \ (\e>0. \d>0. \x'\ s. dist x' x < d --> dist (f x') (f x) < e)" - unfolding continuous_within and Lim_within - apply auto unfolding dist_nz[THEN sym] apply(auto elim!:allE) apply(rule_tac x=d in exI) by auto - -lemma continuous_at_eps_delta: "continuous (at x) f \ (\e>0. \d>0. - \x'. dist x' x < d --> dist(f x')(f x) < e)" - using continuous_within_eps_delta[of x UNIV f] - unfolding within_UNIV by blast - -text{* Versions in terms of open balls. *} - -lemma continuous_within_ball: - "continuous (at x within s) f \ (\e>0. \d>0. - f ` (ball x d \ s) \ ball (f x) e)" (is "?lhs = ?rhs") -proof - assume ?lhs - { fix e::real assume "e>0" - then obtain d where d: "d>0" "\xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" - using `?lhs`[unfolded continuous_within Lim_within] by auto - { fix y assume "y\f ` (ball x d \ s)" - hence "y \ ball (f x) e" using d(2) unfolding dist_nz[THEN sym] - apply (auto simp add: dist_commute mem_ball) apply(erule_tac x=xa in ballE) apply auto using `e>0` by auto - } - hence "\d>0. f ` (ball x d \ s) \ ball (f x) e" using `d>0` unfolding subset_eq ball_def by (auto simp add: dist_commute) } - thus ?rhs by auto -next - assume ?rhs thus ?lhs unfolding continuous_within Lim_within ball_def subset_eq - apply (auto simp add: dist_commute) apply(erule_tac x=e in allE) by auto -qed - -lemma continuous_at_ball: - "continuous (at x) f \ (\e>0. \d>0. f ` (ball x d) \ ball (f x) e)" (is "?lhs = ?rhs") -proof - assume ?lhs thus ?rhs unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball - apply auto apply(erule_tac x=e in allE) apply auto apply(rule_tac x=d in exI) apply auto apply(erule_tac x=xa in allE) apply (auto simp add: dist_commute dist_nz) - unfolding dist_nz[THEN sym] by auto -next - assume ?rhs thus ?lhs unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball - apply auto apply(erule_tac x=e in allE) apply auto apply(rule_tac x=d in exI) apply auto apply(erule_tac x="f xa" in allE) by (auto simp add: dist_commute dist_nz) -qed - -text{* For setwise continuity, just start from the epsilon-delta definitions. *} - -definition - continuous_on :: "'a::metric_space set \ ('a \ 'b::metric_space) \ bool" where - "continuous_on s f \ (\x \ s. \e>0. \d::real>0. \x' \ s. dist x' x < d --> dist (f x') (f x) < e)" - - -definition - uniformly_continuous_on :: - "'a::metric_space set \ ('a \ 'b::metric_space) \ bool" where - "uniformly_continuous_on s f \ - (\e>0. \d>0. \x\s. \ x'\s. dist x' x < d - --> dist (f x') (f x) < e)" - -text{* Some simple consequential lemmas. *} - -lemma uniformly_continuous_imp_continuous: - " uniformly_continuous_on s f ==> continuous_on s f" - unfolding uniformly_continuous_on_def continuous_on_def by blast - -lemma continuous_at_imp_continuous_within: - "continuous (at x) f ==> continuous (at x within s) f" - unfolding continuous_within continuous_at using Lim_at_within by auto - -lemma continuous_at_imp_continuous_on: assumes "(\x \ s. continuous (at x) f)" - shows "continuous_on s f" -proof(simp add: continuous_at continuous_on_def, rule, rule, rule) - fix x and e::real assume "x\s" "e>0" - hence "eventually (\xa. dist (f xa) (f x) < e) (at x)" using assms unfolding continuous_at tendsto_iff by auto - then obtain d where d:"d>0" "\xa. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" unfolding eventually_at by auto - { fix x' assume "\ 0 < dist x' x" - hence "x=x'" - using dist_nz[of x' x] by auto - hence "dist (f x') (f x) < e" using `e>0` by auto - } - thus "\d>0. \x'\s. dist x' x < d \ dist (f x') (f x) < e" using d by auto -qed - -lemma continuous_on_eq_continuous_within: - "continuous_on s f \ (\x \ s. continuous (at x within s) f)" (is "?lhs = ?rhs") -proof - assume ?rhs - { fix x assume "x\s" - fix e::real assume "e>0" - assume "\d>0. \xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" - then obtain d where "d>0" and d:"\xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" by auto - { fix x' assume as:"x'\s" "dist x' x < d" - hence "dist (f x') (f x) < e" using `e>0` d `x'\s` dist_eq_0_iff[of x' x] zero_le_dist[of x' x] as(2) by (metis dist_eq_0_iff dist_nz) } - hence "\d>0. \x'\s. dist x' x < d \ dist (f x') (f x) < e" using `d>0` by auto - } - thus ?lhs using `?rhs` unfolding continuous_on_def continuous_within Lim_within by auto -next - assume ?lhs - thus ?rhs unfolding continuous_on_def continuous_within Lim_within by blast -qed - -lemma continuous_on: - "continuous_on s f \ (\x \ s. (f ---> f(x)) (at x within s))" - by (auto simp add: continuous_on_eq_continuous_within continuous_within) - -lemma continuous_on_eq_continuous_at: - "open s ==> (continuous_on s f \ (\x \ s. continuous (at x) f))" - by (auto simp add: continuous_on continuous_at Lim_within_open) - -lemma continuous_within_subset: - "continuous (at x within s) f \ t \ s - ==> continuous (at x within t) f" - unfolding continuous_within by(metis Lim_within_subset) - -lemma continuous_on_subset: - "continuous_on s f \ t \ s ==> continuous_on t f" - unfolding continuous_on by (metis subset_eq Lim_within_subset) - -lemma continuous_on_interior: - "continuous_on s f \ x \ interior s ==> continuous (at x) f" -unfolding interior_def -apply simp -by (meson continuous_on_eq_continuous_at continuous_on_subset) - -lemma continuous_on_eq: - "(\x \ s. f x = g x) \ continuous_on s f - ==> continuous_on s g" - by (simp add: continuous_on_def) - -text{* Characterization of various kinds of continuity in terms of sequences. *} - -(* \ could be generalized, but \ requires metric space *) -lemma continuous_within_sequentially: - fixes f :: "'a::metric_space \ 'b::metric_space" - shows "continuous (at a within s) f \ - (\x. (\n::nat. x n \ s) \ (x ---> a) sequentially - --> ((f o x) ---> f a) sequentially)" (is "?lhs = ?rhs") -proof - assume ?lhs - { fix x::"nat \ 'a" assume x:"\n. x n \ s" "\e>0. \N. \n\N. dist (x n) a < e" - fix e::real assume "e>0" - from `?lhs` obtain d where "d>0" and d:"\x\s. 0 < dist x a \ dist x a < d \ dist (f x) (f a) < e" unfolding continuous_within Lim_within using `e>0` by auto - from x(2) `d>0` obtain N where N:"\n\N. dist (x n) a < d" by auto - hence "\N. \n\N. dist ((f \ x) n) (f a) < e" - apply(rule_tac x=N in exI) using N d apply auto using x(1) - apply(erule_tac x=n in allE) apply(erule_tac x=n in allE) - apply(erule_tac x="x n" in ballE) apply auto unfolding dist_nz[THEN sym] apply auto using `e>0` by auto - } - thus ?rhs unfolding continuous_within unfolding Lim_sequentially by simp -next - assume ?rhs - { fix e::real assume "e>0" - assume "\ (\d>0. \x\s. 0 < dist x a \ dist x a < d \ dist (f x) (f a) < e)" - hence "\d. \x. d>0 \ x\s \ (0 < dist x a \ dist x a < d \ \ dist (f x) (f a) < e)" by blast - then obtain x where x:"\d>0. x d \ s \ (0 < dist (x d) a \ dist (x d) a < d \ \ dist (f (x d)) (f a) < e)" - using choice[of "\d x.0 x\s \ (0 < dist x a \ dist x a < d \ \ dist (f x) (f a) < e)"] by auto - { fix d::real assume "d>0" - hence "\N::nat. inverse (real (N + 1)) < d" using real_arch_inv[of d] by (auto, rule_tac x="n - 1" in exI)auto - then obtain N::nat where N:"inverse (real (N + 1)) < d" by auto - { fix n::nat assume n:"n\N" - hence "dist (x (inverse (real (n + 1)))) a < inverse (real (n + 1))" using x[THEN spec[where x="inverse (real (n + 1))"]] by auto - moreover have "inverse (real (n + 1)) < d" using N n 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) - ultimately have "dist (x (inverse (real (n + 1)))) a < d" by auto - } - hence "\N::nat. \n\N. dist (x (inverse (real (n + 1)))) a < d" by auto - } - hence "(\n::nat. x (inverse (real (n + 1))) \ s) \ (\e>0. \N::nat. \n\N. dist (x (inverse (real (n + 1)))) a < e)" using x by auto - hence "\e>0. \N::nat. \n\N. dist (f (x (inverse (real (n + 1))))) (f a) < e" using `?rhs`[THEN spec[where x="\n::nat. x (inverse (real (n+1)))"], unfolded Lim_sequentially] by auto - hence "False" apply(erule_tac x=e in allE) using `e>0` using x by auto - } - thus ?lhs unfolding continuous_within unfolding Lim_within unfolding Lim_sequentially by blast -qed - -lemma continuous_at_sequentially: - fixes f :: "'a::metric_space \ 'b::metric_space" - shows "continuous (at a) f \ (\x. (x ---> a) sequentially - --> ((f o x) ---> f a) sequentially)" - using continuous_within_sequentially[of a UNIV f] unfolding within_UNIV by auto - -lemma continuous_on_sequentially: - "continuous_on s f \ (\x. \a \ s. (\n. x(n) \ s) \ (x ---> a) sequentially - --> ((f o x) ---> f(a)) sequentially)" (is "?lhs = ?rhs") -proof - assume ?rhs thus ?lhs using continuous_within_sequentially[of _ s f] unfolding continuous_on_eq_continuous_within by auto -next - assume ?lhs thus ?rhs unfolding continuous_on_eq_continuous_within using continuous_within_sequentially[of _ s f] by auto -qed - -lemma uniformly_continuous_on_sequentially: - fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" - shows "uniformly_continuous_on s f \ (\x y. (\n. x n \ s) \ (\n. y n \ s) \ - ((\n. x n - y n) ---> 0) sequentially - \ ((\n. f(x n) - f(y n)) ---> 0) sequentially)" (is "?lhs = ?rhs") -proof - assume ?lhs - { fix x y assume x:"\n. x n \ s" and y:"\n. y n \ s" and xy:"((\n. x n - y n) ---> 0) sequentially" - { fix e::real assume "e>0" - then obtain d where "d>0" and d:"\x\s. \x'\s. dist x' x < d \ dist (f x') (f x) < e" - using `?lhs`[unfolded uniformly_continuous_on_def, THEN spec[where x=e]] by auto - obtain N where N:"\n\N. norm (x n - y n - 0) < d" using xy[unfolded Lim_sequentially dist_norm] and `d>0` by auto - { fix n assume "n\N" - hence "norm (f (x n) - f (y n) - 0) < e" - using N[THEN spec[where x=n]] using d[THEN bspec[where x="x n"], THEN bspec[where x="y n"]] using x and y - unfolding dist_commute and dist_norm by simp } - hence "\N. \n\N. norm (f (x n) - f (y n) - 0) < e" by auto } - hence "((\n. f(x n) - f(y n)) ---> 0) sequentially" unfolding Lim_sequentially and dist_norm by auto } - thus ?rhs by auto -next - assume ?rhs - { assume "\ ?lhs" - then obtain e where "e>0" "\d>0. \x\s. \x'\s. dist x' x < d \ \ dist (f x') (f x) < e" unfolding uniformly_continuous_on_def by auto - then obtain fa where fa:"\x. 0 < x \ fst (fa x) \ s \ snd (fa x) \ s \ dist (fst (fa x)) (snd (fa x)) < x \ \ dist (f (fst (fa x))) (f (snd (fa x))) < e" - using choice[of "\d x. d>0 \ fst x \ s \ snd x \ s \ dist (snd x) (fst x) < d \ \ dist (f (snd x)) (f (fst x)) < e"] unfolding Bex_def - by (auto simp add: dist_commute) - def x \ "\n::nat. fst (fa (inverse (real n + 1)))" - def y \ "\n::nat. snd (fa (inverse (real n + 1)))" - have xyn:"\n. x n \ s \ y n \ s" and xy0:"\n. dist (x n) (y n) < inverse (real n + 1)" and fxy:"\n. \ dist (f (x n)) (f (y n)) < e" - unfolding x_def and y_def using fa by auto - have 1:"\(x::'a) y. dist (x - y) 0 = dist x y" unfolding dist_norm by auto - have 2:"\(x::'b) y. dist (x - y) 0 = dist x y" unfolding dist_norm by auto - { fix e::real assume "e>0" - then obtain N::nat where "N \ 0" and N:"0 < inverse (real N) \ inverse (real N) < e" unfolding real_arch_inv[of e] by auto - { fix n::nat assume "n\N" - hence "inverse (real n + 1) < inverse (real N)" using real_of_nat_ge_zero and `N\0` by auto - also have "\ < e" using N by auto - finally have "inverse (real n + 1) < e" by auto - hence "dist (x n - y n) 0 < e" unfolding 1 using xy0[THEN spec[where x=n]] by auto } - hence "\N. \n\N. dist (x n - y n) 0 < e" by auto } - hence "\e>0. \N. \n\N. dist (f (x n) - f (y n)) 0 < e" using `?rhs`[THEN spec[where x=x], THEN spec[where x=y]] and xyn unfolding Lim_sequentially by auto - hence False unfolding 2 using fxy and `e>0` by auto } - thus ?lhs unfolding uniformly_continuous_on_def by blast -qed - -text{* The usual transformation theorems. *} - -lemma continuous_transform_within: - fixes f g :: "'a::metric_space \ 'b::metric_space" - assumes "0 < d" "x \ s" "\x' \ s. dist x' x < d --> f x' = g x'" - "continuous (at x within s) f" - shows "continuous (at x within s) g" -proof- - { fix e::real assume "e>0" - then obtain d' where d':"d'>0" "\xa\s. 0 < dist xa x \ dist xa x < d' \ dist (f xa) (f x) < e" using assms(4) unfolding continuous_within Lim_within by auto - { fix x' assume "x'\s" "0 < dist x' x" "dist x' x < (min d d')" - hence "dist (f x') (g x) < e" using assms(2,3) apply(erule_tac x=x in ballE) using d' by auto } - hence "\xa\s. 0 < dist xa x \ dist xa x < (min d d') \ dist (f xa) (g x) < e" by blast - hence "\d>0. \xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (g x) < e" using `d>0` `d'>0` by(rule_tac x="min d d'" in exI)auto } - hence "(f ---> g x) (at x within s)" unfolding Lim_within using assms(1) by auto - thus ?thesis unfolding continuous_within using Lim_transform_within[of d s x f g "g x"] using assms by blast -qed - -lemma continuous_transform_at: - fixes f g :: "'a::metric_space \ 'b::metric_space" - assumes "0 < d" "\x'. dist x' x < d --> f x' = g x'" - "continuous (at x) f" - shows "continuous (at x) g" -proof- - { fix e::real assume "e>0" - then obtain d' where d':"d'>0" "\xa. 0 < dist xa x \ dist xa x < d' \ dist (f xa) (f x) < e" using assms(3) unfolding continuous_at Lim_at by auto - { fix x' assume "0 < dist x' x" "dist x' x < (min d d')" - hence "dist (f x') (g x) < e" using assms(2) apply(erule_tac x=x in allE) using d' by auto - } - hence "\xa. 0 < dist xa x \ dist xa x < (min d d') \ dist (f xa) (g x) < e" by blast - hence "\d>0. \xa. 0 < dist xa x \ dist xa x < d \ dist (f xa) (g x) < e" using `d>0` `d'>0` by(rule_tac x="min d d'" in exI)auto - } - hence "(f ---> g x) (at x)" unfolding Lim_at using assms(1) by auto - thus ?thesis unfolding continuous_at using Lim_transform_at[of d x f g "g x"] using assms by blast -qed - -text{* Combination results for pointwise continuity. *} - -lemma continuous_const: "continuous net (\x. c)" - by (auto simp add: continuous_def Lim_const) - -lemma continuous_cmul: - fixes f :: "'a::t2_space \ 'b::real_normed_vector" - shows "continuous net f ==> continuous net (\x. c *\<^sub>R f x)" - by (auto simp add: continuous_def Lim_cmul) - -lemma continuous_neg: - fixes f :: "'a::t2_space \ 'b::real_normed_vector" - shows "continuous net f ==> continuous net (\x. -(f x))" - by (auto simp add: continuous_def Lim_neg) - -lemma continuous_add: - fixes f g :: "'a::t2_space \ 'b::real_normed_vector" - shows "continuous net f \ continuous net g \ continuous net (\x. f x + g x)" - by (auto simp add: continuous_def Lim_add) - -lemma continuous_sub: - fixes f g :: "'a::t2_space \ 'b::real_normed_vector" - shows "continuous net f \ continuous net g \ continuous net (\x. f x - g x)" - by (auto simp add: continuous_def Lim_sub) - -text{* Same thing for setwise continuity. *} - -lemma continuous_on_const: - "continuous_on s (\x. c)" - unfolding continuous_on_eq_continuous_within using continuous_const by blast - -lemma continuous_on_cmul: - fixes f :: "'a::metric_space \ 'b::real_normed_vector" - 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: - fixes f :: "'a::metric_space \ 'b::real_normed_vector" - shows "continuous_on s f \ continuous_on s (\x. - f x)" - unfolding continuous_on_eq_continuous_within using continuous_neg by blast - -lemma continuous_on_add: - fixes f g :: "'a::metric_space \ 'b::real_normed_vector" - shows "continuous_on s f \ continuous_on s g - \ continuous_on s (\x. f x + g x)" - unfolding continuous_on_eq_continuous_within using continuous_add by blast - -lemma continuous_on_sub: - fixes f g :: "'a::metric_space \ 'b::real_normed_vector" - shows "continuous_on s f \ continuous_on s g - \ continuous_on s (\x. f x - g x)" - unfolding continuous_on_eq_continuous_within using continuous_sub by blast - -text{* Same thing for uniform continuity, using sequential formulations. *} - -lemma uniformly_continuous_on_const: - "uniformly_continuous_on s (\x. c)" - unfolding uniformly_continuous_on_def by simp - -lemma uniformly_continuous_on_cmul: - fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" - (* FIXME: generalize 'a to metric_space *) - assumes "uniformly_continuous_on s f" - 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 *\<^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 scaleR_zero_right scaleR_right_diff_distrib by auto - } - thus ?thesis using assms unfolding uniformly_continuous_on_sequentially by auto -qed - -lemma dist_minus: - fixes x y :: "'a::real_normed_vector" - shows "dist (- x) (- y) = dist x y" - unfolding dist_norm minus_diff_minus norm_minus_cancel .. - -lemma uniformly_continuous_on_neg: - fixes f :: "'a::metric_space \ 'b::real_normed_vector" - shows "uniformly_continuous_on s f - ==> uniformly_continuous_on s (\x. -(f x))" - unfolding uniformly_continuous_on_def dist_minus . - -lemma uniformly_continuous_on_add: - fixes f g :: "'a::real_normed_vector \ 'b::real_normed_vector" (* FIXME: generalize 'a *) - assumes "uniformly_continuous_on s f" "uniformly_continuous_on s g" - shows "uniformly_continuous_on s (\x. f x + g x)" -proof- - { fix x y assume "((\n. f (x n) - f (y n)) ---> 0) sequentially" - "((\n. g (x n) - g (y n)) ---> 0) sequentially" - hence "((\xa. f (x xa) - f (y xa) + (g (x xa) - g (y xa))) ---> 0 + 0) sequentially" - using Lim_add[of "\ n. f (x n) - f (y n)" 0 sequentially "\ n. g (x n) - g (y n)" 0] by auto - hence "((\n. f (x n) + g (x n) - (f (y n) + g (y n))) ---> 0) sequentially" unfolding Lim_sequentially and add_diff_add [symmetric] by auto } - thus ?thesis using assms unfolding uniformly_continuous_on_sequentially by auto -qed - -lemma uniformly_continuous_on_sub: - fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" (* FIXME: generalize 'a *) - shows "uniformly_continuous_on s f \ uniformly_continuous_on s g - ==> uniformly_continuous_on s (\x. f x - g x)" - unfolding ab_diff_minus - using uniformly_continuous_on_add[of s f "\x. - g x"] - using uniformly_continuous_on_neg[of s g] by auto - -text{* Identity function is continuous in every sense. *} - -lemma continuous_within_id: - "continuous (at a within s) (\x. x)" - unfolding continuous_within by (rule Lim_at_within [OF Lim_ident_at]) - -lemma continuous_at_id: - "continuous (at a) (\x. x)" - unfolding continuous_at by (rule Lim_ident_at) - -lemma continuous_on_id: - "continuous_on s (\x. x)" - unfolding continuous_on Lim_within by auto - -lemma uniformly_continuous_on_id: - "uniformly_continuous_on s (\x. x)" - unfolding uniformly_continuous_on_def by auto - -text{* Continuity of all kinds is preserved under composition. *} - -lemma continuous_within_compose: - fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) - fixes g :: "'b::metric_space \ 'c::metric_space" - assumes "continuous (at x within s) f" "continuous (at (f x) within f ` s) g" - shows "continuous (at x within s) (g o f)" -proof- - { fix e::real assume "e>0" - with assms(2)[unfolded continuous_within Lim_within] obtain d where "d>0" and d:"\xa\f ` s. 0 < dist xa (f x) \ dist xa (f x) < d \ dist (g xa) (g (f x)) < e" by auto - from assms(1)[unfolded continuous_within Lim_within] obtain d' where "d'>0" and d':"\xa\s. 0 < dist xa x \ dist xa x < d' \ dist (f xa) (f x) < d" using `d>0` by auto - { fix y assume as:"y\s" "0 < dist y x" "dist y x < d'" - hence "dist (f y) (f x) < d" using d'[THEN bspec[where x=y]] by (auto simp add:dist_commute) - hence "dist (g (f y)) (g (f x)) < e" using as(1) d[THEN bspec[where x="f y"]] unfolding dist_nz[THEN sym] using `e>0` by auto } - hence "\d>0. \xa\s. 0 < dist xa x \ dist xa x < d \ dist (g (f xa)) (g (f x)) < e" using `d'>0` by auto } - thus ?thesis unfolding continuous_within Lim_within by auto -qed - -lemma continuous_at_compose: - fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) - fixes g :: "'b::metric_space \ 'c::metric_space" - assumes "continuous (at x) f" "continuous (at (f x)) g" - shows "continuous (at x) (g o f)" -proof- - have " continuous (at (f x) within range f) g" using assms(2) using continuous_within_subset[of "f x" UNIV g "range f", unfolded within_UNIV] by auto - thus ?thesis using assms(1) using continuous_within_compose[of x UNIV f g, unfolded within_UNIV] by auto -qed - -lemma continuous_on_compose: - "continuous_on s f \ continuous_on (f ` s) g \ continuous_on s (g o f)" - unfolding continuous_on_eq_continuous_within using continuous_within_compose[of _ s f g] by auto - -lemma uniformly_continuous_on_compose: - assumes "uniformly_continuous_on s f" "uniformly_continuous_on (f ` s) g" - shows "uniformly_continuous_on s (g o f)" -proof- - { fix e::real assume "e>0" - then obtain d where "d>0" and d:"\x\f ` s. \x'\f ` s. dist x' x < d \ dist (g x') (g x) < e" using assms(2) unfolding uniformly_continuous_on_def by auto - obtain d' where "d'>0" "\x\s. \x'\s. dist x' x < d' \ dist (f x') (f x) < d" using `d>0` using assms(1) unfolding uniformly_continuous_on_def by auto - hence "\d>0. \x\s. \x'\s. dist x' x < d \ dist ((g \ f) x') ((g \ f) x) < e" using `d>0` using d by auto } - thus ?thesis using assms unfolding uniformly_continuous_on_def by auto -qed - -text{* Continuity in terms of open preimages. *} - -lemma continuous_at_open: - fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) - shows "continuous (at x) f \ (\t. open t \ f x \ t --> (\s. open s \ x \ s \ (\x' \ s. (f x') \ t)))" (is "?lhs = ?rhs") -proof - assume ?lhs - { fix t assume as: "open t" "f x \ t" - then obtain e where "e>0" and e:"ball (f x) e \ t" unfolding open_contains_ball by auto - - obtain d where "d>0" and d:"\y. 0 < dist y x \ dist y x < d \ dist (f y) (f x) < e" using `e>0` using `?lhs`[unfolded continuous_at Lim_at open_dist] by auto - - have "open (ball x d)" using open_ball by auto - moreover have "x \ ball x d" unfolding centre_in_ball using `d>0` by simp - moreover - { fix x' assume "x'\ball x d" hence "f x' \ t" - using e[unfolded subset_eq Ball_def mem_ball, THEN spec[where x="f x'"]] d[THEN spec[where x=x']] - unfolding mem_ball apply (auto simp add: dist_commute) - unfolding dist_nz[THEN sym] using as(2) by auto } - hence "\x'\ball x d. f x' \ t" by auto - ultimately have "\s. open s \ x \ s \ (\x'\s. f x' \ t)" - apply(rule_tac x="ball x d" in exI) by simp } - thus ?rhs by auto -next - assume ?rhs - { fix e::real assume "e>0" - then obtain s where s: "open s" "x \ s" "\x'\s. f x' \ ball (f x) e" using `?rhs`[unfolded continuous_at Lim_at, THEN spec[where x="ball (f x) e"]] - unfolding centre_in_ball[of "f x" e, THEN sym] by auto - then obtain d where "d>0" and d:"ball x d \ s" unfolding open_contains_ball by auto - { fix y assume "0 < dist y x \ dist y x < d" - hence "dist (f y) (f x) < e" using d[unfolded subset_eq Ball_def mem_ball, THEN spec[where x=y]] - using s(3)[THEN bspec[where x=y], unfolded mem_ball] by (auto simp add: dist_commute) } - hence "\d>0. \xa. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" using `d>0` by auto } - thus ?lhs unfolding continuous_at Lim_at by auto -qed - -lemma continuous_on_open: - "continuous_on s f \ - (\t. openin (subtopology euclidean (f ` s)) t - --> openin (subtopology euclidean s) {x \ s. f x \ t})" (is "?lhs = ?rhs") -proof - assume ?lhs - { fix t assume as:"openin (subtopology euclidean (f ` s)) t" - have "{x \ s. f x \ t} \ s" using as[unfolded openin_euclidean_subtopology_iff] by auto - moreover - { fix x assume as':"x\{x \ s. f x \ t}" - then obtain e where e: "e>0" "\x'\f ` s. dist x' (f x) < e \ x' \ t" using as[unfolded openin_euclidean_subtopology_iff, THEN conjunct2, THEN bspec[where x="f x"]] by auto - from this(1) obtain d where d: "d>0" "\xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" using `?lhs`[unfolded continuous_on Lim_within, THEN bspec[where x=x]] using as' by auto - have "\e>0. \x'\s. dist x' x < e \ x' \ {x \ s. f x \ t}" using d e unfolding dist_nz[THEN sym] by (rule_tac x=d in exI, auto) } - ultimately have "openin (subtopology euclidean s) {x \ s. f x \ t}" unfolding openin_euclidean_subtopology_iff by auto } - thus ?rhs unfolding continuous_on Lim_within using openin by auto -next - assume ?rhs - { fix e::real and x assume "x\s" "e>0" - { fix xa x' assume "dist (f xa) (f x) < e" "xa \ s" "x' \ s" "dist (f xa) (f x') < e - dist (f xa) (f x)" - hence "dist (f x') (f x) < e" using dist_triangle[of "f x'" "f x" "f xa"] - by (auto simp add: dist_commute) } - hence "ball (f x) e \ f ` s \ f ` s \ (\xa\ball (f x) e \ f ` s. \ea>0. \x'\f ` s. dist x' xa < ea \ x' \ ball (f x) e \ f ` s)" apply auto - apply(rule_tac x="e - dist (f xa) (f x)" in exI) using `e>0` by (auto simp add: dist_commute) - hence "\xa\{xa \ s. f xa \ ball (f x) e \ f ` s}. \ea>0. \x'\s. dist x' xa < ea \ x' \ {xa \ s. f xa \ ball (f x) e \ f ` s}" - using `?rhs`[unfolded openin_euclidean_subtopology_iff, THEN spec[where x="ball (f x) e \ f ` s"]] by auto - hence "\d>0. \xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" apply(erule_tac x=x in ballE) apply auto using `e>0` `x\s` by (auto simp add: dist_commute) } - thus ?lhs unfolding continuous_on Lim_within by auto -qed - -(* ------------------------------------------------------------------------- *) -(* Similarly in terms of closed sets. *) -(* ------------------------------------------------------------------------- *) - -lemma continuous_on_closed: - "continuous_on s f \ (\t. closedin (subtopology euclidean (f ` s)) t --> closedin (subtopology euclidean s) {x \ s. f x \ t})" (is "?lhs = ?rhs") -proof - assume ?lhs - { fix t - have *:"s - {x \ s. f x \ f ` s - t} = {x \ s. f x \ t}" by auto - have **:"f ` s - (f ` s - (f ` s - t)) = f ` s - t" by auto - assume as:"closedin (subtopology euclidean (f ` s)) t" - hence "closedin (subtopology euclidean (f ` s)) (f ` s - (f ` s - t))" unfolding closedin_def topspace_euclidean_subtopology unfolding ** by auto - hence "closedin (subtopology euclidean s) {x \ s. f x \ t}" using `?lhs`[unfolded continuous_on_open, THEN spec[where x="(f ` s) - t"]] - unfolding openin_closedin_eq topspace_euclidean_subtopology unfolding * by auto } - thus ?rhs by auto -next - assume ?rhs - { fix t - have *:"s - {x \ s. f x \ f ` s - t} = {x \ s. f x \ t}" by auto - assume as:"openin (subtopology euclidean (f ` s)) t" - hence "openin (subtopology euclidean s) {x \ s. f x \ t}" using `?rhs`[THEN spec[where x="(f ` s) - t"]] - unfolding openin_closedin_eq topspace_euclidean_subtopology *[THEN sym] closedin_subtopology by auto } - thus ?lhs unfolding continuous_on_open by auto -qed - -text{* Half-global and completely global cases. *} - -lemma continuous_open_in_preimage: - assumes "continuous_on s f" "open t" - shows "openin (subtopology euclidean s) {x \ s. f x \ t}" -proof- - have *:"\x. x \ s \ f x \ t \ x \ s \ f x \ (t \ f ` s)" by auto - have "openin (subtopology euclidean (f ` s)) (t \ f ` s)" - using openin_open_Int[of t "f ` s", OF assms(2)] unfolding openin_open by auto - thus ?thesis using assms(1)[unfolded continuous_on_open, THEN spec[where x="t \ f ` s"]] using * by auto -qed - -lemma continuous_closed_in_preimage: - assumes "continuous_on s f" "closed t" - shows "closedin (subtopology euclidean s) {x \ s. f x \ t}" -proof- - have *:"\x. x \ s \ f x \ t \ x \ s \ f x \ (t \ f ` s)" by auto - have "closedin (subtopology euclidean (f ` s)) (t \ f ` s)" - using closedin_closed_Int[of t "f ` s", OF assms(2)] unfolding Int_commute by auto - thus ?thesis - using assms(1)[unfolded continuous_on_closed, THEN spec[where x="t \ f ` s"]] using * by auto -qed - -lemma continuous_open_preimage: - assumes "continuous_on s f" "open s" "open t" - shows "open {x \ s. f x \ t}" -proof- - obtain T where T: "open T" "{x \ s. f x \ t} = s \ T" - using continuous_open_in_preimage[OF assms(1,3)] unfolding openin_open by auto - thus ?thesis using open_Int[of s T, OF assms(2)] by auto -qed - -lemma continuous_closed_preimage: - assumes "continuous_on s f" "closed s" "closed t" - shows "closed {x \ s. f x \ t}" -proof- - obtain T where T: "closed T" "{x \ s. f x \ t} = s \ T" - using continuous_closed_in_preimage[OF assms(1,3)] unfolding closedin_closed by auto - thus ?thesis using closed_Int[of s T, OF assms(2)] by auto -qed - -lemma continuous_open_preimage_univ: - fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) - shows "\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 :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) - shows "(\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 - -lemma continuous_open_vimage: - fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) - shows "\x. continuous (at x) f \ open s \ open (f -` s)" - unfolding vimage_def by (rule continuous_open_preimage_univ) - -lemma continuous_closed_vimage: - fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) - shows "\x. continuous (at x) f \ closed s \ closed (f -` s)" - unfolding vimage_def by (rule continuous_closed_preimage_univ) - -text{* Equality of continuous functions on closure and related results. *} - -lemma continuous_closed_in_preimage_constant: - "continuous_on s f ==> closedin (subtopology euclidean s) {x \ s. f x = a}" - using continuous_closed_in_preimage[of s f "{a}"] closed_sing by auto - -lemma continuous_closed_preimage_constant: - "continuous_on s f \ closed s ==> closed {x \ s. f x = a}" - using continuous_closed_preimage[of s f "{a}"] closed_sing by auto - -lemma continuous_constant_on_closure: - assumes "continuous_on (closure s) f" - "\x \ s. f x = a" - shows "\x \ (closure s). f x = a" - using continuous_closed_preimage_constant[of "closure s" f a] - assms closure_minimal[of s "{x \ closure s. f x = a}"] closure_subset unfolding subset_eq by auto - -lemma image_closure_subset: - assumes "continuous_on (closure s) f" "closed t" "(f ` s) \ t" - shows "f ` (closure s) \ t" -proof- - have "s \ {x \ closure s. f x \ t}" using assms(3) closure_subset by auto - moreover have "closed {x \ closure s. f x \ t}" - using continuous_closed_preimage[OF assms(1)] and assms(2) by auto - ultimately have "closure s = {x \ closure s . f x \ t}" - using closure_minimal[of s "{x \ closure s. f x \ t}"] by auto - thus ?thesis by auto -qed - -lemma continuous_on_closure_norm_le: - fixes f :: "'a::metric_space \ 'b::real_normed_vector" - assumes "continuous_on (closure s) f" "\y \ s. norm(f y) \ b" "x \ (closure s)" - shows "norm(f x) \ b" -proof- - have *:"f ` s \ cball 0 b" using assms(2)[unfolded mem_cball_0[THEN sym]] by auto - show ?thesis - using image_closure_subset[OF assms(1) closed_cball[of 0 b] *] assms(3) - unfolding subset_eq apply(erule_tac x="f x" in ballE) by (auto simp add: dist_norm) -qed - -text{* Making a continuous function avoid some value in a neighbourhood. *} - -lemma continuous_within_avoid: - fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) - assumes "continuous (at x within s) f" "x \ s" "f x \ a" - shows "\e>0. \y \ s. dist x y < e --> f y \ a" -proof- - obtain d where "d>0" and d:"\xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < dist (f x) a" - using assms(1)[unfolded continuous_within Lim_within, THEN spec[where x="dist (f x) a"]] assms(3)[unfolded dist_nz] by auto - { fix y assume " y\s" "dist x y < d" - hence "f y \ a" using d[THEN bspec[where x=y]] assms(3)[unfolded dist_nz] - apply auto unfolding dist_nz[THEN sym] by (auto simp add: dist_commute) } - thus ?thesis using `d>0` by auto -qed - -lemma continuous_at_avoid: - fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) - assumes "continuous (at x) f" "f x \ a" - shows "\e>0. \y. dist x y < e \ f y \ a" -using assms using continuous_within_avoid[of x UNIV f a, unfolded within_UNIV] by auto - -lemma continuous_on_avoid: - assumes "continuous_on s f" "x \ s" "f x \ a" - shows "\e>0. \y \ s. dist x y < e \ f y \ a" -using assms(1)[unfolded continuous_on_eq_continuous_within, THEN bspec[where x=x], OF assms(2)] continuous_within_avoid[of x s f a] assms(2,3) by auto - -lemma continuous_on_open_avoid: - assumes "continuous_on s f" "open s" "x \ s" "f x \ a" - shows "\e>0. \y. dist x y < e \ f y \ a" -using assms(1)[unfolded continuous_on_eq_continuous_at[OF assms(2)], THEN bspec[where x=x], OF assms(3)] continuous_at_avoid[of x f a] assms(3,4) by auto - -text{* Proving a function is constant by proving open-ness of level set. *} - -lemma continuous_levelset_open_in_cases: - "connected s \ continuous_on s f \ - openin (subtopology euclidean s) {x \ s. f x = a} - ==> (\x \ s. f x \ a) \ (\x \ s. f x = a)" -unfolding connected_clopen using continuous_closed_in_preimage_constant by auto - -lemma continuous_levelset_open_in: - "connected s \ continuous_on s f \ - openin (subtopology euclidean s) {x \ s. f x = a} \ - (\x \ s. f x = a) ==> (\x \ s. f x = a)" -using continuous_levelset_open_in_cases[of s f ] -by meson - -lemma continuous_levelset_open: - assumes "connected s" "continuous_on s f" "open {x \ s. f x = a}" "\x \ s. f x = a" - shows "\x \ s. f x = a" -using continuous_levelset_open_in[OF assms(1,2), of a, unfolded openin_open] using assms (3,4) by auto - -text{* Some arithmetical combinations (more to prove). *} - -lemma open_scaling[intro]: - fixes s :: "'a::real_normed_vector set" - assumes "c \ 0" "open 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 *\<^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 *\<^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 - -lemma minus_image_eq_vimage: - fixes A :: "'a::ab_group_add set" - shows "(\x. - x) ` A = (\x. - x) -` A" - by (auto intro!: image_eqI [where f="\x. - x"]) - -lemma open_negations: - fixes s :: "'a::real_normed_vector set" - shows "open s ==> open ((\ x. -x) ` s)" - unfolding scaleR_minus1_left [symmetric] - by (rule open_scaling, auto) - -lemma open_translation: - 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 } - moreover have "{x. x - a \ s} = op + a ` s" apply auto unfolding image_iff apply(rule_tac x="x - a" in bexI) by auto - ultimately show ?thesis using continuous_open_preimage_univ[of "\x. x - a" s] using assms by auto -qed - -lemma open_affinity: - fixes s :: "'a::real_normed_vector set" - assumes "open s" "c \ 0" - shows "open ((\x. a + c *\<^sub>R x) ` s)" -proof- - 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: - fixes s :: "'a::real_normed_vector set" - shows "interior ((\x. a + x) ` s) = (\x. a + x) ` (interior s)" -proof (rule set_ext, rule) - fix x assume "x \ interior (op + a ` s)" - then obtain e where "e>0" and e:"ball x e \ op + a ` s" unfolding mem_interior by auto - hence "ball (x - a) e \ s" unfolding subset_eq Ball_def mem_ball dist_norm apply auto apply(erule_tac x="a + xa" in allE) unfolding ab_group_add_class.diff_diff_eq[THEN sym] by auto - thus "x \ op + a ` interior s" unfolding image_iff apply(rule_tac x="x - a" in bexI) unfolding mem_interior using `e > 0` by auto -next - fix x assume "x \ op + a ` interior s" - then obtain y e where "e>0" and e:"ball y e \ s" and y:"x = a + y" unfolding image_iff Bex_def mem_interior by auto - { fix z have *:"a + y - z = y + a - z" by auto - assume "z\ball x e" - hence "z - a \ s" using e[unfolded subset_eq, THEN bspec[where x="z - a"]] unfolding mem_ball dist_norm y ab_group_add_class.diff_diff_eq2 * by auto - hence "z \ op + a ` s" unfolding image_iff by(auto intro!: bexI[where x="z - a"]) } - hence "ball x e \ op + a ` s" unfolding subset_eq by auto - thus "x \ interior (op + a ` s)" unfolding mem_interior using `e>0` by auto -qed - -subsection {* Preservation of compactness and connectedness under continuous function. *} - -lemma compact_continuous_image: - assumes "continuous_on s f" "compact s" - shows "compact(f ` s)" -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:"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. subseq r \ ((x \ r) ---> l) sequentially" unfolding Lim_sequentially using r lr `l\s` by auto } - thus ?thesis unfolding compact_def by auto -qed - -lemma connected_continuous_image: - assumes "continuous_on s f" "connected s" - shows "connected(f ` s)" -proof- - { fix T assume as: "T \ {}" "T \ f ` s" "openin (subtopology euclidean (f ` s)) T" "closedin (subtopology euclidean (f ` s)) T" - have "{x \ s. f x \ T} = {} \ {x \ s. f x \ T} = s" - using assms(1)[unfolded continuous_on_open, THEN spec[where x=T]] - using assms(1)[unfolded continuous_on_closed, THEN spec[where x=T]] - using assms(2)[unfolded connected_clopen, THEN spec[where x="{x \ s. f x \ T}"]] as(3,4) by auto - hence False using as(1,2) - using as(4)[unfolded closedin_def topspace_euclidean_subtopology] by auto } - thus ?thesis unfolding connected_clopen by auto -qed - -text{* Continuity implies uniform continuity on a compact domain. *} - -lemma compact_uniformly_continuous: - assumes "continuous_on s f" "compact s" - shows "uniformly_continuous_on s f" -proof- - { fix x assume x:"x\s" - hence "\xa. \y. 0 < xa \ (y > 0 \ (\x'\s. dist x' x < y \ dist (f x') (f x) < xa))" using assms(1)[unfolded continuous_on_def, THEN bspec[where x=x]] by auto - hence "\fa. \xa>0. \x'\s. fa xa > 0 \ (dist x' x < fa xa \ dist (f x') (f x) < xa)" using choice[of "\e d. e>0 \ d>0 \(\x'\s. (dist x' x < d \ dist (f x') (f x) < e))"] by auto } - then have "\x\s. \y. \xa. 0 < xa \ (\x'\s. y xa > 0 \ (dist x' x < y xa \ dist (f x') (f x) < xa))" by auto - then obtain d where d:"\e>0. \x\s. \x'\s. d x e > 0 \ (dist x' x < d x e \ dist (f x') (f x) < e)" - using bchoice[of s "\x fa. \xa>0. \x'\s. fa xa > 0 \ (dist x' x < fa xa \ dist (f x') (f x) < xa)"] by blast - - { fix e::real assume "e>0" - - { fix x assume "x\s" hence "x \ ball x (d x (e / 2))" unfolding centre_in_ball using d[THEN spec[where x="e/2"]] using `e>0` by auto } - hence "s \ \{ball x (d x (e / 2)) |x. x \ s}" unfolding subset_eq by auto - moreover - { fix b assume "b\{ball x (d x (e / 2)) |x. x \ s}" hence "open b" by auto } - ultimately obtain ea where "ea>0" and ea:"\x\s. \b\{ball x (d x (e / 2)) |x. x \ s}. ball x ea \ b" using heine_borel_lemma[OF assms(2), of "{ball x (d x (e / 2)) | x. x\s }"] by auto - - { fix x y assume "x\s" "y\s" and as:"dist y x < ea" - obtain z where "z\s" and z:"ball x ea \ ball z (d z (e / 2))" using ea[THEN bspec[where x=x]] and `x\s` by auto - hence "x\ball z (d z (e / 2))" using `ea>0` unfolding subset_eq by auto - hence "dist (f z) (f x) < e / 2" using d[THEN spec[where x="e/2"]] and `e>0` and `x\s` and `z\s` - by (auto simp add: dist_commute) - moreover have "y\ball z (d z (e / 2))" using as and `ea>0` and z[unfolded subset_eq] - by (auto simp add: dist_commute) - hence "dist (f z) (f y) < e / 2" using d[THEN spec[where x="e/2"]] and `e>0` and `y\s` and `z\s` - by (auto simp add: dist_commute) - ultimately have "dist (f y) (f x) < e" using dist_triangle_half_r[of "f z" "f x" e "f y"] - by (auto simp add: dist_commute) } - then have "\d>0. \x\s. \x'\s. dist x' x < d \ dist (f x') (f x) < e" using `ea>0` by auto } - thus ?thesis unfolding uniformly_continuous_on_def by auto -qed - -text{* Continuity of inverse function on compact domain. *} - -lemma continuous_on_inverse: - 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- - have *:"g ` f ` s = s" using assms(3) by (auto simp add: image_iff) - { fix t assume t:"closedin (subtopology euclidean (g ` f ` s)) t" - then obtain T where T: "closed T" "t = s \ T" unfolding closedin_closed unfolding * by auto - have "continuous_on (s \ T) f" using continuous_on_subset[OF assms(1), of "s \ t"] - unfolding T(2) and Int_left_absorb by auto - moreover have "compact (s \ T)" - using assms(2) unfolding compact_eq_bounded_closed - using bounded_subset[of s "s \ T"] and T(1) by auto - ultimately have "closed (f ` t)" using T(1) unfolding T(2) - using compact_continuous_image [of "s \ T" f] unfolding compact_eq_bounded_closed by auto - moreover have "{x \ f ` s. g x \ t} = f ` s \ f ` t" using assms(3) unfolding T(2) by auto - ultimately have "closedin (subtopology euclidean (f ` s)) {x \ f ` s. g x \ t}" - unfolding closedin_closed by auto } - thus ?thesis unfolding continuous_on_closed by auto -qed - -subsection{* A uniformly convergent limit of continuous functions is continuous. *} - -lemma norm_triangle_lt: - fixes x y :: "'a::real_normed_vector" - shows "norm x + norm y < e \ norm (x + y) < e" -by (rule le_less_trans [OF norm_triangle_ineq]) - -lemma continuous_uniform_limit: - fixes f :: "'a \ 'b::metric_space \ 'c::real_normed_vector" - assumes "\ (trivial_limit net)" "eventually (\n. continuous_on s (f n)) net" - "\e>0. eventually (\n. \x \ s. norm(f n x - g x) < e) net" - shows "continuous_on s g" -proof- - { fix x and e::real assume "x\s" "e>0" - have "eventually (\n. \x\s. norm (f n x - g x) < e / 3) net" using `e>0` assms(3)[THEN spec[where x="e/3"]] by auto - then obtain n where n:"\xa\s. norm (f n xa - g xa) < e / 3" "continuous_on s (f n)" - using eventually_and[of "(\n. \x\s. norm (f n x - g x) < e / 3)" "(\n. continuous_on s (f n))" net] assms(1,2) eventually_happens by blast - have "e / 3 > 0" using `e>0` by auto - then obtain d where "d>0" and d:"\x'\s. dist x' x < d \ dist (f n x') (f n x) < e / 3" - using n(2)[unfolded continuous_on_def, THEN bspec[where x=x], OF `x\s`, THEN spec[where x="e/3"]] by blast - { fix y assume "y\s" "dist y x < d" - hence "dist (f n y) (f n x) < e / 3" using d[THEN bspec[where x=y]] by auto - hence "norm (f n y - g x) < 2 * e / 3" using norm_triangle_lt[of "f n y - f n x" "f n x - g x" "2*e/3"] - using n(1)[THEN bspec[where x=x], OF `x\s`] unfolding dist_norm unfolding ab_group_add_class.ab_diff_minus by auto - hence "dist (g y) (g x) < e" unfolding dist_norm using n(1)[THEN bspec[where x=y], OF `y\s`] - unfolding norm_minus_cancel[of "f n y - g y", THEN sym] using norm_triangle_lt[of "f n y - g x" "g y - f n y" e] by (auto simp add: uminus_add_conv_diff) } - hence "\d>0. \x'\s. dist x' x < d \ dist (g x') (g x) < e" using `d>0` by auto } - thus ?thesis unfolding continuous_on_def by auto -qed - -subsection{* Topological properties of linear functions. *} - -lemma linear_lim_0: - assumes "bounded_linear f" shows "(f ---> 0) (at (0))" -proof- - interpret f: bounded_linear f by fact - have "(f ---> f 0) (at 0)" - using tendsto_ident_at by (rule f.tendsto) - thus ?thesis unfolding f.zero . -qed - -lemma linear_continuous_at: - assumes "bounded_linear f" shows "continuous (at a) f" - unfolding continuous_at using assms - apply (rule bounded_linear.tendsto) - apply (rule tendsto_ident_at) - done - -lemma linear_continuous_within: - shows "bounded_linear f ==> continuous (at x within s) f" - using continuous_at_imp_continuous_within[of x f s] using linear_continuous_at[of f] by auto - -lemma linear_continuous_on: - shows "bounded_linear f ==> continuous_on s f" - using continuous_at_imp_continuous_on[of s f] using linear_continuous_at[of f] by auto - -text{* Also bilinear functions, in composition form. *} - -lemma bilinear_continuous_at_compose: - shows "continuous (at x) f \ continuous (at x) g \ bounded_bilinear h - ==> continuous (at x) (\x. h (f x) (g x))" - unfolding continuous_at using Lim_bilinear[of f "f x" "(at x)" g "g x" h] by auto - -lemma bilinear_continuous_within_compose: - shows "continuous (at x within s) f \ continuous (at x within s) g \ bounded_bilinear h - ==> continuous (at x within s) (\x. h (f x) (g x))" - unfolding continuous_within using Lim_bilinear[of f "f x"] by auto - -lemma bilinear_continuous_on_compose: - shows "continuous_on s f \ continuous_on s g \ bounded_bilinear h - ==> continuous_on s (\x. h (f x) (g x))" - unfolding continuous_on_eq_continuous_within apply auto apply(erule_tac x=x in ballE) apply auto apply(erule_tac x=x in ballE) apply auto - using bilinear_continuous_within_compose[of _ s f g h] by auto - -subsection{* Topological stuff lifted from and dropped to R *} - - -lemma open_real: - fixes s :: "real set" shows - "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 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 - 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_real_range: - fixes f :: "'a::real_normed_vector \ real" - 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)) - -text{* Hence some handy theorems on distance, diameter etc. of/from a set. *} - -lemma compact_attains_sup: - fixes s :: "real set" - assumes "compact s" "s \ {}" - shows "\x \ s. \y \ s. y \ x" -proof- - 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_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 s" "s \ {}" shows "\x \ s. \y \ s. x \ y" -proof- - 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 - moreover - { fix x assume "x \ s" - hence *:"abs (x - rinf s) = x - rinf s" using as(1)[THEN bspec[where x=x]] by auto - 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_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 f - ==> (\x \ s. \y \ s. f y \ f x)" - using compact_attains_sup[of "f ` s"] - 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 f - \ (\x \ s. \y \ s. f x \ f y)" - using compact_attains_inf[of "f ` s"] - using compact_continuous_image[of s f] by auto - -lemma distance_attains_sup: - assumes "compact s" "s \ {}" - shows "\x \ s. \y \ s. dist a y \ dist a x" -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 :: "'a::heine_borel" - assumes "closed s" "s \ {}" - shows "\x \ s. \y \ s. dist a x \ dist a y" -proof- - from assms(2) obtain b where "b\s" by auto - let ?B = "cball a (dist b a) \ s" - have "b \ ?B" using `b\s` by (simp add: dist_commute) - hence "?B \ {}" by auto - moreover - { fix x assume "x\?B" - fix e::real assume "e>0" - { 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 - -subsection{* We can now extend limit compositions to consider the scalar multiplier. *} - -lemma Lim_mul: - fixes f :: "'a \ 'b::real_normed_vector" - assumes "(c ---> d) net" "(f ---> l) net" - shows "((\x. c(x) *\<^sub>R f x) ---> (d *\<^sub>R l)) net" - using assms by (rule scaleR.tendsto) - -lemma Lim_vmul: - 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" and v :: "'b::real_normed_vector" - shows "continuous net c ==> continuous net (\x. c(x) *\<^sub>R v)" - unfolding continuous_def using Lim_vmul[of c] by auto - -lemma continuous_mul: - fixes c :: "'a::metric_space \ real" - fixes f :: "'a::metric_space \ 'b::real_normed_vector" - shows "continuous net c \ continuous net f - ==> continuous net (\x. c(x) *\<^sub>R f x) " - unfolding continuous_def by (intro tendsto_intros) - -lemma continuous_on_vmul: - fixes c :: "'a::metric_space \ real" and v :: "'b::real_normed_vector" - shows "continuous_on s c ==> continuous_on s (\x. c(x) *\<^sub>R v)" - unfolding continuous_on_eq_continuous_within using continuous_vmul[of _ c] by auto - -lemma continuous_on_mul: - fixes c :: "'a::metric_space \ real" - fixes f :: "'a::metric_space \ 'b::real_normed_vector" - shows "continuous_on s c \ continuous_on s f - ==> continuous_on s (\x. c(x) *\<^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 "(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 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 :: "'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 :: "'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. *} - -lemma bounded_pastecart: - fixes s :: "('a::real_normed_vector ^ _) set" (* FIXME: generalize to metric_space *) - assumes "bounded s" "bounded t" - shows "bounded { pastecart x y | x y . (x \ s \ y \ t)}" -proof- - obtain a b where ab:"\x\s. norm x \ a" "\x\t. norm x \ b" using assms[unfolded bounded_iff] by auto - { fix x y assume "x\s" "y\t" - hence "norm x \ a" "norm y \ b" using ab by auto - hence "norm (pastecart x y) \ a + b" using norm_pastecart[of x y] by auto } - 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" - shows "closed {pastecart x y | x y . x \ s \ y \ t}" -proof- - { fix x l assume as:"\n::nat. x n \ {pastecart x y |x y. x \ s \ y \ t}" "(x ---> l) sequentially" - { fix n::nat have "fstcart (x n) \ s" "sndcart (x n) \ t" using as(1)[THEN spec[where x=n]] by auto } note * = this - moreover - { fix e::real assume "e>0" - then obtain N::nat where N:"\n\N. dist (x n) l < e" using as(2)[unfolded Lim_sequentially, THEN spec[where x=e]] by auto - { fix n::nat assume "n\N" - hence "dist (fstcart (x n)) (fstcart l) < e" "dist (sndcart (x n)) (sndcart l) < e" - using N[THEN spec[where x=n]] dist_fstcart[of "x n" l] dist_sndcart[of "x n" l] by auto } - hence "\N. \n\N. dist (fstcart (x n)) (fstcart l) < e" "\N. \n\N. dist (sndcart (x n)) (sndcart l) < e" by auto } - ultimately have "fstcart l \ s" "sndcart l \ t" - using assms(1)[unfolded closed_sequential_limits, THEN spec[where x="\n. fstcart (x n)"], THEN spec[where x="fstcart l"]] - using assms(2)[unfolded closed_sequential_limits, THEN spec[where x="\n. sndcart (x n)"], THEN spec[where x="sndcart l"]] - unfolding Lim_sequentially by auto - hence "l \ {pastecart x y |x y. x \ s \ y \ t}" using pastecart_fst_snd[THEN sym, of l] by auto } - thus ?thesis unfolding closed_sequential_limits by auto -qed - -lemma compact_pastecart: - fixes s t :: "(real ^ _) set" - 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_scaling: - fixes s :: "'a::real_normed_vector set" - assumes "compact s" shows "compact ((\x. c *\<^sub>R x) ` s)" -proof- - let ?f = "\x. scaleR c x" - have *:"bounded_linear ?f" by (rule scaleR.bounded_linear_right) - show ?thesis using compact_continuous_image[of s ?f] continuous_at_imp_continuous_on[of s ?f] - using linear_continuous_at[OF *] assms by auto -qed - -lemma compact_negations: - fixes s :: "'a::real_normed_vector set" - assumes "compact s" shows "compact ((\x. -x) ` s)" - using compact_scaling [OF assms, of "- 1"] by auto - -lemma compact_sums: - 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. 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 :: "'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} = {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 :: "'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 - thus ?thesis using compact_sums[OF assms compact_sing[of a]] by auto -qed - -lemma compact_affinity: - fixes s :: "'a::real_normed_vector set" - assumes "compact s" shows "compact ((\x. a + c *\<^sub>R x) ` s)" -proof- - 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 :: "'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[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 - -text{* We can state this in terms of diameter of a set. *} - -definition "diameter s = (if s = {} then 0::real else rsup {norm(x - y) | x y. x \ s \ y \ s})" - (* TODO: generalize to class metric_space *) - -lemma diameter_bounded: - assumes "bounded s" - shows "\x\s. \y\s. norm(x - y) \ diameter s" - "\d>0. d < diameter s --> (\x\s. \y\s. norm(x - y) > d)" -proof- - let ?D = "{norm (x - y) |x y. x \ s \ y \ s}" - obtain a where a:"\x\s. norm x \ a" using assms[unfolded bounded_iff] by auto - { fix x y assume "x \ s" "y \ s" - hence "norm (x - y) \ 2 * a" using norm_triangle_ineq[of x "-y", unfolded norm_minus_cancel] a[THEN bspec[where x=x]] a[THEN bspec[where x=y]] by (auto simp add: ring_simps) } - note * = this - { fix x y assume "x\s" "y\s" hence "s \ {}" by auto - have lub:"isLub UNIV ?D (rsup ?D)" using * rsup[of ?D] using `s\{}` unfolding setle_def by auto - have "norm(x - y) \ diameter s" unfolding diameter_def using `s\{}` *[OF `x\s` `y\s`] `x\s` `y\s` isLubD1[OF lub] unfolding setle_def by auto } - moreover - { fix d::real assume "d>0" "d < diameter s" - hence "s\{}" unfolding diameter_def by auto - hence lub:"isLub UNIV ?D (rsup ?D)" using * rsup[of ?D] unfolding setle_def by auto - have "\d' \ ?D. d' > d" - proof(rule ccontr) - assume "\ (\d'\{norm (x - y) |x y. x \ s \ y \ s}. d < d')" - hence as:"\d'\?D. d' \ d" apply auto apply(erule_tac x="norm (x - y)" in allE) by auto - hence "isUb UNIV ?D d" unfolding isUb_def unfolding setle_def by auto - thus False using `d < diameter s` `s\{}` isLub_le_isUb[OF lub, of d] unfolding diameter_def by auto - qed - hence "\x\s. \y\s. norm(x - y) > d" by auto } - ultimately show "\x\s. \y\s. norm(x - y) \ diameter s" - "\d>0. d < diameter s --> (\x\s. \y\s. norm(x - y) > d)" by auto -qed - -lemma diameter_bounded_bound: - "bounded s \ x \ s \ y \ s ==> norm(x - y) \ diameter s" - using diameter_bounded by blast - -lemma diameter_compact_attained: - 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) 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 - thus ?thesis using diameter_bounded(1)[OF b, THEN bspec[where x=x], THEN bspec[where x=y], OF xys] and xys by auto -qed - -text{* Related results with closure as the conclusion. *} - -lemma closed_scaling: - fixes s :: "'a::real_normed_vector set" - assumes "closed s" shows "closed ((\x. c *\<^sub>R x) ` s)" -proof(cases "s={}") - case True thus ?thesis by auto -next - case False - show ?thesis - proof(cases "c=0") - have *:"(\x. 0) ` s = {0}" using `s\{}` by auto - case True thus ?thesis apply auto unfolding * using closed_sing by auto - next - case False - { fix x l assume as:"\n::nat. x n \ scaleR c ` s" "(x ---> l) sequentially" - { fix n::nat have "scaleR (1 / c) (x n) \ s" - using as(1)[THEN spec[where x=n]] - using `c\0` by (auto simp add: vector_smult_assoc) - } - moreover - { fix e::real assume "e>0" - hence "0 < e *\c\" using `c\0` mult_pos_pos[of e "abs c"] by auto - 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] - 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" - using assms[unfolded closed_sequential_limits, THEN spec[where x="\n. scaleR (1/c) (x n)"], THEN spec[where x="scaleR (1/c) l"]] - unfolding image_iff using `c\0` apply(rule_tac x="scaleR (1 / c) l" in bexI) by auto } - thus ?thesis unfolding closed_sequential_limits by fast - qed -qed - -lemma closed_negations: - fixes s :: "'a::real_normed_vector set" - assumes "closed s" shows "closed ((\x. -x) ` s)" - using closed_scaling[OF assms, of "- 1"] by simp - -lemma compact_closed_sums: - fixes s :: "'a::real_normed_vector set" - assumes "compact s" "closed t" shows "closed {x + y | x y. x \ s \ y \ t}" -proof- - let ?S = "{x + y |x y. x \ s \ y \ t}" - { 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:"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_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 - hence "l \ ?S" using `l' \ s` apply auto apply(rule_tac x=l' in exI) apply(rule_tac x="l - l'" in exI) by auto - } - thus ?thesis unfolding closed_sequential_limits by fast -qed - -lemma closed_compact_sums: - fixes s t :: "'a::real_normed_vector set" - assumes "closed s" "compact t" - shows "closed {x + y | x y. x \ s \ y \ t}" -proof- - have "{x + y |x y. x \ t \ y \ s} = {x + y |x y. x \ s \ y \ t}" apply auto - apply(rule_tac x=y in exI) apply auto apply(rule_tac x=y in exI) by auto - thus ?thesis using compact_closed_sums[OF assms(2,1)] by simp -qed - -lemma compact_closed_differences: - fixes s t :: "'a::real_normed_vector set" - assumes "compact s" "closed t" - shows "closed {x - y | x y. x \ s \ y \ t}" -proof- - have "{x + y |x y. x \ s \ y \ uminus ` t} = {x - y |x y. x \ s \ y \ 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_closed_sums[OF assms(1) closed_negations[OF assms(2)]] by auto -qed - -lemma closed_compact_differences: - fixes s t :: "'a::real_normed_vector set" - assumes "closed s" "compact t" - shows "closed {x - y | x y. x \ s \ y \ t}" -proof- - have "{x + y |x y. x \ s \ y \ uminus ` t} = {x - y |x y. x \ s \ y \ t}" - apply auto apply(rule_tac x=xa in exI) apply auto apply(rule_tac x=xa in exI) by auto - thus ?thesis using closed_compact_sums[OF assms(1) compact_negations[OF assms(2)]] by simp -qed - -lemma closed_translation: - fixes a :: "'a::real_normed_vector" - assumes "closed s" shows "closed ((\x. a + x) ` s)" -proof- - have "{a + y |y. y \ s} = (op + a ` s)" by auto - thus ?thesis using compact_closed_sums[OF compact_sing[of a] assms] by auto -qed - -lemma translation_UNIV: - fixes a :: "'a::ab_group_add" shows "range (\x. a + x) = UNIV" - apply (auto simp add: image_iff) apply(rule_tac x="x - a" in exI) by auto - -lemma translation_diff: - fixes a :: "'a::ab_group_add" - shows "(\x. a + x) ` (s - t) = ((\x. a + x) ` s) - ((\x. a + x) ` t)" - by auto - -lemma closure_translation: - fixes a :: "'a::real_normed_vector" - shows "closure ((\x. a + x) ` s) = (\x. a + x) ` (closure s)" -proof- - have *:"op + a ` (UNIV - s) = UNIV - op + a ` s" - apply auto unfolding image_iff apply(rule_tac x="x - a" in bexI) by auto - show ?thesis unfolding closure_interior translation_diff translation_UNIV - using interior_translation[of a "UNIV - s"] unfolding * by auto -qed - -lemma frontier_translation: - fixes a :: "'a::real_normed_vector" - shows "frontier((\x. a + x) ` s) = (\x. a + x) ` (frontier s)" - unfolding frontier_def translation_diff interior_translation closure_translation by auto - -subsection{* Separation between points and sets. *} - -lemma separate_point_closed: - fixes s :: "'a::heine_borel set" - shows "closed s \ a \ s ==> (\d>0. \x\s. d \ dist a x)" -proof(cases "s = {}") - case True - thus ?thesis by(auto intro!: exI[where x=1]) -next - case False - assume "closed s" "a \ s" - then obtain x where "x\s" "\y\s. dist a x \ dist a y" using `s \ {}` distance_attains_inf [of s a] by blast - with `x\s` show ?thesis using dist_pos_lt[of a x] and`a \ s` by blast -qed - -lemma separate_compact_closed: - 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- - have "0 \ {x - y |x y. x \ s \ y \ t}" using assms(3) by auto - then obtain d where "d>0" and d:"\x\{x - y |x y. x \ s \ y \ t}. d \ dist 0 x" - using separate_point_closed[OF compact_closed_differences[OF assms(1,2)], of 0] by auto - { fix x y assume "x\s" "y\t" - hence "x - y \ {x - y |x y. x \ s \ y \ t}" by auto - hence "d \ dist (x - y) 0" using d[THEN bspec[where x="x - y"]] using dist_commute - by (auto simp add: dist_commute) - hence "d \ dist x y" unfolding dist_norm by auto } - thus ?thesis using `d>0` by auto -qed - -lemma separate_closed_compact: - 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- - have *:"t \ s = {}" using assms(3) by auto - show ?thesis using separate_compact_closed[OF assms(2,1) *] - apply auto apply(rule_tac x=d in exI) apply auto apply (erule_tac x=y in ballE) - by (auto simp add: dist_commute) -qed - -(* A cute way of denoting open and closed intervals using overloading. *) - -lemma interval: fixes a :: "'a::ord^'n::finite" shows - "{a <..< b} = {x::'a^'n. \i. a$i < x$i \ x$i < b$i}" and - "{a .. b} = {x::'a^'n. \i. a$i \ x$i \ x$i \ b$i}" - by (auto simp add: expand_set_eq vector_less_def vector_less_eq_def) - -lemma mem_interval: fixes a :: "'a::ord^'n::finite" shows - "x \ {a<.. (\i. a$i < x$i \ x$i < b$i)" - "x \ {a .. b} \ (\i. a$i \ x$i \ x$i \ b$i)" - using interval[of a b] by(auto simp add: expand_set_eq vector_less_def vector_less_eq_def) - -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 forall_1) - -lemma interval_eq_empty: fixes a :: "real^'n::finite" shows - "({a <..< b} = {} \ (\i. b$i \ a$i))" (is ?th1) and - "({a .. b} = {} \ (\i. b$i < a$i))" (is ?th2) -proof- - { fix i x assume as:"b$i \ a$i" and x:"x\{a <..< b}" - hence "a $ i < x $ i \ x $ i < b $ i" unfolding mem_interval by auto - hence "a$i < b$i" by auto - hence False using as by auto } - moreover - { assume as:"\i. \ (b$i \ a$i)" - 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) *\<^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 } - ultimately show ?th1 by blast - - { fix i x assume as:"b$i < a$i" and x:"x\{a .. b}" - hence "a $ i \ x $ i \ x $ i \ b $ i" unfolding mem_interval by auto - hence "a$i \ b$i" by auto - hence False using as by auto } - moreover - { assume as:"\i. \ (b$i < a$i)" - 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) *\<^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 } - ultimately show ?th2 by blast -qed - -lemma interval_ne_empty: fixes a :: "real^'n::finite" shows - "{a .. b} \ {} \ (\i. a$i \ b$i)" and - "{a <..< b} \ {} \ (\i. a$i < b$i)" - unfolding interval_eq_empty[of a b] by (auto simp add: not_less not_le) (* BH: Why doesn't just "auto" work here? *) - -lemma subset_interval_imp: fixes a :: "real^'n::finite" shows - "(\i. a$i \ c$i \ d$i \ b$i) \ {c .. d} \ {a .. b}" and - "(\i. a$i < c$i \ d$i < b$i) \ {c .. d} \ {a<..i. a$i \ c$i \ d$i \ b$i) \ {c<.. {a .. b}" and - "(\i. a$i \ c$i \ d$i \ b$i) \ {c<.. {a<.. {a<.. {a .. b}" -proof(simp add: subset_eq, rule) - fix x - assume x:"x \{a<.. x $ i" - using x order_less_imp_le[of "a$i" "x$i"] - by(simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq) - } - moreover - { fix i - have "x $ i \ b $ i" - using x order_less_imp_le[of "x$i" "b$i"] - by(simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq) - } - ultimately - show "a \ x \ x \ b" - by(simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq) -qed - -lemma subset_interval: fixes a :: "real^'n::finite" shows - "{c .. d} \ {a .. b} \ (\i. c$i \ d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th1) and - "{c .. d} \ {a<.. (\i. c$i \ d$i) --> (\i. a$i < c$i \ d$i < b$i)" (is ?th2) and - "{c<.. {a .. b} \ (\i. c$i < d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th3) and - "{c<.. {a<.. (\i. c$i < d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th4) -proof- - show ?th1 unfolding subset_eq and Ball_def and mem_interval by (auto intro: order_trans) - show ?th2 unfolding subset_eq and Ball_def and mem_interval by (auto intro: le_less_trans less_le_trans order_trans less_imp_le) - { assume as: "{c<.. {a .. b}" "\i. c$i < d$i" - hence "{c<.. {}" unfolding interval_eq_empty by (auto, drule_tac x=i in spec, simp) (* BH: Why doesn't just "auto" work? *) - fix i - (** TODO combine the following two parts as done in the HOL_light version. **) - { let ?x = "(\ j. (if j=i then ((min (a$j) (d$j))+c$j)/2 else (c$j+d$j)/2))::real^'n" - assume as2: "a$i > c$i" - { fix j - have "c $ j < ?x $ j \ ?x $ j < d $ j" unfolding Cart_lambda_beta - apply(cases "j=i") using as(2)[THEN spec[where x=j]] - by (auto simp add: less_divide_eq_number_of1 as2) } - hence "?x\{c<..{a .. b}" - unfolding mem_interval apply auto apply(rule_tac x=i in exI) - using as(2)[THEN spec[where x=i]] and as2 - by (auto simp add: less_divide_eq_number_of1) - ultimately have False using as by auto } - hence "a$i \ c$i" by(rule ccontr)auto - moreover - { let ?x = "(\ j. (if j=i then ((max (b$j) (c$j))+d$j)/2 else (c$j+d$j)/2))::real^'n" - assume as2: "b$i < d$i" - { fix j - have "d $ j > ?x $ j \ ?x $ j > c $ j" unfolding Cart_lambda_beta - apply(cases "j=i") using as(2)[THEN spec[where x=j]] - by (auto simp add: less_divide_eq_number_of1 as2) } - hence "?x\{c<..{a .. b}" - unfolding mem_interval apply auto apply(rule_tac x=i in exI) - using as(2)[THEN spec[where x=i]] and as2 - by (auto simp add: less_divide_eq_number_of1) - ultimately have False using as by auto } - hence "b$i \ d$i" by(rule ccontr)auto - ultimately - have "a$i \ c$i \ d$i \ b$i" by auto - } note part1 = this - thus ?th3 unfolding subset_eq and Ball_def and mem_interval apply auto apply (erule_tac x=ia in allE, simp)+ by (erule_tac x=i in allE, erule_tac x=i in allE, simp)+ - { assume as:"{c<.. {a<..i. c$i < d$i" - fix i - from as(1) have "{c<.. {a..b}" using interval_open_subset_closed[of a b] by auto - hence "a$i \ c$i \ d$i \ b$i" using part1 and as(2) by auto } note * = this - thus ?th4 unfolding subset_eq and Ball_def and mem_interval apply auto apply (erule_tac x=ia in allE, simp)+ by (erule_tac x=i in allE, erule_tac x=i in allE, simp)+ -qed - -lemma disjoint_interval: fixes a::"real^'n::finite" shows - "{a .. b} \ {c .. d} = {} \ (\i. (b$i < a$i \ d$i < c$i \ b$i < c$i \ d$i < a$i))" (is ?th1) and - "{a .. b} \ {c<.. (\i. (b$i < a$i \ d$i \ c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th2) and - "{a<.. {c .. d} = {} \ (\i. (b$i \ a$i \ d$i < c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th3) and - "{a<.. {c<.. (\i. (b$i \ a$i \ d$i \ c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th4) -proof- - let ?z = "(\ i. ((max (a$i) (c$i)) + (min (b$i) (d$i))) / 2)::real^'n" - show ?th1 ?th2 ?th3 ?th4 - unfolding expand_set_eq and Int_iff and empty_iff and mem_interval and all_conj_distrib[THEN sym] and eq_False - apply (auto elim!: allE[where x="?z"]) - apply ((rule_tac x=x in exI, force) | (rule_tac x=i in exI, force))+ - done -qed - -lemma inter_interval: fixes a :: "'a::linorder^'n::finite" shows - "{a .. b} \ {c .. d} = {(\ i. max (a$i) (c$i)) .. (\ i. min (b$i) (d$i))}" - unfolding expand_set_eq and Int_iff and mem_interval - by (auto simp add: less_divide_eq_number_of1 intro!: bexI) - -(* Moved interval_open_subset_closed a bit upwards *) - -lemma open_interval_lemma: fixes x :: "real" shows - "a < x \ x < b ==> (\d>0. \x'. abs(x' - x) < d --> a < x' \ x' < b)" - by(rule_tac x="min (x - a) (b - x)" in exI, auto) - -lemma open_interval: fixes a :: "real^'n::finite" shows "open {a<..{a<..d>0. \x'. abs (x' - (x$i)) < d \ a$i < x' \ x' < b$i" - using x[unfolded mem_interval, THEN spec[where x=i]] - using open_interval_lemma[of "a$i" "x$i" "b$i"] by auto } - - hence "\i. \d>0. \x'. abs (x' - (x$i)) < d \ a$i < x' \ x' < b$i" by auto - then obtain d where d:"\i. 0 < d i \ (\x'. \x' - x $ i\ < d i \ a $ i < x' \ x' < b $ i)" - using bchoice[of "UNIV" "\i d. d>0 \ (\x'. \x' - x $ i\ < d \ a $ i < x' \ x' < b $ i)"] by auto - - let ?d = "Min (range d)" - have **:"finite (range d)" "range d \ {}" by auto - have "?d>0" unfolding Min_gr_iff[OF **] using d by auto - moreover - { fix x' assume as:"dist x' x < ?d" - { fix i - have "\x'$i - x $ i\ < d i" - using norm_bound_component_lt[OF as[unfolded dist_norm], of i] - unfolding vector_minus_component and Min_gr_iff[OF **] by auto - hence "a $ i < x' $ i" "x' $ i < b $ i" using d[THEN spec[where x=i]] by auto } - hence "a < x' \ x' < b" unfolding vector_less_def by auto } - ultimately have "\e>0. \x'. dist x' x < e \ x' \ {a<..e>0. \x'\{a..b}. x' \ x \ dist x' x < e"(* and xab:"a$i > x$i \ b$i < x$i"*) - { assume xa:"a$i > x$i" - with as obtain y where y:"y\{a..b}" "y \ x" "dist y x < a$i - x$i" by(erule_tac x="a$i - x$i" in allE)auto - hence False unfolding mem_interval and dist_norm - using component_le_norm[of "y-x" i, unfolded vector_minus_component] and xa by(auto elim!: allE[where x=i]) - } hence "a$i \ x$i" by(rule ccontr)auto - moreover - { assume xb:"b$i < x$i" - with as obtain y where y:"y\{a..b}" "y \ x" "dist y x < x$i - b$i" by(erule_tac x="x$i - b$i" in allE)auto - hence False unfolding mem_interval and dist_norm - using component_le_norm[of "y-x" i, unfolded vector_minus_component] and xb by(auto elim!: allE[where x=i]) - } hence "x$i \ b$i" by(rule ccontr)auto - ultimately - have "a $ i \ x $ i \ x $ i \ b $ i" by auto } - thus ?thesis unfolding closed_limpt islimpt_approachable mem_interval by auto -qed - -lemma interior_closed_interval: fixes a :: "real^'n::finite" shows - "interior {a .. b} = {a<.. ?L" using interior_maximal[OF interval_open_subset_closed open_interval] by auto -next - { fix x assume "\T. open T \ x \ T \ T \ {a..b}" - 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) *\<^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 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 - unfolding vector_smult_component and basis_component using `e>0` by auto } - hence "x \ {a<.. ?R" unfolding interior_def and subset_eq by auto -qed - -lemma bounded_closed_interval: fixes a :: "real^'n::finite" shows - "bounded {a .. b}" -proof- - let ?b = "\i\UNIV. \a$i\ + \b$i\" - { fix x::"real^'n" assume x:"\i. a $ i \ x $ i \ x $ i \ b $ i" - { fix i - have "\x$i\ \ \a$i\ + \b$i\" using x[THEN spec[where x=i]] by auto } - hence "(\i\UNIV. \x $ i\) \ ?b" by(rule setsum_mono) - hence "norm x \ ?b" using norm_le_l1[of x] by auto } - thus ?thesis unfolding interval and bounded_iff by auto -qed - -lemma bounded_interval: fixes a :: "real^'n::finite" shows - "bounded {a .. b} \ bounded {a<.. UNIV) \ ({a<.. UNIV)" - using bounded_interval[of a b] - by auto - -lemma compact_interval: fixes a :: "real^'n::finite" shows - "compact {a .. b}" - 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) *\<^sub>R (a + b)) \ {a<..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) } - thus ?thesis unfolding mem_interval by auto -qed - -lemma open_closed_interval_convex: fixes x :: "real^'n::finite" - assumes x:"x \ {a<.. {a .. b}" and e:"0 < e" "e \ 1" - shows "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \ {a<.. < e * x $ i + (1 - e) * y $ i" apply(rule add_less_le_mono) - using e unfolding mult_less_cancel_left and mult_le_cancel_left apply simp_all - using x unfolding mem_interval apply simp - using y unfolding mem_interval apply simp - done - finally have "a $ i < (e *\<^sub>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) - using e unfolding mult_less_cancel_left and mult_le_cancel_left apply simp_all - using x unfolding mem_interval apply simp - using y unfolding mem_interval apply simp - done - 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 - -lemma closure_open_interval: fixes a :: "real^'n::finite" - assumes "{a<.. {}" - shows "closure {a<.. {a .. b}" - 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)) *\<^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 - { assume "\ (f ---> x) sequentially" - { fix e::real assume "e>0" - hence "\N::nat. inverse (real (N + 1)) < e" using real_arch_inv[of e] apply (auto simp add: Suc_pred') apply(rule_tac x="n - 1" in exI) by auto - 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 "((\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)) *\<^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<..a. s \ {-a<..0" and b:"\x\s. norm x \ b" using assms[unfolded bounded_pos] by auto - def a \ "(\ i. b+1)::real^'n" - { fix x assume "x\s" - fix i - have "(-a)$i < x$i" and "x$i < a$i" using b[THEN bspec[where x=x], OF `x\s`] and component_le_norm[of x i] - unfolding vector_uminus_component and a_def and Cart_lambda_beta by auto - } - thus ?thesis by(auto intro: exI[where x=a] simp add: vector_less_def) -qed - -lemma bounded_subset_open_interval: - fixes s :: "(real ^ 'n::finite) set" - shows "bounded s ==> (\a b. s \ {a<..a. s \ {-a .. a}" -proof- - obtain a where "s \ {- a<.. (\a b. s \ {a .. b})" - using bounded_subset_closed_interval_symmetric[of s] by auto - -lemma frontier_closed_interval: - fixes a b :: "real ^ _" - shows "frontier {a .. b} = {a .. b} - {a<.. {}" shows "{a<.. {c .. d} = {} \ {a<.. {c<..x::1. P x) \ P 1" - by (metis num1_eq_iff) - -lemma ex_1: "(\x::1. P x) \ P 1" - by auto (metis num1_eq_iff) - -lemma interval_cases_1: fixes x :: "real^1" shows - "x \ {a .. b} ==> x \ {a<.. (x = a) \ (x = b)" - by(simp add: Cart_eq vector_less_def vector_less_eq_def all_1, auto) - -lemma in_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 add: Cart_eq vector_less_def vector_less_eq_def all_1 dest_vec1_def) - -lemma interval_eq_empty_1: fixes a :: "real^1" shows - "{a .. b} = {} \ dest_vec1 b < dest_vec1 a" - "{a<.. dest_vec1 b \ dest_vec1 a" - unfolding interval_eq_empty and ex_1 and dest_vec1_def by auto - -lemma subset_interval_1: fixes a :: "real^1" shows - "({a .. b} \ {c .. d} \ dest_vec1 b < dest_vec1 a \ - dest_vec1 c \ dest_vec1 a \ dest_vec1 a \ dest_vec1 b \ dest_vec1 b \ dest_vec1 d)" - "({a .. b} \ {c<.. dest_vec1 b < dest_vec1 a \ - dest_vec1 c < dest_vec1 a \ dest_vec1 a \ dest_vec1 b \ dest_vec1 b < dest_vec1 d)" - "({a<.. {c .. d} \ dest_vec1 b \ dest_vec1 a \ - dest_vec1 c \ dest_vec1 a \ dest_vec1 a < dest_vec1 b \ dest_vec1 b \ dest_vec1 d)" - "({a<.. {c<.. dest_vec1 b \ dest_vec1 a \ - dest_vec1 c \ dest_vec1 a \ dest_vec1 a < dest_vec1 b \ dest_vec1 b \ dest_vec1 d)" - unfolding subset_interval[of a b c d] unfolding all_1 and dest_vec1_def by auto - -lemma eq_interval_1: fixes a :: "real^1" shows - "{a .. b} = {c .. d} \ - dest_vec1 b < dest_vec1 a \ dest_vec1 d < dest_vec1 c \ - dest_vec1 a = dest_vec1 c \ dest_vec1 b = dest_vec1 d" -using set_eq_subset[of "{a .. b}" "{c .. d}"] -using subset_interval_1(1)[of a b c d] -using subset_interval_1(1)[of c d a b] -by auto (* FIXME: slow *) - -lemma disjoint_interval_1: fixes a :: "real^1" shows - "{a .. b} \ {c .. d} = {} \ dest_vec1 b < dest_vec1 a \ dest_vec1 d < dest_vec1 c \ dest_vec1 b < dest_vec1 c \ dest_vec1 d < dest_vec1 a" - "{a .. b} \ {c<.. dest_vec1 b < dest_vec1 a \ dest_vec1 d \ dest_vec1 c \ dest_vec1 b \ dest_vec1 c \ dest_vec1 d \ dest_vec1 a" - "{a<.. {c .. d} = {} \ dest_vec1 b \ dest_vec1 a \ dest_vec1 d < dest_vec1 c \ dest_vec1 b \ dest_vec1 c \ dest_vec1 d \ dest_vec1 a" - "{a<.. {c<.. dest_vec1 b \ dest_vec1 a \ dest_vec1 d \ dest_vec1 c \ dest_vec1 b \ dest_vec1 c \ dest_vec1 d \ dest_vec1 a" - unfolding disjoint_interval and dest_vec1_def ex_1 by auto - -lemma open_closed_interval_1: fixes a :: "real^1" shows - "{a<.. dest_vec1 b ==> {a .. b} = {a<.. {a,b}" - unfolding expand_set_eq apply simp unfolding vector_less_def and vector_less_eq_def and all_1 and dest_vec1_eq[THEN sym] and dest_vec1_def by auto - -(* Some stuff for half-infinite intervals too; FIXME: notation? *) - -lemma closed_interval_left: fixes b::"real^'n::finite" - shows "closed {x::real^'n. \i. x$i \ b$i}" -proof- - { fix i - fix x::"real^'n" assume x:"\e>0. \x'\{x. \i. x $ i \ b $ i}. x' \ x \ dist x' x < e" - { assume "x$i > b$i" - then obtain y where "y $ i \ b $ i" "y \ x" "dist y x < x$i - b$i" using x[THEN spec[where x="x$i - b$i"]] by auto - hence False using component_le_norm[of "y - x" i] unfolding dist_norm and vector_minus_component by auto } - hence "x$i \ b$i" by(rule ccontr)auto } - thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast -qed - -lemma closed_interval_right: fixes a::"real^'n::finite" - shows "closed {x::real^'n. \i. a$i \ x$i}" -proof- - { fix i - fix x::"real^'n" assume x:"\e>0. \x'\{x. \i. a $ i \ x $ i}. x' \ x \ dist x' x < e" - { assume "a$i > x$i" - then obtain y where "a $ i \ y $ i" "y \ x" "dist y x < a$i - x$i" using x[THEN spec[where x="a$i - x$i"]] by auto - hence False using component_le_norm[of "y - x" i] unfolding dist_norm and vector_minus_component by auto } - hence "a$i \ x$i" by(rule ccontr)auto } - thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast -qed - -subsection{* Intervals in general, including infinite and mixtures of open and closed. *} - -definition "is_interval s \ (\a\s. \b\s. \x. (\i. ((a$i \ x$i \ x$i \ b$i) \ (b$i \ x$i \ x$i \ a$i))) \ x \ s)" - -lemma is_interval_interval: "is_interval {a .. b::real^'n::finite}" (is ?th1) "is_interval {a<..x y z::real. x < y \ y < z \ x < z" by auto - show ?th1 ?th2 unfolding is_interval_def mem_interval Ball_def atLeastAtMost_iff - by(meson real_le_trans le_less_trans less_le_trans *)+ qed - -lemma is_interval_empty: - "is_interval {}" - unfolding is_interval_def - by simp - -lemma is_interval_univ: - "is_interval UNIV" - unfolding is_interval_def - by simp - -subsection{* Closure of halfspaces and hyperplanes. *} - -lemma Lim_inner: - assumes "(f ---> l) net" shows "((\y. inner a (f y)) ---> inner a l) net" - by (intro tendsto_intros assms) - -lemma continuous_at_inner: "continuous (at x) (inner a)" - unfolding continuous_at by (intro tendsto_intros) - -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. continuous (at x) (inner a)" - unfolding continuous_at by (rule allI) (intro tendsto_intros) - hence "closed (inner a -` {..b})" - using closed_real_atMost by (rule continuous_closed_vimage) - moreover have "{x. inner a x \ b} = inner a -` {..b}" by auto - ultimately show ?thesis by simp -qed - -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. 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 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 inner_basis[OF assms] by auto - -text{* Openness of halfspaces. *} - -lemma open_halfspace_lt: "open {x. inner a x < b}" -proof- - 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. inner a x > b}" -proof- - 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 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 inner_basis[OF assms] by auto - -text{* This gives a simple derivation of limit component bounds. *} - -lemma Lim_component_le: fixes f :: "'a \ real^'n::finite" - 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. 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 - -lemma Lim_component_ge: fixes f :: "'a \ real^'n::finite" - 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. 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 - -lemma Lim_component_eq: fixes f :: "'a \ real^'n::finite" - assumes net:"(f ---> l) net" "~(trivial_limit net)" and ev:"eventually (\x. f(x)$i = b) net" - shows "l$i = b" - using ev[unfolded order_eq_iff eventually_and] using Lim_component_ge[OF net, of b i] and Lim_component_le[OF net, of i b] by auto - -lemma Lim_drop_le: fixes f :: "'a \ real^1" shows - "(f ---> l) net \ ~(trivial_limit net) \ eventually (\x. dest_vec1 (f x) \ b) net ==> dest_vec1 l \ b" - using Lim_component_le[of f l net 1 b] unfolding dest_vec1_def by auto - -lemma Lim_drop_ge: fixes f :: "'a \ real^1" shows - "(f ---> l) net \ ~(trivial_limit net) \ eventually (\x. b \ dest_vec1 (f x)) net ==> b \ dest_vec1 l" - using Lim_component_ge[of f l net b 1] unfolding dest_vec1_def by auto - -text{* Limits relative to a union. *} - -lemma eventually_within_Un: - "eventually P (net within (s \ t)) \ - eventually P (net within s) \ eventually P (net within t)" - unfolding Limits.eventually_within - by (auto elim!: eventually_rev_mp) - -lemma Lim_within_union: - "(f ---> l) (net within (s \ t)) \ - (f ---> l) (net within s) \ (f ---> l) (net within t)" - unfolding tendsto_def - by (auto simp add: eventually_within_Un) - -lemma continuous_on_union: - assumes "closed s" "closed t" "continuous_on s f" "continuous_on t f" - shows "continuous_on (s \ t) f" - using assms unfolding continuous_on unfolding Lim_within_union - unfolding Lim unfolding trivial_limit_within unfolding closed_limpt by auto - -lemma continuous_on_cases: - assumes "closed s" "closed t" "continuous_on s f" "continuous_on t g" - "\x. (x\s \ \ P x) \ (x \ t \ P x) \ f x = g x" - shows "continuous_on (s \ t) (\x. if P x then f x else g x)" -proof- - let ?h = "(\x. if P x then f x else g x)" - have "\x\s. f x = (if P x then f x else g x)" using assms(5) by auto - hence "continuous_on s ?h" using continuous_on_eq[of s f ?h] using assms(3) by auto - moreover - have "\x\t. g x = (if P x then f x else g x)" using assms(5) by auto - hence "continuous_on t ?h" using continuous_on_eq[of t g ?h] using assms(4) by auto - ultimately show ?thesis using continuous_on_union[OF assms(1,2), of ?h] by auto -qed - - -text{* Some more convenient intermediate-value theorem formulations. *} - -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. 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 - ultimately show False using assms(1)[unfolded connected_def not_ex, THEN spec[where x="?A"], THEN spec[where x="?B"]] and assms(2-5) by auto -qed - -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: inner_basis) - -text{* Also more convenient formulations of monotone convergence. *} - -lemma bounded_increasing_convergent: fixes s::"nat \ real^1" - assumes "bounded {s n| n::nat. True}" "\n. dest_vec1(s n) \ dest_vec1(s(Suc n))" - shows "\l. (s ---> l) sequentially" -proof- - obtain a where a:"\n. \dest_vec1 (s n)\ \ a" using assms(1)[unfolded bounded_iff abs_dest_vec1] by auto - { fix m::nat - 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] 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 - -subsection{* Basic homeomorphism definitions. *} - -definition "homeomorphism s t f g \ - (\x\s. (g(f x) = x)) \ (f ` s = t) \ continuous_on s f \ - (\y\t. (f(g y) = y)) \ (g ` t = s) \ continuous_on t g" - -definition - homeomorphic :: "'a::metric_space set \ 'b::metric_space set \ bool" - (infixr "homeomorphic" 60) where - homeomorphic_def: "s homeomorphic t \ (\f g. homeomorphism s t f g)" - -lemma homeomorphic_refl: "s homeomorphic s" - unfolding homeomorphic_def - unfolding homeomorphism_def - using continuous_on_id - apply(rule_tac x = "(\x. x)" in exI) - apply(rule_tac x = "(\x. x)" in exI) - by blast - -lemma homeomorphic_sym: - "s homeomorphic t \ t homeomorphic s" -unfolding homeomorphic_def -unfolding homeomorphism_def -by blast (* FIXME: slow *) - -lemma homeomorphic_trans: - assumes "s homeomorphic t" "t homeomorphic u" shows "s homeomorphic u" -proof- - obtain f1 g1 where fg1:"\x\s. g1 (f1 x) = x" "f1 ` s = t" "continuous_on s f1" "\y\t. f1 (g1 y) = y" "g1 ` t = s" "continuous_on t g1" - using assms(1) unfolding homeomorphic_def homeomorphism_def by auto - obtain f2 g2 where fg2:"\x\t. g2 (f2 x) = x" "f2 ` t = u" "continuous_on t f2" "\y\u. f2 (g2 y) = y" "g2 ` u = t" "continuous_on u g2" - using assms(2) unfolding homeomorphic_def homeomorphism_def by auto - - { fix x assume "x\s" hence "(g1 \ g2) ((f2 \ f1) x) = x" using fg1(1)[THEN bspec[where x=x]] and fg2(1)[THEN bspec[where x="f1 x"]] and fg1(2) by auto } - moreover have "(f2 \ f1) ` s = u" using fg1(2) fg2(2) by auto - moreover have "continuous_on s (f2 \ f1)" using continuous_on_compose[OF fg1(3)] and fg2(3) unfolding fg1(2) by auto - moreover { fix y assume "y\u" hence "(f2 \ f1) ((g1 \ g2) y) = y" using fg2(4)[THEN bspec[where x=y]] and fg1(4)[THEN bspec[where x="g2 y"]] and fg2(5) by auto } - moreover have "(g1 \ g2) ` u = s" using fg1(5) fg2(5) by auto - moreover have "continuous_on u (g1 \ g2)" using continuous_on_compose[OF fg2(6)] and fg1(6) unfolding fg2(5) by auto - ultimately show ?thesis unfolding homeomorphic_def homeomorphism_def apply(rule_tac x="f2 \ f1" in exI) apply(rule_tac x="g1 \ g2" in exI) by auto -qed - -lemma homeomorphic_minimal: - "s homeomorphic t \ - (\f g. (\x\s. f(x) \ t \ (g(f(x)) = x)) \ - (\y\t. g(y) \ s \ (f(g(y)) = y)) \ - continuous_on s f \ continuous_on t g)" -unfolding homeomorphic_def homeomorphism_def -apply auto apply (rule_tac x=f in exI) apply (rule_tac x=g in exI) -apply auto apply (rule_tac x=f in exI) apply (rule_tac x=g in exI) apply auto -unfolding image_iff -apply(erule_tac x="g x" in ballE) apply(erule_tac x="x" in ballE) -apply auto apply(rule_tac x="g x" in bexI) apply auto -apply(erule_tac x="f x" in ballE) apply(erule_tac x="x" in ballE) -apply auto apply(rule_tac x="f x" in bexI) by auto - -subsection{* Relatively weak hypotheses if a set is compact. *} - -definition "inv_on f s = (\x. SOME y. y\s \ f y = x)" - -lemma assumes "inj_on f s" "x\s" - shows "inv_on f s (f x) = x" - using assms unfolding inj_on_def inv_on_def by auto - -lemma homeomorphism_compact: - fixes f :: "'a::heine_borel \ 'b::heine_borel" - (* class constraint due to continuous_on_inverse *) - assumes "compact s" "continuous_on s f" "f ` s = t" "inj_on f s" - shows "\g. homeomorphism s t f g" -proof- - def g \ "\x. SOME y. y\s \ f y = x" - have g:"\x\s. g (f x) = x" using assms(3) assms(4)[unfolded inj_on_def] unfolding g_def by auto - { fix y assume "y\t" - then obtain x where x:"f x = y" "x\s" using assms(3) by auto - hence "g (f x) = x" using g by auto - hence "f (g y) = y" unfolding x(1)[THEN sym] by auto } - hence g':"\x\t. f (g x) = x" by auto - moreover - { fix x - have "x\s \ x \ g ` t" using g[THEN bspec[where x=x]] unfolding image_iff using assms(3) by(auto intro!: bexI[where x="f x"]) - moreover - { assume "x\g ` t" - then obtain y where y:"y\t" "g y = x" by auto - then obtain x' where x':"x'\s" "f x' = y" using assms(3) by auto - hence "x \ s" unfolding g_def using someI2[of "\b. b\s \ f b = y" x' "\x. x\s"] unfolding y(2)[THEN sym] and g_def by auto } - ultimately have "x\s \ x \ g ` t" by auto } - hence "g ` t = s" by auto - ultimately - show ?thesis unfolding homeomorphism_def homeomorphic_def - apply(rule_tac x=g in exI) using g and assms(3) and continuous_on_inverse[OF assms(2,1), of g, unfolded assms(3)] and assms(2) by auto -qed - -lemma homeomorphic_compact: - fixes f :: "'a::heine_borel \ 'b::heine_borel" - (* class constraint due to continuous_on_inverse *) - shows "compact s \ continuous_on s f \ (f ` s = t) \ inj_on f s - \ s homeomorphic t" - unfolding homeomorphic_def by(metis homeomorphism_compact) - -text{* Preservation of topological properties. *} - -lemma homeomorphic_compactness: - "s homeomorphic t ==> (compact s \ compact t)" -unfolding homeomorphic_def homeomorphism_def -by (metis compact_continuous_image) - -text{* Results on translation, scaling etc. *} - -lemma homeomorphic_scaling: - fixes s :: "'a::real_normed_vector set" - assumes "c \ 0" shows "s homeomorphic ((\x. c *\<^sub>R x) ` s)" - unfolding homeomorphic_minimal - 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: - fixes s :: "'a::real_normed_vector set" - shows "s homeomorphic ((\x. a + x) ` s)" - unfolding homeomorphic_minimal - apply(rule_tac x="\x. a + x" in exI) - apply(rule_tac x="\x. -a + x" in exI) - using continuous_on_add[OF continuous_on_const continuous_on_id] by auto - -lemma homeomorphic_affinity: - fixes s :: "'a::real_normed_vector set" - assumes "c \ 0" shows "s homeomorphic ((\x. a + c *\<^sub>R x) ` s)" -proof- - 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 *\<^sub>R x) ` s" a] unfolding * by auto -qed - -lemma homeomorphic_balls: - fixes a b ::"'a::real_normed_vector" (* FIXME: generalize to metric_space *) - assumes "0 < d" "0 < e" - shows "(ball a d) homeomorphic (ball b e)" (is ?th) - "(cball a d) homeomorphic (cball b e)" (is ?cth) -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) *\<^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) *\<^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. *} - -lemma cauchy_isometric: - fixes x :: "nat \ real ^ 'n::finite" - assumes e:"0 < e" and s:"subspace s" and f:"bounded_linear f" and normf:"\x\s. norm(f x) \ e * norm(x)" and xs:"\n::nat. x n \ s" and cf:"Cauchy(f o x)" - shows "Cauchy x" -proof- - interpret f: bounded_linear f by fact - { fix d::real assume "d>0" - then obtain N where N:"\n\N. norm (f (x n) - f (x N)) < e * d" - using cf[unfolded cauchy o_def dist_norm, THEN spec[where x="e*d"]] and e and mult_pos_pos[of e d] by auto - { fix n assume "n\N" - hence "norm (f (x n - x N)) < e * d" using N[THEN spec[where x=n]] unfolding f.diff[THEN sym] by auto - moreover have "e * norm (x n - x N) \ norm (f (x n - x N))" - using subspace_sub[OF s, of "x n" "x N"] using xs[THEN spec[where x=N]] and xs[THEN spec[where x=n]] - using normf[THEN bspec[where x="x n - x N"]] by auto - ultimately have "norm (x n - x N) < d" using `e>0` - using mult_left_less_imp_less[of e "norm (x n - x N)" d] by auto } - hence "\N. \n\N. norm (x n - x N) < d" by auto } - thus ?thesis unfolding cauchy and dist_norm by auto -qed - -lemma complete_isometric_image: - fixes f :: "real ^ _ \ real ^ _" - assumes "0 < e" and s:"subspace s" and f:"bounded_linear f" and normf:"\x\s. norm(f x) \ e * norm(x)" and cs:"complete s" - shows "complete(f ` s)" -proof- - { fix g assume as:"\n::nat. g n \ f ` s" and cfg:"Cauchy g" - then obtain x where "\n. x n \ s \ g n = f (x n)" unfolding image_iff and Bex_def - using choice[of "\ n xa. xa \ s \ g n = f xa"] by auto - hence x:"\n. x n \ s" "\n. g n = f (x n)" by auto - hence "f \ x = g" unfolding expand_fun_eq by auto - then obtain l where "l\s" and l:"(x ---> l) sequentially" - using cs[unfolded complete_def, THEN spec[where x="x"]] - using cauchy_isometric[OF `0l\f ` s. (g ---> l) sequentially" - using linear_continuous_at[OF f, unfolded continuous_at_sequentially, THEN spec[where x=x], of l] - unfolding `f \ x = g` by auto } - thus ?thesis unfolding complete_def by auto -qed - -lemma dist_0_norm: - fixes x :: "'a::real_normed_vector" - shows "dist 0 x = norm x" -unfolding dist_norm by simp - -lemma injective_imp_isometric: fixes f::"real^'m::finite \ real^'n::finite" - assumes s:"closed s" "subspace s" and f:"bounded_linear f" "\x\s. (f x = 0) \ (x = 0)" - shows "\e>0. \x\s. norm (f x) \ e * norm(x)" -proof(cases "s \ {0::real^'m}") - case True - { fix x assume "x \ s" - hence "x = 0" using True by auto - hence "norm x \ norm (f x)" by auto } - thus ?thesis by(auto intro!: exI[where x=1]) -next - interpret f: bounded_linear f by fact - case False - then obtain a where a:"a\0" "a\s" by auto - from False have "s \ {}" by auto - let ?S = "{f x| x. (x \ s \ norm x = norm a)}" - let ?S' = "{x::real^'m. x\s \ norm x = norm a}" - let ?S'' = "{x::real^'m. norm x = norm a}" - - have "?S'' = frontier(cball 0 (norm a))" unfolding frontier_cball and dist_norm by (auto simp add: norm_minus_cancel) - hence "compact ?S''" using compact_frontier[OF compact_cball, of 0 "norm a"] by auto - moreover have "?S' = s \ ?S''" by auto - ultimately have "compact ?S'" using closed_inter_compact[of s ?S''] using s(1) by auto - moreover have *:"f ` ?S' = ?S" by auto - ultimately have "compact ?S" using compact_continuous_image[OF linear_continuous_on[OF f(1)], of ?S'] by auto - hence "closed ?S" using compact_imp_closed by auto - moreover have "?S \ {}" using a by auto - ultimately obtain b' where "b'\?S" "\y\?S. norm b' \ norm y" using distance_attains_inf[of ?S 0] unfolding dist_0_norm by auto - then obtain b where "b\s" and ba:"norm b = norm a" and b:"\x\{x \ s. norm x = norm a}. norm (f b) \ norm (f x)" unfolding *[THEN sym] unfolding image_iff by auto - - let ?e = "norm (f b) / norm b" - have "norm b > 0" using ba and a and norm_ge_zero by auto - moreover have "norm (f b) > 0" using f(2)[THEN bspec[where x=b], OF `b\s`] using `norm b >0` unfolding zero_less_norm_iff by auto - ultimately have "0 < norm (f b) / norm b" by(simp only: divide_pos_pos) - moreover - { fix x assume "x\s" - hence "norm (f b) / norm b * norm x \ norm (f x)" - proof(cases "x=0") - case True thus "norm (f b) / norm b * norm x \ norm (f x)" by auto - 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 *\<^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 f.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 - show ?thesis by auto -qed - -lemma closed_injective_image_subspace: - fixes f :: "real ^ _ \ real ^ _" - assumes "subspace s" "bounded_linear f" "\x\s. f x = 0 --> x = 0" "closed s" - shows "closed(f ` s)" -proof- - obtain e where "e>0" and e:"\x\s. e * norm x \ norm (f x)" using injective_imp_isometric[OF assms(4,1,2,3)] by auto - show ?thesis using complete_isometric_image[OF `e>0` assms(1,2) e] and assms(4) - unfolding complete_eq_closed[THEN sym] by auto -qed - -subsection{* Some properties of a canonical subspace. *} - -lemma subspace_substandard: - "subspace {x::real^'n. (\i. P i \ x$i = 0)}" - unfolding subspace_def by(auto simp add: vector_add_component vector_smult_component elim!: ballE) - -lemma closed_substandard: - "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. 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: 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. 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 - thus ?thesis by(auto simp add: closed_Inter closed_hyperplane) -qed - -lemma dim_substandard: - shows "dim {x::real^'n::finite. \i. i \ d \ x$i = 0} = card d" (is "dim ?A = _") -proof- - let ?D = "UNIV::'n set" - let ?B = "(basis::'n\real^'n) ` d" - - let ?bas = "basis::'n \ real^'n" - - have "?B \ ?A" by auto - - moreover - { fix x::"real^'n" assume "x\?A" - with finite[of d] - have "x\ span ?B" - proof(induct d arbitrary: x) - case empty hence "x=0" unfolding Cart_eq by auto - thus ?case using subspace_0[OF subspace_span[of "{}"]] by auto - next - 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 *\<^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 - using *[THEN spec[where x=i]] by auto } - hence "y \ span (basis ` (insert k F))" using insert(3) - using span_mono[of "?bas ` F" "?bas ` (insert k F)"] - 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 *\<^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 *\<^sub>R basis k \ span (?bas ` (insert k F))" - using span_add by auto - thus ?case using y by auto - qed - } - hence "?A \ span ?B" by auto - - moreover - { fix x assume "x \ ?B" - hence "x\{(basis i)::real^'n |i. i \ ?D}" using assms by auto } - hence "independent ?B" using independent_mono[OF independent_stdbasis, of ?B] and assms by auto - - moreover - have "d \ ?D" unfolding subset_eq using assms by auto - hence *:"inj_on (basis::'n\real^'n) d" using subset_inj_on[OF basis_inj, of "d"] by auto - have "?B hassize (card d)" unfolding hassize_def and card_image[OF *] by auto - - ultimately show ?thesis using dim_unique[of "basis ` d" ?A] by auto -qed - -text{* Hence closure and completeness of all subspaces. *} - -lemma closed_subspace_lemma: "n \ card (UNIV::'n::finite set) \ \A::'n set. card A = n" -apply (induct n) -apply (rule_tac x="{}" in exI, simp) -apply clarsimp -apply (subgoal_tac "\x. x \ A") -apply (erule exE) -apply (rule_tac x="insert x A" in exI, simp) -apply (subgoal_tac "A \ UNIV", auto) -done - -lemma closed_subspace: fixes s::"(real^'n::finite) set" - assumes "subspace s" shows "closed s" -proof- - have "dim s \ card (UNIV :: 'n set)" using dim_subset_univ by auto - then obtain d::"'n set" where t: "card d = dim s" - using closed_subspace_lemma by auto - let ?t = "{x::real^'n. \i. i \ d \ x$i = 0}" - obtain f where f:"bounded_linear f" "f ` ?t = s" "inj_on f ?t" - using subspace_isomorphism[unfolded linear_conv_bounded_linear, OF subspace_substandard[of "\i. i \ d"] assms] - using dim_substandard[of d] and t by auto - interpret f: bounded_linear f by fact - have "\x\?t. f x = 0 \ x = 0" using f.zero using f(3)[unfolded inj_on_def] - by(erule_tac x=0 in ballE) auto - moreover have "closed ?t" using closed_substandard . - moreover have "subspace ?t" using subspace_substandard . - ultimately show ?thesis using closed_injective_image_subspace[of ?t f] - unfolding f(2) using f(1) by auto -qed - -lemma complete_subspace: - fixes s :: "(real ^ _) set" shows "subspace s ==> complete s" - using complete_eq_closed closed_subspace - by auto - -lemma dim_closure: - fixes s :: "(real ^ _) set" - shows "dim(closure s) = dim s" (is "?dc = ?d") -proof- - have "?dc \ ?d" using closure_minimal[OF span_inc, of s] - using closed_subspace[OF subspace_span, of s] - using dim_subset[of "closure s" "span s"] unfolding dim_span by auto - thus ?thesis using dim_subset[OF closure_subset, of s] by auto -qed - -text{* Affine transformations of intervals. *} - -lemma affinity_inverses: - assumes m0: "m \ (0::'a::field)" - shows "(\x. m *s x + c) o (\x. inverse(m) *s x + (-(inverse(m) *s c))) = id" - "(\x. inverse(m) *s x + (-(inverse(m) *s c))) o (\x. m *s x + c) = id" - using m0 -apply (auto simp add: expand_fun_eq vector_add_ldistrib vector_smult_assoc) -by (simp add: vector_smult_lneg[symmetric] vector_smult_assoc vector_sneg_minus1[symmetric]) - -lemma real_affinity_le: - "0 < (m::'a::ordered_field) ==> (m * x + c \ y \ x \ inverse(m) * y + -(c / m))" - by (simp add: field_simps inverse_eq_divide) - -lemma real_le_affinity: - "0 < (m::'a::ordered_field) ==> (y \ m * x + c \ inverse(m) * y + -(c / m) \ x)" - by (simp add: field_simps inverse_eq_divide) - -lemma real_affinity_lt: - "0 < (m::'a::ordered_field) ==> (m * x + c < y \ x < inverse(m) * y + -(c / m))" - by (simp add: field_simps inverse_eq_divide) - -lemma real_lt_affinity: - "0 < (m::'a::ordered_field) ==> (y < m * x + c \ inverse(m) * y + -(c / m) < x)" - by (simp add: field_simps inverse_eq_divide) - -lemma real_affinity_eq: - "(m::'a::ordered_field) \ 0 ==> (m * x + c = y \ x = inverse(m) * y + -(c / m))" - by (simp add: field_simps inverse_eq_divide) - -lemma real_eq_affinity: - "(m::'a::ordered_field) \ 0 ==> (y = m * x + c \ inverse(m) * y + -(c / m) = x)" - by (simp add: field_simps inverse_eq_divide) - -lemma vector_affinity_eq: - assumes m0: "(m::'a::field) \ 0" - shows "m *s x + c = y \ x = inverse m *s y + -(inverse m *s c)" -proof - assume h: "m *s x + c = y" - hence "m *s x = y - c" by (simp add: ring_simps) - hence "inverse m *s (m *s x) = inverse m *s (y - c)" by simp - then show "x = inverse m *s y + - (inverse m *s c)" - using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib) -next - assume h: "x = inverse m *s y + - (inverse m *s c)" - show "m *s x + c = y" unfolding h diff_minus[symmetric] - using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib) -qed - -lemma vector_eq_affinity: - "(m::'a::field) \ 0 ==> (y = m *s x + c \ inverse(m) *s y + -(inverse(m) *s c) = x)" - using vector_affinity_eq[where m=m and x=x and y=y and c=c] - by metis - -lemma image_affinity_interval: fixes m::real - fixes a b c :: "real^'n::finite" - shows "(\x. m *\<^sub>R x + c) ` {a .. b} = - (if {a .. b} = {} then {} - 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 *\<^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 *\<^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 *\<^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 *\<^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) *\<^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 *\<^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) *\<^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 *\<^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...) *} - -lemma banach_fix: - assumes s:"complete s" "s \ {}" and c:"0 \ c" "c < 1" and f:"(f ` s) \ s" and - lipschitz:"\x\s. \y\s. dist (f x) (f y) \ c * dist x y" - shows "\! x\s. (f x = x)" -proof- - have "1 - c > 0" using c by auto - - from s(2) obtain z0 where "z0 \ s" by auto - def z \ "\n. (f ^^ n) z0" - { fix n::nat - have "z n \ s" unfolding z_def - proof(induct n) case 0 thus ?case using `z0 \s` by auto - next case Suc thus ?case using f by auto qed } - note z_in_s = this - - def d \ "dist (z 0) (z 1)" - - have fzn:"\n. f (z n) = z (Suc n)" unfolding z_def by auto - { fix n::nat - have "dist (z n) (z (Suc n)) \ (c ^ n) * d" - proof(induct n) - case 0 thus ?case unfolding d_def by auto - next - case (Suc m) - hence "c * dist (z m) (z (Suc m)) \ c ^ Suc m * d" - using `0 \ c` using mult_mono1_class.mult_mono1[of "dist (z m) (z (Suc m))" "c ^ m * d" c] by auto - thus ?case using lipschitz[THEN bspec[where x="z m"], OF z_in_s, THEN bspec[where x="z (Suc m)"], OF z_in_s] - unfolding fzn and mult_le_cancel_left by auto - qed - } note cf_z = this - - { fix n m::nat - have "(1 - c) * dist (z m) (z (m+n)) \ (c ^ m) * d * (1 - c ^ n)" - proof(induct n) - case 0 show ?case by auto - next - case (Suc k) - have "(1 - c) * dist (z m) (z (m + Suc k)) \ (1 - c) * (dist (z m) (z (m + k)) + dist (z (m + k)) (z (Suc (m + k))))" - using dist_triangle and c by(auto simp add: dist_triangle) - also have "\ \ (1 - c) * (dist (z m) (z (m + k)) + c ^ (m + k) * d)" - using cf_z[of "m + k"] and c by auto - also have "\ \ c ^ m * d * (1 - c ^ k) + (1 - c) * c ^ (m + k) * d" - using Suc by (auto simp add: ring_simps) - also have "\ = (c ^ m) * (d * (1 - c ^ k) + (1 - c) * c ^ k * d)" - unfolding power_add by (auto simp add: ring_simps) - also have "\ \ (c ^ m) * d * (1 - c ^ Suc k)" - using c by (auto simp add: ring_simps) - finally show ?case by auto - qed - } note cf_z2 = this - { fix e::real assume "e>0" - hence "\N. \m n. N \ m \ N \ n \ dist (z m) (z n) < e" - proof(cases "d = 0") - case True - hence "\n. z n = z0" using cf_z2[of 0] and c unfolding z_def by (auto simp add: pos_prod_le[OF `1 - c > 0`]) - thus ?thesis using `e>0` by auto - next - case False hence "d>0" unfolding d_def using zero_le_dist[of "z 0" "z 1"] - by (metis False d_def real_less_def) - hence "0 < e * (1 - c) / d" using `e>0` and `1-c>0` - using divide_pos_pos[of "e * (1 - c)" d] and mult_pos_pos[of e "1 - c"] by auto - then obtain N where N:"c ^ N < e * (1 - c) / d" using real_arch_pow_inv[of "e * (1 - c) / d" c] and c by auto - { fix m n::nat assume "m>n" and as:"m\N" "n\N" - have *:"c ^ n \ c ^ N" using `n\N` and c using power_decreasing[OF `n\N`, of c] by auto - have "1 - c ^ (m - n) > 0" using c and power_strict_mono[of c 1 "m - n"] using `m>n` by auto - hence **:"d * (1 - c ^ (m - n)) / (1 - c) > 0" - using real_mult_order[OF `d>0`, of "1 - c ^ (m - n)"] - using divide_pos_pos[of "d * (1 - c ^ (m - n))" "1 - c"] - using `0 < 1 - c` by auto - - have "dist (z m) (z n) \ c ^ n * d * (1 - c ^ (m - n)) / (1 - c)" - using cf_z2[of n "m - n"] and `m>n` unfolding pos_le_divide_eq[OF `1-c>0`] - by (auto simp add: real_mult_commute dist_commute) - also have "\ \ c ^ N * d * (1 - c ^ (m - n)) / (1 - c)" - using mult_right_mono[OF * order_less_imp_le[OF **]] - unfolding real_mult_assoc by auto - also have "\ < (e * (1 - c) / d) * d * (1 - c ^ (m - n)) / (1 - c)" - using mult_strict_right_mono[OF N **] unfolding real_mult_assoc by auto - also have "\ = e * (1 - c ^ (m - n))" using c and `d>0` and `1 - c > 0` by auto - also have "\ \ e" using c and `1 - c ^ (m - n) > 0` and `e>0` using mult_right_le_one_le[of e "1 - c ^ (m - n)"] by auto - finally have "dist (z m) (z n) < e" by auto - } note * = this - { fix m n::nat assume as:"N\m" "N\n" - hence "dist (z n) (z m) < e" - proof(cases "n = m") - case True thus ?thesis using `e>0` by auto - next - case False thus ?thesis using as and *[of n m] *[of m n] unfolding nat_neq_iff by (auto simp add: dist_commute) - qed } - thus ?thesis by auto - qed - } - hence "Cauchy z" unfolding cauchy_def by auto - then obtain x where "x\s" and x:"(z ---> x) sequentially" using s(1)[unfolded compact_def complete_def, THEN spec[where x=z]] and z_in_s by auto - - def e \ "dist (f x) x" - have "e = 0" proof(rule ccontr) - assume "e \ 0" hence "e>0" unfolding e_def using zero_le_dist[of "f x" x] - by (metis dist_eq_0_iff dist_nz e_def) - then obtain N where N:"\n\N. dist (z n) x < e / 2" - using x[unfolded Lim_sequentially, THEN spec[where x="e/2"]] by auto - hence N':"dist (z N) x < e / 2" by auto - - have *:"c * dist (z N) x \ dist (z N) x" unfolding mult_le_cancel_right2 - using zero_le_dist[of "z N" x] and c - by (metis dist_eq_0_iff dist_nz order_less_asym real_less_def) - have "dist (f (z N)) (f x) \ c * dist (z N) x" using lipschitz[THEN bspec[where x="z N"], THEN bspec[where x=x]] - using z_in_s[of N] `x\s` using c by auto - also have "\ < e / 2" using N' and c using * by auto - finally show False unfolding fzn - using N[THEN spec[where x="Suc N"]] and dist_triangle_half_r[of "z (Suc N)" "f x" e x] - unfolding e_def by auto - qed - hence "f x = x" unfolding e_def by auto - moreover - { fix y assume "f y = y" "y\s" - hence "dist x y \ c * dist x y" using lipschitz[THEN bspec[where x=x], THEN bspec[where x=y]] - using `x\s` and `f x = x` by auto - hence "dist x y = 0" unfolding mult_le_cancel_right1 - using c and zero_le_dist[of x y] by auto - hence "y = x" by auto - } - ultimately show ?thesis unfolding Bex1_def using `x\s` by blast+ -qed - -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\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 - { fix y assume "y\s" - hence "x = y" using `x\s` and dist[THEN bspec[where x=x], THEN bspec[where x=y]] - unfolding g[THEN bspec[where x=x], OF `x\s`] - unfolding g[THEN bspec[where x=y], OF `y\s`] by auto } - thus ?thesis unfolding Bex1_def using `x\s` and g by blast+ -next - case True - then obtain x where [simp]:"x\s" and "g x \ x" by auto - { fix x y assume "x \ s" "y \ s" - hence "dist (g x) (g y) \ dist x y" - using dist[THEN bspec[where x=x], THEN bspec[where x=y]] by auto } note dist' = this - def y \ "g x" - have [simp]:"y\s" unfolding y_def using gs[unfolded image_subset_iff] and `x\s` by blast - def f \ "\n. g ^^ n" - have [simp]:"\n z. g (f n z) = f (Suc n) z" unfolding f_def by auto - have [simp]:"\z. f 0 z = z" unfolding f_def by auto - { fix n::nat and z assume "z\s" - have "f n z \ s" unfolding f_def - proof(induct n) - case 0 thus ?case using `z\s` by simp - next - case (Suc n) thus ?case using gs[unfolded image_subset_iff] by auto - qed } note fs = this - { fix m n ::nat assume "m\n" - fix w z assume "w\s" "z\s" - have "dist (f n w) (f n z) \ dist (f m w) (f m z)" using `m\n` - proof(induct n) - case 0 thus ?case by auto - next - case (Suc n) - thus ?case proof(cases "m\n") - case True thus ?thesis using Suc(1) - using dist'[OF fs fs, OF `w\s` `z\s`, of n n] by auto - next - case False hence mn:"m = Suc n" using Suc(2) by simp - show ?thesis unfolding mn by auto - qed - qed } note distf = this - - 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 \ "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 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::'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 - - { assume as:"dist a b > dist (f n x) (f n y)" - then obtain Na Nb where "\m\Na. dist (f (r m) x) a < (dist a b - dist (f n x) (f n y)) / 2" - and "\m\Nb. dist (f (r m) y) b < (dist a b - dist (f n x) (f n y)) / 2" - using lima limb unfolding h_def Lim_sequentially by (fastsimp simp del: less_divide_eq_number_of1) - hence "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)" - apply(erule_tac x="Na+Nb+n" in allE) - apply(erule_tac x="Na+Nb+n" in allE) apply simp - using dist_triangle_add_half[of a "f (r (Na + Nb + n)) x" "dist a b - dist (f n x) (f n y)" - "-b" "- f (r (Na + Nb + n)) y"] - unfolding ** unfolding group_simps(12) by (auto simp add: dist_commute) - 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 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 - } - hence "dist a b \ dist (f n x) (f n y)" by(rule ccontr)auto } - note ab_fn = this - - have [simp]:"a = b" proof(rule ccontr) - def e \ "dist a b - dist (g a) (g b)" - assume "a\b" hence "e > 0" unfolding e_def using dist by fastsimp - hence "\n. dist (f n x) a < e/2 \ dist (f n y) b < e/2" - using lima limb unfolding Lim_sequentially - apply (auto elim!: allE[where x="e/2"]) apply(rule_tac x="r (max N Na)" in exI) unfolding h_def by fastsimp - then obtain n where n:"dist (f n x) a < e/2 \ dist (f n y) b < e/2" by auto - have "dist (f (Suc n) x) (g a) \ dist (f n x) a" - using dist[THEN bspec[where x="f n x"], THEN bspec[where x="a"]] and fs by auto - moreover have "dist (f (Suc n) y) (g b) \ dist (f n y) b" - using dist[THEN bspec[where x="f n y"], THEN bspec[where x="b"]] and fs by auto - ultimately have "dist (f (Suc n) x) (g a) + dist (f (Suc n) y) (g b) < e" using n by auto - thus False unfolding e_def using ab_fn[of "Suc n"] by norm - qed - - have [simp]:"\n. f (Suc n) x = f n y" unfolding f_def y_def by(induct_tac n)auto - { fix x y assume "x\s" "y\s" moreover - fix e::real assume "e>0" ultimately - 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 "((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 - moreover - { fix x assume "x\s" "g x = x" "x\a" - hence "False" using dist[THEN bspec[where x=a], THEN bspec[where x=x]] - using `g a = a` and `a\s` by auto } - ultimately show "\!x\s. g x = x" unfolding Bex1_def using `a\s` by blast -qed - -end diff -r 1fad3160d873 -r 2083bde13ce1 src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Fri Oct 23 13:23:18 2009 +0200 @@ -0,0 +1,3371 @@ +(* Title: HOL/Library/Convex_Euclidean_Space.thy + Author: Robert Himmelmann, TU Muenchen +*) + +header {* Convex sets, functions and related things. *} + +theory Convex_Euclidean_Space +imports Topology_Euclidean_Space +begin + + +(* ------------------------------------------------------------------------- *) +(* To be moved elsewhere *) +(* ------------------------------------------------------------------------- *) + +declare vector_add_ldistrib[simp] vector_ssub_ldistrib[simp] vector_smult_assoc[simp] vector_smult_rneg[simp] +declare vector_sadd_rdistrib[simp] vector_sub_rdistrib[simp] +declare dot_ladd[simp] dot_radd[simp] dot_lsub[simp] dot_rsub[simp] +declare dot_lmult[simp] dot_rmult[simp] dot_lneg[simp] dot_rneg[simp] +declare UNIV_1[simp] + +term "(x::real^'n \ real) 0" + +lemma dim1in[intro]:"Suc 0 \ {1::nat .. CARD(1)}" by auto + +lemmas vector_component_simps = vector_minus_component vector_smult_component vector_add_component vector_less_eq_def Cart_lambda_beta dest_vec1_def basis_component vector_uminus_component + +lemmas continuous_intros = continuous_add continuous_vmul continuous_cmul continuous_const continuous_sub continuous_at_id continuous_within_id + +lemmas continuous_on_intros = continuous_on_add continuous_on_const continuous_on_id continuous_on_compose continuous_on_cmul continuous_on_neg continuous_on_sub + uniformly_continuous_on_add uniformly_continuous_on_const uniformly_continuous_on_id uniformly_continuous_on_compose uniformly_continuous_on_cmul uniformly_continuous_on_neg uniformly_continuous_on_sub + +lemma dest_vec1_simps[simp]: fixes a::"real^1" + shows "a$1 = 0 \ a = 0" (*"a \ 1 \ dest_vec1 a \ 1" "0 \ a \ 0 \ dest_vec1 a"*) + "a \ b \ dest_vec1 a \ dest_vec1 b" "dest_vec1 (1::real^1) = 1" + by(auto simp add:vector_component_simps all_1 Cart_eq) + +lemma nequals0I:"x\A \ A \ {}" by auto + +lemma norm_not_0:"(x::real^'n::finite)\0 \ norm x \ 0" by auto + +lemma setsum_delta_notmem: assumes "x\s" + shows "setsum (\y. if (y = x) then P x else Q y) s = setsum Q s" + "setsum (\y. if (x = y) then P x else Q y) s = setsum Q s" + "setsum (\y. if (y = x) then P y else Q y) s = setsum Q s" + "setsum (\y. if (x = y) then P y else Q y) s = setsum Q s" + apply(rule_tac [!] setsum_cong2) using assms by auto + +lemma setsum_delta'': + fixes s::"'a::real_vector set" assumes "finite s" + 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)) *\<^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)) *\<^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 *\<^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: + "dest_vec1 ` {a .. b} = {dest_vec1 a .. dest_vec1 b}" + "dest_vec1 ` {a<.. b} = {dest_vec1 a<.. dest_vec1 b}" + "dest_vec1 ` {a ..x. dest_vec1 (f x)) S" + using dest_vec1_sum[OF assms] by auto + +lemma dist_triangle_eq: + 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 smult_conv_scaleR *] + by(auto simp add:norm_minus_commute) qed + +lemma norm_eqI:"x = y \ norm x = norm y" by auto +lemma norm_minus_eqI:"(x::real^'n::finite) = - y \ norm x = norm y" by auto + +lemma Min_grI: assumes "finite A" "A \ {}" "\a\A. x < a" shows "x < Min A" + unfolding Min_gr_iff[OF assms(1,2)] using assms(3) by auto + +lemma dimindex_ge_1:"CARD(_::finite) \ 1" + using one_le_card_finite by auto + +lemma real_dimindex_ge_1:"real (CARD('n::finite)) \ 1" + by(metis dimindex_ge_1 linorder_not_less real_eq_of_nat real_le_trans real_of_nat_1 real_of_nat_le_iff) + +lemma real_dimindex_gt_0:"real (CARD('n::finite)) > 0" apply(rule less_le_trans[OF _ real_dimindex_ge_1]) by auto + +subsection {* Affine set and affine hull.*} + +definition + affine :: "'a::real_vector set \ bool" where + "affine s \ (\x\s. \y\s. \u v. 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 *\<^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 + +lemma affine_empty[intro]: "affine {}" + unfolding affine_def by auto + +lemma affine_sing[intro]: "affine {x}" + unfolding affine_alt by (auto simp add: scaleR_left_distrib [symmetric]) + +lemma affine_UNIV[intro]: "affine UNIV" + unfolding affine_def by auto + +lemma affine_Inter: "(\s\f. affine s) \ affine (\ f)" + unfolding affine_def by auto + +lemma affine_Int: "affine s \ affine t \ affine (s \ t)" + unfolding affine_def by auto + +lemma affine_affine_hull: "affine(affine hull s)" + unfolding hull_def using affine_Inter[of "{t \ affine. s \ t}"] + unfolding mem_def by auto + +lemma affine_hull_eq[simp]: "(affine hull s = s) \ affine s" +proof- + { fix f assume "f \ affine" + hence "affine (\f)" using affine_Inter[of f] unfolding subset_eq mem_def by auto } + thus ?thesis using hull_eq[unfolded mem_def, of affine s] by auto +qed + +lemma setsum_restrict_set'': assumes "finite A" + shows "setsum f {x \ A. P x} = (\x\A. if P x then f x else 0)" + unfolding mem_def[of _ P, symmetric] unfolding setsum_restrict_set'[OF assms] .. + +subsection {* Some explicit formulations (from Lars Schewe). *} + +lemma affine: fixes V::"'a::real_vector 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 *\<^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: scaleR_left_distrib[THEN sym]) +next + fix s u assume as:"\x\V. \y\V. \u v. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ V" + "finite s" "s \ {}" "s \ V" "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 *\<^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::"'a set" and u::"'a \ real" + assume IA:"\u s. \2 < card s; \x\V. \y\V. \u v. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ V; finite s; + s \ {}; s \ V; 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 + thus False using as(7) and `card s > 2` by (metis Numeral1_eq1_nat less_0_number_of less_int_code(15) + less_nat_number_of not_less_iff_gr_or_eq of_nat_1 of_nat_eq_iff pos2 rel_simps(4)) qed + then obtain x where x:"x\s" "u x \ 1" by auto + + have c:"card (s - {x}) = card s - 1" apply(rule card_Diff_singleton) using `x\s` as(4) by auto + have *:"s = insert x (s - {x})" "finite (s - {x})" using `x\s` and as(4) by auto + have **:"setsum u (s - {x}) = 1 - u x" + using setsum_clauses(2)[OF *(2), of u x, unfolded *(1)[THEN sym] as(7)] by auto + have ***:"inverse (1 - u x) * setsum u (s - {x}) = 1" unfolding ** using `u x \ 1` by auto + have "(\xa\s - {x}. (inverse (1 - u x) * u xa) *\<^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 + thus ?thesis apply(rule_tac IA[of "s - {x}" "\y. (inverse (1 - u x) * u y)"]) + unfolding setsum_right_distrib[THEN sym] using as and *** and True by auto + next case False hence "card (s - {x}) = Suc (Suc 0)" using as(2) and c by auto + then obtain a b where "(s - {x}) = {a, b}" "a\b" unfolding card_Suc_eq by auto + thus ?thesis using as(3)[THEN bspec[where x=a], THEN bspec[where x=b]] + using *** *(2) and `s \ V` unfolding setsum_right_distrib by(auto simp add: setsum_clauses(2)) qed + thus "(\x\s. u x *\<^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)) *\<^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) + thus ?thesis using as(4,5) by simp + qed(insert `s\{}` `finite s`, auto) +qed + +lemma affine_hull_explicit: + "affine hull p = {y. \s u. finite s \ s \ {} \ s \ p \ setsum u s = 1 \ setsum (\v. (u v) *\<^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 *\<^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 *\<^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 *\<^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 *\<^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 *\<^sub>R v) = u *\<^sub>R x + v *\<^sub>R y" + apply(rule_tac x="sx \ sy" in exI) + apply(rule_tac x="\a. (if a\sx then u * ux a else 0) + (if a\sy then v * uy a else 0)" in exI) + unfolding scaleR_left_distrib 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 *\<^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 *\<^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 *\<^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. *} + +lemma affine_hull_empty[simp]: "affine hull {} = {}" + apply(rule hull_unique) unfolding mem_def by auto + +lemma affine_hull_finite_step: + fixes y :: "'a::real_vector" + shows "(\u. 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 *\<^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 + next + case False thus ?thesis apply(rule_tac x="u a" in exI) using u and `?as` by auto + qed } moreover + { assume ?rhs + then obtain v u where vu:"setsum u s = w - v" "(\x\s. u x *\<^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 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 *\<^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 *\<^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: + fixes a b :: "'a::real_vector" + shows "affine hull {a,b} = {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::'a)" by auto + 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 *\<^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: + fixes a b c :: "'a::real_vector" + shows "affine hull {a,b,c} = { u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c| u v w. u + v + w = 1}" (is "?lhs = ?rhs") +proof- + have *:"\x y z. z = x - y \ y + z = (x::real)" + "\x y z. z = x - y \ y + z = (x::'a)" by auto + show ?thesis apply(simp add: affine_hull_finite affine_hull_finite_step) + unfolding * apply auto + apply(rule_tac x=v in exI) apply(rule_tac x=va in exI) apply auto + apply(rule_tac x=u in exI) by(auto intro!: exI) +qed + +subsection {* Some relations between affine hull and subspaces. *} + +lemma affine_hull_insert_subset_span: + fixes a :: "real ^ _" + shows "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 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 *\<^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 *\<^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: + fixes a :: "real ^ _" + assumes "a \ s" + shows "affine hull (insert a s) = + {a + v | v . v \ span {x - a | x. x \ s}}" + apply(rule, rule affine_hull_insert_subset_span) unfolding subset_eq Ball_def + unfolding affine_hull_explicit and mem_Collect_eq proof(rule,rule,erule exE,erule conjE) + fix y v assume "y = a + v" "v \ span {x - a |x. x \ s}" + then obtain t u where obt:"finite t" "t \ {x - a |x. x \ s}" "a + (\v\t. u v *\<^sub>R v) = y" unfolding span_explicit smult_conv_scaleR by auto + def f \ "(\x. x + a) ` t" + have f:"finite f" "f \ s" "(\v\f. u (v - a) *\<^sub>R (v - a)) = y - a" unfolding f_def using obt + by(auto simp add: 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 *\<^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 scaleR_left.setsum algebra_simps) qed + +lemma affine_hull_span: + fixes a :: "real ^ _" + assumes "a \ s" + shows "affine hull s = {a + v | v. v \ span {x - a | x. x \ s - {a}}}" + using affine_hull_insert_span[of a "s - {a}", unfolded insert_Diff[OF assms]] by auto + +subsection {* Convexity. *} + +definition + convex :: "'a::real_vector set \ bool" where + "convex s \ (\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) + by (auto simp add: *) qed + +lemma mem_convex: + assumes "convex s" "a \ s" "b \ s" "0 \ u" "u \ 1" + 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:scaleR_left_distrib[THEN sym]) + +lemma convex_UNIV[intro]: "convex UNIV" + unfolding convex_def by auto + +lemma convex_Inter: "(\s\f. convex s) ==> convex(\ f)" + unfolding convex_def by auto + +lemma convex_Int: "convex s \ convex t \ convex (s \ t)" + unfolding convex_def by auto + +lemma convex_halfspace_le: "convex {x. inner a x \ b}" + unfolding convex_def apply 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. inner a x = b}" +proof- + 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. 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) + +subsection {* Explicit expressions for convexity in terms of arbitrary sums. *} + +lemma convex: "convex s \ + (\(k::nat) u x. (\i. 1\i \ i\k \ 0 \ u i \ x i \s) \ (setsum u {1..k} = 1) + \ setsum (\i. u i *\<^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 *\<^sub>R x i) \ s" + "x \ s" "y \ s" "0 \ u" "0 \ v" "u + v = (1::real)" + 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 *\<^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 *\<^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) + hence "setsum u {1 .. k} > 0" using ui apply(rule_tac less_le_trans[of _ "u i"]) using Suc(2)[THEN spec[where x=i]] and i(1) by auto + thus False using Suc(3) unfolding setsum_cl_ivl_Suc and True by simp qed + thus ?thesis unfolding setsum_cl_ivl_Suc using True and Suc(2) by auto + next + have *:"setsum u {1..k} = 1 - u (Suc k)" using Suc(3)[unfolded setsum_cl_ivl_Suc] by auto + have **:"u (Suc k) \ 1" apply(rule ccontr) unfolding not_le using Suc(3) using setsum_nonneg[of "{1..k}" u] using Suc(2) by auto + have ***:"\i k. (u i / (1 - u (Suc k))) *\<^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))) *\<^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)) *\<^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 scaleR_right.setsum [symmetric] using nn by auto qed qed auto qed + + +lemma convex_explicit: + fixes s :: "'a::real_vector set" + shows "convex s \ + (\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 *\<^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 *\<^sub>R x + v *\<^sub>R y \ s" "finite (t::'a 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 *\<^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 *\<^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 *\<^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) + hence "setsum u f > 0" using uy apply(rule_tac less_le_trans[of _ "u y"]) using as(4) and y(1) by auto + thus False using as(2,5) unfolding setsum_clauses(2)[OF as(1)] and True by auto qed + thus ?thesis unfolding setsum_clauses(2)[OF as(1)] using as(2,3) unfolding True by auto + next + have *:"setsum u f = setsum u (insert x f) - u x" using as(2) unfolding setsum_clauses(2)[OF as(1)] by auto + have **:"u x \ 1" apply(rule ccontr) unfolding not_le using as(5)[unfolded setsum_clauses(2)[OF as(1)]] and as(2) + using setsum_nonneg[of f u] and as(4) by auto + case False hence "inverse (1 - u x) *\<^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 *\<^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 *\<^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 *\<^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 *\<^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 *\<^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 :: "'a::real_vector set \ bool" where + "cone s \ (\x\s. \c\0. (c *\<^sub>R x) \ s)" + +lemma cone_empty[intro, simp]: "cone {}" + unfolding cone_def by auto + +lemma cone_univ[intro, simp]: "cone UNIV" + unfolding cone_def by auto + +lemma cone_Inter[intro]: "(\s\f. cone s) \ cone(\ f)" + unfolding cone_def by auto + +subsection {* Conic hull. *} + +lemma cone_cone_hull: "cone (cone hull s)" + unfolding hull_def using cone_Inter[of "{t \ conic. s \ t}"] + by (auto simp add: mem_def) + +lemma cone_hull_eq: "(cone hull s = s) \ cone s" + apply(rule hull_eq[unfolded mem_def]) + using cone_Inter unfolding subset_eq by (auto simp add: mem_def) + +subsection {* Affine dependence and consequential theorems (from Lars Schewe). *} + +definition + affine_dependent :: "'a::real_vector 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 *\<^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 *\<^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 *\<^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 *\<^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 *\<^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 scaleR_scaleR[THEN sym] and scaleR_right.setsum [symmetric] unfolding setsum_right_distrib[THEN sym] and setsum_diff1[OF as(1)] using as by auto +qed + +lemma affine_dependent_explicit_finite: + fixes s :: "'a::real_vector set" assumes "finite s" + 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) *\<^sub>R v = (if vt then (u v) *\<^sub>R v else (0::'a))" by auto + assume ?lhs + then obtain t u v where "finite t" "t \ s" "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_absorb1[OF `t\s`] by auto +next + assume ?rhs + 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 + +subsection {* A general lemma. *} + +lemma convex_connected: + fixes s :: "'a::real_normed_vector set" + assumes "convex s" shows "connected s" +proof- + { fix e1 e2 assume as:"open e1" "open e2" "e1 \ e2 \ s = {}" "s \ e1 \ e2" + assume "e1 \ s \ {}" "e2 \ s \ {}" + then obtain x1 x2 where x1:"x1\e1" "x1\s" and x2:"x2\e2" "x2\s" by auto + hence n:"norm (x1 - x2) > 0" unfolding zero_less_norm_iff using as(3) by auto + + { fix x e::real assume as:"0 \ x" "x \ 1" "0 < e" + { fix y have *:"(1 - x) *\<^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) *\<^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) *\<^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) *\<^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) *\<^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 } + thus ?thesis unfolding connected_def by auto +qed + +subsection {* One rather trivial consequence. *} + +lemma connected_UNIV: "connected (UNIV :: 'a::real_normed_vector set)" + by(simp add: convex_connected convex_UNIV) + +subsection {* Convex functions into the reals. *} + +definition + convex_on :: "'a::real_vector set \ ('a \ real) \ bool" where + "convex_on s f \ + (\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 + +lemma convex_add: + assumes "convex_on s f" "convex_on s g" + shows "convex_on s (\x. f x + g x)" +proof- + { fix x y assume "x\s" "y\s" moreover + fix u v ::real assume "0 \ u" "0 \ v" "u + v = 1" + ultimately have "f (u *\<^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 *\<^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 + +lemma convex_cmul: + assumes "0 \ (c::real)" "convex_on s f" + shows "convex_on s (\x. c * f x)" +proof- + have *:"\u c fx v fy ::real. u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" by (simp add: ring_simps) + show ?thesis using assms(2) and mult_mono1[OF _ assms(1)] unfolding convex_on_def and * by auto +qed + +lemma convex_lower: + assumes "convex_on s f" "x\s" "y \ s" "0 \ u" "0 \ v" "u + v = 1" + shows "f (u *\<^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) + using assms(4,5) by(auto simp add: mult_mono1) + also have "\ = max (f x) (f y)" using assms(6) unfolding distrib[THEN sym] by auto + finally show ?thesis using assms(1)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x=u]] + using assms(2-6) by auto +qed + +lemma convex_local_global_minimum: + fixes s :: "'a::real_normed_vector set" + assumes "0 s" "\y\ball x e. f x \ f y" + shows "\y\s. f x \ f y" +proof(rule ccontr) + have "x\s" using assms(1,3) by auto + assume "\ (\y\s. f x \ f y)" + then obtain y where "y\s" and y:"f x > f y" by auto + hence xy:"0 < dist x y" by (auto simp add: dist_nz[THEN sym]) + + then obtain u where "0 < u" "u \ 1" and u:"u < e / dist x y" + using real_lbound_gt_zero[of 1 "e / dist x y"] using xy `e>0` and divide_pos_pos[of e "dist x y"] by auto + hence "f ((1-u) *\<^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) *\<^sub>R 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) *\<^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 + +lemma convex_distance: + fixes s :: "'a::real_normed_vector set" + shows "convex_on s (\x. dist a x)" +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 *\<^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 *\<^sub>R x) ` s)" + unfolding convex_def and image_iff apply auto + 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 *\<^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 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 *\<^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 +qed + +lemma convex_differences: + assumes "convex s" "convex t" + shows "convex {x - y| x y. x \ s \ y \ t}" +proof- + have "{x - y| x y. x \ s \ y \ t} = {x + y |x y. x \ s \ y \ uminus ` t}" unfolding image_iff apply auto + apply(rule_tac x=xa in exI) apply(rule_tac x="-y" in exI) apply simp + apply(rule_tac x=xa in exI) apply(rule_tac x=xb in exI) by simp + thus ?thesis using convex_sums[OF assms(1) convex_negations[OF assms(2)]] by auto +qed + +lemma convex_translation: assumes "convex s" shows "convex ((\x. a + x) ` s)" +proof- have "{a + y |y. y \ s} = (\x. a + x) ` s" by auto + thus ?thesis using convex_sums[OF convex_singleton[of a] assms] by auto qed + +lemma convex_affinity: assumes "convex s" 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:"bounded_linear f" + shows "convex(f ` s)" +proof(auto simp add: convex_def) + interpret f: bounded_linear f by fact + fix x y assume xy:"x \ s" "y \ s" + fix u v ::real assume uv:"0 \ u" "0 \ v" "u + v = 1" + 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 f.add f.scaleR + using c[unfolded convex_def] xy uv by auto +qed + +subsection {* Balls, being convex, are connected. *} + +lemma convex_ball: + fixes x :: "'a::real_normed_vector" + shows "convex (ball x e)" +proof(auto simp add: convex_def) + fix y z assume yz:"dist x y < e" "dist x z < e" + fix u v ::real assume uv:"0 \ u" "0 \ v" "u + v = 1" + have "dist x (u *\<^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 *\<^sub>R y + v *\<^sub>R z) < e" using real_convex_bound_lt[OF yz uv] by auto +qed + +lemma convex_cball: + fixes x :: "'a::real_normed_vector" + shows "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 *\<^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 *\<^sub>R y + v *\<^sub>R z) \ e" using real_convex_bound_le[OF yz uv] by auto +qed + +lemma connected_ball: + fixes x :: "'a::real_normed_vector" + shows "connected (ball x e)" + using convex_connected convex_ball by auto + +lemma connected_cball: + fixes x :: "'a::real_normed_vector" + shows "connected(cball x e)" + using convex_connected convex_cball by auto + +subsection {* Convex hull. *} + +lemma convex_convex_hull: "convex(convex hull s)" + unfolding hull_def using convex_Inter[of "{t\convex. s\t}"] + unfolding mem_def by auto + +lemma convex_hull_eq: "(convex hull s = s) \ convex s" apply(rule hull_eq[unfolded mem_def]) + using convex_Inter[unfolded Ball_def mem_def] by auto + +lemma bounded_convex_hull: + fixes s :: "'a::real_normed_vector set" + assumes "bounded s" shows "bounded(convex hull s)" +proof- from assms obtain B where B:"\x\s. norm x \ B" unfolding bounded_iff by auto + show ?thesis apply(rule bounded_subset[OF bounded_cball, of _ 0 B]) + unfolding subset_hull[unfolded mem_def, of convex, OF convex_cball] + unfolding subset_eq mem_cball dist_norm using B by auto qed + +lemma finite_imp_bounded_convex_hull: + fixes s :: "'a::real_normed_vector set" + shows "finite s \ bounded(convex hull s)" + using bounded_convex_hull finite_imp_bounded by auto + +subsection {* Stepping theorems for convex hulls of finite sets. *} + +lemma convex_hull_empty[simp]: "convex hull {} = {}" + apply(rule hull_unique) unfolding mem_def by auto + +lemma convex_hull_singleton[simp]: "convex hull {a} = {a}" + apply(rule hull_unique) unfolding mem_def by auto + +lemma convex_hull_insert: + fixes s :: "'a::real_vector set" + assumes "s \ {}" + shows "convex hull (insert a s) = {x. \u\0. \v\0. \b. (u + v = 1) \ + b \ (convex hull s) \ (x = u *\<^sub>R a + v *\<^sub>R b)}" (is "?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 *\<^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] + apply(erule_tac x=a in ballE) apply(erule_tac x=b in ballE) apply(erule_tac x=u in allE) using obt by auto +next + show "convex ?hull" unfolding convex_def apply(rule,rule,rule,rule,rule,rule,rule) proof- + fix x y u v assume as:"(0::real) \ u" "0 \ v" "u + v = 1" "x\?hull" "y\?hull" + from as(4) obtain u1 v1 b1 where obt1:"u1\0" "v1\0" "u1 + v1 = 1" "b1 \ convex hull s" "x = u1 *\<^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::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" by (auto simp add: algebra_simps) + have "\b \ convex hull s. u *\<^sub>R x + v *\<^sub>R y = (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::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" by (auto simp add: algebra_simps) + 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: ** 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) + also have "\ = u * v1 + v * v2" by simp finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" by auto + case False have "0 \ u * v1 + v * v2" "0 \ u * v1" "0 \ u * v1 + v * v2" "0 \ v * v2" apply - + apply(rule add_nonneg_nonneg) prefer 4 apply(rule add_nonneg_nonneg) apply(rule_tac [!] mult_nonneg_nonneg) + using as(1,2) obt1(1,2) obt2(1,2) by auto + thus ?thesis unfolding obt1(5) obt2(5) unfolding * and ** using False + apply(rule_tac x="((u * v1) / (u * v1 + v * v2)) *\<^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 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 + have "u1 * u + u2 * v \ (max u1 u2) * u + (max u1 u2) * v" apply(rule add_mono) + apply(rule_tac [!] mult_right_mono) using as(1,2) obt1(1,2) obt2(1,2) by auto + also have "\ \ 1" unfolding mult.add_right[THEN sym] and as(3) using u1 u2 by auto + finally + show "u *\<^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: algebra_simps) + qed +qed + + +subsection {* Explicit expression for convex hull. *} + +lemma convex_hull_indexed: + fixes s :: "'a::real_vector set" + shows "convex hull s = {y. \k u x. (\i\{1::nat .. k}. 0 \ u i \ x i \ s) \ + (setsum u {1..k} = 1) \ + (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- + fix x assume "x\s" + thus "x \ ?hull" unfolding mem_Collect_eq apply(rule_tac x=1 in exI, rule_tac x="\x. 1" in exI) by auto +next + fix t assume as:"s \ t" "convex t" + show "?hull \ t" apply(rule) unfolding mem_Collect_eq apply(erule exE | erule conjE)+ proof- + fix x k u y assume assm:"\i\{1::nat..k}. 0 \ u i \ y i \ s" "setsum u {1..k} = 1" "(\i = 1..k. u i *\<^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 *\<^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::'a) x2 s1 s2 i.(if P i then s1 else s2) *\<^sub>R (if P i then x1 else x2) = (if P i then s1 *\<^sub>R x1 else s2 *\<^sub>R x2)" + "{1..k1 + k2} \ {1..k1} = {1..k1}" "{1..k1 + k2} \ - {1..k1} = (\i. i + k1) ` {1..k2}" + prefer 3 apply(rule,rule) unfolding image_iff apply(rule_tac x="x - k1" in bexI) by(auto simp add: not_le) + have inj:"inj_on (\i. i + k1) {1..k2}" unfolding inj_on_def by auto + 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 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}") + case True thus ?thesis using mult_nonneg_nonneg[of u "u1 i"] and uv(1) x(1)[THEN bspec[where x=i]] by auto + next def j \ "i - k1" + case False with i have "j \ {1..k2}" unfolding j_def by auto + thus ?thesis unfolding j_def[symmetric] using False + using mult_nonneg_nonneg[of v "u2 j"] and uv(2) y(1)[THEN bspec[where x=j]] by auto qed + qed(auto simp add: not_le x(2,3) y(2,3) uv(3)) +qed + +lemma convex_hull_finite: + fixes s :: "'a::real_vector set" + assumes "finite s" + shows "convex hull s = {y. \u. (\x\s. 0 \ u x) \ + 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 *\<^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 + fix u v ::real assume uv:"0 \ u" "0 \ v" "u + v = 1" + fix ux assume ux:"\x\s. 0 \ ux x" "setsum ux s = (1::real)" + fix uy assume uy:"\x\s. 0 \ uy x" "setsum uy s = (1::real)" + { fix x assume "x\s" + hence "0 \ u * ux x + v * uy x" using ux(1)[THEN bspec[where x=x]] uy(1)[THEN bspec[where x=x]] and uv(1,2) + by (auto, metis add_nonneg_nonneg mult_nonneg_nonneg uv(1) uv(2)) } + moreover have "(\x\s. u * ux x + v * uy x) = 1" + unfolding setsum_addf and setsum_right_distrib[THEN sym] and ux(2) uy(2) using uv(3) by auto + moreover have "(\x\s. (u * ux x + v * uy x) *\<^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 *\<^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 + +subsection {* Another formulation from Lars Schewe. *} + +lemma setsum_constant_scaleR: + fixes y :: "'a::real_vector" + shows "(\x\A. y) = of_nat (card A) *\<^sub>R y" +apply (cases "finite A") +apply (induct set: finite) +apply (simp_all add: algebra_simps) +done + +lemma convex_hull_explicit: + fixes p :: "'a::real_vector set" + shows "convex hull p = {y. \s u. finite s \ s \ p \ + (\x\s. 0 \ u x) \ 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 *\<^sub>R y i) = x" + unfolding convex_hull_indexed by auto + + have fin:"finite {1..k}" by auto + have fin':"\v. finite {i \ {1..k}. y i = v}" by auto + { fix j assume "j\{1..k}" + hence "y j \ p" "0 \ setsum u {i. Suc 0 \ i \ i \ k \ y i = y j}" + using obt(1)[THEN bspec[where x=j]] and obt(2) apply simp + apply(rule setsum_nonneg) using obt(1) by auto } + moreover + have "(\v\y ` {1..k}. setsum u {i \ {1..k}. y i = v}) = 1" + unfolding setsum_image_gen[OF fin, THEN sym] using obt(2) by auto + moreover have "(\v\y ` {1..k}. setsum u {i \ {1..k}. y i = v} *\<^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 *\<^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 + + { fix i::nat assume "i\{1..card s}" + hence "f i \ s" apply(subst f(2)[THEN sym]) by auto + hence "0 \ u (f i)" "f i \ p" using obt(2,3) by auto } + moreover have *:"finite {1..card s}" by auto + { fix y assume "y\s" + then obtain i where "i\{1..card s}" "f i = y" using f using image_iff[of y f "{1..card s}"] by auto + hence "{x. Suc 0 \ x \ x \ card s \ f x = y} = {i}" apply auto using f(1)[unfolded inj_on_def] apply(erule_tac x=x in ballE) by auto + hence "card {x. Suc 0 \ x \ x \ card s \ f x = y} = 1" by auto + hence "(\x\{x \ {1..card s}. f x = y}. u (f x)) = u y" + "(\x\{x \ {1..card s}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y" + by (auto simp add: setsum_constant_scaleR) } + + 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 *\<^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 +qed + +subsection {* A stepping theorem for that expansion. *} + +lemma convex_hull_finite_step: + fixes s :: "'a::real_vector set" assumes "finite s" + 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 *\<^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 *\<^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 *\<^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 + +subsection {* Hence some special cases. *} + +lemma convex_hull_2: + "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 *\<^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 *\<^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,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" + "\x y z ::real^'n. x + y + z = 1 \ x = 1 - y - z" by (auto simp add: ring_simps) + show ?thesis unfolding convex_hull_finite[OF fin(1)] and Collect_def and convex_hull_finite_step[OF fin(2)] and * + unfolding convex_hull_finite_step[OF fin(3)] apply(rule ext) apply simp apply auto + apply(rule_tac x=va in exI) apply (rule_tac x="u c" in exI) apply simp + apply(rule_tac x="1 - v - w" in exI) apply simp apply(rule_tac x=v in exI) apply simp apply(rule_tac x="\x. w" in exI) by simp qed + +lemma convex_hull_3_alt: + "convex hull {a,b,c} = {a + u *\<^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 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. *} + +text {* TODO: Generalize linear algebra concepts defined in @{text +Euclidean_Space.thy} so that we can generalize these lemmas. *} + +lemma subspace_imp_affine: + fixes s :: "(real ^ _) set" shows "subspace s \ affine s" + 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 + +lemma subspace_imp_convex: + fixes s :: "(real ^ _) set" shows "subspace s \ convex s" + using subspace_imp_affine affine_imp_convex by auto + +lemma affine_hull_subset_span: + fixes s :: "(real ^ _) set" shows "(affine hull s) \ (span s)" + unfolding span_def apply(rule hull_antimono) unfolding subset_eq Ball_def mem_def + using subspace_imp_affine by auto + +lemma convex_hull_subset_span: + fixes s :: "(real ^ _) set" shows "(convex hull s) \ (span s)" + unfolding span_def apply(rule hull_antimono) unfolding subset_eq Ball_def mem_def + using subspace_imp_convex by auto + +lemma convex_hull_subset_affine_hull: "(convex hull s) \ (affine hull s)" + unfolding span_def apply(rule hull_antimono) unfolding subset_eq Ball_def mem_def + using affine_imp_convex by auto + +lemma affine_dependent_imp_dependent: + fixes s :: "(real ^ _) set" shows "affine_dependent s \ dependent s" + unfolding affine_dependent_def dependent_def + using affine_hull_subset_span by auto + +lemma dependent_imp_affine_dependent: + fixes s :: "(real ^ _) set" + assumes "dependent {x - a| x . x \ s}" "a \ s" + shows "affine_dependent (insert a s)" +proof- + 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 + have "0\S" using obt(2) assms(2) unfolding subset_eq by auto + have fin:"finite t" and "t\s" unfolding t_def using obt(1,2) by auto + + hence "finite (insert a t)" and "insert a t \ insert a s" by auto + moreover have *:"\P Q. (\x\t. (if x = a then P x else Q x)) = (\x\t. Q x)" + apply(rule setsum_cong2) using `a\s` `t\s` by auto + have "(\x\insert a t. if x = a then - (\x\t. u (x - a)) else u (x - a)) = 0" + unfolding setsum_clauses(2)[OF fin] using `a\s` `t\s` apply auto unfolding * by auto + moreover have "\v\insert a t. (if v = a then - (\x\t. u (x - a)) else u (v - a)) \ 0" + apply(rule_tac x="v + a" in bexI) using obt(3,4) and `0\S` unfolding t_def by auto + moreover have *:"\P Q. (\x\t. (if x = a then P x else Q x) *\<^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)) *\<^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 *\<^sub>R x) \ s)" (is "?lhs = ?rhs") +proof- + { fix x y assume "x\s" "y\s" and ?lhs + 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*\<^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 + +lemma affine_dependent_biggerset: fixes s::"(real^'n::finite) set" + assumes "finite s" "card s \ CARD('n) + 2" + shows "affine_dependent s" +proof- + have "s\{}" using assms by auto then obtain a where "a\s" by auto + have *:"{x - a |x. x \ s - {a}} = (\x. x - a) ` (s - {a})" by auto + have "card {x - a |x. x \ s - {a}} = card (s - {a})" unfolding * + apply(rule card_image) unfolding inj_on_def by auto + also have "\ > CARD('n)" using assms(2) + unfolding card_Diff_singleton[OF assms(1) `a\s`] by auto + finally show ?thesis apply(subst insert_Diff[OF `a\s`, THEN sym]) + apply(rule dependent_imp_affine_dependent) + apply(rule dependent_biggerset) by auto qed + +lemma affine_dependent_biggerset_general: + assumes "finite (s::(real^'n::finite) set)" "card s \ dim s + 2" + shows "affine_dependent s" +proof- + from assms(2) have "s \ {}" by auto + then obtain a where "a\s" by auto + have *:"{x - a |x. x \ s - {a}} = (\x. x - a) ` (s - {a})" by auto + have **:"card {x - a |x. x \ s - {a}} = card (s - {a})" unfolding * + apply(rule card_image) unfolding inj_on_def by auto + have "dim {x - a |x. x \ s - {a}} \ dim s" + apply(rule subset_le_dim) unfolding subset_eq + using `a\s` by (auto simp add:span_superset span_sub) + also have "\ < dim s + 1" by auto + also have "\ \ card (s - {a})" using assms + using card_Diff_singleton[OF assms(1) `a\s`] by auto + finally show ?thesis apply(subst insert_Diff[OF `a\s`, THEN sym]) + apply(rule dependent_imp_affine_dependent) apply(rule dependent_biggerset_general) unfolding ** by auto qed + +subsection {* Caratheodory's theorem. *} + +lemma convex_hull_caratheodory: fixes p::"(real^'n::finite) set" + shows "convex hull p = {y. \s u. finite s \ s \ p \ card s \ CARD('n) + 1 \ + (\x\s. 0 \ u x) \ setsum u s = 1 \ setsum (\v. u v *\<^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 *\<^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 *\<^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 *\<^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) + assume as:"\x\s. 0 \ w x" + hence "setsum w (s - {v}) \ 0" apply(rule_tac setsum_nonneg) by auto + hence "setsum w s > 0" unfolding setsum_diff1'[OF obt(1) `v\s`] + using as[THEN bspec[where x=v]] and `v\s` using `w v \ 0` by auto + thus False using wv(1) by auto + qed hence "i\{}" unfolding i_def by auto + + hence "t \ 0" using Min_ge_iff[of i 0 ] and obt(1) unfolding t_def i_def + using obt(4)[unfolded le_less] apply auto unfolding divide_le_0_iff by auto + have t:"\v\s. u v + t * w v \ 0" proof + fix v assume "v\s" hence v:"0\u v" using obt(4)[THEN bspec[where x=v]] by auto + show"0 \ u v + t * w v" proof(cases "w v < 0") + case False thus ?thesis apply(rule_tac add_nonneg_nonneg) + using v apply simp apply(rule mult_nonneg_nonneg) using `t\0` by auto next + case True hence "t \ u v / (- w v)" using `v\s` + unfolding t_def i_def apply(rule_tac Min_le) using obt(1) by auto + thus ?thesis unfolding real_0_le_add_iff + using pos_le_divide_eq[OF True[unfolded neg_0_less_iff_less[THEN sym]]] by auto + qed qed + + obtain a where "a\s" and "t = (\v. (u v) / (- w v)) a" and "w a < 0" + using Min_in[OF _ `i\{}`] and obt(1) unfolding i_def t_def by auto + hence a:"a\s" "u a + t * w a = 0" by auto + have *:"\f. setsum f (s - {a}) = setsum f s - ((f a)::'a::ring)" unfolding setsum_diff1'[OF obt(1) `a\s`] by auto + have "(\v\s. u v + t * w v) = 1" + unfolding setsum_addf wv(1) setsum_right_distrib[THEN sym] obt(5) by auto + moreover have "(\v\s. u v *\<^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: * 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 *\<^sub>R v) = y" using obt by auto +qed auto + +lemma caratheodory: + "convex hull p = {x::real^'n::finite. \s. finite s \ s \ p \ + card s \ CARD('n) + 1 \ x \ convex hull s}" + unfolding expand_set_eq apply(rule, rule) unfolding mem_Collect_eq proof- + fix x assume "x \ convex hull p" + then obtain s u where "finite s" "s \ p" "card s \ CARD('n) + 1" + "\x\s. 0 \ u x" "setsum u s = 1" "(\v\s. u v *\<^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 +next + fix x assume "\s. finite s \ s \ p \ card s \ CARD('n) + 1 \ x \ convex hull s" + then obtain s where "finite s" "s \ p" "card s \ CARD('n) + 1" "x \ convex hull s" by auto + thus "x \ convex hull p" using hull_mono[OF `s\p`] by auto +qed + +subsection {* Openness and compactness are preserved by convex hull operation. *} + +lemma open_convex_hull: + fixes s :: "'a::real_normed_vector set" + assumes "open s" + shows "open(convex hull s)" + unfolding open_contains_cball convex_hull_explicit unfolding mem_Collect_eq ball_simps(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 *\<^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 *\<^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\{}`] + using b apply simp apply rule apply(erule_tac x=x in ballE) using `t\s` by auto + next fix y assume "y \ cball a (Min i)" + hence y:"norm (a - y) \ Min i" unfolding dist_norm[THEN sym] by auto + { fix x assume "x\t" + hence "Min i \ b x" unfolding i_def apply(rule_tac Min_le) using obt(1) by auto + hence "x + (y - a) \ cball x (b x)" using y unfolding mem_cball dist_norm by auto + moreover from `x\t` have "x\s" using obt(2) by auto + ultimately have "x + (y - a) \ s" using y and b[THEN bspec[where x=x]] unfolding subset_eq by auto } + moreover + have *:"inj_on (\v. v + (y - a)) t" unfolding inj_on_def by auto + have "(\v\(\v. v + (y - a)) ` t. u (v - (y - a))) = 1" + unfolding setsum_reindex[OF *] o_def using obt(4) by auto + moreover have "(\v\(\v. v + (y - a)) ` t. u (v - (y - a)) *\<^sub>R v) = y" + unfolding setsum_reindex[OF *] o_def using obt(4,5) + 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 :: "'a::real_normed_vector set" + assumes "compact s" "compact t" + shows "compact { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ s \ y \ t}" +proof- + let ?X = "{0..1} \ s \ t" + let ?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" + assumes "compact s" shows "compact(convex hull s)" +proof(cases "s={}") + case True thus ?thesis using compact_empty by simp +next + case False then obtain w where "w\s" by auto + show ?thesis unfolding caratheodory[of s] + proof(induct "CARD('n) + 1") + have *:"{x.\sa. finite sa \ sa \ s \ card sa \ 0 \ x \ convex hull sa} = {}" + using compact_empty by (auto simp add: convex_hull_empty) + case 0 thus ?case unfolding * by simp + next + case (Suc n) + show ?case proof(cases "n=0") + case True have "{x. \t. finite t \ t \ s \ card t \ Suc n \ x \ convex hull t} = s" + unfolding expand_set_eq and mem_Collect_eq proof(rule, rule) + fix x assume "\t. finite t \ t \ s \ card t \ Suc n \ x \ convex hull t" + then obtain t where t:"finite t" "t \ s" "card t \ Suc n" "x \ convex hull t" by auto + show "x\s" proof(cases "card t = 0") + case True thus ?thesis using t(4) unfolding card_0_eq[OF t(1)] by(simp add: convex_hull_empty) + next + case False hence "card t = Suc 0" using t(3) `n=0` by auto + then obtain a where "t = {a}" unfolding card_Suc_eq by auto + thus ?thesis using t(2,4) by (simp add: convex_hull_singleton) + qed + next + fix x assume "x\s" + thus "\t. finite t \ t \ s \ card t \ Suc n \ x \ convex hull t" + apply(rule_tac x="{x}" in exI) unfolding convex_hull_singleton by auto + qed thus ?thesis using assms by simp + next + case False have "{x. \t. finite t \ t \ s \ card t \ Suc n \ x \ convex hull t} = + { (1 - u) *\<^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) *\<^sub>R u + c *\<^sub>R v \ + 0 \ c \ c \ 1 \ u \ s \ (\t. finite t \ t \ s \ card t \ n \ v \ convex hull t)" + then obtain u v c t where obt:"x = (1 - c) *\<^sub>R u + c *\<^sub>R v" + "0 \ c \ c \ 1" "u \ s" "finite t" "t \ s" "card t \ n" "v \ convex hull t" by auto + 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" + apply(rule_tac x="insert u t" in exI) by (auto simp add: card_insert_if) + next + fix x assume "\t. finite t \ t \ s \ card t \ Suc n \ x \ convex hull t" + then obtain t where t:"finite t" "t \ s" "card t \ Suc n" "x \ convex hull t" by auto + let ?P = "\u v c. x = (1 - c) *\<^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 + thus ?P apply(rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI) using `w\s` and t + by(auto intro!: exI[where x=t]) + next + case True then obtain a u where au:"t = insert a u" "a\u" apply(drule_tac card_eq_SucD) by auto + show ?P proof(cases "u={}") + case True hence "x=a" using t(4)[unfolded au] by auto + show ?P unfolding `x=a` apply(rule_tac x=a in exI, rule_tac x=a in exI, rule_tac x=1 in exI) + using t and `n\0` unfolding au by(auto intro!: exI[where x="{a}"] simp add: convex_hull_singleton) + next + case False obtain ux vx b where obt:"ux\0" "vx\0" "ux + vx = 1" "b \ convex hull u" "x = ux *\<^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) + using obt and t(1-3) unfolding au and * using card_insert_disjoint[OF _ au(2)] + by(auto intro!: exI[where x=u]) + qed + qed + qed + thus ?thesis using compact_convex_combinations[OF assms Suc] by simp + qed + qed +qed + +lemma finite_imp_compact_convex_hull: + fixes s :: "(real ^ _) set" + shows "finite s \ compact(convex hull s)" + apply(drule finite_imp_compact, drule compact_convex_hull) by assumption + +subsection {* Extremal points of a simplex are some vertices. *} + +lemma dist_increases_online: + fixes a b d :: "'a::real_inner" + assumes "d \ 0" + shows "dist a (b + d) > dist a b \ dist a (b - d) > dist a b" +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 norm_eq_sqrt_inner and real_sqrt_less_iff + by (simp add: algebra_simps inner_commute) +next + 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 norm_eq_sqrt_inner and real_sqrt_less_iff + by (simp add: algebra_simps inner_commute) +qed + +lemma norm_increases_online: + fixes d :: "'a::real_inner" + shows "d \ 0 \ norm(a + d) > norm a \ norm(a - d) > norm a" + using dist_increases_online[of d a 0] unfolding dist_norm by auto + +lemma simplex_furthest_lt: + fixes s::"'a::real_inner set" assumes "finite s" + shows "\x \ (convex hull s). x \ s \ (\y\(convex hull s). norm(x - a) < norm(y - a))" +proof(induct_tac rule: finite_induct[of s]) + fix x s assume as:"finite s" "x\s" "\x\convex hull s. x \ s \ (\y\convex hull s. norm (x - a) < norm (y - a))" + show "\xa\convex hull insert x s. xa \ insert x s \ (\y\convex hull insert x s. norm (xa - a) < norm (y - a))" + proof(rule,rule,cases "s = {}") + case False fix y assume y:"y \ convex hull insert x s" "y \ insert x s" + obtain u v b where obt:"u\0" "v\0" "u + v = 1" "b \ convex hull s" "y = u *\<^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") + case True then obtain z where "z\convex hull s" "norm (y - a) < norm (z - a)" + using as(3)[THEN bspec[where x=y]] and y(2) by auto + thus ?thesis apply(rule_tac x=z in bexI) unfolding convex_hull_insert[OF False] by auto + next + case False show ?thesis using obt(3) proof(cases "u=0", case_tac[!] "v=0") + assume "u=0" "v\0" hence "y = b" using obt by auto + thus ?thesis using False and obt(4) by auto + next + assume "u\0" "v=0" hence "y = x" using obt by auto + thus ?thesis using y(2) by auto + next + assume "u\0" "v\0" + then obtain w where w:"w>0" "wb" proof(rule ccontr) + assume "\ x\b" hence "y=b" unfolding obt(5) + using obt(3) by(auto simp add: scaleR_left_distrib[THEN sym]) + thus False using obt(4) and False by simp qed + 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 *\<^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 *\<^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 + qed + qed auto + qed + qed auto +qed (auto simp add: assms) + +lemma simplex_furthest_le: + fixes s :: "(real ^ _) set" + assumes "finite s" "s \ {}" + shows "\y\s. \x\(convex hull s). norm(x - a) \ norm(y - a)" +proof- + have "convex hull s \ {}" using hull_subset[of s convex] and assms(2) by auto + then obtain x where x:"x\convex hull s" "\y\convex hull s. norm (y - a) \ norm (x - a)" + using distance_attains_sup[OF finite_imp_compact_convex_hull[OF assms(1)], of a] + unfolding dist_commute[of a] unfolding dist_norm by auto + thus ?thesis proof(cases "x\s") + case False then obtain y where "y\convex hull s" "norm (x - a) < norm (y - a)" + using simplex_furthest_lt[OF assms(1), THEN bspec[where x=x]] and x(1) by auto + thus ?thesis using x(2)[THEN bspec[where x=y]] by auto + qed auto +qed + +lemma simplex_furthest_le_exists: + fixes s :: "(real ^ _) set" + shows "finite s \ (\x\(convex hull s). \y\s. norm(x - a) \ norm(y - a))" + using simplex_furthest_le[of s] by (cases "s={}")auto + +lemma simplex_extremal_le: + fixes s :: "(real ^ _) set" + assumes "finite s" "s \ {}" + shows "\u\s. \v\s. \x\convex hull s. \y \ convex hull s. norm(x - y) \ norm(u - v)" +proof- + have "convex hull s \ {}" using hull_subset[of s convex] and assms(2) by auto + then obtain u v where obt:"u\convex hull s" "v\convex hull s" + "\x\convex hull s. \y\convex hull s. norm (x - y) \ norm (u - v)" + using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]] by auto + thus ?thesis proof(cases "u\s \ v\s", erule_tac disjE) + assume "u\s" then obtain y where "y\convex hull s" "norm (u - v) < norm (y - v)" + using simplex_furthest_lt[OF assms(1), THEN bspec[where x=u]] and obt(1) by auto + thus ?thesis using obt(3)[THEN bspec[where x=y], THEN bspec[where x=v]] and obt(2) by auto + next + assume "v\s" then obtain y where "y\convex hull s" "norm (v - u) < norm (y - u)" + using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2) by auto + thus ?thesis using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1) + by (auto simp add: norm_minus_commute) + qed auto +qed + +lemma simplex_extremal_le_exists: + fixes s :: "(real ^ _) set" + shows "finite s \ x \ convex hull s \ y \ convex hull s + \ (\u\s. \v\s. norm(x - y) \ norm(u - v))" + using convex_hull_empty simplex_extremal_le[of s] by(cases "s={}")auto + +subsection {* Closest point of a convex set is unique, with a continuous projection. *} + +definition + closest_point :: "(real ^ 'n::finite) set \ real ^ 'n \ real ^ 'n" where + "closest_point s a = (SOME x. x \ s \ (\y\s. dist a x \ dist a y))" + +lemma closest_point_exists: + assumes "closed s" "s \ {}" + shows "closest_point s a \ s" "\y\s. dist a (closest_point s a) \ dist a y" + unfolding closest_point_def apply(rule_tac[!] someI2_ex) + using distance_attains_inf[OF assms(1,2), of a] by auto + +lemma closest_point_in_set: + "closed s \ s \ {} \ (closest_point s a) \ s" + by(meson closest_point_exists) + +lemma closest_point_le: + "closed s \ x \ s \ dist a (closest_point s a) \ dist a x" + using closest_point_exists[of s] by auto + +lemma closest_point_self: + assumes "x \ s" shows "closest_point s x = x" + unfolding closest_point_def apply(rule some1_equality, rule ex1I[of _ x]) + using assms by auto + +lemma closest_point_refl: + "closed s \ s \ {} \ (closest_point s x = x \ x \ s)" + using closest_point_in_set[of s x] closest_point_self[of x s] by auto + +(* 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 "inner y z > 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 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: + fixes s :: "(real ^ _) set" + assumes "convex s" "closed s" "x \ s" "y \ s" "\z\s. dist a x \ dist a z" + shows "inner (a - x) (y - x) \ 0" +proof(rule ccontr) assume "\ 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: + fixes s :: "(real ^ _) set" + assumes "convex s" "closed s" "x \ s" "y \ s" + "\z\s. dist a x \ dist a z" "\z\s. dist a y \ dist a z" + shows "x = y" using any_closest_point_dot[OF assms(1-4,5)] and any_closest_point_dot[OF assms(1-2,4,3,6)] + unfolding norm_pths(1) and norm_le_square + by (auto simp add: algebra_simps) + +lemma closest_point_unique: + assumes "convex s" "closed s" "x \ s" "\z\s. dist a x \ dist a z" + shows "x = closest_point s a" + using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point s a"] + using closest_point_exists[OF assms(2)] and assms(3) by auto + +lemma closest_point_dot: + assumes "convex s" "closed s" "x \ s" + shows "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 + +lemma closest_point_lt: + assumes "convex s" "closed s" "x \ s" "x \ closest_point s a" + shows "dist a (closest_point s a) < dist a x" + apply(rule ccontr) apply(rule_tac notE[OF assms(4)]) + apply(rule closest_point_unique[OF assms(1-3), of a]) + using closest_point_le[OF assms(2), of _ a] by fastsimp + +lemma closest_point_lipschitz: + assumes "convex s" "closed s" "s \ {}" + shows "dist (closest_point s x) (closest_point s y) \ dist x y" +proof- + have "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 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 \ {}" + shows "continuous (at x) (closest_point s)" + unfolding continuous_at_eps_delta + using le_less_trans[OF closest_point_lipschitz[OF assms]] by auto + +lemma continuous_on_closest_point: + assumes "convex s" "closed s" "s \ {}" + shows "continuous_on t (closest_point s)" + apply(rule continuous_at_imp_continuous_on) using continuous_at_closest_point[OF assms] by auto + +subsection {* Various point-to-set separating/supporting hyperplane theorems. *} + +lemma supporting_hyperplane_closed_point: + fixes s :: "(real ^ _) set" + assumes "convex s" "closed s" "s \ {}" "z \ s" + 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="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 "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) *\<^sub>R y + u *\<^sub>R x)" + using assms(1)[unfolded convex_alt] and y and `x\s` and `y\s` by auto + assume "\ inner (y - z) y \ inner (y - z) x" then obtain v where + "v>0" "v\1" "dist (y + v *\<^sub>R (x - y)) z < dist y z" using closer_point_lemma[of z y x] 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: + fixes s :: "(real ^ _) set" + assumes "convex s" "closed s" "z \ s" + 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 _ 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="inner (y - z) z + (norm(y - z))\ / 2" in exI) + apply rule defer apply rule proof- + fix x assume "x\s" + 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 algebra_simps) qed + moreover have "0 < norm (y - z) ^ 2" using `y\s` `z\s` by auto + 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. 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 + thus ?thesis apply(rule_tac x="basis a" in exI, rule_tac x=1 in exI) using True by auto +next case False thus ?thesis using False using separating_hyperplane_closed_point[OF assms] + apply - apply(erule exE)+ unfolding dot_rzero apply(rule_tac x=a in exI, rule_tac x=b in exI) by auto qed + +subsection {* Now set-to-set for closed/compact sets. *} + +lemma separating_hyperplane_closed_compact: + assumes "convex (s::(real^'n::finite) set)" "closed s" "convex t" "compact t" "t \ {}" "s \ t = {}" + shows "\a b. (\x\s. 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:"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 < 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 + 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 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. 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 \ 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 < inner a x" using `0 < b` by auto + qed +qed + +lemma separating_hyperplane_compact_closed: + fixes s :: "(real ^ _) set" + assumes "convex s" "compact s" "s \ {}" "convex t" "closed t" "s \ t = {}" + 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 + +subsection {* General case without assuming closure and getting non-strict separation. *} + +lemma separating_hyperplane_set_0: + assumes "convex s" "(0::real^'n::finite) \ s" + shows "\a. a \ 0 \ (\x\s. 0 \ 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 < 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 \ 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: 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: 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. inner a x \ b) \ (\x\t. inner a x \ b)" +proof- from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]] + obtain a where "a\0" "\x\{x - y |x y. x \ t \ y \ s}. 0 \ inner a x" 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 + +subsection {* More convexity generalities. *} + +lemma convex_closure: + fixes s :: "'a::real_normed_vector set" + 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 *\<^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 + +lemma convex_interior: + fixes s :: "'a::real_normed_vector set" + assumes "convex s" shows "convex(interior s)" + unfolding convex_alt Ball_def mem_interior apply(rule,rule,rule,rule,rule,rule) apply(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) *\<^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) *\<^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: 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 + +subsection {* Moving and scaling convex hulls. *} + +lemma convex_hull_translation_lemma: + "convex hull ((\x. a + x) ` s) \ (\x. a + x) ` (convex hull s)" + apply(rule hull_minimal, rule image_mono, rule hull_subset) unfolding mem_def + using convex_translation[OF convex_convex_hull, of a s] by assumption + +lemma convex_hull_bilemma: fixes neg + assumes "(\s a. (convex hull (up a s)) \ up a (convex hull s))" + shows "(\s. up a (up (neg a) s) = s) \ (\s. up (neg a) (up a s) = s) \ (\s t a. s \ t \ up a s \ up a t) + \ \s. (convex hull (up a s)) = up a (convex hull s)" + using assms by(metis subset_antisym) + +lemma convex_hull_translation: + "convex hull ((\x. a + x) ` s) = (\x. a + x) ` (convex hull s)" + apply(rule convex_hull_bilemma[rule_format, of _ _ "\a. -a"], rule convex_hull_translation_lemma) unfolding image_image by auto + +lemma convex_hull_scaling_lemma: + "(convex hull ((\x. c *\<^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 *\<^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 scaleR_scaleR by(auto simp add:image_constant_conv convex_hull_eq_empty) + +lemma convex_hull_affinity: + "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: + fixes s :: "(real ^ _) set" + assumes "closed s" "convex s" + 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. 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 + +subsection {* Radon's theorem (from Lars Schewe). *} + +lemma radon_ex_lemma: + assumes "finite c" "affine_dependent c" + shows "\u. setsum u c = 0 \ (\v\c. u v \ 0) \ setsum (\v. u v *\<^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 scaleR_zero_left + and setsum_restrict_set[OF assms(1), THEN sym] by(auto simp add: Int_absorb1) qed + +lemma radon_s_lemma: + assumes "finite s" "setsum f s = (0::real)" + shows "setsum f {x\s. 0 < f x} = - setsum f {x\s. f x < 0}" +proof- have *:"\x. (if f x < 0 then f x else 0) + (if 0 < f x then f x else 0) = f x" by auto + show ?thesis unfolding real_add_eq_0_iff[THEN sym] and setsum_restrict_set''[OF assms(1)] and setsum_addf[THEN sym] and * + using assms(2) by assumption qed + +lemma radon_v_lemma: + assumes "finite s" "setsum f s = 0" "\x. g x = (0::real) \ f x = (0::real^'n)" + shows "(setsum f {x\s. 0 < g x}) = - setsum f {x\s. g x < 0}" +proof- + have *:"\x. (if 0 < g x then f x else 0) + (if g x < 0 then f x else 0) = f x" using assms(3) by auto + show ?thesis unfolding eq_neg_iff_add_eq_0 and setsum_restrict_set''[OF assms(1)] and setsum_addf[THEN sym] and * + using assms(2) by assumption qed + +lemma radon_partition: + assumes "finite c" "affine_dependent c" + shows "\m p. m \ p = {} \ m \ p = c \ (convex hull m) \ (convex hull p) \ {}" proof- + obtain u v where uv:"setsum u c = 0" "v\c" "u v \ 0" "(\v\c. u v *\<^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})) *\<^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") + case True thus ?thesis using setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] by auto + next + case False hence "setsum u c \ setsum (\x. if x=v then u v else 0) c" apply(rule_tac setsum_mono) by auto + thus ?thesis unfolding setsum_delta[OF assms(1)] using uv(2) and `u v < 0` and uv(1) by auto qed + qed (insert setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] uv(2-3), auto) + + hence *:"setsum u {x\c. u x > 0} > 0" unfolding real_less_def apply(rule_tac conjI, rule_tac setsum_nonneg) by auto + moreover have "setsum u ({x \ c. 0 < u x} \ {x \ c. u x < 0}) = setsum u c" + "(\x\{x \ c. 0 < u x} \ {x \ c. u x < 0}. u x *\<^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 *\<^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 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 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 + +lemma radon: assumes "affine_dependent c" + obtains m p where "m\c" "p\c" "m \ p = {}" "(convex hull m) \ (convex hull p) \ {}" +proof- from assms[unfolded affine_dependent_explicit] guess s .. then guess u .. + hence *:"finite s" "affine_dependent s" and s:"s \ c" unfolding affine_dependent_explicit by auto + from radon_partition[OF *] guess m .. then guess p .. + thus ?thesis apply(rule_tac that[of p m]) using s by auto qed + +subsection {* Helly's theorem. *} + +lemma helly_induct: fixes f::"(real^'n::finite) set set" + assumes "f hassize n" "n \ CARD('n) + 1" + "\s\f. convex s" "\t\f. card t = CARD('n) + 1 \ \ t \ {}" + shows "\ f \ {}" + using assms unfolding hassize_def apply(erule_tac conjE) proof(induct n arbitrary: f) +case (Suc n) +show "\ f \ {}" apply(cases "n = CARD('n)") apply(rule Suc(4)[rule_format]) + unfolding card_Diff_singleton_if[OF Suc(5)] and Suc(6) proof- + assume ng:"n \ CARD('n)" hence "\X. \s\f. X s \ \(f - {s})" apply(rule_tac bchoice) unfolding ex_in_conv + apply(rule, rule Suc(1)[rule_format]) unfolding card_Diff_singleton_if[OF Suc(5)] and Suc(6) + defer apply(rule Suc(3)[rule_format]) defer apply(rule Suc(4)[rule_format]) using Suc(2,5) by auto + then obtain X where X:"\s\f. X s \ \(f - {s})" by auto + show ?thesis proof(cases "inj_on X f") + case False then obtain s t where st:"s\t" "s\f" "t\f" "X s = X t" unfolding inj_on_def by auto + hence *:"\ f = \ (f - {s}) \ \ (f - {t})" by auto + show ?thesis unfolding * unfolding ex_in_conv[THEN sym] apply(rule_tac x="X s" in exI) + apply(rule, rule X[rule_format]) using X st by auto + next case True then obtain m p where mp:"m \ p = {}" "m \ p = X ` f" "convex hull m \ convex hull p \ {}" + using radon_partition[of "X ` f"] and affine_dependent_biggerset[of "X ` f"] + unfolding card_image[OF True] and Suc(6) using Suc(2,5) and ng by auto + have "m \ X ` f" "p \ X ` f" using mp(2) by auto + then obtain g h where gh:"m = X ` g" "p = X ` h" "g \ f" "h \ f" unfolding subset_image_iff by auto + hence "f \ (g \ h) = f" by auto + hence f:"f = g \ h" using inj_on_Un_image_eq_iff[of X f "g \ h"] and True + unfolding mp(2)[unfolded image_Un[THEN sym] gh] by auto + have *:"g \ h = {}" using mp(1) unfolding gh using inj_on_image_Int[OF True gh(3,4)] by auto + have "convex hull (X ` h) \ \ g" "convex hull (X ` g) \ \ h" + apply(rule_tac [!] hull_minimal) using Suc(3) gh(3-4) unfolding mem_def unfolding subset_eq + apply(rule_tac [2] convex_Inter, rule_tac [4] convex_Inter) apply rule prefer 3 apply rule proof- + fix x assume "x\X ` g" then guess y unfolding image_iff .. + thus "x\\h" using X[THEN bspec[where x=y]] using * f by auto next + fix x assume "x\X ` h" then guess y unfolding image_iff .. + thus "x\\g" using X[THEN bspec[where x=y]] using * f by auto + qed(auto) + thus ?thesis unfolding f using mp(3)[unfolded gh] by blast qed +qed(insert dimindex_ge_1, auto) qed(auto) + +lemma helly: fixes f::"(real^'n::finite) set set" + assumes "finite f" "card f \ CARD('n) + 1" "\s\f. convex s" + "\t\f. card t = CARD('n) + 1 \ \ t \ {}" + shows "\ f \{}" + apply(rule helly_induct) unfolding hassize_def using assms by auto + +subsection {* Convex hull is "preserved" by a linear function. *} + +lemma convex_hull_linear_image: + assumes "bounded_linear f" + shows "f ` (convex hull s) = convex hull (f ` s)" + apply rule unfolding subset_eq ball_simps apply(rule_tac[!] hull_induct, rule hull_inc) prefer 3 + apply(erule imageE)apply(rule_tac x=xa in image_eqI) apply assumption + apply(rule hull_subset[unfolded subset_eq, rule_format]) apply assumption +proof- + interpret f: bounded_linear f by fact + show "convex {x. f x \ convex hull f ` s}" + unfolding convex_def by(auto simp add: f.scaleR f.add convex_convex_hull[unfolded convex_def, rule_format]) next + interpret f: bounded_linear f by fact + show "convex {x. x \ f ` (convex hull s)}" using convex_convex_hull[unfolded convex_def, of s] + unfolding convex_def by (auto simp add: f.scaleR [symmetric] f.add [symmetric]) +qed auto + +lemma in_convex_hull_linear_image: + assumes "bounded_linear f" "x \ convex hull s" + shows "(f x) \ convex hull (f ` s)" +using convex_hull_linear_image[OF assms(1)] assms(2) by auto + +subsection {* Homeomorphism of all convex compact sets with nonempty interior. *} + +lemma compact_frontier_line_lemma: + fixes s :: "(real ^ _) set" + assumes "compact s" "0 \ s" "x \ 0" + 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 *\<^sub>R x)}" + have A:"?A = (\u. dest_vec1 u *\<^sub>R x) ` {0 .. vec1 (b / norm x)}" + unfolding image_image[of "\u. u *\<^sub>R x" "\x. dest_vec1 x", THEN sym] + unfolding dest_vec1_inverval vec1_dest_vec1 by auto + have "compact ?A" unfolding A apply(rule compact_continuous_image, rule continuous_at_imp_continuous_on) + apply(rule, rule continuous_vmul) + apply (rule continuous_dest_vec1) + apply(rule continuous_at_id) by(rule compact_interval) + moreover have "{y. \u\0. u \ b / norm x \ y = u *\<^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 *\<^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 *\<^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`] 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) using `u\0` `norm x > 0` `v>u` by(auto simp add:field_simps) + } note u_max = this + + 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 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 *\<^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) *\<^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) + 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 *\<^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 *\<^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 *\<^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 *\<^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 *\<^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" + apply(rule homeomorphism_compact) apply(rule compact_frontier[OF assms(1)]) + apply(rule continuous_on_subset[OF contpi]) defer apply(rule set_ext,rule) + unfolding inj_on_def prefer 3 apply(rule,rule,rule) + proof- fix x assume "x\pi ` frontier s" then obtain y where "y\frontier s" "x = pi y" by auto + thus "x \ sphere" using pi(1)[of y] and `0 \ frontier s` by auto + next fix x assume "x\sphere" hence "norm x = 1" "x\0" unfolding sphere_def by auto + then obtain u where "0 \ u" "u *\<^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 *\<^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 *\<^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 + using x y and front_smul[THEN bspec, OF as(1), THEN spec[where x="norm y * (inverse (norm x))"]] + using front_smul[THEN bspec, OF as(2), THEN spec[where x="norm x * (inverse (norm y))"]] + using xys nor by(auto simp add:field_simps divide_le_eq_1 divide_inverse[THEN sym]) + thus "x = y" apply(subst injpi[THEN sym]) using as(3) by auto + qed(insert `0 \ frontier s`, auto) + then obtain surf where surf:"\x\frontier s. surf (pi x) = x" "pi ` frontier s = sphere" "continuous_on (frontier s) pi" + "\y\sphere. pi (surf y) = y" "surf ` sphere = frontier s" "continuous_on sphere surf" unfolding homeomorphism_def by auto + + have cont_surfpi:"continuous_on (UNIV - {0}) (surf \ pi)" apply(rule continuous_on_compose, rule contpi) + apply(rule continuous_on_subset[of sphere], rule surf(6)) using pi(1) by auto + + { fix x assume as:"x \ cball (0::real^'n) 1" + have "norm x *\<^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]) + unfolding surf(5)[THEN sym] by auto + next case True thus ?thesis apply rule defer unfolding pi_def apply(rule fs[unfolded subset_eq, rule_format]) + unfolding surf(5)[unfolded sphere_def, THEN sym] using `0\s` by auto qed } note hom = this + + { fix x assume "x\s" + hence "x \ (\x. norm x *\<^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 *\<^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) *\<^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))) *\<^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))) *\<^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 *\<^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 *\<^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 + hence "B > 0" using assms(2) unfolding subset_eq apply(erule_tac x="basis a" in ballE) defer apply(erule_tac x="basis a" in ballE) + unfolding Ball_def mem_cball dist_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 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 + hence "norm x * norm (surf (pi x)) \ norm x * B" using as(2) by auto + also have "\ < e / B * B" apply(rule mult_strict_right_mono) using as(1) `B>0` by auto + also have "\ = e" using `B>0` by auto + finally show "norm x * norm (surf (pi x)) < e" by assumption + qed(insert `B>0`, auto) qed + next { fix x assume as:"surf (pi x) = 0" + have "x = 0" proof(rule ccontr) + assume "x\0" hence "pi x \ sphere" using pi(1) by auto + hence "surf (pi x) \ frontier s" using surf(5) by auto + thus False using `0\frontier s` unfolding as by simp qed + } note surf_0 = this + show "inj_on (\x. norm x *\<^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 + hence "pi (surf (pi x)) = pi (surf (pi y))" using as(3) + using pi(2)[of "norm x" "surf (pi x)"] pi(2)[of "norm y" "surf (pi y)"] by auto + moreover have "pi x \ sphere" "pi y \ sphere" using pi(1) False by auto + ultimately have *:"pi x = pi y" using surf(4)[THEN bspec[where x="pi x"]] surf(4)[THEN bspec[where x="pi y"]] by auto + moreover have "norm x = norm y" using as(3)[unfolded *] using False by(auto dest:surf_0) + ultimately show ?thesis using injpi by auto qed qed + qed auto qed + +lemma homeomorphic_convex_compact_lemma: fixes s::"(real^'n::finite) set" + assumes "convex s" "compact s" "cball 0 1 \ s" + shows "s homeomorphic (cball (0::real^'n) 1)" + apply(rule starlike_compact_projective[OF assms(2-3)]) proof(rule,rule,rule,erule conjE) + fix x u assume as:"x \ s" "0 \ u" "u < (1::real)" + hence "u *\<^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 *\<^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_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) *\<^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" + shows "s homeomorphic (cball (b::real^'n::finite) e)" +proof- obtain a where "a\interior s" using assms(3) by auto + then obtain d where "d>0" and d:"cball a d \ s" unfolding mem_interior_cball by auto + let ?d = "inverse d" and ?n = "0::real^'n" + have "cball ?n 1 \ (\x. inverse d *\<^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 *\<^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 *\<^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 \ {}" + "convex t" "compact t" "interior t \ {}" + shows "s homeomorphic t" + using assms by(meson zero_less_one homeomorphic_trans homeomorphic_convex_compact_cball homeomorphic_sym) + +subsection {* Epigraphs of convex functions. *} + +definition "epigraph s (f::real^'n \ real) = {xy. fstcart xy \ s \ f(fstcart xy) \ dest_vec1 (sndcart xy)}" + +lemma mem_epigraph: "(pastecart x (vec1 y)) \ epigraph s f \ x \ s \ f x \ y" unfolding epigraph_def by auto + +lemma convex_epigraph: + "convex(epigraph s f) \ convex_on s f \ convex s" + unfolding convex_def convex_on_def unfolding Ball_def forall_pastecart epigraph_def + unfolding mem_Collect_eq fstcart_pastecart sndcart_pastecart sndcart_add sndcart_cmul [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" + shows "convex(epigraph s f)" using assms unfolding convex_epigraph by auto + +lemma convex_epigraph_convex: "convex s \ (convex_on s f \ convex(epigraph s f))" + using convex_epigraph by auto + +subsection {* Use this to derive general bound property of convex function. *} + +lemma forall_of_pastecart: + "(\p. P (\x. fstcart (p x)) (\x. sndcart (p x))) \ (\x y. P x y)" apply meson + apply(erule_tac x="\a. pastecart (x a) (y a)" in allE) unfolding o_def by auto + +lemma forall_of_pastecart': + "(\p. P (fstcart p) (sndcart p)) \ (\x y. P x y)" apply meson + apply(erule_tac x="pastecart x y" in allE) unfolding o_def by auto + +lemma forall_of_dest_vec1: "(\v. P (\x. dest_vec1 (v x))) \ (\x. P x)" + apply rule apply rule apply(erule_tac x="(vec1 \ x)" in allE) unfolding o_def vec1_dest_vec1 by auto + +lemma forall_of_dest_vec1': "(\v. P (dest_vec1 v)) \ (\x. P x)" + apply rule apply rule apply(erule_tac x="(vec1 x)" in allE) defer apply rule + apply(erule_tac x="dest_vec1 v" in allE) unfolding o_def vec1_dest_vec1 by auto + +lemma convex_on: + fixes s :: "(real ^ _) set" + 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 *\<^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 [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) + defer apply(rule setsum_mono) apply(erule conjE)+ apply(erule_tac x=i in allE)apply(rule mult_left_mono) + using assms[unfolded convex] by auto + +subsection {* Convexity of general and special intervals. *} + +lemma is_interval_convex: + fixes s :: "(real ^ _) set" + assumes "is_interval s" shows "convex s" + unfolding convex_def apply(rule,rule,rule,rule,rule,rule,rule) proof- + fix x y u v assume as:"x \ s" "y \ s" "0 \ u" "0 \ v" "u + v = (1::real)" + hence *:"u = 1 - v" "1 - v \ 0" and **:"v = 1 - u" "1 - u \ 0" by auto + { fix a b assume "\ b \ u * a + v * b" + hence "u * a < (1 - v) * b" unfolding not_le using as(4) by(auto simp add: field_simps) + hence "a < b" unfolding * using as(4) *(2) apply(rule_tac mult_left_less_imp_less[of "1 - v"]) by(auto simp add: field_simps) + hence "a \ u * a + v * b" unfolding * using as(4) by (auto simp add: field_simps intro!:mult_right_mono) + } moreover + { fix a b assume "\ u * a + v * b \ a" + hence "v * b > (1 - u) * a" unfolding not_le using as(4) by(auto simp add: field_simps) + hence "a < b" unfolding * using as(4) apply(rule_tac mult_left_less_imp_less) by(auto simp add: ring_simps) + hence "u * a + v * b \ b" unfolding ** using **(2) as(3) by(auto simp add: field_simps intro!:mult_right_mono) } + ultimately show "u *\<^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: + fixes s :: "(real ^ _) set" + shows "is_interval s \ connected s" + using is_interval_convex convex_connected by auto + +lemma convex_interval: "convex {a .. b}" "convex {a<.. (\a\s. \b\s. \ x. dest_vec1 a \ dest_vec1 x \ dest_vec1 x \ dest_vec1 b \ x \ s)" + unfolding is_interval_def dest_vec1_def forall_1 by auto + +lemma is_interval_connected_1: "is_interval s \ connected (s::(real^1) set)" + apply(rule, rule is_interval_connected, assumption) unfolding is_interval_1 + apply(rule,rule,rule,rule,erule conjE,rule ccontr) proof- + fix a b x assume as:"connected s" "a \ s" "b \ s" "dest_vec1 a \ dest_vec1 x" "dest_vec1 x \ dest_vec1 b" "x\s" + hence *:"dest_vec1 a < dest_vec1 x" "dest_vec1 x < dest_vec1 b" apply(rule_tac [!] ccontr) unfolding not_less by auto + let ?halfl = "{z. 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: 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) + apply(rule, rule open_halfspace_lt, rule, rule open_halfspace_gt) apply(rule, rule, rule ccontr) + by(auto simp add: basis_component field_simps) qed + +lemma is_interval_convex_1: + "is_interval s \ convex (s::(real^1) set)" + using is_interval_convex convex_connected is_interval_connected_1 by auto + +lemma convex_connected_1: + "connected s \ convex (s::(real^1) set)" + using is_interval_convex convex_connected is_interval_connected_1 by auto + +subsection {* Another intermediate value theorem formulation. *} + +lemma ivt_increasing_component_on_1: fixes f::"real^1 \ real^'n::finite" + assumes "dest_vec1 a \ dest_vec1 b" "continuous_on {a .. b} f" "(f a)$k \ y" "y \ (f b)$k" + shows "\x\{a..b}. (f x)$k = y" +proof- have "f a \ f ` {a..b}" "f b \ f ` {a..b}" apply(rule_tac[!] imageI) + using assms(1) by(auto simp add: vector_less_eq_def dest_vec1_def) + thus ?thesis using connected_ivt_component[of "f ` {a..b}" "f a" "f b" k y] + using connected_continuous_image[OF assms(2) convex_connected[OF convex_interval(1)]] + using assms by(auto intro!: imageI) qed + +lemma ivt_increasing_component_1: fixes f::"real^1 \ real^'n::finite" + assumes "dest_vec1 a \ dest_vec1 b" + "\x\{a .. b}. continuous (at x) f" "f a$k \ y" "y \ f b$k" + shows "\x\{a..b}. (f x)$k = y" + apply(rule ivt_increasing_component_on_1) using assms using continuous_at_imp_continuous_on by auto + +lemma ivt_decreasing_component_on_1: fixes f::"real^1 \ real^'n::finite" + assumes "dest_vec1 a \ dest_vec1 b" "continuous_on {a .. b} f" "(f b)$k \ y" "y \ (f a)$k" + shows "\x\{a..b}. (f x)$k = y" + apply(subst neg_equal_iff_equal[THEN sym]) unfolding vector_uminus_component[THEN sym] + apply(rule ivt_increasing_component_on_1) using assms using continuous_on_neg + by(auto simp add:vector_uminus_component) + +lemma ivt_decreasing_component_1: fixes f::"real^1 \ real^'n::finite" + assumes "dest_vec1 a \ dest_vec1 b" "\x\{a .. b}. continuous (at x) f" "f b$k \ y" "y \ f a$k" + shows "\x\{a..b}. (f x)$k = y" + apply(rule ivt_decreasing_component_on_1) using assms using continuous_at_imp_continuous_on by auto + +subsection {* A bound within a convex hull, and so an interval. *} + +lemma convex_on_convex_hull_bound: + fixes s :: "(real ^ _) set" + 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 *\<^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) + using assms(2) obt(1) by auto + thus "f x \ b" using assms(1)[unfolded convex_on[OF convex_convex_hull], rule_format, of k u v] + unfolding obt(2-3) using obt(1) and hull_subset[unfolded subset_eq, rule_format, of _ s] by auto qed + +lemma unit_interval_convex_hull: + "{0::real^'n::finite .. 1} = convex hull {x. \i. (x$i = 0) \ (x$i = 1)}" (is "?int = convex hull ?points") +proof- have 01:"{0,1} \ convex hull ?points" apply rule apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) by auto + { fix n x assume "x\{0::real^'n .. 1}" "n \ CARD('n)" "card {i. x$i \ 0} \ n" + hence "x\convex hull ?points" proof(induct n arbitrary: x) + case 0 hence "x = 0" apply(subst Cart_eq) apply rule by auto + thus "x\convex hull ?points" using 01 by auto + next + case (Suc n) show "x\convex hull ?points" proof(cases "{i. x$i \ 0} = {}") + case True hence "x = 0" unfolding Cart_eq by auto + thus "x\convex hull ?points" using 01 by auto + next + case False def xi \ "Min ((\i. x$i) ` {i. x$i \ 0})" + have "xi \ (\i. x$i) ` {i. x$i \ 0}" unfolding xi_def apply(rule Min_in) using False by auto + then obtain i where i':"x$i = xi" "x$i \ 0" by auto + have i:"\j. x$j > 0 \ x$i \ x$j" + unfolding i'(1) xi_def apply(rule_tac Min_le) unfolding image_iff + defer apply(rule_tac x=j in bexI) using i' by auto + have i01:"x$i \ 1" "x$i > 0" using Suc(2)[unfolded mem_interval,rule_format,of i] using i'(2) `x$i \ 0` + by(auto simp add: Cart_lambda_beta) + show ?thesis proof(cases "x$i=1") + case True have "\j\{i. x$i \ 0}. x$j = 1" apply(rule, rule ccontr) unfolding mem_Collect_eq proof- + fix j assume "x $ j \ 0" "x $ j \ 1" + hence j:"x$j \ {0<..<1}" using Suc(2) by(auto simp add: vector_less_eq_def elim!:allE[where x=j]) + hence "x$j \ op $ x ` {i. x $ i \ 0}" by auto + hence "x$j \ x$i" unfolding i'(1) xi_def apply(rule_tac Min_le) by auto + thus False using True Suc(2) j by(auto simp add: vector_less_eq_def elim!:ballE[where x=j]) qed + thus "x\convex hull ?points" apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) + by(auto simp add: Cart_lambda_beta) + next let ?y = "\j. if x$j = 0 then 0 else (x$j - x$i) / (1 - x$i)" + case False hence *:"x = x$i *\<^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 + using Suc(2)[unfolded mem_interval, rule_format, of j] by(auto simp add:field_simps Cart_lambda_beta) + hence "0 \ ?y j \ ?y j \ 1" by auto } + moreover have "i\{j. x$j \ 0} - {j. ((\ j. ?y j)::real^'n) $ j \ 0}" using i01 by(auto simp add: Cart_lambda_beta) + hence "{j. x$j \ 0} \ {j. ((\ j. ?y j)::real^'n::finite) $ j \ 0}" by auto + hence **:"{j. ((\ j. ?y j)::real^'n::finite) $ j \ 0} \ {j. x$j \ 0}" apply - apply rule by(auto simp add: Cart_lambda_beta) + have "card {j. ((\ j. ?y j)::real^'n) $ j \ 0} \ n" using less_le_trans[OF psubset_card_mono[OF _ **] Suc(4)] by auto + ultimately show ?thesis apply(subst *) apply(rule convex_convex_hull[unfolded convex_def, rule_format]) + apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) defer apply(rule Suc(1)) + unfolding mem_interval using i01 Suc(3) by (auto simp add: Cart_lambda_beta) + qed qed qed } note * = this + show ?thesis apply rule defer apply(rule hull_minimal) unfolding subset_eq prefer 3 apply rule + apply(rule_tac n2="CARD('n)" in *) prefer 3 apply(rule card_mono) using 01 and convex_interval(1) prefer 5 apply - apply rule + unfolding mem_interval apply rule unfolding mem_Collect_eq apply(erule_tac x=i in allE) + by(auto simp add: vector_less_eq_def mem_def[of _ convex]) qed + +subsection {* And this is a finite set of vertices. *} + +lemma unit_cube_convex_hull: obtains s where "finite s" "{0 .. 1::real^'n::finite} = convex hull s" + apply(rule that[of "{x::real^'n::finite. \i. x$i=0 \ x$i=1}"]) + apply(rule finite_subset[of _ "(\s. (\ i. if i\s then 1::real else 0)::real^'n::finite) ` UNIV"]) + prefer 3 apply(rule unit_interval_convex_hull) apply rule unfolding mem_Collect_eq proof- + fix x::"real^'n" assume as:"\i. x $ i = 0 \ x $ i = 1" + show "x \ (\s. \ i. if i \ s then 1 else 0) ` UNIV" apply(rule image_eqI[where x="{i. x$i = 1}"]) + unfolding Cart_eq using as by(auto simp add:Cart_lambda_beta) qed auto + +subsection {* Hence any cube (could do any nonempty interval). *} + +lemma cube_convex_hull: + assumes "0 < d" obtains s::"(real^'n::finite) set" where "finite s" "{x - (\ i. d) .. x + (\ i. d)} = convex hull s" proof- + let ?d = "(\ i. d)::real^'n" + have *:"{x - ?d .. x + ?d} = (\y. x - ?d + (2 * d) *\<^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]] + by(auto simp add: vector_component) + hence "1 \ inverse d * (x $ i - y $ i)" "1 \ inverse d * (y $ i - x $ i)" + apply(rule_tac[!] mult_left_le_imp_le[OF _ assms]) unfolding mult_assoc[THEN sym] + using assms by(auto simp add: field_simps right_inverse) + hence "inverse d * (x $ i * 2) \ 2 + inverse d * (y $ i * 2)" + "inverse d * (y $ i * 2) \ 2 + inverse d * (x $ i * 2)" by(auto simp add:field_simps) } + hence "inverse (2 * d) *\<^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) *\<^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) *\<^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) *\<^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: + fixes s :: "(real ^ _) set" + assumes "open s" "convex_on s f" "\x\s. abs(f x) \ b" + shows "continuous_on s f" + apply(rule continuous_at_imp_continuous_on) unfolding continuous_at_real_range proof(rule,rule,rule) + fix x e assume "x\s" "(0::real) < e" + def B \ "abs b + 1" + have B:"0 < B" "\x. x\s \ abs (f x) \ B" + unfolding B_def defer apply(drule assms(3)[rule_format]) by auto + obtain k where "k>0"and k:"cball x k \ s" using assms(1)[unfolded open_contains_cball, THEN bspec[where x=x]] using `x\s` by auto + show "\d>0. \x'. norm (x' - x) < d \ \f x' - f x\ < e" + apply(rule_tac x="min (k / 2) (e / (2 * B) * k)" in exI) apply rule defer proof(rule,rule) + fix y assume as:"norm (y - x) < min (k / 2) (e / (2 * B) * k)" + show "\f y - f x\ < e" proof(cases "y=x") + case False def t \ "k / norm (y - x)" + have "2 < t" "00` by(auto simp add:field_simps) + have "y\s" apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_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 *\<^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 + 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) + hence th1:"f y - f x < e" apply- apply(rule le_less_trans) defer apply assumption + using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w] + using `0s` `w\s` by(auto simp add:field_simps) } + moreover + { def w \ "x - t *\<^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 + 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" + using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w] + using `0s` `w\s` by (auto simp add:field_simps) + also have "\ = (f w + t * f y) / (1 + t)" using `t>0` unfolding real_divide_def by (auto simp add:field_simps) + also have "\ < e + f y" using `t>0` * `e>0` by(auto simp add:field_simps) + finally have "f x - f y < e" by auto } + ultimately show ?thesis by auto + qed(insert `0y \ cball x e. f y \ b" + shows "\y \ cball x e. abs(f y) \ b + 2 * abs(f x)" + apply(rule) proof(cases "0 \ e") case True + fix y assume y:"y\cball x e" def z \ "2 *\<^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) *\<^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" + hence "dist x y < 0" using False unfolding mem_cball not_le by (auto simp del: dist_not_less_zero) + thus "\f y\ \ b + 2 * \f x\" using zero_le_dist[of x y] by auto qed + +subsection {* Hence a convex function on an open set is continuous. *} + +lemma convex_on_continuous: + assumes "open (s::(real^'n::finite) set)" "convex_on s f" + shows "continuous_on s f" + unfolding continuous_on_eq_continuous_at[OF assms(1)] proof + note dimge1 = dimindex_ge_1[where 'a='n] + fix x assume "x\s" + then obtain e where e:"cball x e \ s" "e>0" using assms(1) unfolding open_contains_cball by auto + def d \ "e / real CARD('n)" + have "0 < d" unfolding d_def using `e>0` dimge1 by(rule_tac divide_pos_pos, auto) + let ?d = "(\ i. d)::real^'n" + obtain c where c:"finite c" "{x - ?d..x + ?d} = convex hull c" using cube_convex_hull[OF `d>0`, of x] by auto + have "x\{x - ?d..x + ?d}" using `d>0` unfolding mem_interval by(auto simp add:vector_component_simps) + hence "c\{}" apply(rule_tac ccontr) using c by(auto simp add:convex_hull_empty) + def k \ "Max (f ` c)" + have "convex_on {x - ?d..x + ?d} f" apply(rule convex_on_subset[OF assms(2)]) + apply(rule subset_trans[OF _ e(1)]) unfolding subset_eq mem_cball proof + fix z assume z:"z\{x - ?d..x + ?d}" + have e:"e = setsum (\i. d) (UNIV::'n set)" unfolding setsum_constant d_def using dimge1 + by (metis card_enum field_simps d_def not_one_le_zero of_nat_le_iff real_eq_of_nat real_of_nat_1) + show "dist x z \ e" unfolding dist_norm e apply(rule_tac order_trans[OF norm_le_l1], rule setsum_mono) + using z[unfolded mem_interval] apply(erule_tac x=i in allE) by(auto simp add:field_simps vector_component_simps) qed + hence k:"\y\{x - ?d..x + ?d}. f y \ k" unfolding c(2) apply(rule_tac convex_on_convex_hull_bound) apply assumption + unfolding k_def apply(rule, rule Max_ge) using c(1) by auto + have "d \ e" unfolding d_def apply(rule mult_imp_div_pos_le) using `e>0` dimge1 unfolding mult_le_cancel_left1 using real_dimindex_ge_1 by auto + hence dsube:"cball x d \ cball x e" unfolding subset_eq Ball_def mem_cball by auto + have conv:"convex_on (cball x d) f" apply(rule convex_on_subset, rule convex_on_subset[OF assms(2)]) apply(rule e(1)) using dsube by auto + hence "\y\cball x d. abs (f y) \ k + 2 * abs (f x)" apply(rule_tac convex_bounds_lemma) apply assumption proof + fix y assume y:"y\cball x d" + { fix i::'n have "x $ i - d \ y $ i" "y $ i \ x $ i + d" + using order_trans[OF component_le_norm y[unfolded mem_cball dist_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) 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) 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 :: "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)" + +lemmas segment = open_segment_def closed_segment_def + +definition "starlike s \ (\a\s. \x\s. closed_segment a x \ s)" + +lemma midpoint_refl: "midpoint x x = x" + unfolding midpoint_def unfolding 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) + "dist b (midpoint a b) = (dist a b) / 2" (is ?t2) + "dist (midpoint a b) a = (dist a b) / 2" (is ?t3) + "dist (midpoint a b) b = (dist a b) / 2" (is ?t4) +proof- + have *: "\x y::real^'n::finite. 2 *\<^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) + show ?t4 unfolding midpoint_def dist_norm apply (rule **) by(auto,vector) qed + +lemma midpoint_eq_endpoint: + "midpoint a b = a \ a = (b::real^'n::finite)" + "midpoint a b = b \ a = b" + unfolding dist_eq_0_iff[where 'a="real^'n", THEN sym] dist_midpoint by auto + +lemma convex_contains_segment: + "convex s \ (\a\s. \b\s. closed_segment a b \ s)" + unfolding convex_alt closed_segment_def by auto + +lemma convex_imp_starlike: + "convex s \ s \ {} \ starlike s" + unfolding convex_contains_segment starlike_def by auto + +lemma segment_convex_hull: + "closed_segment a b = convex hull {a,b}" proof- + have *:"\x. {x} \ {}" by auto + have **:"\u v. u + v = 1 \ u = 1 - (v::real)" by auto + show ?thesis unfolding segment convex_hull_insert[OF *] convex_hull_singleton apply(rule set_ext) + unfolding mem_Collect_eq apply(rule,erule exE) + apply(rule_tac x="1 - u" in exI) apply rule defer apply(rule_tac x=u in exI) defer + apply(erule exE, (erule conjE)?)+ apply(rule_tac x="1 - u" in exI) unfolding ** by auto qed + +lemma convex_segment: "convex (closed_segment a b)" + unfolding segment_convex_hull by(rule convex_convex_hull) + +lemma ends_in_segment: "a \ closed_segment a b" "b \ closed_segment a b" + unfolding segment_convex_hull apply(rule_tac[!] hull_subset[unfolded subset_eq, rule_format]) by auto + +lemma segment_furthest_le: + assumes "x \ closed_segment a b" shows "norm(y - x) \ norm(y - a) \ norm(y - x) \ norm(y - b)" proof- + obtain z where "z\{a, b}" "norm (x - y) \ norm (z - y)" using simplex_furthest_le[of "{a, b}" y] + using assms[unfolded segment_convex_hull] by auto + thus ?thesis by(auto simp add:norm_minus_commute) qed + +lemma segment_bound: + assumes "x \ closed_segment a b" + shows "norm(x - a) \ norm(b - a)" "norm(x - b) \ norm(b - a)" + using segment_furthest_le[OF assms, of a] + using segment_furthest_le[OF assms, of b] + by (auto simp add:norm_minus_commute) + +lemma segment_refl:"closed_segment a a = {a}" unfolding segment by (auto simp add: algebra_simps) + +lemma between_mem_segment: "between (a,b) x \ x \ closed_segment a b" + unfolding between_def mem_def by auto + +lemma between:"between (a,b) (x::real^'n::finite) \ dist a b = (dist a x) + (dist x b)" +proof(cases "a = b") + case True thus ?thesis unfolding between_def split_conv mem_def[of x, symmetric] + by(auto simp add:segment_refl dist_commute) next + case False hence Fal:"norm (a - b) \ 0" and Fal2: "norm (a - b) > 0" 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) *\<^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) *\<^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)) *\<^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)) *\<^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) *\<^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 + +lemma between_mem_convex_hull: + "between (a,b) x \ x \ convex hull {a,b}" + unfolding between_mem_segment segment_convex_hull .. + +subsection {* Shrinking towards the interior of a convex set. *} + +lemma mem_interior_convex_shrink: + fixes s :: "(real ^ _) set" + assumes "convex s" "c \ interior s" "x \ s" "0 < e" "e \ 1" + 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 *\<^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 *\<^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]) + apply(rule d[unfolded subset_eq,rule_format]) unfolding mem_ball using assms(3-5) by auto + qed(rule mult_pos_pos, insert `e>0` `d>0`, auto) qed + +lemma mem_interior_closure_convex_shrink: + fixes s :: "(real ^ _) set" + assumes "convex s" "c \ interior s" "x \ closure s" "0 < e" "e \ 1" + 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 + case False hence x:"x islimpt s" using assms(3)[unfolded closure_def] by auto + show ?thesis proof(cases "e=1") + case True obtain y where "y\s" "y \ x" "dist y x < 1" + using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto + thus ?thesis apply(rule_tac x=y in bexI) unfolding True using `d>0` by auto next + case False hence "0 < e * d / (1 - e)" and *:"1 - e > 0" + using `e\1` `e>0` `d>0` by(auto intro!:mult_pos_pos divide_pos_pos) + then obtain y where "y\s" "y \ x" "dist y x < e * d / (1 - e)" + using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto + thus ?thesis apply(rule_tac x=y in bexI) unfolding dist_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) *\<^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 add:field_simps norm_minus_commute) + thus ?thesis unfolding * apply - apply(rule mem_interior_convex_shrink) + using assms(1,4-5) `y\s` by auto qed + +subsection {* Some obvious but surprisingly hard simplex lemmas. *} + +lemma simplex: + assumes "finite s" "0 \ s" + shows "convex hull (insert 0 s) = { y. (\u. (\x\s. 0 \ u x) \ setsum u s \ 1 \ setsum (\x. u x *\<^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) + unfolding if_smult and setsum_delta_notmem[OF assms(2)] by auto + +lemma std_simplex: + "convex hull (insert 0 { basis i | i. i\UNIV}) = + {x::real^'n::finite . (\i. 0 \ x$i) \ setsum (\i. x$i) UNIV \ 1 }" (is "convex hull (insert 0 ?p) = ?s") +proof- let ?D = "UNIV::'n set" + have "0\?p" by(auto simp add: basis_nonzero) + have "{(basis i)::real^'n |i. i \ ?D} = basis ` ?D" by auto + note sumbas = this setsum_reindex[OF basis_inj, unfolded o_def] + show ?thesis unfolding simplex[OF finite_stdbasis `0\?p`] apply(rule set_ext) unfolding mem_Collect_eq apply rule + apply(erule exE, (erule conjE)+) apply(erule_tac[2] conjE)+ proof- + fix x::"real^'n" and u assume as: "\x\{basis i |i. i \?D}. 0 \ u x" "setsum u {basis i |i. i \ ?D} \ 1" "(\x\{basis i |i. i \?D}. u x *\<^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 *\<^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})) = + {x::real^'n::finite. (\i. 0 < x$i) \ setsum (\i. x$i) UNIV < 1 }" + apply(rule set_ext) unfolding mem_interior std_simplex unfolding subset_eq mem_Collect_eq Ball_def mem_ball + unfolding Ball_def[symmetric] apply rule apply(erule exE, (erule conjE)+) defer apply(erule conjE) proof- + fix x::"real^'n" and e assume "0xa. dist x xa < e \ (\x. 0 \ xa $ x) \ setsum (op $ xa) UNIV \ 1" + show "(\xa. 0 < x $ xa) \ setsum (op $ x) UNIV < 1" apply(rule,rule) proof- + fix i::'n show "0 < x $ i" using as[THEN spec[where x="x - (e / 2) *\<^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) *\<^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) *\<^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 +next + fix x::"real^'n::finite" assume as:"\i. 0 < x $ i" "setsum (op $ x) UNIV < 1" + guess a using UNIV_witness[where 'a='b] .. + let ?d = "(1 - setsum (op $ x) UNIV) / real (CARD('n))" + have "Min ((op $ x) ` UNIV) > 0" apply(rule Min_grI) using as(1) dimindex_ge_1 by auto + moreover have"?d > 0" apply(rule divide_pos_pos) using as(2) using dimindex_ge_1 by(auto simp add: Suc_le_eq) + ultimately show "\e>0. \y. dist x y < e \ (\i. 0 \ y $ i) \ setsum (op $ y) UNIV \ 1" + apply(rule_tac x="min (Min ((op $ x) ` UNIV)) ?D" in exI) apply rule defer apply(rule,rule) proof- + fix y assume y:"dist x y < min (Min (op $ x ` UNIV)) ?d" + have "setsum (op $ y) UNIV \ setsum (\i. x$i + ?d) UNIV" proof(rule setsum_mono) + fix i::'n have "abs (y$i - x$i) < ?d" apply(rule le_less_trans) using component_le_norm[of "y - x" i] + using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] by(auto simp add:vector_component_simps norm_minus_commute) + thus "y $ i \ x $ i + ?d" by auto qed + also have "\ \ 1" unfolding setsum_addf setsum_constant card_enum real_eq_of_nat using dimindex_ge_1 by(auto simp add: Suc_le_eq) + finally show "(\i. 0 \ y $ i) \ setsum (op $ y) UNIV \ 1" apply- proof(rule,rule) + fix i::'n have "norm (x - y) < x$i" using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1] + using Min_gr_iff[of "op $ x ` dimset x"] dimindex_ge_1 by auto + thus "0 \ y$i" using component_le_norm[of "x - y" i] and as(1)[rule_format, of i] by(auto simp add: vector_component_simps) + qed auto qed auto qed + +lemma interior_std_simplex_nonempty: obtains a::"real^'n::finite" where + "a \ interior(convex hull (insert 0 {basis i | i . i \ UNIV}))" proof- + let ?D = "UNIV::'n set" let ?a = "setsum (\b::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) + unfolding setsum_delta'[OF finite_UNIV[where 'a='n]] and real_dimindex_ge_1[where 'n='n] by(auto simp add: basis_component[of i]) } + note ** = this + show ?thesis apply(rule that[of ?a]) unfolding interior_std_simplex mem_Collect_eq proof(rule,rule) + fix i::'n show "0 < ?a $ i" unfolding ** using dimindex_ge_1 by(auto simp add: Suc_le_eq) next + have "setsum (op $ ?a) ?D = setsum (\i. inverse (2 * real CARD('n))) ?D" by(rule setsum_cong2, rule **) + also have "\ < 1" unfolding setsum_constant card_enum real_eq_of_nat real_divide_def[THEN sym] by (auto simp add:field_simps) + finally show "setsum (op $ ?a) ?D < 1" by auto qed qed + +subsection {* Paths. *} + +definition "path (g::real^1 \ real^'n::finite) \ continuous_on {0 .. 1} g" + +definition "pathstart (g::real^1 \ real^'n) = g 0" + +definition "pathfinish (g::real^1 \ real^'n) = g 1" + +definition "path_image (g::real^1 \ real^'n) = g ` {0 .. 1}" + +definition "reversepath (g::real^1 \ real^'n) = (\x. g(1 - x))" + +definition joinpaths:: "(real^1 \ real^'n) \ (real^1 \ real^'n) \ (real^1 \ real^'n)" (infixr "+++" 75) + where "joinpaths g1 g2 = (\x. if dest_vec1 x \ ((1 / 2)::real) then g1 (2 *\<^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)" + +definition "injective_path (g::real^1 \ real^'n) \ + (\x\{0..1}. \y\{0..1}. g x = g y \ x = y)" + +subsection {* Some lemmas about these concepts. *} + +lemma injective_imp_simple_path: + "injective_path g \ simple_path g" + unfolding injective_path_def simple_path_def by auto + +lemma path_image_nonempty: "path_image g \ {}" + unfolding path_image_def image_is_empty interval_eq_empty by auto + +lemma pathstart_in_path_image[intro]: "(pathstart g) \ path_image g" + unfolding pathstart_def path_image_def apply(rule imageI) + unfolding mem_interval_1 vec_1[THEN sym] dest_vec1_vec by auto + +lemma pathfinish_in_path_image[intro]: "(pathfinish g) \ path_image g" + unfolding pathfinish_def path_image_def apply(rule imageI) + unfolding mem_interval_1 vec_1[THEN sym] dest_vec1_vec by auto + +lemma connected_path_image[intro]: "path g \ connected(path_image g)" + unfolding path_def path_image_def apply(rule connected_continuous_image, assumption) + by(rule convex_connected, rule convex_interval) + +lemma compact_path_image[intro]: "path g \ compact(path_image g)" + unfolding path_def path_image_def apply(rule compact_continuous_image, assumption) + by(rule compact_interval) + +lemma reversepath_reversepath[simp]: "reversepath(reversepath g) = g" + unfolding reversepath_def by auto + +lemma pathstart_reversepath[simp]: "pathstart(reversepath g) = pathfinish g" + unfolding pathstart_def reversepath_def pathfinish_def by auto + +lemma pathfinish_reversepath[simp]: "pathfinish(reversepath g) = pathstart g" + unfolding pathstart_def reversepath_def pathfinish_def by auto + +lemma pathstart_join[simp]: "pathstart(g1 +++ g2) = pathstart g1" + unfolding pathstart_def joinpaths_def pathfinish_def by auto + +lemma pathfinish_join[simp]:"pathfinish(g1 +++ g2) = pathfinish g2" proof- + have "2 *\<^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 + +lemma path_image_reversepath[simp]: "path_image(reversepath g) = path_image g" proof- + have *:"\g. path_image(reversepath g) \ path_image g" + unfolding path_image_def subset_eq reversepath_def Ball_def image_iff apply(rule,rule,erule bexE) + apply(rule_tac x="1 - xa" in bexI) by(auto simp add:vector_less_eq_def vector_component_simps elim!:ballE) + show ?thesis using *[of g] *[of "reversepath g"] unfolding reversepath_reversepath by auto qed + +lemma path_reversepath[simp]: "path(reversepath g) \ path g" proof- + have *:"\g. path g \ path(reversepath g)" unfolding path_def reversepath_def + apply(rule continuous_on_compose[unfolded o_def, of _ "\x. 1 - x"]) + apply(rule continuous_on_sub, rule continuous_on_const, rule continuous_on_id) + apply(rule continuous_on_subset[of "{0..1}"], assumption) + by (auto, auto simp add:vector_less_eq_def vector_component_simps elim!:ballE) + show ?thesis using *[of g] *[of "reversepath g"] unfolding reversepath_reversepath by auto qed + +lemmas reversepath_simps = path_reversepath path_image_reversepath pathstart_reversepath pathfinish_reversepath + +lemma path_join[simp]: assumes "pathfinish g1 = pathstart g2" shows "path (g1 +++ g2) \ path g1 \ path g2" + unfolding path_def pathfinish_def pathstart_def apply rule defer apply(erule conjE) proof- + assume as:"continuous_on {0..1} (g1 +++ g2)" + have *:"g1 = (\x. g1 (2 *\<^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) + apply (rule continuous_on_cmul, rule continuous_on_add, rule continuous_on_id, rule continuous_on_const) defer + apply (rule continuous_on_cmul, rule continuous_on_id) apply(rule_tac[!] continuous_on_eq[of _ "g1 +++ g2"]) defer prefer 3 + apply(rule_tac[1-2] continuous_on_subset[of "{0 .. 1}"]) apply(rule as, assumption, rule as, assumption) + apply(rule) defer apply rule proof- + fix x assume "x \ op *\<^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 *\<^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 *\<^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)*\<^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) *\<^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) *\<^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)*\<^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 *\<^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) + by(auto intro!: imageI simp add: vector_component_simps) qed + +lemma subset_path_image_join: + assumes "path_image g1 \ s" "path_image g2 \ s" shows "path_image(g1 +++ g2) \ s" + using path_image_join_subset[of g1 g2] and assms by auto + +lemma path_image_join: + assumes "path g1" "path g2" "pathfinish g1 = pathstart g2" + shows "path_image(g1 +++ g2) = (path_image g1) \ (path_image g2)" +apply(rule, rule path_image_join_subset, rule) unfolding Un_iff proof(erule disjE) + fix x assume "x \ path_image g1" + then obtain y where y:"y\{0..1}" "x = g1 y" unfolding path_image_def image_iff by auto + thus "x \ path_image (g1 +++ g2)" unfolding joinpaths_def path_image_def image_iff + apply(rule_tac x="(1/2) *\<^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) *\<^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: + assumes "x \ path_image g1" "x \ path_image g2" shows "x \ path_image(g1 +++ g2)" + using assms and path_image_join_subset[of g1 g2] by auto + +lemma simple_path_reversepath: assumes "simple_path g" shows "simple_path (reversepath g)" + using assms unfolding simple_path_def reversepath_def apply- apply(rule ballI)+ + apply(erule_tac x="1-x" in ballE, erule_tac x="1-y" in ballE) + unfolding mem_interval_1 by(auto simp add:vector_component_simps) + +lemma 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}" + shows "simple_path(g1 +++ g2)" +unfolding simple_path_def proof((rule ballI)+, rule impI) let ?g = "g1 +++ g2" + note inj = assms(1,2)[unfolded injective_path_def, rule_format] + fix x y::"real^1" assume xy:"x \ {0..1}" "y \ {0..1}" "?g x = ?g y" + show "x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" proof(case_tac "x$1 \ 1/2",case_tac[!] "y$1 \ 1/2", unfold not_le) + assume as:"x $ 1 \ 1 / 2" "y $ 1 \ 1 / 2" + hence "g1 (2 *\<^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*\<^sub>R x" "2*\<^sub>R y"] by auto + next assume as:"x $ 1 > 1 / 2" "y $ 1 > 1 / 2" + 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*\<^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 *\<^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 *\<^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 *\<^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 *\<^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 *\<^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 *\<^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: + assumes "injective_path g1" "injective_path g2" "pathfinish g1 = pathstart g2" + "(path_image g1 \ path_image g2) \ {pathstart g2}" + shows "injective_path(g1 +++ g2)" + unfolding injective_path_def proof(rule,rule,rule) let ?g = "g1 +++ g2" + note inj = assms(1,2)[unfolded injective_path_def, rule_format] + fix x y assume xy:"x \ {0..1}" "y \ {0..1}" "(g1 +++ g2) x = (g1 +++ g2) y" + show "x = y" proof(cases "x$1 \ 1/2", case_tac[!] "y$1 \ 1/2", unfold not_le) + assume "x $ 1 \ 1 / 2" "y $ 1 \ 1 / 2" thus ?thesis using inj(1)[of "2*\<^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*\<^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 *\<^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 *\<^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 + +lemmas join_paths_simps = path_join path_image_join pathstart_join pathfinish_join + +subsection {* Reparametrizing a closed curve to start at some chosen point. *} + +definition "shiftpath a (f::real^1 \ real^'n) = + (\x. if dest_vec1 (a + x) \ 1 then f(a + x) else f(a + x - 1))" + +lemma pathstart_shiftpath: "a \ 1 \ pathstart(shiftpath a g) = g a" + unfolding pathstart_def shiftpath_def by auto + +(** move this **) +declare forall_1[simp] ex_1[simp] + +lemma pathfinish_shiftpath: assumes "0 \ a" "pathfinish g = pathstart g" + shows "pathfinish(shiftpath a g) = g a" + using assms unfolding pathstart_def pathfinish_def shiftpath_def + by(auto simp add: vector_component_simps) + +lemma endpoints_shiftpath: + assumes "pathfinish g = pathstart g" "a \ {0 .. 1}" + shows "pathfinish(shiftpath a g) = g a" "pathstart(shiftpath a g) = g a" + using assms by(auto intro!:pathfinish_shiftpath pathstart_shiftpath) + +lemma closed_shiftpath: + assumes "pathfinish g = pathstart g" "a \ {0..1}" + shows "pathfinish(shiftpath a g) = pathstart(shiftpath a g)" + using endpoints_shiftpath[OF assms] by auto + +lemma path_shiftpath: + assumes "path g" "pathfinish g = pathstart g" "a \ {0..1}" + shows "path(shiftpath a g)" proof- + have *:"{0 .. 1} = {0 .. 1-a} \ {1-a .. 1}" using assms(3) by(auto simp add: vector_component_simps) + have **:"\x. x + a = 1 \ g (x + a - 1) = g (x + a)" + using assms(2)[unfolded pathfinish_def pathstart_def] by auto + show ?thesis unfolding path_def shiftpath_def * apply(rule continuous_on_union) + apply(rule closed_interval)+ apply(rule continuous_on_eq[of _ "g \ (\x. a + x)"]) prefer 3 + apply(rule continuous_on_eq[of _ "g \ (\x. a - 1 + x)"]) defer prefer 3 + apply(rule continuous_on_intros)+ prefer 2 apply(rule continuous_on_intros)+ + apply(rule_tac[1-2] continuous_on_subset[OF assms(1)[unfolded path_def]]) + using assms(3) and ** by(auto simp add:vector_component_simps field_simps Cart_eq) qed + +lemma shiftpath_shiftpath: assumes "pathfinish g = pathstart g" "a \ {0..1}" "x \ {0..1}" + shows "shiftpath (1 - a) (shiftpath a g) x = g x" + using assms unfolding pathfinish_def pathstart_def shiftpath_def + by(auto simp add: vector_component_simps) + +lemma path_image_shiftpath: + assumes "a \ {0..1}" "pathfinish g = pathstart g" + shows "path_image(shiftpath a g) = path_image g" proof- + { fix x assume as:"g 1 = g 0" "x \ {0..1::real^1}" " \y\{0..1} \ {x. \ a $ 1 + x $ 1 \ 1}. g x \ g (a + y - 1)" + hence "\y\{0..1} \ {x. a $ 1 + x $ 1 \ 1}. g x = g (a + y)" proof(cases "a \ x") + case False thus ?thesis apply(rule_tac x="1 + x - a" in bexI) + using as(1,2) and as(3)[THEN bspec[where x="1 + x - a"]] and assms(1) + by(auto simp add:vector_component_simps field_simps atomize_not) next + case True thus ?thesis using as(1-2) and assms(1) apply(rule_tac x="x - a" in bexI) + by(auto simp add:vector_component_simps field_simps) qed } + thus ?thesis using assms unfolding shiftpath_def path_image_def pathfinish_def pathstart_def + by(auto simp add:vector_component_simps image_iff) qed + +subsection {* Special case of straight-line paths. *} + +definition + linepath :: "real ^ 'n::finite \ real ^ 'n \ real ^ 1 \ real ^ 'n" where + "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 + +lemma pathfinish_linepath[simp]: "pathfinish(linepath a b) = b" + unfolding pathfinish_def linepath_def by auto + +lemma continuous_linepath_at[intro]: "continuous (at x) (linepath a b)" + unfolding linepath_def + by (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) + +lemma path_linepath[intro]: "path(linepath a b)" + unfolding path_def by(rule continuous_on_linepath) + +lemma path_image_linepath[simp]: "path_image(linepath a b) = (closed_segment a b)" + unfolding path_image_def segment linepath_def apply (rule set_ext, rule) defer + unfolding mem_Collect_eq image_iff apply(erule exE) apply(rule_tac x="u *\<^sub>R 1" in bexI) + by(auto simp add:vector_component_simps) + +lemma reversepath_linepath[simp]: "reversepath(linepath a b) = linepath b a" + unfolding reversepath_def linepath_def by(rule ext, auto simp add:vector_component_simps) + +lemma injective_path_linepath: assumes "a \ b" shows "injective_path(linepath a b)" proof- + { obtain i where i:"a$i \ b$i" using assms[unfolded Cart_eq] by auto + fix x y::"real^1" assume "x $ 1 *\<^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 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) + +subsection {* Bounding a point away from a path. *} + +lemma not_on_path_ball: assumes "path g" "z \ path_image g" + shows "\e>0. ball z e \ (path_image g) = {}" proof- + obtain a where "a\path_image g" "\y\path_image g. dist z a \ dist z y" + using distance_attains_inf[OF _ path_image_nonempty, of g z] + using compact_path_image[THEN compact_imp_closed, OF assms(1)] by auto + thus ?thesis apply(rule_tac x="dist z a" in exI) using assms(2) by(auto intro!: dist_pos_lt) qed + +lemma not_on_path_cball: assumes "path g" "z \ path_image g" + shows "\e>0. cball z e \ (path_image g) = {}" proof- + obtain e where "ball z e \ path_image g = {}" "e>0" using not_on_path_ball[OF assms] by auto + moreover have "cball z (e/2) \ ball z e" using `e>0` by auto + ultimately show ?thesis apply(rule_tac x="e/2" in exI) by auto qed + +subsection {* Path component, considered as a "joinability" relation (from Tom Hales). *} + +definition "path_component s x y \ (\g. path g \ path_image g \ s \ pathstart g = x \ pathfinish g = y)" + +lemmas path_defs = path_def pathstart_def pathfinish_def path_image_def path_component_def + +lemma path_component_mem: assumes "path_component s x y" shows "x \ s" "y \ s" + using assms unfolding path_defs by auto + +lemma path_component_refl: assumes "x \ s" shows "path_component s x x" + unfolding path_defs apply(rule_tac x="\u. x" in exI) using assms + by(auto intro!:continuous_on_intros) + +lemma path_component_refl_eq: "path_component s x x \ x \ s" + by(auto intro!: path_component_mem path_component_refl) + +lemma path_component_sym: "path_component s x y \ path_component s y x" + using assms unfolding path_component_def apply(erule exE) apply(rule_tac x="reversepath g" in exI) + by(auto simp add: reversepath_simps) + +lemma path_component_trans: assumes "path_component s x y" "path_component s y z" shows "path_component s x z" + using assms unfolding path_component_def apply- apply(erule exE)+ apply(rule_tac x="g +++ ga" in exI) by(auto simp add: path_image_join) + +lemma path_component_of_subset: "s \ t \ path_component s x y \ path_component t x y" + unfolding path_component_def by auto + +subsection {* Can also consider it as a set, as the name suggests. *} + +lemma path_component_set: "path_component s x = { y. (\g. path g \ path_image g \ s \ pathstart g = x \ pathfinish g = y )}" + apply(rule set_ext) unfolding mem_Collect_eq unfolding mem_def path_component_def by auto + +lemma mem_path_component_set:"x \ path_component s y \ path_component s y x" unfolding mem_def by auto + +lemma path_component_subset: "(path_component s x) \ s" + apply(rule, rule path_component_mem(2)) by(auto simp add:mem_def) + +lemma path_component_eq_empty: "path_component s x = {} \ x \ s" + apply rule apply(drule equals0D[of _ x]) defer apply(rule equals0I) unfolding mem_path_component_set + apply(drule path_component_mem(1)) using path_component_refl by auto + +subsection {* Path connectedness of a space. *} + +definition "path_connected s \ (\x\s. \y\s. \g. path g \ (path_image g) \ s \ pathstart g = x \ pathfinish g = y)" + +lemma path_connected_component: "path_connected s \ (\x\s. \y\s. path_component s x y)" + unfolding path_connected_def path_component_def by auto + +lemma path_connected_component_set: "path_connected s \ (\x\s. path_component s x = s)" + unfolding path_connected_component apply(rule, rule, rule, rule path_component_subset) + unfolding subset_eq mem_path_component_set Ball_def mem_def by auto + +subsection {* Some useful lemmas about path-connectedness. *} + +lemma convex_imp_path_connected: assumes "convex s" shows "path_connected s" + unfolding path_connected_def apply(rule,rule,rule_tac x="linepath x y" in exI) + unfolding path_image_linepath using assms[unfolded convex_contains_segment] by auto + +lemma path_connected_imp_connected: assumes "path_connected s" shows "connected s" + unfolding connected_def not_ex apply(rule,rule,rule ccontr) unfolding not_not apply(erule conjE)+ proof- + fix e1 e2 assume as:"open e1" "open e2" "s \ e1 \ e2" "e1 \ e2 \ s = {}" "e1 \ s \ {}" "e2 \ s \ {}" + then obtain x1 x2 where obt:"x1\e1\s" "x2\e2\s" by auto + then obtain g where g:"path g" "path_image g \ s" "pathstart g = x1" "pathfinish g = x2" + using assms[unfolded path_connected_def,rule_format,of x1 x2] by auto + have *:"connected {0..1::real^1}" by(auto intro!: convex_connected convex_interval) + have "{0..1} \ {x \ {0..1}. g x \ e1} \ {x \ {0..1}. g x \ e2}" using as(3) g(2)[unfolded path_defs] by blast + moreover have "{x \ {0..1}. g x \ e1} \ {x \ {0..1}. g x \ e2} = {}" using as(4) g(2)[unfolded path_defs] unfolding subset_eq by auto + moreover have "{x \ {0..1}. g x \ e1} \ {} \ {x \ {0..1}. g x \ e2} \ {}" using g(3,4)[unfolded path_defs] using obt by(auto intro!: exI) + ultimately show False using *[unfolded connected_local not_ex,rule_format, of "{x\{0..1}. g x \ e1}" "{x\{0..1}. g x \ e2}"] + using continuous_open_in_preimage[OF g(1)[unfolded path_def] as(1)] + using continuous_open_in_preimage[OF g(1)[unfolded path_def] as(2)] by auto qed + +lemma open_path_component: assumes "open s" shows "open(path_component s x)" + unfolding open_contains_ball proof + fix y assume as:"y \ path_component s x" + hence "y\s" apply- apply(rule path_component_mem(2)) unfolding mem_def by auto + then obtain e where e:"e>0" "ball y e \ s" using assms[unfolded open_contains_ball] by auto + show "\e>0. ball y e \ path_component s x" apply(rule_tac x=e in exI) apply(rule,rule `e>0`,rule) unfolding mem_ball mem_path_component_set proof- + fix z assume "dist y z < e" thus "path_component s x z" apply(rule_tac path_component_trans[of _ _ y]) defer + apply(rule path_component_of_subset[OF e(2)]) apply(rule convex_imp_path_connected[OF convex_ball, unfolded path_connected_component, rule_format]) using `e>0` + using as[unfolded mem_def] by auto qed qed + +lemma open_non_path_component: assumes "open s" shows "open(s - path_component s x)" unfolding open_contains_ball proof + fix y assume as:"y\s - path_component s x" + then obtain e where e:"e>0" "ball y e \ s" using assms[unfolded open_contains_ball] by auto + show "\e>0. ball y e \ s - path_component s x" apply(rule_tac x=e in exI) apply(rule,rule `e>0`,rule,rule) defer proof(rule ccontr) + fix z assume "z\ball y e" "\ z \ path_component s x" + hence "y \ path_component s x" unfolding not_not mem_path_component_set using `e>0` + apply- apply(rule path_component_trans,assumption) apply(rule path_component_of_subset[OF e(2)]) + apply(rule convex_imp_path_connected[OF convex_ball, unfolded path_connected_component, rule_format]) by auto + thus False using as by auto qed(insert e(2), auto) qed + +lemma connected_open_path_connected: assumes "open s" "connected s" shows "path_connected s" + unfolding path_connected_component_set proof(rule,rule,rule path_component_subset, rule) + fix x y assume "x \ s" "y \ s" show "y \ path_component s x" proof(rule ccontr) + assume "y \ path_component s x" moreover + have "path_component s x \ s \ {}" using `x\s` path_component_eq_empty path_component_subset[of s x] by auto + ultimately show False using `y\s` open_non_path_component[OF assms(1)] open_path_component[OF assms(1)] + using assms(2)[unfolded connected_def not_ex, rule_format, of"path_component s x" "s - path_component s x"] by auto +qed qed + +lemma path_connected_continuous_image: + assumes "continuous_on s f" "path_connected s" shows "path_connected (f ` s)" + unfolding path_connected_def proof(rule,rule) + fix x' y' assume "x' \ f ` s" "y' \ f ` s" + then obtain x y where xy:"x\s" "y\s" "x' = f x" "y' = f y" by auto + guess g using assms(2)[unfolded path_connected_def,rule_format,OF xy(1,2)] .. + thus "\g. path g \ path_image g \ f ` s \ pathstart g = x' \ pathfinish g = y'" + unfolding xy apply(rule_tac x="f \ g" in exI) unfolding path_defs + using assms(1) by(auto intro!: continuous_on_compose continuous_on_subset[of _ _ "g ` {0..1}"]) qed + +lemma homeomorphic_path_connectedness: + "s homeomorphic t \ (path_connected s \ path_connected t)" + unfolding homeomorphic_def homeomorphism_def apply(erule exE|erule conjE)+ apply rule + apply(drule_tac f=f in path_connected_continuous_image) prefer 3 + apply(drule_tac f=g in path_connected_continuous_image) by auto + +lemma path_connected_empty: "path_connected {}" + unfolding path_connected_def by auto + +lemma path_connected_singleton: "path_connected {a}" + unfolding path_connected_def apply(rule,rule) + apply(rule_tac x="linepath a a" in exI) by(auto simp add:segment 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) + fix x y assume as:"x \ s \ t" "y \ s \ t" + from assms(3) obtain z where "z \ s \ t" by auto + thus "path_component (s \ t) x y" using as using assms(1-2)[unfolded path_connected_component] apply- + apply(erule_tac[!] UnE)+ apply(rule_tac[2-3] path_component_trans[of _ _ z]) + by(auto simp add:path_component_of_subset [OF Un_upper1] path_component_of_subset[OF Un_upper2]) qed + +subsection {* sphere is path-connected. *} + +lemma path_connected_punctured_universe: + assumes "2 \ CARD('n::finite)" shows "path_connected((UNIV::(real^'n::finite) set) - {a})" proof- + obtain \ where \:"bij_betw \ {1..CARD('n)} (UNIV::'n set)" using ex_bij_betw_nat_finite_1[OF finite_UNIV] by auto + let ?U = "UNIV::(real^'n) set" let ?u = "?U - {0}" + let ?basis = "\k. basis (\ k)" + let ?A = "\k. {x::real^'n. \i\{1..k}. inner (basis (\ i)) x \ 0}" + have "\k\{2..CARD('n)}. path_connected (?A k)" proof + 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 < 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 + next case True hence d:"1\{1..CARD('n)}" "2\{1..CARD('n)}" using Suc(2) by auto + have ***:"Suc 1 = 2" by auto + have **:"\s t P Q. s \ t \ {x. P x \ Q x} = (s \ {x. P x}) \ (t \ {x. Q x})" by auto + have "\ 2 \ \ (Suc 0)" apply(rule ccontr) using \[unfolded bij_betw_def inj_on_def, THEN conjunct1, THEN bspec[where x=2]] using assms by auto + thus ?thesis unfolding * True unfolding ** neq_iff bex_disj_distrib apply - + apply(rule path_connected_Un, rule_tac[1-2] path_connected_Un) defer 3 apply(rule_tac[1-4] convex_imp_path_connected) + apply(rule_tac[5] x=" ?basis 1 + ?basis 2" in nequals0I) + apply(rule_tac[6] x="-?basis 1 + ?basis 2" in nequals0I) + apply(rule_tac[7] x="-?basis 1 - ?basis 2" in nequals0I) + using d unfolding *** by(auto intro!: convex_halfspace_gt convex_halfspace_lt, auto simp add:vector_component_simps inner_basis) + qed qed auto qed note lem = this + + 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:"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)}. 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. 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 + +lemma path_connected_sphere: assumes "2 \ CARD('n::finite)" shows "path_connected {x::real^'n::finite. norm(x - a) = r}" proof(cases "r\0") + case True thus ?thesis proof(cases "r=0") + case False hence "{x::real^'n. norm(x - a) = r} = {}" using True by auto + thus ?thesis using path_connected_empty by auto + qed(auto intro!:path_connected_singleton) next + case False hence *:"{x::real^'n. norm(x - a) = r} = (\x. a + r *\<^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 **:"{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 split:split_if_asm) + have "continuous_on (UNIV - {0}) (\x::real^'n. 1 / norm x)" unfolding o_def continuous_on_eq_continuous_within + apply(rule, rule continuous_at_within_inv[unfolded o_def inverse_eq_divide]) apply(rule continuous_at_within) + apply(rule continuous_at_norm[unfolded o_def]) by auto + thus ?thesis unfolding * ** using path_connected_punctured_universe[OF assms] + by(auto intro!: path_connected_continuous_image continuous_on_intros continuous_on_mul) qed + +lemma connected_sphere: "2 \ CARD('n) \ connected {x::real^'n::finite. norm(x - a) = r}" + using path_connected_sphere path_connected_imp_connected by auto + +(** In continuous_at_vec1_norm : Use \ instead of \. **) + +end diff -r 1fad3160d873 -r 2083bde13ce1 src/HOL/Multivariate_Analysis/Determinants.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Multivariate_Analysis/Determinants.thy Fri Oct 23 13:23:18 2009 +0200 @@ -0,0 +1,1087 @@ +(* Title: Determinants + Author: Amine Chaieb, University of Cambridge +*) + +header {* Traces, Determinant of square matrices and some properties *} + +theory Determinants +imports Euclidean_Space Permutations +begin + +subsection{* First some facts about products*} +lemma setprod_insert_eq: "finite A \ setprod f (insert a A) = (if a \ A then setprod f A else f a * setprod f A)" +apply clarsimp +by(subgoal_tac "insert a A = A", auto) + +lemma setprod_add_split: + assumes mn: "(m::nat) <= n + 1" + shows "setprod f {m.. n+p} = setprod f {m .. n} * setprod f {n+1..n+p}" +proof- + let ?A = "{m .. n+p}" + let ?B = "{m .. n}" + let ?C = "{n+1..n+p}" + from mn have un: "?B \ ?C = ?A" by auto + from mn have dj: "?B \ ?C = {}" by auto + have f: "finite ?B" "finite ?C" by simp_all + from setprod_Un_disjoint[OF f dj, of f, unfolded un] show ?thesis . +qed + + +lemma setprod_offset: "setprod f {(m::nat) + p .. n + p} = setprod (\i. f (i + p)) {m..n}" +apply (rule setprod_reindex_cong[where f="op + p"]) +apply (auto simp add: image_iff Bex_def inj_on_def) +apply arith +apply (rule ext) +apply (simp add: add_commute) +done + +lemma setprod_singleton: "setprod f {x} = f x" by simp + +lemma setprod_singleton_nat_seg: "setprod f {n..n} = f (n::'a::order)" by simp + +lemma setprod_numseg: "setprod f {m..0} = (if m=0 then f 0 else 1)" + "setprod f {m .. Suc n} = (if m \ Suc n then f (Suc n) * setprod f {m..n} + else setprod f {m..n})" + by (auto simp add: atLeastAtMostSuc_conv) + +lemma setprod_le: assumes fS: "finite S" and fg: "\x\S. f x \ 0 \ f x \ (g x :: 'a::ordered_idom)" + shows "setprod f S \ setprod g S" +using fS fg +apply(induct S) +apply simp +apply auto +apply (rule mult_mono) +apply (auto intro: setprod_nonneg) +done + + (* FIXME: In Finite_Set there is a useless further assumption *) +lemma setprod_inversef: "finite A ==> setprod (inverse \ f) A = (inverse (setprod f A) :: 'a:: {division_by_zero, field})" + apply (erule finite_induct) + apply (simp) + apply simp + done + +lemma setprod_le_1: assumes fS: "finite S" and f: "\x\S. f x \ 0 \ f x \ (1::'a::ordered_idom)" + shows "setprod f S \ 1" +using setprod_le[OF fS f] unfolding setprod_1 . + +subsection{* Trace *} + +definition trace :: "'a::semiring_1^'n^'n \ 'a" where + "trace A = setsum (\i. ((A$i)$i)) (UNIV::'n set)" + +lemma trace_0: "trace(mat 0) = 0" + by (simp add: trace_def mat_def) + +lemma trace_I: "trace(mat 1 :: 'a::semiring_1^'n^'n) = of_nat(CARD('n))" + by (simp add: trace_def mat_def) + +lemma trace_add: "trace ((A::'a::comm_semiring_1^'n^'n) + B) = trace A + trace B" + by (simp add: trace_def setsum_addf) + +lemma trace_sub: "trace ((A::'a::comm_ring_1^'n^'n) - B) = trace A - trace B" + by (simp add: trace_def setsum_subtractf) + +lemma trace_mul_sym:"trace ((A::'a::comm_semiring_1^'n^'n) ** B) = trace (B**A)" + apply (simp add: trace_def matrix_matrix_mult_def) + apply (subst setsum_commute) + by (simp add: mult_commute) + +(* ------------------------------------------------------------------------- *) +(* Definition of determinant. *) +(* ------------------------------------------------------------------------- *) + +definition det:: "'a::comm_ring_1^'n^'n \ 'a" where + "det A = setsum (\p. of_int (sign p) * setprod (\i. A$i$p i) (UNIV :: 'n set)) {p. p permutes (UNIV :: 'n set)}" + +(* ------------------------------------------------------------------------- *) +(* A few general lemmas we need below. *) +(* ------------------------------------------------------------------------- *) + +lemma setprod_permute: + assumes p: "p permutes S" + shows "setprod f S = setprod (f o p) S" +proof- + {assume "\ finite S" hence ?thesis by simp} + moreover + {assume fS: "finite S" + then have ?thesis + apply (simp add: setprod_def cong del:strong_setprod_cong) + apply (rule ab_semigroup_mult.fold_image_permute) + apply (auto simp add: p) + apply unfold_locales + done} + ultimately show ?thesis by blast +qed + +lemma setproduct_permute_nat_interval: "p permutes {m::nat .. n} ==> setprod f {m..n} = setprod (f o p) {m..n}" + by (blast intro!: setprod_permute) + +(* ------------------------------------------------------------------------- *) +(* Basic determinant properties. *) +(* ------------------------------------------------------------------------- *) + +lemma det_transp: "det (transp A) = det (A::'a::comm_ring_1 ^'n^'n::finite)" +proof- + let ?di = "\A i j. A$i$j" + let ?U = "(UNIV :: 'n set)" + have fU: "finite ?U" by simp + {fix p assume p: "p \ {p. p permutes ?U}" + from p have pU: "p permutes ?U" by blast + have sth: "sign (inv p) = sign p" + by (metis sign_inverse fU p mem_def Collect_def permutation_permutes) + from permutes_inj[OF pU] + have pi: "inj_on p ?U" by (blast intro: subset_inj_on) + from permutes_image[OF pU] + have "setprod (\i. ?di (transp A) i (inv p i)) ?U = setprod (\i. ?di (transp A) i (inv p i)) (p ` ?U)" by simp + also have "\ = setprod ((\i. ?di (transp A) i (inv p i)) o p) ?U" + unfolding setprod_reindex[OF pi] .. + also have "\ = setprod (\i. ?di A i (p i)) ?U" + proof- + {fix i assume i: "i \ ?U" + from i permutes_inv_o[OF pU] permutes_in_image[OF pU] + have "((\i. ?di (transp A) i (inv p i)) o p) i = ?di A i (p i)" + unfolding transp_def by (simp add: expand_fun_eq)} + then show "setprod ((\i. ?di (transp A) i (inv p i)) o p) ?U = setprod (\i. ?di A i (p i)) ?U" by (auto intro: setprod_cong) + qed + finally have "of_int (sign (inv p)) * (setprod (\i. ?di (transp A) i (inv p i)) ?U) = of_int (sign p) * (setprod (\i. ?di A i (p i)) ?U)" using sth + by simp} + then show ?thesis unfolding det_def apply (subst setsum_permutations_inverse) + apply (rule setsum_cong2) by blast +qed + +lemma det_lowerdiagonal: + fixes A :: "'a::comm_ring_1^'n^'n::{finite,wellorder}" + assumes ld: "\i j. i < j \ A$i$j = 0" + shows "det A = setprod (\i. A$i$i) (UNIV:: 'n set)" +proof- + let ?U = "UNIV:: 'n set" + let ?PU = "{p. p permutes ?U}" + let ?pp = "\p. of_int (sign p) * setprod (\i. A$i$p i) (UNIV :: 'n set)" + have fU: "finite ?U" by simp + from finite_permutations[OF fU] have fPU: "finite ?PU" . + have id0: "{id} \ ?PU" by (auto simp add: permutes_id) + {fix p assume p: "p \ ?PU -{id}" + from p have pU: "p permutes ?U" and pid: "p \ id" by blast+ + from permutes_natset_le[OF pU] pid obtain i where + i: "p i > i" by (metis not_le) + from ld[OF i] have ex:"\i \ ?U. A$i$p i = 0" by blast + from setprod_zero[OF fU ex] have "?pp p = 0" by simp} + then have p0: "\p \ ?PU -{id}. ?pp p = 0" by blast + from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis + unfolding det_def by (simp add: sign_id) +qed + +lemma det_upperdiagonal: + fixes A :: "'a::comm_ring_1^'n^'n::{finite,wellorder}" + assumes ld: "\i j. i > j \ A$i$j = 0" + shows "det A = setprod (\i. A$i$i) (UNIV:: 'n set)" +proof- + let ?U = "UNIV:: 'n set" + let ?PU = "{p. p permutes ?U}" + let ?pp = "(\p. of_int (sign p) * setprod (\i. A$i$p i) (UNIV :: 'n set))" + have fU: "finite ?U" by simp + from finite_permutations[OF fU] have fPU: "finite ?PU" . + have id0: "{id} \ ?PU" by (auto simp add: permutes_id) + {fix p assume p: "p \ ?PU -{id}" + from p have pU: "p permutes ?U" and pid: "p \ id" by blast+ + from permutes_natset_ge[OF pU] pid obtain i where + i: "p i < i" by (metis not_le) + from ld[OF i] have ex:"\i \ ?U. A$i$p i = 0" by blast + from setprod_zero[OF fU ex] have "?pp p = 0" by simp} + then have p0: "\p \ ?PU -{id}. ?pp p = 0" by blast + from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis + unfolding det_def by (simp add: sign_id) +qed + +lemma det_diagonal: + fixes A :: "'a::comm_ring_1^'n^'n::finite" + assumes ld: "\i j. i \ j \ A$i$j = 0" + shows "det A = setprod (\i. A$i$i) (UNIV::'n set)" +proof- + let ?U = "UNIV:: 'n set" + let ?PU = "{p. p permutes ?U}" + let ?pp = "\p. of_int (sign p) * setprod (\i. A$i$p i) (UNIV :: 'n set)" + have fU: "finite ?U" by simp + from finite_permutations[OF fU] have fPU: "finite ?PU" . + have id0: "{id} \ ?PU" by (auto simp add: permutes_id) + {fix p assume p: "p \ ?PU - {id}" + then have "p \ id" by simp + then obtain i where i: "p i \ i" unfolding expand_fun_eq by auto + from ld [OF i [symmetric]] have ex:"\i \ ?U. A$i$p i = 0" by blast + from setprod_zero [OF fU ex] have "?pp p = 0" by simp} + then have p0: "\p \ ?PU - {id}. ?pp p = 0" by blast + from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis + unfolding det_def by (simp add: sign_id) +qed + +lemma det_I: "det (mat 1 :: 'a::comm_ring_1^'n^'n::finite) = 1" +proof- + let ?A = "mat 1 :: 'a::comm_ring_1^'n^'n" + let ?U = "UNIV :: 'n set" + let ?f = "\i j. ?A$i$j" + {fix i assume i: "i \ ?U" + have "?f i i = 1" using i by (vector mat_def)} + hence th: "setprod (\i. ?f i i) ?U = setprod (\x. 1) ?U" + by (auto intro: setprod_cong) + {fix i j assume i: "i \ ?U" and j: "j \ ?U" and ij: "i \ j" + have "?f i j = 0" using i j ij by (vector mat_def) } + then have "det ?A = setprod (\i. ?f i i) ?U" using det_diagonal + by blast + also have "\ = 1" unfolding th setprod_1 .. + finally show ?thesis . +qed + +lemma det_0: "det (mat 0 :: 'a::comm_ring_1^'n^'n::finite) = 0" + by (simp add: det_def setprod_zero) + +lemma det_permute_rows: + fixes A :: "'a::comm_ring_1^'n^'n::finite" + assumes p: "p permutes (UNIV :: 'n::finite set)" + shows "det(\ i. A$p i :: 'a^'n^'n) = of_int (sign p) * det A" + apply (simp add: det_def setsum_right_distrib mult_assoc[symmetric]) + apply (subst sum_permutations_compose_right[OF p]) +proof(rule setsum_cong2) + let ?U = "UNIV :: 'n set" + let ?PU = "{p. p permutes ?U}" + fix q assume qPU: "q \ ?PU" + have fU: "finite ?U" by simp + from qPU have q: "q permutes ?U" by blast + from p q have pp: "permutation p" and qp: "permutation q" + by (metis fU permutation_permutes)+ + from permutes_inv[OF p] have ip: "inv p permutes ?U" . + have "setprod (\i. A$p i$ (q o p) i) ?U = setprod ((\i. A$p i$(q o p) i) o inv p) ?U" + by (simp only: setprod_permute[OF ip, symmetric]) + also have "\ = setprod (\i. A $ (p o inv p) i $ (q o (p o inv p)) i) ?U" + by (simp only: o_def) + also have "\ = setprod (\i. A$i$q i) ?U" by (simp only: o_def permutes_inverses[OF p]) + finally have thp: "setprod (\i. A$p i$ (q o p) i) ?U = setprod (\i. A$i$q i) ?U" + by blast + show "of_int (sign (q o p)) * setprod (\i. A$ p i$ (q o p) i) ?U = of_int (sign p) * of_int (sign q) * setprod (\i. A$i$q i) ?U" + by (simp only: thp sign_compose[OF qp pp] mult_commute of_int_mult) +qed + +lemma det_permute_columns: + fixes A :: "'a::comm_ring_1^'n^'n::finite" + assumes p: "p permutes (UNIV :: 'n set)" + shows "det(\ i j. A$i$ p j :: 'a^'n^'n) = of_int (sign p) * det A" +proof- + let ?Ap = "\ i j. A$i$ p j :: 'a^'n^'n" + let ?At = "transp A" + have "of_int (sign p) * det A = det (transp (\ i. transp A $ p i))" + unfolding det_permute_rows[OF p, of ?At] det_transp .. + moreover + have "?Ap = transp (\ i. transp A $ p i)" + by (simp add: transp_def Cart_eq) + ultimately show ?thesis by simp +qed + +lemma det_identical_rows: + fixes A :: "'a::ordered_idom^'n^'n::finite" + assumes ij: "i \ j" + and r: "row i A = row j A" + shows "det A = 0" +proof- + have tha: "\(a::'a) b. a = b ==> b = - a ==> a = 0" + by simp + have th1: "of_int (-1) = - 1" by (metis of_int_1 of_int_minus number_of_Min) + let ?p = "Fun.swap i j id" + let ?A = "\ i. A $ ?p i" + from r have "A = ?A" by (simp add: Cart_eq row_def swap_def) + hence "det A = det ?A" by simp + moreover have "det A = - det ?A" + by (simp add: det_permute_rows[OF permutes_swap_id] sign_swap_id ij th1) + ultimately show "det A = 0" by (metis tha) +qed + +lemma det_identical_columns: + fixes A :: "'a::ordered_idom^'n^'n::finite" + assumes ij: "i \ j" + and r: "column i A = column j A" + shows "det A = 0" +apply (subst det_transp[symmetric]) +apply (rule det_identical_rows[OF ij]) +by (metis row_transp r) + +lemma det_zero_row: + fixes A :: "'a::{idom, ring_char_0}^'n^'n::finite" + assumes r: "row i A = 0" + shows "det A = 0" +using r +apply (simp add: row_def det_def Cart_eq) +apply (rule setsum_0') +apply (auto simp: sign_nz) +done + +lemma det_zero_column: + fixes A :: "'a::{idom,ring_char_0}^'n^'n::finite" + assumes r: "column i A = 0" + shows "det A = 0" + apply (subst det_transp[symmetric]) + apply (rule det_zero_row [of i]) + by (metis row_transp r) + +lemma det_row_add: + fixes a b c :: "'n::finite \ _ ^ 'n" + shows "det((\ i. if i = k then a i + b i else c i)::'a::comm_ring_1^'n^'n) = + det((\ i. if i = k then a i else c i)::'a::comm_ring_1^'n^'n) + + det((\ i. if i = k then b i else c i)::'a::comm_ring_1^'n^'n)" +unfolding det_def Cart_lambda_beta setsum_addf[symmetric] +proof (rule setsum_cong2) + let ?U = "UNIV :: 'n set" + let ?pU = "{p. p permutes ?U}" + let ?f = "(\i. if i = k then a i + b i else c i)::'n \ 'a::comm_ring_1^'n" + let ?g = "(\ i. if i = k then a i else c i)::'n \ 'a::comm_ring_1^'n" + let ?h = "(\ i. if i = k then b i else c i)::'n \ 'a::comm_ring_1^'n" + fix p assume p: "p \ ?pU" + let ?Uk = "?U - {k}" + from p have pU: "p permutes ?U" by blast + have kU: "?U = insert k ?Uk" by blast + {fix j assume j: "j \ ?Uk" + from j have "?f j $ p j = ?g j $ p j" and "?f j $ p j= ?h j $ p j" + by simp_all} + then have th1: "setprod (\i. ?f i $ p i) ?Uk = setprod (\i. ?g i $ p i) ?Uk" + and th2: "setprod (\i. ?f i $ p i) ?Uk = setprod (\i. ?h i $ p i) ?Uk" + apply - + apply (rule setprod_cong, simp_all)+ + done + have th3: "finite ?Uk" "k \ ?Uk" by auto + have "setprod (\i. ?f i $ p i) ?U = setprod (\i. ?f i $ p i) (insert k ?Uk)" + unfolding kU[symmetric] .. + also have "\ = ?f k $ p k * setprod (\i. ?f i $ p i) ?Uk" + apply (rule setprod_insert) + apply simp + by blast + also have "\ = (a k $ p k * setprod (\i. ?f i $ p i) ?Uk) + (b k$ p k * setprod (\i. ?f i $ p i) ?Uk)" by (simp add: ring_simps) + also have "\ = (a k $ p k * setprod (\i. ?g i $ p i) ?Uk) + (b k$ p k * setprod (\i. ?h i $ p i) ?Uk)" by (metis th1 th2) + also have "\ = setprod (\i. ?g i $ p i) (insert k ?Uk) + setprod (\i. ?h i $ p i) (insert k ?Uk)" + unfolding setprod_insert[OF th3] by simp + finally have "setprod (\i. ?f i $ p i) ?U = setprod (\i. ?g i $ p i) ?U + setprod (\i. ?h i $ p i) ?U" unfolding kU[symmetric] . + then show "of_int (sign p) * setprod (\i. ?f i $ p i) ?U = of_int (sign p) * setprod (\i. ?g i $ p i) ?U + of_int (sign p) * setprod (\i. ?h i $ p i) ?U" + by (simp add: ring_simps) +qed + +lemma det_row_mul: + fixes a b :: "'n::finite \ _ ^ 'n" + shows "det((\ i. if i = k then c *s a i else b i)::'a::comm_ring_1^'n^'n) = + c* det((\ i. if i = k then a i else b i)::'a::comm_ring_1^'n^'n)" + +unfolding det_def Cart_lambda_beta setsum_right_distrib +proof (rule setsum_cong2) + let ?U = "UNIV :: 'n set" + let ?pU = "{p. p permutes ?U}" + let ?f = "(\i. if i = k then c*s a i else b i)::'n \ 'a::comm_ring_1^'n" + let ?g = "(\ i. if i = k then a i else b i)::'n \ 'a::comm_ring_1^'n" + fix p assume p: "p \ ?pU" + let ?Uk = "?U - {k}" + from p have pU: "p permutes ?U" by blast + have kU: "?U = insert k ?Uk" by blast + {fix j assume j: "j \ ?Uk" + from j have "?f j $ p j = ?g j $ p j" by simp} + then have th1: "setprod (\i. ?f i $ p i) ?Uk = setprod (\i. ?g i $ p i) ?Uk" + apply - + apply (rule setprod_cong, simp_all) + done + have th3: "finite ?Uk" "k \ ?Uk" by auto + have "setprod (\i. ?f i $ p i) ?U = setprod (\i. ?f i $ p i) (insert k ?Uk)" + unfolding kU[symmetric] .. + also have "\ = ?f k $ p k * setprod (\i. ?f i $ p i) ?Uk" + apply (rule setprod_insert) + apply simp + by blast + also have "\ = (c*s a k) $ p k * setprod (\i. ?f i $ p i) ?Uk" by (simp add: ring_simps) + also have "\ = c* (a k $ p k * setprod (\i. ?g i $ p i) ?Uk)" + unfolding th1 by (simp add: mult_ac) + also have "\ = c* (setprod (\i. ?g i $ p i) (insert k ?Uk))" + unfolding setprod_insert[OF th3] by simp + finally have "setprod (\i. ?f i $ p i) ?U = c* (setprod (\i. ?g i $ p i) ?U)" unfolding kU[symmetric] . + then show "of_int (sign p) * setprod (\i. ?f i $ p i) ?U = c * (of_int (sign p) * setprod (\i. ?g i $ p i) ?U)" + by (simp add: ring_simps) +qed + +lemma det_row_0: + fixes b :: "'n::finite \ _ ^ 'n" + shows "det((\ i. if i = k then 0 else b i)::'a::comm_ring_1^'n^'n) = 0" +using det_row_mul[of k 0 "\i. 1" b] +apply (simp) + unfolding vector_smult_lzero . + +lemma det_row_operation: + fixes A :: "'a::ordered_idom^'n^'n::finite" + assumes ij: "i \ j" + shows "det (\ k. if k = i then row i A + c *s row j A else row k A) = det A" +proof- + let ?Z = "(\ k. if k = i then row j A else row k A) :: 'a ^'n^'n" + have th: "row i ?Z = row j ?Z" by (vector row_def) + have th2: "((\ k. if k = i then row i A else row k A) :: 'a^'n^'n) = A" + by (vector row_def) + show ?thesis + unfolding det_row_add [of i] det_row_mul[of i] det_identical_rows[OF ij th] th2 + by simp +qed + +lemma det_row_span: + fixes A :: "'a:: ordered_idom^'n^'n::finite" + assumes x: "x \ span {row j A |j. j \ i}" + shows "det (\ k. if k = i then row i A + x else row k A) = det A" +proof- + let ?U = "UNIV :: 'n set" + let ?S = "{row j A |j. j \ i}" + let ?d = "\x. det (\ k. if k = i then x else row k A)" + let ?P = "\x. ?d (row i A + x) = det A" + {fix k + + have "(if k = i then row i A + 0 else row k A) = row k A" by simp} + then have P0: "?P 0" + apply - + apply (rule cong[of det, OF refl]) + by (vector row_def) + moreover + {fix c z y assume zS: "z \ ?S" and Py: "?P y" + from zS obtain j where j: "z = row j A" "i \ j" by blast + let ?w = "row i A + y" + have th0: "row i A + (c*s z + y) = ?w + c*s z" by vector + have thz: "?d z = 0" + apply (rule det_identical_rows[OF j(2)]) + using j by (vector row_def) + have "?d (row i A + (c*s z + y)) = ?d (?w + c*s z)" unfolding th0 .. + then have "?P (c*s z + y)" unfolding thz Py det_row_mul[of i] det_row_add[of i] + by simp } + + ultimately show ?thesis + apply - + apply (rule span_induct_alt[of ?P ?S, OF P0]) + apply blast + apply (rule x) + done +qed + +(* ------------------------------------------------------------------------- *) +(* May as well do this, though it's a bit unsatisfactory since it ignores *) +(* exact duplicates by considering the rows/columns as a set. *) +(* ------------------------------------------------------------------------- *) + +lemma det_dependent_rows: + fixes A:: "'a::ordered_idom^'n^'n::finite" + assumes d: "dependent (rows A)" + shows "det A = 0" +proof- + let ?U = "UNIV :: 'n set" + from d obtain i where i: "row i A \ span (rows A - {row i A})" + unfolding dependent_def rows_def by blast + {fix j k assume jk: "j \ k" + and c: "row j A = row k A" + from det_identical_rows[OF jk c] have ?thesis .} + moreover + {assume H: "\ i j. i \ j \ row i A \ row j A" + have th0: "- row i A \ span {row j A|j. j \ i}" + apply (rule span_neg) + apply (rule set_rev_mp) + apply (rule i) + apply (rule span_mono) + using H i by (auto simp add: rows_def) + from det_row_span[OF th0] + have "det A = det (\ k. if k = i then 0 *s 1 else row k A)" + unfolding right_minus vector_smult_lzero .. + with det_row_mul[of i "0::'a" "\i. 1"] + have "det A = 0" by simp} + ultimately show ?thesis by blast +qed + +lemma det_dependent_columns: assumes d: "dependent(columns (A::'a::ordered_idom^'n^'n::finite))" shows "det A = 0" +by (metis d det_dependent_rows rows_transp det_transp) + +(* ------------------------------------------------------------------------- *) +(* Multilinearity and the multiplication formula. *) +(* ------------------------------------------------------------------------- *) + +lemma Cart_lambda_cong: "(\x. f x = g x) \ (Cart_lambda f::'a^'n) = (Cart_lambda g :: 'a^'n)" + apply (rule iffD1[OF Cart_lambda_unique]) by vector + +lemma det_linear_row_setsum: + assumes fS: "finite S" + shows "det ((\ i. if i = k then setsum (a i) S else c i)::'a::comm_ring_1^'n^'n::finite) = setsum (\j. det ((\ i. if i = k then a i j else c i)::'a^'n^'n)) S" +proof(induct rule: finite_induct[OF fS]) + case 1 thus ?case apply simp unfolding setsum_empty det_row_0[of k] .. +next + case (2 x F) + then show ?case by (simp add: det_row_add cong del: if_weak_cong) +qed + +lemma finite_bounded_functions: + assumes fS: "finite S" + shows "finite {f. (\i \ {1.. (k::nat)}. f i \ S) \ (\i. i \ {1 .. k} \ f i = i)}" +proof(induct k) + case 0 + have th: "{f. \i. f i = i} = {id}" by (auto intro: ext) + show ?case by (auto simp add: th) +next + case (Suc k) + let ?f = "\(y::nat,g) i. if i = Suc k then y else g i" + let ?S = "?f ` (S \ {f. (\i\{1..k}. f i \ S) \ (\i. i \ {1..k} \ f i = i)})" + have "?S = {f. (\i\{1.. Suc k}. f i \ S) \ (\i. i \ {1.. Suc k} \ f i = i)}" + apply (auto simp add: image_iff) + apply (rule_tac x="x (Suc k)" in bexI) + apply (rule_tac x = "\i. if i = Suc k then i else x i" in exI) + apply (auto intro: ext) + done + with finite_imageI[OF finite_cartesian_product[OF fS Suc.hyps(1)], of ?f] + show ?case by metis +qed + + +lemma eq_id_iff[simp]: "(\x. f x = x) = (f = id)" by (auto intro: ext) + +lemma det_linear_rows_setsum_lemma: + assumes fS: "finite S" and fT: "finite T" + shows "det((\ i. if i \ T then setsum (a i) S else c i):: 'a::comm_ring_1^'n^'n::finite) = + setsum (\f. det((\ i. if i \ T then a i (f i) else c i)::'a^'n^'n)) + {f. (\i \ T. f i \ S) \ (\i. i \ T \ f i = i)}" +using fT +proof(induct T arbitrary: a c set: finite) + case empty + have th0: "\x y. (\ i. if i \ {} then x i else y i) = (\ i. y i)" by vector + from "empty.prems" show ?case unfolding th0 by simp +next + case (insert z T a c) + let ?F = "\T. {f. (\i \ T. f i \ S) \ (\i. i \ T \ f i = i)}" + let ?h = "\(y,g) i. if i = z then y else g i" + let ?k = "\h. (h(z),(\i. if i = z then i else h i))" + let ?s = "\ k a c f. det((\ i. if i \ T then a i (f i) else c i)::'a^'n^'n)" + let ?c = "\i. if i = z then a i j else c i" + have thif: "\a b c d. (if a \ b then c else d) = (if a then c else if b then c else d)" by simp + have thif2: "\a b c d e. (if a then b else if c then d else e) = + (if c then (if a then b else d) else (if a then b else e))" by simp + from `z \ T` have nz: "\i. i \ T \ i = z \ False" by auto + have "det (\ i. if i \ insert z T then setsum (a i) S else c i) = + det (\ i. if i = z then setsum (a i) S + else if i \ T then setsum (a i) S else c i)" + unfolding insert_iff thif .. + also have "\ = (\j\S. det (\ i. if i \ T then setsum (a i) S + else if i = z then a i j else c i))" + unfolding det_linear_row_setsum[OF fS] + apply (subst thif2) + using nz by (simp cong del: if_weak_cong cong add: if_cong) + finally have tha: + "det (\ i. if i \ insert z T then setsum (a i) S else c i) = + (\(j, f)\S \ ?F T. det (\ i. if i \ T then a i (f i) + else if i = z then a i j + else c i))" + unfolding insert.hyps unfolding setsum_cartesian_product by blast + show ?case unfolding tha + apply(rule setsum_eq_general_reverses[where h= "?h" and k= "?k"], + blast intro: finite_cartesian_product fS finite, + blast intro: finite_cartesian_product fS finite) + using `z \ T` + apply (auto intro: ext) + apply (rule cong[OF refl[of det]]) + by vector +qed + +lemma det_linear_rows_setsum: + assumes fS: "finite (S::'n::finite set)" + shows "det (\ i. setsum (a i) S) = setsum (\f. det (\ i. a i (f i) :: 'a::comm_ring_1 ^ 'n^'n::finite)) {f. \i. f i \ S}" +proof- + have th0: "\x y. ((\ i. if i \ (UNIV:: 'n set) then x i else y i) :: 'a^'n^'n) = (\ i. x i)" by vector + + from det_linear_rows_setsum_lemma[OF fS, of "UNIV :: 'n set" a, unfolded th0, OF finite] show ?thesis by simp +qed + +lemma matrix_mul_setsum_alt: + fixes A B :: "'a::comm_ring_1^'n^'n::finite" + shows "A ** B = (\ i. setsum (\k. A$i$k *s B $ k) (UNIV :: 'n set))" + by (vector matrix_matrix_mult_def setsum_component) + +lemma det_rows_mul: + "det((\ i. c i *s a i)::'a::comm_ring_1^'n^'n::finite) = + setprod (\i. c i) (UNIV:: 'n set) * det((\ i. a i)::'a^'n^'n)" +proof (simp add: det_def setsum_right_distrib cong add: setprod_cong, rule setsum_cong2) + let ?U = "UNIV :: 'n set" + let ?PU = "{p. p permutes ?U}" + fix p assume pU: "p \ ?PU" + let ?s = "of_int (sign p)" + from pU have p: "p permutes ?U" by blast + have "setprod (\i. c i * a i $ p i) ?U = setprod c ?U * setprod (\i. a i $ p i) ?U" + unfolding setprod_timesf .. + then show "?s * (\xa\?U. c xa * a xa $ p xa) = + setprod c ?U * (?s* (\xa\?U. a xa $ p xa))" by (simp add: ring_simps) +qed + +lemma det_mul: + fixes A B :: "'a::ordered_idom^'n^'n::finite" + shows "det (A ** B) = det A * det B" +proof- + let ?U = "UNIV :: 'n set" + let ?F = "{f. (\i\ ?U. f i \ ?U) \ (\i. i \ ?U \ f i = i)}" + let ?PU = "{p. p permutes ?U}" + have fU: "finite ?U" by simp + have fF: "finite ?F" by (rule finite) + {fix p assume p: "p permutes ?U" + + have "p \ ?F" unfolding mem_Collect_eq permutes_in_image[OF p] + using p[unfolded permutes_def] by simp} + then have PUF: "?PU \ ?F" by blast + {fix f assume fPU: "f \ ?F - ?PU" + have fUU: "f ` ?U \ ?U" using fPU by auto + from fPU have f: "\i \ ?U. f i \ ?U" + "\i. i \ ?U \ f i = i" "\(\y. \!x. f x = y)" unfolding permutes_def + by auto + + let ?A = "(\ i. A$i$f i *s B$f i) :: 'a^'n^'n" + let ?B = "(\ i. B$f i) :: 'a^'n^'n" + {assume fni: "\ inj_on f ?U" + then obtain i j where ij: "f i = f j" "i \ j" + unfolding inj_on_def by blast + from ij + have rth: "row i ?B = row j ?B" by (vector row_def) + from det_identical_rows[OF ij(2) rth] + have "det (\ i. A$i$f i *s B$f i) = 0" + unfolding det_rows_mul by simp} + moreover + {assume fi: "inj_on f ?U" + from f fi have fith: "\i j. f i = f j \ i = j" + unfolding inj_on_def by metis + note fs = fi[unfolded surjective_iff_injective_gen[OF fU fU refl fUU, symmetric]] + + {fix y + from fs f have "\x. f x = y" by blast + then obtain x where x: "f x = y" by blast + {fix z assume z: "f z = y" from fith x z have "z = x" by metis} + with x have "\!x. f x = y" by blast} + with f(3) have "det (\ i. A$i$f i *s B$f i) = 0" by blast} + ultimately have "det (\ i. A$i$f i *s B$f i) = 0" by blast} + hence zth: "\ f\ ?F - ?PU. det (\ i. A$i$f i *s B$f i) = 0" by simp + {fix p assume pU: "p \ ?PU" + from pU have p: "p permutes ?U" by blast + let ?s = "\p. of_int (sign p)" + let ?f = "\q. ?s p * (\i\ ?U. A $ i $ p i) * + (?s q * (\i\ ?U. B $ i $ q i))" + have "(setsum (\q. ?s q * + (\i\ ?U. (\ i. A $ i $ p i *s B $ p i :: 'a^'n^'n) $ i $ q i)) ?PU) = + (setsum (\q. ?s p * (\i\ ?U. A $ i $ p i) * + (?s q * (\i\ ?U. B $ i $ q i))) ?PU)" + unfolding sum_permutations_compose_right[OF permutes_inv[OF p], of ?f] + proof(rule setsum_cong2) + fix q assume qU: "q \ ?PU" + hence q: "q permutes ?U" by blast + from p q have pp: "permutation p" and pq: "permutation q" + unfolding permutation_permutes by auto + have th00: "of_int (sign p) * of_int (sign p) = (1::'a)" + "\a. of_int (sign p) * (of_int (sign p) * a) = a" + unfolding mult_assoc[symmetric] unfolding of_int_mult[symmetric] + by (simp_all add: sign_idempotent) + have ths: "?s q = ?s p * ?s (q o inv p)" + using pp pq permutation_inverse[OF pp] sign_inverse[OF pp] + by (simp add: th00 mult_ac sign_idempotent sign_compose) + have th001: "setprod (\i. B$i$ q (inv p i)) ?U = setprod ((\i. B$i$ q (inv p i)) o p) ?U" + by (rule setprod_permute[OF p]) + have thp: "setprod (\i. (\ i. A$i$p i *s B$p i :: 'a^'n^'n) $i $ q i) ?U = setprod (\i. A$i$p i) ?U * setprod (\i. B$i$ q (inv p i)) ?U" + unfolding th001 setprod_timesf[symmetric] o_def permutes_inverses[OF p] + apply (rule setprod_cong[OF refl]) + using permutes_in_image[OF q] by vector + show "?s q * setprod (\i. (((\ i. A$i$p i *s B$p i) :: 'a^'n^'n)$i$q i)) ?U = ?s p * (setprod (\i. A$i$p i) ?U) * (?s (q o inv p) * setprod (\i. B$i$(q o inv p) i) ?U)" + using ths thp pp pq permutation_inverse[OF pp] sign_inverse[OF pp] + by (simp add: sign_nz th00 ring_simps sign_idempotent sign_compose) + qed + } + then have th2: "setsum (\f. det (\ i. A$i$f i *s B$f i)) ?PU = det A * det B" + unfolding det_def setsum_product + by (rule setsum_cong2) + have "det (A**B) = setsum (\f. det (\ i. A $ i $ f i *s B $ f i)) ?F" + unfolding matrix_mul_setsum_alt det_linear_rows_setsum[OF fU] by simp + also have "\ = setsum (\f. det (\ i. A$i$f i *s B$f i)) ?PU" + using setsum_mono_zero_cong_left[OF fF PUF zth, symmetric] + unfolding det_rows_mul by auto + finally show ?thesis unfolding th2 . +qed + +(* ------------------------------------------------------------------------- *) +(* Relation to invertibility. *) +(* ------------------------------------------------------------------------- *) + +lemma invertible_left_inverse: + fixes A :: "real^'n^'n::finite" + shows "invertible A \ (\(B::real^'n^'n). B** A = mat 1)" + by (metis invertible_def matrix_left_right_inverse) + +lemma invertible_righ_inverse: + fixes A :: "real^'n^'n::finite" + shows "invertible A \ (\(B::real^'n^'n). A** B = mat 1)" + by (metis invertible_def matrix_left_right_inverse) + +lemma invertible_det_nz: + fixes A::"real ^'n^'n::finite" + shows "invertible A \ det A \ 0" +proof- + {assume "invertible A" + then obtain B :: "real ^'n^'n" where B: "A ** B = mat 1" + unfolding invertible_righ_inverse by blast + hence "det (A ** B) = det (mat 1 :: real ^'n^'n)" by simp + hence "det A \ 0" + apply (simp add: det_mul det_I) by algebra } + moreover + {assume H: "\ invertible A" + let ?U = "UNIV :: 'n set" + have fU: "finite ?U" by simp + from H obtain c i where c: "setsum (\i. c i *s row i A) ?U = 0" + and iU: "i \ ?U" and ci: "c i \ 0" + unfolding invertible_righ_inverse + unfolding matrix_right_invertible_independent_rows by blast + have stupid: "\(a::real^'n) b. a + b = 0 \ -a = b" + apply (drule_tac f="op + (- a)" in cong[OF refl]) + apply (simp only: ab_left_minus add_assoc[symmetric]) + apply simp + done + from c ci + have thr0: "- row i A = setsum (\j. (1/ c i) *s (c j *s row j A)) (?U - {i})" + unfolding setsum_diff1'[OF fU iU] setsum_cmul + apply - + apply (rule vector_mul_lcancel_imp[OF ci]) + apply (auto simp add: vector_smult_assoc vector_smult_rneg field_simps) + unfolding stupid .. + have thr: "- row i A \ span {row j A| j. j \ i}" + unfolding thr0 + apply (rule span_setsum) + apply simp + apply (rule ballI) + apply (rule span_mul)+ + apply (rule span_superset) + apply auto + done + let ?B = "(\ k. if k = i then 0 else row k A) :: real ^'n^'n" + have thrb: "row i ?B = 0" using iU by (vector row_def) + have "det A = 0" + unfolding det_row_span[OF thr, symmetric] right_minus + unfolding det_zero_row[OF thrb] ..} + ultimately show ?thesis by blast +qed + +(* ------------------------------------------------------------------------- *) +(* Cramer's rule. *) +(* ------------------------------------------------------------------------- *) + +lemma cramer_lemma_transp: + fixes A:: "'a::ordered_idom^'n^'n::finite" and x :: "'a ^'n::finite" + shows "det ((\ i. if i = k then setsum (\i. x$i *s row i A) (UNIV::'n set) + else row i A)::'a^'n^'n) = x$k * det A" + (is "?lhs = ?rhs") +proof- + let ?U = "UNIV :: 'n set" + let ?Uk = "?U - {k}" + have U: "?U = insert k ?Uk" by blast + have fUk: "finite ?Uk" by simp + have kUk: "k \ ?Uk" by simp + have th00: "\k s. x$k *s row k A + s = (x$k - 1) *s row k A + row k A + s" + by (vector ring_simps) + have th001: "\f k . (\x. if x = k then f k else f x) = f" by (auto intro: ext) + have "(\ i. row i A) = A" by (vector row_def) + then have thd1: "det (\ i. row i A) = det A" by simp + have thd0: "det (\ i. if i = k then row k A + (\i \ ?Uk. x $ i *s row i A) else row i A) = det A" + apply (rule det_row_span) + apply (rule span_setsum[OF fUk]) + apply (rule ballI) + apply (rule span_mul) + apply (rule span_superset) + apply auto + done + show "?lhs = x$k * det A" + apply (subst U) + unfolding setsum_insert[OF fUk kUk] + apply (subst th00) + unfolding add_assoc + apply (subst det_row_add) + unfolding thd0 + unfolding det_row_mul + unfolding th001[of k "\i. row i A"] + unfolding thd1 by (simp add: ring_simps) +qed + +lemma cramer_lemma: + fixes A :: "'a::ordered_idom ^'n^'n::finite" + shows "det((\ i j. if j = k then (A *v x)$i else A$i$j):: 'a^'n^'n) = x$k * det A" +proof- + let ?U = "UNIV :: 'n set" + have stupid: "\c. setsum (\i. c i *s row i (transp A)) ?U = setsum (\i. c i *s column i A) ?U" + by (auto simp add: row_transp intro: setsum_cong2) + show ?thesis unfolding matrix_mult_vsum + unfolding cramer_lemma_transp[of k x "transp A", unfolded det_transp, symmetric] + unfolding stupid[of "\i. x$i"] + apply (subst det_transp[symmetric]) + apply (rule cong[OF refl[of det]]) by (vector transp_def column_def row_def) +qed + +lemma cramer: + fixes A ::"real^'n^'n::finite" + assumes d0: "det A \ 0" + shows "A *v x = b \ x = (\ k. det(\ i j. if j=k then b$i else A$i$j :: real^'n^'n) / det A)" +proof- + from d0 obtain B where B: "A ** B = mat 1" "B ** A = mat 1" + unfolding invertible_det_nz[symmetric] invertible_def by blast + have "(A ** B) *v b = b" by (simp add: B matrix_vector_mul_lid) + hence "A *v (B *v b) = b" by (simp add: matrix_vector_mul_assoc) + then have xe: "\x. A*v x = b" by blast + {fix x assume x: "A *v x = b" + have "x = (\ k. det(\ i j. if j=k then b$i else A$i$j :: real^'n^'n) / det A)" + unfolding x[symmetric] + using d0 by (simp add: Cart_eq cramer_lemma field_simps)} + with xe show ?thesis by auto +qed + +(* ------------------------------------------------------------------------- *) +(* Orthogonality of a transformation and matrix. *) +(* ------------------------------------------------------------------------- *) + +definition "orthogonal_transformation f \ linear f \ (\v w. f v \ f w = v \ w)" + +lemma orthogonal_transformation: "orthogonal_transformation f \ linear f \ (\(v::real ^_). norm (f v) = norm v)" + unfolding orthogonal_transformation_def + apply auto + apply (erule_tac x=v in allE)+ + apply (simp add: real_vector_norm_def) + by (simp add: dot_norm linear_add[symmetric]) + +definition "orthogonal_matrix (Q::'a::semiring_1^'n^'n) \ transp Q ** Q = mat 1 \ Q ** transp Q = mat 1" + +lemma orthogonal_matrix: "orthogonal_matrix (Q:: real ^'n^'n::finite) \ transp Q ** Q = mat 1" + by (metis matrix_left_right_inverse orthogonal_matrix_def) + +lemma orthogonal_matrix_id: "orthogonal_matrix (mat 1 :: _^'n^'n::finite)" + by (simp add: orthogonal_matrix_def transp_mat matrix_mul_lid) + +lemma orthogonal_matrix_mul: + fixes A :: "real ^'n^'n::finite" + assumes oA : "orthogonal_matrix A" + and oB: "orthogonal_matrix B" + shows "orthogonal_matrix(A ** B)" + using oA oB + unfolding orthogonal_matrix matrix_transp_mul + apply (subst matrix_mul_assoc) + apply (subst matrix_mul_assoc[symmetric]) + by (simp add: matrix_mul_rid) + +lemma orthogonal_transformation_matrix: + fixes f:: "real^'n \ real^'n::finite" + shows "orthogonal_transformation f \ linear f \ orthogonal_matrix(matrix f)" + (is "?lhs \ ?rhs") +proof- + let ?mf = "matrix f" + let ?ot = "orthogonal_transformation f" + let ?U = "UNIV :: 'n set" + have fU: "finite ?U" by simp + let ?m1 = "mat 1 :: real ^'n^'n" + {assume ot: ?ot + from ot have lf: "linear f" and fd: "\v w. f v \ f w = v \ w" + unfolding orthogonal_transformation_def orthogonal_matrix by blast+ + {fix i j + let ?A = "transp ?mf ** ?mf" + have th0: "\b (x::'a::comm_ring_1). (if b then 1 else 0)*x = (if b then x else 0)" + "\b (x::'a::comm_ring_1). x*(if b then 1 else 0) = (if b then x else 0)" + by simp_all + from fd[rule_format, of "basis i" "basis j", unfolded matrix_works[OF lf, symmetric] dot_matrix_vector_mul] + have "?A$i$j = ?m1 $ i $ j" + by (simp add: dot_def matrix_matrix_mult_def columnvector_def rowvector_def basis_def th0 setsum_delta[OF fU] mat_def)} + hence "orthogonal_matrix ?mf" unfolding orthogonal_matrix by vector + with lf have ?rhs by blast} + moreover + {assume lf: "linear f" and om: "orthogonal_matrix ?mf" + from lf om have ?lhs + unfolding orthogonal_matrix_def norm_eq orthogonal_transformation + unfolding matrix_works[OF lf, symmetric] + apply (subst dot_matrix_vector_mul) + by (simp add: dot_matrix_product matrix_mul_lid)} + ultimately show ?thesis by blast +qed + +lemma det_orthogonal_matrix: + fixes Q:: "'a::ordered_idom^'n^'n::finite" + assumes oQ: "orthogonal_matrix Q" + shows "det Q = 1 \ det Q = - 1" +proof- + + have th: "\x::'a. x = 1 \ x = - 1 \ x*x = 1" (is "\x::'a. ?ths x") + proof- + fix x:: 'a + have th0: "x*x - 1 = (x - 1)*(x + 1)" by (simp add: ring_simps) + have th1: "\(x::'a) y. x = - y \ x + y = 0" + apply (subst eq_iff_diff_eq_0) by simp + have "x*x = 1 \ x*x - 1 = 0" by simp + also have "\ \ x = 1 \ x = - 1" unfolding th0 th1 by simp + finally show "?ths x" .. + qed + from oQ have "Q ** transp Q = mat 1" by (metis orthogonal_matrix_def) + hence "det (Q ** transp Q) = det (mat 1:: 'a^'n^'n)" by simp + hence "det Q * det Q = 1" by (simp add: det_mul det_I det_transp) + then show ?thesis unfolding th . +qed + +(* ------------------------------------------------------------------------- *) +(* Linearity of scaling, and hence isometry, that preserves origin. *) +(* ------------------------------------------------------------------------- *) +lemma scaling_linear: + fixes f :: "real ^'n \ real ^'n::finite" + assumes f0: "f 0 = 0" and fd: "\x y. dist (f x) (f y) = c * dist x y" + shows "linear f" +proof- + {fix v w + {fix x note fd[rule_format, of x 0, unfolded dist_norm f0 diff_0_right] } + note th0 = this + have "f v \ f w = c^2 * (v \ w)" + unfolding dot_norm_neg dist_norm[symmetric] + unfolding th0 fd[rule_format] by (simp add: power2_eq_square field_simps)} + note fc = this + show ?thesis unfolding linear_def vector_eq + by (simp add: dot_lmult dot_ladd dot_rmult dot_radd fc ring_simps) +qed + +lemma isometry_linear: + "f (0:: real^'n) = (0:: real^'n::finite) \ \x y. dist(f x) (f y) = dist x y + \ linear f" +by (rule scaling_linear[where c=1]) simp_all + +(* ------------------------------------------------------------------------- *) +(* Hence another formulation of orthogonal transformation. *) +(* ------------------------------------------------------------------------- *) + +lemma orthogonal_transformation_isometry: + "orthogonal_transformation f \ f(0::real^'n) = (0::real^'n::finite) \ (\x y. dist(f x) (f y) = dist x y)" + unfolding orthogonal_transformation + apply (rule iffI) + apply clarify + apply (clarsimp simp add: linear_0 linear_sub[symmetric] dist_norm) + apply (rule conjI) + apply (rule isometry_linear) + apply simp + apply simp + apply clarify + apply (erule_tac x=v in allE) + apply (erule_tac x=0 in allE) + by (simp add: dist_norm) + +(* ------------------------------------------------------------------------- *) +(* Can extend an isometry from unit sphere. *) +(* ------------------------------------------------------------------------- *) + +lemma isometry_sphere_extend: + fixes f:: "real ^'n \ real ^'n::finite" + assumes f1: "\x. norm x = 1 \ norm (f x) = 1" + and fd1: "\ x y. norm x = 1 \ norm y = 1 \ dist (f x) (f y) = dist x y" + shows "\g. orthogonal_transformation g \ (\x. norm x = 1 \ g x = f x)" +proof- + {fix x y x' y' x0 y0 x0' y0' :: "real ^'n" + assume H: "x = norm x *s x0" "y = norm y *s y0" + "x' = norm x *s x0'" "y' = norm y *s y0'" + "norm x0 = 1" "norm x0' = 1" "norm y0 = 1" "norm y0' = 1" + "norm(x0' - y0') = norm(x0 - y0)" + + have "norm(x' - y') = norm(x - y)" + apply (subst H(1)) + apply (subst H(2)) + apply (subst H(3)) + apply (subst H(4)) + using H(5-9) + apply (simp add: norm_eq norm_eq_1) + apply (simp add: dot_lsub dot_rsub dot_lmult dot_rmult) + apply (simp add: ring_simps) + by (simp only: right_distrib[symmetric])} + note th0 = this + let ?g = "\x. if x = 0 then 0 else norm x *s f (inverse (norm x) *s x)" + {fix x:: "real ^'n" assume nx: "norm x = 1" + have "?g x = f x" using nx by auto} + hence thfg: "\x. norm x = 1 \ ?g x = f x" by blast + have g0: "?g 0 = 0" by simp + {fix x y :: "real ^'n" + {assume "x = 0" "y = 0" + then have "dist (?g x) (?g y) = dist x y" by simp } + moreover + {assume "x = 0" "y \ 0" + then have "dist (?g x) (?g y) = dist x y" + apply (simp add: dist_norm norm_mul) + apply (rule f1[rule_format]) + by(simp add: norm_mul field_simps)} + moreover + {assume "x \ 0" "y = 0" + then have "dist (?g x) (?g y) = dist x y" + apply (simp add: dist_norm norm_mul) + apply (rule f1[rule_format]) + by(simp add: norm_mul field_simps)} + moreover + {assume z: "x \ 0" "y \ 0" + have th00: "x = norm x *s (inverse (norm x) *s x)" "y = norm y *s (inverse (norm y) *s y)" "norm x *s f ((inverse (norm x) *s x)) = norm x *s f (inverse (norm x) *s x)" + "norm y *s f (inverse (norm y) *s y) = norm y *s f (inverse (norm y) *s y)" + "norm (inverse (norm x) *s x) = 1" + "norm (f (inverse (norm x) *s x)) = 1" + "norm (inverse (norm y) *s y) = 1" + "norm (f (inverse (norm y) *s y)) = 1" + "norm (f (inverse (norm x) *s x) - f (inverse (norm y) *s y)) = + norm (inverse (norm x) *s x - inverse (norm y) *s y)" + using z + by (auto simp add: vector_smult_assoc field_simps norm_mul intro: f1[rule_format] fd1[rule_format, unfolded dist_norm]) + from z th0[OF th00] have "dist (?g x) (?g y) = dist x y" + by (simp add: dist_norm)} + ultimately have "dist (?g x) (?g y) = dist x y" by blast} + note thd = this + show ?thesis + apply (rule exI[where x= ?g]) + unfolding orthogonal_transformation_isometry + using g0 thfg thd by metis +qed + +(* ------------------------------------------------------------------------- *) +(* Rotation, reflection, rotoinversion. *) +(* ------------------------------------------------------------------------- *) + +definition "rotation_matrix Q \ orthogonal_matrix Q \ det Q = 1" +definition "rotoinversion_matrix Q \ orthogonal_matrix Q \ det Q = - 1" + +lemma orthogonal_rotation_or_rotoinversion: + fixes Q :: "'a::ordered_idom^'n^'n::finite" + shows " orthogonal_matrix Q \ rotation_matrix Q \ rotoinversion_matrix Q" + by (metis rotoinversion_matrix_def rotation_matrix_def det_orthogonal_matrix) +(* ------------------------------------------------------------------------- *) +(* Explicit formulas for low dimensions. *) +(* ------------------------------------------------------------------------- *) + +lemma setprod_1: "setprod f {(1::nat)..1} = f 1" by simp + +lemma setprod_2: "setprod f {(1::nat)..2} = f 1 * f 2" + by (simp add: nat_number setprod_numseg mult_commute) +lemma setprod_3: "setprod f {(1::nat)..3} = f 1 * f 2 * f 3" + by (simp add: nat_number setprod_numseg mult_commute) + +lemma det_1: "det (A::'a::comm_ring_1^1^1) = A$1$1" + by (simp add: det_def permutes_sing sign_id UNIV_1) + +lemma det_2: "det (A::'a::comm_ring_1^2^2) = A$1$1 * A$2$2 - A$1$2 * A$2$1" +proof- + have f12: "finite {2::2}" "1 \ {2::2}" by auto + show ?thesis + unfolding det_def UNIV_2 + unfolding setsum_over_permutations_insert[OF f12] + unfolding permutes_sing + apply (simp add: sign_swap_id sign_id swap_id_eq) + by (simp add: arith_simps(31)[symmetric] of_int_minus of_int_1 del: arith_simps(31)) +qed + +lemma det_3: "det (A::'a::comm_ring_1^3^3) = + A$1$1 * A$2$2 * A$3$3 + + A$1$2 * A$2$3 * A$3$1 + + A$1$3 * A$2$1 * A$3$2 - + A$1$1 * A$2$3 * A$3$2 - + A$1$2 * A$2$1 * A$3$3 - + A$1$3 * A$2$2 * A$3$1" +proof- + have f123: "finite {2::3, 3}" "1 \ {2::3, 3}" by auto + have f23: "finite {3::3}" "2 \ {3::3}" by auto + + show ?thesis + unfolding det_def UNIV_3 + unfolding setsum_over_permutations_insert[OF f123] + unfolding setsum_over_permutations_insert[OF f23] + + unfolding permutes_sing + apply (simp add: sign_swap_id permutation_swap_id sign_compose sign_id swap_id_eq) + apply (simp add: arith_simps(31)[symmetric] of_int_minus of_int_1 del: arith_simps(31)) + by (simp add: ring_simps) +qed + +end diff -r 1fad3160d873 -r 2083bde13ce1 src/HOL/Multivariate_Analysis/Euclidean_Space.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Fri Oct 23 13:23:18 2009 +0200 @@ -0,0 +1,5372 @@ +(* Title: Library/Euclidean_Space + Author: Amine Chaieb, University of Cambridge +*) + +header {* (Real) Vectors in Euclidean space, and elementary linear algebra.*} + +theory Euclidean_Space +imports + Complex_Main "~~/src/HOL/Decision_Procs/Dense_Linear_Order" + Finite_Cartesian_Product Glbs Infinite_Set Numeral_Type + Inner_Product +uses "positivstellensatz.ML" ("normarith.ML") +begin + +text{* Some common special cases.*} + +lemma forall_1: "(\i::1. P i) \ P 1" + by (metis num1_eq_iff) + +lemma exhaust_2: + fixes x :: 2 shows "x = 1 \ x = 2" +proof (induct x) + case (of_int z) + then have "0 <= z" and "z < 2" by simp_all + then have "z = 0 | z = 1" by arith + then show ?case by auto +qed + +lemma forall_2: "(\i::2. P i) \ P 1 \ P 2" + by (metis exhaust_2) + +lemma exhaust_3: + fixes x :: 3 shows "x = 1 \ x = 2 \ x = 3" +proof (induct x) + case (of_int z) + then have "0 <= z" and "z < 3" by simp_all + then have "z = 0 \ z = 1 \ z = 2" by arith + then show ?case by auto +qed + +lemma forall_3: "(\i::3. P i) \ P 1 \ P 2 \ P 3" + by (metis exhaust_3) + +lemma UNIV_1: "UNIV = {1::1}" + by (auto simp add: num1_eq_iff) + +lemma UNIV_2: "UNIV = {1::2, 2::2}" + using exhaust_2 by auto + +lemma UNIV_3: "UNIV = {1::3, 2::3, 3::3}" + using exhaust_3 by auto + +lemma setsum_1: "setsum f (UNIV::1 set) = f 1" + unfolding UNIV_1 by simp + +lemma setsum_2: "setsum f (UNIV::2 set) = f 1 + f 2" + unfolding UNIV_2 by simp + +lemma setsum_3: "setsum f (UNIV::3 set) = f 1 + f 2 + f 3" + unfolding UNIV_3 by (simp add: add_ac) + +subsection{* Basic componentwise operations on vectors. *} + +instantiation "^" :: (plus,type) plus +begin +definition vector_add_def : "op + \ (\ x y. (\ i. (x$i) + (y$i)))" +instance .. +end + +instantiation "^" :: (times,type) times +begin + definition vector_mult_def : "op * \ (\ x y. (\ i. (x$i) * (y$i)))" + instance .. +end + +instantiation "^" :: (minus,type) minus begin + definition vector_minus_def : "op - \ (\ x y. (\ i. (x$i) - (y$i)))" +instance .. +end + +instantiation "^" :: (uminus,type) uminus begin + definition vector_uminus_def : "uminus \ (\ x. (\ i. - (x$i)))" +instance .. +end +instantiation "^" :: (zero,type) zero begin + definition vector_zero_def : "0 \ (\ i. 0)" +instance .. +end + +instantiation "^" :: (one,type) one begin + definition vector_one_def : "1 \ (\ i. 1)" +instance .. +end + +instantiation "^" :: (ord,type) ord + begin +definition vector_less_eq_def: + "less_eq (x :: 'a ^'b) y = (ALL i. x$i <= y$i)" +definition vector_less_def: "less (x :: 'a ^'b) y = (ALL i. x$i < y$i)" + +instance by (intro_classes) +end + +instantiation "^" :: (scaleR, type) scaleR +begin +definition vector_scaleR_def: "scaleR = (\ r x. (\ i. scaleR r (x$i)))" +instance .. +end + +text{* Also the scalar-vector multiplication. *} + +definition vector_scalar_mult:: "'a::times \ 'a ^'n \ 'a ^ 'n" (infixl "*s" 70) + where "c *s x = (\ i. c * (x$i))" + +text{* Constant Vectors *} + +definition "vec x = (\ i. x)" + +text{* Dot products. *} + +definition dot :: "'a::{comm_monoid_add, times} ^ 'n \ 'a ^ 'n \ 'a" (infix "\" 70) where + "x \ y = setsum (\i. x$i * y$i) UNIV" + +lemma dot_1[simp]: "(x::'a::{comm_monoid_add, times}^1) \ y = (x$1) * (y$1)" + by (simp add: dot_def setsum_1) + +lemma dot_2[simp]: "(x::'a::{comm_monoid_add, times}^2) \ y = (x$1) * (y$1) + (x$2) * (y$2)" + by (simp add: dot_def setsum_2) + +lemma dot_3[simp]: "(x::'a::{comm_monoid_add, times}^3) \ y = (x$1) * (y$1) + (x$2) * (y$2) + (x$3) * (y$3)" + by (simp add: dot_def setsum_3) + +subsection {* A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space. *} + +method_setup vector = {* +let + val ss1 = HOL_basic_ss addsimps [@{thm dot_def}, @{thm setsum_addf} RS sym, + @{thm setsum_subtractf} RS sym, @{thm setsum_right_distrib}, + @{thm setsum_left_distrib}, @{thm setsum_negf} RS sym] + val ss2 = @{simpset} addsimps + [@{thm vector_add_def}, @{thm vector_mult_def}, + @{thm vector_minus_def}, @{thm vector_uminus_def}, + @{thm vector_one_def}, @{thm vector_zero_def}, @{thm vec_def}, + @{thm vector_scaleR_def}, + @{thm Cart_lambda_beta}, @{thm vector_scalar_mult_def}] + fun vector_arith_tac ths = + simp_tac ss1 + THEN' (fn i => rtac @{thm setsum_cong2} i + ORELSE rtac @{thm setsum_0'} i + ORELSE simp_tac (HOL_basic_ss addsimps [@{thm "Cart_eq"}]) i) + (* THEN' TRY o clarify_tac HOL_cs THEN' (TRY o rtac @{thm iffI}) *) + THEN' asm_full_simp_tac (ss2 addsimps ths) + in + Attrib.thms >> (fn ths => K (SIMPLE_METHOD' (vector_arith_tac ths))) + end +*} "Lifts trivial vector statements to real arith statements" + +lemma vec_0[simp]: "vec 0 = 0" by (vector vector_zero_def) +lemma vec_1[simp]: "vec 1 = 1" by (vector vector_one_def) + + + +text{* Obvious "component-pushing". *} + +lemma vec_component [simp]: "(vec x :: 'a ^ 'n)$i = x" + by (vector vec_def) + +lemma vector_add_component [simp]: + fixes x y :: "'a::{plus} ^ 'n" + shows "(x + y)$i = x$i + y$i" + by vector + +lemma vector_minus_component [simp]: + fixes x y :: "'a::{minus} ^ 'n" + shows "(x - y)$i = x$i - y$i" + by vector + +lemma vector_mult_component [simp]: + fixes x y :: "'a::{times} ^ 'n" + shows "(x * y)$i = x$i * y$i" + by vector + +lemma vector_smult_component [simp]: + fixes y :: "'a::{times} ^ 'n" + shows "(c *s y)$i = c * (y$i)" + by vector + +lemma vector_uminus_component [simp]: + fixes x :: "'a::{uminus} ^ 'n" + shows "(- x)$i = - (x$i)" + by vector + +lemma vector_scaleR_component [simp]: + fixes x :: "'a::scaleR ^ 'n" + shows "(scaleR r x)$i = scaleR r (x$i)" + by vector + +lemma cond_component: "(if b then x else y)$i = (if b then x$i else y$i)" by vector + +lemmas vector_component = + vec_component vector_add_component vector_mult_component + vector_smult_component vector_minus_component vector_uminus_component + vector_scaleR_component cond_component + +subsection {* Some frequently useful arithmetic lemmas over vectors. *} + +instance "^" :: (semigroup_add,type) semigroup_add + apply (intro_classes) by (vector add_assoc) + + +instance "^" :: (monoid_add,type) monoid_add + apply (intro_classes) by vector+ + +instance "^" :: (group_add,type) group_add + apply (intro_classes) by (vector algebra_simps)+ + +instance "^" :: (ab_semigroup_add,type) ab_semigroup_add + apply (intro_classes) by (vector add_commute) + +instance "^" :: (comm_monoid_add,type) comm_monoid_add + apply (intro_classes) by vector + +instance "^" :: (ab_group_add,type) ab_group_add + apply (intro_classes) by vector+ + +instance "^" :: (cancel_semigroup_add,type) cancel_semigroup_add + apply (intro_classes) + by (vector Cart_eq)+ + +instance "^" :: (cancel_ab_semigroup_add,type) cancel_ab_semigroup_add + apply (intro_classes) + by (vector Cart_eq) + +instance "^" :: (real_vector, type) real_vector + by default (vector scaleR_left_distrib scaleR_right_distrib)+ + +instance "^" :: (semigroup_mult,type) semigroup_mult + apply (intro_classes) by (vector mult_assoc) + +instance "^" :: (monoid_mult,type) monoid_mult + apply (intro_classes) by vector+ + +instance "^" :: (ab_semigroup_mult,type) ab_semigroup_mult + apply (intro_classes) by (vector mult_commute) + +instance "^" :: (ab_semigroup_idem_mult,type) ab_semigroup_idem_mult + apply (intro_classes) by (vector mult_idem) + +instance "^" :: (comm_monoid_mult,type) comm_monoid_mult + apply (intro_classes) by vector + +fun vector_power :: "('a::{one,times} ^'n) \ nat \ 'a^'n" where + "vector_power x 0 = 1" + | "vector_power x (Suc n) = x * vector_power x n" + +instance "^" :: (semiring,type) semiring + apply (intro_classes) by (vector ring_simps)+ + +instance "^" :: (semiring_0,type) semiring_0 + apply (intro_classes) by (vector ring_simps)+ +instance "^" :: (semiring_1,type) semiring_1 + apply (intro_classes) by vector +instance "^" :: (comm_semiring,type) comm_semiring + apply (intro_classes) by (vector ring_simps)+ + +instance "^" :: (comm_semiring_0,type) comm_semiring_0 by (intro_classes) +instance "^" :: (cancel_comm_monoid_add, type) cancel_comm_monoid_add .. +instance "^" :: (semiring_0_cancel,type) semiring_0_cancel by (intro_classes) +instance "^" :: (comm_semiring_0_cancel,type) comm_semiring_0_cancel by (intro_classes) +instance "^" :: (ring,type) ring by (intro_classes) +instance "^" :: (semiring_1_cancel,type) semiring_1_cancel by (intro_classes) +instance "^" :: (comm_semiring_1,type) comm_semiring_1 by (intro_classes) + +instance "^" :: (ring_1,type) ring_1 .. + +instance "^" :: (real_algebra,type) real_algebra + apply intro_classes + apply (simp_all add: vector_scaleR_def ring_simps) + apply vector + apply vector + done + +instance "^" :: (real_algebra_1,type) real_algebra_1 .. + +lemma of_nat_index: + "(of_nat n :: 'a::semiring_1 ^'n)$i = of_nat n" + apply (induct n) + apply vector + apply vector + done +lemma zero_index[simp]: + "(0 :: 'a::zero ^'n)$i = 0" by vector + +lemma one_index[simp]: + "(1 :: 'a::one ^'n)$i = 1" by vector + +lemma one_plus_of_nat_neq_0: "(1::'a::semiring_char_0) + of_nat n \ 0" +proof- + have "(1::'a) + of_nat n = 0 \ of_nat 1 + of_nat n = (of_nat 0 :: 'a)" by simp + also have "\ \ 1 + n = 0" by (simp only: of_nat_add[symmetric] of_nat_eq_iff) + finally show ?thesis by simp +qed + +instance "^" :: (semiring_char_0,type) semiring_char_0 +proof (intro_classes) + fix m n ::nat + show "(of_nat m :: 'a^'b) = of_nat n \ m = n" + by (simp add: Cart_eq of_nat_index) +qed + +instance "^" :: (comm_ring_1,type) comm_ring_1 by intro_classes +instance "^" :: (ring_char_0,type) ring_char_0 by intro_classes + +lemma vector_smult_assoc: "a *s (b *s x) = ((a::'a::semigroup_mult) * b) *s x" + by (vector mult_assoc) +lemma vector_sadd_rdistrib: "((a::'a::semiring) + b) *s x = a *s x + b *s x" + by (vector ring_simps) +lemma vector_add_ldistrib: "(c::'a::semiring) *s (x + y) = c *s x + c *s y" + by (vector ring_simps) +lemma vector_smult_lzero[simp]: "(0::'a::mult_zero) *s x = 0" by vector +lemma vector_smult_lid[simp]: "(1::'a::monoid_mult) *s x = x" by vector +lemma vector_ssub_ldistrib: "(c::'a::ring) *s (x - y) = c *s x - c *s y" + by (vector ring_simps) +lemma vector_smult_rneg: "(c::'a::ring) *s -x = -(c *s x)" by vector +lemma vector_smult_lneg: "- (c::'a::ring) *s x = -(c *s x)" by vector +lemma vector_sneg_minus1: "-x = (- (1::'a::ring_1)) *s x" by vector +lemma vector_smult_rzero[simp]: "c *s 0 = (0::'a::mult_zero ^ 'n)" by vector +lemma vector_sub_rdistrib: "((a::'a::ring) - b) *s x = a *s x - b *s x" + by (vector ring_simps) + +lemma vec_eq[simp]: "(vec m = vec n) \ (m = n)" + by (simp add: Cart_eq) + +subsection {* Topological space *} + +instantiation "^" :: (topological_space, finite) topological_space +begin + +definition open_vector_def: + "open (S :: ('a ^ 'b) set) \ + (\x\S. \A. (\i. open (A i) \ x$i \ A i) \ + (\y. (\i. y$i \ A i) \ y \ S))" + +instance proof + show "open (UNIV :: ('a ^ 'b) set)" + unfolding open_vector_def by auto +next + fix S T :: "('a ^ 'b) set" + assume "open S" "open T" thus "open (S \ T)" + unfolding open_vector_def + apply clarify + apply (drule (1) bspec)+ + apply (clarify, rename_tac Sa Ta) + apply (rule_tac x="\i. Sa i \ Ta i" in exI) + apply (simp add: open_Int) + done +next + fix K :: "('a ^ 'b) set set" + assume "\S\K. open S" thus "open (\K)" + unfolding open_vector_def + apply clarify + apply (drule (1) bspec) + apply (drule (1) bspec) + apply clarify + apply (rule_tac x=A in exI) + apply fast + done +qed + +end + +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 assume "open S" "a $ i \ S" + then have "open ((\y. y $ i) -` S)" "a \ ((\y. y $ i) -` S)" + 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" + by simp +qed + +subsection {* Square root of sum of squares *} + +definition + "setL2 f A = sqrt (\i\A. (f i)\)" + +lemma setL2_cong: + "\A = B; \x. x \ B \ f x = g x\ \ setL2 f A = setL2 g B" + unfolding setL2_def by simp + +lemma strong_setL2_cong: + "\A = B; \x. x \ B =simp=> f x = g x\ \ setL2 f A = setL2 g B" + unfolding setL2_def simp_implies_def by simp + +lemma setL2_infinite [simp]: "\ finite A \ setL2 f A = 0" + unfolding setL2_def by simp + +lemma setL2_empty [simp]: "setL2 f {} = 0" + unfolding setL2_def by simp + +lemma setL2_insert [simp]: + "\finite F; a \ F\ \ + setL2 f (insert a F) = sqrt ((f a)\ + (setL2 f F)\)" + unfolding setL2_def by (simp add: setsum_nonneg) + +lemma setL2_nonneg [simp]: "0 \ setL2 f A" + unfolding setL2_def by (simp add: setsum_nonneg) + +lemma setL2_0': "\a\A. f a = 0 \ setL2 f A = 0" + unfolding setL2_def by simp + +lemma setL2_constant: "setL2 (\x. y) A = sqrt (of_nat (card A)) * \y\" + unfolding setL2_def by (simp add: real_sqrt_mult) + +lemma setL2_mono: + assumes "\i. i \ K \ f i \ g i" + assumes "\i. i \ K \ 0 \ f i" + shows "setL2 f K \ setL2 g K" + unfolding setL2_def + by (simp add: setsum_nonneg setsum_mono power_mono prems) + +lemma setL2_strict_mono: + assumes "finite K" and "K \ {}" + assumes "\i. i \ K \ f i < g i" + assumes "\i. i \ K \ 0 \ f i" + shows "setL2 f K < setL2 g K" + unfolding setL2_def + by (simp add: setsum_strict_mono power_strict_mono assms) + +lemma setL2_right_distrib: + "0 \ r \ r * setL2 f A = setL2 (\x. r * f x) A" + unfolding setL2_def + apply (simp add: power_mult_distrib) + apply (simp add: setsum_right_distrib [symmetric]) + apply (simp add: real_sqrt_mult setsum_nonneg) + done + +lemma setL2_left_distrib: + "0 \ r \ setL2 f A * r = setL2 (\x. f x * r) A" + unfolding setL2_def + apply (simp add: power_mult_distrib) + apply (simp add: setsum_left_distrib [symmetric]) + apply (simp add: real_sqrt_mult setsum_nonneg) + done + +lemma setsum_nonneg_eq_0_iff: + fixes f :: "'a \ 'b::pordered_ab_group_add" + shows "\finite A; \x\A. 0 \ f x\ \ setsum f A = 0 \ (\x\A. f x = 0)" + apply (induct set: finite, simp) + apply (simp add: add_nonneg_eq_0_iff setsum_nonneg) + done + +lemma setL2_eq_0_iff: "finite A \ setL2 f A = 0 \ (\x\A. f x = 0)" + unfolding setL2_def + by (simp add: setsum_nonneg setsum_nonneg_eq_0_iff) + +lemma setL2_triangle_ineq: + shows "setL2 (\i. f i + g i) A \ setL2 f A + setL2 g A" +proof (cases "finite A") + case False + thus ?thesis by simp +next + case True + thus ?thesis + proof (induct set: finite) + case empty + show ?case by simp + next + case (insert x F) + hence "sqrt ((f x + g x)\ + (setL2 (\i. f i + g i) F)\) \ + sqrt ((f x + g x)\ + (setL2 f F + setL2 g F)\)" + by (intro real_sqrt_le_mono add_left_mono power_mono insert + setL2_nonneg add_increasing zero_le_power2) + also have + "\ \ sqrt ((f x)\ + (setL2 f F)\) + sqrt ((g x)\ + (setL2 g F)\)" + by (rule real_sqrt_sum_squares_triangle_ineq) + finally show ?case + using insert by simp + qed +qed + +lemma sqrt_sum_squares_le_sum: + "\0 \ x; 0 \ y\ \ sqrt (x\ + y\) \ x + y" + apply (rule power2_le_imp_le) + apply (simp add: power2_sum) + apply (simp add: mult_nonneg_nonneg) + apply (simp add: add_nonneg_nonneg) + done + +lemma setL2_le_setsum [rule_format]: + "(\i\A. 0 \ f i) \ setL2 f A \ setsum f A" + apply (cases "finite A") + apply (induct set: finite) + apply simp + apply clarsimp + apply (erule order_trans [OF sqrt_sum_squares_le_sum]) + apply simp + apply simp + apply simp + done + +lemma sqrt_sum_squares_le_sum_abs: "sqrt (x\ + y\) \ \x\ + \y\" + apply (rule power2_le_imp_le) + apply (simp add: power2_sum) + apply (simp add: mult_nonneg_nonneg) + apply (simp add: add_nonneg_nonneg) + done + +lemma setL2_le_setsum_abs: "setL2 f A \ (\i\A. \f i\)" + apply (cases "finite A") + apply (induct set: finite) + apply simp + apply simp + apply (rule order_trans [OF sqrt_sum_squares_le_sum_abs]) + apply simp + apply simp + done + +lemma setL2_mult_ineq_lemma: + fixes a b c d :: real + shows "2 * (a * c) * (b * d) \ a\ * d\ + b\ * c\" +proof - + have "0 \ (a * d - b * c)\" by simp + also have "\ = a\ * d\ + b\ * c\ - 2 * (a * d) * (b * c)" + by (simp only: power2_diff power_mult_distrib) + also have "\ = a\ * d\ + b\ * c\ - 2 * (a * c) * (b * d)" + by simp + finally show "2 * (a * c) * (b * d) \ a\ * d\ + b\ * c\" + by simp +qed + +lemma setL2_mult_ineq: "(\i\A. \f i\ * \g i\) \ setL2 f A * setL2 g A" + apply (cases "finite A") + apply (induct set: finite) + apply simp + apply (rule power2_le_imp_le, simp) + apply (rule order_trans) + apply (rule power_mono) + apply (erule add_left_mono) + apply (simp add: add_nonneg_nonneg mult_nonneg_nonneg setsum_nonneg) + apply (simp add: power2_sum) + apply (simp add: power_mult_distrib) + apply (simp add: right_distrib left_distrib) + apply (rule ord_le_eq_trans) + apply (rule setL2_mult_ineq_lemma) + apply simp + apply (intro mult_nonneg_nonneg setL2_nonneg) + apply simp + done + +lemma member_le_setL2: "\finite A; i \ A\ \ f i \ setL2 f A" + apply (rule_tac s="insert i (A - {i})" and t="A" in subst) + apply fast + apply (subst setL2_insert) + apply simp + apply simp + apply simp + done + +subsection {* Metric *} + +(* TODO: move somewhere else *) +lemma finite_choice: "finite A \ \x\A. \y. P x y \ \f. \x\A. P x (f x)" +apply (induct set: finite, simp_all) +apply (clarify, rename_tac y) +apply (rule_tac x="f(x:=y)" in exI, simp) +done + +instantiation "^" :: (metric_space, finite) metric_space +begin + +definition dist_vector_def: + "dist (x::'a^'b) (y::'a^'b) = setL2 (\i. dist (x$i) (y$i)) UNIV" + +lemma dist_nth_le: "dist (x $ i) (y $ i) \ dist x y" +unfolding dist_vector_def +by (rule member_le_setL2) simp_all + +instance proof + fix x y :: "'a ^ 'b" + show "dist x y = 0 \ x = y" + unfolding dist_vector_def + by (simp add: setL2_eq_0_iff Cart_eq) +next + fix x y z :: "'a ^ 'b" + show "dist x y \ dist x z + dist y z" + unfolding dist_vector_def + apply (rule order_trans [OF _ setL2_triangle_ineq]) + apply (simp add: setL2_mono dist_triangle2) + done +next + (* FIXME: long proof! *) + fix S :: "('a ^ 'b) set" + show "open S \ (\x\S. \e>0. \y. dist y x < e \ y \ S)" + unfolding open_vector_def open_dist + apply safe + apply (drule (1) bspec) + apply clarify + apply (subgoal_tac "\e>0. \i y. dist y (x$i) < e \ y \ A i") + apply clarify + apply (rule_tac x=e in exI, clarify) + apply (drule spec, erule mp, clarify) + apply (drule spec, drule spec, erule mp) + apply (erule le_less_trans [OF dist_nth_le]) + apply (subgoal_tac "\i\UNIV. \e>0. \y. dist y (x$i) < e \ y \ A i") + apply (drule finite_choice [OF finite], clarify) + apply (rule_tac x="Min (range f)" in exI, simp) + apply clarify + apply (drule_tac x=i in spec, clarify) + apply (erule (1) bspec) + apply (drule (1) bspec, clarify) + apply (subgoal_tac "\r. (\i::'b. 0 < r i) \ e = setL2 r UNIV") + apply clarify + apply (rule_tac x="\i. {y. dist y (x$i) < r i}" in exI) + apply (rule conjI) + apply clarify + apply (rule conjI) + apply (clarify, rename_tac y) + apply (rule_tac x="r i - dist y (x$i)" in exI, rule conjI, simp) + apply clarify + apply (simp only: less_diff_eq) + apply (erule le_less_trans [OF dist_triangle]) + apply simp + apply clarify + apply (drule spec, erule mp) + apply (simp add: dist_vector_def setL2_strict_mono) + apply (rule_tac x="\i. e / sqrt (of_nat CARD('b))" in exI) + apply (simp add: divide_pos_pos setL2_constant) + done +qed + +end + +lemma LIMSEQ_Cart_nth: + "(X ----> a) \ (\n. X n $ i) ----> a $ i" +unfolding LIMSEQ_conv_tendsto by (rule tendsto_Cart_nth) + +lemma LIM_Cart_nth: + "(f -- x --> y) \ (\x. f x $ i) -- x --> y $ i" +unfolding LIM_conv_tendsto by (rule tendsto_Cart_nth) + +lemma Cauchy_Cart_nth: + "Cauchy (\n. X n) \ Cauchy (\n. X n $ i)" +unfolding Cauchy_def by (fast intro: le_less_trans [OF dist_nth_le]) + +lemma LIMSEQ_vector: + fixes X :: "nat \ 'a::metric_space ^ 'n::finite" + assumes X: "\i. (\n. X n $ i) ----> (a $ i)" + shows "X ----> a" +proof (rule metric_LIMSEQ_I) + fix r :: real assume "0 < r" + then have "0 < r / of_nat CARD('n)" (is "0 < ?s") + by (simp add: divide_pos_pos) + def N \ "\i. LEAST N. \n\N. dist (X n $ i) (a $ i) < ?s" + def M \ "Max (range N)" + have "\i. \N. \n\N. dist (X n $ i) (a $ i) < ?s" + using X `0 < ?s` by (rule metric_LIMSEQ_D) + hence "\i. \n\N i. dist (X n $ i) (a $ i) < ?s" + unfolding N_def by (rule LeastI_ex) + hence M: "\i. \n\M. dist (X n $ i) (a $ i) < ?s" + unfolding M_def by simp + { + fix n :: nat assume "M \ n" + have "dist (X n) a = setL2 (\i. dist (X n $ i) (a $ i)) UNIV" + unfolding dist_vector_def .. + also have "\ \ setsum (\i. dist (X n $ i) (a $ i)) UNIV" + by (rule setL2_le_setsum [OF zero_le_dist]) + also have "\ < setsum (\i::'n. ?s) UNIV" + by (rule setsum_strict_mono, simp_all add: M `M \ n`) + also have "\ = r" + by simp + finally have "dist (X n) a < r" . + } + hence "\n\M. dist (X n) a < r" + by simp + then show "\M. \n\M. dist (X n) a < r" .. +qed + +lemma Cauchy_vector: + fixes X :: "nat \ 'a::metric_space ^ 'n::finite" + assumes X: "\i. Cauchy (\n. X n $ i)" + shows "Cauchy (\n. X n)" +proof (rule metric_CauchyI) + fix r :: real assume "0 < r" + then have "0 < r / of_nat CARD('n)" (is "0 < ?s") + by (simp add: divide_pos_pos) + def N \ "\i. LEAST N. \m\N. \n\N. dist (X m $ i) (X n $ i) < ?s" + def M \ "Max (range N)" + have "\i. \N. \m\N. \n\N. dist (X m $ i) (X n $ i) < ?s" + using X `0 < ?s` by (rule metric_CauchyD) + hence "\i. \m\N i. \n\N i. dist (X m $ i) (X n $ i) < ?s" + unfolding N_def by (rule LeastI_ex) + hence M: "\i. \m\M. \n\M. dist (X m $ i) (X n $ i) < ?s" + unfolding M_def by simp + { + fix m n :: nat + assume "M \ m" "M \ n" + have "dist (X m) (X n) = setL2 (\i. dist (X m $ i) (X n $ i)) UNIV" + unfolding dist_vector_def .. + also have "\ \ setsum (\i. dist (X m $ i) (X n $ i)) UNIV" + by (rule setL2_le_setsum [OF zero_le_dist]) + also have "\ < setsum (\i::'n. ?s) UNIV" + by (rule setsum_strict_mono, simp_all add: M `M \ m` `M \ n`) + also have "\ = r" + by simp + finally have "dist (X m) (X n) < r" . + } + hence "\m\M. \n\M. dist (X m) (X n) < r" + by simp + then show "\M. \m\M. \n\M. dist (X m) (X n) < r" .. +qed + +instance "^" :: (complete_space, finite) complete_space +proof + fix X :: "nat \ 'a ^ 'b" assume "Cauchy X" + have "\i. (\n. X n $ i) ----> lim (\n. X n $ i)" + using Cauchy_Cart_nth [OF `Cauchy X`] + by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) + hence "X ----> Cart_lambda (\i. lim (\n. X n $ i))" + by (simp add: LIMSEQ_vector) + then show "convergent X" + by (rule convergentI) +qed + +subsection {* Norms *} + +instantiation "^" :: (real_normed_vector, finite) real_normed_vector +begin + +definition norm_vector_def: + "norm (x::'a^'b) = setL2 (\i. norm (x$i)) UNIV" + +definition vector_sgn_def: + "sgn (x::'a^'b) = scaleR (inverse (norm x)) x" + +instance proof + fix a :: real and x y :: "'a ^ 'b" + show "0 \ norm x" + unfolding norm_vector_def + by (rule setL2_nonneg) + show "norm x = 0 \ x = 0" + unfolding norm_vector_def + by (simp add: setL2_eq_0_iff Cart_eq) + show "norm (x + y) \ norm x + norm y" + 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 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 norm_vector_def + by (simp add: dist_norm) +qed + +end + +lemma norm_nth_le: "norm (x $ i) \ norm x" +unfolding norm_vector_def +by (rule member_le_setL2) simp_all + +interpretation Cart_nth: bounded_linear "\x. x $ i" +apply default +apply (rule vector_add_component) +apply (rule vector_scaleR_component) +apply (rule_tac x="1" in exI, simp add: norm_nth_le) +done + +instance "^" :: (banach, finite) banach .. + +subsection {* Inner products *} + +instantiation "^" :: (real_inner, finite) real_inner +begin + +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 inner_vector_def + by (simp add: inner_commute) + show "inner (x + y) z = inner x z + inner y z" + unfolding inner_vector_def + by (simp add: inner_add_left setsum_addf) + show "inner (scaleR r x) y = r * inner x y" + unfolding inner_vector_def + by (simp add: setsum_right_distrib) + show "0 \ inner x x" + unfolding inner_vector_def + by (simp add: setsum_nonneg) + show "inner x x = 0 \ x = 0" + unfolding inner_vector_def + by (simp add: Cart_eq setsum_nonneg_eq_0_iff) + show "norm x = sqrt (inner x x)" + unfolding inner_vector_def norm_vector_def setL2_def + by (simp add: power2_norm_eq_inner) +qed + +end + +subsection{* Properties of the dot product. *} + +lemma dot_sym: "(x::'a:: {comm_monoid_add, ab_semigroup_mult} ^ 'n) \ y = y \ x" + by (vector mult_commute) +lemma dot_ladd: "((x::'a::ring ^ 'n) + y) \ z = (x \ z) + (y \ z)" + by (vector ring_simps) +lemma dot_radd: "x \ (y + (z::'a::ring ^ 'n)) = (x \ y) + (x \ z)" + by (vector ring_simps) +lemma dot_lsub: "((x::'a::ring ^ 'n) - y) \ z = (x \ z) - (y \ z)" + by (vector ring_simps) +lemma dot_rsub: "(x::'a::ring ^ 'n) \ (y - z) = (x \ y) - (x \ z)" + by (vector ring_simps) +lemma dot_lmult: "(c *s x) \ y = (c::'a::ring) * (x \ y)" by (vector ring_simps) +lemma dot_rmult: "x \ (c *s y) = (c::'a::comm_ring) * (x \ y)" by (vector ring_simps) +lemma dot_lneg: "(-x) \ (y::'a::ring ^ 'n) = -(x \ y)" by vector +lemma dot_rneg: "(x::'a::ring ^ 'n) \ (-y) = -(x \ y)" by vector +lemma dot_lzero[simp]: "0 \ x = (0::'a::{comm_monoid_add, mult_zero})" by vector +lemma dot_rzero[simp]: "x \ 0 = (0::'a::{comm_monoid_add, mult_zero})" by vector +lemma dot_pos_le[simp]: "(0::'a\ordered_ring_strict) <= x \ x" + by (simp add: dot_def setsum_nonneg) + +lemma setsum_squares_eq_0_iff: assumes fS: "finite F" and fp: "\x \ F. f x \ (0 ::'a::pordered_ab_group_add)" shows "setsum f F = 0 \ (ALL x:F. f x = 0)" +using fS fp setsum_nonneg[OF fp] +proof (induct set: finite) + case empty thus ?case by simp +next + case (insert x F) + from insert.prems have Fx: "f x \ 0" and Fp: "\ a \ F. f a \ 0" by simp_all + from insert.hyps Fp setsum_nonneg[OF Fp] + have h: "setsum f F = 0 \ (\a \F. f a = 0)" by metis + from add_nonneg_eq_0_iff[OF Fx setsum_nonneg[OF Fp]] insert.hyps(1,2) + show ?case by (simp add: h) +qed + +lemma dot_eq_0: "x \ x = 0 \ (x::'a::{ordered_ring_strict,ring_no_zero_divisors} ^ 'n::finite) = 0" + by (simp add: dot_def setsum_squares_eq_0_iff Cart_eq) + +lemma dot_pos_lt[simp]: "(0 < x \ x) \ (x::'a::{ordered_ring_strict,ring_no_zero_divisors} ^ 'n::finite) \ 0" using dot_eq_0[of x] dot_pos_le[of x] + by (auto simp add: le_less) + +subsection{* The collapse of the general concepts to dimension one. *} + +lemma vector_one: "(x::'a ^1) = (\ i. (x$1))" + by (simp add: Cart_eq forall_1) + +lemma forall_one: "(\(x::'a ^1). P x) \ (\x. P(\ i. x))" + apply auto + apply (erule_tac x= "x$1" in allE) + apply (simp only: vector_one[symmetric]) + done + +lemma norm_vector_1: "norm (x :: _^1) = norm (x$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) + +lemma dist_real: "dist(x::real ^ 1) y = abs((x$1) - (y$1))" + by (auto simp add: norm_real dist_norm) + +subsection {* A connectedness or intermediate value lemma with several applications. *} + +lemma connected_real_lemma: + fixes f :: "real \ 'a::metric_space" + assumes ab: "a \ b" and fa: "f a \ e1" and fb: "f b \ e2" + and dst: "\e x. a <= x \ x <= b \ 0 < e ==> \d > 0. \y. abs(y - x) < d \ dist(f y) (f x) < e" + and e1: "\y \ e1. \e > 0. \y'. dist y' y < e \ y' \ e1" + and e2: "\y \ e2. \e > 0. \y'. dist y' y < e \ y' \ e2" + and e12: "~(\x \ a. x <= b \ f x \ e1 \ f x \ e2)" + shows "\x \ a. x <= b \ f x \ e1 \ f x \ e2" (is "\ x. ?P x") +proof- + let ?S = "{c. \x \ a. x <= c \ f x \ e1}" + have Se: " \x. x \ ?S" apply (rule exI[where x=a]) by (auto simp add: fa) + have Sub: "\y. isUb UNIV ?S y" + apply (rule exI[where x= b]) + using ab fb e12 by (auto simp add: isUb_def setle_def) + from reals_complete[OF Se Sub] obtain l where + l: "isLub UNIV ?S l"by blast + have alb: "a \ l" "l \ b" using l ab fa fb e12 + apply (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def) + by (metis linorder_linear) + have ale1: "\z \ a. z < l \ f z \ e1" using l + apply (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def) + by (metis linorder_linear not_le) + have th1: "\z x e d :: real. z <= x + e \ e < d ==> z < x \ abs(z - x) < d" by arith + have th2: "\e x:: real. 0 < e ==> ~(x + e <= x)" by arith + have th3: "\d::real. d > 0 \ \e > 0. e < d" by dlo + {assume le2: "f l \ e2" + from le2 fa fb e12 alb have la: "l \ a" by metis + hence lap: "l - a > 0" using alb by arith + from e2[rule_format, OF le2] obtain e where + e: "e > 0" "\y. dist y (f l) < e \ y \ e2" by metis + from dst[OF alb e(1)] obtain d where + d: "d > 0" "\y. \y - l\ < d \ dist (f y) (f l) < e" by metis + have "\d'. d' < d \ d' >0 \ l - d' > a" using lap d(1) + apply ferrack by arith + then obtain d' where d': "d' > 0" "d' < d" "l - d' > a" by metis + from d e have th0: "\y. \y - l\ < d \ f y \ e2" by metis + from th0[rule_format, of "l - d'"] d' have "f (l - d') \ e2" by auto + moreover + have "f (l - d') \ e1" using ale1[rule_format, of "l -d'"] d' by auto + ultimately have False using e12 alb d' by auto} + moreover + {assume le1: "f l \ e1" + from le1 fa fb e12 alb have lb: "l \ b" by metis + hence blp: "b - l > 0" using alb by arith + from e1[rule_format, OF le1] obtain e where + e: "e > 0" "\y. dist y (f l) < e \ y \ e1" by metis + from dst[OF alb e(1)] obtain d where + d: "d > 0" "\y. \y - l\ < d \ dist (f y) (f l) < e" by metis + have "\d'. d' < d \ d' >0" using d(1) by dlo + then obtain d' where d': "d' > 0" "d' < d" by metis + from d e have th0: "\y. \y - l\ < d \ f y \ e1" by auto + hence "\y. l \ y \ y \ l + d' \ f y \ e1" using d' by auto + with ale1 have "\y. a \ y \ y \ l + d' \ f y \ e1" by auto + with l d' have False + by (auto simp add: isLub_def isUb_def setle_def setge_def leastP_def) } + ultimately show ?thesis using alb by metis +qed + +text{* One immediately useful corollary is the existence of square roots! --- Should help to get rid of all the development of square-root for reals as a special case @{typ "real^1"} *} + +lemma square_bound_lemma: "(x::real) < (1 + x) * (1 + x)" +proof- + have "(x + 1/2)^2 + 3/4 > 0" using zero_le_power2[of "x+1/2"] by arith + thus ?thesis by (simp add: ring_simps power2_eq_square) +qed + +lemma square_continuous: "0 < (e::real) ==> \d. 0 < d \ (\y. abs(y - x) < d \ abs(y * y - x * x) < e)" + using isCont_power[OF isCont_ident, of 2, unfolded isCont_def LIM_eq, rule_format, of e x] apply (auto simp add: power2_eq_square) + apply (rule_tac x="s" in exI) + apply auto + apply (erule_tac x=y in allE) + apply auto + done + +lemma real_le_lsqrt: "0 <= x \ 0 <= y \ x <= y^2 ==> sqrt x <= y" + using real_sqrt_le_iff[of x "y^2"] by simp + +lemma real_le_rsqrt: "x^2 \ y \ x \ sqrt y" + using real_sqrt_le_mono[of "x^2" y] by simp + +lemma real_less_rsqrt: "x^2 < y \ x < sqrt y" + using real_sqrt_less_mono[of "x^2" y] by simp + +lemma sqrt_even_pow2: assumes n: "even n" + shows "sqrt(2 ^ n) = 2 ^ (n div 2)" +proof- + from n obtain m where m: "n = 2*m" unfolding even_nat_equiv_def2 + by (auto simp add: nat_number) + from m have "sqrt(2 ^ n) = sqrt ((2 ^ m) ^ 2)" + by (simp only: power_mult[symmetric] mult_commute) + then show ?thesis using m by simp +qed + +lemma real_div_sqrt: "0 <= x ==> x / sqrt(x) = sqrt(x)" + apply (cases "x = 0", simp_all) + using sqrt_divide_self_eq[of x] + apply (simp add: inverse_eq_divide real_sqrt_ge_0_iff field_simps) + done + +text{* Hence derive more interesting properties of the norm. *} + +text {* + This type-specific version is only here + to make @{text normarith.ML} happy. +*} +lemma norm_0: "norm (0::real ^ _) = 0" + by (rule norm_zero) + +lemma norm_mul[simp]: "norm(a *s x) = abs(a) * norm x" + 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: norm_vector_def dot_def setL2_def power2_eq_square) +lemma real_vector_norm_def: "norm x = sqrt (x \ x)" + 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) +lemma vector_mul_eq_0[simp]: "(a *s x = 0) \ a = (0::'a::idom) \ x = 0" + by vector +lemma vector_mul_lcancel[simp]: "a *s x = a *s y \ a = (0::real) \ x = y" + by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_ssub_ldistrib) +lemma vector_mul_rcancel[simp]: "a *s x = b *s x \ (a::real) = b \ x = 0" + by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_sub_rdistrib) +lemma vector_mul_lcancel_imp: "a \ (0::real) ==> a *s x = a *s y ==> (x = y)" + by (metis vector_mul_lcancel) +lemma vector_mul_rcancel_imp: "x \ 0 \ (a::real) *s x = b *s x ==> a = b" + by (metis vector_mul_rcancel) +lemma norm_cauchy_schwarz: + fixes x y :: "real ^ 'n::finite" + shows "x \ y <= norm x * norm y" +proof- + {assume "norm x = 0" + hence ?thesis by (simp add: dot_lzero dot_rzero)} + moreover + {assume "norm y = 0" + hence ?thesis by (simp add: dot_lzero dot_rzero)} + moreover + {assume h: "norm x \ 0" "norm y \ 0" + let ?z = "norm y *s x - norm x *s y" + from h have p: "norm x * norm y > 0" by (metis norm_ge_zero le_less zero_compare_simps) + from dot_pos_le[of ?z] + have "(norm x * norm y) * (x \ y) \ norm x ^2 * norm y ^2" + apply (simp add: dot_rsub dot_lsub dot_lmult dot_rmult ring_simps) + by (simp add: norm_pow_2[symmetric] power2_eq_square dot_sym) + hence "x\y \ (norm x ^2 * norm y ^2) / (norm x * norm y)" using p + by (simp add: field_simps) + hence ?thesis using h by (simp add: power2_eq_square)} + ultimately show ?thesis by metis +qed + +lemma norm_cauchy_schwarz_abs: + fixes x y :: "real ^ 'n::finite" + shows "\x \ y\ \ norm x * norm y" + using norm_cauchy_schwarz[of x y] norm_cauchy_schwarz[of x "-y"] + by (simp add: real_abs_def dot_rneg) + +lemma norm_triangle_sub: + fixes x y :: "'a::real_normed_vector" + shows "norm x \ norm y + norm (x - y)" + using norm_triangle_ineq[of "y" "x - y"] by (simp add: ring_simps) + +lemma norm_triangle_le: "norm(x::real ^'n::finite) + norm y <= e ==> norm(x + y) <= e" + by (metis order_trans norm_triangle_ineq) +lemma norm_triangle_lt: "norm(x::real ^'n::finite) + norm(y) < e ==> norm(x + y) < e" + by (metis basic_trans_rules(21) norm_triangle_ineq) + +lemma component_le_norm: "\x$i\ <= norm (x::real ^ 'n::finite)" + apply (simp add: norm_vector_def) + apply (rule member_le_setL2, simp_all) + done + +lemma norm_bound_component_le: "norm(x::real ^ 'n::finite) <= e + ==> \x$i\ <= e" + by (metis component_le_norm order_trans) + +lemma norm_bound_component_lt: "norm(x::real ^ 'n::finite) < e + ==> \x$i\ < e" + 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: norm_vector_def setL2_le_setsum) + +lemma real_abs_norm: "\norm x\ = norm (x :: real ^ _)" + by (rule abs_norm_cancel) +lemma real_abs_sub_norm: "\norm(x::real ^'n::finite) - norm y\ <= norm(x - y)" + by (rule norm_triangle_ineq3) +lemma norm_le: "norm(x::real ^ _) <= norm(y) \ x \ x <= y \ y" + by (simp add: real_vector_norm_def) +lemma norm_lt: "norm(x::real ^ _) < norm(y) \ x \ x < y \ y" + by (simp add: real_vector_norm_def) +lemma norm_eq: "norm (x::real ^ _) = norm y \ x \ x = y \ y" + by (simp add: order_eq_iff norm_le) +lemma norm_eq_1: "norm(x::real ^ _) = 1 \ x \ x = 1" + by (simp add: real_vector_norm_def) + +text{* Squaring equations and inequalities involving norms. *} + +lemma dot_square_norm: "x \ x = norm(x)^2" + by (simp add: real_vector_norm_def) + +lemma norm_eq_square: "norm(x) = a \ 0 <= a \ x \ x = a^2" + by (auto simp add: real_vector_norm_def) + +lemma real_abs_le_square_iff: "\x\ \ \y\ \ (x::real)^2 \ y^2" +proof- + have "x^2 \ y^2 \ (x -y) * (y + x) \ 0" by (simp add: ring_simps power2_eq_square) + also have "\ \ \x\ \ \y\" apply (simp add: zero_compare_simps real_abs_def not_less) by arith +finally show ?thesis .. +qed + +lemma norm_le_square: "norm(x) <= a \ 0 <= a \ x \ x <= a^2" + apply (simp add: dot_square_norm real_abs_le_square_iff[symmetric]) + using norm_ge_zero[of x] + apply arith + done + +lemma norm_ge_square: "norm(x) >= a \ a <= 0 \ x \ x >= a ^ 2" + apply (simp add: dot_square_norm real_abs_le_square_iff[symmetric]) + using norm_ge_zero[of x] + apply arith + done + +lemma norm_lt_square: "norm(x) < a \ 0 < a \ x \ x < a^2" + by (metis not_le norm_ge_square) +lemma norm_gt_square: "norm(x) > a \ a < 0 \ x \ x > a^2" + by (metis norm_le_square not_less) + +text{* Dot product in terms of the norm rather than conversely. *} + +lemma dot_norm: "x \ y = (norm(x + y) ^2 - norm x ^ 2 - norm y ^ 2) / 2" + by (simp add: norm_pow_2 dot_ladd dot_radd dot_sym) + +lemma dot_norm_neg: "x \ y = ((norm x ^ 2 + norm y ^ 2) - norm(x - y) ^ 2) / 2" + by (simp add: norm_pow_2 dot_ladd dot_radd dot_lsub dot_rsub dot_sym) + + +text{* Equality of vectors in terms of @{term "op \"} products. *} + +lemma vector_eq: "(x:: real ^ 'n::finite) = y \ x \ x = x \ y\ y \ y = x \ x" (is "?lhs \ ?rhs") +proof + assume "?lhs" then show ?rhs by simp +next + assume ?rhs + then have "x \ x - x \ y = 0 \ x \ y - y\ y = 0" by simp + hence "x \ (x - y) = 0 \ y \ (x - y) = 0" + by (simp add: dot_rsub dot_lsub dot_sym) + then have "(x - y) \ (x - y) = 0" by (simp add: ring_simps dot_lsub dot_rsub) + then show "x = y" by (simp add: dot_eq_0) +qed + + +subsection{* General linear decision procedure for normed spaces. *} + +lemma norm_cmul_rule_thm: + fixes x :: "'a::real_normed_vector" + shows "b >= norm(x) ==> \c\ * b >= norm(scaleR c x)" + unfolding norm_scaleR + apply (erule mult_mono1) + apply simp + done + + (* FIXME: Move all these theorems into the ML code using lemma antiquotation *) +lemma norm_add_rule_thm: + fixes x1 x2 :: "'a::real_normed_vector" + shows "norm x1 \ b1 \ norm x2 \ b2 \ norm (x1 + x2) \ b1 + b2" + by (rule order_trans [OF norm_triangle_ineq add_mono]) + +lemma ge_iff_diff_ge_0: "(a::'a::ordered_ring) \ b == a - b \ 0" + by (simp add: ring_simps) + +lemma pth_1: + fixes x :: "'a::real_normed_vector" + shows "x == scaleR 1 x" by simp + +lemma pth_2: + fixes x :: "'a::real_normed_vector" + shows "x - y == x + -y" by (atomize (full)) simp + +lemma pth_3: + fixes x :: "'a::real_normed_vector" + shows "- x == scaleR (-1) x" by simp + +lemma pth_4: + fixes x :: "'a::real_normed_vector" + shows "scaleR 0 x == 0" and "scaleR c 0 = (0::'a)" by simp_all + +lemma pth_5: + fixes x :: "'a::real_normed_vector" + shows "scaleR c (scaleR d x) == scaleR (c * d) x" by simp + +lemma pth_6: + fixes x :: "'a::real_normed_vector" + shows "scaleR c (x + y) == scaleR c x + scaleR c y" + by (simp add: scaleR_right_distrib) + +lemma pth_7: + fixes x :: "'a::real_normed_vector" + shows "0 + x == x" and "x + 0 == x" by simp_all + +lemma pth_8: + fixes x :: "'a::real_normed_vector" + shows "scaleR c x + scaleR d x == scaleR (c + d) x" + by (simp add: scaleR_left_distrib) + +lemma pth_9: + fixes x :: "'a::real_normed_vector" shows + "(scaleR c x + z) + scaleR d x == scaleR (c + d) x + z" + "scaleR c x + (scaleR d x + z) == scaleR (c + d) x + z" + "(scaleR c x + w) + (scaleR d x + z) == scaleR (c + d) x + (w + z)" + by (simp_all add: algebra_simps) + +lemma pth_a: + fixes x :: "'a::real_normed_vector" + shows "scaleR 0 x + y == y" by simp + +lemma pth_b: + fixes x :: "'a::real_normed_vector" shows + "scaleR c x + scaleR d y == scaleR c x + scaleR d y" + "(scaleR c x + z) + scaleR d y == scaleR c x + (z + scaleR d y)" + "scaleR c x + (scaleR d y + z) == scaleR c x + (scaleR d y + z)" + "(scaleR c x + w) + (scaleR d y + z) == scaleR c x + (w + (scaleR d y + z))" + by (simp_all add: algebra_simps) + +lemma pth_c: + fixes x :: "'a::real_normed_vector" shows + "scaleR c x + scaleR d y == scaleR d y + scaleR c x" + "(scaleR c x + z) + scaleR d y == scaleR d y + (scaleR c x + z)" + "scaleR c x + (scaleR d y + z) == scaleR d y + (scaleR c x + z)" + "(scaleR c x + w) + (scaleR d y + z) == scaleR d y + ((scaleR c x + w) + z)" + by (simp_all add: algebra_simps) + +lemma pth_d: + fixes x :: "'a::real_normed_vector" + shows "x + 0 == x" by simp + +lemma norm_imp_pos_and_ge: + fixes x :: "'a::real_normed_vector" + shows "norm x == n \ norm x \ 0 \ n \ norm x" + by atomize auto + +lemma real_eq_0_iff_le_ge_0: "(x::real) = 0 == x \ 0 \ -x \ 0" by arith + +lemma norm_pths: + fixes x :: "'a::real_normed_vector" shows + "x = y \ norm (x - y) \ 0" + "x \ y \ \ (norm (x - y) \ 0)" + using norm_ge_zero[of "x - y"] by auto + +lemma vector_dist_norm: + fixes x :: "'a::real_normed_vector" + shows "dist x y = norm (x - y)" + by (rule dist_norm) + +use "normarith.ML" + +method_setup norm = {* Scan.succeed (SIMPLE_METHOD' o NormArith.norm_arith_tac) +*} "Proves simple linear statements about vector norms" + + +text{* Hence more metric properties. *} + +lemma dist_triangle_alt: + fixes x y z :: "'a::metric_space" + shows "dist y z <= dist x y + dist x z" +using dist_triangle [of y z x] by (simp add: dist_commute) + +lemma dist_pos_lt: + fixes x y :: "'a::metric_space" + shows "x \ y ==> 0 < dist x y" +by (simp add: zero_less_dist_iff) + +lemma dist_nz: + fixes x y :: "'a::metric_space" + shows "x \ y \ 0 < dist x y" +by (simp add: zero_less_dist_iff) + +lemma dist_triangle_le: + fixes x y z :: "'a::metric_space" + shows "dist x z + dist y z <= e \ dist x y <= e" +by (rule order_trans [OF dist_triangle2]) + +lemma dist_triangle_lt: + fixes x y z :: "'a::metric_space" + shows "dist x z + dist y z < e ==> dist x y < e" +by (rule le_less_trans [OF dist_triangle2]) + +lemma dist_triangle_half_l: + fixes x1 x2 y :: "'a::metric_space" + shows "dist x1 y < e / 2 \ dist x2 y < e / 2 \ dist x1 x2 < e" +by (rule dist_triangle_lt [where z=y], simp) + +lemma dist_triangle_half_r: + fixes x1 x2 y :: "'a::metric_space" + shows "dist y x1 < e / 2 \ dist y x2 < e / 2 \ dist x1 x2 < e" +by (rule dist_triangle_half_l, simp_all add: dist_commute) + +lemma dist_triangle_add: + fixes x y x' y' :: "'a::real_normed_vector" + shows "dist (x + y) (x' + y') <= dist x x' + dist y y'" + by norm + +lemma dist_mul[simp]: "dist (c *s x) (c *s y) = \c\ * dist x y" + unfolding dist_norm vector_ssub_ldistrib[symmetric] norm_mul .. + +lemma dist_triangle_add_half: + fixes x x' y y' :: "'a::real_normed_vector" + shows "dist x x' < e / 2 \ dist y y' < e / 2 \ dist(x + y) (x' + y') < e" + by norm + +lemma setsum_component [simp]: + fixes f:: " 'a \ ('b::comm_monoid_add) ^'n" + shows "(setsum f S)$i = setsum (\x. (f x)$i) S" + by (cases "finite S", induct S set: finite, simp_all) + +lemma setsum_eq: "setsum f S = (\ i. setsum (\x. (f x)$i ) S)" + by (simp add: Cart_eq) + +lemma setsum_clauses: + shows "setsum f {} = 0" + and "finite S \ setsum f (insert x S) = + (if x \ S then setsum f S else f x + setsum f S)" + by (auto simp add: insert_absorb) + +lemma setsum_cmul: + fixes f:: "'c \ ('a::semiring_1)^'n" + shows "setsum (\x. c *s f x) S = c *s setsum f S" + by (simp add: Cart_eq setsum_right_distrib) + +lemma setsum_norm: + fixes f :: "'a \ 'b::real_normed_vector" + assumes fS: "finite S" + shows "norm (setsum f S) <= setsum (\x. norm(f x)) S" +proof(induct rule: finite_induct[OF fS]) + case 1 thus ?case by simp +next + case (2 x S) + from "2.hyps" have "norm (setsum f (insert x S)) \ norm (f x) + norm (setsum f S)" by (simp add: norm_triangle_ineq) + also have "\ \ norm (f x) + setsum (\x. norm(f x)) S" + using "2.hyps" by simp + finally show ?case using "2.hyps" by simp +qed + +lemma real_setsum_norm: + fixes f :: "'a \ real ^'n::finite" + assumes fS: "finite S" + shows "norm (setsum f S) <= setsum (\x. norm(f x)) S" +proof(induct rule: finite_induct[OF fS]) + case 1 thus ?case by simp +next + case (2 x S) + from "2.hyps" have "norm (setsum f (insert x S)) \ norm (f x) + norm (setsum f S)" by (simp add: norm_triangle_ineq) + also have "\ \ norm (f x) + setsum (\x. norm(f x)) S" + using "2.hyps" by simp + finally show ?case using "2.hyps" by simp +qed + +lemma setsum_norm_le: + fixes f :: "'a \ 'b::real_normed_vector" + assumes fS: "finite S" + and fg: "\x \ S. norm (f x) \ g x" + shows "norm (setsum f S) \ setsum g S" +proof- + from fg have "setsum (\x. norm(f x)) S <= setsum g S" + by - (rule setsum_mono, simp) + then show ?thesis using setsum_norm[OF fS, of f] fg + by arith +qed + +lemma real_setsum_norm_le: + fixes f :: "'a \ real ^ 'n::finite" + assumes fS: "finite S" + and fg: "\x \ S. norm (f x) \ g x" + shows "norm (setsum f S) \ setsum g S" +proof- + from fg have "setsum (\x. norm(f x)) S <= setsum g S" + by - (rule setsum_mono, simp) + then show ?thesis using real_setsum_norm[OF fS, of f] fg + by arith +qed + +lemma setsum_norm_bound: + fixes f :: "'a \ 'b::real_normed_vector" + assumes fS: "finite S" + and K: "\x \ S. norm (f x) \ K" + shows "norm (setsum f S) \ of_nat (card S) * K" + using setsum_norm_le[OF fS K] setsum_constant[symmetric] + by simp + +lemma real_setsum_norm_bound: + fixes f :: "'a \ real ^ 'n::finite" + assumes fS: "finite S" + and K: "\x \ S. norm (f x) \ K" + shows "norm (setsum f S) \ of_nat (card S) * K" + using real_setsum_norm_le[OF fS K] setsum_constant[symmetric] + by simp + +lemma setsum_vmul: + fixes f :: "'a \ 'b::{real_normed_vector,semiring, mult_zero}" + assumes fS: "finite S" + shows "setsum f S *s v = setsum (\x. f x *s v) S" +proof(induct rule: finite_induct[OF fS]) + case 1 then show ?case by (simp add: vector_smult_lzero) +next + case (2 x F) + from "2.hyps" have "setsum f (insert x F) *s v = (f x + setsum f F) *s v" + by simp + also have "\ = f x *s v + setsum f F *s v" + by (simp add: vector_sadd_rdistrib) + also have "\ = setsum (\x. f x *s v) (insert x F)" using "2.hyps" by simp + finally show ?case . +qed + +(* FIXME : Problem thm setsum_vmul[of _ "f:: 'a \ real ^'n"] --- + Get rid of *s and use real_vector instead! Also prove that ^ creates a real_vector !! *) + + (* FIXME: Here too need stupid finiteness assumption on T!!! *) +lemma setsum_group: + assumes fS: "finite S" and fT: "finite T" and fST: "f ` S \ T" + shows "setsum (\y. setsum g {x. x\ S \ f x = y}) T = setsum g S" + +apply (subst setsum_image_gen[OF fS, of g f]) +apply (rule setsum_mono_zero_right[OF fT fST]) +by (auto intro: setsum_0') + +lemma vsum_norm_allsubsets_bound: + fixes f:: "'a \ real ^'n::finite" + assumes fP: "finite P" and fPs: "\Q. Q \ P \ norm (setsum f Q) \ e" + shows "setsum (\x. norm (f x)) P \ 2 * real CARD('n) * e" +proof- + let ?d = "real CARD('n)" + let ?nf = "\x. norm (f x)" + let ?U = "UNIV :: 'n set" + have th0: "setsum (\x. setsum (\i. \f x $ i\) ?U) P = setsum (\i. setsum (\x. \f x $ i\) P) ?U" + by (rule setsum_commute) + have th1: "2 * ?d * e = of_nat (card ?U) * (2 * e)" by (simp add: real_of_nat_def) + have "setsum ?nf P \ setsum (\x. setsum (\i. \f x $ i\) ?U) P" + apply (rule setsum_mono) + by (rule norm_le_l1) + also have "\ \ 2 * ?d * e" + unfolding th0 th1 + proof(rule setsum_bounded) + fix i assume i: "i \ ?U" + let ?Pp = "{x. x\ P \ f x $ i \ 0}" + let ?Pn = "{x. x \ P \ f x $ i < 0}" + have thp: "P = ?Pp \ ?Pn" by auto + have thp0: "?Pp \ ?Pn ={}" by auto + have PpP: "?Pp \ P" and PnP: "?Pn \ P" by blast+ + have Ppe:"setsum (\x. \f x $ i\) ?Pp \ e" + using component_le_norm[of "setsum (\x. f x) ?Pp" i] fPs[OF PpP] + by (auto intro: abs_le_D1) + have Pne: "setsum (\x. \f x $ i\) ?Pn \ e" + using component_le_norm[of "setsum (\x. - f x) ?Pn" i] fPs[OF PnP] + by (auto simp add: setsum_negf intro: abs_le_D1) + have "setsum (\x. \f x $ i\) P = setsum (\x. \f x $ i\) ?Pp + setsum (\x. \f x $ i\) ?Pn" + apply (subst thp) + apply (rule setsum_Un_zero) + using fP thp0 by auto + also have "\ \ 2*e" using Pne Ppe by arith + finally show "setsum (\x. \f x $ i\) P \ 2*e" . + qed + finally show ?thesis . +qed + +lemma dot_lsum: "finite S \ setsum f S \ (y::'a::{comm_ring}^'n) = setsum (\x. f x \ y) S " + by (induct rule: finite_induct, auto simp add: dot_lzero dot_ladd dot_radd) + +lemma dot_rsum: "finite S \ (y::'a::{comm_ring}^'n) \ setsum f S = setsum (\x. y \ f x) S " + by (induct rule: finite_induct, auto simp add: dot_rzero dot_radd) + +subsection{* Basis vectors in coordinate directions. *} + + +definition "basis k = (\ i. if i = k then 1 else 0)" + +lemma basis_component [simp]: "basis k $ i = (if k=i then 1 else 0)" + unfolding basis_def by simp + +lemma delta_mult_idempotent: + "(if k=a then 1 else (0::'a::semiring_1)) * (if k=a then 1 else 0) = (if k=a then 1 else 0)" by (cases "k=a", auto) + +lemma norm_basis: + shows "norm (basis k :: real ^'n::finite) = 1" + apply (simp add: basis_def real_vector_norm_def dot_def) + apply (vector delta_mult_idempotent) + using setsum_delta[of "UNIV :: 'n set" "k" "\k. 1::real"] + apply auto + done + +lemma norm_basis_1: "norm(basis 1 :: real ^'n::{finite,one}) = 1" + by (rule norm_basis) + +lemma vector_choose_size: "0 <= c ==> \(x::real^'n::finite). norm x = c" + apply (rule exI[where x="c *s basis arbitrary"]) + by (simp only: norm_mul norm_basis) + +lemma vector_choose_dist: assumes e: "0 <= e" + shows "\(y::real^'n::finite). dist x y = e" +proof- + from vector_choose_size[OF e] obtain c:: "real ^'n" where "norm c = e" + by blast + then have "dist x (x - c) = e" by (simp add: dist_norm) + then show ?thesis by blast +qed + +lemma basis_inj: "inj (basis :: 'n \ real ^'n::finite)" + by (simp add: inj_on_def Cart_eq) + +lemma cond_value_iff: "f (if b then x else y) = (if b then f x else f y)" + by auto + +lemma basis_expansion: + "setsum (\i. (x$i) *s basis i) UNIV = (x::('a::ring_1) ^'n::finite)" (is "?lhs = ?rhs" is "setsum ?f ?S = _") + by (auto simp add: Cart_eq cond_value_iff setsum_delta[of "?S", where ?'b = "'a", simplified] cong del: if_weak_cong) + +lemma basis_expansion_unique: + "setsum (\i. f i *s basis i) UNIV = (x::('a::comm_ring_1) ^'n::finite) \ (\i. f i = x$i)" + by (simp add: Cart_eq setsum_delta cond_value_iff cong del: if_weak_cong) + +lemma cond_application_beta: "(if b then f else g) x = (if b then f x else g x)" + by auto + +lemma dot_basis: + 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) + +lemma basis_nonzero: + shows "basis k \ (0:: 'a::semiring_1 ^'n)" + by (simp add: basis_eq_0) + +lemma vector_eq_ldot: "(\x. x \ y = x \ z) \ y = (z::'a::semiring_1^'n::finite)" + apply (auto simp add: Cart_eq dot_basis) + apply (erule_tac x="basis i" in allE) + apply (simp add: dot_basis) + apply (subgoal_tac "y = z") + apply simp + apply (simp add: Cart_eq) + done + +lemma vector_eq_rdot: "(\z. x \ z = y \ z) \ x = (y::'a::semiring_1^'n::finite)" + apply (auto simp add: Cart_eq dot_basis) + apply (erule_tac x="basis i" in allE) + apply (simp add: dot_basis) + apply (subgoal_tac "x = y") + apply simp + apply (simp add: Cart_eq) + done + +subsection{* Orthogonality. *} + +definition "orthogonal x y \ (x \ y = 0)" + +lemma orthogonal_basis: + shows "orthogonal (basis i :: 'a^'n::finite) x \ x$i = (0::'a::ring_1)" + by (auto simp add: orthogonal_def dot_def basis_def cond_value_iff cond_application_beta setsum_delta cong del: if_weak_cong) + +lemma orthogonal_basis_basis: + shows "orthogonal (basis i :: 'a::ring_1^'n::finite) (basis j) \ i \ j" + unfolding orthogonal_basis[of i] basis_component[of j] by simp + + (* FIXME : Maybe some of these require less than comm_ring, but not all*) +lemma orthogonal_clauses: + "orthogonal a (0::'a::comm_ring ^'n)" + "orthogonal a x ==> orthogonal a (c *s x)" + "orthogonal a x ==> orthogonal a (-x)" + "orthogonal a x \ orthogonal a y ==> orthogonal a (x + y)" + "orthogonal a x \ orthogonal a y ==> orthogonal a (x - y)" + "orthogonal 0 a" + "orthogonal x a ==> orthogonal (c *s x) a" + "orthogonal x a ==> orthogonal (-x) a" + "orthogonal x a \ orthogonal y a ==> orthogonal (x + y) a" + "orthogonal x a \ orthogonal y a ==> orthogonal (x - y) a" + unfolding orthogonal_def dot_rneg dot_rmult dot_radd dot_rsub + dot_lzero dot_rzero dot_lneg dot_lmult dot_ladd dot_lsub + by simp_all + +lemma orthogonal_commute: "orthogonal (x::'a::{ab_semigroup_mult,comm_monoid_add} ^'n)y \ orthogonal y x" + by (simp add: orthogonal_def dot_sym) + +subsection{* Explicit vector construction from lists. *} + +primrec from_nat :: "nat \ 'a::{monoid_add,one}" +where "from_nat 0 = 0" | "from_nat (Suc n) = 1 + from_nat n" + +lemma from_nat [simp]: "from_nat = of_nat" +by (rule ext, induct_tac x, simp_all) + +primrec + list_fun :: "nat \ _ list \ _ \ _" +where + "list_fun n [] = (\x. 0)" +| "list_fun n (x # xs) = fun_upd (list_fun (Suc n) xs) (from_nat n) x" + +definition "vector l = (\ i. list_fun 1 l i)" +(*definition "vector l = (\ i. if i <= length l then l ! (i - 1) else 0)"*) + +lemma vector_1: "(vector[x]) $1 = x" + unfolding vector_def by simp + +lemma vector_2: + "(vector[x,y]) $1 = x" + "(vector[x,y] :: 'a^2)$2 = (y::'a::zero)" + unfolding vector_def by simp_all + +lemma vector_3: + "(vector [x,y,z] ::('a::zero)^3)$1 = x" + "(vector [x,y,z] ::('a::zero)^3)$2 = y" + "(vector [x,y,z] ::('a::zero)^3)$3 = z" + unfolding vector_def by simp_all + +lemma forall_vector_1: "(\v::'a::zero^1. P v) \ (\x. P(vector[x]))" + apply auto + apply (erule_tac x="v$1" in allE) + apply (subgoal_tac "vector [v$1] = v") + apply simp + apply (vector vector_def) + apply (simp add: forall_1) + done + +lemma forall_vector_2: "(\v::'a::zero^2. P v) \ (\x y. P(vector[x, y]))" + apply auto + apply (erule_tac x="v$1" in allE) + apply (erule_tac x="v$2" in allE) + apply (subgoal_tac "vector [v$1, v$2] = v") + apply simp + apply (vector vector_def) + apply (simp add: forall_2) + done + +lemma forall_vector_3: "(\v::'a::zero^3. P v) \ (\x y z. P(vector[x, y, z]))" + apply auto + apply (erule_tac x="v$1" in allE) + apply (erule_tac x="v$2" in allE) + apply (erule_tac x="v$3" in allE) + apply (subgoal_tac "vector [v$1, v$2, v$3] = v") + apply simp + apply (vector vector_def) + apply (simp add: forall_3) + done + +subsection{* Linear functions. *} + +definition "linear f \ (\x y. f(x + y) = f x + f y) \ (\c x. f(c *s x) = c *s f x)" + +lemma linear_compose_cmul: "linear f ==> linear (\x. (c::'a::comm_semiring) *s f x)" + by (vector linear_def Cart_eq ring_simps) + +lemma linear_compose_neg: "linear (f :: 'a ^'n \ 'a::comm_ring ^'m) ==> linear (\x. -(f(x)))" by (vector linear_def Cart_eq) + +lemma linear_compose_add: "linear (f :: 'a ^'n \ 'a::semiring_1 ^'m) \ linear g ==> linear (\x. f(x) + g(x))" + by (vector linear_def Cart_eq ring_simps) + +lemma linear_compose_sub: "linear (f :: 'a ^'n \ 'a::ring_1 ^'m) \ linear g ==> linear (\x. f x - g x)" + by (vector linear_def Cart_eq ring_simps) + +lemma linear_compose: "linear f \ linear g ==> linear (g o f)" + by (simp add: linear_def) + +lemma linear_id: "linear id" by (simp add: linear_def id_def) + +lemma linear_zero: "linear (\x. 0::'a::semiring_1 ^ 'n)" by (simp add: linear_def) + +lemma linear_compose_setsum: + assumes fS: "finite S" and lS: "\a \ S. linear (f a :: 'a::semiring_1 ^ 'n \ 'a ^ 'm)" + shows "linear(\x. setsum (\a. f a x :: 'a::semiring_1 ^'m) S)" + using lS + apply (induct rule: finite_induct[OF fS]) + by (auto simp add: linear_zero intro: linear_compose_add) + +lemma linear_vmul_component: + fixes f:: "'a::semiring_1^'m \ 'a^'n" + assumes lf: "linear f" + shows "linear (\x. f x $ k *s v)" + using lf + apply (auto simp add: linear_def ) + by (vector ring_simps)+ + +lemma linear_0: "linear f ==> f 0 = (0::'a::semiring_1 ^'n)" + unfolding linear_def + apply clarsimp + apply (erule allE[where x="0::'a"]) + apply simp + done + +lemma linear_cmul: "linear f ==> f(c*s x) = c *s f x" by (simp add: linear_def) + +lemma linear_neg: "linear (f :: 'a::ring_1 ^'n \ _) ==> f (-x) = - f x" + unfolding vector_sneg_minus1 + using linear_cmul[of f] by auto + +lemma linear_add: "linear f ==> f(x + y) = f x + f y" by (metis linear_def) + +lemma linear_sub: "linear (f::'a::ring_1 ^'n \ _) ==> f(x - y) = f x - f y" + by (simp add: diff_def linear_add linear_neg) + +lemma linear_setsum: + fixes f:: "'a::semiring_1^'n \ _" + assumes lf: "linear f" and fS: "finite S" + shows "f (setsum g S) = setsum (f o g) S" +proof (induct rule: finite_induct[OF fS]) + case 1 thus ?case by (simp add: linear_0[OF lf]) +next + case (2 x F) + have "f (setsum g (insert x F)) = f (g x + setsum g F)" using "2.hyps" + by simp + also have "\ = f (g x) + f (setsum g F)" using linear_add[OF lf] by simp + also have "\ = setsum (f o g) (insert x F)" using "2.hyps" by simp + finally show ?case . +qed + +lemma linear_setsum_mul: + fixes f:: "'a ^'n \ 'a::semiring_1^'m" + assumes lf: "linear f" and fS: "finite S" + shows "f (setsum (\i. c i *s v i) S) = setsum (\i. c i *s f (v i)) S" + using linear_setsum[OF lf fS, of "\i. c i *s v i" , unfolded o_def] + linear_cmul[OF lf] by simp + +lemma linear_injective_0: + assumes lf: "linear (f:: 'a::ring_1 ^ 'n \ _)" + shows "inj f \ (\x. f x = 0 \ x = 0)" +proof- + have "inj f \ (\ x y. f x = f y \ x = y)" by (simp add: inj_on_def) + also have "\ \ (\ x y. f x - f y = 0 \ x - y = 0)" by simp + also have "\ \ (\ x y. f (x - y) = 0 \ x - y = 0)" + by (simp add: linear_sub[OF lf]) + also have "\ \ (\ x. f x = 0 \ x = 0)" by auto + finally show ?thesis . +qed + +lemma linear_bounded: + fixes f:: "real ^'m::finite \ real ^'n::finite" + assumes lf: "linear f" + shows "\B. \x. norm (f x) \ B * norm x" +proof- + let ?S = "UNIV:: 'm set" + let ?B = "setsum (\i. norm(f(basis i))) ?S" + have fS: "finite ?S" by simp + {fix x:: "real ^ 'm" + let ?g = "(\i. (x$i) *s (basis i) :: real ^ 'm)" + have "norm (f x) = norm (f (setsum (\i. (x$i) *s (basis i)) ?S))" + by (simp only: basis_expansion) + also have "\ = norm (setsum (\i. (x$i) *s f (basis i))?S)" + using linear_setsum[OF lf fS, of ?g, unfolded o_def] linear_cmul[OF lf] + by auto + finally have th0: "norm (f x) = norm (setsum (\i. (x$i) *s f (basis i))?S)" . + {fix i assume i: "i \ ?S" + from component_le_norm[of x i] + have "norm ((x$i) *s f (basis i :: real ^'m)) \ norm (f (basis i)) * norm x" + unfolding norm_mul + apply (simp only: mult_commute) + apply (rule mult_mono) + by (auto simp add: ring_simps norm_ge_zero) } + then have th: "\i\ ?S. norm ((x$i) *s f (basis i :: real ^'m)) \ norm (f (basis i)) * norm x" by metis + from real_setsum_norm_le[OF fS, of "\i. (x$i) *s (f (basis i))", OF th] + have "norm (f x) \ ?B * norm x" unfolding th0 setsum_left_distrib by metis} + then show ?thesis by blast +qed + +lemma linear_bounded_pos: + fixes f:: "real ^'n::finite \ real ^ 'm::finite" + assumes lf: "linear f" + shows "\B > 0. \x. norm (f x) \ B * norm x" +proof- + from linear_bounded[OF lf] obtain B where + B: "\x. norm (f x) \ B * norm x" by blast + let ?K = "\B\ + 1" + have Kp: "?K > 0" by arith + {assume C: "B < 0" + have "norm (1::real ^ 'n) > 0" by (simp add: zero_less_norm_iff) + with C have "B * norm (1:: real ^ 'n) < 0" + by (simp add: zero_compare_simps) + with B[rule_format, of 1] norm_ge_zero[of "f 1"] have False by simp + } + then have Bp: "B \ 0" by ferrack + {fix x::"real ^ 'n" + have "norm (f x) \ ?K * norm x" + using B[rule_format, of x] norm_ge_zero[of x] norm_ge_zero[of "f x"] Bp + apply (auto simp add: ring_simps split add: abs_split) + apply (erule order_trans, simp) + done + } + then show ?thesis using Kp by blast +qed + +lemma smult_conv_scaleR: "c *s x = scaleR c x" + unfolding vector_scalar_mult_def vector_scaleR_def by simp + +lemma linear_conv_bounded_linear: + fixes f :: "real ^ _ \ real ^ _" + shows "linear f \ bounded_linear f" +proof + assume "linear f" + show "bounded_linear f" + proof + fix x y show "f (x + y) = f x + f y" + using `linear f` unfolding linear_def by simp + next + fix r x show "f (scaleR r x) = scaleR r (f x)" + using `linear f` unfolding linear_def + by (simp add: smult_conv_scaleR) + next + have "\B. \x. norm (f x) \ B * norm x" + using `linear f` by (rule linear_bounded) + thus "\K. \x. norm (f x) \ norm x * K" + by (simp add: mult_commute) + qed +next + assume "bounded_linear f" + then interpret f: bounded_linear f . + show "linear f" + unfolding linear_def smult_conv_scaleR + by (simp add: f.add f.scaleR) +qed + +subsection{* Bilinear functions. *} + +definition "bilinear f \ (\x. linear(\y. f x y)) \ (\y. linear(\x. f x y))" + +lemma bilinear_ladd: "bilinear h ==> h (x + y) z = (h x z) + (h y z)" + by (simp add: bilinear_def linear_def) +lemma bilinear_radd: "bilinear h ==> h x (y + z) = (h x y) + (h x z)" + by (simp add: bilinear_def linear_def) + +lemma bilinear_lmul: "bilinear h ==> h (c *s x) y = c *s (h x y)" + by (simp add: bilinear_def linear_def) + +lemma bilinear_rmul: "bilinear h ==> h x (c *s y) = c *s (h x y)" + by (simp add: bilinear_def linear_def) + +lemma bilinear_lneg: "bilinear h ==> h (- (x:: 'a::ring_1 ^ 'n)) y = -(h x y)" + by (simp only: vector_sneg_minus1 bilinear_lmul) + +lemma bilinear_rneg: "bilinear h ==> h x (- (y:: 'a::ring_1 ^ 'n)) = - h x y" + by (simp only: vector_sneg_minus1 bilinear_rmul) + +lemma (in ab_group_add) eq_add_iff: "x = x + y \ y = 0" + using add_imp_eq[of x y 0] by auto + +lemma bilinear_lzero: + fixes h :: "'a::ring^'n \ _" assumes bh: "bilinear h" shows "h 0 x = 0" + using bilinear_ladd[OF bh, of 0 0 x] + by (simp add: eq_add_iff ring_simps) + +lemma bilinear_rzero: + fixes h :: "'a::ring^'n \ _" assumes bh: "bilinear h" shows "h x 0 = 0" + using bilinear_radd[OF bh, of x 0 0 ] + by (simp add: eq_add_iff ring_simps) + +lemma bilinear_lsub: "bilinear h ==> h (x - (y:: 'a::ring_1 ^ 'n)) z = h x z - h y z" + by (simp add: diff_def bilinear_ladd bilinear_lneg) + +lemma bilinear_rsub: "bilinear h ==> h z (x - (y:: 'a::ring_1 ^ 'n)) = h z x - h z y" + by (simp add: diff_def bilinear_radd bilinear_rneg) + +lemma bilinear_setsum: + fixes h:: "'a ^'n \ 'a::semiring_1^'m \ 'a ^ 'k" + assumes bh: "bilinear h" and fS: "finite S" and fT: "finite T" + shows "h (setsum f S) (setsum g T) = setsum (\(i,j). h (f i) (g j)) (S \ T) " +proof- + have "h (setsum f S) (setsum g T) = setsum (\x. h (f x) (setsum g T)) S" + apply (rule linear_setsum[unfolded o_def]) + using bh fS by (auto simp add: bilinear_def) + also have "\ = setsum (\x. setsum (\y. h (f x) (g y)) T) S" + apply (rule setsum_cong, simp) + apply (rule linear_setsum[unfolded o_def]) + using bh fT by (auto simp add: bilinear_def) + finally show ?thesis unfolding setsum_cartesian_product . +qed + +lemma bilinear_bounded: + fixes h:: "real ^'m::finite \ real^'n::finite \ real ^ 'k::finite" + assumes bh: "bilinear h" + shows "\B. \x y. norm (h x y) \ B * norm x * norm y" +proof- + let ?M = "UNIV :: 'm set" + let ?N = "UNIV :: 'n set" + let ?B = "setsum (\(i,j). norm (h (basis i) (basis j))) (?M \ ?N)" + have fM: "finite ?M" and fN: "finite ?N" by simp_all + {fix x:: "real ^ 'm" and y :: "real^'n" + have "norm (h x y) = norm (h (setsum (\i. (x$i) *s basis i) ?M) (setsum (\i. (y$i) *s basis i) ?N))" unfolding basis_expansion .. + also have "\ = norm (setsum (\ (i,j). h ((x$i) *s basis i) ((y$j) *s basis j)) (?M \ ?N))" unfolding bilinear_setsum[OF bh fM fN] .. + finally have th: "norm (h x y) = \" . + have "norm (h x y) \ ?B * norm x * norm y" + apply (simp add: setsum_left_distrib th) + apply (rule real_setsum_norm_le) + using fN fM + apply simp + apply (auto simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh] norm_mul ring_simps) + apply (rule mult_mono) + apply (auto simp add: norm_ge_zero zero_le_mult_iff component_le_norm) + apply (rule mult_mono) + apply (auto simp add: norm_ge_zero zero_le_mult_iff component_le_norm) + done} + then show ?thesis by metis +qed + +lemma bilinear_bounded_pos: + fixes h:: "real ^'m::finite \ real^'n::finite \ real ^ 'k::finite" + assumes bh: "bilinear h" + shows "\B > 0. \x y. norm (h x y) \ B * norm x * norm y" +proof- + from bilinear_bounded[OF bh] obtain B where + B: "\x y. norm (h x y) \ B * norm x * norm y" by blast + let ?K = "\B\ + 1" + have Kp: "?K > 0" by arith + have KB: "B < ?K" by arith + {fix x::"real ^'m" and y :: "real ^'n" + from KB Kp + have "B * norm x * norm y \ ?K * norm x * norm y" + apply - + apply (rule mult_right_mono, rule mult_right_mono) + by (auto simp add: norm_ge_zero) + then have "norm (h x y) \ ?K * norm x * norm y" + using B[rule_format, of x y] by simp} + with Kp show ?thesis by blast +qed + +lemma bilinear_conv_bounded_bilinear: + fixes h :: "real ^ _ \ real ^ _ \ real ^ _" + shows "bilinear h \ bounded_bilinear h" +proof + assume "bilinear h" + show "bounded_bilinear h" + proof + fix x y z show "h (x + y) z = h x z + h y z" + using `bilinear h` unfolding bilinear_def linear_def by simp + next + fix x y z show "h x (y + z) = h x y + h x z" + using `bilinear h` unfolding bilinear_def linear_def by simp + next + fix r x y show "h (scaleR r x) y = scaleR r (h x y)" + using `bilinear h` unfolding bilinear_def linear_def + by (simp add: smult_conv_scaleR) + next + fix r x y show "h x (scaleR r y) = scaleR r (h x y)" + using `bilinear h` unfolding bilinear_def linear_def + by (simp add: smult_conv_scaleR) + next + have "\B. \x y. norm (h x y) \ B * norm x * norm y" + using `bilinear h` by (rule bilinear_bounded) + thus "\K. \x y. norm (h x y) \ norm x * norm y * K" + by (simp add: mult_ac) + qed +next + assume "bounded_bilinear h" + then interpret h: bounded_bilinear h . + show "bilinear h" + unfolding bilinear_def linear_conv_bounded_linear + using h.bounded_linear_left h.bounded_linear_right + by simp +qed + +subsection{* Adjoints. *} + +definition "adjoint f = (SOME f'. \x y. f x \ y = x \ f' y)" + +lemma choice_iff: "(\x. \y. P x y) \ (\f. \x. P x (f x))" by metis + +lemma adjoint_works_lemma: + fixes f:: "'a::ring_1 ^'n::finite \ 'a ^ 'm::finite" + assumes lf: "linear f" + shows "\x y. f x \ y = x \ adjoint f y" +proof- + let ?N = "UNIV :: 'n set" + let ?M = "UNIV :: 'm set" + have fN: "finite ?N" by simp + have fM: "finite ?M" by simp + {fix y:: "'a ^ 'm" + let ?w = "(\ i. (f (basis i) \ y)) :: 'a ^ 'n" + {fix x + have "f x \ y = f (setsum (\i. (x$i) *s basis i) ?N) \ y" + by (simp only: basis_expansion) + also have "\ = (setsum (\i. (x$i) *s f (basis i)) ?N) \ y" + unfolding linear_setsum[OF lf fN] + by (simp add: linear_cmul[OF lf]) + finally have "f x \ y = x \ ?w" + apply (simp only: ) + apply (simp add: dot_def setsum_left_distrib setsum_right_distrib setsum_commute[of _ ?M ?N] ring_simps) + done} + } + then show ?thesis unfolding adjoint_def + some_eq_ex[of "\f'. \x y. f x \ y = x \ f' y"] + using choice_iff[of "\a b. \x. f x \ a = x \ b "] + by metis +qed + +lemma adjoint_works: + fixes f:: "'a::ring_1 ^'n::finite \ 'a ^ 'm::finite" + assumes lf: "linear f" + shows "x \ adjoint f y = f x \ y" + using adjoint_works_lemma[OF lf] by metis + + +lemma adjoint_linear: + fixes f :: "'a::comm_ring_1 ^'n::finite \ 'a ^ 'm::finite" + assumes lf: "linear f" + shows "linear (adjoint f)" + by (simp add: linear_def vector_eq_ldot[symmetric] dot_radd dot_rmult adjoint_works[OF lf]) + +lemma adjoint_clauses: + fixes f:: "'a::comm_ring_1 ^'n::finite \ 'a ^ 'm::finite" + assumes lf: "linear f" + shows "x \ adjoint f y = f x \ y" + and "adjoint f y \ x = y \ f x" + by (simp_all add: adjoint_works[OF lf] dot_sym ) + +lemma adjoint_adjoint: + fixes f:: "'a::comm_ring_1 ^ 'n::finite \ 'a ^ 'm::finite" + assumes lf: "linear f" + shows "adjoint (adjoint f) = f" + apply (rule ext) + by (simp add: vector_eq_ldot[symmetric] adjoint_clauses[OF adjoint_linear[OF lf]] adjoint_clauses[OF lf]) + +lemma adjoint_unique: + fixes f:: "'a::comm_ring_1 ^ 'n::finite \ 'a ^ 'm::finite" + assumes lf: "linear f" and u: "\x y. f' x \ y = x \ f y" + shows "f' = adjoint f" + apply (rule ext) + using u + by (simp add: vector_eq_rdot[symmetric] adjoint_clauses[OF lf]) + +text{* Matrix notation. NB: an MxN matrix is of type @{typ "'a^'n^'m"}, not @{typ "'a^'m^'n"} *} + +consts generic_mult :: "'a \ 'b \ 'c" (infixr "\" 75) + +defs (overloaded) +matrix_matrix_mult_def: "(m:: ('a::semiring_1) ^'n^'m) \ (m' :: 'a ^'p^'n) \ (\ i j. setsum (\k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m" + +abbreviation + matrix_matrix_mult' :: "('a::semiring_1) ^'n^'m \ 'a ^'p^'n \ 'a ^ 'p ^'m" (infixl "**" 70) + where "m ** m' == m\ m'" + +defs (overloaded) + matrix_vector_mult_def: "(m::('a::semiring_1) ^'n^'m) \ (x::'a ^'n) \ (\ i. setsum (\j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m" + +abbreviation + matrix_vector_mult' :: "('a::semiring_1) ^'n^'m \ 'a ^'n \ 'a ^ 'm" (infixl "*v" 70) + where + "m *v v == m \ v" + +defs (overloaded) + vector_matrix_mult_def: "(x::'a^'m) \ (m::('a::semiring_1) ^'n^'m) \ (\ j. setsum (\i. ((m$i)$j) * (x$i)) (UNIV :: 'm set)) :: 'a^'n" + +abbreviation + vactor_matrix_mult' :: "'a ^ 'm \ ('a::semiring_1) ^'n^'m \ 'a ^'n " (infixl "v*" 70) + where + "v v* m == v \ m" + +definition "(mat::'a::zero => 'a ^'n^'n) k = (\ i j. if i = j then k else 0)" +definition "(transp::'a^'n^'m \ 'a^'m^'n) A = (\ i j. ((A$j)$i))" +definition "(row::'m => 'a ^'n^'m \ 'a ^'n) i A = (\ j. ((A$i)$j))" +definition "(column::'n =>'a^'n^'m =>'a^'m) j A = (\ i. ((A$i)$j))" +definition "rows(A::'a^'n^'m) = { row i A | i. i \ (UNIV :: 'm set)}" +definition "columns(A::'a^'n^'m) = { column i A | i. i \ (UNIV :: 'n set)}" + +lemma mat_0[simp]: "mat 0 = 0" by (vector mat_def) +lemma matrix_add_ldistrib: "(A ** (B + C)) = (A \ B) + (A \ C)" + by (vector matrix_matrix_mult_def setsum_addf[symmetric] ring_simps) + +lemma matrix_mul_lid: + fixes A :: "'a::semiring_1 ^ 'm ^ 'n::finite" + shows "mat 1 ** A = A" + apply (simp add: matrix_matrix_mult_def mat_def) + apply vector + by (auto simp only: cond_value_iff cond_application_beta setsum_delta'[OF finite] mult_1_left mult_zero_left if_True UNIV_I) + + +lemma matrix_mul_rid: + fixes A :: "'a::semiring_1 ^ 'm::finite ^ 'n" + shows "A ** mat 1 = A" + apply (simp add: matrix_matrix_mult_def mat_def) + apply vector + by (auto simp only: cond_value_iff cond_application_beta setsum_delta[OF finite] mult_1_right mult_zero_right if_True UNIV_I cong: if_cong) + +lemma matrix_mul_assoc: "A ** (B ** C) = (A ** B) ** C" + apply (vector matrix_matrix_mult_def setsum_right_distrib setsum_left_distrib mult_assoc) + apply (subst setsum_commute) + apply simp + done + +lemma matrix_vector_mul_assoc: "A *v (B *v x) = (A ** B) *v x" + apply (vector matrix_matrix_mult_def matrix_vector_mult_def setsum_right_distrib setsum_left_distrib mult_assoc) + apply (subst setsum_commute) + apply simp + done + +lemma matrix_vector_mul_lid: "mat 1 *v x = (x::'a::semiring_1 ^ 'n::finite)" + apply (vector matrix_vector_mult_def mat_def) + by (simp add: cond_value_iff cond_application_beta + setsum_delta' cong del: if_weak_cong) + +lemma matrix_transp_mul: "transp(A ** B) = transp B ** transp (A::'a::comm_semiring_1^'m^'n)" + by (simp add: matrix_matrix_mult_def transp_def Cart_eq mult_commute) + +lemma matrix_eq: + fixes A B :: "'a::semiring_1 ^ 'n::finite ^ 'm" + shows "A = B \ (\x. A *v x = B *v x)" (is "?lhs \ ?rhs") + apply auto + apply (subst Cart_eq) + apply clarify + apply (clarsimp simp add: matrix_vector_mult_def basis_def cond_value_iff cond_application_beta Cart_eq cong del: if_weak_cong) + apply (erule_tac x="basis ia" in allE) + apply (erule_tac x="i" in allE) + by (auto simp add: basis_def cond_value_iff cond_application_beta setsum_delta[OF finite] cong del: if_weak_cong) + +lemma matrix_vector_mul_component: + shows "((A::'a::semiring_1^'n'^'m) *v x)$k = (A$k) \ x" + by (simp add: matrix_vector_mult_def dot_def) + +lemma dot_lmul_matrix: "((x::'a::comm_semiring_1 ^'n) v* A) \ y = x \ (A *v y)" + apply (simp add: dot_def matrix_vector_mult_def vector_matrix_mult_def setsum_left_distrib setsum_right_distrib mult_ac) + apply (subst setsum_commute) + by simp + +lemma transp_mat: "transp (mat n) = mat n" + by (vector transp_def mat_def) + +lemma transp_transp: "transp(transp A) = A" + by (vector transp_def) + +lemma row_transp: + fixes A:: "'a::semiring_1^'n^'m" + shows "row i (transp A) = column i A" + by (simp add: row_def column_def transp_def Cart_eq) + +lemma column_transp: + fixes A:: "'a::semiring_1^'n^'m" + shows "column i (transp A) = row i A" + by (simp add: row_def column_def transp_def Cart_eq) + +lemma rows_transp: "rows(transp (A::'a::semiring_1^'n^'m)) = columns A" +by (auto simp add: rows_def columns_def row_transp intro: set_ext) + +lemma columns_transp: "columns(transp (A::'a::semiring_1^'n^'m)) = rows A" by (metis transp_transp rows_transp) + +text{* Two sometimes fruitful ways of looking at matrix-vector multiplication. *} + +lemma matrix_mult_dot: "A *v x = (\ i. A$i \ x)" + by (simp add: matrix_vector_mult_def dot_def) + +lemma matrix_mult_vsum: "(A::'a::comm_semiring_1^'n^'m) *v x = setsum (\i. (x$i) *s column i A) (UNIV:: 'n set)" + by (simp add: matrix_vector_mult_def Cart_eq column_def mult_commute) + +lemma vector_componentwise: + "(x::'a::ring_1^'n::finite) = (\ j. setsum (\i. (x$i) * (basis i :: 'a^'n)$j) (UNIV :: 'n set))" + apply (subst basis_expansion[symmetric]) + by (vector Cart_eq setsum_component) + +lemma linear_componentwise: + fixes f:: "'a::ring_1 ^ 'm::finite \ 'a ^ 'n" + assumes lf: "linear f" + shows "(f x)$j = setsum (\i. (x$i) * (f (basis i)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs") +proof- + let ?M = "(UNIV :: 'm set)" + let ?N = "(UNIV :: 'n set)" + have fM: "finite ?M" by simp + have "?rhs = (setsum (\i.(x$i) *s f (basis i) ) ?M)$j" + unfolding vector_smult_component[symmetric] + unfolding setsum_component[of "(\i.(x$i) *s f (basis i :: 'a^'m))" ?M] + .. + then show ?thesis unfolding linear_setsum_mul[OF lf fM, symmetric] basis_expansion .. +qed + +text{* Inverse matrices (not necessarily square) *} + +definition "invertible(A::'a::semiring_1^'n^'m) \ (\A'::'a^'m^'n. A ** A' = mat 1 \ A' ** A = mat 1)" + +definition "matrix_inv(A:: 'a::semiring_1^'n^'m) = + (SOME A'::'a^'m^'n. A ** A' = mat 1 \ A' ** A = mat 1)" + +text{* Correspondence between matrices and linear operators. *} + +definition matrix:: "('a::{plus,times, one, zero}^'m \ 'a ^ 'n) \ 'a^'m^'n" +where "matrix f = (\ i j. (f(basis j))$i)" + +lemma matrix_vector_mul_linear: "linear(\x. A *v (x::'a::comm_semiring_1 ^ 'n))" + by (simp add: linear_def matrix_vector_mult_def Cart_eq ring_simps setsum_right_distrib setsum_addf) + +lemma matrix_works: assumes lf: "linear f" shows "matrix f *v x = f (x::'a::comm_ring_1 ^ 'n::finite)" +apply (simp add: matrix_def matrix_vector_mult_def Cart_eq mult_commute) +apply clarify +apply (rule linear_componentwise[OF lf, symmetric]) +done + +lemma matrix_vector_mul: "linear f ==> f = (\x. matrix f *v (x::'a::comm_ring_1 ^ 'n::finite))" by (simp add: ext matrix_works) + +lemma matrix_of_matrix_vector_mul: "matrix(\x. A *v (x :: 'a:: comm_ring_1 ^ 'n::finite)) = A" + by (simp add: matrix_eq matrix_vector_mul_linear matrix_works) + +lemma matrix_compose: + assumes lf: "linear (f::'a::comm_ring_1^'n::finite \ 'a^'m::finite)" + and lg: "linear (g::'a::comm_ring_1^'m::finite \ 'a^'k)" + shows "matrix (g o f) = matrix g ** matrix f" + using lf lg linear_compose[OF lf lg] matrix_works[OF linear_compose[OF lf lg]] + by (simp add: matrix_eq matrix_works matrix_vector_mul_assoc[symmetric] o_def) + +lemma matrix_vector_column:"(A::'a::comm_semiring_1^'n^'m) *v x = setsum (\i. (x$i) *s ((transp A)$i)) (UNIV:: 'n set)" + by (simp add: matrix_vector_mult_def transp_def Cart_eq mult_commute) + +lemma adjoint_matrix: "adjoint(\x. (A::'a::comm_ring_1^'n::finite^'m::finite) *v x) = (\x. transp A *v x)" + apply (rule adjoint_unique[symmetric]) + apply (rule matrix_vector_mul_linear) + apply (simp add: transp_def dot_def matrix_vector_mult_def setsum_left_distrib setsum_right_distrib) + apply (subst setsum_commute) + apply (auto simp add: mult_ac) + done + +lemma matrix_adjoint: assumes lf: "linear (f :: 'a::comm_ring_1^'n::finite \ 'a ^ 'm::finite)" + shows "matrix(adjoint f) = transp(matrix f)" + apply (subst matrix_vector_mul[OF lf]) + unfolding adjoint_matrix matrix_of_matrix_vector_mul .. + +subsection{* Interlude: Some properties of real sets *} + +lemma seq_mono_lemma: assumes "\(n::nat) \ m. (d n :: real) < e n" and "\n \ m. e n <= e m" + shows "\n \ m. d n < e m" + using prems apply auto + apply (erule_tac x="n" in allE) + apply (erule_tac x="n" in allE) + apply auto + done + + +lemma real_convex_bound_lt: + assumes xa: "(x::real) < a" and ya: "y < a" and u: "0 <= u" and v: "0 <= v" + and uv: "u + v = 1" + shows "u * x + v * y < a" +proof- + have uv': "u = 0 \ v \ 0" using u v uv by arith + have "a = a * (u + v)" unfolding uv by simp + hence th: "u * a + v * a = a" by (simp add: ring_simps) + from xa u have "u \ 0 \ u*x < u*a" by (simp add: mult_compare_simps) + from ya v have "v \ 0 \ v * y < v * a" by (simp add: mult_compare_simps) + from xa ya u v have "u * x + v * y < u * a + v * a" + apply (cases "u = 0", simp_all add: uv') + apply(rule mult_strict_left_mono) + using uv' apply simp_all + + apply (rule add_less_le_mono) + apply(rule mult_strict_left_mono) + apply simp_all + apply (rule mult_left_mono) + apply simp_all + done + thus ?thesis unfolding th . +qed + +lemma real_convex_bound_le: + assumes xa: "(x::real) \ a" and ya: "y \ a" and u: "0 <= u" and v: "0 <= v" + and uv: "u + v = 1" + shows "u * x + v * y \ a" +proof- + from xa ya u v have "u * x + v * y \ u * a + v * a" by (simp add: add_mono mult_left_mono) + also have "\ \ (u + v) * a" by (simp add: ring_simps) + finally show ?thesis unfolding uv by simp +qed + +lemma infinite_enumerate: assumes fS: "infinite S" + shows "\r. subseq r \ (\n. r n \ S)" +unfolding subseq_def +using enumerate_in_set[OF fS] enumerate_mono[of _ _ S] fS by auto + +lemma approachable_lt_le: "(\(d::real)>0. \x. f x < d \ P x) \ (\d>0. \x. f x \ d \ P x)" +apply auto +apply (rule_tac x="d/2" in exI) +apply auto +done + + +lemma triangle_lemma: + assumes x: "0 <= (x::real)" and y:"0 <= y" and z: "0 <= z" and xy: "x^2 <= y^2 + z^2" + shows "x <= y + z" +proof- + have "y^2 + z^2 \ y^2 + 2*y*z + z^2" using z y by (simp add: zero_compare_simps) + with xy have th: "x ^2 \ (y+z)^2" by (simp add: power2_eq_square ring_simps) + from y z have yz: "y + z \ 0" by arith + from power2_le_imp_le[OF th yz] show ?thesis . +qed + + +lemma lambda_skolem: "(\i. \x. P i x) \ + (\x::'a ^ 'n. \i. P i (x$i))" (is "?lhs \ ?rhs") +proof- + let ?S = "(UNIV :: 'n set)" + {assume H: "?rhs" + then have ?lhs by auto} + moreover + {assume H: "?lhs" + then obtain f where f:"\i. P i (f i)" unfolding choice_iff by metis + let ?x = "(\ i. (f i)) :: 'a ^ 'n" + {fix i + from f have "P i (f i)" by metis + then have "P i (?x$i)" by auto + } + hence "\i. P i (?x$i)" by metis + hence ?rhs by metis } + ultimately show ?thesis by metis +qed + +(* Supremum and infimum of real sets *) + + +definition rsup:: "real set \ real" where + "rsup S = (SOME a. isLub UNIV S a)" + +lemma rsup_alt: "rsup S = (SOME a. (\x \ S. x \ a) \ (\b. (\x \ S. x \ b) \ a \ b))" by (auto simp add: isLub_def rsup_def leastP_def isUb_def setle_def setge_def) + +lemma rsup: assumes Se: "S \ {}" and b: "\b. S *<= b" + shows "isLub UNIV S (rsup S)" +using Se b +unfolding rsup_def +apply clarify +apply (rule someI_ex) +apply (rule reals_complete) +by (auto simp add: isUb_def setle_def) + +lemma rsup_le: assumes Se: "S \ {}" and Sb: "S *<= b" shows "rsup S \ b" +proof- + from Sb have bu: "isUb UNIV S b" by (simp add: isUb_def setle_def) + from rsup[OF Se] Sb have "isLub UNIV S (rsup S)" by blast + then show ?thesis using bu by (auto simp add: isLub_def leastP_def setle_def setge_def) +qed + +lemma rsup_finite_Max: assumes fS: "finite S" and Se: "S \ {}" + shows "rsup S = Max S" +using fS Se +proof- + let ?m = "Max S" + from Max_ge[OF fS] have Sm: "\ x\ S. x \ ?m" by metis + with rsup[OF Se] have lub: "isLub UNIV S (rsup S)" by (metis setle_def) + from Max_in[OF fS Se] lub have mrS: "?m \ rsup S" + by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def) + moreover + have "rsup S \ ?m" using Sm lub + by (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def) + ultimately show ?thesis by arith +qed + +lemma rsup_finite_in: assumes fS: "finite S" and Se: "S \ {}" + shows "rsup S \ S" + using rsup_finite_Max[OF fS Se] Max_in[OF fS Se] by metis + +lemma rsup_finite_Ub: assumes fS: "finite S" and Se: "S \ {}" + shows "isUb S S (rsup S)" + using rsup_finite_Max[OF fS Se] rsup_finite_in[OF fS Se] Max_ge[OF fS] + unfolding isUb_def setle_def by metis + +lemma rsup_finite_ge_iff: assumes fS: "finite S" and Se: "S \ {}" + shows "a \ rsup S \ (\ x \ S. a \ x)" +using rsup_finite_Ub[OF fS Se] by (auto simp add: isUb_def setle_def) + +lemma rsup_finite_le_iff: assumes fS: "finite S" and Se: "S \ {}" + shows "a \ rsup S \ (\ x \ S. a \ x)" +using rsup_finite_Ub[OF fS Se] by (auto simp add: isUb_def setle_def) + +lemma rsup_finite_gt_iff: assumes fS: "finite S" and Se: "S \ {}" + shows "a < rsup S \ (\ x \ S. a < x)" +using rsup_finite_Ub[OF fS Se] by (auto simp add: isUb_def setle_def) + +lemma rsup_finite_lt_iff: assumes fS: "finite S" and Se: "S \ {}" + shows "a > rsup S \ (\ x \ S. a > x)" +using rsup_finite_Ub[OF fS Se] by (auto simp add: isUb_def setle_def) + +lemma rsup_unique: assumes b: "S *<= b" and S: "\b' < b. \x \ S. b' < x" + shows "rsup S = b" +using b S +unfolding setle_def rsup_alt +apply - +apply (rule some_equality) +apply (metis linorder_not_le order_eq_iff[symmetric])+ +done + +lemma rsup_le_subset: "S\{} \ S \ T \ (\b. T *<= b) \ rsup S \ rsup T" + apply (rule rsup_le) + apply simp + using rsup[of T] by (auto simp add: isLub_def leastP_def setge_def setle_def isUb_def) + +lemma isUb_def': "isUb R S = (\x. S *<= x \ x \ R)" + apply (rule ext) + by (metis isUb_def) + +lemma UNIV_trivial: "UNIV x" using UNIV_I[of x] by (metis mem_def) +lemma rsup_bounds: assumes Se: "S \ {}" and l: "a <=* S" and u: "S *<= b" + shows "a \ rsup S \ rsup S \ b" +proof- + from rsup[OF Se] u have lub: "isLub UNIV S (rsup S)" by blast + hence b: "rsup S \ b" using u by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def') + from Se obtain y where y: "y \ S" by blast + from lub l have "a \ rsup S" apply (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def') + apply (erule ballE[where x=y]) + apply (erule ballE[where x=y]) + apply arith + using y apply auto + done + with b show ?thesis by blast +qed + +lemma rsup_abs_le: "S \ {} \ (\x\S. \x\ \ a) \ \rsup S\ \ a" + unfolding abs_le_interval_iff using rsup_bounds[of S "-a" a] + by (auto simp add: setge_def setle_def) + +lemma rsup_asclose: assumes S:"S \ {}" and b: "\x\S. \x - l\ \ e" shows "\rsup S - l\ \ e" +proof- + have th: "\(x::real) l e. \x - l\ \ e \ l - e \ x \ x \ l + e" by arith + show ?thesis using S b rsup_bounds[of S "l - e" "l+e"] unfolding th + by (auto simp add: setge_def setle_def) +qed + +definition rinf:: "real set \ real" where + "rinf S = (SOME a. isGlb UNIV S a)" + +lemma rinf_alt: "rinf S = (SOME a. (\x \ S. x \ a) \ (\b. (\x \ S. x \ b) \ a \ b))" by (auto simp add: isGlb_def rinf_def greatestP_def isLb_def setle_def setge_def) + +lemma reals_complete_Glb: assumes Se: "\x. x \ S" and lb: "\ y. isLb UNIV S y" + shows "\(t::real). isGlb UNIV S t" +proof- + let ?M = "uminus ` S" + from lb have th: "\y. isUb UNIV ?M y" apply (auto simp add: isUb_def isLb_def setle_def setge_def) + by (rule_tac x="-y" in exI, auto) + from Se have Me: "\x. x \ ?M" by blast + from reals_complete[OF Me th] obtain t where t: "isLub UNIV ?M t" by blast + have "isGlb UNIV S (- t)" using t + apply (auto simp add: isLub_def isGlb_def leastP_def greatestP_def setle_def setge_def isUb_def isLb_def) + apply (erule_tac x="-y" in allE) + apply auto + done + then show ?thesis by metis +qed + +lemma rinf: assumes Se: "S \ {}" and b: "\b. b <=* S" + shows "isGlb UNIV S (rinf S)" +using Se b +unfolding rinf_def +apply clarify +apply (rule someI_ex) +apply (rule reals_complete_Glb) +apply (auto simp add: isLb_def setle_def setge_def) +done + +lemma rinf_ge: assumes Se: "S \ {}" and Sb: "b <=* S" shows "rinf S \ b" +proof- + from Sb have bu: "isLb UNIV S b" by (simp add: isLb_def setge_def) + from rinf[OF Se] Sb have "isGlb UNIV S (rinf S)" by blast + then show ?thesis using bu by (auto simp add: isGlb_def greatestP_def setle_def setge_def) +qed + +lemma rinf_finite_Min: assumes fS: "finite S" and Se: "S \ {}" + shows "rinf S = Min S" +using fS Se +proof- + let ?m = "Min S" + from Min_le[OF fS] have Sm: "\ x\ S. x \ ?m" by metis + with rinf[OF Se] have glb: "isGlb UNIV S (rinf S)" by (metis setge_def) + from Min_in[OF fS Se] glb have mrS: "?m \ rinf S" + by (auto simp add: isGlb_def greatestP_def setle_def setge_def isLb_def) + moreover + have "rinf S \ ?m" using Sm glb + by (auto simp add: isGlb_def greatestP_def isLb_def setle_def setge_def) + ultimately show ?thesis by arith +qed + +lemma rinf_finite_in: assumes fS: "finite S" and Se: "S \ {}" + shows "rinf S \ S" + using rinf_finite_Min[OF fS Se] Min_in[OF fS Se] by metis + +lemma rinf_finite_Lb: assumes fS: "finite S" and Se: "S \ {}" + shows "isLb S S (rinf S)" + using rinf_finite_Min[OF fS Se] rinf_finite_in[OF fS Se] Min_le[OF fS] + unfolding isLb_def setge_def by metis + +lemma rinf_finite_ge_iff: assumes fS: "finite S" and Se: "S \ {}" + shows "a \ rinf S \ (\ x \ S. a \ x)" +using rinf_finite_Lb[OF fS Se] by (auto simp add: isLb_def setge_def) + +lemma rinf_finite_le_iff: assumes fS: "finite S" and Se: "S \ {}" + shows "a \ rinf S \ (\ x \ S. a \ x)" +using rinf_finite_Lb[OF fS Se] by (auto simp add: isLb_def setge_def) + +lemma rinf_finite_gt_iff: assumes fS: "finite S" and Se: "S \ {}" + shows "a < rinf S \ (\ x \ S. a < x)" +using rinf_finite_Lb[OF fS Se] by (auto simp add: isLb_def setge_def) + +lemma rinf_finite_lt_iff: assumes fS: "finite S" and Se: "S \ {}" + shows "a > rinf S \ (\ x \ S. a > x)" +using rinf_finite_Lb[OF fS Se] by (auto simp add: isLb_def setge_def) + +lemma rinf_unique: assumes b: "b <=* S" and S: "\b' > b. \x \ S. b' > x" + shows "rinf S = b" +using b S +unfolding setge_def rinf_alt +apply - +apply (rule some_equality) +apply (metis linorder_not_le order_eq_iff[symmetric])+ +done + +lemma rinf_ge_subset: "S\{} \ S \ T \ (\b. b <=* T) \ rinf S >= rinf T" + apply (rule rinf_ge) + apply simp + using rinf[of T] by (auto simp add: isGlb_def greatestP_def setge_def setle_def isLb_def) + +lemma isLb_def': "isLb R S = (\x. x <=* S \ x \ R)" + apply (rule ext) + by (metis isLb_def) + +lemma rinf_bounds: assumes Se: "S \ {}" and l: "a <=* S" and u: "S *<= b" + shows "a \ rinf S \ rinf S \ b" +proof- + from rinf[OF Se] l have lub: "isGlb UNIV S (rinf S)" by blast + hence b: "a \ rinf S" using l by (auto simp add: isGlb_def greatestP_def setle_def setge_def isLb_def') + from Se obtain y where y: "y \ S" by blast + from lub u have "b \ rinf S" apply (auto simp add: isGlb_def greatestP_def setle_def setge_def isLb_def') + apply (erule ballE[where x=y]) + apply (erule ballE[where x=y]) + apply arith + using y apply auto + done + with b show ?thesis by blast +qed + +lemma rinf_abs_ge: "S \ {} \ (\x\S. \x\ \ a) \ \rinf S\ \ a" + unfolding abs_le_interval_iff using rinf_bounds[of S "-a" a] + by (auto simp add: setge_def setle_def) + +lemma rinf_asclose: assumes S:"S \ {}" and b: "\x\S. \x - l\ \ e" shows "\rinf S - l\ \ e" +proof- + have th: "\(x::real) l e. \x - l\ \ e \ l - e \ x \ x \ l + e" by arith + show ?thesis using S b rinf_bounds[of S "l - e" "l+e"] unfolding th + by (auto simp add: setge_def setle_def) +qed + + + +subsection{* Operator norm. *} + +definition "onorm f = rsup {norm (f x)| x. norm x = 1}" + +lemma norm_bound_generalize: + fixes f:: "real ^'n::finite \ real^'m::finite" + assumes lf: "linear f" + shows "(\x. norm x = 1 \ norm (f x) \ b) \ (\x. norm (f x) \ b * norm x)" (is "?lhs \ ?rhs") +proof- + {assume H: ?rhs + {fix x :: "real^'n" assume x: "norm x = 1" + from H[rule_format, of x] x have "norm (f x) \ b" by simp} + then have ?lhs by blast } + + moreover + {assume H: ?lhs + from H[rule_format, of "basis arbitrary"] + have bp: "b \ 0" using norm_ge_zero[of "f (basis arbitrary)"] + by (auto simp add: norm_basis elim: order_trans [OF norm_ge_zero]) + {fix x :: "real ^'n" + {assume "x = 0" + then have "norm (f x) \ b * norm x" by (simp add: linear_0[OF lf] bp)} + moreover + {assume x0: "x \ 0" + hence n0: "norm x \ 0" by (metis norm_eq_zero) + let ?c = "1/ norm x" + have "norm (?c*s x) = 1" using x0 by (simp add: n0 norm_mul) + with H have "norm (f(?c*s x)) \ b" by blast + hence "?c * norm (f x) \ b" + by (simp add: linear_cmul[OF lf] norm_mul) + hence "norm (f x) \ b * norm x" + using n0 norm_ge_zero[of x] by (auto simp add: field_simps)} + ultimately have "norm (f x) \ b * norm x" by blast} + then have ?rhs by blast} + ultimately show ?thesis by blast +qed + +lemma onorm: + fixes f:: "real ^'n::finite \ real ^'m::finite" + assumes lf: "linear f" + shows "norm (f x) <= onorm f * norm x" + and "\x. norm (f x) <= b * norm x \ onorm f <= b" +proof- + { + let ?S = "{norm (f x) |x. norm x = 1}" + have Se: "?S \ {}" using norm_basis by auto + from linear_bounded[OF lf] have b: "\ b. ?S *<= b" + unfolding norm_bound_generalize[OF lf, symmetric] by (auto simp add: setle_def) + {from rsup[OF Se b, unfolded onorm_def[symmetric]] + show "norm (f x) <= onorm f * norm x" + apply - + apply (rule spec[where x = x]) + unfolding norm_bound_generalize[OF lf, symmetric] + by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)} + { + show "\x. norm (f x) <= b * norm x \ onorm f <= b" + using rsup[OF Se b, unfolded onorm_def[symmetric]] + unfolding norm_bound_generalize[OF lf, symmetric] + by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)} + } +qed + +lemma onorm_pos_le: assumes lf: "linear (f::real ^'n::finite \ real ^'m::finite)" shows "0 <= onorm f" + using order_trans[OF norm_ge_zero onorm(1)[OF lf, of "basis arbitrary"], unfolded norm_basis] by simp + +lemma onorm_eq_0: assumes lf: "linear (f::real ^'n::finite \ real ^'m::finite)" + shows "onorm f = 0 \ (\x. f x = 0)" + using onorm[OF lf] + apply (auto simp add: onorm_pos_le) + apply atomize + apply (erule allE[where x="0::real"]) + using onorm_pos_le[OF lf] + apply arith + done + +lemma onorm_const: "onorm(\x::real^'n::finite. (y::real ^ 'm::finite)) = norm y" +proof- + let ?f = "\x::real^'n. (y::real ^ 'm)" + have th: "{norm (?f x)| x. norm x = 1} = {norm y}" + by(auto intro: vector_choose_size set_ext) + show ?thesis + unfolding onorm_def th + apply (rule rsup_unique) by (simp_all add: setle_def) +qed + +lemma onorm_pos_lt: assumes lf: "linear (f::real ^ 'n::finite \ real ^'m::finite)" + shows "0 < onorm f \ ~(\x. f x = 0)" + unfolding onorm_eq_0[OF lf, symmetric] + using onorm_pos_le[OF lf] by arith + +lemma onorm_compose: + assumes lf: "linear (f::real ^'n::finite \ real ^'m::finite)" + and lg: "linear (g::real^'k::finite \ real^'n::finite)" + shows "onorm (f o g) <= onorm f * onorm g" + apply (rule onorm(2)[OF linear_compose[OF lg lf], rule_format]) + unfolding o_def + apply (subst mult_assoc) + apply (rule order_trans) + apply (rule onorm(1)[OF lf]) + apply (rule mult_mono1) + apply (rule onorm(1)[OF lg]) + apply (rule onorm_pos_le[OF lf]) + done + +lemma onorm_neg_lemma: assumes lf: "linear (f::real ^'n::finite \ real^'m::finite)" + shows "onorm (\x. - f x) \ onorm f" + using onorm[OF linear_compose_neg[OF lf]] onorm[OF lf] + unfolding norm_minus_cancel by metis + +lemma onorm_neg: assumes lf: "linear (f::real ^'n::finite \ real^'m::finite)" + shows "onorm (\x. - f x) = onorm f" + using onorm_neg_lemma[OF lf] onorm_neg_lemma[OF linear_compose_neg[OF lf]] + by simp + +lemma onorm_triangle: + assumes lf: "linear (f::real ^'n::finite \ real ^'m::finite)" and lg: "linear g" + shows "onorm (\x. f x + g x) <= onorm f + onorm g" + apply(rule onorm(2)[OF linear_compose_add[OF lf lg], rule_format]) + apply (rule order_trans) + apply (rule norm_triangle_ineq) + apply (simp add: distrib) + apply (rule add_mono) + apply (rule onorm(1)[OF lf]) + apply (rule onorm(1)[OF lg]) + done + +lemma onorm_triangle_le: "linear (f::real ^'n::finite \ real ^'m::finite) \ linear g \ onorm(f) + onorm(g) <= e + \ onorm(\x. f x + g x) <= e" + apply (rule order_trans) + apply (rule onorm_triangle) + apply assumption+ + done + +lemma onorm_triangle_lt: "linear (f::real ^'n::finite \ real ^'m::finite) \ linear g \ onorm(f) + onorm(g) < e + ==> onorm(\x. f x + g x) < e" + apply (rule order_le_less_trans) + apply (rule onorm_triangle) + by assumption+ + +(* "lift" from 'a to 'a^1 and "drop" from 'a^1 to 'a -- FIXME: potential use of transfer *) + +definition vec1:: "'a \ 'a ^ 1" where "vec1 x = (\ i. x)" + +definition dest_vec1:: "'a ^1 \ 'a" + where "dest_vec1 x = (x$1)" + +lemma vec1_component[simp]: "(vec1 x)$1 = x" + by (simp add: vec1_def) + +lemma vec1_dest_vec1[simp]: "vec1(dest_vec1 x) = x" "dest_vec1(vec1 y) = y" + by (simp_all add: vec1_def dest_vec1_def Cart_eq forall_1) + +lemma forall_vec1: "(\x. P x) \ (\x. P (vec1 x))" by (metis vec1_dest_vec1) + +lemma exists_vec1: "(\x. P x) \ (\x. P(vec1 x))" by (metis vec1_dest_vec1) + +lemma forall_dest_vec1: "(\x. P x) \ (\x. P(dest_vec1 x))" by (metis vec1_dest_vec1) + +lemma exists_dest_vec1: "(\x. P x) \ (\x. P(dest_vec1 x))"by (metis vec1_dest_vec1) + +lemma vec1_eq[simp]: "vec1 x = vec1 y \ x = y" by (metis vec1_dest_vec1) + +lemma dest_vec1_eq[simp]: "dest_vec1 x = dest_vec1 y \ x = y" by (metis vec1_dest_vec1) + +lemma vec1_in_image_vec1: "vec1 x \ (vec1 ` S) \ x \ S" by auto + +lemma vec1_vec: "vec1 x = vec x" by (vector vec1_def) + +lemma vec1_add: "vec1(x + y) = vec1 x + vec1 y" by (vector vec1_def) +lemma vec1_sub: "vec1(x - y) = vec1 x - vec1 y" by (vector vec1_def) +lemma vec1_cmul: "vec1(c* x) = c *s vec1 x " by (vector vec1_def) +lemma vec1_neg: "vec1(- x) = - vec1 x " by (vector vec1_def) + +lemma vec1_setsum: assumes fS: "finite S" + shows "vec1(setsum f S) = setsum (vec1 o f) S" + apply (induct rule: finite_induct[OF fS]) + apply (simp add: vec1_vec) + apply (auto simp add: vec1_add) + done + +lemma dest_vec1_lambda: "dest_vec1(\ i. x i) = x 1" + by (simp add: dest_vec1_def) + +lemma dest_vec1_vec: "dest_vec1(vec x) = x" + by (simp add: vec1_vec[symmetric]) + +lemma dest_vec1_add: "dest_vec1(x + y) = dest_vec1 x + dest_vec1 y" + by (metis vec1_dest_vec1 vec1_add) + +lemma dest_vec1_sub: "dest_vec1(x - y) = dest_vec1 x - dest_vec1 y" + by (metis vec1_dest_vec1 vec1_sub) + +lemma dest_vec1_cmul: "dest_vec1(c*sx) = c * dest_vec1 x" + by (metis vec1_dest_vec1 vec1_cmul) + +lemma dest_vec1_neg: "dest_vec1(- x) = - dest_vec1 x" + by (metis vec1_dest_vec1 vec1_neg) + +lemma dest_vec1_0[simp]: "dest_vec1 0 = 0" by (metis vec_0 dest_vec1_vec) + +lemma dest_vec1_sum: assumes fS: "finite S" + shows "dest_vec1(setsum f S) = setsum (dest_vec1 o f) S" + apply (induct rule: finite_induct[OF fS]) + apply (simp add: dest_vec1_vec) + apply (auto simp add: dest_vec1_add) + done + +lemma norm_vec1: "norm(vec1 x) = abs(x)" + by (simp add: vec1_def norm_real) + +lemma dist_vec1: "dist(vec1 x) (vec1 y) = abs(x - y)" + by (simp only: dist_real vec1_component) +lemma abs_dest_vec1: "norm x = \dest_vec1 x\" + by (metis vec1_dest_vec1 norm_vec1) + +lemma linear_vmul_dest_vec1: + fixes f:: "'a::semiring_1^'n \ 'a^1" + shows "linear f \ linear (\x. dest_vec1(f x) *s v)" + unfolding dest_vec1_def + apply (rule linear_vmul_component) + by auto + +lemma linear_from_scalars: + assumes lf: "linear (f::'a::comm_ring_1 ^1 \ 'a^'n)" + shows "f = (\x. dest_vec1 x *s column 1 (matrix f))" + apply (rule ext) + apply (subst matrix_works[OF lf, symmetric]) + apply (auto simp add: Cart_eq matrix_vector_mult_def dest_vec1_def column_def mult_commute UNIV_1) + done + +lemma linear_to_scalars: assumes lf: "linear (f::'a::comm_ring_1 ^'n::finite \ 'a^1)" + shows "f = (\x. vec1(row 1 (matrix f) \ x))" + apply (rule ext) + apply (subst matrix_works[OF lf, symmetric]) + apply (simp add: Cart_eq matrix_vector_mult_def vec1_def row_def dot_def mult_commute forall_1) + done + +lemma dest_vec1_eq_0: "dest_vec1 x = 0 \ x = 0" + by (simp add: dest_vec1_eq[symmetric]) + +lemma setsum_scalars: assumes fS: "finite S" + shows "setsum f S = vec1 (setsum (dest_vec1 o f) S)" + unfolding vec1_setsum[OF fS] by simp + +lemma dest_vec1_wlog_le: "(\(x::'a::linorder ^ 1) y. P x y \ P y x) \ (\x y. dest_vec1 x <= dest_vec1 y ==> P x y) \ P x y" + apply (cases "dest_vec1 x \ dest_vec1 y") + apply simp + apply (subgoal_tac "dest_vec1 y \ dest_vec1 x") + apply (auto) + done + +text{* Pasting vectors. *} + +lemma linear_fstcart: "linear fstcart" + by (auto simp add: linear_def Cart_eq) + +lemma linear_sndcart: "linear sndcart" + by (auto simp add: linear_def Cart_eq) + +lemma fstcart_vec[simp]: "fstcart(vec x) = vec x" + by (simp add: Cart_eq) + +lemma fstcart_add[simp]:"fstcart(x + y) = fstcart (x::'a::{plus,times}^('b + 'c)) + fstcart y" + by (simp add: Cart_eq) + +lemma fstcart_cmul[simp]:"fstcart(c*s x) = c*s fstcart (x::'a::{plus,times}^('b + 'c))" + by (simp add: Cart_eq) + +lemma fstcart_neg[simp]:"fstcart(- x) = - fstcart (x::'a::ring_1^('b + 'c))" + by (simp add: Cart_eq) + +lemma fstcart_sub[simp]:"fstcart(x - y) = fstcart (x::'a::ring_1^('b + 'c)) - fstcart y" + by (simp add: Cart_eq) + +lemma fstcart_setsum: + fixes f:: "'d \ 'a::semiring_1^_" + assumes fS: "finite S" + shows "fstcart (setsum f S) = setsum (\i. fstcart (f i)) S" + by (induct rule: finite_induct[OF fS], simp_all add: vec_0[symmetric] del: vec_0) + +lemma sndcart_vec[simp]: "sndcart(vec x) = vec x" + by (simp add: Cart_eq) + +lemma sndcart_add[simp]:"sndcart(x + y) = sndcart (x::'a::{plus,times}^('b + 'c)) + sndcart y" + by (simp add: Cart_eq) + +lemma sndcart_cmul[simp]:"sndcart(c*s x) = c*s sndcart (x::'a::{plus,times}^('b + 'c))" + by (simp add: Cart_eq) + +lemma sndcart_neg[simp]:"sndcart(- x) = - sndcart (x::'a::ring_1^('b + 'c))" + by (simp add: Cart_eq) + +lemma sndcart_sub[simp]:"sndcart(x - y) = sndcart (x::'a::ring_1^('b + 'c)) - sndcart y" + by (simp add: Cart_eq) + +lemma sndcart_setsum: + fixes f:: "'d \ 'a::semiring_1^_" + assumes fS: "finite S" + shows "sndcart (setsum f S) = setsum (\i. sndcart (f i)) S" + by (induct rule: finite_induct[OF fS], simp_all add: vec_0[symmetric] del: vec_0) + +lemma pastecart_vec[simp]: "pastecart (vec x) (vec x) = vec x" + by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart) + +lemma pastecart_add[simp]:"pastecart (x1::'a::{plus,times}^_) y1 + pastecart x2 y2 = pastecart (x1 + x2) (y1 + y2)" + by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart) + +lemma pastecart_cmul[simp]: "pastecart (c *s (x1::'a::{plus,times}^_)) (c *s y1) = c *s pastecart x1 y1" + by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart) + +lemma pastecart_neg[simp]: "pastecart (- (x::'a::ring_1^_)) (- y) = - pastecart x y" + unfolding vector_sneg_minus1 pastecart_cmul .. + +lemma pastecart_sub: "pastecart (x1::'a::ring_1^_) y1 - pastecart x2 y2 = pastecart (x1 - x2) (y1 - y2)" + by (simp add: diff_def pastecart_neg[symmetric] del: pastecart_neg) + +lemma pastecart_setsum: + fixes f:: "'d \ 'a::semiring_1^_" + assumes fS: "finite S" + shows "pastecart (setsum f S) (setsum g S) = setsum (\i. pastecart (f i) (g i)) S" + by (simp add: pastecart_eq fstcart_setsum[OF fS] sndcart_setsum[OF fS] fstcart_pastecart sndcart_pastecart) + +lemma setsum_Plus: + "\finite A; finite B\ \ + (\x\A <+> B. g x) = (\x\A. g (Inl x)) + (\x\B. g (Inr x))" + unfolding Plus_def + by (subst setsum_Un_disjoint, auto simp add: setsum_reindex) + +lemma setsum_UNIV_sum: + fixes g :: "'a::finite + 'b::finite \ _" + shows "(\x\UNIV. g x) = (\x\UNIV. g (Inl x)) + (\x\UNIV. g (Inr x))" + apply (subst UNIV_Plus_UNIV [symmetric]) + apply (rule setsum_Plus [OF finite finite]) + done + +lemma norm_fstcart: "norm(fstcart x) <= norm (x::real ^('n::finite + 'm::finite))" +proof- + have th0: "norm x = norm (pastecart (fstcart x) (sndcart x))" + by (simp add: pastecart_fst_snd) + have th1: "fstcart x \ fstcart x \ pastecart (fstcart x) (sndcart x) \ pastecart (fstcart x) (sndcart x)" + by (simp add: dot_def setsum_UNIV_sum pastecart_def setsum_nonneg) + then show ?thesis + unfolding th0 + unfolding real_vector_norm_def real_sqrt_le_iff id_def + by (simp add: dot_def) +qed + +lemma dist_fstcart: "dist(fstcart (x::real^_)) (fstcart y) <= dist x y" + unfolding dist_norm by (metis fstcart_sub[symmetric] norm_fstcart) + +lemma norm_sndcart: "norm(sndcart x) <= norm (x::real ^('n::finite + 'm::finite))" +proof- + have th0: "norm x = norm (pastecart (fstcart x) (sndcart x))" + by (simp add: pastecart_fst_snd) + have th1: "sndcart x \ sndcart x \ pastecart (fstcart x) (sndcart x) \ pastecart (fstcart x) (sndcart x)" + by (simp add: dot_def setsum_UNIV_sum pastecart_def setsum_nonneg) + then show ?thesis + unfolding th0 + unfolding real_vector_norm_def real_sqrt_le_iff id_def + by (simp add: dot_def) +qed + +lemma dist_sndcart: "dist(sndcart (x::real^_)) (sndcart y) <= dist x y" + unfolding dist_norm by (metis sndcart_sub[symmetric] norm_sndcart) + +lemma dot_pastecart: "(pastecart (x1::'a::{times,comm_monoid_add}^'n::finite) (x2::'a::{times,comm_monoid_add}^'m::finite)) \ (pastecart y1 y2) = x1 \ y1 + x2 \ y2" + by (simp add: dot_def setsum_UNIV_sum pastecart_def) + +text {* TODO: move to NthRoot *} +lemma sqrt_add_le_add_sqrt: + assumes x: "0 \ x" and y: "0 \ y" + shows "sqrt (x + y) \ sqrt x + sqrt y" +apply (rule power2_le_imp_le) +apply (simp add: real_sum_squared_expand add_nonneg_nonneg x y) +apply (simp add: mult_nonneg_nonneg x y) +apply (simp add: add_nonneg_nonneg x y) +done + +lemma norm_pastecart: "norm (pastecart x y) <= norm x + norm y" + 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). *} + +definition hull :: "'a set set \ 'a set \ 'a set" (infixl "hull" 75) where + "S hull s = Inter {t. t \ S \ s \ t}" + +lemma hull_same: "s \ S \ S hull s = s" + unfolding hull_def by auto + +lemma hull_in: "(\T. T \ S ==> Inter T \ S) ==> (S hull s) \ S" +unfolding hull_def subset_iff by auto + +lemma hull_eq: "(\T. T \ S ==> Inter T \ S) ==> (S hull s) = s \ s \ S" +using hull_same[of s S] hull_in[of S s] by metis + + +lemma hull_hull: "S hull (S hull s) = S hull s" + unfolding hull_def by blast + +lemma hull_subset: "s \ (S hull s)" + unfolding hull_def by blast + +lemma hull_mono: " s \ t ==> (S hull s) \ (S hull t)" + unfolding hull_def by blast + +lemma hull_antimono: "S \ T ==> (T hull s) \ (S hull s)" + unfolding hull_def by blast + +lemma hull_minimal: "s \ t \ t \ S ==> (S hull s) \ t" + unfolding hull_def by blast + +lemma subset_hull: "t \ S ==> S hull s \ t \ s \ t" + unfolding hull_def by blast + +lemma hull_unique: "s \ t \ t \ S \ (\t'. s \ t' \ t' \ S ==> t \ t') + ==> (S hull s = t)" +unfolding hull_def by auto + +lemma hull_induct: "(\x. x\ S \ P x) \ Q {x. P x} \ \x\ Q hull S. P x" + using hull_minimal[of S "{x. P x}" Q] + by (auto simp add: subset_eq Collect_def mem_def) + +lemma hull_inc: "x \ S \ x \ P hull S" by (metis hull_subset subset_eq) + +lemma hull_union_subset: "(S hull s) \ (S hull t) \ (S hull (s \ t))" +unfolding Un_subset_iff by (metis hull_mono Un_upper1 Un_upper2) + +lemma hull_union: assumes T: "\T. T \ S ==> Inter T \ S" + shows "S hull (s \ t) = S hull (S hull s \ S hull t)" +apply rule +apply (rule hull_mono) +unfolding Un_subset_iff +apply (metis hull_subset Un_upper1 Un_upper2 subset_trans) +apply (rule hull_minimal) +apply (metis hull_union_subset) +apply (metis hull_in T) +done + +lemma hull_redundant_eq: "a \ (S hull s) \ (S hull (insert a s) = S hull s)" + unfolding hull_def by blast + +lemma hull_redundant: "a \ (S hull s) ==> (S hull (insert a s) = S hull s)" +by (metis hull_redundant_eq) + +text{* Archimedian properties and useful consequences. *} + +lemma real_arch_simple: "\n. x <= real (n::nat)" + using reals_Archimedean2[of x] apply auto by (rule_tac x="Suc n" in exI, auto) +lemmas real_arch_lt = reals_Archimedean2 + +lemmas real_arch = reals_Archimedean3 + +lemma real_arch_inv: "0 < e \ (\n::nat. n \ 0 \ 0 < inverse (real n) \ inverse (real n) < e)" + using reals_Archimedean + apply (auto simp add: field_simps inverse_positive_iff_positive) + apply (subgoal_tac "inverse (real n) > 0") + apply arith + apply simp + done + +lemma real_pow_lbound: "0 <= x ==> 1 + real n * x <= (1 + x) ^ n" +proof(induct n) + case 0 thus ?case by simp +next + case (Suc n) + hence h: "1 + real n * x \ (1 + x) ^ n" by simp + from h have p: "1 \ (1 + x) ^ n" using Suc.prems by simp + from h have "1 + real n * x + x \ (1 + x) ^ n + x" by simp + also have "\ \ (1 + x) ^ Suc n" apply (subst diff_le_0_iff_le[symmetric]) + apply (simp add: ring_simps) + using mult_left_mono[OF p Suc.prems] by simp + finally show ?case by (simp add: real_of_nat_Suc ring_simps) +qed + +lemma real_arch_pow: assumes x: "1 < (x::real)" shows "\n. y < x^n" +proof- + from x have x0: "x - 1 > 0" by arith + from real_arch[OF x0, rule_format, of y] + obtain n::nat where n:"y < real n * (x - 1)" by metis + from x0 have x00: "x- 1 \ 0" by arith + from real_pow_lbound[OF x00, of n] n + have "y < x^n" by auto + then show ?thesis by metis +qed + +lemma real_arch_pow2: "\n. (x::real) < 2^ n" + using real_arch_pow[of 2 x] by simp + +lemma real_arch_pow_inv: assumes y: "(y::real) > 0" and x1: "x < 1" + shows "\n. x^n < y" +proof- + {assume x0: "x > 0" + from x0 x1 have ix: "1 < 1/x" by (simp add: field_simps) + from real_arch_pow[OF ix, of "1/y"] + obtain n where n: "1/y < (1/x)^n" by blast + then + have ?thesis using y x0 by (auto simp add: field_simps power_divide) } + moreover + {assume "\ x > 0" with y x1 have ?thesis apply auto by (rule exI[where x=1], auto)} + ultimately show ?thesis by metis +qed + +lemma forall_pos_mono: "(\d e::real. d < e \ P d ==> P e) \ (\n::nat. n \ 0 ==> P(inverse(real n))) \ (\e. 0 < e ==> P e)" + by (metis real_arch_inv) + +lemma forall_pos_mono_1: "(\d e::real. d < e \ P d ==> P e) \ (\n. P(inverse(real (Suc n)))) ==> 0 < e ==> P e" + apply (rule forall_pos_mono) + apply auto + apply (atomize) + apply (erule_tac x="n - 1" in allE) + apply auto + done + +lemma real_archimedian_rdiv_eq_0: assumes x0: "x \ 0" and c: "c \ 0" and xc: "\(m::nat)>0. real m * x \ c" + shows "x = 0" +proof- + {assume "x \ 0" with x0 have xp: "x > 0" by arith + from real_arch[OF xp, rule_format, of c] obtain n::nat where n: "c < real n * x" by blast + with xc[rule_format, of n] have "n = 0" by arith + with n c have False by simp} + then show ?thesis by blast +qed + +(* ------------------------------------------------------------------------- *) +(* Relate max and min to sup and inf. *) +(* ------------------------------------------------------------------------- *) + +lemma real_max_rsup: "max x y = rsup {x,y}" +proof- + have f: "finite {x, y}" "{x,y} \ {}" by simp_all + from rsup_finite_le_iff[OF f, of "max x y"] have "rsup {x,y} \ max x y" by simp + moreover + have "max x y \ rsup {x,y}" using rsup_finite_ge_iff[OF f, of "max x y"] + by (simp add: linorder_linear) + ultimately show ?thesis by arith +qed + +lemma real_min_rinf: "min x y = rinf {x,y}" +proof- + have f: "finite {x, y}" "{x,y} \ {}" by simp_all + from rinf_finite_le_iff[OF f, of "min x y"] have "rinf {x,y} \ min x y" + by (simp add: linorder_linear) + moreover + have "min x y \ rinf {x,y}" using rinf_finite_ge_iff[OF f, of "min x y"] + by simp + ultimately show ?thesis by arith +qed + +(* ------------------------------------------------------------------------- *) +(* Geometric progression. *) +(* ------------------------------------------------------------------------- *) + +lemma sum_gp_basic: "((1::'a::{field}) - x) * setsum (\i. x^i) {0 .. n} = (1 - x^(Suc n))" + (is "?lhs = ?rhs") +proof- + {assume x1: "x = 1" hence ?thesis by simp} + moreover + {assume x1: "x\1" + hence x1': "x - 1 \ 0" "1 - x \ 0" "x - 1 = - (1 - x)" "- (1 - x) \ 0" by auto + from geometric_sum[OF x1, of "Suc n", unfolded x1'] + have "(- (1 - x)) * setsum (\i. x^i) {0 .. n} = - (1 - x^(Suc n))" + unfolding atLeastLessThanSuc_atLeastAtMost + using x1' apply (auto simp only: field_simps) + apply (simp add: ring_simps) + done + then have ?thesis by (simp add: ring_simps) } + ultimately show ?thesis by metis +qed + +lemma sum_gp_multiplied: assumes mn: "m <= n" + shows "((1::'a::{field}) - x) * setsum (op ^ x) {m..n} = x^m - x^ Suc n" + (is "?lhs = ?rhs") +proof- + let ?S = "{0..(n - m)}" + from mn have mn': "n - m \ 0" by arith + let ?f = "op + m" + have i: "inj_on ?f ?S" unfolding inj_on_def by auto + have f: "?f ` ?S = {m..n}" + using mn apply (auto simp add: image_iff Bex_def) by arith + have th: "op ^ x o op + m = (\i. x^m * x^i)" + by (rule ext, simp add: power_add power_mult) + from setsum_reindex[OF i, of "op ^ x", unfolded f th setsum_right_distrib[symmetric]] + have "?lhs = x^m * ((1 - x) * setsum (op ^ x) {0..n - m})" by simp + then show ?thesis unfolding sum_gp_basic using mn + by (simp add: ring_simps power_add[symmetric]) +qed + +lemma sum_gp: "setsum (op ^ (x::'a::{field})) {m .. n} = + (if n < m then 0 else if x = 1 then of_nat ((n + 1) - m) + else (x^ m - x^ (Suc n)) / (1 - x))" +proof- + {assume nm: "n < m" hence ?thesis by simp} + moreover + {assume "\ n < m" hence nm: "m \ n" by arith + {assume x: "x = 1" hence ?thesis by simp} + moreover + {assume x: "x \ 1" hence nz: "1 - x \ 0" by simp + from sum_gp_multiplied[OF nm, of x] nz have ?thesis by (simp add: field_simps)} + ultimately have ?thesis by metis + } + ultimately show ?thesis by metis +qed + +lemma sum_gp_offset: "setsum (op ^ (x::'a::{field})) {m .. m+n} = + (if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))" + unfolding sum_gp[of x m "m + n"] power_Suc + by (simp add: ring_simps power_add) + + +subsection{* A bit of linear algebra. *} + +definition "subspace S \ 0 \ S \ (\x\ S. \y \S. x + y \ S) \ (\c. \x \S. c *s x \S )" +definition "span S = (subspace hull S)" +definition "dependent S \ (\a \ S. a \ span(S - {a}))" +abbreviation "independent s == ~(dependent s)" + +(* Closure properties of subspaces. *) + +lemma subspace_UNIV[simp]: "subspace(UNIV)" by (simp add: subspace_def) + +lemma subspace_0: "subspace S ==> 0 \ S" by (metis subspace_def) + +lemma subspace_add: "subspace S \ x \ S \ y \ S ==> x + y \ S" + by (metis subspace_def) + +lemma subspace_mul: "subspace S \ x \ S \ c *s x \ S" + by (metis subspace_def) + +lemma subspace_neg: "subspace S \ (x::'a::ring_1^'n) \ S \ - x \ S" + by (metis vector_sneg_minus1 subspace_mul) + +lemma subspace_sub: "subspace S \ (x::'a::ring_1^'n) \ S \ y \ S \ x - y \ S" + by (metis diff_def subspace_add subspace_neg) + +lemma subspace_setsum: + assumes sA: "subspace A" and fB: "finite B" + and f: "\x\ B. f x \ A" + shows "setsum f B \ A" + using fB f sA + apply(induct rule: finite_induct[OF fB]) + by (simp add: subspace_def sA, auto simp add: sA subspace_add) + +lemma subspace_linear_image: + assumes lf: "linear (f::'a::semiring_1^'n \ _)" and sS: "subspace S" + shows "subspace(f ` S)" + using lf sS linear_0[OF lf] + unfolding linear_def subspace_def + apply (auto simp add: image_iff) + apply (rule_tac x="x + y" in bexI, auto) + apply (rule_tac x="c*s x" in bexI, auto) + done + +lemma subspace_linear_preimage: "linear (f::'a::semiring_1^'n \ _) ==> subspace S ==> subspace {x. f x \ S}" + by (auto simp add: subspace_def linear_def linear_0[of f]) + +lemma subspace_trivial: "subspace {0::'a::semiring_1 ^_}" + by (simp add: subspace_def) + +lemma subspace_inter: "subspace A \ subspace B ==> subspace (A \ B)" + by (simp add: subspace_def) + + +lemma span_mono: "A \ B ==> span A \ span B" + by (metis span_def hull_mono) + +lemma subspace_span: "subspace(span S)" + unfolding span_def + apply (rule hull_in[unfolded mem_def]) + apply (simp only: subspace_def Inter_iff Int_iff subset_eq) + apply auto + apply (erule_tac x="X" in ballE) + apply (simp add: mem_def) + apply blast + apply (erule_tac x="X" in ballE) + apply (erule_tac x="X" in ballE) + apply (erule_tac x="X" in ballE) + apply (clarsimp simp add: mem_def) + apply simp + apply simp + apply simp + apply (erule_tac x="X" in ballE) + apply (erule_tac x="X" in ballE) + apply (simp add: mem_def) + apply simp + apply simp + done + +lemma span_clauses: + "a \ S ==> a \ span S" + "0 \ span S" + "x\ span S \ y \ span S ==> x + y \ span S" + "x \ span S \ c *s x \ span S" + by (metis span_def hull_subset subset_eq subspace_span subspace_def)+ + +lemma span_induct: assumes SP: "\x. x \ S ==> P x" + and P: "subspace P" and x: "x \ span S" shows "P x" +proof- + from SP have SP': "S \ P" by (simp add: mem_def subset_eq) + from P have P': "P \ subspace" by (simp add: mem_def) + from x hull_minimal[OF SP' P', unfolded span_def[symmetric]] + show "P x" by (metis mem_def subset_eq) +qed + +lemma span_empty: "span {} = {(0::'a::semiring_0 ^ 'n)}" + apply (simp add: span_def) + apply (rule hull_unique) + apply (auto simp add: mem_def subspace_def) + unfolding mem_def[of "0::'a^'n", symmetric] + apply simp + done + +lemma independent_empty: "independent {}" + by (simp add: dependent_def) + +lemma independent_mono: "independent A \ B \ A ==> independent B" + apply (clarsimp simp add: dependent_def span_mono) + apply (subgoal_tac "span (B - {a}) \ span (A - {a})") + apply force + apply (rule span_mono) + apply auto + done + +lemma span_subspace: "A \ B \ B \ span A \ subspace B \ span A = B" + by (metis order_antisym span_def hull_minimal mem_def) + +lemma span_induct': assumes SP: "\x \ S. P x" + and P: "subspace P" shows "\x \ span S. P x" + using span_induct SP P by blast + +inductive span_induct_alt_help for S:: "'a::semiring_1^'n \ bool" + where + span_induct_alt_help_0: "span_induct_alt_help S 0" + | span_induct_alt_help_S: "x \ S \ span_induct_alt_help S z \ span_induct_alt_help S (c *s x + z)" + +lemma span_induct_alt': + assumes h0: "h (0::'a::semiring_1^'n)" and hS: "\c x y. x \ S \ h y \ h (c*s x + y)" shows "\x \ span S. h x" +proof- + {fix x:: "'a^'n" assume x: "span_induct_alt_help S x" + have "h x" + apply (rule span_induct_alt_help.induct[OF x]) + apply (rule h0) + apply (rule hS, assumption, assumption) + done} + note th0 = this + {fix x assume x: "x \ span S" + + have "span_induct_alt_help S x" + proof(rule span_induct[where x=x and S=S]) + show "x \ span S" using x . + next + fix x assume xS : "x \ S" + from span_induct_alt_help_S[OF xS span_induct_alt_help_0, of 1] + show "span_induct_alt_help S x" by simp + next + have "span_induct_alt_help S 0" by (rule span_induct_alt_help_0) + moreover + {fix x y assume h: "span_induct_alt_help S x" "span_induct_alt_help S y" + from h + have "span_induct_alt_help S (x + y)" + apply (induct rule: span_induct_alt_help.induct) + apply simp + unfolding add_assoc + apply (rule span_induct_alt_help_S) + apply assumption + apply simp + done} + moreover + {fix c x assume xt: "span_induct_alt_help S x" + then have "span_induct_alt_help S (c*s x)" + apply (induct rule: span_induct_alt_help.induct) + apply (simp add: span_induct_alt_help_0) + apply (simp add: vector_smult_assoc vector_add_ldistrib) + apply (rule span_induct_alt_help_S) + apply assumption + apply simp + done + } + ultimately show "subspace (span_induct_alt_help S)" + unfolding subspace_def mem_def Ball_def by blast + qed} + with th0 show ?thesis by blast +qed + +lemma span_induct_alt: + assumes h0: "h (0::'a::semiring_1^'n)" and hS: "\c x y. x \ S \ h y \ h (c*s x + y)" and x: "x \ span S" + shows "h x" +using span_induct_alt'[of h S] h0 hS x by blast + +(* Individual closure properties. *) + +lemma span_superset: "x \ S ==> x \ span S" by (metis span_clauses) + +lemma span_0: "0 \ span S" by (metis subspace_span subspace_0) + +lemma span_add: "x \ span S \ y \ span S ==> x + y \ span S" + by (metis subspace_add subspace_span) + +lemma span_mul: "x \ span S ==> (c *s x) \ span S" + by (metis subspace_span subspace_mul) + +lemma span_neg: "x \ span S ==> - (x::'a::ring_1^'n) \ span S" + by (metis subspace_neg subspace_span) + +lemma span_sub: "(x::'a::ring_1^'n) \ span S \ y \ span S ==> x - y \ span S" + by (metis subspace_span subspace_sub) + +lemma span_setsum: "finite A \ \x \ A. f x \ span S ==> setsum f A \ span S" + apply (rule subspace_setsum) + by (metis subspace_span subspace_setsum)+ + +lemma span_add_eq: "(x::'a::ring_1^'n) \ span S \ x + y \ span S \ y \ span S" + apply (auto simp only: span_add span_sub) + apply (subgoal_tac "(x + y) - x \ span S", simp) + by (simp only: span_add span_sub) + +(* Mapping under linear image. *) + +lemma span_linear_image: assumes lf: "linear (f::'a::semiring_1 ^ 'n => _)" + shows "span (f ` S) = f ` (span S)" +proof- + {fix x + assume x: "x \ span (f ` S)" + have "x \ f ` span S" + apply (rule span_induct[where x=x and S = "f ` S"]) + apply (clarsimp simp add: image_iff) + apply (frule span_superset) + apply blast + apply (simp only: mem_def) + apply (rule subspace_linear_image[OF lf]) + apply (rule subspace_span) + apply (rule x) + done} + moreover + {fix x assume x: "x \ span S" + have th0:"(\a. f a \ span (f ` S)) = {x. f x \ span (f ` S)}" apply (rule set_ext) + unfolding mem_def Collect_def .. + have "f x \ span (f ` S)" + apply (rule span_induct[where S=S]) + apply (rule span_superset) + apply simp + apply (subst th0) + apply (rule subspace_linear_preimage[OF lf subspace_span, of "f ` S"]) + apply (rule x) + done} + ultimately show ?thesis by blast +qed + +(* The key breakdown property. *) + +lemma span_breakdown: + assumes bS: "(b::'a::ring_1 ^ 'n) \ S" and aS: "a \ span S" + shows "\k. a - k*s b \ span (S - {b})" (is "?P a") +proof- + {fix x assume xS: "x \ S" + {assume ab: "x = b" + then have "?P x" + apply simp + apply (rule exI[where x="1"], simp) + by (rule span_0)} + moreover + {assume ab: "x \ b" + then have "?P x" using xS + apply - + apply (rule exI[where x=0]) + apply (rule span_superset) + by simp} + ultimately have "?P x" by blast} + moreover have "subspace ?P" + unfolding subspace_def + apply auto + apply (simp add: mem_def) + apply (rule exI[where x=0]) + using span_0[of "S - {b}"] + apply (simp add: mem_def) + apply (clarsimp simp add: mem_def) + apply (rule_tac x="k + ka" in exI) + apply (subgoal_tac "x + y - (k + ka) *s b = (x - k*s b) + (y - ka *s b)") + apply (simp only: ) + apply (rule span_add[unfolded mem_def]) + apply assumption+ + apply (vector ring_simps) + apply (clarsimp simp add: mem_def) + apply (rule_tac x= "c*k" in exI) + apply (subgoal_tac "c *s x - (c * k) *s b = c*s (x - k*s b)") + apply (simp only: ) + apply (rule span_mul[unfolded mem_def]) + apply assumption + by (vector ring_simps) + ultimately show "?P a" using aS span_induct[where S=S and P= "?P"] by metis +qed + +lemma span_breakdown_eq: + "(x::'a::ring_1^'n) \ span (insert a S) \ (\k. (x - k *s a) \ span S)" (is "?lhs \ ?rhs") +proof- + {assume x: "x \ span (insert a S)" + from x span_breakdown[of "a" "insert a S" "x"] + have ?rhs apply clarsimp + apply (rule_tac x= "k" in exI) + apply (rule set_rev_mp[of _ "span (S - {a})" _]) + apply assumption + apply (rule span_mono) + apply blast + done} + moreover + { fix k assume k: "x - k *s a \ span S" + have eq: "x = (x - k *s a) + k *s a" by vector + have "(x - k *s a) + k *s a \ span (insert a S)" + apply (rule span_add) + apply (rule set_rev_mp[of _ "span S" _]) + apply (rule k) + apply (rule span_mono) + apply blast + apply (rule span_mul) + apply (rule span_superset) + apply blast + done + then have ?lhs using eq by metis} + ultimately show ?thesis by blast +qed + +(* Hence some "reversal" results.*) + +lemma in_span_insert: + assumes a: "(a::'a::field^'n) \ span (insert b S)" and na: "a \ span S" + shows "b \ span (insert a S)" +proof- + from span_breakdown[of b "insert b S" a, OF insertI1 a] + obtain k where k: "a - k*s b \ span (S - {b})" by auto + {assume k0: "k = 0" + with k have "a \ span S" + apply (simp) + apply (rule set_rev_mp) + apply assumption + apply (rule span_mono) + apply blast + done + with na have ?thesis by blast} + moreover + {assume k0: "k \ 0" + have eq: "b = (1/k) *s a - ((1/k) *s a - b)" by vector + from k0 have eq': "(1/k) *s (a - k*s b) = (1/k) *s a - b" + by (vector field_simps) + from k have "(1/k) *s (a - k*s b) \ span (S - {b})" + by (rule span_mul) + hence th: "(1/k) *s a - b \ span (S - {b})" + unfolding eq' . + + from k + have ?thesis + apply (subst eq) + apply (rule span_sub) + apply (rule span_mul) + apply (rule span_superset) + apply blast + apply (rule set_rev_mp) + apply (rule th) + apply (rule span_mono) + using na by blast} + ultimately show ?thesis by blast +qed + +lemma in_span_delete: + assumes a: "(a::'a::field^'n) \ span S" + and na: "a \ span (S-{b})" + shows "b \ span (insert a (S - {b}))" + apply (rule in_span_insert) + apply (rule set_rev_mp) + apply (rule a) + apply (rule span_mono) + apply blast + apply (rule na) + done + +(* Transitivity property. *) + +lemma span_trans: + assumes x: "(x::'a::ring_1^'n) \ span S" and y: "y \ span (insert x S)" + shows "y \ span S" +proof- + from span_breakdown[of x "insert x S" y, OF insertI1 y] + obtain k where k: "y -k*s x \ span (S - {x})" by auto + have eq: "y = (y - k *s x) + k *s x" by vector + show ?thesis + apply (subst eq) + apply (rule span_add) + apply (rule set_rev_mp) + apply (rule k) + apply (rule span_mono) + apply blast + apply (rule span_mul) + by (rule x) +qed + +(* ------------------------------------------------------------------------- *) +(* An explicit expansion is sometimes needed. *) +(* ------------------------------------------------------------------------- *) + +lemma span_explicit: + "span P = {y::'a::semiring_1^'n. \S u. finite S \ S \ P \ setsum (\v. u v *s v) S = y}" + (is "_ = ?E" is "_ = {y. ?h y}" is "_ = {y. \S u. ?Q S u y}") +proof- + {fix x assume x: "x \ ?E" + then obtain S u where fS: "finite S" and SP: "S\P" and u: "setsum (\v. u v *s v) S = x" + by blast + have "x \ span P" + unfolding u[symmetric] + apply (rule span_setsum[OF fS]) + using span_mono[OF SP] + by (auto intro: span_superset span_mul)} + moreover + have "\x \ span P. x \ ?E" + unfolding mem_def Collect_def + proof(rule span_induct_alt') + show "?h 0" + apply (rule exI[where x="{}"]) by simp + next + fix c x y + assume x: "x \ P" and hy: "?h y" + from hy obtain S u where fS: "finite S" and SP: "S\P" + and u: "setsum (\v. u v *s v) S = y" by blast + let ?S = "insert x S" + let ?u = "\y. if y = x then (if x \ S then u y + c else c) + else u y" + from fS SP x have th0: "finite (insert x S)" "insert x S \ P" by blast+ + {assume xS: "x \ S" + have S1: "S = (S - {x}) \ {x}" + and Sss:"finite (S - {x})" "finite {x}" "(S -{x}) \ {x} = {}" using xS fS by auto + have "setsum (\v. ?u v *s v) ?S =(\v\S - {x}. u v *s v) + (u x + c) *s x" + using xS + by (simp add: setsum_Un_disjoint[OF Sss, unfolded S1[symmetric]] + setsum_clauses(2)[OF fS] cong del: if_weak_cong) + also have "\ = (\v\S. u v *s v) + c *s x" + apply (simp add: setsum_Un_disjoint[OF Sss, unfolded S1[symmetric]]) + by (vector ring_simps) + also have "\ = c*s x + y" + by (simp add: add_commute u) + finally have "setsum (\v. ?u v *s v) ?S = c*s x + y" . + then have "?Q ?S ?u (c*s x + y)" using th0 by blast} + moreover + {assume xS: "x \ S" + have th00: "(\v\S. (if v = x then c else u v) *s v) = y" + unfolding u[symmetric] + apply (rule setsum_cong2) + using xS by auto + have "?Q ?S ?u (c*s x + y)" using fS xS th0 + by (simp add: th00 setsum_clauses add_commute cong del: if_weak_cong)} + ultimately have "?Q ?S ?u (c*s x + y)" + by (cases "x \ S", simp, simp) + then show "?h (c*s x + y)" + apply - + apply (rule exI[where x="?S"]) + apply (rule exI[where x="?u"]) by metis + qed + ultimately show ?thesis by blast +qed + +lemma dependent_explicit: + "dependent P \ (\S u. finite S \ S \ P \ (\(v::'a::{idom,field}^'n) \S. u v \ 0 \ setsum (\v. u v *s v) S = 0))" (is "?lhs = ?rhs") +proof- + {assume dP: "dependent P" + then obtain a S u where aP: "a \ P" and fS: "finite S" + and SP: "S \ P - {a}" and ua: "setsum (\v. u v *s v) S = a" + unfolding dependent_def span_explicit by blast + let ?S = "insert a S" + let ?u = "\y. if y = a then - 1 else u y" + let ?v = a + from aP SP have aS: "a \ S" by blast + from fS SP aP have th0: "finite ?S" "?S \ P" "?v \ ?S" "?u ?v \ 0" by auto + have s0: "setsum (\v. ?u v *s v) ?S = 0" + using fS aS + apply (simp add: vector_smult_lneg vector_smult_lid setsum_clauses ring_simps ) + apply (subst (2) ua[symmetric]) + apply (rule setsum_cong2) + by auto + with th0 have ?rhs + apply - + apply (rule exI[where x= "?S"]) + apply (rule exI[where x= "?u"]) + by clarsimp} + moreover + {fix S u v assume fS: "finite S" + and SP: "S \ P" and vS: "v \ S" and uv: "u v \ 0" + and u: "setsum (\v. u v *s v) S = 0" + let ?a = v + let ?S = "S - {v}" + let ?u = "\i. (- u i) / u v" + have th0: "?a \ P" "finite ?S" "?S \ P" using fS SP vS by auto + have "setsum (\v. ?u v *s v) ?S = setsum (\v. (- (inverse (u ?a))) *s (u v *s v)) S - ?u v *s v" + using fS vS uv + by (simp add: setsum_diff1 vector_smult_lneg divide_inverse + vector_smult_assoc field_simps) + also have "\ = ?a" + unfolding setsum_cmul u + using uv by (simp add: vector_smult_lneg) + finally have "setsum (\v. ?u v *s v) ?S = ?a" . + with th0 have ?lhs + unfolding dependent_def span_explicit + apply - + apply (rule bexI[where x= "?a"]) + apply simp_all + apply (rule exI[where x= "?S"]) + by auto} + ultimately show ?thesis by blast +qed + + +lemma span_finite: + assumes fS: "finite S" + shows "span S = {(y::'a::semiring_1^'n). \u. setsum (\v. u v *s v) S = y}" + (is "_ = ?rhs") +proof- + {fix y assume y: "y \ span S" + from y obtain S' u where fS': "finite S'" and SS': "S' \ S" and + u: "setsum (\v. u v *s v) S' = y" unfolding span_explicit by blast + let ?u = "\x. if x \ S' then u x else 0" + from setsum_restrict_set[OF fS, of "\v. u v *s v" S', symmetric] SS' + have "setsum (\v. ?u v *s v) S = setsum (\v. u v *s v) S'" + unfolding cond_value_iff cond_application_beta + by (simp add: cond_value_iff inf_absorb2 cong del: if_weak_cong) + hence "setsum (\v. ?u v *s v) S = y" by (metis u) + hence "y \ ?rhs" by auto} + moreover + {fix y u assume u: "setsum (\v. u v *s v) S = y" + then have "y \ span S" using fS unfolding span_explicit by auto} + ultimately show ?thesis by blast +qed + + +(* Standard bases are a spanning set, and obviously finite. *) + +lemma span_stdbasis:"span {basis i :: 'a::ring_1^'n::finite | i. i \ (UNIV :: 'n set)} = UNIV" +apply (rule set_ext) +apply auto +apply (subst basis_expansion[symmetric]) +apply (rule span_setsum) +apply simp +apply auto +apply (rule span_mul) +apply (rule span_superset) +apply (auto simp add: Collect_def mem_def) +done + +lemma has_size_stdbasis: "{basis i ::real ^'n::finite | i. i \ (UNIV :: 'n set)} hassize CARD('n)" (is "?S hassize ?n") +proof- + have eq: "?S = basis ` UNIV" by blast + show ?thesis unfolding eq + apply (rule hassize_image_inj[OF basis_inj]) + by (simp add: hassize_def) +qed + +lemma finite_stdbasis: "finite {basis i ::real^'n::finite |i. i\ (UNIV:: 'n set)}" + using has_size_stdbasis[unfolded hassize_def] + .. + +lemma card_stdbasis: "card {basis i ::real^'n::finite |i. i\ (UNIV :: 'n set)} = CARD('n)" + using has_size_stdbasis[unfolded hassize_def] + .. + +lemma independent_stdbasis_lemma: + assumes x: "(x::'a::semiring_1 ^ 'n) \ span (basis ` S)" + and iS: "i \ S" + shows "(x$i) = 0" +proof- + let ?U = "UNIV :: 'n set" + let ?B = "basis ` S" + let ?P = "\(x::'a^'n). \i\ ?U. i \ S \ x$i =0" + {fix x::"'a^'n" assume xS: "x\ ?B" + from xS have "?P x" by auto} + moreover + have "subspace ?P" + by (auto simp add: subspace_def Collect_def mem_def) + ultimately show ?thesis + using x span_induct[of ?B ?P x] iS by blast +qed + +lemma independent_stdbasis: "independent {basis i ::real^'n::finite |i. i\ (UNIV :: 'n set)}" +proof- + let ?I = "UNIV :: 'n set" + let ?b = "basis :: _ \ real ^'n" + let ?B = "?b ` ?I" + have eq: "{?b i|i. i \ ?I} = ?B" + by auto + {assume d: "dependent ?B" + then obtain k where k: "k \ ?I" "?b k \ span (?B - {?b k})" + unfolding dependent_def by auto + have eq1: "?B - {?b k} = ?B - ?b ` {k}" by simp + have eq2: "?B - {?b k} = ?b ` (?I - {k})" + unfolding eq1 + apply (rule inj_on_image_set_diff[symmetric]) + apply (rule basis_inj) using k(1) by auto + from k(2) have th0: "?b k \ span (?b ` (?I - {k}))" unfolding eq2 . + from independent_stdbasis_lemma[OF th0, of k, simplified] + have False by simp} + then show ?thesis unfolding eq dependent_def .. +qed + +(* This is useful for building a basis step-by-step. *) + +lemma independent_insert: + "independent(insert (a::'a::field ^'n) S) \ + (if a \ S then independent S + else independent S \ a \ span S)" (is "?lhs \ ?rhs") +proof- + {assume aS: "a \ S" + hence ?thesis using insert_absorb[OF aS] by simp} + moreover + {assume aS: "a \ S" + {assume i: ?lhs + then have ?rhs using aS + apply simp + apply (rule conjI) + apply (rule independent_mono) + apply assumption + apply blast + by (simp add: dependent_def)} + moreover + {assume i: ?rhs + have ?lhs using i aS + apply simp + apply (auto simp add: dependent_def) + apply (case_tac "aa = a", auto) + apply (subgoal_tac "insert a S - {aa} = insert a (S - {aa})") + apply simp + apply (subgoal_tac "a \ span (insert aa (S - {aa}))") + apply (subgoal_tac "insert aa (S - {aa}) = S") + apply simp + apply blast + apply (rule in_span_insert) + apply assumption + apply blast + apply blast + done} + ultimately have ?thesis by blast} + ultimately show ?thesis by blast +qed + +(* The degenerate case of the Exchange Lemma. *) + +lemma mem_delete: "x \ (A - {a}) \ x \ a \ x \ A" + by blast + +lemma span_span: "span (span A) = span A" + unfolding span_def hull_hull .. + +lemma span_inc: "S \ span S" + by (metis subset_eq span_superset) + +lemma spanning_subset_independent: + assumes BA: "B \ A" and iA: "independent (A::('a::field ^'n) set)" + and AsB: "A \ span B" + shows "A = B" +proof + from BA show "B \ A" . +next + from span_mono[OF BA] span_mono[OF AsB] + have sAB: "span A = span B" unfolding span_span by blast + + {fix x assume x: "x \ A" + from iA have th0: "x \ span (A - {x})" + unfolding dependent_def using x by blast + from x have xsA: "x \ span A" by (blast intro: span_superset) + have "A - {x} \ A" by blast + hence th1:"span (A - {x}) \ span A" by (metis span_mono) + {assume xB: "x \ B" + from xB BA have "B \ A -{x}" by blast + hence "span B \ span (A - {x})" by (metis span_mono) + with th1 th0 sAB have "x \ span A" by blast + with x have False by (metis span_superset)} + then have "x \ B" by blast} + then show "A \ B" by blast +qed + +(* The general case of the Exchange Lemma, the key to what follows. *) + +lemma exchange_lemma: + assumes f:"finite (t:: ('a::field^'n) set)" and i: "independent s" + and sp:"s \ span t" + shows "\t'. (t' hassize card t) \ s \ t' \ t' \ s \ t \ s \ span t'" +using f i sp +proof(induct c\"card(t - s)" arbitrary: s t rule: nat_less_induct) + fix n:: nat and s t :: "('a ^'n) set" + assume H: " \m(x:: ('a ^'n) set) xa. + finite xa \ + independent x \ + x \ span xa \ + m = card (xa - x) \ + (\t'. (t' hassize card xa) \ + x \ t' \ t' \ x \ xa \ x \ span t')" + and ft: "finite t" and s: "independent s" and sp: "s \ span t" + and n: "n = card (t - s)" + let ?P = "\t'. (t' hassize card t) \ s \ t' \ t' \ s \ t \ s \ span t'" + let ?ths = "\t'. ?P t'" + {assume st: "s \ t" + from st ft span_mono[OF st] have ?ths apply - apply (rule exI[where x=t]) + by (auto simp add: hassize_def intro: span_superset)} + moreover + {assume st: "t \ s" + + from spanning_subset_independent[OF st s sp] + st ft span_mono[OF st] have ?ths apply - apply (rule exI[where x=t]) + by (auto simp add: hassize_def intro: span_superset)} + moreover + {assume st: "\ s \ t" "\ t \ s" + from st(2) obtain b where b: "b \ t" "b \ s" by blast + from b have "t - {b} - s \ t - s" by blast + then have cardlt: "card (t - {b} - s) < n" using n ft + by (auto intro: psubset_card_mono) + from b ft have ct0: "card t \ 0" by auto + {assume stb: "s \ span(t -{b})" + from ft have ftb: "finite (t -{b})" by auto + from H[rule_format, OF cardlt ftb s stb] + obtain u where u: "u hassize card (t-{b})" "s \ u" "u \ s \ (t - {b})" "s \ span u" by blast + let ?w = "insert b u" + have th0: "s \ insert b u" using u by blast + from u(3) b have "u \ s \ t" by blast + then have th1: "insert b u \ s \ t" using u b by blast + have bu: "b \ u" using b u by blast + from u(1) have fu: "finite u" by (simp add: hassize_def) + from u(1) ft b have "u hassize (card t - 1)" by auto + then + have th2: "insert b u hassize card t" + using card_insert_disjoint[OF fu bu] ct0 by (auto simp add: hassize_def) + from u(4) have "s \ span u" . + also have "\ \ span (insert b u)" apply (rule span_mono) by blast + finally have th3: "s \ span (insert b u)" . from th0 th1 th2 th3 have th: "?P ?w" by blast + from th have ?ths by blast} + moreover + {assume stb: "\ s \ span(t -{b})" + from stb obtain a where a: "a \ s" "a \ span (t - {b})" by blast + have ab: "a \ b" using a b by blast + have at: "a \ t" using a ab span_superset[of a "t- {b}"] by auto + have mlt: "card ((insert a (t - {b})) - s) < n" + using cardlt ft n a b by auto + have ft': "finite (insert a (t - {b}))" using ft by auto + {fix x assume xs: "x \ s" + have t: "t \ (insert b (insert a (t -{b})))" using b by auto + from b(1) have "b \ span t" by (simp add: span_superset) + have bs: "b \ span (insert a (t - {b}))" + by (metis in_span_delete a sp mem_def subset_eq) + from xs sp have "x \ span t" by blast + with span_mono[OF t] + have x: "x \ span (insert b (insert a (t - {b})))" .. + from span_trans[OF bs x] have "x \ span (insert a (t - {b}))" .} + then have sp': "s \ span (insert a (t - {b}))" by blast + + from H[rule_format, OF mlt ft' s sp' refl] obtain u where + u: "u hassize card (insert a (t -{b}))" "s \ u" "u \ s \ insert a (t -{b})" + "s \ span u" by blast + from u a b ft at ct0 have "?P u" by (auto simp add: hassize_def) + then have ?ths by blast } + ultimately have ?ths by blast + } + ultimately + show ?ths by blast +qed + +(* This implies corresponding size bounds. *) + +lemma independent_span_bound: + assumes f: "finite t" and i: "independent (s::('a::field^'n) set)" and sp:"s \ span t" + shows "finite s \ card s \ card t" + by (metis exchange_lemma[OF f i sp] hassize_def finite_subset card_mono) + + +lemma finite_Atleast_Atmost_nat[simp]: "finite {f x |x. x\ (UNIV::'a::finite set)}" +proof- + have eq: "{f x |x. x\ UNIV} = f ` UNIV" by auto + show ?thesis unfolding eq + apply (rule finite_imageI) + apply (rule finite) + done +qed + + +lemma independent_bound: + fixes S:: "(real^'n::finite) set" + shows "independent S \ finite S \ card S <= CARD('n)" + apply (subst card_stdbasis[symmetric]) + apply (rule independent_span_bound) + apply (rule finite_Atleast_Atmost_nat) + apply assumption + unfolding span_stdbasis + apply (rule subset_UNIV) + done + +lemma dependent_biggerset: "(finite (S::(real ^'n::finite) set) ==> card S > CARD('n)) ==> dependent S" + by (metis independent_bound not_less) + +(* Hence we can create a maximal independent subset. *) + +lemma maximal_independent_subset_extend: + assumes sv: "(S::(real^'n::finite) set) \ V" and iS: "independent S" + shows "\B. S \ B \ B \ V \ independent B \ V \ span B" + using sv iS +proof(induct d\ "CARD('n) - card S" arbitrary: S rule: nat_less_induct) + fix n and S:: "(real^'n) set" + assume H: "\mS \ V. independent S \ m = CARD('n) - card S \ + (\B. S \ B \ B \ V \ independent B \ V \ span B)" + and sv: "S \ V" and i: "independent S" and n: "n = CARD('n) - card S" + let ?P = "\B. S \ B \ B \ V \ independent B \ V \ span B" + let ?ths = "\x. ?P x" + let ?d = "CARD('n)" + {assume "V \ span S" + then have ?ths using sv i by blast } + moreover + {assume VS: "\ V \ span S" + from VS obtain a where a: "a \ V" "a \ span S" by blast + from a have aS: "a \ S" by (auto simp add: span_superset) + have th0: "insert a S \ V" using a sv by blast + from independent_insert[of a S] i a + have th1: "independent (insert a S)" by auto + have mlt: "?d - card (insert a S) < n" + using aS a n independent_bound[OF th1] + by auto + + from H[rule_format, OF mlt th0 th1 refl] + obtain B where B: "insert a S \ B" "B \ V" "independent B" " V \ span B" + by blast + from B have "?P B" by auto + then have ?ths by blast} + ultimately show ?ths by blast +qed + +lemma maximal_independent_subset: + "\(B:: (real ^'n::finite) set). B\ V \ independent B \ V \ span B" + by (metis maximal_independent_subset_extend[of "{}:: (real ^'n) set"] empty_subsetI independent_empty) + +(* Notion of dimension. *) + +definition "dim V = (SOME n. \B. B \ V \ independent B \ V \ span B \ (B hassize n))" + +lemma basis_exists: "\B. (B :: (real ^'n::finite) set) \ V \ independent B \ V \ span B \ (B hassize dim V)" +unfolding dim_def some_eq_ex[of "\n. \B. B \ V \ independent B \ V \ span B \ (B hassize n)"] +unfolding hassize_def +using maximal_independent_subset[of V] independent_bound +by auto + +(* Consequences of independence or spanning for cardinality. *) + +lemma independent_card_le_dim: "(B::(real ^'n::finite) set) \ V \ independent B \ finite B \ card B \ dim V" +by (metis basis_exists[of V] independent_span_bound[where ?'a=real] hassize_def subset_trans) + +lemma span_card_ge_dim: "(B::(real ^'n::finite) set) \ V \ V \ span B \ finite B \ dim V \ card B" + by (metis basis_exists[of V] independent_span_bound hassize_def subset_trans) + +lemma basis_card_eq_dim: + "B \ (V:: (real ^'n::finite) set) \ V \ span B \ independent B \ finite B \ card B = dim V" + by (metis order_eq_iff independent_card_le_dim span_card_ge_dim independent_mono) + +lemma dim_unique: "(B::(real ^'n::finite) set) \ V \ V \ span B \ independent B \ B hassize n \ dim V = n" + by (metis basis_card_eq_dim hassize_def) + +(* More lemmas about dimension. *) + +lemma dim_univ: "dim (UNIV :: (real^'n::finite) set) = CARD('n)" + apply (rule dim_unique[of "{basis i |i. i\ (UNIV :: 'n set)}"]) + by (auto simp only: span_stdbasis has_size_stdbasis independent_stdbasis) + +lemma dim_subset: + "(S:: (real ^'n::finite) set) \ T \ dim S \ dim T" + using basis_exists[of T] basis_exists[of S] + by (metis independent_span_bound[where ?'a = real and ?'n = 'n] subset_eq hassize_def) + +lemma dim_subset_univ: "dim (S:: (real^'n::finite) set) \ CARD('n)" + by (metis dim_subset subset_UNIV dim_univ) + +(* Converses to those. *) + +lemma card_ge_dim_independent: + assumes BV:"(B::(real ^'n::finite) set) \ V" and iB:"independent B" and dVB:"dim V \ card B" + shows "V \ span B" +proof- + {fix a assume aV: "a \ V" + {assume aB: "a \ span B" + then have iaB: "independent (insert a B)" using iB aV BV by (simp add: independent_insert) + from aV BV have th0: "insert a B \ V" by blast + from aB have "a \B" by (auto simp add: span_superset) + with independent_card_le_dim[OF th0 iaB] dVB have False by auto} + then have "a \ span B" by blast} + then show ?thesis by blast +qed + +lemma card_le_dim_spanning: + assumes BV: "(B:: (real ^'n::finite) set) \ V" and VB: "V \ span B" + and fB: "finite B" and dVB: "dim V \ card B" + shows "independent B" +proof- + {fix a assume a: "a \ B" "a \ span (B -{a})" + from a fB have c0: "card B \ 0" by auto + from a fB have cb: "card (B -{a}) = card B - 1" by auto + from BV a have th0: "B -{a} \ V" by blast + {fix x assume x: "x \ V" + from a have eq: "insert a (B -{a}) = B" by blast + from x VB have x': "x \ span B" by blast + from span_trans[OF a(2), unfolded eq, OF x'] + have "x \ span (B -{a})" . } + then have th1: "V \ span (B -{a})" by blast + have th2: "finite (B -{a})" using fB by auto + from span_card_ge_dim[OF th0 th1 th2] + have c: "dim V \ card (B -{a})" . + from c c0 dVB cb have False by simp} + then show ?thesis unfolding dependent_def by blast +qed + +lemma card_eq_dim: "(B:: (real ^'n::finite) set) \ V \ B hassize dim V \ independent B \ V \ span B" + by (metis hassize_def order_eq_iff card_le_dim_spanning + card_ge_dim_independent) + +(* ------------------------------------------------------------------------- *) +(* More general size bound lemmas. *) +(* ------------------------------------------------------------------------- *) + +lemma independent_bound_general: + "independent (S:: (real^'n::finite) set) \ finite S \ card S \ dim S" + by (metis independent_card_le_dim independent_bound subset_refl) + +lemma dependent_biggerset_general: "(finite (S:: (real^'n::finite) set) \ card S > dim S) \ dependent S" + using independent_bound_general[of S] by (metis linorder_not_le) + +lemma dim_span: "dim (span (S:: (real ^'n::finite) set)) = dim S" +proof- + have th0: "dim S \ dim (span S)" + by (auto simp add: subset_eq intro: dim_subset span_superset) + from basis_exists[of S] + obtain B where B: "B \ S" "independent B" "S \ span B" "B hassize dim S" by blast + from B have fB: "finite B" "card B = dim S" unfolding hassize_def by blast+ + have bSS: "B \ span S" using B(1) by (metis subset_eq span_inc) + have sssB: "span S \ span B" using span_mono[OF B(3)] by (simp add: span_span) + from span_card_ge_dim[OF bSS sssB fB(1)] th0 show ?thesis + using fB(2) by arith +qed + +lemma subset_le_dim: "(S:: (real ^'n::finite) set) \ span T \ dim S \ dim T" + by (metis dim_span dim_subset) + +lemma span_eq_dim: "span (S:: (real ^'n::finite) set) = span T ==> dim S = dim T" + by (metis dim_span) + +lemma spans_image: + assumes lf: "linear (f::'a::semiring_1^'n \ _)" and VB: "V \ span B" + shows "f ` V \ span (f ` B)" + unfolding span_linear_image[OF lf] + by (metis VB image_mono) + +lemma dim_image_le: + fixes f :: "real^'n::finite \ real^'m::finite" + assumes lf: "linear f" shows "dim (f ` S) \ dim (S:: (real ^'n::finite) set)" +proof- + from basis_exists[of S] obtain B where + B: "B \ S" "independent B" "S \ span B" "B hassize dim S" by blast + from B have fB: "finite B" "card B = dim S" unfolding hassize_def by blast+ + have "dim (f ` S) \ card (f ` B)" + apply (rule span_card_ge_dim) + using lf B fB by (auto simp add: span_linear_image spans_image subset_image_iff) + also have "\ \ dim S" using card_image_le[OF fB(1)] fB by simp + finally show ?thesis . +qed + +(* Relation between bases and injectivity/surjectivity of map. *) + +lemma spanning_surjective_image: + assumes us: "UNIV \ span (S:: ('a::semiring_1 ^'n) set)" + and lf: "linear f" and sf: "surj f" + shows "UNIV \ span (f ` S)" +proof- + have "UNIV \ f ` UNIV" using sf by (auto simp add: surj_def) + also have " \ \ span (f ` S)" using spans_image[OF lf us] . +finally show ?thesis . +qed + +lemma independent_injective_image: + assumes iS: "independent (S::('a::semiring_1^'n) set)" and lf: "linear f" and fi: "inj f" + shows "independent (f ` S)" +proof- + {fix a assume a: "a \ S" "f a \ span (f ` S - {f a})" + have eq: "f ` S - {f a} = f ` (S - {a})" using fi + by (auto simp add: inj_on_def) + from a have "f a \ f ` span (S -{a})" + unfolding eq span_linear_image[OF lf, of "S - {a}"] by blast + hence "a \ span (S -{a})" using fi by (auto simp add: inj_on_def) + with a(1) iS have False by (simp add: dependent_def) } + then show ?thesis unfolding dependent_def by blast +qed + +(* ------------------------------------------------------------------------- *) +(* Picking an orthogonal replacement for a spanning set. *) +(* ------------------------------------------------------------------------- *) + (* FIXME : Move to some general theory ?*) +definition "pairwise R S \ (\x \ S. \y\ S. x\y \ R x y)" + +lemma vector_sub_project_orthogonal: "(b::'a::ordered_field^'n::finite) \ (x - ((b \ x) / (b\b)) *s b) = 0" + apply (cases "b = 0", simp) + apply (simp add: dot_rsub dot_rmult) + unfolding times_divide_eq_right[symmetric] + by (simp add: field_simps dot_eq_0) + +lemma basis_orthogonal: + fixes B :: "(real ^'n::finite) set" + assumes fB: "finite B" + shows "\C. finite C \ card C \ card B \ span C = span B \ pairwise orthogonal C" + (is " \C. ?P B C") +proof(induct rule: finite_induct[OF fB]) + case 1 thus ?case apply (rule exI[where x="{}"]) by (auto simp add: pairwise_def) +next + case (2 a B) + note fB = `finite B` and aB = `a \ B` + from `\C. finite C \ card C \ card B \ span C = span B \ pairwise orthogonal C` + obtain C where C: "finite C" "card C \ card B" + "span C = span B" "pairwise orthogonal C" by blast + let ?a = "a - setsum (\x. (x\a / (x\x)) *s x) C" + let ?C = "insert ?a C" + from C(1) have fC: "finite ?C" by simp + from fB aB C(1,2) have cC: "card ?C \ card (insert a B)" by (simp add: card_insert_if) + {fix x k + have th0: "\(a::'b::comm_ring) b c. a - (b - c) = c + (a - b)" by (simp add: ring_simps) + have "x - k *s (a - (\x\C. (x \ a / (x \ x)) *s x)) \ span C \ x - k *s a \ span C" + apply (simp only: vector_ssub_ldistrib th0) + apply (rule span_add_eq) + apply (rule span_mul) + apply (rule span_setsum[OF C(1)]) + apply clarify + apply (rule span_mul) + by (rule span_superset)} + then have SC: "span ?C = span (insert a B)" + unfolding expand_set_eq span_breakdown_eq C(3)[symmetric] by auto + thm pairwise_def + {fix x y assume xC: "x \ ?C" and yC: "y \ ?C" and xy: "x \ y" + {assume xa: "x = ?a" and ya: "y = ?a" + have "orthogonal x y" using xa ya xy by blast} + moreover + {assume xa: "x = ?a" and ya: "y \ ?a" "y \ C" + from ya have Cy: "C = insert y (C - {y})" by blast + have fth: "finite (C - {y})" using C by simp + have "orthogonal x y" + using xa ya + unfolding orthogonal_def xa dot_lsub dot_rsub diff_eq_0_iff_eq + apply simp + apply (subst Cy) + using C(1) fth + apply (simp only: setsum_clauses) + thm dot_ladd + apply (auto simp add: dot_ladd dot_radd dot_lmult dot_rmult dot_eq_0 dot_sym[of y a] dot_lsum[OF fth]) + apply (rule setsum_0') + apply clarsimp + apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format]) + by auto} + moreover + {assume xa: "x \ ?a" "x \ C" and ya: "y = ?a" + from xa have Cx: "C = insert x (C - {x})" by blast + have fth: "finite (C - {x})" using C by simp + have "orthogonal x y" + using xa ya + unfolding orthogonal_def ya dot_rsub dot_lsub diff_eq_0_iff_eq + apply simp + apply (subst Cx) + using C(1) fth + apply (simp only: setsum_clauses) + apply (subst dot_sym[of x]) + apply (auto simp add: dot_radd dot_rmult dot_eq_0 dot_sym[of x a] dot_rsum[OF fth]) + apply (rule setsum_0') + apply clarsimp + apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format]) + by auto} + moreover + {assume xa: "x \ C" and ya: "y \ C" + have "orthogonal x y" using xa ya xy C(4) unfolding pairwise_def by blast} + ultimately have "orthogonal x y" using xC yC by blast} + then have CPO: "pairwise orthogonal ?C" unfolding pairwise_def by blast + from fC cC SC CPO have "?P (insert a B) ?C" by blast + then show ?case by blast +qed + +lemma orthogonal_basis_exists: + fixes V :: "(real ^'n::finite) set" + shows "\B. independent B \ B \ span V \ V \ span B \ (B hassize dim V) \ pairwise orthogonal B" +proof- + from basis_exists[of V] obtain B where B: "B \ V" "independent B" "V \ span B" "B hassize dim V" by blast + from B have fB: "finite B" "card B = dim V" by (simp_all add: hassize_def) + from basis_orthogonal[OF fB(1)] obtain C where + C: "finite C" "card C \ card B" "span C = span B" "pairwise orthogonal C" by blast + from C B + have CSV: "C \ span V" by (metis span_inc span_mono subset_trans) + from span_mono[OF B(3)] C have SVC: "span V \ span C" by (simp add: span_span) + from card_le_dim_spanning[OF CSV SVC C(1)] C(2,3) fB + have iC: "independent C" by (simp add: dim_span) + from C fB have "card C \ dim V" by simp + moreover have "dim V \ card C" using span_card_ge_dim[OF CSV SVC C(1)] + by (simp add: dim_span) + ultimately have CdV: "C hassize dim V" unfolding hassize_def using C(1) by simp + from C B CSV CdV iC show ?thesis by auto +qed + +lemma span_eq: "span S = span T \ S \ span T \ T \ span S" + by (metis set_eq_subset span_mono span_span span_inc) (* FIXME: slow *) + +(* ------------------------------------------------------------------------- *) +(* Low-dimensional subset is in a hyperplane (weak orthogonal complement). *) +(* ------------------------------------------------------------------------- *) + +lemma span_not_univ_orthogonal: + assumes sU: "span S \ UNIV" + shows "\(a:: real ^'n::finite). a \0 \ (\x \ span S. a \ x = 0)" +proof- + from sU obtain a where a: "a \ span S" by blast + from orthogonal_basis_exists obtain B where + B: "independent B" "B \ span S" "S \ span B" "B hassize dim S" "pairwise orthogonal B" + by blast + from B have fB: "finite B" "card B = dim S" by (simp_all add: hassize_def) + from span_mono[OF B(2)] span_mono[OF B(3)] + have sSB: "span S = span B" by (simp add: span_span) + let ?a = "a - setsum (\b. (a\b / (b\b)) *s b) B" + have "setsum (\b. (a\b / (b\b)) *s b) B \ span S" + unfolding sSB + apply (rule span_setsum[OF fB(1)]) + apply clarsimp + apply (rule span_mul) + by (rule span_superset) + with a have a0:"?a \ 0" by auto + have "\x\span B. ?a \ x = 0" + proof(rule span_induct') + show "subspace (\x. ?a \ x = 0)" + by (auto simp add: subspace_def mem_def dot_radd dot_rmult) + next + {fix x assume x: "x \ B" + from x have B': "B = insert x (B - {x})" by blast + have fth: "finite (B - {x})" using fB by simp + have "?a \ x = 0" + apply (subst B') using fB fth + unfolding setsum_clauses(2)[OF fth] + apply simp + apply (clarsimp simp add: dot_lsub dot_ladd dot_lmult dot_lsum dot_eq_0) + apply (rule setsum_0', rule ballI) + unfolding dot_sym + by (auto simp add: x field_simps dot_eq_0 intro: B(5)[unfolded pairwise_def orthogonal_def, rule_format])} + then show "\x \ B. ?a \ x = 0" by blast + qed + with a0 show ?thesis unfolding sSB by (auto intro: exI[where x="?a"]) +qed + +lemma span_not_univ_subset_hyperplane: + assumes SU: "span S \ (UNIV ::(real^'n::finite) set)" + shows "\ a. a \0 \ span S \ {x. a \ x = 0}" + using span_not_univ_orthogonal[OF SU] by auto + +lemma lowdim_subset_hyperplane: + assumes d: "dim S < CARD('n::finite)" + shows "\(a::real ^'n::finite). a \ 0 \ span S \ {x. a \ x = 0}" +proof- + {assume "span S = UNIV" + hence "dim (span S) = dim (UNIV :: (real ^'n) set)" by simp + hence "dim S = CARD('n)" by (simp add: dim_span dim_univ) + with d have False by arith} + hence th: "span S \ UNIV" by blast + from span_not_univ_subset_hyperplane[OF th] show ?thesis . +qed + +(* We can extend a linear basis-basis injection to the whole set. *) + +lemma linear_indep_image_lemma: + assumes lf: "linear f" and fB: "finite B" + and ifB: "independent (f ` B)" + and fi: "inj_on f B" and xsB: "x \ span B" + and fx: "f (x::'a::field^'n) = 0" + shows "x = 0" + using fB ifB fi xsB fx +proof(induct arbitrary: x rule: finite_induct[OF fB]) + case 1 thus ?case by (auto simp add: span_empty) +next + case (2 a b x) + have fb: "finite b" using "2.prems" by simp + have th0: "f ` b \ f ` (insert a b)" + apply (rule image_mono) by blast + from independent_mono[ OF "2.prems"(2) th0] + have ifb: "independent (f ` b)" . + have fib: "inj_on f b" + apply (rule subset_inj_on [OF "2.prems"(3)]) + by blast + from span_breakdown[of a "insert a b", simplified, OF "2.prems"(4)] + obtain k where k: "x - k*s a \ span (b -{a})" by blast + have "f (x - k*s a) \ span (f ` b)" + unfolding span_linear_image[OF lf] + apply (rule imageI) + using k span_mono[of "b-{a}" b] by blast + hence "f x - k*s f a \ span (f ` b)" + by (simp add: linear_sub[OF lf] linear_cmul[OF lf]) + hence th: "-k *s f a \ span (f ` b)" + using "2.prems"(5) by (simp add: vector_smult_lneg) + {assume k0: "k = 0" + from k0 k have "x \ span (b -{a})" by simp + then have "x \ span b" using span_mono[of "b-{a}" b] + by blast} + moreover + {assume k0: "k \ 0" + from span_mul[OF th, of "- 1/ k"] k0 + have th1: "f a \ span (f ` b)" + by (auto simp add: vector_smult_assoc) + from inj_on_image_set_diff[OF "2.prems"(3), of "insert a b " "{a}", symmetric] + have tha: "f ` insert a b - f ` {a} = f ` (insert a b - {a})" by blast + from "2.prems"(2)[unfolded dependent_def bex_simps(10), rule_format, of "f a"] + have "f a \ span (f ` b)" using tha + using "2.hyps"(2) + "2.prems"(3) by auto + with th1 have False by blast + then have "x \ span b" by blast} + ultimately have xsb: "x \ span b" by blast + from "2.hyps"(3)[OF fb ifb fib xsb "2.prems"(5)] + show "x = 0" . +qed + +(* We can extend a linear mapping from basis. *) + +lemma linear_independent_extend_lemma: + assumes fi: "finite B" and ib: "independent B" + shows "\g. (\x\ span B. \y\ span B. g ((x::'a::field^'n) + y) = g x + g y) + \ (\x\ span B. \c. g (c*s x) = c *s g x) + \ (\x\ B. g x = f x)" +using ib fi +proof(induct rule: finite_induct[OF fi]) + case 1 thus ?case by (auto simp add: span_empty) +next + case (2 a b) + from "2.prems" "2.hyps" have ibf: "independent b" "finite b" + by (simp_all add: independent_insert) + from "2.hyps"(3)[OF ibf] obtain g where + g: "\x\span b. \y\span b. g (x + y) = g x + g y" + "\x\span b. \c. g (c *s x) = c *s g x" "\x\b. g x = f x" by blast + let ?h = "\z. SOME k. (z - k *s a) \ span b" + {fix z assume z: "z \ span (insert a b)" + have th0: "z - ?h z *s a \ span b" + apply (rule someI_ex) + unfolding span_breakdown_eq[symmetric] + using z . + {fix k assume k: "z - k *s a \ span b" + have eq: "z - ?h z *s a - (z - k*s a) = (k - ?h z) *s a" + by (simp add: ring_simps vector_sadd_rdistrib[symmetric]) + from span_sub[OF th0 k] + have khz: "(k - ?h z) *s a \ span b" by (simp add: eq) + {assume "k \ ?h z" hence k0: "k - ?h z \ 0" by simp + from k0 span_mul[OF khz, of "1 /(k - ?h z)"] + have "a \ span b" by (simp add: vector_smult_assoc) + with "2.prems"(1) "2.hyps"(2) have False + by (auto simp add: dependent_def)} + then have "k = ?h z" by blast} + with th0 have "z - ?h z *s a \ span b \ (\k. z - k *s a \ span b \ k = ?h z)" by blast} + note h = this + let ?g = "\z. ?h z *s f a + g (z - ?h z *s a)" + {fix x y assume x: "x \ span (insert a b)" and y: "y \ span (insert a b)" + have tha: "\(x::'a^'n) y a k l. (x + y) - (k + l) *s a = (x - k *s a) + (y - l *s a)" + by (vector ring_simps) + have addh: "?h (x + y) = ?h x + ?h y" + apply (rule conjunct2[OF h, rule_format, symmetric]) + apply (rule span_add[OF x y]) + unfolding tha + by (metis span_add x y conjunct1[OF h, rule_format]) + have "?g (x + y) = ?g x + ?g y" + unfolding addh tha + g(1)[rule_format,OF conjunct1[OF h, OF x] conjunct1[OF h, OF y]] + by (simp add: vector_sadd_rdistrib)} + moreover + {fix x:: "'a^'n" and c:: 'a assume x: "x \ span (insert a b)" + have tha: "\(x::'a^'n) c k a. c *s x - (c * k) *s a = c *s (x - k *s a)" + by (vector ring_simps) + have hc: "?h (c *s x) = c * ?h x" + apply (rule conjunct2[OF h, rule_format, symmetric]) + apply (metis span_mul x) + by (metis tha span_mul x conjunct1[OF h]) + have "?g (c *s x) = c*s ?g x" + unfolding hc tha g(2)[rule_format, OF conjunct1[OF h, OF x]] + by (vector ring_simps)} + moreover + {fix x assume x: "x \ (insert a b)" + {assume xa: "x = a" + have ha1: "1 = ?h a" + apply (rule conjunct2[OF h, rule_format]) + apply (metis span_superset insertI1) + using conjunct1[OF h, OF span_superset, OF insertI1] + by (auto simp add: span_0) + + from xa ha1[symmetric] have "?g x = f x" + apply simp + using g(2)[rule_format, OF span_0, of 0] + by simp} + moreover + {assume xb: "x \ b" + have h0: "0 = ?h x" + apply (rule conjunct2[OF h, rule_format]) + apply (metis span_superset insertI1 xb x) + apply simp + apply (metis span_superset xb) + done + have "?g x = f x" + by (simp add: h0[symmetric] g(3)[rule_format, OF xb])} + ultimately have "?g x = f x" using x by blast } + ultimately show ?case apply - apply (rule exI[where x="?g"]) by blast +qed + +lemma linear_independent_extend: + assumes iB: "independent (B:: (real ^'n::finite) set)" + shows "\g. linear g \ (\x\B. g x = f x)" +proof- + from maximal_independent_subset_extend[of B UNIV] iB + obtain C where C: "B \ C" "independent C" "\x. x \ span C" by auto + + from C(2) independent_bound[of C] linear_independent_extend_lemma[of C f] + obtain g where g: "(\x\ span C. \y\ span C. g (x + y) = g x + g y) + \ (\x\ span C. \c. g (c*s x) = c *s g x) + \ (\x\ C. g x = f x)" by blast + from g show ?thesis unfolding linear_def using C + apply clarsimp by blast +qed + +(* Can construct an isomorphism between spaces of same dimension. *) + +lemma card_le_inj: assumes fA: "finite A" and fB: "finite B" + and c: "card A \ card B" shows "(\f. f ` A \ B \ inj_on f A)" +using fB c +proof(induct arbitrary: B rule: finite_induct[OF fA]) + case 1 thus ?case by simp +next + case (2 x s t) + thus ?case + proof(induct rule: finite_induct[OF "2.prems"(1)]) + case 1 then show ?case by simp + next + case (2 y t) + from "2.prems"(1,2,5) "2.hyps"(1,2) have cst:"card s \ card t" by simp + from "2.prems"(3) [OF "2.hyps"(1) cst] obtain f where + f: "f ` s \ t \ inj_on f s" by blast + from f "2.prems"(2) "2.hyps"(2) show ?case + apply - + apply (rule exI[where x = "\z. if z = x then y else f z"]) + by (auto simp add: inj_on_def) + qed +qed + +lemma card_subset_eq: assumes fB: "finite B" and AB: "A \ B" and + c: "card A = card B" + shows "A = B" +proof- + from fB AB have fA: "finite A" by (auto intro: finite_subset) + from fA fB have fBA: "finite (B - A)" by auto + have e: "A \ (B - A) = {}" by blast + have eq: "A \ (B - A) = B" using AB by blast + from card_Un_disjoint[OF fA fBA e, unfolded eq c] + have "card (B - A) = 0" by arith + hence "B - A = {}" unfolding card_eq_0_iff using fA fB by simp + with AB show "A = B" by blast +qed + +lemma subspace_isomorphism: + assumes s: "subspace (S:: (real ^'n::finite) set)" + and t: "subspace (T :: (real ^ 'm::finite) set)" + and d: "dim S = dim T" + shows "\f. linear f \ f ` S = T \ inj_on f S" +proof- + from basis_exists[of S] obtain B where + B: "B \ S" "independent B" "S \ span B" "B hassize dim S" by blast + from basis_exists[of T] obtain C where + C: "C \ T" "independent C" "T \ span C" "C hassize dim T" by blast + from B(4) C(4) card_le_inj[of B C] d obtain f where + f: "f ` B \ C" "inj_on f B" unfolding hassize_def by auto + from linear_independent_extend[OF B(2)] obtain g where + g: "linear g" "\x\ B. g x = f x" by blast + from B(4) have fB: "finite B" by (simp add: hassize_def) + from C(4) have fC: "finite C" by (simp add: hassize_def) + from inj_on_iff_eq_card[OF fB, of f] f(2) + have "card (f ` B) = card B" by simp + with B(4) C(4) have ceq: "card (f ` B) = card C" using d + by (simp add: hassize_def) + have "g ` B = f ` B" using g(2) + by (auto simp add: image_iff) + also have "\ = C" using card_subset_eq[OF fC f(1) ceq] . + finally have gBC: "g ` B = C" . + have gi: "inj_on g B" using f(2) g(2) + by (auto simp add: inj_on_def) + note g0 = linear_indep_image_lemma[OF g(1) fB, unfolded gBC, OF C(2) gi] + {fix x y assume x: "x \ S" and y: "y \ S" and gxy:"g x = g y" + from B(3) x y have x': "x \ span B" and y': "y \ span B" by blast+ + from gxy have th0: "g (x - y) = 0" by (simp add: linear_sub[OF g(1)]) + have th1: "x - y \ span B" using x' y' by (metis span_sub) + have "x=y" using g0[OF th1 th0] by simp } + then have giS: "inj_on g S" + unfolding inj_on_def by blast + from span_subspace[OF B(1,3) s] + have "g ` S = span (g ` B)" by (simp add: span_linear_image[OF g(1)]) + also have "\ = span C" unfolding gBC .. + also have "\ = T" using span_subspace[OF C(1,3) t] . + finally have gS: "g ` S = T" . + from g(1) gS giS show ?thesis by blast +qed + +(* linear functions are equal on a subspace if they are on a spanning set. *) + +lemma subspace_kernel: + assumes lf: "linear (f::'a::semiring_1 ^'n \ _)" + shows "subspace {x. f x = 0}" +apply (simp add: subspace_def) +by (simp add: linear_add[OF lf] linear_cmul[OF lf] linear_0[OF lf]) + +lemma linear_eq_0_span: + assumes lf: "linear f" and f0: "\x\B. f x = 0" + shows "\x \ span B. f x = (0::'a::semiring_1 ^'n)" +proof + fix x assume x: "x \ span B" + let ?P = "\x. f x = 0" + from subspace_kernel[OF lf] have "subspace ?P" unfolding Collect_def . + with x f0 span_induct[of B "?P" x] show "f x = 0" by blast +qed + +lemma linear_eq_0: + assumes lf: "linear f" and SB: "S \ span B" and f0: "\x\B. f x = 0" + shows "\x \ S. f x = (0::'a::semiring_1^'n)" + by (metis linear_eq_0_span[OF lf] subset_eq SB f0) + +lemma linear_eq: + assumes lf: "linear (f::'a::ring_1^'n \ _)" and lg: "linear g" and S: "S \ span B" + and fg: "\ x\ B. f x = g x" + shows "\x\ S. f x = g x" +proof- + let ?h = "\x. f x - g x" + from fg have fg': "\x\ B. ?h x = 0" by simp + from linear_eq_0[OF linear_compose_sub[OF lf lg] S fg'] + show ?thesis by simp +qed + +lemma linear_eq_stdbasis: + assumes lf: "linear (f::'a::ring_1^'m::finite \ 'a^'n::finite)" and lg: "linear g" + and fg: "\i. f (basis i) = g(basis i)" + shows "f = g" +proof- + let ?U = "UNIV :: 'm set" + let ?I = "{basis i:: 'a^'m|i. i \ ?U}" + {fix x assume x: "x \ (UNIV :: ('a^'m) set)" + from equalityD2[OF span_stdbasis] + have IU: " (UNIV :: ('a^'m) set) \ span ?I" by blast + from linear_eq[OF lf lg IU] fg x + have "f x = g x" unfolding Collect_def Ball_def mem_def by metis} + then show ?thesis by (auto intro: ext) +qed + +(* Similar results for bilinear functions. *) + +lemma bilinear_eq: + assumes bf: "bilinear (f:: 'a::ring^'m \ 'a^'n \ 'a^'p)" + and bg: "bilinear g" + and SB: "S \ span B" and TC: "T \ span C" + and fg: "\x\ B. \y\ C. f x y = g x y" + shows "\x\S. \y\T. f x y = g x y " +proof- + let ?P = "\x. \y\ span C. f x y = g x y" + from bf bg have sp: "subspace ?P" + unfolding bilinear_def linear_def subspace_def bf bg + by(auto simp add: span_0 mem_def bilinear_lzero[OF bf] bilinear_lzero[OF bg] span_add Ball_def intro: bilinear_ladd[OF bf]) + + have "\x \ span B. \y\ span C. f x y = g x y" + apply - + apply (rule ballI) + apply (rule span_induct[of B ?P]) + defer + apply (rule sp) + apply assumption + apply (clarsimp simp add: Ball_def) + apply (rule_tac P="\y. f xa y = g xa y" and S=C in span_induct) + using fg + apply (auto simp add: subspace_def) + using bf bg unfolding bilinear_def linear_def + by(auto simp add: span_0 mem_def bilinear_rzero[OF bf] bilinear_rzero[OF bg] span_add Ball_def intro: bilinear_ladd[OF bf]) + then show ?thesis using SB TC by (auto intro: ext) +qed + +lemma bilinear_eq_stdbasis: + assumes bf: "bilinear (f:: 'a::ring_1^'m::finite \ 'a^'n::finite \ 'a^'p)" + and bg: "bilinear g" + and fg: "\i j. f (basis i) (basis j) = g (basis i) (basis j)" + shows "f = g" +proof- + from fg have th: "\x \ {basis i| i. i\ (UNIV :: 'm set)}. \y\ {basis j |j. j \ (UNIV :: 'n set)}. f x y = g x y" by blast + from bilinear_eq[OF bf bg equalityD2[OF span_stdbasis] equalityD2[OF span_stdbasis] th] show ?thesis by (blast intro: ext) +qed + +(* Detailed theorems about left and right invertibility in general case. *) + +lemma left_invertible_transp: + "(\(B::'a^'n^'m). B ** transp (A::'a^'n^'m) = mat (1::'a::comm_semiring_1)) \ (\(B::'a^'m^'n). A ** B = mat 1)" + by (metis matrix_transp_mul transp_mat transp_transp) + +lemma right_invertible_transp: + "(\(B::'a^'n^'m). transp (A::'a^'n^'m) ** B = mat (1::'a::comm_semiring_1)) \ (\(B::'a^'m^'n). B ** A = mat 1)" + by (metis matrix_transp_mul transp_mat transp_transp) + +lemma linear_injective_left_inverse: + assumes lf: "linear (f::real ^'n::finite \ real ^'m::finite)" and fi: "inj f" + shows "\g. linear g \ g o f = id" +proof- + from linear_independent_extend[OF independent_injective_image, OF independent_stdbasis, OF lf fi] + obtain h:: "real ^'m \ real ^'n" where h: "linear h" " \x \ f ` {basis i|i. i \ (UNIV::'n set)}. h x = inv f x" by blast + from h(2) + have th: "\i. (h \ f) (basis i) = id (basis i)" + using inv_o_cancel[OF fi, unfolded stupid_ext[symmetric] id_def o_def] + by auto + + from linear_eq_stdbasis[OF linear_compose[OF lf h(1)] linear_id th] + have "h o f = id" . + then show ?thesis using h(1) by blast +qed + +lemma linear_surjective_right_inverse: + assumes lf: "linear (f:: real ^'m::finite \ real ^'n::finite)" and sf: "surj f" + shows "\g. linear g \ f o g = id" +proof- + from linear_independent_extend[OF independent_stdbasis] + obtain h:: "real ^'n \ real ^'m" where + h: "linear h" "\ x\ {basis i| i. i\ (UNIV :: 'n set)}. h x = inv f x" by blast + from h(2) + have th: "\i. (f o h) (basis i) = id (basis i)" + using sf + apply (auto simp add: surj_iff o_def stupid_ext[symmetric]) + apply (erule_tac x="basis i" in allE) + by auto + + from linear_eq_stdbasis[OF linear_compose[OF h(1) lf] linear_id th] + have "f o h = id" . + then show ?thesis using h(1) by blast +qed + +lemma matrix_left_invertible_injective: +"(\B. (B::real^'m^'n) ** (A::real^'n::finite^'m::finite) = mat 1) \ (\x y. A *v x = A *v y \ x = y)" +proof- + {fix B:: "real^'m^'n" and x y assume B: "B ** A = mat 1" and xy: "A *v x = A*v y" + from xy have "B*v (A *v x) = B *v (A*v y)" by simp + hence "x = y" + unfolding matrix_vector_mul_assoc B matrix_vector_mul_lid .} + moreover + {assume A: "\x y. A *v x = A *v y \ x = y" + hence i: "inj (op *v A)" unfolding inj_on_def by auto + from linear_injective_left_inverse[OF matrix_vector_mul_linear i] + obtain g where g: "linear g" "g o op *v A = id" by blast + have "matrix g ** A = mat 1" + unfolding matrix_eq matrix_vector_mul_lid matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)] + using g(2) by (simp add: o_def id_def stupid_ext) + then have "\B. (B::real ^'m^'n) ** A = mat 1" by blast} + ultimately show ?thesis by blast +qed + +lemma matrix_left_invertible_ker: + "(\B. (B::real ^'m::finite^'n::finite) ** (A::real^'n^'m) = mat 1) \ (\x. A *v x = 0 \ x = 0)" + unfolding matrix_left_invertible_injective + using linear_injective_0[OF matrix_vector_mul_linear, of A] + by (simp add: inj_on_def) + +lemma matrix_right_invertible_surjective: +"(\B. (A::real^'n::finite^'m::finite) ** (B::real^'m^'n) = mat 1) \ surj (\x. A *v x)" +proof- + {fix B :: "real ^'m^'n" assume AB: "A ** B = mat 1" + {fix x :: "real ^ 'm" + have "A *v (B *v x) = x" + by (simp add: matrix_vector_mul_lid matrix_vector_mul_assoc AB)} + hence "surj (op *v A)" unfolding surj_def by metis } + moreover + {assume sf: "surj (op *v A)" + from linear_surjective_right_inverse[OF matrix_vector_mul_linear sf] + obtain g:: "real ^'m \ real ^'n" where g: "linear g" "op *v A o g = id" + by blast + + have "A ** (matrix g) = mat 1" + unfolding matrix_eq matrix_vector_mul_lid + matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)] + using g(2) unfolding o_def stupid_ext[symmetric] id_def + . + hence "\B. A ** (B::real^'m^'n) = mat 1" by blast + } + ultimately show ?thesis unfolding surj_def by blast +qed + +lemma matrix_left_invertible_independent_columns: + fixes A :: "real^'n::finite^'m::finite" + shows "(\(B::real ^'m^'n). B ** A = mat 1) \ (\c. setsum (\i. c i *s column i A) (UNIV :: 'n set) = 0 \ (\i. c i = 0))" + (is "?lhs \ ?rhs") +proof- + let ?U = "UNIV :: 'n set" + {assume k: "\x. A *v x = 0 \ x = 0" + {fix c i assume c: "setsum (\i. c i *s column i A) ?U = 0" + and i: "i \ ?U" + let ?x = "\ i. c i" + have th0:"A *v ?x = 0" + using c + unfolding matrix_mult_vsum Cart_eq + by auto + from k[rule_format, OF th0] i + have "c i = 0" by (vector Cart_eq)} + hence ?rhs by blast} + moreover + {assume H: ?rhs + {fix x assume x: "A *v x = 0" + let ?c = "\i. ((x$i ):: real)" + from H[rule_format, of ?c, unfolded matrix_mult_vsum[symmetric], OF x] + have "x = 0" by vector}} + ultimately show ?thesis unfolding matrix_left_invertible_ker by blast +qed + +lemma matrix_right_invertible_independent_rows: + fixes A :: "real^'n::finite^'m::finite" + shows "(\(B::real^'m^'n). A ** B = mat 1) \ (\c. setsum (\i. c i *s row i A) (UNIV :: 'm set) = 0 \ (\i. c i = 0))" + unfolding left_invertible_transp[symmetric] + matrix_left_invertible_independent_columns + by (simp add: column_transp) + +lemma matrix_right_invertible_span_columns: + "(\(B::real ^'n::finite^'m::finite). (A::real ^'m^'n) ** B = mat 1) \ span (columns A) = UNIV" (is "?lhs = ?rhs") +proof- + let ?U = "UNIV :: 'm set" + have fU: "finite ?U" by simp + have lhseq: "?lhs \ (\y. \(x::real^'m). setsum (\i. (x$i) *s column i A) ?U = y)" + unfolding matrix_right_invertible_surjective matrix_mult_vsum surj_def + apply (subst eq_commute) .. + have rhseq: "?rhs \ (\x. x \ span (columns A))" by blast + {assume h: ?lhs + {fix x:: "real ^'n" + from h[unfolded lhseq, rule_format, of x] obtain y:: "real ^'m" + where y: "setsum (\i. (y$i) *s column i A) ?U = x" by blast + have "x \ span (columns A)" + unfolding y[symmetric] + apply (rule span_setsum[OF fU]) + apply clarify + apply (rule span_mul) + apply (rule span_superset) + unfolding columns_def + by blast} + then have ?rhs unfolding rhseq by blast} + moreover + {assume h:?rhs + let ?P = "\(y::real ^'n). \(x::real^'m). setsum (\i. (x$i) *s column i A) ?U = y" + {fix y have "?P y" + proof(rule span_induct_alt[of ?P "columns A"]) + show "\x\real ^ 'm. setsum (\i. (x$i) *s column i A) ?U = 0" + apply (rule exI[where x=0]) + by (simp add: zero_index vector_smult_lzero) + next + fix c y1 y2 assume y1: "y1 \ columns A" and y2: "?P y2" + from y1 obtain i where i: "i \ ?U" "y1 = column i A" + unfolding columns_def by blast + from y2 obtain x:: "real ^'m" where + x: "setsum (\i. (x$i) *s column i A) ?U = y2" by blast + let ?x = "(\ j. if j = i then c + (x$i) else (x$j))::real^'m" + show "?P (c*s y1 + y2)" + proof(rule exI[where x= "?x"], vector, auto simp add: i x[symmetric] cond_value_iff right_distrib cond_application_beta cong del: if_weak_cong) + fix j + have th: "\xa \ ?U. (if xa = i then (c + (x$i)) * ((column xa A)$j) + else (x$xa) * ((column xa A$j))) = (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))" using i(1) + by (simp add: ring_simps) + have "setsum (\xa. if xa = i then (c + (x$i)) * ((column xa A)$j) + else (x$xa) * ((column xa A$j))) ?U = setsum (\xa. (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))) ?U" + apply (rule setsum_cong[OF refl]) + using th by blast + also have "\ = setsum (\xa. if xa = i then c * ((column i A)$j) else 0) ?U + setsum (\xa. ((x$xa) * ((column xa A)$j))) ?U" + by (simp add: setsum_addf) + also have "\ = c * ((column i A)$j) + setsum (\xa. ((x$xa) * ((column xa A)$j))) ?U" + unfolding setsum_delta[OF fU] + using i(1) by simp + finally show "setsum (\xa. if xa = i then (c + (x$i)) * ((column xa A)$j) + else (x$xa) * ((column xa A$j))) ?U = c * ((column i A)$j) + setsum (\xa. ((x$xa) * ((column xa A)$j))) ?U" . + qed + next + show "y \ span (columns A)" unfolding h by blast + qed} + then have ?lhs unfolding lhseq ..} + ultimately show ?thesis by blast +qed + +lemma matrix_left_invertible_span_rows: + "(\(B::real^'m::finite^'n::finite). B ** (A::real^'n^'m) = mat 1) \ span (rows A) = UNIV" + unfolding right_invertible_transp[symmetric] + unfolding columns_transp[symmetric] + unfolding matrix_right_invertible_span_columns + .. + +(* An injective map real^'n->real^'n is also surjective. *) + +lemma linear_injective_imp_surjective: + assumes lf: "linear (f:: real ^'n::finite \ real ^'n)" and fi: "inj f" + shows "surj f" +proof- + let ?U = "UNIV :: (real ^'n) set" + from basis_exists[of ?U] obtain B + where B: "B \ ?U" "independent B" "?U \ span B" "B hassize dim ?U" + by blast + from B(4) have d: "dim ?U = card B" by (simp add: hassize_def) + have th: "?U \ span (f ` B)" + apply (rule card_ge_dim_independent) + apply blast + apply (rule independent_injective_image[OF B(2) lf fi]) + apply (rule order_eq_refl) + apply (rule sym) + unfolding d + apply (rule card_image) + apply (rule subset_inj_on[OF fi]) + by blast + from th show ?thesis + unfolding span_linear_image[OF lf] surj_def + using B(3) by blast +qed + +(* And vice versa. *) + +lemma surjective_iff_injective_gen: + assumes fS: "finite S" and fT: "finite T" and c: "card S = card T" + and ST: "f ` S \ T" + shows "(\y \ T. \x \ S. f x = y) \ inj_on f S" (is "?lhs \ ?rhs") +proof- + {assume h: "?lhs" + {fix x y assume x: "x \ S" and y: "y \ S" and f: "f x = f y" + from x fS have S0: "card S \ 0" by auto + {assume xy: "x \ y" + have th: "card S \ card (f ` (S - {y}))" + unfolding c + apply (rule card_mono) + apply (rule finite_imageI) + using fS apply simp + using h xy x y f unfolding subset_eq image_iff + apply auto + apply (case_tac "xa = f x") + apply (rule bexI[where x=x]) + apply auto + done + also have " \ \ card (S -{y})" + apply (rule card_image_le) + using fS by simp + also have "\ \ card S - 1" using y fS by simp + finally have False using S0 by arith } + then have "x = y" by blast} + then have ?rhs unfolding inj_on_def by blast} + moreover + {assume h: ?rhs + have "f ` S = T" + apply (rule card_subset_eq[OF fT ST]) + unfolding card_image[OF h] using c . + then have ?lhs by blast} + ultimately show ?thesis by blast +qed + +lemma linear_surjective_imp_injective: + assumes lf: "linear (f::real ^'n::finite => real ^'n)" and sf: "surj f" + shows "inj f" +proof- + let ?U = "UNIV :: (real ^'n) set" + from basis_exists[of ?U] obtain B + where B: "B \ ?U" "independent B" "?U \ span B" "B hassize dim ?U" + by blast + {fix x assume x: "x \ span B" and fx: "f x = 0" + from B(4) have fB: "finite B" by (simp add: hassize_def) + from B(4) have d: "dim ?U = card B" by (simp add: hassize_def) + have fBi: "independent (f ` B)" + apply (rule card_le_dim_spanning[of "f ` B" ?U]) + apply blast + using sf B(3) + unfolding span_linear_image[OF lf] surj_def subset_eq image_iff + apply blast + using fB apply (blast intro: finite_imageI) + unfolding d + apply (rule card_image_le) + apply (rule fB) + done + have th0: "dim ?U \ card (f ` B)" + apply (rule span_card_ge_dim) + apply blast + unfolding span_linear_image[OF lf] + apply (rule subset_trans[where B = "f ` UNIV"]) + using sf unfolding surj_def apply blast + apply (rule image_mono) + apply (rule B(3)) + apply (metis finite_imageI fB) + done + + moreover have "card (f ` B) \ card B" + by (rule card_image_le, rule fB) + ultimately have th1: "card B = card (f ` B)" unfolding d by arith + have fiB: "inj_on f B" + unfolding surjective_iff_injective_gen[OF fB finite_imageI[OF fB] th1 subset_refl, symmetric] by blast + from linear_indep_image_lemma[OF lf fB fBi fiB x] fx + have "x = 0" by blast} + note th = this + from th show ?thesis unfolding linear_injective_0[OF lf] + using B(3) by blast +qed + +(* Hence either is enough for isomorphism. *) + +lemma left_right_inverse_eq: + assumes fg: "f o g = id" and gh: "g o h = id" + shows "f = h" +proof- + have "f = f o (g o h)" unfolding gh by simp + also have "\ = (f o g) o h" by (simp add: o_assoc) + finally show "f = h" unfolding fg by simp +qed + +lemma isomorphism_expand: + "f o g = id \ g o f = id \ (\x. f(g x) = x) \ (\x. g(f x) = x)" + by (simp add: expand_fun_eq o_def id_def) + +lemma linear_injective_isomorphism: + assumes lf: "linear (f :: real^'n::finite \ real ^'n)" and fi: "inj f" + shows "\f'. linear f' \ (\x. f' (f x) = x) \ (\x. f (f' x) = x)" +unfolding isomorphism_expand[symmetric] +using linear_surjective_right_inverse[OF lf linear_injective_imp_surjective[OF lf fi]] linear_injective_left_inverse[OF lf fi] +by (metis left_right_inverse_eq) + +lemma linear_surjective_isomorphism: + assumes lf: "linear (f::real ^'n::finite \ real ^'n)" and sf: "surj f" + shows "\f'. linear f' \ (\x. f' (f x) = x) \ (\x. f (f' x) = x)" +unfolding isomorphism_expand[symmetric] +using linear_surjective_right_inverse[OF lf sf] linear_injective_left_inverse[OF lf linear_surjective_imp_injective[OF lf sf]] +by (metis left_right_inverse_eq) + +(* Left and right inverses are the same for R^N->R^N. *) + +lemma linear_inverse_left: + assumes lf: "linear (f::real ^'n::finite \ real ^'n)" and lf': "linear f'" + shows "f o f' = id \ f' o f = id" +proof- + {fix f f':: "real ^'n \ real ^'n" + assume lf: "linear f" "linear f'" and f: "f o f' = id" + from f have sf: "surj f" + + apply (auto simp add: o_def stupid_ext[symmetric] id_def surj_def) + by metis + from linear_surjective_isomorphism[OF lf(1) sf] lf f + have "f' o f = id" unfolding stupid_ext[symmetric] o_def id_def + by metis} + then show ?thesis using lf lf' by metis +qed + +(* Moreover, a one-sided inverse is automatically linear. *) + +lemma left_inverse_linear: + assumes lf: "linear (f::real ^'n::finite \ real ^'n)" and gf: "g o f = id" + shows "linear g" +proof- + from gf have fi: "inj f" apply (auto simp add: inj_on_def o_def id_def stupid_ext[symmetric]) + by metis + from linear_injective_isomorphism[OF lf fi] + obtain h:: "real ^'n \ real ^'n" where + h: "linear h" "\x. h (f x) = x" "\x. f (h x) = x" by blast + have "h = g" apply (rule ext) using gf h(2,3) + apply (simp add: o_def id_def stupid_ext[symmetric]) + by metis + with h(1) show ?thesis by blast +qed + +lemma right_inverse_linear: + assumes lf: "linear (f:: real ^'n::finite \ real ^'n)" and gf: "f o g = id" + shows "linear g" +proof- + from gf have fi: "surj f" apply (auto simp add: surj_def o_def id_def stupid_ext[symmetric]) + by metis + from linear_surjective_isomorphism[OF lf fi] + obtain h:: "real ^'n \ real ^'n" where + h: "linear h" "\x. h (f x) = x" "\x. f (h x) = x" by blast + have "h = g" apply (rule ext) using gf h(2,3) + apply (simp add: o_def id_def stupid_ext[symmetric]) + by metis + with h(1) show ?thesis by blast +qed + +(* The same result in terms of square matrices. *) + +lemma matrix_left_right_inverse: + fixes A A' :: "real ^'n::finite^'n" + shows "A ** A' = mat 1 \ A' ** A = mat 1" +proof- + {fix A A' :: "real ^'n^'n" assume AA': "A ** A' = mat 1" + have sA: "surj (op *v A)" + unfolding surj_def + apply clarify + apply (rule_tac x="(A' *v y)" in exI) + by (simp add: matrix_vector_mul_assoc AA' matrix_vector_mul_lid) + from linear_surjective_isomorphism[OF matrix_vector_mul_linear sA] + obtain f' :: "real ^'n \ real ^'n" + where f': "linear f'" "\x. f' (A *v x) = x" "\x. A *v f' x = x" by blast + have th: "matrix f' ** A = mat 1" + by (simp add: matrix_eq matrix_works[OF f'(1)] matrix_vector_mul_assoc[symmetric] matrix_vector_mul_lid f'(2)[rule_format]) + hence "(matrix f' ** A) ** A' = mat 1 ** A'" by simp + hence "matrix f' = A'" by (simp add: matrix_mul_assoc[symmetric] AA' matrix_mul_rid matrix_mul_lid) + hence "matrix f' ** A = A' ** A" by simp + hence "A' ** A = mat 1" by (simp add: th)} + then show ?thesis by blast +qed + +(* Considering an n-element vector as an n-by-1 or 1-by-n matrix. *) + +definition "rowvector v = (\ i j. (v$j))" + +definition "columnvector v = (\ i j. (v$i))" + +lemma transp_columnvector: + "transp(columnvector v) = rowvector v" + by (simp add: transp_def rowvector_def columnvector_def Cart_eq) + +lemma transp_rowvector: "transp(rowvector v) = columnvector v" + by (simp add: transp_def columnvector_def rowvector_def Cart_eq) + +lemma dot_rowvector_columnvector: + "columnvector (A *v v) = A ** columnvector v" + by (vector columnvector_def matrix_matrix_mult_def matrix_vector_mult_def) + +lemma dot_matrix_product: "(x::'a::semiring_1^'n::finite) \ y = (((rowvector x ::'a^'n^1) ** (columnvector y :: 'a^1^'n))$1)$1" + by (vector matrix_matrix_mult_def rowvector_def columnvector_def dot_def) + +lemma dot_matrix_vector_mul: + fixes A B :: "real ^'n::finite ^'n" and x y :: "real ^'n" + shows "(A *v x) \ (B *v y) = + (((rowvector x :: real^'n^1) ** ((transp A ** B) ** (columnvector y :: real ^1^'n)))$1)$1" +unfolding dot_matrix_product transp_columnvector[symmetric] + dot_rowvector_columnvector matrix_transp_mul matrix_mul_assoc .. + +(* Infinity norm. *) + +definition "infnorm (x::real^'n::finite) = rsup {abs(x$i) |i. i\ (UNIV :: 'n set)}" + +lemma numseg_dimindex_nonempty: "\i. i \ (UNIV :: 'n set)" + by auto + +lemma infnorm_set_image: + "{abs(x$i) |i. i\ (UNIV :: 'n set)} = + (\i. abs(x$i)) ` (UNIV :: 'n set)" by blast + +lemma infnorm_set_lemma: + shows "finite {abs((x::'a::abs ^'n::finite)$i) |i. i\ (UNIV :: 'n set)}" + and "{abs(x$i) |i. i\ (UNIV :: 'n::finite set)} \ {}" + unfolding infnorm_set_image + by (auto intro: finite_imageI) + +lemma infnorm_pos_le: "0 \ infnorm (x::real^'n::finite)" + unfolding infnorm_def + unfolding rsup_finite_ge_iff[ OF infnorm_set_lemma] + unfolding infnorm_set_image + by auto + +lemma infnorm_triangle: "infnorm ((x::real^'n::finite) + y) \ infnorm x + infnorm y" +proof- + have th: "\x y (z::real). x - y <= z \ x - z <= y" by arith + have th1: "\S f. f ` S = { f i| i. i \ S}" by blast + have th2: "\x (y::real). abs(x + y) - abs(x) <= abs(y)" by arith + show ?thesis + unfolding infnorm_def + unfolding rsup_finite_le_iff[ OF infnorm_set_lemma] + apply (subst diff_le_eq[symmetric]) + unfolding rsup_finite_ge_iff[ OF infnorm_set_lemma] + unfolding infnorm_set_image bex_simps + apply (subst th) + unfolding th1 + unfolding rsup_finite_ge_iff[ OF infnorm_set_lemma] + + unfolding infnorm_set_image ball_simps bex_simps + apply simp + apply (metis th2) + done +qed + +lemma infnorm_eq_0: "infnorm x = 0 \ (x::real ^'n::finite) = 0" +proof- + have "infnorm x <= 0 \ x = 0" + unfolding infnorm_def + unfolding rsup_finite_le_iff[OF infnorm_set_lemma] + unfolding infnorm_set_image ball_simps + by vector + then show ?thesis using infnorm_pos_le[of x] by simp +qed + +lemma infnorm_0: "infnorm 0 = 0" + by (simp add: infnorm_eq_0) + +lemma infnorm_neg: "infnorm (- x) = infnorm x" + unfolding infnorm_def + apply (rule cong[of "rsup" "rsup"]) + apply blast + apply (rule set_ext) + apply auto + done + +lemma infnorm_sub: "infnorm (x - y) = infnorm (y - x)" +proof- + have "y - x = - (x - y)" by simp + then show ?thesis by (metis infnorm_neg) +qed + +lemma real_abs_sub_infnorm: "\ infnorm x - infnorm y\ \ infnorm (x - y)" +proof- + have th: "\(nx::real) n ny. nx <= n + ny \ ny <= n + nx ==> \nx - ny\ <= n" + by arith + from infnorm_triangle[of "x - y" " y"] infnorm_triangle[of "x - y" "-x"] + have ths: "infnorm x \ infnorm (x - y) + infnorm y" + "infnorm y \ infnorm (x - y) + infnorm x" + by (simp_all add: ring_simps infnorm_neg diff_def[symmetric]) + from th[OF ths] show ?thesis . +qed + +lemma real_abs_infnorm: " \infnorm x\ = infnorm x" + using infnorm_pos_le[of x] by arith + +lemma component_le_infnorm: + shows "\x$i\ \ infnorm (x::real^'n::finite)" +proof- + let ?U = "UNIV :: 'n set" + let ?S = "{\x$i\ |i. i\ ?U}" + have fS: "finite ?S" unfolding image_Collect[symmetric] + apply (rule finite_imageI) unfolding Collect_def mem_def by simp + have S0: "?S \ {}" by blast + have th1: "\S f. f ` S = { f i| i. i \ S}" by blast + from rsup_finite_in[OF fS S0] rsup_finite_Ub[OF fS S0] + show ?thesis unfolding infnorm_def isUb_def setle_def + unfolding infnorm_set_image ball_simps by auto +qed + +lemma infnorm_mul_lemma: "infnorm(a *s x) <= \a\ * infnorm x" + apply (subst infnorm_def) + unfolding rsup_finite_le_iff[OF infnorm_set_lemma] + unfolding infnorm_set_image ball_simps + apply (simp add: abs_mult) + apply (rule allI) + apply (cut_tac component_le_infnorm[of x]) + apply (rule mult_mono) + apply auto + done + +lemma infnorm_mul: "infnorm(a *s x) = abs a * infnorm x" +proof- + {assume a0: "a = 0" hence ?thesis by (simp add: infnorm_0) } + moreover + {assume a0: "a \ 0" + from a0 have th: "(1/a) *s (a *s x) = x" + by (simp add: vector_smult_assoc) + from a0 have ap: "\a\ > 0" by arith + from infnorm_mul_lemma[of "1/a" "a *s x"] + have "infnorm x \ 1/\a\ * infnorm (a*s x)" + unfolding th by simp + with ap have "\a\ * infnorm x \ \a\ * (1/\a\ * infnorm (a *s x))" by (simp add: field_simps) + then have "\a\ * infnorm x \ infnorm (a*s x)" + using ap by (simp add: field_simps) + with infnorm_mul_lemma[of a x] have ?thesis by arith } + ultimately show ?thesis by blast +qed + +lemma infnorm_pos_lt: "infnorm x > 0 \ x \ 0" + using infnorm_pos_le[of x] infnorm_eq_0[of x] by arith + +(* Prove that it differs only up to a bound from Euclidean norm. *) + +lemma infnorm_le_norm: "infnorm x \ norm x" + unfolding infnorm_def rsup_finite_le_iff[OF infnorm_set_lemma] + unfolding infnorm_set_image ball_simps + by (metis component_le_norm) +lemma card_enum: "card {1 .. n} = n" by auto +lemma norm_le_infnorm: "norm(x) <= sqrt(real CARD('n)) * infnorm(x::real ^'n::finite)" +proof- + let ?d = "CARD('n)" + have "real ?d \ 0" by simp + hence d2: "(sqrt (real ?d))^2 = real ?d" + by (auto intro: real_sqrt_pow2) + have th: "sqrt (real ?d) * infnorm x \ 0" + by (simp add: zero_le_mult_iff real_sqrt_ge_0_iff infnorm_pos_le) + have th1: "x\x \ (sqrt (real ?d) * infnorm x)^2" + unfolding power_mult_distrib d2 + apply (subst power2_abs[symmetric]) + unfolding real_of_nat_def dot_def power2_eq_square[symmetric] + apply (subst power2_abs[symmetric]) + apply (rule setsum_bounded) + apply (rule power_mono) + unfolding abs_of_nonneg[OF infnorm_pos_le] + unfolding infnorm_def rsup_finite_ge_iff[OF infnorm_set_lemma] + unfolding infnorm_set_image bex_simps + apply blast + by (rule abs_ge_zero) + from real_le_lsqrt[OF dot_pos_le th th1] + show ?thesis unfolding real_vector_norm_def id_def . +qed + +(* Equality in Cauchy-Schwarz and triangle inequalities. *) + +lemma norm_cauchy_schwarz_eq: "(x::real ^'n::finite) \ y = norm x * norm y \ norm x *s y = norm y *s x" (is "?lhs \ ?rhs") +proof- + {assume h: "x = 0" + hence ?thesis by simp} + moreover + {assume h: "y = 0" + hence ?thesis by simp} + moreover + {assume x: "x \ 0" and y: "y \ 0" + from dot_eq_0[of "norm y *s x - norm x *s y"] + have "?rhs \ (norm y * (norm y * norm x * norm x - norm x * (x \ y)) - norm x * (norm y * (y \ x) - norm x * norm y * norm y) = 0)" + using x y + unfolding dot_rsub dot_lsub dot_lmult dot_rmult + unfolding norm_pow_2[symmetric] power2_eq_square diff_eq_0_iff_eq apply (simp add: dot_sym) + apply (simp add: ring_simps) + apply metis + done + also have "\ \ (2 * norm x * norm y * (norm x * norm y - x \ y) = 0)" using x y + by (simp add: ring_simps dot_sym) + also have "\ \ ?lhs" using x y + apply simp + by metis + finally have ?thesis by blast} + ultimately show ?thesis by blast +qed + +lemma norm_cauchy_schwarz_abs_eq: + fixes x y :: "real ^ 'n::finite" + shows "abs(x \ y) = norm x * norm y \ + norm x *s y = norm y *s x \ norm(x) *s y = - norm y *s x" (is "?lhs \ ?rhs") +proof- + have th: "\(x::real) a. a \ 0 \ abs x = a \ x = a \ x = - a" by arith + have "?rhs \ norm x *s y = norm y *s x \ norm (- x) *s y = norm y *s (- x)" + apply simp by vector + also have "\ \(x \ y = norm x * norm y \ + (-x) \ y = norm x * norm y)" + unfolding norm_cauchy_schwarz_eq[symmetric] + unfolding norm_minus_cancel + norm_mul by blast + also have "\ \ ?lhs" + unfolding th[OF mult_nonneg_nonneg, OF norm_ge_zero[of x] norm_ge_zero[of y]] dot_lneg + by arith + finally show ?thesis .. +qed + +lemma norm_triangle_eq: + fixes x y :: "real ^ 'n::finite" + shows "norm(x + y) = norm x + norm y \ norm x *s y = norm y *s x" +proof- + {assume x: "x =0 \ y =0" + hence ?thesis by (cases "x=0", simp_all)} + moreover + {assume x: "x \ 0" and y: "y \ 0" + hence "norm x \ 0" "norm y \ 0" + by simp_all + hence n: "norm x > 0" "norm y > 0" + using norm_ge_zero[of x] norm_ge_zero[of y] + by arith+ + have th: "\(a::real) b c. a + b + c \ 0 ==> (a = b + c \ a^2 = (b + c)^2)" by algebra + have "norm(x + y) = norm x + norm y \ norm(x + y)^ 2 = (norm x + norm y) ^2" + apply (rule th) using n norm_ge_zero[of "x + y"] + by arith + also have "\ \ norm x *s y = norm y *s x" + unfolding norm_cauchy_schwarz_eq[symmetric] + unfolding norm_pow_2 dot_ladd dot_radd + by (simp add: norm_pow_2[symmetric] power2_eq_square dot_sym ring_simps) + finally have ?thesis .} + ultimately show ?thesis by blast +qed + +(* Collinearity.*) + +definition "collinear S \ (\u. \x \ S. \ y \ S. \c. x - y = c *s u)" + +lemma collinear_empty: "collinear {}" by (simp add: collinear_def) + +lemma collinear_sing: "collinear {(x::'a::ring_1^'n)}" + apply (simp add: collinear_def) + apply (rule exI[where x=0]) + by simp + +lemma collinear_2: "collinear {(x::'a::ring_1^'n),y}" + apply (simp add: collinear_def) + apply (rule exI[where x="x - y"]) + apply auto + apply (rule exI[where x=0], simp) + apply (rule exI[where x=1], simp) + apply (rule exI[where x="- 1"], simp add: vector_sneg_minus1[symmetric]) + apply (rule exI[where x=0], simp) + done + +lemma collinear_lemma: "collinear {(0::real^'n),x,y} \ x = 0 \ y = 0 \ (\c. y = c *s x)" (is "?lhs \ ?rhs") +proof- + {assume "x=0 \ y = 0" hence ?thesis + by (cases "x = 0", simp_all add: collinear_2 insert_commute)} + moreover + {assume x: "x \ 0" and y: "y \ 0" + {assume h: "?lhs" + then obtain u where u: "\ x\ {0,x,y}. \y\ {0,x,y}. \c. x - y = c *s u" unfolding collinear_def by blast + from u[rule_format, of x 0] u[rule_format, of y 0] + obtain cx and cy where + cx: "x = cx*s u" and cy: "y = cy*s u" + by auto + from cx x have cx0: "cx \ 0" by auto + from cy y have cy0: "cy \ 0" by auto + let ?d = "cy / cx" + from cx cy cx0 have "y = ?d *s x" + by (simp add: vector_smult_assoc) + hence ?rhs using x y by blast} + moreover + {assume h: "?rhs" + then obtain c where c: "y = c*s x" using x y by blast + have ?lhs unfolding collinear_def c + apply (rule exI[where x=x]) + apply auto + apply (rule exI[where x="- 1"], simp only: vector_smult_lneg vector_smult_lid) + apply (rule exI[where x= "-c"], simp only: vector_smult_lneg) + apply (rule exI[where x=1], simp) + apply (rule exI[where x="1 - c"], simp add: vector_smult_lneg vector_sub_rdistrib) + apply (rule exI[where x="c - 1"], simp add: vector_smult_lneg vector_sub_rdistrib) + done} + ultimately have ?thesis by blast} + ultimately show ?thesis by blast +qed + +lemma norm_cauchy_schwarz_equal: + fixes x y :: "real ^ 'n::finite" + shows "abs(x \ y) = norm x * norm y \ collinear {(0::real^'n),x,y}" +unfolding norm_cauchy_schwarz_abs_eq +apply (cases "x=0", simp_all add: collinear_2) +apply (cases "y=0", simp_all add: collinear_2 insert_commute) +unfolding collinear_lemma +apply simp +apply (subgoal_tac "norm x \ 0") +apply (subgoal_tac "norm y \ 0") +apply (rule iffI) +apply (cases "norm x *s y = norm y *s x") +apply (rule exI[where x="(1/norm x) * norm y"]) +apply (drule sym) +unfolding vector_smult_assoc[symmetric] +apply (simp add: vector_smult_assoc field_simps) +apply (rule exI[where x="(1/norm x) * - norm y"]) +apply clarify +apply (drule sym) +unfolding vector_smult_assoc[symmetric] +apply (simp add: vector_smult_assoc field_simps) +apply (erule exE) +apply (erule ssubst) +unfolding vector_smult_assoc +unfolding norm_mul +apply (subgoal_tac "norm x * c = \c\ * norm x \ norm x * c = - \c\ * norm x") +apply (case_tac "c <= 0", simp add: ring_simps) +apply (simp add: ring_simps) +apply (case_tac "c <= 0", simp add: ring_simps) +apply (simp add: ring_simps) +apply simp +apply simp +done + +end diff -r 1fad3160d873 -r 2083bde13ce1 src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy Fri Oct 23 13:23:18 2009 +0200 @@ -0,0 +1,95 @@ +(* Title: HOL/Library/Finite_Cartesian_Product + Author: Amine Chaieb, University of Cambridge +*) + +header {* Definition of finite Cartesian product types. *} + +theory Finite_Cartesian_Product +imports Main (*FIXME: ATP_Linkup is only needed for metis at a few places. We could dispense of that by changing the proofs.*) +begin + +definition hassize (infixr "hassize" 12) where + "(S hassize n) = (finite S \ card S = n)" + +lemma hassize_image_inj: assumes f: "inj_on f S" and S: "S hassize n" + shows "f ` S hassize n" + using f S card_image[OF f] + by (simp add: hassize_def inj_on_def) + + +subsection {* Finite Cartesian products, with indexing and lambdas. *} + +typedef (open Cart) + ('a, 'b) "^" (infixl "^" 15) + = "UNIV :: ('b \ 'a) set" + morphisms Cart_nth Cart_lambda .. + +notation Cart_nth (infixl "$" 90) + +notation (xsymbols) Cart_lambda (binder "\" 10) + +lemma stupid_ext: "(\x. f x = g x) \ (f = g)" + apply auto + apply (rule ext) + apply auto + done + +lemma Cart_eq: "((x:: 'a ^ 'b) = y) \ (\i. x$i = y$i)" + by (simp add: Cart_nth_inject [symmetric] expand_fun_eq) + +lemma Cart_lambda_beta [simp]: "Cart_lambda g $ i = g i" + by (simp add: Cart_lambda_inverse) + +lemma Cart_lambda_unique: + fixes f :: "'a ^ 'b" + shows "(\i. f$i = g i) \ Cart_lambda g = f" + by (auto simp add: Cart_eq) + +lemma Cart_lambda_eta: "(\ i. (g$i)) = g" + by (simp add: Cart_eq) + +text{* A non-standard sum to "paste" Cartesian products. *} + +definition pastecart :: "'a ^ 'm \ 'a ^ 'n \ 'a ^ ('m + 'n)" where + "pastecart f g = (\ i. case i of Inl a \ f$a | Inr b \ g$b)" + +definition fstcart:: "'a ^('m + 'n) \ 'a ^ 'm" where + "fstcart f = (\ i. (f$(Inl i)))" + +definition sndcart:: "'a ^('m + 'n) \ 'a ^ 'n" where + "sndcart f = (\ i. (f$(Inr i)))" + +lemma nth_pastecart_Inl [simp]: "pastecart f g $ Inl a = f$a" + unfolding pastecart_def by simp + +lemma nth_pastecart_Inr [simp]: "pastecart f g $ Inr b = g$b" + unfolding pastecart_def by simp + +lemma nth_fstcart [simp]: "fstcart f $ i = f $ Inl i" + unfolding fstcart_def by simp + +lemma nth_sndtcart [simp]: "sndcart f $ i = f $ Inr i" + unfolding sndcart_def by simp + +lemma finite_sum_image: "(UNIV::('a + 'b) set) = range Inl \ range Inr" +by (auto, case_tac x, auto) + +lemma fstcart_pastecart: "fstcart (pastecart (x::'a ^'m ) (y:: 'a ^ 'n)) = x" + by (simp add: Cart_eq) + +lemma sndcart_pastecart: "sndcart (pastecart (x::'a ^'m ) (y:: 'a ^ 'n)) = y" + by (simp add: Cart_eq) + +lemma pastecart_fst_snd: "pastecart (fstcart z) (sndcart z) = z" + by (simp add: Cart_eq pastecart_def fstcart_def sndcart_def split: sum.split) + +lemma pastecart_eq: "(x = y) \ (fstcart x = fstcart y) \ (sndcart x = sndcart y)" + using pastecart_fst_snd[of x] pastecart_fst_snd[of y] by metis + +lemma forall_pastecart: "(\p. P p) \ (\x y. P (pastecart x y))" + by (metis pastecart_fst_snd fstcart_pastecart sndcart_pastecart) + +lemma exists_pastecart: "(\p. P p) \ (\x y. P (pastecart x y))" + by (metis pastecart_fst_snd fstcart_pastecart sndcart_pastecart) + +end diff -r 1fad3160d873 -r 2083bde13ce1 src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy Fri Oct 23 13:23:18 2009 +0200 @@ -0,0 +1,6 @@ +theory Multivariate_Analysis imports + Convex_Euclidean_Space + Determinants +begin + +end diff -r 1fad3160d873 -r 2083bde13ce1 src/HOL/Multivariate_Analysis/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Multivariate_Analysis/ROOT.ML Fri Oct 23 13:23:18 2009 +0200 @@ -0,0 +1,6 @@ +(* + no_document use_thy "ThisTheory"; + use_thy "ThatTheory"; +*) + +use_thy "Multivariate_Analysis"; diff -r 1fad3160d873 -r 2083bde13ce1 src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Fri Oct 23 13:23:18 2009 +0200 @@ -0,0 +1,6027 @@ +(* Title: HOL/Library/Topology_Euclidian_Space.thy + Author: Amine Chaieb, University of Cambridge + Author: Robert Himmelmann, TU Muenchen +*) + +header {* Elementary topology in Euclidean space. *} + +theory Topology_Euclidean_Space +imports SEQ Euclidean_Space Product_Vector +begin + +declare fstcart_pastecart[simp] sndcart_pastecart[simp] + +subsection{* General notion of a topology *} + +definition "istopology L \ {} \ L \ (\S \L. \T \L. S \ T \ L) \ (\K. K \L \ \ K \ L)" +typedef (open) 'a topology = "{L::('a set) set. istopology L}" + morphisms "openin" "topology" + unfolding istopology_def by blast + +lemma istopology_open_in[intro]: "istopology(openin U)" + using openin[of U] by blast + +lemma topology_inverse': "istopology U \ openin (topology U) = U" + using topology_inverse[unfolded mem_def Collect_def] . + +lemma topology_inverse_iff: "istopology U \ openin (topology U) = U" + using topology_inverse[of U] istopology_open_in[of "topology U"] by auto + +lemma topology_eq: "T1 = T2 \ (\S. openin T1 S \ openin T2 S)" +proof- + {assume "T1=T2" hence "\S. openin T1 S \ openin T2 S" by simp} + moreover + {assume H: "\S. openin T1 S \ openin T2 S" + hence "openin T1 = openin T2" by (metis mem_def set_ext) + hence "topology (openin T1) = topology (openin T2)" by simp + hence "T1 = T2" unfolding openin_inverse .} + ultimately show ?thesis by blast +qed + +text{* Infer the "universe" from union of all sets in the topology. *} + +definition "topspace T = \{S. openin T S}" + +subsection{* Main properties of open sets *} + +lemma openin_clauses: + fixes U :: "'a topology" + shows "openin U {}" + "\S T. openin U S \ openin U T \ openin U (S\T)" + "\K. (\S \ K. openin U S) \ openin U (\K)" + using openin[of U] unfolding istopology_def Collect_def mem_def + by (metis mem_def subset_eq)+ + +lemma openin_subset[intro]: "openin U S \ S \ topspace U" + unfolding topspace_def by blast +lemma openin_empty[simp]: "openin U {}" by (simp add: openin_clauses) + +lemma openin_Int[intro]: "openin U S \ openin U T \ openin U (S \ T)" + by (simp add: openin_clauses) + +lemma openin_Union[intro]: "(\S \K. openin U S) \ openin U (\ K)" by (simp add: openin_clauses) + +lemma openin_Un[intro]: "openin U S \ openin U T \ openin U (S \ T)" + using openin_Union[of "{S,T}" U] by auto + +lemma openin_topspace[intro, simp]: "openin U (topspace U)" by (simp add: openin_Union topspace_def) + +lemma openin_subopen: "openin U S \ (\x \ S. \T. openin U T \ x \ T \ T \ S)" (is "?lhs \ ?rhs") +proof- + {assume ?lhs then have ?rhs by auto } + moreover + {assume H: ?rhs + then obtain t where t: "\x\S. openin U (t x) \ x \ t x \ t x \ S" + unfolding Ball_def ex_simps(6)[symmetric] choice_iff by blast + from t have th0: "\x\ t`S. openin U x" by auto + have "\ t`S = S" using t by auto + with openin_Union[OF th0] have "openin U S" by simp } + ultimately show ?thesis by blast +qed + +subsection{* Closed sets *} + +definition "closedin U S \ S \ topspace U \ openin U (topspace U - S)" + +lemma closedin_subset: "closedin U S \ S \ topspace U" by (metis closedin_def) +lemma closedin_empty[simp]: "closedin U {}" by (simp add: closedin_def) +lemma closedin_topspace[intro,simp]: + "closedin U (topspace U)" by (simp add: closedin_def) +lemma closedin_Un[intro]: "closedin U S \ closedin U T \ closedin U (S \ T)" + by (auto simp add: Diff_Un closedin_def) + +lemma Diff_Inter[intro]: "A - \S = \ {A - s|s. s\S}" by auto +lemma closedin_Inter[intro]: assumes Ke: "K \ {}" and Kc: "\S \K. closedin U S" + shows "closedin U (\ K)" using Ke Kc unfolding closedin_def Diff_Inter by auto + +lemma closedin_Int[intro]: "closedin U S \ closedin U T \ closedin U (S \ T)" + using closedin_Inter[of "{S,T}" U] by auto + +lemma Diff_Diff_Int: "A - (A - B) = A \ B" by blast +lemma openin_closedin_eq: "openin U S \ S \ topspace U \ closedin U (topspace U - S)" + apply (auto simp add: closedin_def Diff_Diff_Int inf_absorb2) + apply (metis openin_subset subset_eq) + done + +lemma openin_closedin: "S \ topspace U \ (openin U S \ closedin U (topspace U - S))" + by (simp add: openin_closedin_eq) + +lemma openin_diff[intro]: assumes oS: "openin U S" and cT: "closedin U T" shows "openin U (S - T)" +proof- + have "S - T = S \ (topspace U - T)" using openin_subset[of U S] oS cT + by (auto simp add: topspace_def openin_subset) + then show ?thesis using oS cT by (auto simp add: closedin_def) +qed + +lemma closedin_diff[intro]: assumes oS: "closedin U S" and cT: "openin U T" shows "closedin U (S - T)" +proof- + have "S - T = S \ (topspace U - T)" using closedin_subset[of U S] oS cT + by (auto simp add: topspace_def ) + then show ?thesis using oS cT by (auto simp add: openin_closedin_eq) +qed + +subsection{* Subspace topology. *} + +definition "subtopology U V = topology {S \ V |S. openin U S}" + +lemma istopology_subtopology: "istopology {S \ V |S. openin U S}" (is "istopology ?L") +proof- + have "{} \ ?L" by blast + {fix A B assume A: "A \ ?L" and B: "B \ ?L" + from A B obtain Sa and Sb where Sa: "openin U Sa" "A = Sa \ V" and Sb: "openin U Sb" "B = Sb \ V" by blast + have "A\B = (Sa \ Sb) \ V" "openin U (Sa \ Sb)" using Sa Sb by blast+ + then have "A \ B \ ?L" by blast} + moreover + {fix K assume K: "K \ ?L" + have th0: "?L = (\S. S \ V) ` openin U " + apply (rule set_ext) + apply (simp add: Ball_def image_iff) + by (metis mem_def) + from K[unfolded th0 subset_image_iff] + obtain Sk where Sk: "Sk \ openin U" "K = (\S. S \ V) ` Sk" by blast + have "\K = (\Sk) \ V" using Sk by auto + moreover have "openin U (\ Sk)" using Sk by (auto simp add: subset_eq mem_def) + ultimately have "\K \ ?L" by blast} + ultimately show ?thesis unfolding istopology_def by blast +qed + +lemma openin_subtopology: + "openin (subtopology U V) S \ (\ T. (openin U T) \ (S = T \ V))" + unfolding subtopology_def topology_inverse'[OF istopology_subtopology] + by (auto simp add: Collect_def) + +lemma topspace_subtopology: "topspace(subtopology U V) = topspace U \ V" + by (auto simp add: topspace_def openin_subtopology) + +lemma closedin_subtopology: + "closedin (subtopology U V) S \ (\T. closedin U T \ S = T \ V)" + unfolding closedin_def topspace_subtopology + apply (simp add: openin_subtopology) + apply (rule iffI) + apply clarify + apply (rule_tac x="topspace U - T" in exI) + by auto + +lemma openin_subtopology_refl: "openin (subtopology U V) V \ V \ topspace U" + unfolding openin_subtopology + apply (rule iffI, clarify) + apply (frule openin_subset[of U]) apply blast + apply (rule exI[where x="topspace U"]) + by auto + +lemma subtopology_superset: assumes UV: "topspace U \ V" + shows "subtopology U V = U" +proof- + {fix S + {fix T assume T: "openin U T" "S = T \ V" + from T openin_subset[OF T(1)] UV have eq: "S = T" by blast + have "openin U S" unfolding eq using T by blast} + moreover + {assume S: "openin U S" + hence "\T. openin U T \ S = T \ V" + using openin_subset[OF S] UV by auto} + ultimately have "(\T. openin U T \ S = T \ V) \ openin U S" by blast} + then show ?thesis unfolding topology_eq openin_subtopology by blast +qed + + +lemma subtopology_topspace[simp]: "subtopology U (topspace U) = U" + by (simp add: subtopology_superset) + +lemma subtopology_UNIV[simp]: "subtopology U UNIV = U" + by (simp add: subtopology_superset) + +subsection{* The universal Euclidean versions are what we use most of the time *} + +definition + euclidean :: "'a::topological_space topology" where + "euclidean = topology open" + +lemma open_openin: "open S \ openin euclidean S" + unfolding euclidean_def + apply (rule cong[where x=S and y=S]) + apply (rule topology_inverse[symmetric]) + apply (auto simp add: istopology_def) + by (auto simp add: mem_def subset_eq) + +lemma topspace_euclidean: "topspace euclidean = UNIV" + apply (simp add: topspace_def) + apply (rule set_ext) + by (auto simp add: open_openin[symmetric]) + +lemma topspace_euclidean_subtopology[simp]: "topspace (subtopology euclidean S) = S" + by (simp add: topspace_euclidean topspace_subtopology) + +lemma closed_closedin: "closed S \ closedin euclidean S" + by (simp add: closed_def closedin_def topspace_euclidean open_openin Compl_eq_Diff_UNIV) + +lemma open_subopen: "open S \ (\x\S. \T. open T \ x \ T \ T \ S)" + by (simp add: open_openin openin_subopen[symmetric]) + +subsection{* Open and closed balls. *} + +definition + ball :: "'a::metric_space \ real \ 'a set" where + "ball x e = {y. dist x y < e}" + +definition + cball :: "'a::metric_space \ real \ 'a set" where + "cball x e = {y. dist x y \ e}" + +lemma mem_ball[simp]: "y \ ball x e \ dist x y < e" by (simp add: ball_def) +lemma mem_cball[simp]: "y \ cball x e \ dist x y \ e" by (simp add: cball_def) + +lemma mem_ball_0 [simp]: + fixes x :: "'a::real_normed_vector" + shows "x \ ball 0 e \ norm x < e" + by (simp add: dist_norm) + +lemma mem_cball_0 [simp]: + fixes x :: "'a::real_normed_vector" + shows "x \ cball 0 e \ norm x \ e" + by (simp add: dist_norm) + +lemma centre_in_cball[simp]: "x \ cball x e \ 0\ e" by simp +lemma ball_subset_cball[simp,intro]: "ball x e \ cball x e" by (simp add: subset_eq) +lemma subset_ball[intro]: "d <= e ==> ball x d \ ball x e" by (simp add: subset_eq) +lemma subset_cball[intro]: "d <= e ==> cball x d \ cball x e" by (simp add: subset_eq) +lemma ball_max_Un: "ball a (max r s) = ball a r \ ball a s" + by (simp add: expand_set_eq) arith + +lemma ball_min_Int: "ball a (min r s) = ball a r \ ball a s" + by (simp add: expand_set_eq) + +subsection{* Topological properties of open balls *} + +lemma diff_less_iff: "(a::real) - b > 0 \ a > b" + "(a::real) - b < 0 \ a < b" + "a - b < c \ a < c +b" "a - b > c \ a > c +b" by arith+ +lemma diff_le_iff: "(a::real) - b \ 0 \ a \ b" "(a::real) - b \ 0 \ a \ b" + "a - b \ c \ a \ c +b" "a - b \ c \ a \ c +b" by arith+ + +lemma open_ball[intro, simp]: "open (ball x e)" + unfolding open_dist ball_def Collect_def Ball_def mem_def + unfolding dist_commute + apply clarify + apply (rule_tac x="e - dist xa x" in exI) + using dist_triangle_alt[where z=x] + apply (clarsimp simp add: diff_less_iff) + apply atomize + apply (erule_tac x="y" in allE) + apply (erule_tac x="xa" in allE) + by arith + +lemma centre_in_ball[simp]: "x \ ball x e \ e > 0" by (metis mem_ball dist_self) +lemma open_contains_ball: "open S \ (\x\S. \e>0. ball x e \ S)" + unfolding open_dist subset_eq mem_ball Ball_def dist_commute .. + +lemma open_contains_ball_eq: "open S \ \x. x\S \ (\e>0. ball x e \ S)" + by (metis open_contains_ball subset_eq centre_in_ball) + +lemma ball_eq_empty[simp]: "ball x e = {} \ e \ 0" + unfolding mem_ball expand_set_eq + apply (simp add: not_less) + by (metis zero_le_dist order_trans dist_self) + +lemma ball_empty[intro]: "e \ 0 ==> ball x e = {}" by simp + +subsection{* Basic "localization" results are handy for connectedness. *} + +lemma openin_open: "openin (subtopology euclidean U) S \ (\T. open T \ (S = U \ T))" + by (auto simp add: openin_subtopology open_openin[symmetric]) + +lemma openin_open_Int[intro]: "open S \ openin (subtopology euclidean U) (U \ S)" + by (auto simp add: openin_open) + +lemma open_openin_trans[trans]: + "open S \ open T \ T \ S \ openin (subtopology euclidean S) T" + by (metis Int_absorb1 openin_open_Int) + +lemma open_subset: "S \ T \ open S \ openin (subtopology euclidean T) S" + by (auto simp add: openin_open) + +lemma closedin_closed: "closedin (subtopology euclidean U) S \ (\T. closed T \ S = U \ T)" + by (simp add: closedin_subtopology closed_closedin Int_ac) + +lemma closedin_closed_Int: "closed S ==> closedin (subtopology euclidean U) (U \ S)" + by (metis closedin_closed) + +lemma closed_closedin_trans: "closed S \ closed T \ T \ S \ closedin (subtopology euclidean S) T" + apply (subgoal_tac "S \ T = T" ) + apply auto + apply (frule closedin_closed_Int[of T S]) + by simp + +lemma closed_subset: "S \ T \ closed S \ closedin (subtopology euclidean T) S" + by (auto simp add: closedin_closed) + +lemma openin_euclidean_subtopology_iff: + fixes S U :: "'a::metric_space set" + shows "openin (subtopology euclidean U) S + \ S \ U \ (\x\S. \e>0. \x'\U. dist x' x < e \ x'\ S)" (is "?lhs \ ?rhs") +proof- + {assume ?lhs hence ?rhs unfolding openin_subtopology open_openin[symmetric] + by (simp add: open_dist) blast} + moreover + {assume SU: "S \ U" and H: "\x. x \ S \ \e>0. \x'\U. dist x' x < e \ x' \ S" + from H obtain d where d: "\x . x\ S \ d x > 0 \ (\x' \ U. dist x' x < d x \ x' \ S)" + by metis + let ?T = "\{B. \x\S. B = ball x (d x)}" + have oT: "open ?T" by auto + { fix x assume "x\S" + hence "x \ \{B. \x\S. B = ball x (d x)}" + apply simp apply(rule_tac x="ball x(d x)" in exI) apply auto + by (rule d [THEN conjunct1]) + hence "x\ ?T \ U" using SU and `x\S` by auto } + moreover + { fix y assume "y\?T" + then obtain B where "y\B" "B\{B. \x\S. B = ball x (d x)}" by auto + then obtain x where "x\S" and x:"y \ ball x (d x)" by auto + assume "y\U" + hence "y\S" using d[OF `x\S`] and x by(auto simp add: dist_commute) } + ultimately have "S = ?T \ U" by blast + with oT have ?lhs unfolding openin_subtopology open_openin[symmetric] by blast} + ultimately show ?thesis by blast +qed + +text{* These "transitivity" results are handy too. *} + +lemma openin_trans[trans]: "openin (subtopology euclidean T) S \ openin (subtopology euclidean U) T + \ openin (subtopology euclidean U) S" + unfolding open_openin openin_open by blast + +lemma openin_open_trans: "openin (subtopology euclidean T) S \ open T \ open S" + by (auto simp add: openin_open intro: openin_trans) + +lemma closedin_trans[trans]: + "closedin (subtopology euclidean T) S \ + closedin (subtopology euclidean U) T + ==> closedin (subtopology euclidean U) S" + by (auto simp add: closedin_closed closed_closedin closed_Inter Int_assoc) + +lemma closedin_closed_trans: "closedin (subtopology euclidean T) S \ closed T \ closed S" + by (auto simp add: closedin_closed intro: closedin_trans) + +subsection{* Connectedness *} + +definition "connected S \ + ~(\e1 e2. open e1 \ open e2 \ S \ (e1 \ e2) \ (e1 \ e2 \ S = {}) + \ ~(e1 \ S = {}) \ ~(e2 \ S = {}))" + +lemma connected_local: + "connected S \ ~(\e1 e2. + openin (subtopology euclidean S) e1 \ + openin (subtopology euclidean S) e2 \ + S \ e1 \ e2 \ + e1 \ e2 = {} \ + ~(e1 = {}) \ + ~(e2 = {}))" +unfolding connected_def openin_open by (safe, blast+) + +lemma exists_diff: "(\S. P(UNIV - S)) \ (\S. P S)" (is "?lhs \ ?rhs") +proof- + + {assume "?lhs" hence ?rhs by blast } + moreover + {fix S assume H: "P S" + have "S = UNIV - (UNIV - S)" by auto + with H have "P (UNIV - (UNIV - S))" by metis } + ultimately show ?thesis by metis +qed + +lemma connected_clopen: "connected S \ + (\T. openin (subtopology euclidean S) T \ + closedin (subtopology euclidean S) T \ T = {} \ T = S)" (is "?lhs \ ?rhs") +proof- + have " \ connected S \ (\e1 e2. open e1 \ open (UNIV - e2) \ S \ e1 \ (UNIV - e2) \ e1 \ (UNIV - e2) \ S = {} \ e1 \ S \ {} \ (UNIV - e2) \ S \ {})" + unfolding connected_def openin_open closedin_closed + apply (subst exists_diff) by blast + hence th0: "connected S \ \ (\e2 e1. closed e2 \ open e1 \ S \ e1 \ (UNIV - e2) \ e1 \ (UNIV - e2) \ S = {} \ e1 \ S \ {} \ (UNIV - e2) \ S \ {})" + (is " _ \ \ (\e2 e1. ?P e2 e1)") apply (simp add: closed_def Compl_eq_Diff_UNIV) by metis + + have th1: "?rhs \ \ (\t' t. closed t'\t = S\t' \ t\{} \ t\S \ (\t'. open t' \ t = S \ t'))" + (is "_ \ \ (\t' t. ?Q t' t)") + unfolding connected_def openin_open closedin_closed by auto + {fix e2 + {fix e1 have "?P e2 e1 \ (\t. closed e2 \ t = S\e2 \ open e1 \ t = S\e1 \ t\{} \ t\S)" + by auto} + then have "(\e1. ?P e2 e1) \ (\t. ?Q e2 t)" by metis} + then have "\e2. (\e1. ?P e2 e1) \ (\t. ?Q e2 t)" by blast + then show ?thesis unfolding th0 th1 by simp +qed + +lemma connected_empty[simp, intro]: "connected {}" + by (simp add: connected_def) + +subsection{* Hausdorff and other separation properties *} + +class t0_space = + assumes t0_space: "x \ y \ \U. open U \ \ (x \ U \ y \ U)" + +class t1_space = + assumes t1_space: "x \ y \ \U V. open U \ open V \ x \ U \ y \ U \ x \ V \ y \ V" +begin + +subclass t0_space +proof +qed (fast dest: t1_space) + +end + +text {* T2 spaces are also known as Hausdorff spaces. *} + +class t2_space = + assumes hausdorff: "x \ y \ \U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" +begin + +subclass t1_space +proof +qed (fast dest: hausdorff) + +end + +instance metric_space \ t2_space +proof + fix x y :: "'a::metric_space" + assume xy: "x \ y" + let ?U = "ball x (dist x y / 2)" + let ?V = "ball y (dist x y / 2)" + have th0: "\d x y z. (d x z :: real) <= d x y + d y z \ d y z = d z y + ==> ~(d x y * 2 < d x z \ d z y * 2 < d x z)" by arith + have "open ?U \ open ?V \ x \ ?U \ y \ ?V \ ?U \ ?V = {}" + using dist_pos_lt[OF xy] th0[of dist,OF dist_triangle dist_commute] + by (auto simp add: expand_set_eq) + then show "\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" + by blast +qed + +lemma separation_t2: + fixes x y :: "'a::t2_space" + shows "x \ y \ (\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {})" + using hausdorff[of x y] by blast + +lemma separation_t1: + fixes x y :: "'a::t1_space" + shows "x \ y \ (\U V. open U \ open V \ x \U \ y\ U \ x\V \ y\V)" + using t1_space[of x y] by blast + +lemma separation_t0: + fixes x y :: "'a::t0_space" + shows "x \ y \ (\U. open U \ ~(x\U \ y\U))" + using t0_space[of x y] by blast + +subsection{* Limit points *} + +definition + islimpt:: "'a::topological_space \ 'a set \ bool" + (infixr "islimpt" 60) where + "x islimpt S \ (\T. x\T \ open T \ (\y\S. y\T \ y\x))" + +lemma islimptI: + assumes "\T. x \ T \ open T \ \y\S. y \ T \ y \ x" + shows "x islimpt S" + using assms unfolding islimpt_def by auto + +lemma islimptE: + assumes "x islimpt S" and "x \ T" and "open T" + obtains y where "y \ S" and "y \ T" and "y \ x" + using assms unfolding islimpt_def by auto + +lemma islimpt_subset: "x islimpt S \ S \ T ==> x islimpt T" by (auto simp add: islimpt_def) + +lemma islimpt_approachable: + fixes x :: "'a::metric_space" + shows "x islimpt S \ (\e>0. \x'\S. x' \ x \ dist x' x < e)" + unfolding islimpt_def + apply auto + apply(erule_tac x="ball x e" in allE) + apply auto + apply(rule_tac x=y in bexI) + apply (auto simp add: dist_commute) + apply (simp add: open_dist, drule (1) bspec) + apply (clarify, drule spec, drule (1) mp, auto) + done + +lemma islimpt_approachable_le: + fixes x :: "'a::metric_space" + shows "x islimpt S \ (\e>0. \x'\ S. x' \ x \ dist x' x <= e)" + unfolding islimpt_approachable + using approachable_lt_le[where f="\x'. dist x' x" and P="\x'. \ (x'\S \ x'\x)"] + by metis (* FIXME: VERY slow! *) + +class perfect_space = + (* FIXME: perfect_space should inherit from topological_space *) + assumes islimpt_UNIV [simp, intro]: "(x::'a::metric_space) islimpt UNIV" + +lemma perfect_choose_dist: + fixes x :: "'a::perfect_space" + shows "0 < r \ \a. a \ x \ dist a x < r" +using islimpt_UNIV [of x] +by (simp add: islimpt_approachable) + +instance real :: perfect_space +apply default +apply (rule islimpt_approachable [THEN iffD2]) +apply (clarify, rule_tac x="x + e/2" in bexI) +apply (auto simp add: dist_norm) +done + +instance "^" :: (perfect_space, finite) perfect_space +proof + fix x :: "'a ^ 'b" + { + fix e :: real assume "0 < e" + def a \ "x $ undefined" + have "a islimpt UNIV" by (rule islimpt_UNIV) + with `0 < e` obtain b where "b \ a" and "dist b a < e" + unfolding islimpt_approachable by auto + def y \ "Cart_lambda ((Cart_nth x)(undefined := b))" + from `b \ a` have "y \ x" + unfolding a_def y_def by (simp add: Cart_eq) + from `dist b a < e` have "dist y x < e" + unfolding dist_vector_def a_def y_def + apply simp + apply (rule le_less_trans [OF setL2_le_setsum [OF zero_le_dist]]) + apply (subst setsum_diff1' [where a=undefined], simp, simp, simp) + done + from `y \ x` and `dist y x < e` + have "\y\UNIV. y \ x \ dist y x < e" by auto + } + then show "x islimpt UNIV" unfolding islimpt_approachable by blast +qed + +lemma closed_limpt: "closed S \ (\x. x islimpt S \ x \ S)" + unfolding closed_def + apply (subst open_subopen) + apply (simp add: islimpt_def subset_eq Compl_eq_Diff_UNIV) + by (metis DiffE DiffI UNIV_I insertCI insert_absorb mem_def) + +lemma islimpt_EMPTY[simp]: "\ x islimpt {}" + unfolding islimpt_def by auto + +lemma closed_positive_orthant: "closed {x::real^'n::finite. \i. 0 \x$i}" +proof- + let ?U = "UNIV :: 'n set" + let ?O = "{x::real^'n. \i. x$i\0}" + {fix x:: "real^'n" and i::'n assume H: "\e>0. \x'\?O. x' \ x \ dist x' x < e" + and xi: "x$i < 0" + from xi have th0: "-x$i > 0" by arith + from H[rule_format, OF th0] obtain x' where x': "x' \?O" "x' \ x" "dist x' x < -x $ i" by blast + have th:" \b a (x::real). abs x <= b \ b <= a ==> ~(a + x < 0)" by arith + have th': "\x (y::real). x < 0 \ 0 <= y ==> abs x <= abs (y - x)" by arith + have th1: "\x$i\ \ \(x' - x)$i\" using x'(1) xi + apply (simp only: vector_component) + by (rule th') auto + have th2: "\dist x x'\ \ \(x' - x)$i\" using component_le_norm[of "x'-x" i] + apply (simp add: dist_norm) by norm + from th[OF th1 th2] x'(3) have False by (simp add: dist_commute) } + then show ?thesis unfolding closed_limpt islimpt_approachable + unfolding not_le[symmetric] by blast +qed + +lemma finite_set_avoid: + fixes a :: "'a::metric_space" + assumes fS: "finite S" shows "\d>0. \x\S. x \ a \ d <= dist a x" +proof(induct rule: finite_induct[OF fS]) + case 1 thus ?case apply auto by ferrack +next + case (2 x F) + from 2 obtain d where d: "d >0" "\x\F. x\a \ d \ dist a x" by blast + {assume "x = a" hence ?case using d by auto } + moreover + {assume xa: "x\a" + let ?d = "min d (dist a x)" + have dp: "?d > 0" using xa d(1) using dist_nz by auto + from d have d': "\x\F. x\a \ ?d \ dist a x" by auto + with dp xa have ?case by(auto intro!: exI[where x="?d"]) } + ultimately show ?case by blast +qed + +lemma islimpt_finite: + fixes S :: "'a::metric_space set" + assumes fS: "finite S" shows "\ a islimpt S" + unfolding islimpt_approachable + using finite_set_avoid[OF fS, of a] by (metis dist_commute not_le) + +lemma islimpt_Un: "x islimpt (S \ T) \ x islimpt S \ x islimpt T" + apply (rule iffI) + defer + apply (metis Un_upper1 Un_upper2 islimpt_subset) + unfolding islimpt_def + apply (rule ccontr, clarsimp, rename_tac A B) + apply (drule_tac x="A \ B" in spec) + apply (auto simp add: open_Int) + done + +lemma discrete_imp_closed: + fixes S :: "'a::metric_space set" + assumes e: "0 < e" and d: "\x \ S. \y \ S. dist y x < e \ y = x" + shows "closed S" +proof- + {fix x assume C: "\e>0. \x'\S. x' \ x \ dist x' x < e" + from e have e2: "e/2 > 0" by arith + from C[rule_format, OF e2] obtain y where y: "y \ S" "y\x" "dist y x < e/2" by blast + let ?m = "min (e/2) (dist x y) " + from e2 y(2) have mp: "?m > 0" by (simp add: dist_nz[THEN sym]) + from C[rule_format, OF mp] obtain z where z: "z \ S" "z\x" "dist z x < ?m" by blast + have th: "dist z y < e" using z y + by (intro dist_triangle_lt [where z=x], simp) + from d[rule_format, OF y(1) z(1) th] y z + have False by (auto simp add: dist_commute)} + then show ?thesis by (metis islimpt_approachable closed_limpt [where 'a='a]) +qed + +subsection{* Interior of a Set *} +definition "interior S = {x. \T. open T \ x \ T \ T \ S}" + +lemma interior_eq: "interior S = S \ open S" + apply (simp add: expand_set_eq interior_def) + apply (subst (2) open_subopen) by (safe, blast+) + +lemma interior_open: "open S ==> (interior S = S)" by (metis interior_eq) + +lemma interior_empty[simp]: "interior {} = {}" by (simp add: interior_def) + +lemma open_interior[simp, intro]: "open(interior S)" + apply (simp add: interior_def) + apply (subst open_subopen) by blast + +lemma interior_interior[simp]: "interior(interior S) = interior S" by (metis interior_eq open_interior) +lemma interior_subset: "interior S \ S" by (auto simp add: interior_def) +lemma subset_interior: "S \ T ==> (interior S) \ (interior T)" by (auto simp add: interior_def) +lemma interior_maximal: "T \ S \ open T ==> T \ (interior S)" by (auto simp add: interior_def) +lemma interior_unique: "T \ S \ open T \ (\T'. T' \ S \ open T' \ T' \ T) \ interior S = T" + by (metis equalityI interior_maximal interior_subset open_interior) +lemma mem_interior: "x \ interior S \ (\e. 0 < e \ ball x e \ S)" + apply (simp add: interior_def) + by (metis open_contains_ball centre_in_ball open_ball subset_trans) + +lemma open_subset_interior: "open S ==> S \ interior T \ S \ T" + by (metis interior_maximal interior_subset subset_trans) + +lemma interior_inter[simp]: "interior(S \ T) = interior S \ interior T" + apply (rule equalityI, simp) + apply (metis Int_lower1 Int_lower2 subset_interior) + by (metis Int_mono interior_subset open_Int open_interior open_subset_interior) + +lemma interior_limit_point [intro]: + fixes x :: "'a::perfect_space" + assumes x: "x \ interior S" shows "x islimpt S" +proof- + from x obtain e where e: "e>0" "\x'. dist x x' < e \ x' \ S" + unfolding mem_interior subset_eq Ball_def mem_ball by blast + { + fix d::real assume d: "d>0" + let ?m = "min d e" + have mde2: "0 < ?m" using e(1) d(1) by simp + from perfect_choose_dist [OF mde2, of x] + obtain y where "y \ x" and "dist y x < ?m" by blast + then have "dist y x < e" "dist y x < d" by simp_all + from `dist y x < e` e(2) have "y \ S" by (simp add: dist_commute) + have "\x'\S. x'\ x \ dist x' x < d" + using `y \ S` `y \ x` `dist y x < d` by fast + } + then show ?thesis unfolding islimpt_approachable by blast +qed + +lemma interior_closed_Un_empty_interior: + assumes cS: "closed S" and iT: "interior T = {}" + shows "interior(S \ T) = interior S" +proof + show "interior S \ interior (S\T)" + by (rule subset_interior, blast) +next + show "interior (S \ T) \ interior S" + proof + fix x assume "x \ interior (S \ T)" + then obtain R where "open R" "x \ R" "R \ S \ T" + unfolding interior_def by fast + show "x \ interior S" + proof (rule ccontr) + assume "x \ interior S" + with `x \ R` `open R` obtain y where "y \ R - S" + unfolding interior_def expand_set_eq by fast + from `open R` `closed S` have "open (R - S)" by (rule open_Diff) + from `R \ S \ T` have "R - S \ T" by fast + from `y \ R - S` `open (R - S)` `R - S \ T` `interior T = {}` + show "False" unfolding interior_def by fast + qed + qed +qed + + +subsection{* Closure of a Set *} + +definition "closure S = S \ {x | x. x islimpt S}" + +lemma closure_interior: "closure S = UNIV - interior (UNIV - S)" +proof- + { fix x + have "x\UNIV - interior (UNIV - S) \ x \ closure S" (is "?lhs = ?rhs") + proof + let ?exT = "\ y. (\T. open T \ y \ T \ T \ UNIV - S)" + assume "?lhs" + hence *:"\ ?exT x" + unfolding interior_def + by simp + { assume "\ ?rhs" + hence False using * + unfolding closure_def islimpt_def + by blast + } + thus "?rhs" + by blast + next + assume "?rhs" thus "?lhs" + unfolding closure_def interior_def islimpt_def + by blast + qed + } + thus ?thesis + by blast +qed + +lemma interior_closure: "interior S = UNIV - (closure (UNIV - S))" +proof- + { fix x + have "x \ interior S \ x \ UNIV - (closure (UNIV - S))" + unfolding interior_def closure_def islimpt_def + by blast (* FIXME: VERY slow! *) + } + thus ?thesis + by blast +qed + +lemma closed_closure[simp, intro]: "closed (closure S)" +proof- + have "closed (UNIV - interior (UNIV -S))" by blast + thus ?thesis using closure_interior[of S] by simp +qed + +lemma closure_hull: "closure S = closed hull S" +proof- + have "S \ closure S" + unfolding closure_def + by blast + moreover + have "closed (closure S)" + using closed_closure[of S] + by assumption + moreover + { fix t + assume *:"S \ t" "closed t" + { fix x + assume "x islimpt S" + hence "x islimpt t" using *(1) + using islimpt_subset[of x, of S, of t] + by blast + } + with * have "closure S \ t" + unfolding closure_def + using closed_limpt[of t] + by auto + } + ultimately show ?thesis + using hull_unique[of S, of "closure S", of closed] + unfolding mem_def + by simp +qed + +lemma closure_eq: "closure S = S \ closed S" + unfolding closure_hull + using hull_eq[of closed, unfolded mem_def, OF closed_Inter, of S] + by (metis mem_def subset_eq) + +lemma closure_closed[simp]: "closed S \ closure S = S" + using closure_eq[of S] + by simp + +lemma closure_closure[simp]: "closure (closure S) = closure S" + unfolding closure_hull + using hull_hull[of closed S] + by assumption + +lemma closure_subset: "S \ closure S" + unfolding closure_hull + using hull_subset[of S closed] + by assumption + +lemma subset_closure: "S \ T \ closure S \ closure T" + unfolding closure_hull + using hull_mono[of S T closed] + by assumption + +lemma closure_minimal: "S \ T \ closed T \ closure S \ T" + using hull_minimal[of S T closed] + unfolding closure_hull mem_def + by simp + +lemma closure_unique: "S \ T \ closed T \ (\ T'. S \ T' \ closed T' \ T \ T') \ closure S = T" + using hull_unique[of S T closed] + unfolding closure_hull mem_def + by simp + +lemma closure_empty[simp]: "closure {} = {}" + using closed_empty closure_closed[of "{}"] + by simp + +lemma closure_univ[simp]: "closure UNIV = UNIV" + using closure_closed[of UNIV] + by simp + +lemma closure_eq_empty: "closure S = {} \ S = {}" + using closure_empty closure_subset[of S] + by blast + +lemma closure_subset_eq: "closure S \ S \ closed S" + using closure_eq[of S] closure_subset[of S] + by simp + +lemma open_inter_closure_eq_empty: + "open S \ (S \ closure T) = {} \ S \ T = {}" + using open_subset_interior[of S "UNIV - T"] + using interior_subset[of "UNIV - T"] + unfolding closure_interior + by auto + +lemma open_inter_closure_subset: + "open S \ (S \ (closure T)) \ closure(S \ T)" +proof + fix x + assume as: "open S" "x \ S \ closure T" + { assume *:"x islimpt T" + have "x islimpt (S \ T)" + proof (rule islimptI) + fix A + assume "x \ A" "open A" + with as have "x \ A \ S" "open (A \ S)" + by (simp_all add: open_Int) + with * obtain y where "y \ T" "y \ A \ S" "y \ x" + by (rule islimptE) + hence "y \ S \ T" "y \ A \ y \ x" + by simp_all + thus "\y\(S \ T). y \ A \ y \ x" .. + qed + } + then show "x \ closure (S \ T)" using as + unfolding closure_def + by blast +qed + +lemma closure_complement: "closure(UNIV - S) = UNIV - interior(S)" +proof- + have "S = UNIV - (UNIV - S)" + by auto + thus ?thesis + unfolding closure_interior + by auto +qed + +lemma interior_complement: "interior(UNIV - S) = UNIV - closure(S)" + unfolding closure_interior + by blast + +subsection{* Frontier (aka boundary) *} + +definition "frontier S = closure S - interior S" + +lemma frontier_closed: "closed(frontier S)" + by (simp add: frontier_def closed_Diff) + +lemma frontier_closures: "frontier S = (closure S) \ (closure(UNIV - S))" + by (auto simp add: frontier_def interior_closure) + +lemma frontier_straddle: + fixes a :: "'a::metric_space" + shows "a \ frontier S \ (\e>0. (\x\S. dist a x < e) \ (\x. x \ S \ dist a x < e))" (is "?lhs \ ?rhs") +proof + assume "?lhs" + { fix e::real + assume "e > 0" + let ?rhse = "(\x\S. dist a x < e) \ (\x. x \ S \ dist a x < e)" + { assume "a\S" + have "\x\S. dist a x < e" using `e>0` `a\S` by(rule_tac x=a in bexI) auto + moreover have "\x. x \ S \ dist a x < e" using `?lhs` `a\S` + unfolding frontier_closures closure_def islimpt_def using `e>0` + by (auto, erule_tac x="ball a e" in allE, auto) + ultimately have ?rhse by auto + } + moreover + { assume "a\S" + hence ?rhse using `?lhs` + unfolding frontier_closures closure_def islimpt_def + using open_ball[of a e] `e > 0` + by (auto, erule_tac x = "ball a e" in allE, auto) (* FIXME: VERY slow! *) + } + ultimately have ?rhse by auto + } + thus ?rhs by auto +next + assume ?rhs + moreover + { fix T assume "a\S" and + as:"\e>0. (\x\S. dist a x < e) \ (\x. x \ S \ dist a x < e)" "a \ S" "a \ T" "open T" + from `open T` `a \ T` have "\e>0. ball a e \ T" unfolding open_contains_ball[of T] by auto + then obtain e where "e>0" "ball a e \ T" by auto + then obtain y where y:"y\S" "dist a y < e" using as(1) by auto + have "\y\S. y \ T \ y \ a" + using `dist a y < e` `ball a e \ T` unfolding ball_def using `y\S` `a\S` by auto + } + hence "a \ closure S" unfolding closure_def islimpt_def using `?rhs` by auto + moreover + { fix T assume "a \ T" "open T" "a\S" + then obtain e where "e>0" and balle: "ball a e \ T" unfolding open_contains_ball using `?rhs` by auto + obtain x where "x \ S" "dist a x < e" using `?rhs` using `e>0` by auto + hence "\y\UNIV - S. y \ T \ y \ a" using balle `a\S` unfolding ball_def by (rule_tac x=x in bexI)auto + } + hence "a islimpt (UNIV - S) \ a\S" unfolding islimpt_def by auto + ultimately show ?lhs unfolding frontier_closures using closure_def[of "UNIV - S"] by auto +qed + +lemma frontier_subset_closed: "closed S \ frontier S \ S" + by (metis frontier_def closure_closed Diff_subset) + +lemma frontier_empty: "frontier {} = {}" + by (simp add: frontier_def closure_empty) + +lemma frontier_subset_eq: "frontier S \ S \ closed S" +proof- + { assume "frontier S \ S" + hence "closure S \ S" using interior_subset unfolding frontier_def by auto + hence "closed S" using closure_subset_eq by auto + } + thus ?thesis using frontier_subset_closed[of S] by auto +qed + +lemma frontier_complement: "frontier(UNIV - S) = frontier S" + by (auto simp add: frontier_def closure_complement interior_complement) + +lemma frontier_disjoint_eq: "frontier S \ S = {} \ open S" + using frontier_complement frontier_subset_eq[of "UNIV - S"] + unfolding open_closed Compl_eq_Diff_UNIV by auto + +subsection{* Common nets and The "within" modifier for nets. *} + +definition + at_infinity :: "'a::real_normed_vector net" where + "at_infinity = Abs_net (range (\r. {x. r \ norm x}))" + +definition + indirection :: "'a::real_normed_vector \ 'a \ 'a net" (infixr "indirection" 70) where + "a indirection v = (at a) within {b. \c\0. b - a = scaleR c v}" + +text{* Prove That They are all nets. *} + +lemma Rep_net_at_infinity: + "Rep_net at_infinity = range (\r. {x. r \ norm x})" +unfolding at_infinity_def +apply (rule Abs_net_inverse') +apply (rule image_nonempty, simp) +apply (clarsimp, rename_tac r s) +apply (rule_tac x="max r s" in exI, auto) +done + +lemma within_UNIV: "net within UNIV = net" + by (simp add: Rep_net_inject [symmetric] Rep_net_within) + +subsection{* Identify Trivial limits, where we can't approach arbitrarily closely. *} + +definition + trivial_limit :: "'a net \ bool" where + "trivial_limit net \ {} \ Rep_net net" + +lemma trivial_limit_within: + shows "trivial_limit (at a within S) \ \ a islimpt S" +proof + assume "trivial_limit (at a within S)" + thus "\ a islimpt S" + unfolding trivial_limit_def + unfolding Rep_net_within Rep_net_at + unfolding islimpt_def + apply (clarsimp simp add: expand_set_eq) + apply (rename_tac T, rule_tac x=T in exI) + apply (clarsimp, drule_tac x=y in spec, simp) + done +next + assume "\ a islimpt S" + thus "trivial_limit (at a within S)" + unfolding trivial_limit_def + unfolding Rep_net_within Rep_net_at + unfolding islimpt_def + apply (clarsimp simp add: image_image) + apply (rule_tac x=T in image_eqI) + apply (auto simp add: expand_set_eq) + done +qed + +lemma trivial_limit_at_iff: "trivial_limit (at a) \ \ a islimpt UNIV" + using trivial_limit_within [of a UNIV] + by (simp add: within_UNIV) + +lemma trivial_limit_at: + fixes a :: "'a::perfect_space" + shows "\ trivial_limit (at a)" + by (simp add: trivial_limit_at_iff) + +lemma trivial_limit_at_infinity: + "\ trivial_limit (at_infinity :: ('a::{real_normed_vector,zero_neq_one}) net)" + (* FIXME: find a more appropriate type class *) + 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_sgn) + done + +lemma trivial_limit_sequentially: "\ trivial_limit sequentially" + by (auto simp add: trivial_limit_def Rep_net_sequentially) + +subsection{* Some property holds "sufficiently close" to the limit point. *} + +lemma eventually_at: (* FIXME: this replaces Limits.eventually_at *) + "eventually P (at a) \ (\d>0. \x. 0 < dist x a \ dist x a < d \ P x)" +unfolding eventually_at dist_nz by auto + +lemma eventually_at_infinity: + "eventually P at_infinity \ (\b. \x. norm x >= b \ P x)" +unfolding eventually_def Rep_net_at_infinity by auto + +lemma eventually_within: "eventually P (at a within S) \ + (\d>0. \x\S. 0 < dist x a \ dist x a < d \ P x)" +unfolding eventually_within eventually_at dist_nz by auto + +lemma eventually_within_le: "eventually P (at a within S) \ + (\d>0. \x\S. 0 < dist x a \ dist x a <= d \ P x)" (is "?lhs = ?rhs") +unfolding eventually_within +apply safe +apply (rule_tac x="d/2" in exI, simp) +apply (rule_tac x="d" in exI, simp) +done + +lemma eventually_happens: "eventually P net ==> trivial_limit net \ (\x. P x)" + unfolding eventually_def trivial_limit_def + using Rep_net_nonempty [of net] by auto + +lemma always_eventually: "(\x. P x) ==> eventually P net" + unfolding eventually_def trivial_limit_def + using Rep_net_nonempty [of net] by auto + +lemma trivial_limit_eventually: "trivial_limit net \ eventually P net" + unfolding trivial_limit_def eventually_def by auto + +lemma eventually_False: "eventually (\x. False) net \ trivial_limit net" + unfolding trivial_limit_def eventually_def by auto + +lemma trivial_limit_eq: "trivial_limit net \ (\P. eventually P net)" + apply (safe elim!: trivial_limit_eventually) + apply (simp add: eventually_False [symmetric]) + done + +text{* Combining theorems for "eventually" *} + +lemma eventually_conjI: + "\eventually (\x. P x) net; eventually (\x. Q x) net\ + \ eventually (\x. P x \ Q x) net" +by (rule eventually_conj) + +lemma eventually_rev_mono: + "eventually P net \ (\x. P x \ Q x) \ eventually Q net" +using eventually_mono [of P Q] by fast + +lemma eventually_and: " eventually (\x. P x \ Q x) net \ eventually P net \ eventually Q net" + by (auto intro!: eventually_conjI elim: eventually_rev_mono) + +lemma eventually_false: "eventually (\x. False) net \ trivial_limit net" + by (auto simp add: eventually_False) + +lemma not_eventually: "(\x. \ P x ) \ ~(trivial_limit net) ==> ~(eventually (\x. P x) net)" + by (simp add: eventually_False) + +subsection{* Limits, defined as vacuously true when the limit is trivial. *} + + text{* Notation Lim to avoid collition with lim defined in analysis *} +definition + Lim :: "'a net \ ('a \ 'b::t2_space) \ 'b" where + "Lim net f = (THE l. (f ---> l) net)" + +lemma Lim: + "(f ---> l) net \ + trivial_limit net \ + (\e>0. eventually (\x. dist (f x) l < e) net)" + unfolding tendsto_iff trivial_limit_eq by auto + + +text{* Show that they yield usual definitions in the various cases. *} + +lemma Lim_within_le: "(f ---> l)(at a within S) \ + (\e>0. \d>0. \x\S. 0 < dist x a \ dist x a <= d \ dist (f x) l < e)" + by (auto simp add: tendsto_iff eventually_within_le) + +lemma Lim_within: "(f ---> l) (at a within S) \ + (\e >0. \d>0. \x \ S. 0 < dist x a \ dist x a < d \ dist (f x) l < e)" + by (auto simp add: tendsto_iff eventually_within) + +lemma Lim_at: "(f ---> l) (at a) \ + (\e >0. \d>0. \x. 0 < dist x a \ dist x a < d \ dist (f x) l < e)" + by (auto simp add: tendsto_iff eventually_at) + +lemma Lim_at_iff_LIM: "(f ---> l) (at a) \ f -- a --> l" + unfolding Lim_at LIM_def by (simp only: zero_less_dist_iff) + +lemma Lim_at_infinity: + "(f ---> l) at_infinity \ (\e>0. \b. \x. norm x >= b \ dist (f x) l < e)" + by (auto simp add: tendsto_iff eventually_at_infinity) + +lemma Lim_sequentially: + "(S ---> l) sequentially \ + (\e>0. \N. \n\N. dist (S n) l < e)" + by (auto simp add: tendsto_iff eventually_sequentially) + +lemma Lim_sequentially_iff_LIMSEQ: "(S ---> l) sequentially \ S ----> l" + unfolding Lim_sequentially LIMSEQ_def .. + +lemma Lim_eventually: "eventually (\x. f x = l) net \ (f ---> l) net" + by (rule topological_tendstoI, auto elim: eventually_rev_mono) + +text{* The expected monotonicity property. *} + +lemma Lim_within_empty: "(f ---> l) (net within {})" + unfolding tendsto_def Limits.eventually_within by simp + +lemma Lim_within_subset: "(f ---> l) (net within S) \ T \ S \ (f ---> l) (net within T)" + unfolding tendsto_def Limits.eventually_within + by (auto elim!: eventually_elim1) + +lemma Lim_Un: assumes "(f ---> l) (net within S)" "(f ---> l) (net within T)" + shows "(f ---> l) (net within (S \ T))" + using assms unfolding tendsto_def Limits.eventually_within + apply clarify + apply (drule spec, drule (1) mp, drule (1) mp) + apply (drule spec, drule (1) mp, drule (1) mp) + apply (auto elim: eventually_elim2) + done + +lemma Lim_Un_univ: + "(f ---> l) (net within S) \ (f ---> l) (net within T) \ S \ T = UNIV + ==> (f ---> l) net" + by (metis Lim_Un within_UNIV) + +text{* Interrelations between restricted and unrestricted limits. *} + +lemma Lim_at_within: "(f ---> l) net ==> (f ---> l)(net within S)" + (* FIXME: rename *) + unfolding tendsto_def Limits.eventually_within + apply (clarify, drule spec, drule (1) mp, drule (1) mp) + by (auto elim!: eventually_elim1) + +lemma Lim_within_open: + fixes f :: "'a::topological_space \ 'b::topological_space" + assumes"a \ S" "open S" + shows "(f ---> l)(at a within S) \ (f ---> l)(at a)" (is "?lhs \ ?rhs") +proof + assume ?lhs + { fix A assume "open A" "l \ A" + with `?lhs` have "eventually (\x. f x \ A) (at a within S)" + by (rule topological_tendstoD) + hence "eventually (\x. x \ S \ f x \ A) (at a)" + unfolding Limits.eventually_within . + then obtain T where "open T" "a \ T" "\x\T. x \ a \ x \ S \ f x \ A" + unfolding eventually_at_topological by fast + hence "open (T \ S)" "a \ T \ S" "\x\(T \ S). x \ a \ f x \ A" + using assms by auto + hence "\T. open T \ a \ T \ (\x\T. x \ a \ f x \ A)" + by fast + hence "eventually (\x. f x \ A) (at a)" + unfolding eventually_at_topological . + } + thus ?rhs by (rule topological_tendstoI) +next + assume ?rhs + thus ?lhs by (rule Lim_at_within) +qed + +text{* Another limit point characterization. *} + +lemma islimpt_sequential: + fixes x :: "'a::metric_space" (* FIXME: generalize to topological_space *) + shows "x islimpt S \ (\f. (\n::nat. f n \ S -{x}) \ (f ---> x) sequentially)" + (is "?lhs = ?rhs") +proof + assume ?lhs + then obtain f where f:"\y. y>0 \ f y \ S \ f y \ x \ dist (f y) x < y" + unfolding islimpt_approachable using choice[of "\e y. e>0 \ y\S \ y\x \ dist y x < e"] by auto + { fix n::nat + have "f (inverse (real n + 1)) \ S - {x}" using f by auto + } + moreover + { fix e::real assume "e>0" + hence "\N::nat. inverse (real (N + 1)) < e" using real_arch_inv[of e] apply (auto simp add: Suc_pred') apply(rule_tac x="n - 1" in exI) by auto + 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) + moreover have "\n\N. dist (f (inverse (real n + 1))) x < (inverse (real n + 1))" using f `e>0` by auto + ultimately have "\N::nat. \n\N. dist (f (inverse (real n + 1))) x < e" apply(rule_tac x=N in exI) apply auto apply(erule_tac x=n in allE)+ by auto + } + hence " ((\n. f (inverse (real n + 1))) ---> x) sequentially" + unfolding Lim_sequentially using f by auto + ultimately show ?rhs apply (rule_tac x="(\n::nat. f (inverse (real n + 1)))" in exI) by auto +next + assume ?rhs + then obtain f::"nat\'a" where f:"(\n. f n \ S - {x})" "(\e>0. \N. \n\N. dist (f n) x < e)" unfolding Lim_sequentially by auto + { fix e::real assume "e>0" + then obtain N where "dist (f N) x < e" using f(2) by auto + moreover have "f N\S" "f N \ x" using f(1) by auto + ultimately have "\x'\S. x' \ x \ dist x' x < e" by auto + } + thus ?lhs unfolding islimpt_approachable by auto +qed + +text{* Basic arithmetical combining theorems for limits. *} + +lemma Lim_linear: + assumes "(f ---> l) net" "bounded_linear h" + shows "((\x. h (f x)) ---> h l) net" +using `bounded_linear h` `(f ---> l) net` +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 \ 'b::real_normed_vector" + 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" + shows "(f ---> l) net ==> ((\x. -(f x)) ---> -l) net" + by (rule tendsto_minus) + +lemma Lim_add: fixes f :: "'a \ 'b::real_normed_vector" shows + "(f ---> l) net \ (g ---> m) net \ ((\x. f(x) + g(x)) ---> l + m) net" + by (rule tendsto_add) + +lemma Lim_sub: + fixes f :: "'a \ 'b::real_normed_vector" + shows "(f ---> l) net \ (g ---> m) net \ ((\x. f(x) - g(x)) ---> l - m) net" + by (rule tendsto_diff) + +lemma Lim_null: + fixes f :: "'a \ 'b::real_normed_vector" + shows "(f ---> l) net \ ((\x. f(x) - l) ---> 0) net" by (simp add: Lim dist_norm) + +lemma Lim_null_norm: + fixes f :: "'a \ 'b::real_normed_vector" + 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" "(g ---> 0) net" + shows "(f ---> 0) net" +proof(simp add: tendsto_iff, rule+) + fix e::real assume "0 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 (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: + fixes f :: "'a \ 'b::metric_space ^ 'n::finite" + shows "(f ---> l) net \ ((\a. f a $i) ---> l$i) net" + unfolding tendsto_iff + 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" + fixes g :: "'a \ 'c::real_normed_vector" + assumes "eventually (\n. norm(f n) <= norm(g n)) net" "(g ---> 0) net" + shows "(f ---> 0) net" +proof (rule tendstoI) + fix e::real assume "e>0" + { fix x + assume "norm (f x) \ norm (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) \ norm (g x)" "\x. dist (g x) 0 < e" net] + using eventually_mono[of "\x. norm (f x) \ norm (g x) \ dist (g x) 0 < e" "\x. dist (f x) 0 < e" net] + using assms `e>0` unfolding tendsto_iff by blast +qed + +text{* Deducing things about the limit from the elements. *} + +lemma Lim_in_closed_set: + assumes "closed S" "eventually (\x. f(x) \ S) net" "\(trivial_limit net)" "(f ---> l) net" + shows "l \ S" +proof (rule ccontr) + assume "l \ S" + with `closed S` have "open (- S)" "l \ - S" + by (simp_all add: open_Compl) + with assms(4) have "eventually (\x. f x \ - S) net" + by (rule topological_tendstoD) + with assms(2) have "eventually (\x. False) net" + by (rule eventually_elim2) simp + with assms(3) show "False" + by (simp add: eventually_False) +qed + +text{* Need to prove closed(cball(x,e)) before deducing this as a corollary. *} + +lemma Lim_dist_ubound: + assumes "\(trivial_limit net)" "(f ---> l) net" "eventually (\x. dist a (f x) <= e) net" + shows "dist a l <= e" +proof (rule ccontr) + assume "\ dist a l \ e" + then have "0 < dist a l - e" by simp + with assms(2) have "eventually (\x. dist (f x) l < dist a l - e) net" + by (rule tendstoD) + with assms(3) have "eventually (\x. dist a (f x) \ e \ dist (f x) l < dist a l - e) net" + by (rule eventually_conjI) + then obtain w where "dist a (f w) \ e" "dist (f w) l < dist a l - e" + using assms(1) eventually_happens by auto + hence "dist a (f w) + dist (f w) l < e + (dist a l - e)" + by (rule add_le_less_mono) + hence "dist a (f w) + dist (f w) l < dist a l" + by simp + also have "\ \ dist a (f w) + dist (f w) l" + by (rule dist_triangle) + finally show False by simp +qed + +lemma Lim_norm_ubound: + fixes f :: "'a \ 'b::real_normed_vector" + assumes "\(trivial_limit net)" "(f ---> l) net" "eventually (\x. norm(f x) <= e) net" + shows "norm(l) <= e" +proof (rule ccontr) + assume "\ norm l \ e" + then have "0 < norm l - e" by simp + with assms(2) have "eventually (\x. dist (f x) l < norm l - e) net" + by (rule tendstoD) + with assms(3) have "eventually (\x. norm (f x) \ e \ dist (f x) l < norm l - e) net" + by (rule eventually_conjI) + then obtain w where "norm (f w) \ e" "dist (f w) l < norm l - e" + using assms(1) eventually_happens by auto + hence "norm (f w - l) < norm l - e" "norm (f w) \ e" by (simp_all add: dist_norm) + hence "norm (f w - l) + norm (f w) < norm l" by simp + hence "norm (f w - l - f w) < norm l" by (rule le_less_trans [OF norm_triangle_ineq4]) + thus False using `\ norm l \ e` by simp +qed + +lemma Lim_norm_lbound: + fixes f :: "'a \ 'b::real_normed_vector" + assumes "\ (trivial_limit net)" "(f ---> l) net" "eventually (\x. e <= norm(f x)) net" + shows "e \ norm l" +proof (rule ccontr) + assume "\ e \ norm l" + then have "0 < e - norm l" by simp + with assms(2) have "eventually (\x. dist (f x) l < e - norm l) net" + by (rule tendstoD) + with assms(3) have "eventually (\x. e \ norm (f x) \ dist (f x) l < e - norm l) net" + by (rule eventually_conjI) + then obtain w where "e \ norm (f w)" "dist (f w) l < e - norm l" + using assms(1) eventually_happens by auto + hence "norm (f w - l) + norm l < e" "e \ norm (f w)" by (simp_all add: dist_norm) + hence "norm (f w - l) + norm l < norm (f w)" by (rule less_le_trans) + hence "norm (f w - l + l) < norm (f w)" by (rule le_less_trans [OF norm_triangle_ineq]) + thus False by simp +qed + +text{* Uniqueness of the limit, when nontrivial. *} + +lemma Lim_unique: + fixes f :: "'a \ 'b::t2_space" + assumes "\ trivial_limit net" "(f ---> l) net" "(f ---> l') net" + shows "l = l'" +proof (rule ccontr) + assume "l \ l'" + obtain U V where "open U" "open V" "l \ U" "l' \ V" "U \ V = {}" + using hausdorff [OF `l \ l'`] by fast + have "eventually (\x. f x \ U) net" + using `(f ---> l) net` `open U` `l \ U` by (rule topological_tendstoD) + moreover + have "eventually (\x. f x \ V) net" + using `(f ---> l') net` `open V` `l' \ V` by (rule topological_tendstoD) + ultimately + have "eventually (\x. False) net" + proof (rule eventually_elim2) + fix x + assume "f x \ U" "f x \ V" + hence "f x \ U \ V" by simp + with `U \ V = {}` show "False" by simp + qed + with `\ trivial_limit net` show "False" + by (simp add: eventually_False) +qed + +lemma tendsto_Lim: + fixes f :: "'a \ 'b::t2_space" + shows "~(trivial_limit net) \ (f ---> l) net ==> Lim net f = l" + unfolding Lim_def using Lim_unique[of net f] by auto + +text{* Limit under bilinear function *} + +lemma Lim_bilinear: + assumes "(f ---> l) net" and "(g ---> m) net" and "bounded_bilinear h" + shows "((\x. h (f x) (g x)) ---> (h l m)) net" +using `bounded_bilinear h` `(f ---> l) net` `(g ---> m) net` +by (rule bounded_bilinear.tendsto) + +text{* These are special for limits out of the same vector space. *} + +lemma Lim_within_id: "(id ---> a) (at a within s)" + unfolding tendsto_def Limits.eventually_within eventually_at_topological + by auto + +lemma Lim_at_id: "(id ---> a) (at a)" +apply (subst within_UNIV[symmetric]) by (simp add: Lim_within_id) + +lemma Lim_at_zero: + fixes a :: "'a::real_normed_vector" + fixes l :: "'b::topological_space" + shows "(f ---> l) (at a) \ ((\x. f(a + x)) ---> l) (at 0)" (is "?lhs = ?rhs") +proof + assume "?lhs" + { fix S assume "open S" "l \ S" + with `?lhs` have "eventually (\x. f x \ S) (at a)" + by (rule topological_tendstoD) + then obtain d where d: "d>0" "\x. x \ a \ dist x a < d \ f x \ S" + unfolding Limits.eventually_at by fast + { fix x::"'a" assume "x \ 0 \ dist x 0 < d" + hence "f (a + x) \ S" using d + apply(erule_tac x="x+a" in allE) + by(auto simp add: comm_monoid_add.mult_commute dist_norm dist_commute) + } + hence "\d>0. \x. x \ 0 \ dist x 0 < d \ f (a + x) \ S" + using d(1) by auto + hence "eventually (\x. f (a + x) \ S) (at 0)" + unfolding Limits.eventually_at . + } + thus "?rhs" by (rule topological_tendstoI) +next + assume "?rhs" + { fix S assume "open S" "l \ S" + with `?rhs` have "eventually (\x. f (a + x) \ S) (at 0)" + by (rule topological_tendstoD) + then obtain d where d: "d>0" "\x. x \ 0 \ dist x 0 < d \ f (a + x) \ S" + unfolding Limits.eventually_at by fast + { fix x::"'a" assume "x \ a \ dist x a < d" + hence "f x \ S" using d apply(erule_tac x="x-a" in allE) + by(auto simp add: comm_monoid_add.mult_commute dist_norm dist_commute) + } + hence "\d>0. \x. x \ a \ dist x a < d \ f x \ S" using d(1) by auto + hence "eventually (\x. f x \ S) (at a)" unfolding Limits.eventually_at . + } + thus "?lhs" by (rule topological_tendstoI) +qed + +text{* It's also sometimes useful to extract the limit point from the net. *} + +definition + netlimit :: "'a::t2_space net \ 'a" where + "netlimit net = (SOME a. ((\x. x) ---> a) net)" + +lemma netlimit_within: + assumes "\ trivial_limit (at a within S)" + shows "netlimit (at a within S) = a" +unfolding netlimit_def +apply (rule some_equality) +apply (rule Lim_at_within) +apply (rule Lim_ident_at) +apply (erule Lim_unique [OF assms]) +apply (rule Lim_at_within) +apply (rule Lim_ident_at) +done + +lemma netlimit_at: + fixes a :: "'a::perfect_space" + shows "netlimit (at a) = a" + apply (subst within_UNIV[symmetric]) + using netlimit_within[of a UNIV] + by (simp add: trivial_limit_at within_UNIV) + +text{* Transformation of limit. *} + +lemma Lim_transform: + fixes f g :: "'a::type \ 'b::real_normed_vector" + assumes "((\x. f x - g x) ---> 0) net" "(f ---> l) net" + shows "(g ---> l) net" +proof- + from assms have "((\x. f x - g x - f x) ---> 0 - l) net" using Lim_sub[of "\x. f x - g x" 0 net f l] by auto + thus "?thesis" using Lim_neg [of "\ x. - g x" "-l" net] by auto +qed + +lemma Lim_transform_eventually: + "eventually (\x. f x = g x) net \ (f ---> l) net ==> (g ---> l) net" + apply (rule topological_tendstoI) + apply (drule (2) topological_tendstoD) + apply (erule (1) eventually_elim2, simp) + done + +lemma Lim_transform_within: + fixes l :: "'b::metric_space" (* TODO: generalize *) + assumes "0 < d" "(\x'\S. 0 < dist x' x \ dist x' x < d \ f x' = g x')" + "(f ---> l) (at x within S)" + shows "(g ---> l) (at x within S)" + using assms(1,3) unfolding Lim_within + apply - + apply (clarify, rename_tac e) + apply (drule_tac x=e in spec, clarsimp, rename_tac r) + apply (rule_tac x="min d r" in exI, clarsimp, rename_tac y) + apply (drule_tac x=y in bspec, assumption, clarsimp) + apply (simp add: assms(2)) + done + +lemma Lim_transform_at: + fixes l :: "'b::metric_space" (* TODO: generalize *) + shows "0 < d \ (\x'. 0 < dist x' x \ dist x' x < d \ f x' = g x') \ + (f ---> l) (at x) ==> (g ---> l) (at x)" + apply (subst within_UNIV[symmetric]) + using Lim_transform_within[of d UNIV x f g l] + by (auto simp add: within_UNIV) + +text{* Common case assuming being away from some crucial point like 0. *} + +lemma Lim_transform_away_within: + fixes a b :: "'a::metric_space" + fixes l :: "'b::metric_space" (* TODO: generalize *) + assumes "a\b" "\x\ S. x \ a \ x \ b \ f x = g x" + and "(f ---> l) (at a within S)" + shows "(g ---> l) (at a within S)" +proof- + have "\x'\S. 0 < dist x' a \ dist x' a < dist a b \ f x' = g x'" using assms(2) + apply auto apply(erule_tac x=x' in ballE) by (auto simp add: dist_commute) + thus ?thesis using Lim_transform_within[of "dist a b" S a f g l] using assms(1,3) unfolding dist_nz by auto +qed + +lemma Lim_transform_away_at: + fixes a b :: "'a::metric_space" + fixes l :: "'b::metric_space" (* TODO: generalize *) + assumes ab: "a\b" and fg: "\x. x \ a \ x \ b \ f x = g x" + and fl: "(f ---> l) (at a)" + shows "(g ---> l) (at a)" + using Lim_transform_away_within[OF ab, of UNIV f g l] fg fl + by (auto simp add: within_UNIV) + +text{* Alternatively, within an open set. *} + +lemma Lim_transform_within_open: + fixes a :: "'a::metric_space" + fixes l :: "'b::metric_space" (* TODO: generalize *) + assumes "open S" "a \ S" "\x\S. x \ a \ f x = g x" "(f ---> l) (at a)" + shows "(g ---> l) (at a)" +proof- + from assms(1,2) obtain e::real where "e>0" and e:"ball a e \ S" unfolding open_contains_ball by auto + hence "\x'. 0 < dist x' a \ dist x' a < e \ f x' = g x'" using assms(3) + unfolding ball_def subset_eq apply auto apply(erule_tac x=x' in allE) apply(erule_tac x=x' in ballE) by(auto simp add: dist_commute) + thus ?thesis using Lim_transform_at[of e a f g l] `e>0` assms(4) by auto +qed + +text{* A congruence rule allowing us to transform limits assuming not at point. *} + +(* FIXME: Only one congruence rule for tendsto can be used at a time! *) + +lemma Lim_cong_within[cong add]: + fixes a :: "'a::metric_space" + fixes l :: "'b::metric_space" (* TODO: generalize *) + shows "(\x. x \ a \ f x = g x) ==> ((\x. f x) ---> l) (at a within S) \ ((g ---> l) (at a within S))" + by (simp add: Lim_within dist_nz[symmetric]) + +lemma Lim_cong_at[cong add]: + fixes a :: "'a::metric_space" + fixes l :: "'b::metric_space" (* TODO: generalize *) + shows "(\x. x \ a ==> f x = g x) ==> (((\x. f x) ---> l) (at a) \ ((g ---> l) (at a)))" + by (simp add: Lim_at dist_nz[symmetric]) + +text{* Useful lemmas on closure and set of possible sequential limits.*} + +lemma closure_sequential: + fixes l :: "'a::metric_space" (* TODO: generalize *) + shows "l \ closure S \ (\x. (\n. x n \ S) \ (x ---> l) sequentially)" (is "?lhs = ?rhs") +proof + assume "?lhs" moreover + { assume "l \ S" + hence "?rhs" using Lim_const[of l sequentially] by auto + } moreover + { assume "l islimpt S" + hence "?rhs" unfolding islimpt_sequential by auto + } ultimately + show "?rhs" unfolding closure_def by auto +next + assume "?rhs" + thus "?lhs" unfolding closure_def unfolding islimpt_sequential by auto +qed + +lemma closed_sequential_limits: + fixes S :: "'a::metric_space set" + shows "closed S \ (\x l. (\n. x n \ S) \ (x ---> l) sequentially \ l \ S)" + unfolding closed_limpt + using closure_sequential [where 'a='a] closure_closed [where 'a='a] closed_limpt [where 'a='a] islimpt_sequential [where 'a='a] mem_delete [where 'a='a] + by metis + +lemma closure_approachable: + fixes S :: "'a::metric_space set" + shows "x \ closure S \ (\e>0. \y\S. dist y x < e)" + apply (auto simp add: closure_def islimpt_approachable) + by (metis dist_self) + +lemma closed_approachable: + fixes S :: "'a::metric_space set" + shows "closed S ==> (\e>0. \y\S. dist y x < e) \ x \ S" + by (metis closure_closed closure_approachable) + +text{* Some other lemmas about sequences. *} + +lemma seq_offset: + fixes l :: "'a::metric_space" (* TODO: generalize *) + shows "(f ---> l) sequentially ==> ((\i. f( i + k)) ---> l) sequentially" + apply (auto simp add: Lim_sequentially) + by (metis trans_le_add1 ) + +lemma seq_offset_neg: + "(f ---> l) sequentially ==> ((\i. f(i - k)) ---> l) sequentially" + apply (rule topological_tendstoI) + apply (drule (2) topological_tendstoD) + apply (simp only: eventually_sequentially) + apply (subgoal_tac "\N k (n::nat). N + k <= n ==> N <= n - k") + apply metis + by arith + +lemma seq_offset_rev: + "((\i. f(i + k)) ---> l) sequentially ==> (f ---> l) sequentially" + apply (rule topological_tendstoI) + apply (drule (2) topological_tendstoD) + apply (simp only: eventually_sequentially) + apply (subgoal_tac "\N k (n::nat). N + k <= n ==> N <= n - k \ (n - k) + k = n") + by metis arith + +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 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 by simp +qed + +text{* More properties of closed balls. *} + +lemma closed_cball: "closed (cball x e)" +unfolding cball_def closed_def +unfolding Collect_neg_eq [symmetric] not_le +apply (clarsimp simp add: open_dist, rename_tac y) +apply (rule_tac x="dist x y - e" in exI, clarsimp) +apply (rename_tac x') +apply (cut_tac x=x and y=x' and z=y in dist_triangle) +apply simp +done + +lemma open_contains_cball: "open S \ (\x\S. \e>0. cball x e \ S)" +proof- + { fix x and e::real assume "x\S" "e>0" "ball x e \ S" + hence "\d>0. cball x d \ S" unfolding subset_eq by (rule_tac x="e/2" in exI, auto) + } moreover + { fix x and e::real assume "x\S" "e>0" "cball x e \ S" + hence "\d>0. ball x d \ S" unfolding subset_eq apply(rule_tac x="e/2" in exI) by auto + } ultimately + show ?thesis unfolding open_contains_ball by auto +qed + +lemma open_contains_cball_eq: "open S ==> (\x. x \ S \ (\e>0. cball x e \ S))" + by (metis open_contains_cball subset_eq order_less_imp_le centre_in_cball mem_def) + +lemma mem_interior_cball: "x \ interior S \ (\e>0. cball x e \ S)" + apply (simp add: interior_def, safe) + apply (force simp add: open_contains_cball) + apply (rule_tac x="ball x e" in exI) + apply (simp add: open_ball centre_in_ball subset_trans [OF ball_subset_cball]) + done + +lemma islimpt_ball: + fixes x y :: "'a::{real_normed_vector,perfect_space}" + shows "y islimpt ball x e \ 0 < e \ y \ cball x e" (is "?lhs = ?rhs") +proof + assume "?lhs" + { assume "e \ 0" + hence *:"ball x e = {}" using ball_eq_empty[of x e] by auto + have False using `?lhs` unfolding * using islimpt_EMPTY[of y] by auto + } + hence "e > 0" by (metis not_less) + moreover + have "y \ cball x e" using closed_cball[of x e] islimpt_subset[of y "ball x e" "cball x e"] ball_subset_cball[of x e] `?lhs` unfolding closed_limpt by auto + ultimately show "?rhs" by auto +next + assume "?rhs" hence "e>0" by auto + { fix d::real assume "d>0" + have "\x'\ball x e. x' \ y \ dist x' y < d" + proof(cases "d \ dist x y") + case True thus "\x'\ball x e. x' \ y \ dist x' y < d" + proof(cases "x=y") + case True hence False using `d \ dist x y` `d>0` by auto + thus "\x'\ball x e. x' \ y \ dist x' y < d" by auto + next + case False + + have "dist x (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) + = norm (x - y + (d / (2 * norm (y - x))) *\<^sub>R (y - x))" + unfolding mem_cball mem_ball dist_norm diff_diff_eq2 diff_add_eq[THEN sym] by auto + 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) + 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 + also have "\ \ e - d/2" using `d \ dist x y` and `d>0` and `?rhs` by(auto simp add: dist_norm) + finally have "y - (d / (2 * dist y x)) *\<^sub>R (y - x) \ ball x e" using `d>0` by auto + + moreover + + 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 + 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 + qed + next + case False hence "d > dist x y" by auto + show "\x'\ball x e. x' \ y \ dist x' y < d" + proof(cases "x=y") + case True + obtain z where **: "z \ y" "dist z y < min e d" + using perfect_choose_dist[of "min e d" y] + using `d > 0` `e>0` by auto + show "\x'\ball x e. x' \ y \ dist x' y < d" + unfolding `x = y` + using `z \ y` ** + by (rule_tac x=z in bexI, auto simp add: dist_commute) + next + case False thus "\x'\ball x e. x' \ y \ dist x' y < d" + using `d>0` `d > dist x y` `?rhs` by(rule_tac x=x in bexI, auto) + qed + qed } + thus "?lhs" unfolding mem_cball islimpt_approachable mem_ball by auto +qed + +lemma closure_ball_lemma: + fixes x y :: "'a::real_normed_vector" + assumes "x \ y" shows "y islimpt ball x (dist x y)" +proof (rule islimptI) + fix T assume "y \ T" "open T" + then obtain r where "0 < r" "\z. dist z y < r \ z \ T" + unfolding open_dist by fast + (* choose point between x and y, within distance r of y. *) + def k \ "min 1 (r / (2 * dist x y))" + def z \ "y + scaleR k (x - y)" + have z_def2: "z = x + scaleR (1 - k) (y - x)" + 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 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_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) + apply (simp add: k_def divide_pos_pos zero_less_dist_iff `0 < r` `x \ y`) + apply (simp add: zero_less_dist_iff `x \ y`) + done + hence "z \ ball x (dist x y)" by simp + have "z \ y" + unfolding z_def k_def using `x \ y` `0 < r` + by (simp add: min_def) + show "\z\ball x (dist x y). z \ T \ z \ y" + using `z \ ball x (dist x y)` `z \ T` `z \ y` + by fast +qed + +lemma closure_ball: + fixes x :: "'a::real_normed_vector" + shows "0 < e \ closure (ball x e) = cball x e" +apply (rule equalityI) +apply (rule closure_minimal) +apply (rule ball_subset_cball) +apply (rule closed_cball) +apply (rule subsetI, rename_tac y) +apply (simp add: le_less [where 'a=real]) +apply (erule disjE) +apply (rule subsetD [OF closure_subset], simp) +apply (simp add: closure_def) +apply clarify +apply (rule closure_ball_lemma) +apply (simp add: zero_less_dist_iff) +done + +(* In a trivial vector space, this fails for e = 0. *) +lemma interior_cball: + 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 + { fix y assume "y \ cball x e" + hence False unfolding mem_cball using dist_nz[of x y] cs by auto } + hence "cball x e = {}" by auto + hence "interior (cball x e) = {}" using interior_empty by auto + ultimately show ?thesis by blast +next + case True note cs = this + have "ball x e \ cball x e" using ball_subset_cball by auto moreover + { 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_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") + case True + hence "e>0" using xa_y[unfolded dist_nz] xa_cball[unfolded mem_cball] by (auto simp add: dist_commute) + thus "y \ ball x e" using `x = y ` by simp + next + case False + 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) *\<^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) *\<^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 + qed } + hence "\S \ cball x e. open S \ S \ ball x e" by auto + ultimately show ?thesis using interior_unique[of "ball x e" "cball x e"] using open_ball[of x e] by auto +qed + +lemma frontier_ball: + 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 :: "'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) + by arith + +lemma cball_eq_empty: "(cball x e = {}) \ e < 0" + apply (simp add: expand_set_eq not_le) + by (metis zero_le_dist dist_self order_less_le_trans) +lemma cball_empty: "e < 0 ==> cball x e = {}" by (simp add: cball_eq_empty) + +lemma cball_eq_sing: + fixes x :: "'a::perfect_space" + shows "(cball x e = {x}) \ e = 0" +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 :: "'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. *} + +lemma eventually_within_interior: + assumes "x \ interior S" + shows "eventually P (at x within S) \ eventually P (at x)" (is "?lhs = ?rhs") +proof- + from assms obtain T where T: "open T" "x \ T" "T \ S" + unfolding interior_def by fast + { assume "?lhs" + then obtain A where "open A" "x \ A" "\y\A. y \ x \ y \ S \ P y" + unfolding Limits.eventually_within Limits.eventually_at_topological + by auto + with T have "open (A \ T)" "x \ A \ T" "\y\(A \ T). y \ x \ P y" + by auto + then have "?rhs" + unfolding Limits.eventually_at_topological by auto + } moreover + { assume "?rhs" hence "?lhs" + unfolding Limits.eventually_within + by (auto elim: eventually_elim1) + } ultimately + show "?thesis" .. +qed + +lemma lim_within_interior: + "x \ interior S \ (f ---> l) (at x within S) \ (f ---> l) (at x)" + unfolding tendsto_def by (simp add: eventually_within_interior) + +lemma netlimit_within_interior: + fixes x :: "'a::{perfect_space, real_normed_vector}" + (* FIXME: generalize to perfect_space *) + assumes "x \ interior S" + shows "netlimit(at x within S) = x" (is "?lhs = ?rhs") +proof- + from assms obtain e::real where e:"e>0" "ball x e \ S" using open_interior[of S] unfolding open_contains_ball using interior_subset[of S] by auto + hence "\ trivial_limit (at x within S)" using islimpt_subset[of x "ball x e" S] unfolding trivial_limit_within islimpt_ball centre_in_cball by auto + thus ?thesis using netlimit_within by auto +qed + +subsection{* Boundedness. *} + + (* FIXME: This has to be unified with BSEQ!! *) +definition + bounded :: "'a::metric_space set \ bool" where + "bounded S \ (\x e. \y\S. dist x y \ e)" + +lemma bounded_any_center: "bounded S \ (\e. \y\S. dist a y \ e)" +unfolding bounded_def +apply safe +apply (rule_tac x="dist a x + e" in exI, clarify) +apply (drule (1) bspec) +apply (erule order_trans [OF dist_triangle add_left_mono]) +apply auto +done + +lemma bounded_iff: "bounded S \ (\a. \x\S. norm x \ a)" +unfolding bounded_any_center [where a=0] +by (simp add: dist_norm) + +lemma bounded_empty[simp]: "bounded {}" by (simp add: bounded_def) +lemma bounded_subset: "bounded T \ S \ T ==> bounded S" + by (metis bounded_def subset_eq) + +lemma bounded_interior[intro]: "bounded S ==> bounded(interior S)" + by (metis bounded_subset interior_subset) + +lemma bounded_closure[intro]: assumes "bounded S" shows "bounded(closure S)" +proof- + from assms obtain x and a where a: "\y\S. dist x y \ a" unfolding bounded_def by auto + { fix y assume "y \ closure S" + then obtain f where f: "\n. f n \ S" "(f ---> y) sequentially" + unfolding closure_sequential by auto + have "\n. f n \ S \ dist x (f n) \ a" using a by simp + hence "eventually (\n. dist x (f n) \ a) sequentially" + by (rule eventually_mono, simp add: f(1)) + have "dist x y \ a" + apply (rule Lim_dist_ubound [of sequentially f]) + apply (rule trivial_limit_sequentially) + apply (rule f(2)) + apply fact + done + } + thus ?thesis unfolding bounded_def by auto +qed + +lemma bounded_cball[simp,intro]: "bounded (cball x e)" + apply (simp add: bounded_def) + apply (rule_tac x=x in exI) + apply (rule_tac x=e in exI) + apply auto + done + +lemma bounded_ball[simp,intro]: "bounded(ball x e)" + by (metis ball_subset_cball bounded_cball bounded_subset) + +lemma finite_imp_bounded[intro]: assumes "finite S" shows "bounded S" +proof- + { fix a F assume as:"bounded F" + then obtain x e where "\y\F. dist x y \ e" unfolding bounded_def by auto + hence "\y\(insert a F). dist x y \ max e (dist x a)" by auto + hence "bounded (insert a F)" unfolding bounded_def by (intro exI) + } + thus ?thesis using finite_induct[of S bounded] using bounded_empty assms by auto +qed + +lemma bounded_Un[simp]: "bounded (S \ T) \ bounded S \ bounded T" + apply (auto simp add: bounded_def) + apply (rename_tac x y r s) + apply (rule_tac x=x in exI) + apply (rule_tac x="max r (dist x y + s)" in exI) + apply (rule ballI, rename_tac z, safe) + apply (drule (1) bspec, simp) + apply (drule (1) bspec) + apply (rule min_max.le_supI2) + apply (erule order_trans [OF dist_triangle add_left_mono]) + done + +lemma bounded_Union[intro]: "finite F \ (\S\F. bounded S) \ bounded(\F)" + by (induct rule: finite_induct[of F], auto) + +lemma bounded_pos: "bounded S \ (\b>0. \x\ S. norm x <= b)" + apply (simp add: bounded_iff) + apply (subgoal_tac "\x (y::real). 0 < 1 + abs y \ (x <= y \ x <= 1 + abs y)") + by metis arith + +lemma bounded_Int[intro]: "bounded S \ bounded T \ bounded (S \ T)" + by (metis Int_lower1 Int_lower2 bounded_subset) + +lemma bounded_diff[intro]: "bounded S ==> bounded (S - T)" +apply (metis Diff_subset bounded_subset) +done + +lemma bounded_insert[intro]:"bounded(insert x S) \ bounded S" + by (metis Diff_cancel Un_empty_right Un_insert_right bounded_Un bounded_subset finite.emptyI finite_imp_bounded infinite_remove subset_insertI) + +lemma not_bounded_UNIV[simp, intro]: + "\ bounded (UNIV :: 'a::{real_normed_vector, perfect_space} set)" +proof(auto simp add: bounded_pos not_le) + obtain x :: 'a where "x \ 0" + using perfect_choose_dist [OF zero_less_one] by fast + 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_sgn) + then show "\x::'a. b < norm x" .. +qed + +lemma bounded_linear_image: + assumes "bounded S" "bounded_linear f" + shows "bounded(f ` S)" +proof- + from assms(1) obtain b where b:"b>0" "\x\S. norm x \ b" unfolding bounded_pos by auto + from assms(2) obtain B where B:"B>0" "\x. norm (f x) \ B * norm x" using bounded_linear.pos_bounded by (auto simp add: mult_ac) + { fix x assume "x\S" + hence "norm x \ b" using b by auto + hence "norm (f x) \ B * b" using B(2) apply(erule_tac x=x in allE) + by (metis B(1) B(2) real_le_trans real_mult_le_cancel_iff2) + } + thus ?thesis unfolding bounded_pos apply(rule_tac x="b*B" in exI) + using b B real_mult_order[of b B] by (auto simp add: real_mult_commute) +qed + +lemma bounded_scaling: + fixes S :: "'a::real_normed_vector set" + shows "bounded S \ bounded ((\x. c *\<^sub>R x) ` S)" + apply (rule bounded_linear_image, assumption) + apply (rule scaleR.bounded_linear_right) + done + +lemma bounded_translation: + fixes S :: "'a::real_normed_vector set" + assumes "bounded S" shows "bounded ((\x. a + x) ` S)" +proof- + from assms obtain b where b:"b>0" "\x\S. norm x \ b" unfolding bounded_pos by auto + { fix x assume "x\S" + hence "norm (a + x) \ b + norm a" using norm_triangle_ineq[of a x] b by auto + } + thus ?thesis unfolding bounded_pos using norm_ge_zero[of a] b(1) using add_strict_increasing[of b 0 "norm a"] + by (auto intro!: add exI[of _ "b + norm a"]) +qed + + +text{* Some theorems on sups and infs using the notion "bounded". *} + +lemma bounded_real: + fixes S :: "real set" + 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_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_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_real) + by (auto simp add: isLub_def isUb_def leastP_def setle_def setge_def) +qed + +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 +next + let ?S = "insert x S" + case False + hence *:"\x\S. x \ rsup S" using bounded_has_rsup(1)[of S] using assms by auto + hence "insert x S *<= max x (rsup S)" unfolding setle_def by auto + hence "isLub UNIV ?S (rsup ?S)" using rsup[of ?S] by auto + moreover + have **:"isUb UNIV ?S (max x (rsup S))" unfolding isUb_def setle_def using * by auto + { fix y assume as:"isUb UNIV (insert x S) y" + hence "max x (rsup S) \ y" unfolding isUb_def using rsup_le[OF `S\{}`] + unfolding setle_def by auto } + hence "max x (rsup S) <=* isUb UNIV (insert x S)" unfolding setge_def Ball_def mem_def by auto + hence "isLub UNIV ?S (max x (rsup S))" using ** isLubI2[of UNIV ?S "max x (rsup S)"] unfolding Collect_def by auto + ultimately show ?thesis using real_isLub_unique[of UNIV ?S] using `S\{}` by auto +qed + +lemma sup_insert_finite: "finite S \ rsup(insert x S) = (if S = {} then x else max x (rsup S))" + apply (rule rsup_insert) + apply (rule finite_imp_bounded) + by simp + +lemma bounded_has_rinf: + 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_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_real) + by (auto simp add: isGlb_def isLb_def greatestP_def setle_def setge_def) +qed + +(* TODO: Move this to RComplete.thy -- would need to include Glb into RComplete *) +lemma real_isGlb_unique: "[| isGlb R S x; isGlb R S y |] ==> x = (y::real)" + apply (frule isGlb_isLb) + apply (frule_tac x = y in isGlb_isLb) + apply (blast intro!: order_antisym dest!: isGlb_le_isLb) + done + +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 +next + let ?S = "insert x S" + case False + hence *:"\x\S. x \ rinf S" using bounded_has_rinf(1)[of S] using assms by auto + hence "min x (rinf S) <=* insert x S" unfolding setge_def by auto + hence "isGlb UNIV ?S (rinf ?S)" using rinf[of ?S] by auto + moreover + have **:"isLb UNIV ?S (min x (rinf S))" unfolding isLb_def setge_def using * by auto + { fix y assume as:"isLb UNIV (insert x S) y" + hence "min x (rinf S) \ y" unfolding isLb_def using rinf_ge[OF `S\{}`] + unfolding setge_def by auto } + hence "isLb UNIV (insert x S) *<= min x (rinf S)" unfolding setle_def Ball_def mem_def by auto + hence "isGlb UNIV ?S (min x (rinf S))" using ** isGlbI2[of UNIV ?S "min x (rinf S)"] unfolding Collect_def by auto + ultimately show ?thesis using real_isGlb_unique[of UNIV ?S] using `S\{}` by auto +qed + +lemma inf_insert_finite: "finite S ==> rinf(insert x S) = (if S = {} then x else min x (rinf S))" + by (rule rinf_insert, rule finite_imp_bounded, simp) + +subsection{* Compactness (the definition is the one based on convegent subsequences). *} + +definition + compact :: "'a::metric_space set \ bool" where (* TODO: generalize *) + "compact S \ + (\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 + Heine-Borel property if every closed and bounded subset is compact. +*} + +class heine_borel = + assumes bounded_imp_convergent_subsequence: + "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: "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. subseq r \ ((f \ r) ---> l) sequentially" + using `l \ s` r l by blast +qed + +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 [unfolded subseq_def] by auto + ultimately show "Suc n \ r (Suc n)" by auto +qed + +lemma eventually_subseq: + assumes r: "subseq r" + shows "eventually P sequentially \ eventually (\n. P (r n)) sequentially" +unfolding eventually_sequentially +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 + apply (rule_tac x="nat_rec e f" in exI) + apply (rule conjI)+ +apply (rule def_nat_rec_0, simp) +apply (rule allI, rule def_nat_rec_Suc, simp) +apply (rule allI, rule impI, rule ext) +apply (erule conjE) +apply (induct_tac x) +apply (simp add: nat_rec_0) +apply (erule_tac x="n" in allE) +apply (simp) +done + +lemma convergent_bounded_increasing: fixes s ::"nat\real" + 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 + then obtain t where t:"isLub UNIV (range s) t" using reals_complete[of "range s" ] by auto + { fix e::real assume "e>0" and as:"\N. \n\N. \ \s n - t\ < e" + { fix n::nat + 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)[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 "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. subseq r \ ((s \ r) ---> l) sequentially" +proof- + 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 + +instance real :: heine_borel +proof + fix s :: "real set" and f :: "nat \ real" + assume s: "bounded s" and f: "\n. f n \ s" + then obtain b where b: "\n. abs (f n) \ b" + unfolding bounded_iff by auto + obtain l :: real and r :: "nat \ nat" where + r: "subseq r" and l: "((f \ r) ---> l) sequentially" + using compact_real_lemma [OF b] by auto + thus "\l r. subseq r \ ((f \ r) ---> l) sequentially" + by auto +qed + +lemma bounded_component: "bounded s \ bounded ((\x. x $ i) ` s)" +unfolding bounded_def +apply clarify +apply (rule_tac x="x $ i" in exI) +apply (rule_tac x="e" in exI) +apply clarify +apply (rule order_trans [OF dist_nth_le], simp) +done + +lemma compact_lemma: + fixes f :: "nat \ 'a::heine_borel ^ 'n::finite" + assumes "bounded s" and "\n. f n \ s" + shows "\d. + \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. subseq r \ + (\e>0. eventually (\n. \i\d. dist (f (r n) $ i) (l $ i) < e) sequentially)" + 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:"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:"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:"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_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) + } + ultimately show ?case by auto + qed +qed + +instance "^" :: (heine_borel, finite) heine_borel +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: "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" + { fix e::real assume "e>0" + hence "0 < e / (real_of_nat (card ?d))" + using zero_less_card_finite using divide_pos_pos[of e, of "real_of_nat (card ?d)"] by auto + with l have "eventually (\n. \i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))) sequentially" + by simp + moreover + { fix n assume n: "\i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))" + have "dist (f (r n)) l \ (\i\?d. dist (f (r n) $ i) (l $ i))" + unfolding dist_vector_def using zero_le_dist by (rule setL2_le_setsum) + also have "\ < (\i\?d. e / (real_of_nat (card ?d)))" + by (rule setsum_strict_mono) (simp_all add: n) + finally have "dist (f (r n)) l < e" by simp + } + ultimately have "eventually (\n. dist (f (r n)) l < e) sequentially" + by (rule eventually_elim1) + } + hence *:"((f \ r) ---> l) sequentially" unfolding o_def tendsto_iff by simp + 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. *} + +lemma cauchy_def: + "Cauchy s \ (\e>0. \N. \m n. m \ N \ n \ N --> dist(s m)(s n) < e)" +unfolding Cauchy_def by blast + +definition + complete :: "'a::metric_space set \ bool" where + "complete s \ (\f. (\n. f n \ s) \ Cauchy f + --> (\l \ s. (f ---> l) sequentially))" + +lemma cauchy: "Cauchy s \ (\e>0.\ N::nat. \n\N. dist(s n)(s N) < e)" (is "?lhs = ?rhs") +proof- + { assume ?rhs + { fix e::real + assume "e>0" + with `?rhs` obtain N where N:"\n\N. dist (s n) (s N) < e/2" + by (erule_tac x="e/2" in allE) auto + { fix n m + assume nm:"N \ m \ N \ n" + hence "dist (s m) (s n) < e" using N + using dist_triangle_half_l[of "s m" "s N" "e" "s n"] + by blast + } + hence "\N. \m n. N \ m \ N \ n \ dist (s m) (s n) < e" + by blast + } + hence ?lhs + unfolding cauchy_def + by blast + } + thus ?thesis + unfolding cauchy_def + using dist_triangle_half_l + by blast +qed + +lemma convergent_imp_cauchy: + "(s ---> l) sequentially ==> Cauchy s" +proof(simp only: cauchy_def, rule, rule) + fix e::real assume "e>0" "(s ---> l) sequentially" + then obtain N::nat where N:"\n\N. dist (s n) l < e/2" unfolding Lim_sequentially by(erule_tac x="e/2" in allE) auto + thus "\N. \m n. N \ m \ N \ n \ dist (s m) (s n) < e" using dist_triangle_half_l[of _ l e _] by (rule_tac x=N in exI) auto +qed + +lemma cauchy_imp_bounded: assumes "Cauchy s" shows "bounded {y. (\n::nat. y = s n)}" +proof- + from assms obtain N::nat where "\m n. N \ m \ N \ n \ dist (s m) (s n) < 1" unfolding cauchy_def apply(erule_tac x= 1 in allE) by auto + hence N:"\n. N \ n \ dist (s N) (s n) < 1" by auto + moreover + have "bounded (s ` {0..N})" using finite_imp_bounded[of "s ` {1..N}"] by auto + then obtain a where a:"\x\s ` {0..N}. dist (s N) x \ a" + unfolding bounded_any_center [where a="s N"] by auto + ultimately show "?thesis" + unfolding bounded_any_center [where a="s N"] + apply(rule_tac x="max a 1" in exI) apply auto + apply(erule_tac x=n in allE) apply(erule_tac x=n in ballE) by auto +qed + +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" "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 + from lr(3)[unfolded Lim_sequentially, THEN spec[where x="e/2"]] obtain M where M:"\n\M. dist ((f \ r) n) l < e/2" using `e>0` by auto + { fix n::nat assume n:"n \ max N M" + have "dist ((f \ r) n) l < e/2" using n M by auto + moreover have "r n \ N" using lr'[of n] n by auto + hence "dist (f n) ((f \ r) n) < e / 2" using N using n by auto + ultimately have "dist (f n) l < e" using dist_triangle_half_r[of "f (r n)" "f n" e l] by (auto simp add: dist_commute) } + hence "\N. \n\N. dist (f n) l < e" by blast } + hence "\l\s. (f ---> l) sequentially" using `l\s` unfolding Lim_sequentially by auto } + thus ?thesis unfolding complete_def by auto +qed + +instance heine_borel < complete_space +proof + fix f :: "nat \ 'a" assume "Cauchy f" + hence "bounded (range f)" unfolding image_def + using cauchy_imp_bounded [of f] by auto + hence "compact (closure (range f))" + using bounded_closed_imp_compact [of "closure (range f)"] by auto + hence "complete (closure (range f))" + using compact_imp_complete by auto + moreover have "\n. f n \ closure (range f)" + using closure_subset [of "range f"] by auto + ultimately have "\l\closure (range f). (f ---> l) sequentially" + using `Cauchy f` unfolding complete_def by auto + then show "convergent f" + unfolding convergent_def LIMSEQ_conv_tendsto [symmetric] by auto +qed + +lemma complete_univ: "complete (UNIV :: 'a::complete_space set)" +proof(simp add: complete_def, rule, rule) + fix f :: "nat \ 'a" assume "Cauchy f" + hence "convergent f" by (rule Cauchy_convergent) + hence "\l. f ----> l" unfolding convergent_def . + thus "\l. (f ---> l) sequentially" unfolding LIMSEQ_conv_tendsto . +qed + +lemma complete_imp_closed: assumes "complete s" shows "closed s" +proof - + { fix x assume "x islimpt s" + then obtain f where f: "\n. f n \ s - {x}" "(f ---> x) sequentially" + unfolding islimpt_sequential by auto + then obtain l where l: "l\s" "(f ---> l) sequentially" + using `complete s`[unfolded complete_def] using convergent_imp_cauchy[of f x] by auto + hence "x \ s" using Lim_unique[of sequentially f l x] trivial_limit_sequentially f(2) by auto + } + thus "closed s" unfolding closed_limpt by auto +qed + +lemma complete_eq_closed: + fixes s :: "'a::complete_space set" + shows "complete s \ closed s" (is "?lhs = ?rhs") +proof + assume ?lhs thus ?rhs by (rule complete_imp_closed) +next + assume ?rhs + { fix f assume as:"\n::nat. f n \ s" "Cauchy f" + then obtain l where "(f ---> l) sequentially" using complete_univ[unfolded complete_def, THEN spec[where x=f]] by auto + hence "\l\s. (f ---> l) sequentially" using `?rhs`[unfolded closed_sequential_limits, THEN spec[where x=f], THEN spec[where x=l]] using as(1) by auto } + thus ?lhs unfolding complete_def by auto +qed + +lemma convergent_eq_cauchy: + fixes s :: "nat \ 'a::complete_space" + shows "(\l. (s ---> l) sequentially) \ Cauchy s" (is "?lhs = ?rhs") +proof + assume ?lhs then obtain l where "(s ---> l) sequentially" by auto + thus ?rhs using convergent_imp_cauchy by auto +next + assume ?rhs thus ?lhs using complete_univ[unfolded complete_def, THEN spec[where x=s]] by auto +qed + +lemma convergent_imp_bounded: + fixes s :: "nat \ 'a::metric_space" + shows "(s ---> l) sequentially ==> bounded (s ` (UNIV::(nat set)))" + using convergent_imp_cauchy[of s] + using cauchy_imp_bounded[of s] + unfolding image_def + by auto + +subsection{* Total boundedness. *} + +fun helper_1::"('a::metric_space set) \ real \ nat \ 'a" where + "helper_1 s e n = (SOME y::'a. y \ s \ (\m (dist (helper_1 s e m) y < e)))" +declare helper_1.simps[simp del] + +lemma compact_imp_totally_bounded: + assumes "compact s" + shows "\e>0. \k. finite k \ k \ s \ s \ (\((\x. ball x e) ` k))" +proof(rule, rule, rule ccontr) + fix e::real assume "e>0" and assm:"\ (\k. finite k \ k \ s \ s \ \(\x. ball x e) ` k)" + def x \ "helper_1 s e" + { fix n + have "x n \ s \ (\m dist (x m) (x n) < e)" + proof(induct_tac rule:nat_less_induct) + fix n def Q \ "(\y. y \ s \ (\m dist (x m) y < e))" + assume as:"\m s \ (\ma dist (x ma) (x m) < e)" + have "\ s \ (\x\x ` {0..s" "z \ (\x\x ` {0.. 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:"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[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 + +subsection{* Heine-Borel theorem (following Burkill \& Burkill vol. 2) *} + +lemma heine_borel_lemma: fixes s::"'a::metric_space set" + assumes "compact s" "s \ (\ t)" "\b \ t. open b" + shows "\e>0. \x \ s. \b \ t. ball x e \ b" +proof(rule ccontr) + assume "\ (\e>0. \x\s. \b\t. ball x e \ b)" + hence cont:"\e>0. \x\s. \xa\t. \ (ball x e \ xa)" by auto + { fix n::nat + have "1 / real (n + 1) > 0" by auto + hence "\x. x\s \ (\xa\t. \ (ball x (inverse (real (n+1))) \ xa))" using cont unfolding Bex_def by auto } + hence "\n::nat. \x. x \ s \ (\xa\t. \ ball x (inverse (real (n + 1))) \ xa)" by auto + 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:"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 + then obtain e where "e>0" and e:"\z. dist z l < e \ z\b" + using assms(3)[THEN bspec[where x=b]] unfolding open_dist by auto + + then obtain N1 where N1:"\n\N1. dist ((f \ r) n) l < e / 2" + using lr[unfolded Lim_sequentially, THEN spec[where x="e/2"]] by auto + + 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 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 + using f[THEN spec[where x="r (N1 + N2)"]] using `b\t` by auto + have "\y\ball x (inverse (real (r (N1 + N2) + 1))). y\b" apply(rule ccontr) using x by auto + then obtain y where y:"y \ ball x (inverse (real (r (N1 + N2) + 1)))" "y \ b" by auto + + have "dist x l < e/2" using N1 unfolding x_def o_def by auto + hence "dist y l < e" using y N2' using dist_triangle[of y l x]by (auto simp add:dist_commute) + + thus False using e and `y\b` by auto +qed + +lemma compact_imp_heine_borel: "compact s ==> (\f. (\t \ f. open t) \ s \ (\ f) + \ (\f'. f' \ f \ finite f' \ s \ (\ f')))" +proof clarify + fix f assume "compact s" " \t\f. open t" "s \ \f" + then obtain e::real where "e>0" and "\x\s. \b\f. ball x e \ b" using heine_borel_lemma[of s f] by auto + hence "\x\s. \b. b\f \ ball x e \ b" by auto + hence "\bb. \x\s. bb x \f \ ball x e \ bb x" using bchoice[of s "\x b. b\f \ ball x e \ b"] by auto + then obtain bb where bb:"\x\s. (bb x) \ f \ ball x e \ (bb x)" by blast + + from `compact s` have "\ k. finite k \ k \ s \ s \ \(\x. ball x e) ` k" using compact_imp_totally_bounded[of s] `e>0` by auto + then obtain k where k:"finite k" "k \ s" "s \ \(\x. ball x e) ` k" by auto + + have "finite (bb ` k)" using k(1) by auto + moreover + { fix x assume "x\s" + hence "x\\(\x. ball x e) ` k" using k(3) unfolding subset_eq by auto + hence "\X\bb ` k. x \ X" using bb k(2) by blast + hence "x \ \(bb ` k)" using Union_iff[of x "bb ` k"] by auto + } + ultimately show "\f'\f. finite f' \ s \ \f'" using bb k(2) by (rule_tac x="bb ` k" in exI) auto +qed + +subsection{* Bolzano-Weierstrass property. *} + +lemma heine_borel_imp_bolzano_weierstrass: + assumes "\f. (\t \ f. open t) \ s \ (\ f) --> (\f'. f' \ f \ finite f' \ s \ (\ f'))" + "infinite t" "t \ s" + shows "\x \ s. x islimpt t" +proof(rule ccontr) + assume "\ (\x \ s. x islimpt t)" + then obtain f where f:"\x\s. x \ f x \ open (f x) \ (\y\t. y \ f x \ y = x)" unfolding islimpt_def + using bchoice[of s "\ x T. x \ T \ open T \ (\y\t. y \ T \ y = x)"] by auto + obtain g where g:"g\{t. \x. x \ s \ t = f x}" "finite g" "s \ \g" + using assms(1)[THEN spec[where x="{t. \x. x\s \ t = f x}"]] using f by auto + from g(1,3) have g':"\x\g. \xa \ s. x = f xa" by auto + { fix x y assume "x\t" "y\t" "f x = f y" + hence "x \ f x" "y \ f x \ y = x" using f[THEN bspec[where x=x]] and `t\s` by auto + hence "x = y" using `f x = f y` and f[THEN bspec[where x=y]] and `y\t` and `t\s` by auto } + hence "infinite (f ` t)" using assms(2) using finite_imageD[unfolded inj_on_def, of f t] by auto + moreover + { fix x assume "x\t" "f x \ g" + from g(3) assms(3) `x\t` obtain h where "h\g" and "x\h" by auto + then obtain y where "y\s" "h = f y" using g'[THEN bspec[where x=h]] by auto + hence "y = x" using f[THEN bspec[where x=y]] and `x\t` and `x\h`[unfolded `h = f y`] by auto + hence False using `f x \ g` `h\g` unfolding `h = f y` by auto } + hence "f ` t \ g" by auto + ultimately show False using g(2) using finite_subset by auto +qed + +subsection{* Complete the chain of compactness variants. *} + +primrec helper_2::"(real \ 'a::metric_space) \ nat \ 'a" where + "helper_2 beyond 0 = beyond 0" | + "helper_2 beyond (Suc n) = beyond (dist undefined (helper_2 beyond n) + 1 )" + +lemma bolzano_weierstrass_imp_bounded: fixes s::"'a::metric_space set" + assumes "\t. infinite t \ t \ s --> (\x \ s. x islimpt t)" + shows "bounded s" +proof(rule ccontr) + assume "\ bounded s" + then obtain beyond where "\a. beyond a \s \ \ dist undefined (beyond a) \ a" + unfolding bounded_any_center [where a=undefined] + apply simp using choice[of "\a x. x\s \ \ dist undefined x \ a"] by auto + hence beyond:"\a. beyond a \s" "\a. dist undefined (beyond a) > a" + unfolding linorder_not_le by auto + def x \ "helper_2 beyond" + + { fix m n ::nat assume "mn" + have "1 < dist (x m) (x n)" + proof(cases "mn` by auto + hence "1 < dist undefined (x m) - dist undefined (x n)" using *[of n m] by auto + thus ?thesis using dist_triangle2 [of undefined "x m" "x n"] by arith + qed } note ** = this + { fix a b assume "x a = x b" "a \ b" + hence False using **[of a b] by auto } + hence "inj x" unfolding inj_on_def by auto + moreover + { fix n::nat + have "x n \ s" + proof(cases "n = 0") + case True thus ?thesis unfolding x_def using beyond by auto + next + case False then obtain z where "n = Suc z" using not0_implies_Suc by auto + thus ?thesis unfolding x_def using beyond by auto + qed } + ultimately have "infinite (range x) \ range x \ s" unfolding x_def using range_inj_infinite[of "helper_2 beyond"] using beyond(1) by auto + + then obtain l where "l\s" and l:"l islimpt range x" using assms[THEN spec[where x="range x"]] by auto + then obtain y where "x y \ l" and y:"dist (x y) l < 1/2" unfolding islimpt_approachable apply(erule_tac x="1/2" in allE) by auto + then obtain z where "x z \ l" and z:"dist (x z) l < dist (x y) l" using l[unfolded islimpt_approachable, THEN spec[where x="dist (x y) l"]] + unfolding dist_nz by auto + show False using y and z and dist_triangle_half_l[of "x y" l 1 "x z"] and **[of y z] by auto +qed + +lemma sequence_infinite_lemma: + fixes l :: "'a::metric_space" (* TODO: generalize *) + assumes "\n::nat. (f n \ l)" "(f ---> l) sequentially" + shows "infinite {y. (\ n. y = f n)}" +proof(rule ccontr) + let ?A = "(\x. dist x l) ` {y. \n. y = f n}" + assume "\ infinite {y. \n. y = f n}" + hence **:"finite ?A" "?A \ {}" by auto + obtain k where k:"dist (f k) l = Min ?A" using Min_in[OF **] by auto + have "0 < Min ?A" using assms(1) unfolding dist_nz unfolding Min_gr_iff[OF **] by auto + then obtain N where "dist (f N) l < Min ?A" using assms(2)[unfolded Lim_sequentially, THEN spec[where x="Min ?A"]] by auto + moreover have "dist (f N) l \ ?A" by auto + ultimately show False using Min_le[OF **(1), of "dist (f N) l"] by auto +qed + +lemma sequence_unique_limpt: + fixes l :: "'a::metric_space" (* TODO: generalize *) + assumes "\n::nat. (f n \ l)" "(f ---> l) sequentially" "l' islimpt {y. (\n. y = f n)}" + shows "l' = l" +proof(rule ccontr) + def e \ "dist l' l" + assume "l' \ l" hence "e>0" unfolding dist_nz e_def by auto + then obtain N::nat where N:"\n\N. dist (f n) l < e / 2" + using assms(2)[unfolded Lim_sequentially, THEN spec[where x="e/2"]] by auto + def d \ "Min (insert (e/2) ((\n. if dist (f n) l' = 0 then e/2 else dist (f n) l') ` {0 .. N}))" + have "d>0" using `e>0` unfolding d_def e_def using zero_le_dist[of _ l', unfolded order_le_less] by auto + obtain k where k:"f k \ l'" "dist (f k) l' < d" using `d>0` and assms(3)[unfolded islimpt_approachable, THEN spec[where x="d"]] by auto + have "k\N" using k(1)[unfolded dist_nz] using k(2)[unfolded d_def] + by force + hence "dist l' l < e" using N[THEN spec[where x=k]] using k(2)[unfolded d_def] and dist_triangle_half_r[of "f k" l' e l] by auto + thus False unfolding e_def by auto +qed + +lemma bolzano_weierstrass_imp_closed: + fixes s :: "'a::metric_space set" (* TODO: can this be generalized? *) + assumes "\t. infinite t \ t \ s --> (\x \ s. x islimpt t)" + shows "closed s" +proof- + { fix x l assume as: "\n::nat. x n \ s" "(x ---> l) sequentially" + hence "l \ s" + proof(cases "\n. x n \ l") + case False thus "l\s" using as(1) by auto + next + case True note cas = this + with as(2) have "infinite {y. \n. y = x n}" using sequence_infinite_lemma[of x l] by auto + then obtain l' where "l'\s" "l' islimpt {y. \n. y = x n}" using assms[THEN spec[where x="{y. \n. y = x n}"]] as(1) by auto + thus "l\s" using sequence_unique_limpt[of x l l'] using as cas by auto + qed } + thus ?thesis unfolding closed_sequential_limits by fast +qed + +text{* Hence express everything as an equivalence. *} + +lemma compact_eq_heine_borel: + fixes s :: "'a::heine_borel set" + shows "compact s \ + (\f. (\t \ f. open t) \ s \ (\ f) + --> (\f'. f' \ f \ finite f' \ s \ (\ f')))" (is "?lhs = ?rhs") +proof + assume ?lhs thus ?rhs using compact_imp_heine_borel[of s] by blast +next + assume ?rhs + hence "\t. infinite t \ t \ s \ (\x\s. x islimpt t)" + by (blast intro: heine_borel_imp_bolzano_weierstrass[of s]) + thus ?lhs using bolzano_weierstrass_imp_bounded[of s] bolzano_weierstrass_imp_closed[of s] bounded_closed_imp_compact[of s] by blast +qed + +lemma compact_eq_bolzano_weierstrass: + fixes s :: "'a::heine_borel set" + shows "compact s \ (\t. infinite t \ t \ s --> (\x \ s. x islimpt t))" (is "?lhs = ?rhs") +proof + assume ?lhs thus ?rhs unfolding compact_eq_heine_borel using heine_borel_imp_bolzano_weierstrass[of s] by auto +next + assume ?rhs thus ?lhs using bolzano_weierstrass_imp_bounded bolzano_weierstrass_imp_closed bounded_closed_imp_compact by auto +qed + +lemma compact_eq_bounded_closed: + fixes s :: "'a::heine_borel set" + shows "compact s \ bounded s \ closed s" (is "?lhs = ?rhs") +proof + assume ?lhs thus ?rhs unfolding compact_eq_bolzano_weierstrass using bolzano_weierstrass_imp_bounded bolzano_weierstrass_imp_closed by auto +next + assume ?rhs thus ?lhs using bounded_closed_imp_compact by auto +qed + +lemma compact_imp_bounded: + fixes s :: "'a::metric_space set" + shows "compact s ==> bounded s" +proof - + assume "compact s" + hence "\f. (\t\f. open t) \ s \ \f \ (\f'\f. finite f' \ s \ \f')" + by (rule compact_imp_heine_borel) + hence "\t. infinite t \ t \ s \ (\x \ s. x islimpt t)" + using heine_borel_imp_bolzano_weierstrass[of s] by auto + thus "bounded s" + by (rule bolzano_weierstrass_imp_bounded) +qed + +lemma compact_imp_closed: + fixes s :: "'a::metric_space set" + shows "compact s ==> closed s" +proof - + assume "compact s" + hence "\f. (\t\f. open t) \ s \ \f \ (\f'\f. finite f' \ s \ \f')" + by (rule compact_imp_heine_borel) + hence "\t. infinite t \ t \ s \ (\x \ s. x islimpt t)" + using heine_borel_imp_bolzano_weierstrass[of s] by auto + thus "closed s" + by (rule bolzano_weierstrass_imp_closed) +qed + +text{* In particular, some common special cases. *} + +lemma compact_empty[simp]: + "compact {}" + unfolding compact_def + by simp + +(* TODO: can any of the next 3 lemmas be generalized to metric spaces? *) + + (* FIXME : Rename *) +lemma compact_union[intro]: + fixes s t :: "'a::heine_borel set" + shows "compact s \ compact t ==> compact (s \ t)" + unfolding compact_eq_bounded_closed + using bounded_Un[of s t] + using closed_Un[of s t] + by simp + +lemma compact_inter[intro]: + fixes s t :: "'a::heine_borel set" + shows "compact s \ compact t ==> compact (s \ t)" + unfolding compact_eq_bounded_closed + using bounded_Int[of s t] + using closed_Int[of s t] + by simp + +lemma compact_inter_closed[intro]: + fixes s t :: "'a::heine_borel set" + shows "compact s \ closed t ==> compact (s \ t)" + unfolding compact_eq_bounded_closed + using closed_Int[of s t] + using bounded_subset[of "s \ t" s] + by blast + +lemma closed_inter_compact[intro]: + fixes s t :: "'a::heine_borel set" + shows "closed s \ compact t ==> compact (s \ t)" +proof- + assume "closed s" "compact t" + moreover + have "s \ t = t \ s" by auto ultimately + show ?thesis + using compact_inter_closed[of t s] + by auto +qed + +lemma closed_sing [simp]: + fixes a :: "'a::metric_space" + shows "closed {a}" + apply (clarsimp simp add: closed_def open_dist) + apply (rule ccontr) + apply (drule_tac x="dist x a" in spec) + apply (simp add: dist_nz dist_commute) + done + +lemma finite_imp_closed: + fixes s :: "'a::metric_space set" + shows "finite s ==> closed s" +proof (induct set: finite) + case empty show "closed {}" by simp +next + case (insert x F) + hence "closed ({x} \ F)" by (simp only: closed_Un closed_sing) + thus "closed (insert x F)" by simp +qed + +lemma finite_imp_compact: + fixes s :: "'a::heine_borel set" + shows "finite s ==> compact s" + unfolding compact_eq_bounded_closed + using finite_imp_closed finite_imp_bounded + by blast + +lemma compact_sing [simp]: "compact {a}" + unfolding compact_def o_def subseq_def + by (auto simp add: tendsto_const) + +lemma compact_cball[simp]: + fixes x :: "'a::heine_borel" + shows "compact(cball x e)" + using compact_eq_bounded_closed bounded_cball closed_cball + by blast + +lemma compact_frontier_bounded[intro]: + fixes s :: "'a::heine_borel set" + shows "bounded s ==> compact(frontier s)" + unfolding frontier_def + using compact_eq_bounded_closed + by blast + +lemma compact_frontier[intro]: + fixes s :: "'a::heine_borel set" + shows "compact s ==> compact (frontier s)" + using compact_eq_bounded_closed compact_frontier_bounded + by blast + +lemma frontier_subset_compact: + fixes s :: "'a::heine_borel set" + shows "compact s ==> frontier s \ s" + using frontier_subset_closed compact_eq_bounded_closed + by blast + +lemma open_delete: + fixes s :: "'a::metric_space set" + shows "open s ==> open(s - {x})" + using open_Diff[of s "{x}"] closed_sing + by blast + +text{* Finite intersection property. I could make it an equivalence in fact. *} + +lemma compact_imp_fip: + fixes s :: "'a::heine_borel set" + assumes "compact s" "\t \ f. closed t" + "\f'. finite f' \ f' \ f --> (s \ (\ f') \ {})" + shows "s \ (\ f) \ {}" +proof + assume as:"s \ (\ f) = {}" + hence "s \ \op - UNIV ` f" by auto + moreover have "Ball (op - UNIV ` f) open" using open_Diff closed_Diff using assms(2) by auto + ultimately obtain f' where f':"f' \ op - UNIV ` f" "finite f'" "s \ \f'" using assms(1)[unfolded compact_eq_heine_borel, THEN spec[where x="(\t. UNIV - t) ` f"]] by auto + hence "finite (op - UNIV ` f') \ op - UNIV ` f' \ f" by(auto simp add: Diff_Diff_Int) + hence "s \ \op - UNIV ` f' \ {}" using assms(3)[THEN spec[where x="op - UNIV ` f'"]] by auto + thus False using f'(3) unfolding subset_eq and Union_iff by blast +qed + +subsection{* Bounded closed nest property (proof does not use Heine-Borel). *} + +lemma bounded_closed_nest: + assumes "\n. closed(s n)" "\n. (s n \ {})" + "(\m n. m \ n --> s n \ s m)" "bounded(s 0)" + shows "\a::'a::heine_borel. \n::nat. a \ s(n)" +proof- + 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" "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 + { fix e::real assume "e>0" + 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 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) + using assms(3) apply(erule_tac x=n in allE)apply( erule_tac x="r (max N n)" in allE) by auto + ultimately have "\y\s n. dist y l < e" by auto + } + hence "l \ s n" using closed_approachable[of "s n" l] assms(1) by blast + } + thus ?thesis by auto +qed + +text{* Decreasing case does not even need compactness, just completeness. *} + +lemma decreasing_closed_nest: + assumes "\n. closed(s n)" + "\n. (s n \ {})" + "\m n. m \ n --> s n \ s m" + "\e>0. \n. \x \ (s n). \ y \ (s n). dist x y < e" + shows "\a::'a::heine_borel. \n::nat. a \ s n" +proof- + have "\n. \ x. x\s n" using assms(2) by auto + hence "\t. \n. t n \ s n" using choice[of "\ n x. x \ s n"] by auto + then obtain t where t: "\n. t n \ s n" by auto + { fix e::real assume "e>0" + then obtain N where N:"\x\s N. \y\s N. dist x y < e" using assms(4) by auto + { fix m n ::nat assume "N \ m \ N \ n" + hence "t m \ s N" "t n \ s N" using assms(3) t unfolding subset_eq t by blast+ + hence "dist (t m) (t n) < e" using N by auto + } + hence "\N. \m n. N \ m \ N \ n \ dist (t m) (t n) < e" by auto + } + hence "Cauchy t" unfolding cauchy_def by auto + then obtain l where l:"(t ---> l) sequentially" using complete_univ unfolding complete_def by auto + { fix n::nat + { fix e::real assume "e>0" + then obtain N::nat where N:"\n\N. dist (t n) l < e" using l[unfolded Lim_sequentially] by auto + have "t (max n N) \ s n" using assms(3) unfolding subset_eq apply(erule_tac x=n in allE) apply (erule_tac x="max n N" in allE) using t by auto + hence "\y\s n. dist y l < e" apply(rule_tac x="t (max n N)" in bexI) using N by auto + } + hence "l \ s n" using closed_approachable[of "s n" l] assms(1) by auto + } + then show ?thesis by auto +qed + +text{* Strengthen it to the intersection actually being a singleton. *} + +lemma decreasing_closed_nest_sing: + assumes "\n. closed(s n)" + "\n. s n \ {}" + "\m n. m \ n --> s n \ s m" + "\e>0. \n. \x \ (s n). \ y\(s n). dist x y < e" + shows "\a::'a::heine_borel. \ {t. (\n::nat. t = s n)} = {a}" +proof- + obtain a where a:"\n. a \ s n" using decreasing_closed_nest[of s] using assms by auto + { fix b assume b:"b \ \{t. \n. t = s n}" + { fix e::real assume "e>0" + hence "dist a b < e" using assms(4 )using b using a by blast + } + hence "dist a b = 0" by (metis dist_eq_0_iff dist_nz real_less_def) + } + with a have "\{t. \n. t = s n} = {a}" by auto + thus ?thesis by auto +qed + +text{* Cauchy-type criteria for uniform convergence. *} + +lemma uniformly_convergent_eq_cauchy: fixes s::"nat \ 'b \ 'a::heine_borel" shows + "(\l. \e>0. \N. \n x. N \ n \ P x --> dist(s n x)(l x) < e) \ + (\e>0. \N. \m n x. N \ m \ N \ n \ P x --> dist (s m x) (s n x) < e)" (is "?lhs = ?rhs") +proof(rule) + assume ?lhs + then obtain l where l:"\e>0. \N. \n x. N \ n \ P x \ dist (s n x) (l x) < e" by auto + { fix e::real assume "e>0" + then obtain N::nat where N:"\n x. N \ n \ P x \ dist (s n x) (l x) < e / 2" using l[THEN spec[where x="e/2"]] by auto + { fix n m::nat and x::"'b" assume "N \ m \ N \ n \ P x" + hence "dist (s m x) (s n x) < e" + using N[THEN spec[where x=m], THEN spec[where x=x]] + using N[THEN spec[where x=n], THEN spec[where x=x]] + using dist_triangle_half_l[of "s m x" "l x" e "s n x"] by auto } + hence "\N. \m n x. N \ m \ N \ n \ P x --> dist (s m x) (s n x) < e" by auto } + thus ?rhs by auto +next + assume ?rhs + hence "\x. P x \ Cauchy (\n. s n x)" unfolding cauchy_def apply auto by (erule_tac x=e in allE)auto + then obtain l where l:"\x. P x \ ((\n. s n x) ---> l x) sequentially" unfolding convergent_eq_cauchy[THEN sym] + using choice[of "\x l. P x \ ((\n. s n x) ---> l) sequentially"] by auto + { fix e::real assume "e>0" + then obtain N where N:"\m n x. N \ m \ N \ n \ P x \ dist (s m x) (s n x) < e/2" + using `?rhs`[THEN spec[where x="e/2"]] by auto + { fix x assume "P x" + then obtain M where M:"\n\M. dist (s n x) (l x) < e/2" + using l[THEN spec[where x=x], unfolded Lim_sequentially] using `e>0` by(auto elim!: allE[where x="e/2"]) + fix n::nat assume "n\N" + hence "dist(s n x)(l x) < e" using `P x`and N[THEN spec[where x=n], THEN spec[where x="N+M"], THEN spec[where x=x]] + using M[THEN spec[where x="N+M"]] and dist_triangle_half_l[of "s n x" "s (N+M) x" e "l x"] by (auto simp add: dist_commute) } + hence "\N. \n x. N \ n \ P x \ dist(s n x)(l x) < e" by auto } + thus ?lhs by auto +qed + +lemma uniformly_cauchy_imp_uniformly_convergent: + fixes s :: "nat \ 'a \ 'b::heine_borel" + assumes "\e>0.\N. \m (n::nat) x. N \ m \ N \ n \ P x --> dist(s m x)(s n x) < e" + "\x. P x --> (\e>0. \N. \n. N \ n --> dist(s n x)(l x) < e)" + shows "\e>0. \N. \n x. N \ n \ P x --> dist(s n x)(l x) < e" +proof- + obtain l' where l:"\e>0. \N. \n x. N \ n \ P x \ dist (s n x) (l' x) < e" + using assms(1) unfolding uniformly_convergent_eq_cauchy[THEN sym] by auto + moreover + { fix x assume "P x" + hence "l x = l' x" using Lim_unique[OF trivial_limit_sequentially, of "\n. s n x" "l x" "l' x"] + using l and assms(2) unfolding Lim_sequentially by blast } + ultimately show ?thesis by auto +qed + +subsection{* Define continuity over a net to take in restrictions of the set. *} + +definition + continuous :: "'a::t2_space net \ ('a \ 'b::topological_space) \ bool" where + "continuous net f \ (f ---> f(netlimit net)) net" + +lemma continuous_trivial_limit: + "trivial_limit net ==> continuous net f" + unfolding continuous_def tendsto_def trivial_limit_eq by auto + +lemma continuous_within: "continuous (at x within s) f \ (f ---> f(x)) (at x within s)" + unfolding continuous_def + unfolding tendsto_def + using netlimit_within[of x s] + by (cases "trivial_limit (at x within s)") (auto simp add: trivial_limit_eventually) + +lemma continuous_at: "continuous (at x) f \ (f ---> f(x)) (at x)" + using continuous_within [of x UNIV f] by (simp add: within_UNIV) + +lemma continuous_at_within: + assumes "continuous (at x) f" shows "continuous (at x within s) f" + using assms unfolding continuous_at continuous_within + by (rule Lim_at_within) + +text{* Derive the epsilon-delta forms, which we often use as "definitions" *} + +lemma continuous_within_eps_delta: + "continuous (at x within s) f \ (\e>0. \d>0. \x'\ s. dist x' x < d --> dist (f x') (f x) < e)" + unfolding continuous_within and Lim_within + apply auto unfolding dist_nz[THEN sym] apply(auto elim!:allE) apply(rule_tac x=d in exI) by auto + +lemma continuous_at_eps_delta: "continuous (at x) f \ (\e>0. \d>0. + \x'. dist x' x < d --> dist(f x')(f x) < e)" + using continuous_within_eps_delta[of x UNIV f] + unfolding within_UNIV by blast + +text{* Versions in terms of open balls. *} + +lemma continuous_within_ball: + "continuous (at x within s) f \ (\e>0. \d>0. + f ` (ball x d \ s) \ ball (f x) e)" (is "?lhs = ?rhs") +proof + assume ?lhs + { fix e::real assume "e>0" + then obtain d where d: "d>0" "\xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" + using `?lhs`[unfolded continuous_within Lim_within] by auto + { fix y assume "y\f ` (ball x d \ s)" + hence "y \ ball (f x) e" using d(2) unfolding dist_nz[THEN sym] + apply (auto simp add: dist_commute mem_ball) apply(erule_tac x=xa in ballE) apply auto using `e>0` by auto + } + hence "\d>0. f ` (ball x d \ s) \ ball (f x) e" using `d>0` unfolding subset_eq ball_def by (auto simp add: dist_commute) } + thus ?rhs by auto +next + assume ?rhs thus ?lhs unfolding continuous_within Lim_within ball_def subset_eq + apply (auto simp add: dist_commute) apply(erule_tac x=e in allE) by auto +qed + +lemma continuous_at_ball: + "continuous (at x) f \ (\e>0. \d>0. f ` (ball x d) \ ball (f x) e)" (is "?lhs = ?rhs") +proof + assume ?lhs thus ?rhs unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball + apply auto apply(erule_tac x=e in allE) apply auto apply(rule_tac x=d in exI) apply auto apply(erule_tac x=xa in allE) apply (auto simp add: dist_commute dist_nz) + unfolding dist_nz[THEN sym] by auto +next + assume ?rhs thus ?lhs unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball + apply auto apply(erule_tac x=e in allE) apply auto apply(rule_tac x=d in exI) apply auto apply(erule_tac x="f xa" in allE) by (auto simp add: dist_commute dist_nz) +qed + +text{* For setwise continuity, just start from the epsilon-delta definitions. *} + +definition + continuous_on :: "'a::metric_space set \ ('a \ 'b::metric_space) \ bool" where + "continuous_on s f \ (\x \ s. \e>0. \d::real>0. \x' \ s. dist x' x < d --> dist (f x') (f x) < e)" + + +definition + uniformly_continuous_on :: + "'a::metric_space set \ ('a \ 'b::metric_space) \ bool" where + "uniformly_continuous_on s f \ + (\e>0. \d>0. \x\s. \ x'\s. dist x' x < d + --> dist (f x') (f x) < e)" + +text{* Some simple consequential lemmas. *} + +lemma uniformly_continuous_imp_continuous: + " uniformly_continuous_on s f ==> continuous_on s f" + unfolding uniformly_continuous_on_def continuous_on_def by blast + +lemma continuous_at_imp_continuous_within: + "continuous (at x) f ==> continuous (at x within s) f" + unfolding continuous_within continuous_at using Lim_at_within by auto + +lemma continuous_at_imp_continuous_on: assumes "(\x \ s. continuous (at x) f)" + shows "continuous_on s f" +proof(simp add: continuous_at continuous_on_def, rule, rule, rule) + fix x and e::real assume "x\s" "e>0" + hence "eventually (\xa. dist (f xa) (f x) < e) (at x)" using assms unfolding continuous_at tendsto_iff by auto + then obtain d where d:"d>0" "\xa. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" unfolding eventually_at by auto + { fix x' assume "\ 0 < dist x' x" + hence "x=x'" + using dist_nz[of x' x] by auto + hence "dist (f x') (f x) < e" using `e>0` by auto + } + thus "\d>0. \x'\s. dist x' x < d \ dist (f x') (f x) < e" using d by auto +qed + +lemma continuous_on_eq_continuous_within: + "continuous_on s f \ (\x \ s. continuous (at x within s) f)" (is "?lhs = ?rhs") +proof + assume ?rhs + { fix x assume "x\s" + fix e::real assume "e>0" + assume "\d>0. \xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" + then obtain d where "d>0" and d:"\xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" by auto + { fix x' assume as:"x'\s" "dist x' x < d" + hence "dist (f x') (f x) < e" using `e>0` d `x'\s` dist_eq_0_iff[of x' x] zero_le_dist[of x' x] as(2) by (metis dist_eq_0_iff dist_nz) } + hence "\d>0. \x'\s. dist x' x < d \ dist (f x') (f x) < e" using `d>0` by auto + } + thus ?lhs using `?rhs` unfolding continuous_on_def continuous_within Lim_within by auto +next + assume ?lhs + thus ?rhs unfolding continuous_on_def continuous_within Lim_within by blast +qed + +lemma continuous_on: + "continuous_on s f \ (\x \ s. (f ---> f(x)) (at x within s))" + by (auto simp add: continuous_on_eq_continuous_within continuous_within) + +lemma continuous_on_eq_continuous_at: + "open s ==> (continuous_on s f \ (\x \ s. continuous (at x) f))" + by (auto simp add: continuous_on continuous_at Lim_within_open) + +lemma continuous_within_subset: + "continuous (at x within s) f \ t \ s + ==> continuous (at x within t) f" + unfolding continuous_within by(metis Lim_within_subset) + +lemma continuous_on_subset: + "continuous_on s f \ t \ s ==> continuous_on t f" + unfolding continuous_on by (metis subset_eq Lim_within_subset) + +lemma continuous_on_interior: + "continuous_on s f \ x \ interior s ==> continuous (at x) f" +unfolding interior_def +apply simp +by (meson continuous_on_eq_continuous_at continuous_on_subset) + +lemma continuous_on_eq: + "(\x \ s. f x = g x) \ continuous_on s f + ==> continuous_on s g" + by (simp add: continuous_on_def) + +text{* Characterization of various kinds of continuity in terms of sequences. *} + +(* \ could be generalized, but \ requires metric space *) +lemma continuous_within_sequentially: + fixes f :: "'a::metric_space \ 'b::metric_space" + shows "continuous (at a within s) f \ + (\x. (\n::nat. x n \ s) \ (x ---> a) sequentially + --> ((f o x) ---> f a) sequentially)" (is "?lhs = ?rhs") +proof + assume ?lhs + { fix x::"nat \ 'a" assume x:"\n. x n \ s" "\e>0. \N. \n\N. dist (x n) a < e" + fix e::real assume "e>0" + from `?lhs` obtain d where "d>0" and d:"\x\s. 0 < dist x a \ dist x a < d \ dist (f x) (f a) < e" unfolding continuous_within Lim_within using `e>0` by auto + from x(2) `d>0` obtain N where N:"\n\N. dist (x n) a < d" by auto + hence "\N. \n\N. dist ((f \ x) n) (f a) < e" + apply(rule_tac x=N in exI) using N d apply auto using x(1) + apply(erule_tac x=n in allE) apply(erule_tac x=n in allE) + apply(erule_tac x="x n" in ballE) apply auto unfolding dist_nz[THEN sym] apply auto using `e>0` by auto + } + thus ?rhs unfolding continuous_within unfolding Lim_sequentially by simp +next + assume ?rhs + { fix e::real assume "e>0" + assume "\ (\d>0. \x\s. 0 < dist x a \ dist x a < d \ dist (f x) (f a) < e)" + hence "\d. \x. d>0 \ x\s \ (0 < dist x a \ dist x a < d \ \ dist (f x) (f a) < e)" by blast + then obtain x where x:"\d>0. x d \ s \ (0 < dist (x d) a \ dist (x d) a < d \ \ dist (f (x d)) (f a) < e)" + using choice[of "\d x.0 x\s \ (0 < dist x a \ dist x a < d \ \ dist (f x) (f a) < e)"] by auto + { fix d::real assume "d>0" + hence "\N::nat. inverse (real (N + 1)) < d" using real_arch_inv[of d] by (auto, rule_tac x="n - 1" in exI)auto + then obtain N::nat where N:"inverse (real (N + 1)) < d" by auto + { fix n::nat assume n:"n\N" + hence "dist (x (inverse (real (n + 1)))) a < inverse (real (n + 1))" using x[THEN spec[where x="inverse (real (n + 1))"]] by auto + moreover have "inverse (real (n + 1)) < d" using N n 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) + ultimately have "dist (x (inverse (real (n + 1)))) a < d" by auto + } + hence "\N::nat. \n\N. dist (x (inverse (real (n + 1)))) a < d" by auto + } + hence "(\n::nat. x (inverse (real (n + 1))) \ s) \ (\e>0. \N::nat. \n\N. dist (x (inverse (real (n + 1)))) a < e)" using x by auto + hence "\e>0. \N::nat. \n\N. dist (f (x (inverse (real (n + 1))))) (f a) < e" using `?rhs`[THEN spec[where x="\n::nat. x (inverse (real (n+1)))"], unfolded Lim_sequentially] by auto + hence "False" apply(erule_tac x=e in allE) using `e>0` using x by auto + } + thus ?lhs unfolding continuous_within unfolding Lim_within unfolding Lim_sequentially by blast +qed + +lemma continuous_at_sequentially: + fixes f :: "'a::metric_space \ 'b::metric_space" + shows "continuous (at a) f \ (\x. (x ---> a) sequentially + --> ((f o x) ---> f a) sequentially)" + using continuous_within_sequentially[of a UNIV f] unfolding within_UNIV by auto + +lemma continuous_on_sequentially: + "continuous_on s f \ (\x. \a \ s. (\n. x(n) \ s) \ (x ---> a) sequentially + --> ((f o x) ---> f(a)) sequentially)" (is "?lhs = ?rhs") +proof + assume ?rhs thus ?lhs using continuous_within_sequentially[of _ s f] unfolding continuous_on_eq_continuous_within by auto +next + assume ?lhs thus ?rhs unfolding continuous_on_eq_continuous_within using continuous_within_sequentially[of _ s f] by auto +qed + +lemma uniformly_continuous_on_sequentially: + fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" + shows "uniformly_continuous_on s f \ (\x y. (\n. x n \ s) \ (\n. y n \ s) \ + ((\n. x n - y n) ---> 0) sequentially + \ ((\n. f(x n) - f(y n)) ---> 0) sequentially)" (is "?lhs = ?rhs") +proof + assume ?lhs + { fix x y assume x:"\n. x n \ s" and y:"\n. y n \ s" and xy:"((\n. x n - y n) ---> 0) sequentially" + { fix e::real assume "e>0" + then obtain d where "d>0" and d:"\x\s. \x'\s. dist x' x < d \ dist (f x') (f x) < e" + using `?lhs`[unfolded uniformly_continuous_on_def, THEN spec[where x=e]] by auto + obtain N where N:"\n\N. norm (x n - y n - 0) < d" using xy[unfolded Lim_sequentially dist_norm] and `d>0` by auto + { fix n assume "n\N" + hence "norm (f (x n) - f (y n) - 0) < e" + using N[THEN spec[where x=n]] using d[THEN bspec[where x="x n"], THEN bspec[where x="y n"]] using x and y + unfolding dist_commute and dist_norm by simp } + hence "\N. \n\N. norm (f (x n) - f (y n) - 0) < e" by auto } + hence "((\n. f(x n) - f(y n)) ---> 0) sequentially" unfolding Lim_sequentially and dist_norm by auto } + thus ?rhs by auto +next + assume ?rhs + { assume "\ ?lhs" + then obtain e where "e>0" "\d>0. \x\s. \x'\s. dist x' x < d \ \ dist (f x') (f x) < e" unfolding uniformly_continuous_on_def by auto + then obtain fa where fa:"\x. 0 < x \ fst (fa x) \ s \ snd (fa x) \ s \ dist (fst (fa x)) (snd (fa x)) < x \ \ dist (f (fst (fa x))) (f (snd (fa x))) < e" + using choice[of "\d x. d>0 \ fst x \ s \ snd x \ s \ dist (snd x) (fst x) < d \ \ dist (f (snd x)) (f (fst x)) < e"] unfolding Bex_def + by (auto simp add: dist_commute) + def x \ "\n::nat. fst (fa (inverse (real n + 1)))" + def y \ "\n::nat. snd (fa (inverse (real n + 1)))" + have xyn:"\n. x n \ s \ y n \ s" and xy0:"\n. dist (x n) (y n) < inverse (real n + 1)" and fxy:"\n. \ dist (f (x n)) (f (y n)) < e" + unfolding x_def and y_def using fa by auto + have 1:"\(x::'a) y. dist (x - y) 0 = dist x y" unfolding dist_norm by auto + have 2:"\(x::'b) y. dist (x - y) 0 = dist x y" unfolding dist_norm by auto + { fix e::real assume "e>0" + then obtain N::nat where "N \ 0" and N:"0 < inverse (real N) \ inverse (real N) < e" unfolding real_arch_inv[of e] by auto + { fix n::nat assume "n\N" + hence "inverse (real n + 1) < inverse (real N)" using real_of_nat_ge_zero and `N\0` by auto + also have "\ < e" using N by auto + finally have "inverse (real n + 1) < e" by auto + hence "dist (x n - y n) 0 < e" unfolding 1 using xy0[THEN spec[where x=n]] by auto } + hence "\N. \n\N. dist (x n - y n) 0 < e" by auto } + hence "\e>0. \N. \n\N. dist (f (x n) - f (y n)) 0 < e" using `?rhs`[THEN spec[where x=x], THEN spec[where x=y]] and xyn unfolding Lim_sequentially by auto + hence False unfolding 2 using fxy and `e>0` by auto } + thus ?lhs unfolding uniformly_continuous_on_def by blast +qed + +text{* The usual transformation theorems. *} + +lemma continuous_transform_within: + fixes f g :: "'a::metric_space \ 'b::metric_space" + assumes "0 < d" "x \ s" "\x' \ s. dist x' x < d --> f x' = g x'" + "continuous (at x within s) f" + shows "continuous (at x within s) g" +proof- + { fix e::real assume "e>0" + then obtain d' where d':"d'>0" "\xa\s. 0 < dist xa x \ dist xa x < d' \ dist (f xa) (f x) < e" using assms(4) unfolding continuous_within Lim_within by auto + { fix x' assume "x'\s" "0 < dist x' x" "dist x' x < (min d d')" + hence "dist (f x') (g x) < e" using assms(2,3) apply(erule_tac x=x in ballE) using d' by auto } + hence "\xa\s. 0 < dist xa x \ dist xa x < (min d d') \ dist (f xa) (g x) < e" by blast + hence "\d>0. \xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (g x) < e" using `d>0` `d'>0` by(rule_tac x="min d d'" in exI)auto } + hence "(f ---> g x) (at x within s)" unfolding Lim_within using assms(1) by auto + thus ?thesis unfolding continuous_within using Lim_transform_within[of d s x f g "g x"] using assms by blast +qed + +lemma continuous_transform_at: + fixes f g :: "'a::metric_space \ 'b::metric_space" + assumes "0 < d" "\x'. dist x' x < d --> f x' = g x'" + "continuous (at x) f" + shows "continuous (at x) g" +proof- + { fix e::real assume "e>0" + then obtain d' where d':"d'>0" "\xa. 0 < dist xa x \ dist xa x < d' \ dist (f xa) (f x) < e" using assms(3) unfolding continuous_at Lim_at by auto + { fix x' assume "0 < dist x' x" "dist x' x < (min d d')" + hence "dist (f x') (g x) < e" using assms(2) apply(erule_tac x=x in allE) using d' by auto + } + hence "\xa. 0 < dist xa x \ dist xa x < (min d d') \ dist (f xa) (g x) < e" by blast + hence "\d>0. \xa. 0 < dist xa x \ dist xa x < d \ dist (f xa) (g x) < e" using `d>0` `d'>0` by(rule_tac x="min d d'" in exI)auto + } + hence "(f ---> g x) (at x)" unfolding Lim_at using assms(1) by auto + thus ?thesis unfolding continuous_at using Lim_transform_at[of d x f g "g x"] using assms by blast +qed + +text{* Combination results for pointwise continuity. *} + +lemma continuous_const: "continuous net (\x. c)" + by (auto simp add: continuous_def Lim_const) + +lemma continuous_cmul: + fixes f :: "'a::t2_space \ 'b::real_normed_vector" + shows "continuous net f ==> continuous net (\x. c *\<^sub>R f x)" + by (auto simp add: continuous_def Lim_cmul) + +lemma continuous_neg: + fixes f :: "'a::t2_space \ 'b::real_normed_vector" + shows "continuous net f ==> continuous net (\x. -(f x))" + by (auto simp add: continuous_def Lim_neg) + +lemma continuous_add: + fixes f g :: "'a::t2_space \ 'b::real_normed_vector" + shows "continuous net f \ continuous net g \ continuous net (\x. f x + g x)" + by (auto simp add: continuous_def Lim_add) + +lemma continuous_sub: + fixes f g :: "'a::t2_space \ 'b::real_normed_vector" + shows "continuous net f \ continuous net g \ continuous net (\x. f x - g x)" + by (auto simp add: continuous_def Lim_sub) + +text{* Same thing for setwise continuity. *} + +lemma continuous_on_const: + "continuous_on s (\x. c)" + unfolding continuous_on_eq_continuous_within using continuous_const by blast + +lemma continuous_on_cmul: + fixes f :: "'a::metric_space \ 'b::real_normed_vector" + 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: + fixes f :: "'a::metric_space \ 'b::real_normed_vector" + shows "continuous_on s f \ continuous_on s (\x. - f x)" + unfolding continuous_on_eq_continuous_within using continuous_neg by blast + +lemma continuous_on_add: + fixes f g :: "'a::metric_space \ 'b::real_normed_vector" + shows "continuous_on s f \ continuous_on s g + \ continuous_on s (\x. f x + g x)" + unfolding continuous_on_eq_continuous_within using continuous_add by blast + +lemma continuous_on_sub: + fixes f g :: "'a::metric_space \ 'b::real_normed_vector" + shows "continuous_on s f \ continuous_on s g + \ continuous_on s (\x. f x - g x)" + unfolding continuous_on_eq_continuous_within using continuous_sub by blast + +text{* Same thing for uniform continuity, using sequential formulations. *} + +lemma uniformly_continuous_on_const: + "uniformly_continuous_on s (\x. c)" + unfolding uniformly_continuous_on_def by simp + +lemma uniformly_continuous_on_cmul: + fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" + (* FIXME: generalize 'a to metric_space *) + assumes "uniformly_continuous_on s f" + 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 *\<^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 scaleR_zero_right scaleR_right_diff_distrib by auto + } + thus ?thesis using assms unfolding uniformly_continuous_on_sequentially by auto +qed + +lemma dist_minus: + fixes x y :: "'a::real_normed_vector" + shows "dist (- x) (- y) = dist x y" + unfolding dist_norm minus_diff_minus norm_minus_cancel .. + +lemma uniformly_continuous_on_neg: + fixes f :: "'a::metric_space \ 'b::real_normed_vector" + shows "uniformly_continuous_on s f + ==> uniformly_continuous_on s (\x. -(f x))" + unfolding uniformly_continuous_on_def dist_minus . + +lemma uniformly_continuous_on_add: + fixes f g :: "'a::real_normed_vector \ 'b::real_normed_vector" (* FIXME: generalize 'a *) + assumes "uniformly_continuous_on s f" "uniformly_continuous_on s g" + shows "uniformly_continuous_on s (\x. f x + g x)" +proof- + { fix x y assume "((\n. f (x n) - f (y n)) ---> 0) sequentially" + "((\n. g (x n) - g (y n)) ---> 0) sequentially" + hence "((\xa. f (x xa) - f (y xa) + (g (x xa) - g (y xa))) ---> 0 + 0) sequentially" + using Lim_add[of "\ n. f (x n) - f (y n)" 0 sequentially "\ n. g (x n) - g (y n)" 0] by auto + hence "((\n. f (x n) + g (x n) - (f (y n) + g (y n))) ---> 0) sequentially" unfolding Lim_sequentially and add_diff_add [symmetric] by auto } + thus ?thesis using assms unfolding uniformly_continuous_on_sequentially by auto +qed + +lemma uniformly_continuous_on_sub: + fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" (* FIXME: generalize 'a *) + shows "uniformly_continuous_on s f \ uniformly_continuous_on s g + ==> uniformly_continuous_on s (\x. f x - g x)" + unfolding ab_diff_minus + using uniformly_continuous_on_add[of s f "\x. - g x"] + using uniformly_continuous_on_neg[of s g] by auto + +text{* Identity function is continuous in every sense. *} + +lemma continuous_within_id: + "continuous (at a within s) (\x. x)" + unfolding continuous_within by (rule Lim_at_within [OF Lim_ident_at]) + +lemma continuous_at_id: + "continuous (at a) (\x. x)" + unfolding continuous_at by (rule Lim_ident_at) + +lemma continuous_on_id: + "continuous_on s (\x. x)" + unfolding continuous_on Lim_within by auto + +lemma uniformly_continuous_on_id: + "uniformly_continuous_on s (\x. x)" + unfolding uniformly_continuous_on_def by auto + +text{* Continuity of all kinds is preserved under composition. *} + +lemma continuous_within_compose: + fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) + fixes g :: "'b::metric_space \ 'c::metric_space" + assumes "continuous (at x within s) f" "continuous (at (f x) within f ` s) g" + shows "continuous (at x within s) (g o f)" +proof- + { fix e::real assume "e>0" + with assms(2)[unfolded continuous_within Lim_within] obtain d where "d>0" and d:"\xa\f ` s. 0 < dist xa (f x) \ dist xa (f x) < d \ dist (g xa) (g (f x)) < e" by auto + from assms(1)[unfolded continuous_within Lim_within] obtain d' where "d'>0" and d':"\xa\s. 0 < dist xa x \ dist xa x < d' \ dist (f xa) (f x) < d" using `d>0` by auto + { fix y assume as:"y\s" "0 < dist y x" "dist y x < d'" + hence "dist (f y) (f x) < d" using d'[THEN bspec[where x=y]] by (auto simp add:dist_commute) + hence "dist (g (f y)) (g (f x)) < e" using as(1) d[THEN bspec[where x="f y"]] unfolding dist_nz[THEN sym] using `e>0` by auto } + hence "\d>0. \xa\s. 0 < dist xa x \ dist xa x < d \ dist (g (f xa)) (g (f x)) < e" using `d'>0` by auto } + thus ?thesis unfolding continuous_within Lim_within by auto +qed + +lemma continuous_at_compose: + fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) + fixes g :: "'b::metric_space \ 'c::metric_space" + assumes "continuous (at x) f" "continuous (at (f x)) g" + shows "continuous (at x) (g o f)" +proof- + have " continuous (at (f x) within range f) g" using assms(2) using continuous_within_subset[of "f x" UNIV g "range f", unfolded within_UNIV] by auto + thus ?thesis using assms(1) using continuous_within_compose[of x UNIV f g, unfolded within_UNIV] by auto +qed + +lemma continuous_on_compose: + "continuous_on s f \ continuous_on (f ` s) g \ continuous_on s (g o f)" + unfolding continuous_on_eq_continuous_within using continuous_within_compose[of _ s f g] by auto + +lemma uniformly_continuous_on_compose: + assumes "uniformly_continuous_on s f" "uniformly_continuous_on (f ` s) g" + shows "uniformly_continuous_on s (g o f)" +proof- + { fix e::real assume "e>0" + then obtain d where "d>0" and d:"\x\f ` s. \x'\f ` s. dist x' x < d \ dist (g x') (g x) < e" using assms(2) unfolding uniformly_continuous_on_def by auto + obtain d' where "d'>0" "\x\s. \x'\s. dist x' x < d' \ dist (f x') (f x) < d" using `d>0` using assms(1) unfolding uniformly_continuous_on_def by auto + hence "\d>0. \x\s. \x'\s. dist x' x < d \ dist ((g \ f) x') ((g \ f) x) < e" using `d>0` using d by auto } + thus ?thesis using assms unfolding uniformly_continuous_on_def by auto +qed + +text{* Continuity in terms of open preimages. *} + +lemma continuous_at_open: + fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) + shows "continuous (at x) f \ (\t. open t \ f x \ t --> (\s. open s \ x \ s \ (\x' \ s. (f x') \ t)))" (is "?lhs = ?rhs") +proof + assume ?lhs + { fix t assume as: "open t" "f x \ t" + then obtain e where "e>0" and e:"ball (f x) e \ t" unfolding open_contains_ball by auto + + obtain d where "d>0" and d:"\y. 0 < dist y x \ dist y x < d \ dist (f y) (f x) < e" using `e>0` using `?lhs`[unfolded continuous_at Lim_at open_dist] by auto + + have "open (ball x d)" using open_ball by auto + moreover have "x \ ball x d" unfolding centre_in_ball using `d>0` by simp + moreover + { fix x' assume "x'\ball x d" hence "f x' \ t" + using e[unfolded subset_eq Ball_def mem_ball, THEN spec[where x="f x'"]] d[THEN spec[where x=x']] + unfolding mem_ball apply (auto simp add: dist_commute) + unfolding dist_nz[THEN sym] using as(2) by auto } + hence "\x'\ball x d. f x' \ t" by auto + ultimately have "\s. open s \ x \ s \ (\x'\s. f x' \ t)" + apply(rule_tac x="ball x d" in exI) by simp } + thus ?rhs by auto +next + assume ?rhs + { fix e::real assume "e>0" + then obtain s where s: "open s" "x \ s" "\x'\s. f x' \ ball (f x) e" using `?rhs`[unfolded continuous_at Lim_at, THEN spec[where x="ball (f x) e"]] + unfolding centre_in_ball[of "f x" e, THEN sym] by auto + then obtain d where "d>0" and d:"ball x d \ s" unfolding open_contains_ball by auto + { fix y assume "0 < dist y x \ dist y x < d" + hence "dist (f y) (f x) < e" using d[unfolded subset_eq Ball_def mem_ball, THEN spec[where x=y]] + using s(3)[THEN bspec[where x=y], unfolded mem_ball] by (auto simp add: dist_commute) } + hence "\d>0. \xa. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" using `d>0` by auto } + thus ?lhs unfolding continuous_at Lim_at by auto +qed + +lemma continuous_on_open: + "continuous_on s f \ + (\t. openin (subtopology euclidean (f ` s)) t + --> openin (subtopology euclidean s) {x \ s. f x \ t})" (is "?lhs = ?rhs") +proof + assume ?lhs + { fix t assume as:"openin (subtopology euclidean (f ` s)) t" + have "{x \ s. f x \ t} \ s" using as[unfolded openin_euclidean_subtopology_iff] by auto + moreover + { fix x assume as':"x\{x \ s. f x \ t}" + then obtain e where e: "e>0" "\x'\f ` s. dist x' (f x) < e \ x' \ t" using as[unfolded openin_euclidean_subtopology_iff, THEN conjunct2, THEN bspec[where x="f x"]] by auto + from this(1) obtain d where d: "d>0" "\xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" using `?lhs`[unfolded continuous_on Lim_within, THEN bspec[where x=x]] using as' by auto + have "\e>0. \x'\s. dist x' x < e \ x' \ {x \ s. f x \ t}" using d e unfolding dist_nz[THEN sym] by (rule_tac x=d in exI, auto) } + ultimately have "openin (subtopology euclidean s) {x \ s. f x \ t}" unfolding openin_euclidean_subtopology_iff by auto } + thus ?rhs unfolding continuous_on Lim_within using openin by auto +next + assume ?rhs + { fix e::real and x assume "x\s" "e>0" + { fix xa x' assume "dist (f xa) (f x) < e" "xa \ s" "x' \ s" "dist (f xa) (f x') < e - dist (f xa) (f x)" + hence "dist (f x') (f x) < e" using dist_triangle[of "f x'" "f x" "f xa"] + by (auto simp add: dist_commute) } + hence "ball (f x) e \ f ` s \ f ` s \ (\xa\ball (f x) e \ f ` s. \ea>0. \x'\f ` s. dist x' xa < ea \ x' \ ball (f x) e \ f ` s)" apply auto + apply(rule_tac x="e - dist (f xa) (f x)" in exI) using `e>0` by (auto simp add: dist_commute) + hence "\xa\{xa \ s. f xa \ ball (f x) e \ f ` s}. \ea>0. \x'\s. dist x' xa < ea \ x' \ {xa \ s. f xa \ ball (f x) e \ f ` s}" + using `?rhs`[unfolded openin_euclidean_subtopology_iff, THEN spec[where x="ball (f x) e \ f ` s"]] by auto + hence "\d>0. \xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" apply(erule_tac x=x in ballE) apply auto using `e>0` `x\s` by (auto simp add: dist_commute) } + thus ?lhs unfolding continuous_on Lim_within by auto +qed + +(* ------------------------------------------------------------------------- *) +(* Similarly in terms of closed sets. *) +(* ------------------------------------------------------------------------- *) + +lemma continuous_on_closed: + "continuous_on s f \ (\t. closedin (subtopology euclidean (f ` s)) t --> closedin (subtopology euclidean s) {x \ s. f x \ t})" (is "?lhs = ?rhs") +proof + assume ?lhs + { fix t + have *:"s - {x \ s. f x \ f ` s - t} = {x \ s. f x \ t}" by auto + have **:"f ` s - (f ` s - (f ` s - t)) = f ` s - t" by auto + assume as:"closedin (subtopology euclidean (f ` s)) t" + hence "closedin (subtopology euclidean (f ` s)) (f ` s - (f ` s - t))" unfolding closedin_def topspace_euclidean_subtopology unfolding ** by auto + hence "closedin (subtopology euclidean s) {x \ s. f x \ t}" using `?lhs`[unfolded continuous_on_open, THEN spec[where x="(f ` s) - t"]] + unfolding openin_closedin_eq topspace_euclidean_subtopology unfolding * by auto } + thus ?rhs by auto +next + assume ?rhs + { fix t + have *:"s - {x \ s. f x \ f ` s - t} = {x \ s. f x \ t}" by auto + assume as:"openin (subtopology euclidean (f ` s)) t" + hence "openin (subtopology euclidean s) {x \ s. f x \ t}" using `?rhs`[THEN spec[where x="(f ` s) - t"]] + unfolding openin_closedin_eq topspace_euclidean_subtopology *[THEN sym] closedin_subtopology by auto } + thus ?lhs unfolding continuous_on_open by auto +qed + +text{* Half-global and completely global cases. *} + +lemma continuous_open_in_preimage: + assumes "continuous_on s f" "open t" + shows "openin (subtopology euclidean s) {x \ s. f x \ t}" +proof- + have *:"\x. x \ s \ f x \ t \ x \ s \ f x \ (t \ f ` s)" by auto + have "openin (subtopology euclidean (f ` s)) (t \ f ` s)" + using openin_open_Int[of t "f ` s", OF assms(2)] unfolding openin_open by auto + thus ?thesis using assms(1)[unfolded continuous_on_open, THEN spec[where x="t \ f ` s"]] using * by auto +qed + +lemma continuous_closed_in_preimage: + assumes "continuous_on s f" "closed t" + shows "closedin (subtopology euclidean s) {x \ s. f x \ t}" +proof- + have *:"\x. x \ s \ f x \ t \ x \ s \ f x \ (t \ f ` s)" by auto + have "closedin (subtopology euclidean (f ` s)) (t \ f ` s)" + using closedin_closed_Int[of t "f ` s", OF assms(2)] unfolding Int_commute by auto + thus ?thesis + using assms(1)[unfolded continuous_on_closed, THEN spec[where x="t \ f ` s"]] using * by auto +qed + +lemma continuous_open_preimage: + assumes "continuous_on s f" "open s" "open t" + shows "open {x \ s. f x \ t}" +proof- + obtain T where T: "open T" "{x \ s. f x \ t} = s \ T" + using continuous_open_in_preimage[OF assms(1,3)] unfolding openin_open by auto + thus ?thesis using open_Int[of s T, OF assms(2)] by auto +qed + +lemma continuous_closed_preimage: + assumes "continuous_on s f" "closed s" "closed t" + shows "closed {x \ s. f x \ t}" +proof- + obtain T where T: "closed T" "{x \ s. f x \ t} = s \ T" + using continuous_closed_in_preimage[OF assms(1,3)] unfolding closedin_closed by auto + thus ?thesis using closed_Int[of s T, OF assms(2)] by auto +qed + +lemma continuous_open_preimage_univ: + fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) + shows "\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 :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) + shows "(\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 + +lemma continuous_open_vimage: + fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) + shows "\x. continuous (at x) f \ open s \ open (f -` s)" + unfolding vimage_def by (rule continuous_open_preimage_univ) + +lemma continuous_closed_vimage: + fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) + shows "\x. continuous (at x) f \ closed s \ closed (f -` s)" + unfolding vimage_def by (rule continuous_closed_preimage_univ) + +text{* Equality of continuous functions on closure and related results. *} + +lemma continuous_closed_in_preimage_constant: + "continuous_on s f ==> closedin (subtopology euclidean s) {x \ s. f x = a}" + using continuous_closed_in_preimage[of s f "{a}"] closed_sing by auto + +lemma continuous_closed_preimage_constant: + "continuous_on s f \ closed s ==> closed {x \ s. f x = a}" + using continuous_closed_preimage[of s f "{a}"] closed_sing by auto + +lemma continuous_constant_on_closure: + assumes "continuous_on (closure s) f" + "\x \ s. f x = a" + shows "\x \ (closure s). f x = a" + using continuous_closed_preimage_constant[of "closure s" f a] + assms closure_minimal[of s "{x \ closure s. f x = a}"] closure_subset unfolding subset_eq by auto + +lemma image_closure_subset: + assumes "continuous_on (closure s) f" "closed t" "(f ` s) \ t" + shows "f ` (closure s) \ t" +proof- + have "s \ {x \ closure s. f x \ t}" using assms(3) closure_subset by auto + moreover have "closed {x \ closure s. f x \ t}" + using continuous_closed_preimage[OF assms(1)] and assms(2) by auto + ultimately have "closure s = {x \ closure s . f x \ t}" + using closure_minimal[of s "{x \ closure s. f x \ t}"] by auto + thus ?thesis by auto +qed + +lemma continuous_on_closure_norm_le: + fixes f :: "'a::metric_space \ 'b::real_normed_vector" + assumes "continuous_on (closure s) f" "\y \ s. norm(f y) \ b" "x \ (closure s)" + shows "norm(f x) \ b" +proof- + have *:"f ` s \ cball 0 b" using assms(2)[unfolded mem_cball_0[THEN sym]] by auto + show ?thesis + using image_closure_subset[OF assms(1) closed_cball[of 0 b] *] assms(3) + unfolding subset_eq apply(erule_tac x="f x" in ballE) by (auto simp add: dist_norm) +qed + +text{* Making a continuous function avoid some value in a neighbourhood. *} + +lemma continuous_within_avoid: + fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) + assumes "continuous (at x within s) f" "x \ s" "f x \ a" + shows "\e>0. \y \ s. dist x y < e --> f y \ a" +proof- + obtain d where "d>0" and d:"\xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < dist (f x) a" + using assms(1)[unfolded continuous_within Lim_within, THEN spec[where x="dist (f x) a"]] assms(3)[unfolded dist_nz] by auto + { fix y assume " y\s" "dist x y < d" + hence "f y \ a" using d[THEN bspec[where x=y]] assms(3)[unfolded dist_nz] + apply auto unfolding dist_nz[THEN sym] by (auto simp add: dist_commute) } + thus ?thesis using `d>0` by auto +qed + +lemma continuous_at_avoid: + fixes f :: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize *) + assumes "continuous (at x) f" "f x \ a" + shows "\e>0. \y. dist x y < e \ f y \ a" +using assms using continuous_within_avoid[of x UNIV f a, unfolded within_UNIV] by auto + +lemma continuous_on_avoid: + assumes "continuous_on s f" "x \ s" "f x \ a" + shows "\e>0. \y \ s. dist x y < e \ f y \ a" +using assms(1)[unfolded continuous_on_eq_continuous_within, THEN bspec[where x=x], OF assms(2)] continuous_within_avoid[of x s f a] assms(2,3) by auto + +lemma continuous_on_open_avoid: + assumes "continuous_on s f" "open s" "x \ s" "f x \ a" + shows "\e>0. \y. dist x y < e \ f y \ a" +using assms(1)[unfolded continuous_on_eq_continuous_at[OF assms(2)], THEN bspec[where x=x], OF assms(3)] continuous_at_avoid[of x f a] assms(3,4) by auto + +text{* Proving a function is constant by proving open-ness of level set. *} + +lemma continuous_levelset_open_in_cases: + "connected s \ continuous_on s f \ + openin (subtopology euclidean s) {x \ s. f x = a} + ==> (\x \ s. f x \ a) \ (\x \ s. f x = a)" +unfolding connected_clopen using continuous_closed_in_preimage_constant by auto + +lemma continuous_levelset_open_in: + "connected s \ continuous_on s f \ + openin (subtopology euclidean s) {x \ s. f x = a} \ + (\x \ s. f x = a) ==> (\x \ s. f x = a)" +using continuous_levelset_open_in_cases[of s f ] +by meson + +lemma continuous_levelset_open: + assumes "connected s" "continuous_on s f" "open {x \ s. f x = a}" "\x \ s. f x = a" + shows "\x \ s. f x = a" +using continuous_levelset_open_in[OF assms(1,2), of a, unfolded openin_open] using assms (3,4) by auto + +text{* Some arithmetical combinations (more to prove). *} + +lemma open_scaling[intro]: + fixes s :: "'a::real_normed_vector set" + assumes "c \ 0" "open 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 *\<^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 *\<^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 + +lemma minus_image_eq_vimage: + fixes A :: "'a::ab_group_add set" + shows "(\x. - x) ` A = (\x. - x) -` A" + by (auto intro!: image_eqI [where f="\x. - x"]) + +lemma open_negations: + fixes s :: "'a::real_normed_vector set" + shows "open s ==> open ((\ x. -x) ` s)" + unfolding scaleR_minus1_left [symmetric] + by (rule open_scaling, auto) + +lemma open_translation: + 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 } + moreover have "{x. x - a \ s} = op + a ` s" apply auto unfolding image_iff apply(rule_tac x="x - a" in bexI) by auto + ultimately show ?thesis using continuous_open_preimage_univ[of "\x. x - a" s] using assms by auto +qed + +lemma open_affinity: + fixes s :: "'a::real_normed_vector set" + assumes "open s" "c \ 0" + shows "open ((\x. a + c *\<^sub>R x) ` s)" +proof- + 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: + fixes s :: "'a::real_normed_vector set" + shows "interior ((\x. a + x) ` s) = (\x. a + x) ` (interior s)" +proof (rule set_ext, rule) + fix x assume "x \ interior (op + a ` s)" + then obtain e where "e>0" and e:"ball x e \ op + a ` s" unfolding mem_interior by auto + hence "ball (x - a) e \ s" unfolding subset_eq Ball_def mem_ball dist_norm apply auto apply(erule_tac x="a + xa" in allE) unfolding ab_group_add_class.diff_diff_eq[THEN sym] by auto + thus "x \ op + a ` interior s" unfolding image_iff apply(rule_tac x="x - a" in bexI) unfolding mem_interior using `e > 0` by auto +next + fix x assume "x \ op + a ` interior s" + then obtain y e where "e>0" and e:"ball y e \ s" and y:"x = a + y" unfolding image_iff Bex_def mem_interior by auto + { fix z have *:"a + y - z = y + a - z" by auto + assume "z\ball x e" + hence "z - a \ s" using e[unfolded subset_eq, THEN bspec[where x="z - a"]] unfolding mem_ball dist_norm y ab_group_add_class.diff_diff_eq2 * by auto + hence "z \ op + a ` s" unfolding image_iff by(auto intro!: bexI[where x="z - a"]) } + hence "ball x e \ op + a ` s" unfolding subset_eq by auto + thus "x \ interior (op + a ` s)" unfolding mem_interior using `e>0` by auto +qed + +subsection {* Preservation of compactness and connectedness under continuous function. *} + +lemma compact_continuous_image: + assumes "continuous_on s f" "compact s" + shows "compact(f ` s)" +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:"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. subseq r \ ((x \ r) ---> l) sequentially" unfolding Lim_sequentially using r lr `l\s` by auto } + thus ?thesis unfolding compact_def by auto +qed + +lemma connected_continuous_image: + assumes "continuous_on s f" "connected s" + shows "connected(f ` s)" +proof- + { fix T assume as: "T \ {}" "T \ f ` s" "openin (subtopology euclidean (f ` s)) T" "closedin (subtopology euclidean (f ` s)) T" + have "{x \ s. f x \ T} = {} \ {x \ s. f x \ T} = s" + using assms(1)[unfolded continuous_on_open, THEN spec[where x=T]] + using assms(1)[unfolded continuous_on_closed, THEN spec[where x=T]] + using assms(2)[unfolded connected_clopen, THEN spec[where x="{x \ s. f x \ T}"]] as(3,4) by auto + hence False using as(1,2) + using as(4)[unfolded closedin_def topspace_euclidean_subtopology] by auto } + thus ?thesis unfolding connected_clopen by auto +qed + +text{* Continuity implies uniform continuity on a compact domain. *} + +lemma compact_uniformly_continuous: + assumes "continuous_on s f" "compact s" + shows "uniformly_continuous_on s f" +proof- + { fix x assume x:"x\s" + hence "\xa. \y. 0 < xa \ (y > 0 \ (\x'\s. dist x' x < y \ dist (f x') (f x) < xa))" using assms(1)[unfolded continuous_on_def, THEN bspec[where x=x]] by auto + hence "\fa. \xa>0. \x'\s. fa xa > 0 \ (dist x' x < fa xa \ dist (f x') (f x) < xa)" using choice[of "\e d. e>0 \ d>0 \(\x'\s. (dist x' x < d \ dist (f x') (f x) < e))"] by auto } + then have "\x\s. \y. \xa. 0 < xa \ (\x'\s. y xa > 0 \ (dist x' x < y xa \ dist (f x') (f x) < xa))" by auto + then obtain d where d:"\e>0. \x\s. \x'\s. d x e > 0 \ (dist x' x < d x e \ dist (f x') (f x) < e)" + using bchoice[of s "\x fa. \xa>0. \x'\s. fa xa > 0 \ (dist x' x < fa xa \ dist (f x') (f x) < xa)"] by blast + + { fix e::real assume "e>0" + + { fix x assume "x\s" hence "x \ ball x (d x (e / 2))" unfolding centre_in_ball using d[THEN spec[where x="e/2"]] using `e>0` by auto } + hence "s \ \{ball x (d x (e / 2)) |x. x \ s}" unfolding subset_eq by auto + moreover + { fix b assume "b\{ball x (d x (e / 2)) |x. x \ s}" hence "open b" by auto } + ultimately obtain ea where "ea>0" and ea:"\x\s. \b\{ball x (d x (e / 2)) |x. x \ s}. ball x ea \ b" using heine_borel_lemma[OF assms(2), of "{ball x (d x (e / 2)) | x. x\s }"] by auto + + { fix x y assume "x\s" "y\s" and as:"dist y x < ea" + obtain z where "z\s" and z:"ball x ea \ ball z (d z (e / 2))" using ea[THEN bspec[where x=x]] and `x\s` by auto + hence "x\ball z (d z (e / 2))" using `ea>0` unfolding subset_eq by auto + hence "dist (f z) (f x) < e / 2" using d[THEN spec[where x="e/2"]] and `e>0` and `x\s` and `z\s` + by (auto simp add: dist_commute) + moreover have "y\ball z (d z (e / 2))" using as and `ea>0` and z[unfolded subset_eq] + by (auto simp add: dist_commute) + hence "dist (f z) (f y) < e / 2" using d[THEN spec[where x="e/2"]] and `e>0` and `y\s` and `z\s` + by (auto simp add: dist_commute) + ultimately have "dist (f y) (f x) < e" using dist_triangle_half_r[of "f z" "f x" e "f y"] + by (auto simp add: dist_commute) } + then have "\d>0. \x\s. \x'\s. dist x' x < d \ dist (f x') (f x) < e" using `ea>0` by auto } + thus ?thesis unfolding uniformly_continuous_on_def by auto +qed + +text{* Continuity of inverse function on compact domain. *} + +lemma continuous_on_inverse: + 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- + have *:"g ` f ` s = s" using assms(3) by (auto simp add: image_iff) + { fix t assume t:"closedin (subtopology euclidean (g ` f ` s)) t" + then obtain T where T: "closed T" "t = s \ T" unfolding closedin_closed unfolding * by auto + have "continuous_on (s \ T) f" using continuous_on_subset[OF assms(1), of "s \ t"] + unfolding T(2) and Int_left_absorb by auto + moreover have "compact (s \ T)" + using assms(2) unfolding compact_eq_bounded_closed + using bounded_subset[of s "s \ T"] and T(1) by auto + ultimately have "closed (f ` t)" using T(1) unfolding T(2) + using compact_continuous_image [of "s \ T" f] unfolding compact_eq_bounded_closed by auto + moreover have "{x \ f ` s. g x \ t} = f ` s \ f ` t" using assms(3) unfolding T(2) by auto + ultimately have "closedin (subtopology euclidean (f ` s)) {x \ f ` s. g x \ t}" + unfolding closedin_closed by auto } + thus ?thesis unfolding continuous_on_closed by auto +qed + +subsection{* A uniformly convergent limit of continuous functions is continuous. *} + +lemma norm_triangle_lt: + fixes x y :: "'a::real_normed_vector" + shows "norm x + norm y < e \ norm (x + y) < e" +by (rule le_less_trans [OF norm_triangle_ineq]) + +lemma continuous_uniform_limit: + fixes f :: "'a \ 'b::metric_space \ 'c::real_normed_vector" + assumes "\ (trivial_limit net)" "eventually (\n. continuous_on s (f n)) net" + "\e>0. eventually (\n. \x \ s. norm(f n x - g x) < e) net" + shows "continuous_on s g" +proof- + { fix x and e::real assume "x\s" "e>0" + have "eventually (\n. \x\s. norm (f n x - g x) < e / 3) net" using `e>0` assms(3)[THEN spec[where x="e/3"]] by auto + then obtain n where n:"\xa\s. norm (f n xa - g xa) < e / 3" "continuous_on s (f n)" + using eventually_and[of "(\n. \x\s. norm (f n x - g x) < e / 3)" "(\n. continuous_on s (f n))" net] assms(1,2) eventually_happens by blast + have "e / 3 > 0" using `e>0` by auto + then obtain d where "d>0" and d:"\x'\s. dist x' x < d \ dist (f n x') (f n x) < e / 3" + using n(2)[unfolded continuous_on_def, THEN bspec[where x=x], OF `x\s`, THEN spec[where x="e/3"]] by blast + { fix y assume "y\s" "dist y x < d" + hence "dist (f n y) (f n x) < e / 3" using d[THEN bspec[where x=y]] by auto + hence "norm (f n y - g x) < 2 * e / 3" using norm_triangle_lt[of "f n y - f n x" "f n x - g x" "2*e/3"] + using n(1)[THEN bspec[where x=x], OF `x\s`] unfolding dist_norm unfolding ab_group_add_class.ab_diff_minus by auto + hence "dist (g y) (g x) < e" unfolding dist_norm using n(1)[THEN bspec[where x=y], OF `y\s`] + unfolding norm_minus_cancel[of "f n y - g y", THEN sym] using norm_triangle_lt[of "f n y - g x" "g y - f n y" e] by (auto simp add: uminus_add_conv_diff) } + hence "\d>0. \x'\s. dist x' x < d \ dist (g x') (g x) < e" using `d>0` by auto } + thus ?thesis unfolding continuous_on_def by auto +qed + +subsection{* Topological properties of linear functions. *} + +lemma linear_lim_0: + assumes "bounded_linear f" shows "(f ---> 0) (at (0))" +proof- + interpret f: bounded_linear f by fact + have "(f ---> f 0) (at 0)" + using tendsto_ident_at by (rule f.tendsto) + thus ?thesis unfolding f.zero . +qed + +lemma linear_continuous_at: + assumes "bounded_linear f" shows "continuous (at a) f" + unfolding continuous_at using assms + apply (rule bounded_linear.tendsto) + apply (rule tendsto_ident_at) + done + +lemma linear_continuous_within: + shows "bounded_linear f ==> continuous (at x within s) f" + using continuous_at_imp_continuous_within[of x f s] using linear_continuous_at[of f] by auto + +lemma linear_continuous_on: + shows "bounded_linear f ==> continuous_on s f" + using continuous_at_imp_continuous_on[of s f] using linear_continuous_at[of f] by auto + +text{* Also bilinear functions, in composition form. *} + +lemma bilinear_continuous_at_compose: + shows "continuous (at x) f \ continuous (at x) g \ bounded_bilinear h + ==> continuous (at x) (\x. h (f x) (g x))" + unfolding continuous_at using Lim_bilinear[of f "f x" "(at x)" g "g x" h] by auto + +lemma bilinear_continuous_within_compose: + shows "continuous (at x within s) f \ continuous (at x within s) g \ bounded_bilinear h + ==> continuous (at x within s) (\x. h (f x) (g x))" + unfolding continuous_within using Lim_bilinear[of f "f x"] by auto + +lemma bilinear_continuous_on_compose: + shows "continuous_on s f \ continuous_on s g \ bounded_bilinear h + ==> continuous_on s (\x. h (f x) (g x))" + unfolding continuous_on_eq_continuous_within apply auto apply(erule_tac x=x in ballE) apply auto apply(erule_tac x=x in ballE) apply auto + using bilinear_continuous_within_compose[of _ s f g h] by auto + +subsection{* Topological stuff lifted from and dropped to R *} + + +lemma open_real: + fixes s :: "real set" shows + "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 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 + 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_real_range: + fixes f :: "'a::real_normed_vector \ real" + 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)) + +text{* Hence some handy theorems on distance, diameter etc. of/from a set. *} + +lemma compact_attains_sup: + fixes s :: "real set" + assumes "compact s" "s \ {}" + shows "\x \ s. \y \ s. y \ x" +proof- + 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_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 s" "s \ {}" shows "\x \ s. \y \ s. x \ y" +proof- + 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 + moreover + { fix x assume "x \ s" + hence *:"abs (x - rinf s) = x - rinf s" using as(1)[THEN bspec[where x=x]] by auto + 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_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 f + ==> (\x \ s. \y \ s. f y \ f x)" + using compact_attains_sup[of "f ` s"] + 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 f + \ (\x \ s. \y \ s. f x \ f y)" + using compact_attains_inf[of "f ` s"] + using compact_continuous_image[of s f] by auto + +lemma distance_attains_sup: + assumes "compact s" "s \ {}" + shows "\x \ s. \y \ s. dist a y \ dist a x" +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 :: "'a::heine_borel" + assumes "closed s" "s \ {}" + shows "\x \ s. \y \ s. dist a x \ dist a y" +proof- + from assms(2) obtain b where "b\s" by auto + let ?B = "cball a (dist b a) \ s" + have "b \ ?B" using `b\s` by (simp add: dist_commute) + hence "?B \ {}" by auto + moreover + { fix x assume "x\?B" + fix e::real assume "e>0" + { 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 + +subsection{* We can now extend limit compositions to consider the scalar multiplier. *} + +lemma Lim_mul: + fixes f :: "'a \ 'b::real_normed_vector" + assumes "(c ---> d) net" "(f ---> l) net" + shows "((\x. c(x) *\<^sub>R f x) ---> (d *\<^sub>R l)) net" + using assms by (rule scaleR.tendsto) + +lemma Lim_vmul: + 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" and v :: "'b::real_normed_vector" + shows "continuous net c ==> continuous net (\x. c(x) *\<^sub>R v)" + unfolding continuous_def using Lim_vmul[of c] by auto + +lemma continuous_mul: + fixes c :: "'a::metric_space \ real" + fixes f :: "'a::metric_space \ 'b::real_normed_vector" + shows "continuous net c \ continuous net f + ==> continuous net (\x. c(x) *\<^sub>R f x) " + unfolding continuous_def by (intro tendsto_intros) + +lemma continuous_on_vmul: + fixes c :: "'a::metric_space \ real" and v :: "'b::real_normed_vector" + shows "continuous_on s c ==> continuous_on s (\x. c(x) *\<^sub>R v)" + unfolding continuous_on_eq_continuous_within using continuous_vmul[of _ c] by auto + +lemma continuous_on_mul: + fixes c :: "'a::metric_space \ real" + fixes f :: "'a::metric_space \ 'b::real_normed_vector" + shows "continuous_on s c \ continuous_on s f + ==> continuous_on s (\x. c(x) *\<^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 "(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 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 :: "'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 :: "'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. *} + +lemma bounded_pastecart: + fixes s :: "('a::real_normed_vector ^ _) set" (* FIXME: generalize to metric_space *) + assumes "bounded s" "bounded t" + shows "bounded { pastecart x y | x y . (x \ s \ y \ t)}" +proof- + obtain a b where ab:"\x\s. norm x \ a" "\x\t. norm x \ b" using assms[unfolded bounded_iff] by auto + { fix x y assume "x\s" "y\t" + hence "norm x \ a" "norm y \ b" using ab by auto + hence "norm (pastecart x y) \ a + b" using norm_pastecart[of x y] by auto } + 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" + shows "closed {pastecart x y | x y . x \ s \ y \ t}" +proof- + { fix x l assume as:"\n::nat. x n \ {pastecart x y |x y. x \ s \ y \ t}" "(x ---> l) sequentially" + { fix n::nat have "fstcart (x n) \ s" "sndcart (x n) \ t" using as(1)[THEN spec[where x=n]] by auto } note * = this + moreover + { fix e::real assume "e>0" + then obtain N::nat where N:"\n\N. dist (x n) l < e" using as(2)[unfolded Lim_sequentially, THEN spec[where x=e]] by auto + { fix n::nat assume "n\N" + hence "dist (fstcart (x n)) (fstcart l) < e" "dist (sndcart (x n)) (sndcart l) < e" + using N[THEN spec[where x=n]] dist_fstcart[of "x n" l] dist_sndcart[of "x n" l] by auto } + hence "\N. \n\N. dist (fstcart (x n)) (fstcart l) < e" "\N. \n\N. dist (sndcart (x n)) (sndcart l) < e" by auto } + ultimately have "fstcart l \ s" "sndcart l \ t" + using assms(1)[unfolded closed_sequential_limits, THEN spec[where x="\n. fstcart (x n)"], THEN spec[where x="fstcart l"]] + using assms(2)[unfolded closed_sequential_limits, THEN spec[where x="\n. sndcart (x n)"], THEN spec[where x="sndcart l"]] + unfolding Lim_sequentially by auto + hence "l \ {pastecart x y |x y. x \ s \ y \ t}" using pastecart_fst_snd[THEN sym, of l] by auto } + thus ?thesis unfolding closed_sequential_limits by auto +qed + +lemma compact_pastecart: + fixes s t :: "(real ^ _) set" + 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_scaling: + fixes s :: "'a::real_normed_vector set" + assumes "compact s" shows "compact ((\x. c *\<^sub>R x) ` s)" +proof- + let ?f = "\x. scaleR c x" + have *:"bounded_linear ?f" by (rule scaleR.bounded_linear_right) + show ?thesis using compact_continuous_image[of s ?f] continuous_at_imp_continuous_on[of s ?f] + using linear_continuous_at[OF *] assms by auto +qed + +lemma compact_negations: + fixes s :: "'a::real_normed_vector set" + assumes "compact s" shows "compact ((\x. -x) ` s)" + using compact_scaling [OF assms, of "- 1"] by auto + +lemma compact_sums: + 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. 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 :: "'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} = {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 :: "'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 + thus ?thesis using compact_sums[OF assms compact_sing[of a]] by auto +qed + +lemma compact_affinity: + fixes s :: "'a::real_normed_vector set" + assumes "compact s" shows "compact ((\x. a + c *\<^sub>R x) ` s)" +proof- + 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 :: "'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[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 + +text{* We can state this in terms of diameter of a set. *} + +definition "diameter s = (if s = {} then 0::real else rsup {norm(x - y) | x y. x \ s \ y \ s})" + (* TODO: generalize to class metric_space *) + +lemma diameter_bounded: + assumes "bounded s" + shows "\x\s. \y\s. norm(x - y) \ diameter s" + "\d>0. d < diameter s --> (\x\s. \y\s. norm(x - y) > d)" +proof- + let ?D = "{norm (x - y) |x y. x \ s \ y \ s}" + obtain a where a:"\x\s. norm x \ a" using assms[unfolded bounded_iff] by auto + { fix x y assume "x \ s" "y \ s" + hence "norm (x - y) \ 2 * a" using norm_triangle_ineq[of x "-y", unfolded norm_minus_cancel] a[THEN bspec[where x=x]] a[THEN bspec[where x=y]] by (auto simp add: ring_simps) } + note * = this + { fix x y assume "x\s" "y\s" hence "s \ {}" by auto + have lub:"isLub UNIV ?D (rsup ?D)" using * rsup[of ?D] using `s\{}` unfolding setle_def by auto + have "norm(x - y) \ diameter s" unfolding diameter_def using `s\{}` *[OF `x\s` `y\s`] `x\s` `y\s` isLubD1[OF lub] unfolding setle_def by auto } + moreover + { fix d::real assume "d>0" "d < diameter s" + hence "s\{}" unfolding diameter_def by auto + hence lub:"isLub UNIV ?D (rsup ?D)" using * rsup[of ?D] unfolding setle_def by auto + have "\d' \ ?D. d' > d" + proof(rule ccontr) + assume "\ (\d'\{norm (x - y) |x y. x \ s \ y \ s}. d < d')" + hence as:"\d'\?D. d' \ d" apply auto apply(erule_tac x="norm (x - y)" in allE) by auto + hence "isUb UNIV ?D d" unfolding isUb_def unfolding setle_def by auto + thus False using `d < diameter s` `s\{}` isLub_le_isUb[OF lub, of d] unfolding diameter_def by auto + qed + hence "\x\s. \y\s. norm(x - y) > d" by auto } + ultimately show "\x\s. \y\s. norm(x - y) \ diameter s" + "\d>0. d < diameter s --> (\x\s. \y\s. norm(x - y) > d)" by auto +qed + +lemma diameter_bounded_bound: + "bounded s \ x \ s \ y \ s ==> norm(x - y) \ diameter s" + using diameter_bounded by blast + +lemma diameter_compact_attained: + 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) 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 + thus ?thesis using diameter_bounded(1)[OF b, THEN bspec[where x=x], THEN bspec[where x=y], OF xys] and xys by auto +qed + +text{* Related results with closure as the conclusion. *} + +lemma closed_scaling: + fixes s :: "'a::real_normed_vector set" + assumes "closed s" shows "closed ((\x. c *\<^sub>R x) ` s)" +proof(cases "s={}") + case True thus ?thesis by auto +next + case False + show ?thesis + proof(cases "c=0") + have *:"(\x. 0) ` s = {0}" using `s\{}` by auto + case True thus ?thesis apply auto unfolding * using closed_sing by auto + next + case False + { fix x l assume as:"\n::nat. x n \ scaleR c ` s" "(x ---> l) sequentially" + { fix n::nat have "scaleR (1 / c) (x n) \ s" + using as(1)[THEN spec[where x=n]] + using `c\0` by (auto simp add: vector_smult_assoc) + } + moreover + { fix e::real assume "e>0" + hence "0 < e *\c\" using `c\0` mult_pos_pos[of e "abs c"] by auto + 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] + 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" + using assms[unfolded closed_sequential_limits, THEN spec[where x="\n. scaleR (1/c) (x n)"], THEN spec[where x="scaleR (1/c) l"]] + unfolding image_iff using `c\0` apply(rule_tac x="scaleR (1 / c) l" in bexI) by auto } + thus ?thesis unfolding closed_sequential_limits by fast + qed +qed + +lemma closed_negations: + fixes s :: "'a::real_normed_vector set" + assumes "closed s" shows "closed ((\x. -x) ` s)" + using closed_scaling[OF assms, of "- 1"] by simp + +lemma compact_closed_sums: + fixes s :: "'a::real_normed_vector set" + assumes "compact s" "closed t" shows "closed {x + y | x y. x \ s \ y \ t}" +proof- + let ?S = "{x + y |x y. x \ s \ y \ t}" + { 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:"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_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 + hence "l \ ?S" using `l' \ s` apply auto apply(rule_tac x=l' in exI) apply(rule_tac x="l - l'" in exI) by auto + } + thus ?thesis unfolding closed_sequential_limits by fast +qed + +lemma closed_compact_sums: + fixes s t :: "'a::real_normed_vector set" + assumes "closed s" "compact t" + shows "closed {x + y | x y. x \ s \ y \ t}" +proof- + have "{x + y |x y. x \ t \ y \ s} = {x + y |x y. x \ s \ y \ t}" apply auto + apply(rule_tac x=y in exI) apply auto apply(rule_tac x=y in exI) by auto + thus ?thesis using compact_closed_sums[OF assms(2,1)] by simp +qed + +lemma compact_closed_differences: + fixes s t :: "'a::real_normed_vector set" + assumes "compact s" "closed t" + shows "closed {x - y | x y. x \ s \ y \ t}" +proof- + have "{x + y |x y. x \ s \ y \ uminus ` t} = {x - y |x y. x \ s \ y \ 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_closed_sums[OF assms(1) closed_negations[OF assms(2)]] by auto +qed + +lemma closed_compact_differences: + fixes s t :: "'a::real_normed_vector set" + assumes "closed s" "compact t" + shows "closed {x - y | x y. x \ s \ y \ t}" +proof- + have "{x + y |x y. x \ s \ y \ uminus ` t} = {x - y |x y. x \ s \ y \ t}" + apply auto apply(rule_tac x=xa in exI) apply auto apply(rule_tac x=xa in exI) by auto + thus ?thesis using closed_compact_sums[OF assms(1) compact_negations[OF assms(2)]] by simp +qed + +lemma closed_translation: + fixes a :: "'a::real_normed_vector" + assumes "closed s" shows "closed ((\x. a + x) ` s)" +proof- + have "{a + y |y. y \ s} = (op + a ` s)" by auto + thus ?thesis using compact_closed_sums[OF compact_sing[of a] assms] by auto +qed + +lemma translation_UNIV: + fixes a :: "'a::ab_group_add" shows "range (\x. a + x) = UNIV" + apply (auto simp add: image_iff) apply(rule_tac x="x - a" in exI) by auto + +lemma translation_diff: + fixes a :: "'a::ab_group_add" + shows "(\x. a + x) ` (s - t) = ((\x. a + x) ` s) - ((\x. a + x) ` t)" + by auto + +lemma closure_translation: + fixes a :: "'a::real_normed_vector" + shows "closure ((\x. a + x) ` s) = (\x. a + x) ` (closure s)" +proof- + have *:"op + a ` (UNIV - s) = UNIV - op + a ` s" + apply auto unfolding image_iff apply(rule_tac x="x - a" in bexI) by auto + show ?thesis unfolding closure_interior translation_diff translation_UNIV + using interior_translation[of a "UNIV - s"] unfolding * by auto +qed + +lemma frontier_translation: + fixes a :: "'a::real_normed_vector" + shows "frontier((\x. a + x) ` s) = (\x. a + x) ` (frontier s)" + unfolding frontier_def translation_diff interior_translation closure_translation by auto + +subsection{* Separation between points and sets. *} + +lemma separate_point_closed: + fixes s :: "'a::heine_borel set" + shows "closed s \ a \ s ==> (\d>0. \x\s. d \ dist a x)" +proof(cases "s = {}") + case True + thus ?thesis by(auto intro!: exI[where x=1]) +next + case False + assume "closed s" "a \ s" + then obtain x where "x\s" "\y\s. dist a x \ dist a y" using `s \ {}` distance_attains_inf [of s a] by blast + with `x\s` show ?thesis using dist_pos_lt[of a x] and`a \ s` by blast +qed + +lemma separate_compact_closed: + 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- + have "0 \ {x - y |x y. x \ s \ y \ t}" using assms(3) by auto + then obtain d where "d>0" and d:"\x\{x - y |x y. x \ s \ y \ t}. d \ dist 0 x" + using separate_point_closed[OF compact_closed_differences[OF assms(1,2)], of 0] by auto + { fix x y assume "x\s" "y\t" + hence "x - y \ {x - y |x y. x \ s \ y \ t}" by auto + hence "d \ dist (x - y) 0" using d[THEN bspec[where x="x - y"]] using dist_commute + by (auto simp add: dist_commute) + hence "d \ dist x y" unfolding dist_norm by auto } + thus ?thesis using `d>0` by auto +qed + +lemma separate_closed_compact: + 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- + have *:"t \ s = {}" using assms(3) by auto + show ?thesis using separate_compact_closed[OF assms(2,1) *] + apply auto apply(rule_tac x=d in exI) apply auto apply (erule_tac x=y in ballE) + by (auto simp add: dist_commute) +qed + +(* A cute way of denoting open and closed intervals using overloading. *) + +lemma interval: fixes a :: "'a::ord^'n::finite" shows + "{a <..< b} = {x::'a^'n. \i. a$i < x$i \ x$i < b$i}" and + "{a .. b} = {x::'a^'n. \i. a$i \ x$i \ x$i \ b$i}" + by (auto simp add: expand_set_eq vector_less_def vector_less_eq_def) + +lemma mem_interval: fixes a :: "'a::ord^'n::finite" shows + "x \ {a<.. (\i. a$i < x$i \ x$i < b$i)" + "x \ {a .. b} \ (\i. a$i \ x$i \ x$i \ b$i)" + using interval[of a b] by(auto simp add: expand_set_eq vector_less_def vector_less_eq_def) + +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 forall_1) + +lemma interval_eq_empty: fixes a :: "real^'n::finite" shows + "({a <..< b} = {} \ (\i. b$i \ a$i))" (is ?th1) and + "({a .. b} = {} \ (\i. b$i < a$i))" (is ?th2) +proof- + { fix i x assume as:"b$i \ a$i" and x:"x\{a <..< b}" + hence "a $ i < x $ i \ x $ i < b $ i" unfolding mem_interval by auto + hence "a$i < b$i" by auto + hence False using as by auto } + moreover + { assume as:"\i. \ (b$i \ a$i)" + 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) *\<^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 } + ultimately show ?th1 by blast + + { fix i x assume as:"b$i < a$i" and x:"x\{a .. b}" + hence "a $ i \ x $ i \ x $ i \ b $ i" unfolding mem_interval by auto + hence "a$i \ b$i" by auto + hence False using as by auto } + moreover + { assume as:"\i. \ (b$i < a$i)" + 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) *\<^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 } + ultimately show ?th2 by blast +qed + +lemma interval_ne_empty: fixes a :: "real^'n::finite" shows + "{a .. b} \ {} \ (\i. a$i \ b$i)" and + "{a <..< b} \ {} \ (\i. a$i < b$i)" + unfolding interval_eq_empty[of a b] by (auto simp add: not_less not_le) (* BH: Why doesn't just "auto" work here? *) + +lemma subset_interval_imp: fixes a :: "real^'n::finite" shows + "(\i. a$i \ c$i \ d$i \ b$i) \ {c .. d} \ {a .. b}" and + "(\i. a$i < c$i \ d$i < b$i) \ {c .. d} \ {a<..i. a$i \ c$i \ d$i \ b$i) \ {c<.. {a .. b}" and + "(\i. a$i \ c$i \ d$i \ b$i) \ {c<.. {a<.. {a<.. {a .. b}" +proof(simp add: subset_eq, rule) + fix x + assume x:"x \{a<.. x $ i" + using x order_less_imp_le[of "a$i" "x$i"] + by(simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq) + } + moreover + { fix i + have "x $ i \ b $ i" + using x order_less_imp_le[of "x$i" "b$i"] + by(simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq) + } + ultimately + show "a \ x \ x \ b" + by(simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq) +qed + +lemma subset_interval: fixes a :: "real^'n::finite" shows + "{c .. d} \ {a .. b} \ (\i. c$i \ d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th1) and + "{c .. d} \ {a<.. (\i. c$i \ d$i) --> (\i. a$i < c$i \ d$i < b$i)" (is ?th2) and + "{c<.. {a .. b} \ (\i. c$i < d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th3) and + "{c<.. {a<.. (\i. c$i < d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th4) +proof- + show ?th1 unfolding subset_eq and Ball_def and mem_interval by (auto intro: order_trans) + show ?th2 unfolding subset_eq and Ball_def and mem_interval by (auto intro: le_less_trans less_le_trans order_trans less_imp_le) + { assume as: "{c<.. {a .. b}" "\i. c$i < d$i" + hence "{c<.. {}" unfolding interval_eq_empty by (auto, drule_tac x=i in spec, simp) (* BH: Why doesn't just "auto" work? *) + fix i + (** TODO combine the following two parts as done in the HOL_light version. **) + { let ?x = "(\ j. (if j=i then ((min (a$j) (d$j))+c$j)/2 else (c$j+d$j)/2))::real^'n" + assume as2: "a$i > c$i" + { fix j + have "c $ j < ?x $ j \ ?x $ j < d $ j" unfolding Cart_lambda_beta + apply(cases "j=i") using as(2)[THEN spec[where x=j]] + by (auto simp add: less_divide_eq_number_of1 as2) } + hence "?x\{c<..{a .. b}" + unfolding mem_interval apply auto apply(rule_tac x=i in exI) + using as(2)[THEN spec[where x=i]] and as2 + by (auto simp add: less_divide_eq_number_of1) + ultimately have False using as by auto } + hence "a$i \ c$i" by(rule ccontr)auto + moreover + { let ?x = "(\ j. (if j=i then ((max (b$j) (c$j))+d$j)/2 else (c$j+d$j)/2))::real^'n" + assume as2: "b$i < d$i" + { fix j + have "d $ j > ?x $ j \ ?x $ j > c $ j" unfolding Cart_lambda_beta + apply(cases "j=i") using as(2)[THEN spec[where x=j]] + by (auto simp add: less_divide_eq_number_of1 as2) } + hence "?x\{c<..{a .. b}" + unfolding mem_interval apply auto apply(rule_tac x=i in exI) + using as(2)[THEN spec[where x=i]] and as2 + by (auto simp add: less_divide_eq_number_of1) + ultimately have False using as by auto } + hence "b$i \ d$i" by(rule ccontr)auto + ultimately + have "a$i \ c$i \ d$i \ b$i" by auto + } note part1 = this + thus ?th3 unfolding subset_eq and Ball_def and mem_interval apply auto apply (erule_tac x=ia in allE, simp)+ by (erule_tac x=i in allE, erule_tac x=i in allE, simp)+ + { assume as:"{c<.. {a<..i. c$i < d$i" + fix i + from as(1) have "{c<.. {a..b}" using interval_open_subset_closed[of a b] by auto + hence "a$i \ c$i \ d$i \ b$i" using part1 and as(2) by auto } note * = this + thus ?th4 unfolding subset_eq and Ball_def and mem_interval apply auto apply (erule_tac x=ia in allE, simp)+ by (erule_tac x=i in allE, erule_tac x=i in allE, simp)+ +qed + +lemma disjoint_interval: fixes a::"real^'n::finite" shows + "{a .. b} \ {c .. d} = {} \ (\i. (b$i < a$i \ d$i < c$i \ b$i < c$i \ d$i < a$i))" (is ?th1) and + "{a .. b} \ {c<.. (\i. (b$i < a$i \ d$i \ c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th2) and + "{a<.. {c .. d} = {} \ (\i. (b$i \ a$i \ d$i < c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th3) and + "{a<.. {c<.. (\i. (b$i \ a$i \ d$i \ c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th4) +proof- + let ?z = "(\ i. ((max (a$i) (c$i)) + (min (b$i) (d$i))) / 2)::real^'n" + show ?th1 ?th2 ?th3 ?th4 + unfolding expand_set_eq and Int_iff and empty_iff and mem_interval and all_conj_distrib[THEN sym] and eq_False + apply (auto elim!: allE[where x="?z"]) + apply ((rule_tac x=x in exI, force) | (rule_tac x=i in exI, force))+ + done +qed + +lemma inter_interval: fixes a :: "'a::linorder^'n::finite" shows + "{a .. b} \ {c .. d} = {(\ i. max (a$i) (c$i)) .. (\ i. min (b$i) (d$i))}" + unfolding expand_set_eq and Int_iff and mem_interval + by (auto simp add: less_divide_eq_number_of1 intro!: bexI) + +(* Moved interval_open_subset_closed a bit upwards *) + +lemma open_interval_lemma: fixes x :: "real" shows + "a < x \ x < b ==> (\d>0. \x'. abs(x' - x) < d --> a < x' \ x' < b)" + by(rule_tac x="min (x - a) (b - x)" in exI, auto) + +lemma open_interval: fixes a :: "real^'n::finite" shows "open {a<..{a<..d>0. \x'. abs (x' - (x$i)) < d \ a$i < x' \ x' < b$i" + using x[unfolded mem_interval, THEN spec[where x=i]] + using open_interval_lemma[of "a$i" "x$i" "b$i"] by auto } + + hence "\i. \d>0. \x'. abs (x' - (x$i)) < d \ a$i < x' \ x' < b$i" by auto + then obtain d where d:"\i. 0 < d i \ (\x'. \x' - x $ i\ < d i \ a $ i < x' \ x' < b $ i)" + using bchoice[of "UNIV" "\i d. d>0 \ (\x'. \x' - x $ i\ < d \ a $ i < x' \ x' < b $ i)"] by auto + + let ?d = "Min (range d)" + have **:"finite (range d)" "range d \ {}" by auto + have "?d>0" unfolding Min_gr_iff[OF **] using d by auto + moreover + { fix x' assume as:"dist x' x < ?d" + { fix i + have "\x'$i - x $ i\ < d i" + using norm_bound_component_lt[OF as[unfolded dist_norm], of i] + unfolding vector_minus_component and Min_gr_iff[OF **] by auto + hence "a $ i < x' $ i" "x' $ i < b $ i" using d[THEN spec[where x=i]] by auto } + hence "a < x' \ x' < b" unfolding vector_less_def by auto } + ultimately have "\e>0. \x'. dist x' x < e \ x' \ {a<..e>0. \x'\{a..b}. x' \ x \ dist x' x < e"(* and xab:"a$i > x$i \ b$i < x$i"*) + { assume xa:"a$i > x$i" + with as obtain y where y:"y\{a..b}" "y \ x" "dist y x < a$i - x$i" by(erule_tac x="a$i - x$i" in allE)auto + hence False unfolding mem_interval and dist_norm + using component_le_norm[of "y-x" i, unfolded vector_minus_component] and xa by(auto elim!: allE[where x=i]) + } hence "a$i \ x$i" by(rule ccontr)auto + moreover + { assume xb:"b$i < x$i" + with as obtain y where y:"y\{a..b}" "y \ x" "dist y x < x$i - b$i" by(erule_tac x="x$i - b$i" in allE)auto + hence False unfolding mem_interval and dist_norm + using component_le_norm[of "y-x" i, unfolded vector_minus_component] and xb by(auto elim!: allE[where x=i]) + } hence "x$i \ b$i" by(rule ccontr)auto + ultimately + have "a $ i \ x $ i \ x $ i \ b $ i" by auto } + thus ?thesis unfolding closed_limpt islimpt_approachable mem_interval by auto +qed + +lemma interior_closed_interval: fixes a :: "real^'n::finite" shows + "interior {a .. b} = {a<.. ?L" using interior_maximal[OF interval_open_subset_closed open_interval] by auto +next + { fix x assume "\T. open T \ x \ T \ T \ {a..b}" + 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) *\<^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 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 + unfolding vector_smult_component and basis_component using `e>0` by auto } + hence "x \ {a<.. ?R" unfolding interior_def and subset_eq by auto +qed + +lemma bounded_closed_interval: fixes a :: "real^'n::finite" shows + "bounded {a .. b}" +proof- + let ?b = "\i\UNIV. \a$i\ + \b$i\" + { fix x::"real^'n" assume x:"\i. a $ i \ x $ i \ x $ i \ b $ i" + { fix i + have "\x$i\ \ \a$i\ + \b$i\" using x[THEN spec[where x=i]] by auto } + hence "(\i\UNIV. \x $ i\) \ ?b" by(rule setsum_mono) + hence "norm x \ ?b" using norm_le_l1[of x] by auto } + thus ?thesis unfolding interval and bounded_iff by auto +qed + +lemma bounded_interval: fixes a :: "real^'n::finite" shows + "bounded {a .. b} \ bounded {a<.. UNIV) \ ({a<.. UNIV)" + using bounded_interval[of a b] + by auto + +lemma compact_interval: fixes a :: "real^'n::finite" shows + "compact {a .. b}" + 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) *\<^sub>R (a + b)) \ {a<..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) } + thus ?thesis unfolding mem_interval by auto +qed + +lemma open_closed_interval_convex: fixes x :: "real^'n::finite" + assumes x:"x \ {a<.. {a .. b}" and e:"0 < e" "e \ 1" + shows "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \ {a<.. < e * x $ i + (1 - e) * y $ i" apply(rule add_less_le_mono) + using e unfolding mult_less_cancel_left and mult_le_cancel_left apply simp_all + using x unfolding mem_interval apply simp + using y unfolding mem_interval apply simp + done + finally have "a $ i < (e *\<^sub>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) + using e unfolding mult_less_cancel_left and mult_le_cancel_left apply simp_all + using x unfolding mem_interval apply simp + using y unfolding mem_interval apply simp + done + 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 + +lemma closure_open_interval: fixes a :: "real^'n::finite" + assumes "{a<.. {}" + shows "closure {a<.. {a .. b}" + 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)) *\<^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 + { assume "\ (f ---> x) sequentially" + { fix e::real assume "e>0" + hence "\N::nat. inverse (real (N + 1)) < e" using real_arch_inv[of e] apply (auto simp add: Suc_pred') apply(rule_tac x="n - 1" in exI) by auto + 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 "((\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)) *\<^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<..a. s \ {-a<..0" and b:"\x\s. norm x \ b" using assms[unfolded bounded_pos] by auto + def a \ "(\ i. b+1)::real^'n" + { fix x assume "x\s" + fix i + have "(-a)$i < x$i" and "x$i < a$i" using b[THEN bspec[where x=x], OF `x\s`] and component_le_norm[of x i] + unfolding vector_uminus_component and a_def and Cart_lambda_beta by auto + } + thus ?thesis by(auto intro: exI[where x=a] simp add: vector_less_def) +qed + +lemma bounded_subset_open_interval: + fixes s :: "(real ^ 'n::finite) set" + shows "bounded s ==> (\a b. s \ {a<..a. s \ {-a .. a}" +proof- + obtain a where "s \ {- a<.. (\a b. s \ {a .. b})" + using bounded_subset_closed_interval_symmetric[of s] by auto + +lemma frontier_closed_interval: + fixes a b :: "real ^ _" + shows "frontier {a .. b} = {a .. b} - {a<.. {}" shows "{a<.. {c .. d} = {} \ {a<.. {c<..x::1. P x) \ P 1" + by (metis num1_eq_iff) + +lemma ex_1: "(\x::1. P x) \ P 1" + by auto (metis num1_eq_iff) + +lemma interval_cases_1: fixes x :: "real^1" shows + "x \ {a .. b} ==> x \ {a<.. (x = a) \ (x = b)" + by(simp add: Cart_eq vector_less_def vector_less_eq_def all_1, auto) + +lemma in_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 add: Cart_eq vector_less_def vector_less_eq_def all_1 dest_vec1_def) + +lemma interval_eq_empty_1: fixes a :: "real^1" shows + "{a .. b} = {} \ dest_vec1 b < dest_vec1 a" + "{a<.. dest_vec1 b \ dest_vec1 a" + unfolding interval_eq_empty and ex_1 and dest_vec1_def by auto + +lemma subset_interval_1: fixes a :: "real^1" shows + "({a .. b} \ {c .. d} \ dest_vec1 b < dest_vec1 a \ + dest_vec1 c \ dest_vec1 a \ dest_vec1 a \ dest_vec1 b \ dest_vec1 b \ dest_vec1 d)" + "({a .. b} \ {c<.. dest_vec1 b < dest_vec1 a \ + dest_vec1 c < dest_vec1 a \ dest_vec1 a \ dest_vec1 b \ dest_vec1 b < dest_vec1 d)" + "({a<.. {c .. d} \ dest_vec1 b \ dest_vec1 a \ + dest_vec1 c \ dest_vec1 a \ dest_vec1 a < dest_vec1 b \ dest_vec1 b \ dest_vec1 d)" + "({a<.. {c<.. dest_vec1 b \ dest_vec1 a \ + dest_vec1 c \ dest_vec1 a \ dest_vec1 a < dest_vec1 b \ dest_vec1 b \ dest_vec1 d)" + unfolding subset_interval[of a b c d] unfolding all_1 and dest_vec1_def by auto + +lemma eq_interval_1: fixes a :: "real^1" shows + "{a .. b} = {c .. d} \ + dest_vec1 b < dest_vec1 a \ dest_vec1 d < dest_vec1 c \ + dest_vec1 a = dest_vec1 c \ dest_vec1 b = dest_vec1 d" +using set_eq_subset[of "{a .. b}" "{c .. d}"] +using subset_interval_1(1)[of a b c d] +using subset_interval_1(1)[of c d a b] +by auto (* FIXME: slow *) + +lemma disjoint_interval_1: fixes a :: "real^1" shows + "{a .. b} \ {c .. d} = {} \ dest_vec1 b < dest_vec1 a \ dest_vec1 d < dest_vec1 c \ dest_vec1 b < dest_vec1 c \ dest_vec1 d < dest_vec1 a" + "{a .. b} \ {c<.. dest_vec1 b < dest_vec1 a \ dest_vec1 d \ dest_vec1 c \ dest_vec1 b \ dest_vec1 c \ dest_vec1 d \ dest_vec1 a" + "{a<.. {c .. d} = {} \ dest_vec1 b \ dest_vec1 a \ dest_vec1 d < dest_vec1 c \ dest_vec1 b \ dest_vec1 c \ dest_vec1 d \ dest_vec1 a" + "{a<.. {c<.. dest_vec1 b \ dest_vec1 a \ dest_vec1 d \ dest_vec1 c \ dest_vec1 b \ dest_vec1 c \ dest_vec1 d \ dest_vec1 a" + unfolding disjoint_interval and dest_vec1_def ex_1 by auto + +lemma open_closed_interval_1: fixes a :: "real^1" shows + "{a<.. dest_vec1 b ==> {a .. b} = {a<.. {a,b}" + unfolding expand_set_eq apply simp unfolding vector_less_def and vector_less_eq_def and all_1 and dest_vec1_eq[THEN sym] and dest_vec1_def by auto + +(* Some stuff for half-infinite intervals too; FIXME: notation? *) + +lemma closed_interval_left: fixes b::"real^'n::finite" + shows "closed {x::real^'n. \i. x$i \ b$i}" +proof- + { fix i + fix x::"real^'n" assume x:"\e>0. \x'\{x. \i. x $ i \ b $ i}. x' \ x \ dist x' x < e" + { assume "x$i > b$i" + then obtain y where "y $ i \ b $ i" "y \ x" "dist y x < x$i - b$i" using x[THEN spec[where x="x$i - b$i"]] by auto + hence False using component_le_norm[of "y - x" i] unfolding dist_norm and vector_minus_component by auto } + hence "x$i \ b$i" by(rule ccontr)auto } + thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast +qed + +lemma closed_interval_right: fixes a::"real^'n::finite" + shows "closed {x::real^'n. \i. a$i \ x$i}" +proof- + { fix i + fix x::"real^'n" assume x:"\e>0. \x'\{x. \i. a $ i \ x $ i}. x' \ x \ dist x' x < e" + { assume "a$i > x$i" + then obtain y where "a $ i \ y $ i" "y \ x" "dist y x < a$i - x$i" using x[THEN spec[where x="a$i - x$i"]] by auto + hence False using component_le_norm[of "y - x" i] unfolding dist_norm and vector_minus_component by auto } + hence "a$i \ x$i" by(rule ccontr)auto } + thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast +qed + +subsection{* Intervals in general, including infinite and mixtures of open and closed. *} + +definition "is_interval s \ (\a\s. \b\s. \x. (\i. ((a$i \ x$i \ x$i \ b$i) \ (b$i \ x$i \ x$i \ a$i))) \ x \ s)" + +lemma is_interval_interval: "is_interval {a .. b::real^'n::finite}" (is ?th1) "is_interval {a<..x y z::real. x < y \ y < z \ x < z" by auto + show ?th1 ?th2 unfolding is_interval_def mem_interval Ball_def atLeastAtMost_iff + by(meson real_le_trans le_less_trans less_le_trans *)+ qed + +lemma is_interval_empty: + "is_interval {}" + unfolding is_interval_def + by simp + +lemma is_interval_univ: + "is_interval UNIV" + unfolding is_interval_def + by simp + +subsection{* Closure of halfspaces and hyperplanes. *} + +lemma Lim_inner: + assumes "(f ---> l) net" shows "((\y. inner a (f y)) ---> inner a l) net" + by (intro tendsto_intros assms) + +lemma continuous_at_inner: "continuous (at x) (inner a)" + unfolding continuous_at by (intro tendsto_intros) + +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. continuous (at x) (inner a)" + unfolding continuous_at by (rule allI) (intro tendsto_intros) + hence "closed (inner a -` {..b})" + using closed_real_atMost by (rule continuous_closed_vimage) + moreover have "{x. inner a x \ b} = inner a -` {..b}" by auto + ultimately show ?thesis by simp +qed + +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. 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 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 inner_basis[OF assms] by auto + +text{* Openness of halfspaces. *} + +lemma open_halfspace_lt: "open {x. inner a x < b}" +proof- + 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. inner a x > b}" +proof- + 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 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 inner_basis[OF assms] by auto + +text{* This gives a simple derivation of limit component bounds. *} + +lemma Lim_component_le: fixes f :: "'a \ real^'n::finite" + 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. 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 + +lemma Lim_component_ge: fixes f :: "'a \ real^'n::finite" + 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. 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 + +lemma Lim_component_eq: fixes f :: "'a \ real^'n::finite" + assumes net:"(f ---> l) net" "~(trivial_limit net)" and ev:"eventually (\x. f(x)$i = b) net" + shows "l$i = b" + using ev[unfolded order_eq_iff eventually_and] using Lim_component_ge[OF net, of b i] and Lim_component_le[OF net, of i b] by auto + +lemma Lim_drop_le: fixes f :: "'a \ real^1" shows + "(f ---> l) net \ ~(trivial_limit net) \ eventually (\x. dest_vec1 (f x) \ b) net ==> dest_vec1 l \ b" + using Lim_component_le[of f l net 1 b] unfolding dest_vec1_def by auto + +lemma Lim_drop_ge: fixes f :: "'a \ real^1" shows + "(f ---> l) net \ ~(trivial_limit net) \ eventually (\x. b \ dest_vec1 (f x)) net ==> b \ dest_vec1 l" + using Lim_component_ge[of f l net b 1] unfolding dest_vec1_def by auto + +text{* Limits relative to a union. *} + +lemma eventually_within_Un: + "eventually P (net within (s \ t)) \ + eventually P (net within s) \ eventually P (net within t)" + unfolding Limits.eventually_within + by (auto elim!: eventually_rev_mp) + +lemma Lim_within_union: + "(f ---> l) (net within (s \ t)) \ + (f ---> l) (net within s) \ (f ---> l) (net within t)" + unfolding tendsto_def + by (auto simp add: eventually_within_Un) + +lemma continuous_on_union: + assumes "closed s" "closed t" "continuous_on s f" "continuous_on t f" + shows "continuous_on (s \ t) f" + using assms unfolding continuous_on unfolding Lim_within_union + unfolding Lim unfolding trivial_limit_within unfolding closed_limpt by auto + +lemma continuous_on_cases: + assumes "closed s" "closed t" "continuous_on s f" "continuous_on t g" + "\x. (x\s \ \ P x) \ (x \ t \ P x) \ f x = g x" + shows "continuous_on (s \ t) (\x. if P x then f x else g x)" +proof- + let ?h = "(\x. if P x then f x else g x)" + have "\x\s. f x = (if P x then f x else g x)" using assms(5) by auto + hence "continuous_on s ?h" using continuous_on_eq[of s f ?h] using assms(3) by auto + moreover + have "\x\t. g x = (if P x then f x else g x)" using assms(5) by auto + hence "continuous_on t ?h" using continuous_on_eq[of t g ?h] using assms(4) by auto + ultimately show ?thesis using continuous_on_union[OF assms(1,2), of ?h] by auto +qed + + +text{* Some more convenient intermediate-value theorem formulations. *} + +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. 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 + ultimately show False using assms(1)[unfolded connected_def not_ex, THEN spec[where x="?A"], THEN spec[where x="?B"]] and assms(2-5) by auto +qed + +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: inner_basis) + +text{* Also more convenient formulations of monotone convergence. *} + +lemma bounded_increasing_convergent: fixes s::"nat \ real^1" + assumes "bounded {s n| n::nat. True}" "\n. dest_vec1(s n) \ dest_vec1(s(Suc n))" + shows "\l. (s ---> l) sequentially" +proof- + obtain a where a:"\n. \dest_vec1 (s n)\ \ a" using assms(1)[unfolded bounded_iff abs_dest_vec1] by auto + { fix m::nat + 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] 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 + +subsection{* Basic homeomorphism definitions. *} + +definition "homeomorphism s t f g \ + (\x\s. (g(f x) = x)) \ (f ` s = t) \ continuous_on s f \ + (\y\t. (f(g y) = y)) \ (g ` t = s) \ continuous_on t g" + +definition + homeomorphic :: "'a::metric_space set \ 'b::metric_space set \ bool" + (infixr "homeomorphic" 60) where + homeomorphic_def: "s homeomorphic t \ (\f g. homeomorphism s t f g)" + +lemma homeomorphic_refl: "s homeomorphic s" + unfolding homeomorphic_def + unfolding homeomorphism_def + using continuous_on_id + apply(rule_tac x = "(\x. x)" in exI) + apply(rule_tac x = "(\x. x)" in exI) + by blast + +lemma homeomorphic_sym: + "s homeomorphic t \ t homeomorphic s" +unfolding homeomorphic_def +unfolding homeomorphism_def +by blast (* FIXME: slow *) + +lemma homeomorphic_trans: + assumes "s homeomorphic t" "t homeomorphic u" shows "s homeomorphic u" +proof- + obtain f1 g1 where fg1:"\x\s. g1 (f1 x) = x" "f1 ` s = t" "continuous_on s f1" "\y\t. f1 (g1 y) = y" "g1 ` t = s" "continuous_on t g1" + using assms(1) unfolding homeomorphic_def homeomorphism_def by auto + obtain f2 g2 where fg2:"\x\t. g2 (f2 x) = x" "f2 ` t = u" "continuous_on t f2" "\y\u. f2 (g2 y) = y" "g2 ` u = t" "continuous_on u g2" + using assms(2) unfolding homeomorphic_def homeomorphism_def by auto + + { fix x assume "x\s" hence "(g1 \ g2) ((f2 \ f1) x) = x" using fg1(1)[THEN bspec[where x=x]] and fg2(1)[THEN bspec[where x="f1 x"]] and fg1(2) by auto } + moreover have "(f2 \ f1) ` s = u" using fg1(2) fg2(2) by auto + moreover have "continuous_on s (f2 \ f1)" using continuous_on_compose[OF fg1(3)] and fg2(3) unfolding fg1(2) by auto + moreover { fix y assume "y\u" hence "(f2 \ f1) ((g1 \ g2) y) = y" using fg2(4)[THEN bspec[where x=y]] and fg1(4)[THEN bspec[where x="g2 y"]] and fg2(5) by auto } + moreover have "(g1 \ g2) ` u = s" using fg1(5) fg2(5) by auto + moreover have "continuous_on u (g1 \ g2)" using continuous_on_compose[OF fg2(6)] and fg1(6) unfolding fg2(5) by auto + ultimately show ?thesis unfolding homeomorphic_def homeomorphism_def apply(rule_tac x="f2 \ f1" in exI) apply(rule_tac x="g1 \ g2" in exI) by auto +qed + +lemma homeomorphic_minimal: + "s homeomorphic t \ + (\f g. (\x\s. f(x) \ t \ (g(f(x)) = x)) \ + (\y\t. g(y) \ s \ (f(g(y)) = y)) \ + continuous_on s f \ continuous_on t g)" +unfolding homeomorphic_def homeomorphism_def +apply auto apply (rule_tac x=f in exI) apply (rule_tac x=g in exI) +apply auto apply (rule_tac x=f in exI) apply (rule_tac x=g in exI) apply auto +unfolding image_iff +apply(erule_tac x="g x" in ballE) apply(erule_tac x="x" in ballE) +apply auto apply(rule_tac x="g x" in bexI) apply auto +apply(erule_tac x="f x" in ballE) apply(erule_tac x="x" in ballE) +apply auto apply(rule_tac x="f x" in bexI) by auto + +subsection{* Relatively weak hypotheses if a set is compact. *} + +definition "inv_on f s = (\x. SOME y. y\s \ f y = x)" + +lemma assumes "inj_on f s" "x\s" + shows "inv_on f s (f x) = x" + using assms unfolding inj_on_def inv_on_def by auto + +lemma homeomorphism_compact: + fixes f :: "'a::heine_borel \ 'b::heine_borel" + (* class constraint due to continuous_on_inverse *) + assumes "compact s" "continuous_on s f" "f ` s = t" "inj_on f s" + shows "\g. homeomorphism s t f g" +proof- + def g \ "\x. SOME y. y\s \ f y = x" + have g:"\x\s. g (f x) = x" using assms(3) assms(4)[unfolded inj_on_def] unfolding g_def by auto + { fix y assume "y\t" + then obtain x where x:"f x = y" "x\s" using assms(3) by auto + hence "g (f x) = x" using g by auto + hence "f (g y) = y" unfolding x(1)[THEN sym] by auto } + hence g':"\x\t. f (g x) = x" by auto + moreover + { fix x + have "x\s \ x \ g ` t" using g[THEN bspec[where x=x]] unfolding image_iff using assms(3) by(auto intro!: bexI[where x="f x"]) + moreover + { assume "x\g ` t" + then obtain y where y:"y\t" "g y = x" by auto + then obtain x' where x':"x'\s" "f x' = y" using assms(3) by auto + hence "x \ s" unfolding g_def using someI2[of "\b. b\s \ f b = y" x' "\x. x\s"] unfolding y(2)[THEN sym] and g_def by auto } + ultimately have "x\s \ x \ g ` t" by auto } + hence "g ` t = s" by auto + ultimately + show ?thesis unfolding homeomorphism_def homeomorphic_def + apply(rule_tac x=g in exI) using g and assms(3) and continuous_on_inverse[OF assms(2,1), of g, unfolded assms(3)] and assms(2) by auto +qed + +lemma homeomorphic_compact: + fixes f :: "'a::heine_borel \ 'b::heine_borel" + (* class constraint due to continuous_on_inverse *) + shows "compact s \ continuous_on s f \ (f ` s = t) \ inj_on f s + \ s homeomorphic t" + unfolding homeomorphic_def by(metis homeomorphism_compact) + +text{* Preservation of topological properties. *} + +lemma homeomorphic_compactness: + "s homeomorphic t ==> (compact s \ compact t)" +unfolding homeomorphic_def homeomorphism_def +by (metis compact_continuous_image) + +text{* Results on translation, scaling etc. *} + +lemma homeomorphic_scaling: + fixes s :: "'a::real_normed_vector set" + assumes "c \ 0" shows "s homeomorphic ((\x. c *\<^sub>R x) ` s)" + unfolding homeomorphic_minimal + 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: + fixes s :: "'a::real_normed_vector set" + shows "s homeomorphic ((\x. a + x) ` s)" + unfolding homeomorphic_minimal + apply(rule_tac x="\x. a + x" in exI) + apply(rule_tac x="\x. -a + x" in exI) + using continuous_on_add[OF continuous_on_const continuous_on_id] by auto + +lemma homeomorphic_affinity: + fixes s :: "'a::real_normed_vector set" + assumes "c \ 0" shows "s homeomorphic ((\x. a + c *\<^sub>R x) ` s)" +proof- + 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 *\<^sub>R x) ` s" a] unfolding * by auto +qed + +lemma homeomorphic_balls: + fixes a b ::"'a::real_normed_vector" (* FIXME: generalize to metric_space *) + assumes "0 < d" "0 < e" + shows "(ball a d) homeomorphic (ball b e)" (is ?th) + "(cball a d) homeomorphic (cball b e)" (is ?cth) +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) *\<^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) *\<^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. *} + +lemma cauchy_isometric: + fixes x :: "nat \ real ^ 'n::finite" + assumes e:"0 < e" and s:"subspace s" and f:"bounded_linear f" and normf:"\x\s. norm(f x) \ e * norm(x)" and xs:"\n::nat. x n \ s" and cf:"Cauchy(f o x)" + shows "Cauchy x" +proof- + interpret f: bounded_linear f by fact + { fix d::real assume "d>0" + then obtain N where N:"\n\N. norm (f (x n) - f (x N)) < e * d" + using cf[unfolded cauchy o_def dist_norm, THEN spec[where x="e*d"]] and e and mult_pos_pos[of e d] by auto + { fix n assume "n\N" + hence "norm (f (x n - x N)) < e * d" using N[THEN spec[where x=n]] unfolding f.diff[THEN sym] by auto + moreover have "e * norm (x n - x N) \ norm (f (x n - x N))" + using subspace_sub[OF s, of "x n" "x N"] using xs[THEN spec[where x=N]] and xs[THEN spec[where x=n]] + using normf[THEN bspec[where x="x n - x N"]] by auto + ultimately have "norm (x n - x N) < d" using `e>0` + using mult_left_less_imp_less[of e "norm (x n - x N)" d] by auto } + hence "\N. \n\N. norm (x n - x N) < d" by auto } + thus ?thesis unfolding cauchy and dist_norm by auto +qed + +lemma complete_isometric_image: + fixes f :: "real ^ _ \ real ^ _" + assumes "0 < e" and s:"subspace s" and f:"bounded_linear f" and normf:"\x\s. norm(f x) \ e * norm(x)" and cs:"complete s" + shows "complete(f ` s)" +proof- + { fix g assume as:"\n::nat. g n \ f ` s" and cfg:"Cauchy g" + then obtain x where "\n. x n \ s \ g n = f (x n)" unfolding image_iff and Bex_def + using choice[of "\ n xa. xa \ s \ g n = f xa"] by auto + hence x:"\n. x n \ s" "\n. g n = f (x n)" by auto + hence "f \ x = g" unfolding expand_fun_eq by auto + then obtain l where "l\s" and l:"(x ---> l) sequentially" + using cs[unfolded complete_def, THEN spec[where x="x"]] + using cauchy_isometric[OF `0l\f ` s. (g ---> l) sequentially" + using linear_continuous_at[OF f, unfolded continuous_at_sequentially, THEN spec[where x=x], of l] + unfolding `f \ x = g` by auto } + thus ?thesis unfolding complete_def by auto +qed + +lemma dist_0_norm: + fixes x :: "'a::real_normed_vector" + shows "dist 0 x = norm x" +unfolding dist_norm by simp + +lemma injective_imp_isometric: fixes f::"real^'m::finite \ real^'n::finite" + assumes s:"closed s" "subspace s" and f:"bounded_linear f" "\x\s. (f x = 0) \ (x = 0)" + shows "\e>0. \x\s. norm (f x) \ e * norm(x)" +proof(cases "s \ {0::real^'m}") + case True + { fix x assume "x \ s" + hence "x = 0" using True by auto + hence "norm x \ norm (f x)" by auto } + thus ?thesis by(auto intro!: exI[where x=1]) +next + interpret f: bounded_linear f by fact + case False + then obtain a where a:"a\0" "a\s" by auto + from False have "s \ {}" by auto + let ?S = "{f x| x. (x \ s \ norm x = norm a)}" + let ?S' = "{x::real^'m. x\s \ norm x = norm a}" + let ?S'' = "{x::real^'m. norm x = norm a}" + + have "?S'' = frontier(cball 0 (norm a))" unfolding frontier_cball and dist_norm by (auto simp add: norm_minus_cancel) + hence "compact ?S''" using compact_frontier[OF compact_cball, of 0 "norm a"] by auto + moreover have "?S' = s \ ?S''" by auto + ultimately have "compact ?S'" using closed_inter_compact[of s ?S''] using s(1) by auto + moreover have *:"f ` ?S' = ?S" by auto + ultimately have "compact ?S" using compact_continuous_image[OF linear_continuous_on[OF f(1)], of ?S'] by auto + hence "closed ?S" using compact_imp_closed by auto + moreover have "?S \ {}" using a by auto + ultimately obtain b' where "b'\?S" "\y\?S. norm b' \ norm y" using distance_attains_inf[of ?S 0] unfolding dist_0_norm by auto + then obtain b where "b\s" and ba:"norm b = norm a" and b:"\x\{x \ s. norm x = norm a}. norm (f b) \ norm (f x)" unfolding *[THEN sym] unfolding image_iff by auto + + let ?e = "norm (f b) / norm b" + have "norm b > 0" using ba and a and norm_ge_zero by auto + moreover have "norm (f b) > 0" using f(2)[THEN bspec[where x=b], OF `b\s`] using `norm b >0` unfolding zero_less_norm_iff by auto + ultimately have "0 < norm (f b) / norm b" by(simp only: divide_pos_pos) + moreover + { fix x assume "x\s" + hence "norm (f b) / norm b * norm x \ norm (f x)" + proof(cases "x=0") + case True thus "norm (f b) / norm b * norm x \ norm (f x)" by auto + 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 *\<^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 f.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 + show ?thesis by auto +qed + +lemma closed_injective_image_subspace: + fixes f :: "real ^ _ \ real ^ _" + assumes "subspace s" "bounded_linear f" "\x\s. f x = 0 --> x = 0" "closed s" + shows "closed(f ` s)" +proof- + obtain e where "e>0" and e:"\x\s. e * norm x \ norm (f x)" using injective_imp_isometric[OF assms(4,1,2,3)] by auto + show ?thesis using complete_isometric_image[OF `e>0` assms(1,2) e] and assms(4) + unfolding complete_eq_closed[THEN sym] by auto +qed + +subsection{* Some properties of a canonical subspace. *} + +lemma subspace_substandard: + "subspace {x::real^'n. (\i. P i \ x$i = 0)}" + unfolding subspace_def by(auto simp add: vector_add_component vector_smult_component elim!: ballE) + +lemma closed_substandard: + "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. 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: 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. 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 + thus ?thesis by(auto simp add: closed_Inter closed_hyperplane) +qed + +lemma dim_substandard: + shows "dim {x::real^'n::finite. \i. i \ d \ x$i = 0} = card d" (is "dim ?A = _") +proof- + let ?D = "UNIV::'n set" + let ?B = "(basis::'n\real^'n) ` d" + + let ?bas = "basis::'n \ real^'n" + + have "?B \ ?A" by auto + + moreover + { fix x::"real^'n" assume "x\?A" + with finite[of d] + have "x\ span ?B" + proof(induct d arbitrary: x) + case empty hence "x=0" unfolding Cart_eq by auto + thus ?case using subspace_0[OF subspace_span[of "{}"]] by auto + next + 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 *\<^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 + using *[THEN spec[where x=i]] by auto } + hence "y \ span (basis ` (insert k F))" using insert(3) + using span_mono[of "?bas ` F" "?bas ` (insert k F)"] + 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 *\<^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 *\<^sub>R basis k \ span (?bas ` (insert k F))" + using span_add by auto + thus ?case using y by auto + qed + } + hence "?A \ span ?B" by auto + + moreover + { fix x assume "x \ ?B" + hence "x\{(basis i)::real^'n |i. i \ ?D}" using assms by auto } + hence "independent ?B" using independent_mono[OF independent_stdbasis, of ?B] and assms by auto + + moreover + have "d \ ?D" unfolding subset_eq using assms by auto + hence *:"inj_on (basis::'n\real^'n) d" using subset_inj_on[OF basis_inj, of "d"] by auto + have "?B hassize (card d)" unfolding hassize_def and card_image[OF *] by auto + + ultimately show ?thesis using dim_unique[of "basis ` d" ?A] by auto +qed + +text{* Hence closure and completeness of all subspaces. *} + +lemma closed_subspace_lemma: "n \ card (UNIV::'n::finite set) \ \A::'n set. card A = n" +apply (induct n) +apply (rule_tac x="{}" in exI, simp) +apply clarsimp +apply (subgoal_tac "\x. x \ A") +apply (erule exE) +apply (rule_tac x="insert x A" in exI, simp) +apply (subgoal_tac "A \ UNIV", auto) +done + +lemma closed_subspace: fixes s::"(real^'n::finite) set" + assumes "subspace s" shows "closed s" +proof- + have "dim s \ card (UNIV :: 'n set)" using dim_subset_univ by auto + then obtain d::"'n set" where t: "card d = dim s" + using closed_subspace_lemma by auto + let ?t = "{x::real^'n. \i. i \ d \ x$i = 0}" + obtain f where f:"bounded_linear f" "f ` ?t = s" "inj_on f ?t" + using subspace_isomorphism[unfolded linear_conv_bounded_linear, OF subspace_substandard[of "\i. i \ d"] assms] + using dim_substandard[of d] and t by auto + interpret f: bounded_linear f by fact + have "\x\?t. f x = 0 \ x = 0" using f.zero using f(3)[unfolded inj_on_def] + by(erule_tac x=0 in ballE) auto + moreover have "closed ?t" using closed_substandard . + moreover have "subspace ?t" using subspace_substandard . + ultimately show ?thesis using closed_injective_image_subspace[of ?t f] + unfolding f(2) using f(1) by auto +qed + +lemma complete_subspace: + fixes s :: "(real ^ _) set" shows "subspace s ==> complete s" + using complete_eq_closed closed_subspace + by auto + +lemma dim_closure: + fixes s :: "(real ^ _) set" + shows "dim(closure s) = dim s" (is "?dc = ?d") +proof- + have "?dc \ ?d" using closure_minimal[OF span_inc, of s] + using closed_subspace[OF subspace_span, of s] + using dim_subset[of "closure s" "span s"] unfolding dim_span by auto + thus ?thesis using dim_subset[OF closure_subset, of s] by auto +qed + +text{* Affine transformations of intervals. *} + +lemma affinity_inverses: + assumes m0: "m \ (0::'a::field)" + shows "(\x. m *s x + c) o (\x. inverse(m) *s x + (-(inverse(m) *s c))) = id" + "(\x. inverse(m) *s x + (-(inverse(m) *s c))) o (\x. m *s x + c) = id" + using m0 +apply (auto simp add: expand_fun_eq vector_add_ldistrib vector_smult_assoc) +by (simp add: vector_smult_lneg[symmetric] vector_smult_assoc vector_sneg_minus1[symmetric]) + +lemma real_affinity_le: + "0 < (m::'a::ordered_field) ==> (m * x + c \ y \ x \ inverse(m) * y + -(c / m))" + by (simp add: field_simps inverse_eq_divide) + +lemma real_le_affinity: + "0 < (m::'a::ordered_field) ==> (y \ m * x + c \ inverse(m) * y + -(c / m) \ x)" + by (simp add: field_simps inverse_eq_divide) + +lemma real_affinity_lt: + "0 < (m::'a::ordered_field) ==> (m * x + c < y \ x < inverse(m) * y + -(c / m))" + by (simp add: field_simps inverse_eq_divide) + +lemma real_lt_affinity: + "0 < (m::'a::ordered_field) ==> (y < m * x + c \ inverse(m) * y + -(c / m) < x)" + by (simp add: field_simps inverse_eq_divide) + +lemma real_affinity_eq: + "(m::'a::ordered_field) \ 0 ==> (m * x + c = y \ x = inverse(m) * y + -(c / m))" + by (simp add: field_simps inverse_eq_divide) + +lemma real_eq_affinity: + "(m::'a::ordered_field) \ 0 ==> (y = m * x + c \ inverse(m) * y + -(c / m) = x)" + by (simp add: field_simps inverse_eq_divide) + +lemma vector_affinity_eq: + assumes m0: "(m::'a::field) \ 0" + shows "m *s x + c = y \ x = inverse m *s y + -(inverse m *s c)" +proof + assume h: "m *s x + c = y" + hence "m *s x = y - c" by (simp add: ring_simps) + hence "inverse m *s (m *s x) = inverse m *s (y - c)" by simp + then show "x = inverse m *s y + - (inverse m *s c)" + using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib) +next + assume h: "x = inverse m *s y + - (inverse m *s c)" + show "m *s x + c = y" unfolding h diff_minus[symmetric] + using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib) +qed + +lemma vector_eq_affinity: + "(m::'a::field) \ 0 ==> (y = m *s x + c \ inverse(m) *s y + -(inverse(m) *s c) = x)" + using vector_affinity_eq[where m=m and x=x and y=y and c=c] + by metis + +lemma image_affinity_interval: fixes m::real + fixes a b c :: "real^'n::finite" + shows "(\x. m *\<^sub>R x + c) ` {a .. b} = + (if {a .. b} = {} then {} + 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 *\<^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 *\<^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 *\<^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 *\<^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) *\<^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 *\<^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) *\<^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 *\<^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...) *} + +lemma banach_fix: + assumes s:"complete s" "s \ {}" and c:"0 \ c" "c < 1" and f:"(f ` s) \ s" and + lipschitz:"\x\s. \y\s. dist (f x) (f y) \ c * dist x y" + shows "\! x\s. (f x = x)" +proof- + have "1 - c > 0" using c by auto + + from s(2) obtain z0 where "z0 \ s" by auto + def z \ "\n. (f ^^ n) z0" + { fix n::nat + have "z n \ s" unfolding z_def + proof(induct n) case 0 thus ?case using `z0 \s` by auto + next case Suc thus ?case using f by auto qed } + note z_in_s = this + + def d \ "dist (z 0) (z 1)" + + have fzn:"\n. f (z n) = z (Suc n)" unfolding z_def by auto + { fix n::nat + have "dist (z n) (z (Suc n)) \ (c ^ n) * d" + proof(induct n) + case 0 thus ?case unfolding d_def by auto + next + case (Suc m) + hence "c * dist (z m) (z (Suc m)) \ c ^ Suc m * d" + using `0 \ c` using mult_mono1_class.mult_mono1[of "dist (z m) (z (Suc m))" "c ^ m * d" c] by auto + thus ?case using lipschitz[THEN bspec[where x="z m"], OF z_in_s, THEN bspec[where x="z (Suc m)"], OF z_in_s] + unfolding fzn and mult_le_cancel_left by auto + qed + } note cf_z = this + + { fix n m::nat + have "(1 - c) * dist (z m) (z (m+n)) \ (c ^ m) * d * (1 - c ^ n)" + proof(induct n) + case 0 show ?case by auto + next + case (Suc k) + have "(1 - c) * dist (z m) (z (m + Suc k)) \ (1 - c) * (dist (z m) (z (m + k)) + dist (z (m + k)) (z (Suc (m + k))))" + using dist_triangle and c by(auto simp add: dist_triangle) + also have "\ \ (1 - c) * (dist (z m) (z (m + k)) + c ^ (m + k) * d)" + using cf_z[of "m + k"] and c by auto + also have "\ \ c ^ m * d * (1 - c ^ k) + (1 - c) * c ^ (m + k) * d" + using Suc by (auto simp add: ring_simps) + also have "\ = (c ^ m) * (d * (1 - c ^ k) + (1 - c) * c ^ k * d)" + unfolding power_add by (auto simp add: ring_simps) + also have "\ \ (c ^ m) * d * (1 - c ^ Suc k)" + using c by (auto simp add: ring_simps) + finally show ?case by auto + qed + } note cf_z2 = this + { fix e::real assume "e>0" + hence "\N. \m n. N \ m \ N \ n \ dist (z m) (z n) < e" + proof(cases "d = 0") + case True + hence "\n. z n = z0" using cf_z2[of 0] and c unfolding z_def by (auto simp add: pos_prod_le[OF `1 - c > 0`]) + thus ?thesis using `e>0` by auto + next + case False hence "d>0" unfolding d_def using zero_le_dist[of "z 0" "z 1"] + by (metis False d_def real_less_def) + hence "0 < e * (1 - c) / d" using `e>0` and `1-c>0` + using divide_pos_pos[of "e * (1 - c)" d] and mult_pos_pos[of e "1 - c"] by auto + then obtain N where N:"c ^ N < e * (1 - c) / d" using real_arch_pow_inv[of "e * (1 - c) / d" c] and c by auto + { fix m n::nat assume "m>n" and as:"m\N" "n\N" + have *:"c ^ n \ c ^ N" using `n\N` and c using power_decreasing[OF `n\N`, of c] by auto + have "1 - c ^ (m - n) > 0" using c and power_strict_mono[of c 1 "m - n"] using `m>n` by auto + hence **:"d * (1 - c ^ (m - n)) / (1 - c) > 0" + using real_mult_order[OF `d>0`, of "1 - c ^ (m - n)"] + using divide_pos_pos[of "d * (1 - c ^ (m - n))" "1 - c"] + using `0 < 1 - c` by auto + + have "dist (z m) (z n) \ c ^ n * d * (1 - c ^ (m - n)) / (1 - c)" + using cf_z2[of n "m - n"] and `m>n` unfolding pos_le_divide_eq[OF `1-c>0`] + by (auto simp add: real_mult_commute dist_commute) + also have "\ \ c ^ N * d * (1 - c ^ (m - n)) / (1 - c)" + using mult_right_mono[OF * order_less_imp_le[OF **]] + unfolding real_mult_assoc by auto + also have "\ < (e * (1 - c) / d) * d * (1 - c ^ (m - n)) / (1 - c)" + using mult_strict_right_mono[OF N **] unfolding real_mult_assoc by auto + also have "\ = e * (1 - c ^ (m - n))" using c and `d>0` and `1 - c > 0` by auto + also have "\ \ e" using c and `1 - c ^ (m - n) > 0` and `e>0` using mult_right_le_one_le[of e "1 - c ^ (m - n)"] by auto + finally have "dist (z m) (z n) < e" by auto + } note * = this + { fix m n::nat assume as:"N\m" "N\n" + hence "dist (z n) (z m) < e" + proof(cases "n = m") + case True thus ?thesis using `e>0` by auto + next + case False thus ?thesis using as and *[of n m] *[of m n] unfolding nat_neq_iff by (auto simp add: dist_commute) + qed } + thus ?thesis by auto + qed + } + hence "Cauchy z" unfolding cauchy_def by auto + then obtain x where "x\s" and x:"(z ---> x) sequentially" using s(1)[unfolded compact_def complete_def, THEN spec[where x=z]] and z_in_s by auto + + def e \ "dist (f x) x" + have "e = 0" proof(rule ccontr) + assume "e \ 0" hence "e>0" unfolding e_def using zero_le_dist[of "f x" x] + by (metis dist_eq_0_iff dist_nz e_def) + then obtain N where N:"\n\N. dist (z n) x < e / 2" + using x[unfolded Lim_sequentially, THEN spec[where x="e/2"]] by auto + hence N':"dist (z N) x < e / 2" by auto + + have *:"c * dist (z N) x \ dist (z N) x" unfolding mult_le_cancel_right2 + using zero_le_dist[of "z N" x] and c + by (metis dist_eq_0_iff dist_nz order_less_asym real_less_def) + have "dist (f (z N)) (f x) \ c * dist (z N) x" using lipschitz[THEN bspec[where x="z N"], THEN bspec[where x=x]] + using z_in_s[of N] `x\s` using c by auto + also have "\ < e / 2" using N' and c using * by auto + finally show False unfolding fzn + using N[THEN spec[where x="Suc N"]] and dist_triangle_half_r[of "z (Suc N)" "f x" e x] + unfolding e_def by auto + qed + hence "f x = x" unfolding e_def by auto + moreover + { fix y assume "f y = y" "y\s" + hence "dist x y \ c * dist x y" using lipschitz[THEN bspec[where x=x], THEN bspec[where x=y]] + using `x\s` and `f x = x` by auto + hence "dist x y = 0" unfolding mult_le_cancel_right1 + using c and zero_le_dist[of x y] by auto + hence "y = x" by auto + } + ultimately show ?thesis unfolding Bex1_def using `x\s` by blast+ +qed + +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\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 + { fix y assume "y\s" + hence "x = y" using `x\s` and dist[THEN bspec[where x=x], THEN bspec[where x=y]] + unfolding g[THEN bspec[where x=x], OF `x\s`] + unfolding g[THEN bspec[where x=y], OF `y\s`] by auto } + thus ?thesis unfolding Bex1_def using `x\s` and g by blast+ +next + case True + then obtain x where [simp]:"x\s" and "g x \ x" by auto + { fix x y assume "x \ s" "y \ s" + hence "dist (g x) (g y) \ dist x y" + using dist[THEN bspec[where x=x], THEN bspec[where x=y]] by auto } note dist' = this + def y \ "g x" + have [simp]:"y\s" unfolding y_def using gs[unfolded image_subset_iff] and `x\s` by blast + def f \ "\n. g ^^ n" + have [simp]:"\n z. g (f n z) = f (Suc n) z" unfolding f_def by auto + have [simp]:"\z. f 0 z = z" unfolding f_def by auto + { fix n::nat and z assume "z\s" + have "f n z \ s" unfolding f_def + proof(induct n) + case 0 thus ?case using `z\s` by simp + next + case (Suc n) thus ?case using gs[unfolded image_subset_iff] by auto + qed } note fs = this + { fix m n ::nat assume "m\n" + fix w z assume "w\s" "z\s" + have "dist (f n w) (f n z) \ dist (f m w) (f m z)" using `m\n` + proof(induct n) + case 0 thus ?case by auto + next + case (Suc n) + thus ?case proof(cases "m\n") + case True thus ?thesis using Suc(1) + using dist'[OF fs fs, OF `w\s` `z\s`, of n n] by auto + next + case False hence mn:"m = Suc n" using Suc(2) by simp + show ?thesis unfolding mn by auto + qed + qed } note distf = this + + 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 \ "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 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::'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 + + { assume as:"dist a b > dist (f n x) (f n y)" + then obtain Na Nb where "\m\Na. dist (f (r m) x) a < (dist a b - dist (f n x) (f n y)) / 2" + and "\m\Nb. dist (f (r m) y) b < (dist a b - dist (f n x) (f n y)) / 2" + using lima limb unfolding h_def Lim_sequentially by (fastsimp simp del: less_divide_eq_number_of1) + hence "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)" + apply(erule_tac x="Na+Nb+n" in allE) + apply(erule_tac x="Na+Nb+n" in allE) apply simp + using dist_triangle_add_half[of a "f (r (Na + Nb + n)) x" "dist a b - dist (f n x) (f n y)" + "-b" "- f (r (Na + Nb + n)) y"] + unfolding ** unfolding group_simps(12) by (auto simp add: dist_commute) + 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 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 + } + hence "dist a b \ dist (f n x) (f n y)" by(rule ccontr)auto } + note ab_fn = this + + have [simp]:"a = b" proof(rule ccontr) + def e \ "dist a b - dist (g a) (g b)" + assume "a\b" hence "e > 0" unfolding e_def using dist by fastsimp + hence "\n. dist (f n x) a < e/2 \ dist (f n y) b < e/2" + using lima limb unfolding Lim_sequentially + apply (auto elim!: allE[where x="e/2"]) apply(rule_tac x="r (max N Na)" in exI) unfolding h_def by fastsimp + then obtain n where n:"dist (f n x) a < e/2 \ dist (f n y) b < e/2" by auto + have "dist (f (Suc n) x) (g a) \ dist (f n x) a" + using dist[THEN bspec[where x="f n x"], THEN bspec[where x="a"]] and fs by auto + moreover have "dist (f (Suc n) y) (g b) \ dist (f n y) b" + using dist[THEN bspec[where x="f n y"], THEN bspec[where x="b"]] and fs by auto + ultimately have "dist (f (Suc n) x) (g a) + dist (f (Suc n) y) (g b) < e" using n by auto + thus False unfolding e_def using ab_fn[of "Suc n"] by norm + qed + + have [simp]:"\n. f (Suc n) x = f n y" unfolding f_def y_def by(induct_tac n)auto + { fix x y assume "x\s" "y\s" moreover + fix e::real assume "e>0" ultimately + 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 "((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 + moreover + { fix x assume "x\s" "g x = x" "x\a" + hence "False" using dist[THEN bspec[where x=a], THEN bspec[where x=x]] + using `g a = a` and `a\s` by auto } + ultimately show "\!x\s. g x = x" unfolding Bex1_def using `a\s` by blast +qed + +end