# HG changeset patch # User wenzelm # Date 1256553020 -3600 # Node ID f35905921049a22f70d46571498e5ea01aa5808c # Parent 70522979c7be305b2f02c36b19ff8be4cd65f51b deleted junk; diff -r 70522979c7be -r f35905921049 src/HOL/Library/#Topology_Euclidean_Space.thy# --- a/src/HOL/Library/#Topology_Euclidean_Space.thy# Mon Oct 26 11:30:08 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,6029 +0,0 @@ -(* 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