# HG changeset patch # User hoelzl # Date 1364296861 -3600 # Node ID f415febf42347e9c4b4e7c087be2bd89635356bb # Parent 609914f0934a8409376c0022841b263f0571fab9 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces diff -r 609914f0934a -r f415febf4234 src/HOL/Limits.thy --- a/src/HOL/Limits.thy Tue Mar 26 12:21:00 2013 +0100 +++ b/src/HOL/Limits.thy Tue Mar 26 12:21:01 2013 +0100 @@ -57,7 +57,21 @@ "at_bot \ (at_infinity :: real filter)" unfolding at_infinity_eq_at_top_bot by simp -subsection {* Boundedness *} +subsubsection {* Boundedness *} + +definition Bfun :: "('a \ 'b::metric_space) \ 'a filter \ bool" where + Bfun_metric_def: "Bfun f F = (\y. \K>0. eventually (\x. dist (f x) y \ K) F)" + +abbreviation Bseq :: "(nat \ 'a::metric_space) \ bool" where + "Bseq X \ Bfun X sequentially" + +lemma Bseq_conv_Bfun: "Bseq X \ Bfun X sequentially" .. + +lemma Bseq_ignore_initial_segment: "Bseq X \ Bseq (\n. X (n + k))" + unfolding Bfun_metric_def by (subst eventually_sequentially_seg) + +lemma Bseq_offset: "Bseq (\n. X (n + k)) \ Bseq X" + unfolding Bfun_metric_def by (subst (asm) eventually_sequentially_seg) lemma Bfun_def: "Bfun f F \ (\K>0. eventually (\x. norm (f x) \ K) F)" @@ -87,6 +101,154 @@ obtains B where "0 < B" and "eventually (\x. norm (f x) \ B) F" using assms unfolding Bfun_def by fast +lemma Cauchy_Bseq: "Cauchy X \ Bseq X" + unfolding Cauchy_def Bfun_metric_def eventually_sequentially + apply (erule_tac x=1 in allE) + apply simp + apply safe + apply (rule_tac x="X M" in exI) + apply (rule_tac x=1 in exI) + apply (erule_tac x=M in allE) + apply simp + apply (rule_tac x=M in exI) + apply (auto simp: dist_commute) + done + + +subsubsection {* Bounded Sequences *} + +lemma BseqI': "(\n. norm (X n) \ K) \ Bseq X" + by (intro BfunI) (auto simp: eventually_sequentially) + +lemma BseqI2': "\n\N. norm (X n) \ K \ Bseq X" + by (intro BfunI) (auto simp: eventually_sequentially) + +lemma Bseq_def: "Bseq X \ (\K>0. \n. norm (X n) \ K)" + unfolding Bfun_def eventually_sequentially +proof safe + fix N K assume "0 < K" "\n\N. norm (X n) \ K" + then show "\K>0. \n. norm (X n) \ K" + by (intro exI[of _ "max (Max (norm ` X ` {..N})) K"] min_max.less_supI2) + (auto intro!: imageI not_less[where 'a=nat, THEN iffD1] Max_ge simp: le_max_iff_disj) +qed auto + +lemma BseqE: "\Bseq X; \K. \0 < K; \n. norm (X n) \ K\ \ Q\ \ Q" +unfolding Bseq_def by auto + +lemma BseqD: "Bseq X ==> \K. 0 < K & (\n. norm (X n) \ K)" +by (simp add: Bseq_def) + +lemma BseqI: "[| 0 < K; \n. norm (X n) \ K |] ==> Bseq X" +by (auto simp add: Bseq_def) + +lemma lemma_NBseq_def: + "(\K > 0. \n. norm (X n) \ K) = (\N. \n. norm (X n) \ real(Suc N))" +proof safe + fix K :: real + from reals_Archimedean2 obtain n :: nat where "K < real n" .. + then have "K \ real (Suc n)" by auto + moreover assume "\m. norm (X m) \ K" + ultimately have "\m. norm (X m) \ real (Suc n)" + by (blast intro: order_trans) + then show "\N. \n. norm (X n) \ real (Suc N)" .. +qed (force simp add: real_of_nat_Suc) + +text{* alternative definition for Bseq *} +lemma Bseq_iff: "Bseq X = (\N. \n. norm (X n) \ real(Suc N))" +apply (simp add: Bseq_def) +apply (simp (no_asm) add: lemma_NBseq_def) +done + +lemma lemma_NBseq_def2: + "(\K > 0. \n. norm (X n) \ K) = (\N. \n. norm (X n) < real(Suc N))" +apply (subst lemma_NBseq_def, auto) +apply (rule_tac x = "Suc N" in exI) +apply (rule_tac [2] x = N in exI) +apply (auto simp add: real_of_nat_Suc) + prefer 2 apply (blast intro: order_less_imp_le) +apply (drule_tac x = n in spec, simp) +done + +(* yet another definition for Bseq *) +lemma Bseq_iff1a: "Bseq X = (\N. \n. norm (X n) < real(Suc N))" +by (simp add: Bseq_def lemma_NBseq_def2) + +subsubsection{*A Few More Equivalence Theorems for Boundedness*} + +text{*alternative formulation for boundedness*} +lemma Bseq_iff2: "Bseq X = (\k > 0. \x. \n. norm (X(n) + -x) \ k)" +apply (unfold Bseq_def, safe) +apply (rule_tac [2] x = "k + norm x" in exI) +apply (rule_tac x = K in exI, simp) +apply (rule exI [where x = 0], auto) +apply (erule order_less_le_trans, simp) +apply (drule_tac x=n in spec, fold diff_minus) +apply (drule order_trans [OF norm_triangle_ineq2]) +apply simp +done + +text{*alternative formulation for boundedness*} +lemma Bseq_iff3: "Bseq X = (\k > 0. \N. \n. norm(X(n) + -X(N)) \ k)" +apply safe +apply (simp add: Bseq_def, safe) +apply (rule_tac x = "K + norm (X N)" in exI) +apply auto +apply (erule order_less_le_trans, simp) +apply (rule_tac x = N in exI, safe) +apply (drule_tac x = n in spec) +apply (rule order_trans [OF norm_triangle_ineq], simp) +apply (auto simp add: Bseq_iff2) +done + +lemma BseqI2: "(\n. k \ f n & f n \ (K::real)) ==> Bseq f" +apply (simp add: Bseq_def) +apply (rule_tac x = " (\k\ + \K\) + 1" in exI, auto) +apply (drule_tac x = n in spec, arith) +done + +subsubsection{*Upper Bounds and Lubs of Bounded Sequences*} + +lemma Bseq_isUb: + "!!(X::nat=>real). Bseq X ==> \U. isUb (UNIV::real set) {x. \n. X n = x} U" +by (auto intro: isUbI setleI simp add: Bseq_def abs_le_iff) + +text{* Use completeness of reals (supremum property) + to show that any bounded sequence has a least upper bound*} + +lemma Bseq_isLub: + "!!(X::nat=>real). Bseq X ==> + \U. isLub (UNIV::real set) {x. \n. X n = x} U" +by (blast intro: reals_complete Bseq_isUb) + +lemma Bseq_minus_iff: "Bseq (%n. -(X n) :: 'a :: real_normed_vector) = Bseq X" + by (simp add: Bseq_def) + +lemma Bseq_eq_bounded: "range f \ {a .. b::real} \ Bseq f" + apply (simp add: subset_eq) + apply (rule BseqI'[where K="max (norm a) (norm b)"]) + apply (erule_tac x=n in allE) + apply auto + done + +lemma incseq_bounded: "incseq X \ \i. X i \ (B::real) \ Bseq X" + by (intro Bseq_eq_bounded[of X "X 0" B]) (auto simp: incseq_def) + +lemma decseq_bounded: "decseq X \ \i. (B::real) \ X i \ Bseq X" + by (intro Bseq_eq_bounded[of X B "X 0"]) (auto simp: decseq_def) + +subsection {* Bounded Monotonic Sequences *} + +subsubsection{*A Bounded and Monotonic Sequence Converges*} + +(* TODO: delete *) +(* FIXME: one use in NSA/HSEQ.thy *) +lemma Bmonoseq_LIMSEQ: "\n. m \ n --> X n = X m ==> \L. (X ----> L)" + apply (rule_tac x="X m" in exI) + apply (rule filterlim_cong[THEN iffD2, OF refl refl _ tendsto_const]) + unfolding eventually_sequentially + apply blast + done + subsection {* Convergence to Zero *} definition Zfun :: "('a \ 'b::real_normed_vector) \ 'a filter \ bool" @@ -242,6 +404,37 @@ subsubsection {* Distance and norms *} +lemma tendsto_dist [tendsto_intros]: + fixes l m :: "'a :: metric_space" + assumes f: "(f ---> l) F" and g: "(g ---> m) F" + shows "((\x. dist (f x) (g x)) ---> dist l m) F" +proof (rule tendstoI) + fix e :: real assume "0 < e" + hence e2: "0 < e/2" by simp + from tendstoD [OF f e2] tendstoD [OF g e2] + show "eventually (\x. dist (dist (f x) (g x)) (dist l m) < e) F" + proof (eventually_elim) + case (elim x) + then show "dist (dist (f x) (g x)) (dist l m) < e" + unfolding dist_real_def + using dist_triangle2 [of "f x" "g x" "l"] + using dist_triangle2 [of "g x" "l" "m"] + using dist_triangle3 [of "l" "m" "f x"] + using dist_triangle [of "f x" "m" "g x"] + by arith + qed +qed + +lemma continuous_dist[continuous_intros]: + fixes f g :: "_ \ 'a :: metric_space" + shows "continuous F f \ continuous F g \ continuous F (\x. dist (f x) (g x))" + unfolding continuous_def by (rule tendsto_dist) + +lemma continuous_on_dist[continuous_on_intros]: + fixes f g :: "_ \ 'a :: metric_space" + shows "continuous_on s f \ continuous_on s g \ continuous_on s (\x. dist (f x) (g x))" + unfolding continuous_on_def by (auto intro: tendsto_dist) + lemma tendsto_norm [tendsto_intros]: "(f ---> a) F \ ((\x. norm (f x)) ---> norm a) F" unfolding norm_conv_dist by (intro tendsto_intros) @@ -1152,123 +1345,6 @@ apply (drule tendsto_minus, auto) done -subsection {* Bounded Monotonic Sequences *} - -subsubsection {* Bounded Sequences *} - -lemma BseqI': "(\n. norm (X n) \ K) \ Bseq X" - by (intro BfunI) (auto simp: eventually_sequentially) - -lemma BseqI2': "\n\N. norm (X n) \ K \ Bseq X" - by (intro BfunI) (auto simp: eventually_sequentially) - -lemma Bseq_def: "Bseq X \ (\K>0. \n. norm (X n) \ K)" - unfolding Bfun_def eventually_sequentially -proof safe - fix N K assume "0 < K" "\n\N. norm (X n) \ K" - then show "\K>0. \n. norm (X n) \ K" - by (intro exI[of _ "max (Max (norm ` X ` {..N})) K"] min_max.less_supI2) - (auto intro!: imageI not_less[where 'a=nat, THEN iffD1] Max_ge simp: le_max_iff_disj) -qed auto - -lemma BseqE: "\Bseq X; \K. \0 < K; \n. norm (X n) \ K\ \ Q\ \ Q" -unfolding Bseq_def by auto - -lemma BseqD: "Bseq X ==> \K. 0 < K & (\n. norm (X n) \ K)" -by (simp add: Bseq_def) - -lemma BseqI: "[| 0 < K; \n. norm (X n) \ K |] ==> Bseq X" -by (auto simp add: Bseq_def) - -lemma lemma_NBseq_def: - "(\K > 0. \n. norm (X n) \ K) = (\N. \n. norm (X n) \ real(Suc N))" -proof safe - fix K :: real - from reals_Archimedean2 obtain n :: nat where "K < real n" .. - then have "K \ real (Suc n)" by auto - moreover assume "\m. norm (X m) \ K" - ultimately have "\m. norm (X m) \ real (Suc n)" - by (blast intro: order_trans) - then show "\N. \n. norm (X n) \ real (Suc N)" .. -qed (force simp add: real_of_nat_Suc) - -text{* alternative definition for Bseq *} -lemma Bseq_iff: "Bseq X = (\N. \n. norm (X n) \ real(Suc N))" -apply (simp add: Bseq_def) -apply (simp (no_asm) add: lemma_NBseq_def) -done - -lemma lemma_NBseq_def2: - "(\K > 0. \n. norm (X n) \ K) = (\N. \n. norm (X n) < real(Suc N))" -apply (subst lemma_NBseq_def, auto) -apply (rule_tac x = "Suc N" in exI) -apply (rule_tac [2] x = N in exI) -apply (auto simp add: real_of_nat_Suc) - prefer 2 apply (blast intro: order_less_imp_le) -apply (drule_tac x = n in spec, simp) -done - -(* yet another definition for Bseq *) -lemma Bseq_iff1a: "Bseq X = (\N. \n. norm (X n) < real(Suc N))" -by (simp add: Bseq_def lemma_NBseq_def2) - -subsubsection{*A Few More Equivalence Theorems for Boundedness*} - -text{*alternative formulation for boundedness*} -lemma Bseq_iff2: "Bseq X = (\k > 0. \x. \n. norm (X(n) + -x) \ k)" -apply (unfold Bseq_def, safe) -apply (rule_tac [2] x = "k + norm x" in exI) -apply (rule_tac x = K in exI, simp) -apply (rule exI [where x = 0], auto) -apply (erule order_less_le_trans, simp) -apply (drule_tac x=n in spec, fold diff_minus) -apply (drule order_trans [OF norm_triangle_ineq2]) -apply simp -done - -text{*alternative formulation for boundedness*} -lemma Bseq_iff3: "Bseq X = (\k > 0. \N. \n. norm(X(n) + -X(N)) \ k)" -apply safe -apply (simp add: Bseq_def, safe) -apply (rule_tac x = "K + norm (X N)" in exI) -apply auto -apply (erule order_less_le_trans, simp) -apply (rule_tac x = N in exI, safe) -apply (drule_tac x = n in spec) -apply (rule order_trans [OF norm_triangle_ineq], simp) -apply (auto simp add: Bseq_iff2) -done - -lemma BseqI2: "(\n. k \ f n & f n \ (K::real)) ==> Bseq f" -apply (simp add: Bseq_def) -apply (rule_tac x = " (\k\ + \K\) + 1" in exI, auto) -apply (drule_tac x = n in spec, arith) -done - -subsubsection{*Upper Bounds and Lubs of Bounded Sequences*} - -lemma Bseq_isUb: - "!!(X::nat=>real). Bseq X ==> \U. isUb (UNIV::real set) {x. \n. X n = x} U" -by (auto intro: isUbI setleI simp add: Bseq_def abs_le_iff) - -text{* Use completeness of reals (supremum property) - to show that any bounded sequence has a least upper bound*} - -lemma Bseq_isLub: - "!!(X::nat=>real). Bseq X ==> - \U. isLub (UNIV::real set) {x. \n. X n = x} U" -by (blast intro: reals_complete Bseq_isUb) - -subsubsection{*A Bounded and Monotonic Sequence Converges*} - -(* TODO: delete *) -(* FIXME: one use in NSA/HSEQ.thy *) -lemma Bmonoseq_LIMSEQ: "\n. m \ n --> X n = X m ==> \L. (X ----> L)" - apply (rule_tac x="X m" in exI) - apply (rule filterlim_cong[THEN iffD2, OF refl refl _ tendsto_const]) - unfolding eventually_sequentially - apply blast - done text {* A monotone sequence converges to its least upper bound. *} @@ -1301,9 +1377,6 @@ "Bseq X \ \m. \n \ m. X m \ X n \ convergent (X::nat=>real)" by (metis Bseq_isLub isLub_mono_imp_LIMSEQ convergentI) -lemma Bseq_minus_iff: "Bseq (%n. -(X n) :: 'a :: real_normed_vector) = Bseq X" - by (simp add: Bseq_def) - text{*Main monotonicity theorem*} lemma Bseq_monoseq_convergent: "Bseq X \ monoseq X \ convergent (X::nat\real)" @@ -1325,19 +1398,6 @@ shows "\Cauchy X; 0 < e\ \ \M. \m\M. \n\M. norm (X m - X n) < e" by (simp add: Cauchy_iff) -lemma Bseq_eq_bounded: "range f \ {a .. b::real} \ Bseq f" - apply (simp add: subset_eq) - apply (rule BseqI'[where K="max (norm a) (norm b)"]) - apply (erule_tac x=n in allE) - apply auto - done - -lemma incseq_bounded: "incseq X \ \i. X i \ (B::real) \ Bseq X" - by (intro Bseq_eq_bounded[of X "X 0" B]) (auto simp: incseq_def) - -lemma decseq_bounded: "decseq X \ \i. (B::real) \ X i \ Bseq X" - by (intro Bseq_eq_bounded[of X B "X 0"]) (auto simp: decseq_def) - lemma incseq_convergent: fixes X :: "nat \ real" assumes "incseq X" and "\i. X i \ B" @@ -1375,10 +1435,6 @@ apply simp done -class banach = real_normed_vector + complete_space - -instance real :: banach by default - subsection {* Power Sequences *} text{*The sequence @{term "x^n"} tends to 0 if @{term "0\x"} and @{term @@ -1615,6 +1671,22 @@ subsection {* Uniform Continuity *} +definition + isUCont :: "['a::metric_space \ 'b::metric_space] \ bool" where + "isUCont f = (\r>0. \s>0. \x y. dist x y < s \ dist (f x) (f y) < r)" + +lemma isUCont_isCont: "isUCont f ==> isCont f x" +by (simp add: isUCont_def isCont_def LIM_def, force) + +lemma isUCont_Cauchy: + "\isUCont f; Cauchy X\ \ Cauchy (\n. f (X n))" +unfolding isUCont_def +apply (rule metric_CauchyI) +apply (drule_tac x=e in spec, safe) +apply (drule_tac e=s in metric_CauchyD, safe) +apply (rule_tac x=M in exI, simp) +done + lemma (in bounded_linear) isUCont: "isUCont f" unfolding isUCont_def dist_norm proof (intro allI impI) @@ -1637,7 +1709,6 @@ lemma (in bounded_linear) Cauchy: "Cauchy X \ Cauchy (\n. f (X n))" by (rule isUCont [THEN isUCont_Cauchy]) - lemma LIM_less_bound: fixes f :: "real \ real" assumes ev: "b < x" "\ x' \ { b <..< x}. 0 \ f x'" and "isCont f x" @@ -1878,5 +1949,6 @@ fixes f :: "real \ real" shows "f -- c --> l \ l \ 0 \ \r. 0 < r \ (\x. x \ c \ \c - x\ < r \ f x \ 0)" using LIM_fun_gt_zero[of f l c] LIM_fun_less_zero[of f l c] by (auto simp add: neq_iff) + end diff -r 609914f0934a -r f415febf4234 src/HOL/Metric_Spaces.thy --- a/src/HOL/Metric_Spaces.thy Tue Mar 26 12:21:00 2013 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,629 +0,0 @@ -(* Title: HOL/Metric_Spaces.thy - Author: Brian Huffman - Author: Johannes Hölzl -*) - -header {* Metric Spaces *} - -theory Metric_Spaces -imports Real Topological_Spaces -begin - - -subsection {* Metric spaces *} - -class dist = - fixes dist :: "'a \ 'a \ real" - -class open_dist = "open" + dist + - assumes open_dist: "open S \ (\x\S. \e>0. \y. dist y x < e \ y \ S)" - -class metric_space = open_dist + - assumes dist_eq_0_iff [simp]: "dist x y = 0 \ x = y" - assumes dist_triangle2: "dist x y \ dist x z + dist y z" -begin - -lemma dist_self [simp]: "dist x x = 0" -by simp - -lemma zero_le_dist [simp]: "0 \ dist x y" -using dist_triangle2 [of x x y] by simp - -lemma zero_less_dist_iff: "0 < dist x y \ x \ y" -by (simp add: less_le) - -lemma dist_not_less_zero [simp]: "\ dist x y < 0" -by (simp add: not_less) - -lemma dist_le_zero_iff [simp]: "dist x y \ 0 \ x = y" -by (simp add: le_less) - -lemma dist_commute: "dist x y = dist y x" -proof (rule order_antisym) - show "dist x y \ dist y x" - using dist_triangle2 [of x y x] by simp - show "dist y x \ dist x y" - using dist_triangle2 [of y x y] by simp -qed - -lemma dist_triangle: "dist x z \ dist x y + dist y z" -using dist_triangle2 [of x z y] by (simp add: dist_commute) - -lemma dist_triangle3: "dist x y \ dist a x + dist a y" -using dist_triangle2 [of x y a] by (simp add: dist_commute) - -lemma dist_triangle_alt: - shows "dist y z <= dist x y + dist x z" -by (rule dist_triangle3) - -lemma dist_pos_lt: - shows "x \ y ==> 0 < dist x y" -by (simp add: zero_less_dist_iff) - -lemma dist_nz: - shows "x \ y \ 0 < dist x y" -by (simp add: zero_less_dist_iff) - -lemma dist_triangle_le: - shows "dist x z + dist y z <= e \ dist x y <= e" -by (rule order_trans [OF dist_triangle2]) - -lemma dist_triangle_lt: - shows "dist x z + dist y z < e ==> dist x y < e" -by (rule le_less_trans [OF dist_triangle2]) - -lemma dist_triangle_half_l: - shows "dist x1 y < e / 2 \ dist x2 y < e / 2 \ dist x1 x2 < e" -by (rule dist_triangle_lt [where z=y], simp) - -lemma dist_triangle_half_r: - shows "dist y x1 < e / 2 \ dist y x2 < e / 2 \ dist x1 x2 < e" -by (rule dist_triangle_half_l, simp_all add: dist_commute) - -subclass topological_space -proof - have "\e::real. 0 < e" - by (fast intro: zero_less_one) - then show "open UNIV" - unfolding open_dist by simp -next - fix S T assume "open S" "open T" - then show "open (S \ T)" - unfolding open_dist - apply clarify - apply (drule (1) bspec)+ - apply (clarify, rename_tac r s) - apply (rule_tac x="min r s" in exI, simp) - done -next - fix K assume "\S\K. open S" thus "open (\K)" - unfolding open_dist by fast -qed - -lemma open_ball: "open {y. dist x y < d}" -proof (unfold open_dist, intro ballI) - fix y assume *: "y \ {y. dist x y < d}" - then show "\e>0. \z. dist z y < e \ z \ {y. dist x y < d}" - by (auto intro!: exI[of _ "d - dist x y"] simp: field_simps dist_triangle_lt) -qed - -subclass first_countable_topology -proof - fix x - show "\A::nat \ 'a set. (\i. x \ A i \ open (A i)) \ (\S. open S \ x \ S \ (\i. A i \ S))" - proof (safe intro!: exI[of _ "\n. {y. dist x y < inverse (Suc n)}"]) - fix S assume "open S" "x \ S" - then obtain e where "0 < e" "{y. dist x y < e} \ S" - by (auto simp: open_dist subset_eq dist_commute) - moreover - then obtain i where "inverse (Suc i) < e" - by (auto dest!: reals_Archimedean) - then have "{y. dist x y < inverse (Suc i)} \ {y. dist x y < e}" - by auto - ultimately show "\i. {y. dist x y < inverse (Suc i)} \ S" - by blast - qed (auto intro: open_ball) -qed - -end - -instance metric_space \ t2_space -proof - fix x y :: "'a::metric_space" - assume xy: "x \ y" - let ?U = "{y'. dist x y' < dist x y / 2}" - let ?V = "{x'. dist y x' < 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] - using open_ball[of _ "dist x y / 2"] by auto - then show "\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" - by blast -qed - -lemma eventually_nhds_metric: - fixes a :: "'a :: metric_space" - shows "eventually P (nhds a) \ (\d>0. \x. dist x a < d \ P x)" -unfolding eventually_nhds open_dist -apply safe -apply fast -apply (rule_tac x="{x. dist x a < d}" in exI, simp) -apply clarsimp -apply (rule_tac x="d - dist x a" in exI, clarsimp) -apply (simp only: less_diff_eq) -apply (erule le_less_trans [OF dist_triangle]) -done - -lemma eventually_at: - fixes a :: "'a::metric_space" - shows "eventually P (at a) \ (\d>0. \x. x \ a \ dist x a < d \ P x)" -unfolding at_def eventually_within eventually_nhds_metric by auto - -lemma eventually_within_less: (* COPY FROM Topo/eventually_within *) - fixes a :: "'a :: metric_space" - shows "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: (* COPY FROM Topo/eventually_within_le *) - fixes a :: "'a :: metric_space" - shows "eventually P (at a within S) \ (\d>0. \x\S. 0 < dist x a \ dist x a <= d \ P x)" - unfolding eventually_within_less by auto (metis dense order_le_less_trans) - -lemma tendstoI: - fixes l :: "'a :: metric_space" - assumes "\e. 0 < e \ eventually (\x. dist (f x) l < e) F" - shows "(f ---> l) F" - apply (rule topological_tendstoI) - apply (simp add: open_dist) - apply (drule (1) bspec, clarify) - apply (drule assms) - apply (erule eventually_elim1, simp) - done - -lemma tendstoD: - fixes l :: "'a :: metric_space" - shows "(f ---> l) F \ 0 < e \ eventually (\x. dist (f x) l < e) F" - apply (drule_tac S="{x. dist x l < e}" in topological_tendstoD) - apply (clarsimp simp add: open_dist) - apply (rule_tac x="e - dist x l" in exI, clarsimp) - apply (simp only: less_diff_eq) - apply (erule le_less_trans [OF dist_triangle]) - apply simp - apply simp - done - -lemma tendsto_iff: - fixes l :: "'a :: metric_space" - shows "(f ---> l) F \ (\e>0. eventually (\x. dist (f x) l < e) F)" - using tendstoI tendstoD by fast - -lemma metric_tendsto_imp_tendsto: - fixes a :: "'a :: metric_space" and b :: "'b :: metric_space" - assumes f: "(f ---> a) F" - assumes le: "eventually (\x. dist (g x) b \ dist (f x) a) F" - shows "(g ---> b) F" -proof (rule tendstoI) - fix e :: real assume "0 < e" - with f have "eventually (\x. dist (f x) a < e) F" by (rule tendstoD) - with le show "eventually (\x. dist (g x) b < e) F" - using le_less_trans by (rule eventually_elim2) -qed - -lemma filterlim_real_sequentially: "LIM x sequentially. real x :> at_top" - unfolding filterlim_at_top - apply (intro allI) - apply (rule_tac c="natceiling (Z + 1)" in eventually_sequentiallyI) - apply (auto simp: natceiling_le_eq) - done - -subsubsection {* Limits of Sequences *} - -lemma LIMSEQ_def: "X ----> (L::'a::metric_space) \ (\r>0. \no. \n\no. dist (X n) L < r)" - unfolding tendsto_iff eventually_sequentially .. - -lemma LIMSEQ_iff_nz: "X ----> (L::'a::metric_space) = (\r>0. \no>0. \n\no. dist (X n) L < r)" - unfolding LIMSEQ_def by (metis Suc_leD zero_less_Suc) - -lemma metric_LIMSEQ_I: - "(\r. 0 < r \ \no. \n\no. dist (X n) L < r) \ X ----> (L::'a::metric_space)" -by (simp add: LIMSEQ_def) - -lemma metric_LIMSEQ_D: - "\X ----> (L::'a::metric_space); 0 < r\ \ \no. \n\no. dist (X n) L < r" -by (simp add: LIMSEQ_def) - - -subsubsection {* Limits of Functions *} - -lemma LIM_def: "f -- (a::'a::metric_space) --> (L::'b::metric_space) = - (\r > 0. \s > 0. \x. x \ a & dist x a < s - --> dist (f x) L < r)" -unfolding tendsto_iff eventually_at .. - -lemma metric_LIM_I: - "(\r. 0 < r \ \s>0. \x. x \ a \ dist x a < s \ dist (f x) L < r) - \ f -- (a::'a::metric_space) --> (L::'b::metric_space)" -by (simp add: LIM_def) - -lemma metric_LIM_D: - "\f -- (a::'a::metric_space) --> (L::'b::metric_space); 0 < r\ - \ \s>0. \x. x \ a \ dist x a < s \ dist (f x) L < r" -by (simp add: LIM_def) - -lemma metric_LIM_imp_LIM: - assumes f: "f -- a --> (l::'a::metric_space)" - assumes le: "\x. x \ a \ dist (g x) m \ dist (f x) l" - shows "g -- a --> (m::'b::metric_space)" - by (rule metric_tendsto_imp_tendsto [OF f]) (auto simp add: eventually_at_topological le) - -lemma metric_LIM_equal2: - assumes 1: "0 < R" - assumes 2: "\x. \x \ a; dist x a < R\ \ f x = g x" - shows "g -- a --> l \ f -- (a::'a::metric_space) --> l" -apply (rule topological_tendstoI) -apply (drule (2) topological_tendstoD) -apply (simp add: eventually_at, safe) -apply (rule_tac x="min d R" in exI, safe) -apply (simp add: 1) -apply (simp add: 2) -done - -lemma metric_LIM_compose2: - assumes f: "f -- (a::'a::metric_space) --> b" - assumes g: "g -- b --> c" - assumes inj: "\d>0. \x. x \ a \ dist x a < d \ f x \ b" - shows "(\x. g (f x)) -- a --> c" - using g f inj [folded eventually_at] - by (rule tendsto_compose_eventually) - -lemma metric_isCont_LIM_compose2: - fixes f :: "'a :: metric_space \ _" - assumes f [unfolded isCont_def]: "isCont f a" - assumes g: "g -- f a --> l" - assumes inj: "\d>0. \x. x \ a \ dist x a < d \ f x \ f a" - shows "(\x. g (f x)) -- a --> l" -by (rule metric_LIM_compose2 [OF f g inj]) - -subsubsection {* Boundedness *} - -definition Bfun :: "('a \ 'b::metric_space) \ 'a filter \ bool" where - Bfun_metric_def: "Bfun f F = (\y. \K>0. eventually (\x. dist (f x) y \ K) F)" - -abbreviation Bseq :: "(nat \ 'a::metric_space) \ bool" where - "Bseq X \ Bfun X sequentially" - -lemma Bseq_conv_Bfun: "Bseq X \ Bfun X sequentially" .. - -lemma Bseq_ignore_initial_segment: "Bseq X \ Bseq (\n. X (n + k))" - unfolding Bfun_metric_def by (subst eventually_sequentially_seg) - -lemma Bseq_offset: "Bseq (\n. X (n + k)) \ Bseq X" - unfolding Bfun_metric_def by (subst (asm) eventually_sequentially_seg) - -subsection {* Complete metric spaces *} - -subsection {* Cauchy sequences *} - -definition (in metric_space) Cauchy :: "(nat \ 'a) \ bool" where - "Cauchy X = (\e>0. \M. \m \ M. \n \ M. dist (X m) (X n) < e)" - -subsection {* Cauchy Sequences *} - -lemma metric_CauchyI: - "(\e. 0 < e \ \M. \m\M. \n\M. dist (X m) (X n) < e) \ Cauchy X" - by (simp add: Cauchy_def) - -lemma metric_CauchyD: - "Cauchy X \ 0 < e \ \M. \m\M. \n\M. dist (X m) (X n) < e" - by (simp add: Cauchy_def) - -lemma metric_Cauchy_iff2: - "Cauchy X = (\j. (\M. \m \ M. \n \ M. dist (X m) (X n) < inverse(real (Suc j))))" -apply (simp add: Cauchy_def, auto) -apply (drule reals_Archimedean, safe) -apply (drule_tac x = n in spec, auto) -apply (rule_tac x = M in exI, auto) -apply (drule_tac x = m in spec, simp) -apply (drule_tac x = na in spec, auto) -done - -lemma Cauchy_subseq_Cauchy: - "\ Cauchy X; subseq f \ \ Cauchy (X o f)" -apply (auto simp add: Cauchy_def) -apply (drule_tac x=e in spec, clarify) -apply (rule_tac x=M in exI, clarify) -apply (blast intro: le_trans [OF _ seq_suble] dest!: spec) -done - -lemma Cauchy_Bseq: "Cauchy X \ Bseq X" - unfolding Cauchy_def Bfun_metric_def eventually_sequentially - apply (erule_tac x=1 in allE) - apply simp - apply safe - apply (rule_tac x="X M" in exI) - apply (rule_tac x=1 in exI) - apply (erule_tac x=M in allE) - apply simp - apply (rule_tac x=M in exI) - apply (auto simp: dist_commute) - done - -subsubsection {* Cauchy Sequences are Convergent *} - -class complete_space = metric_space + - assumes Cauchy_convergent: "Cauchy X \ convergent X" - -theorem LIMSEQ_imp_Cauchy: - assumes X: "X ----> a" shows "Cauchy X" -proof (rule metric_CauchyI) - fix e::real assume "0 < e" - hence "0 < e/2" by simp - with X have "\N. \n\N. dist (X n) a < e/2" by (rule metric_LIMSEQ_D) - then obtain N where N: "\n\N. dist (X n) a < e/2" .. - show "\N. \m\N. \n\N. dist (X m) (X n) < e" - proof (intro exI allI impI) - fix m assume "N \ m" - hence m: "dist (X m) a < e/2" using N by fast - fix n assume "N \ n" - hence n: "dist (X n) a < e/2" using N by fast - have "dist (X m) (X n) \ dist (X m) a + dist (X n) a" - by (rule dist_triangle2) - also from m n have "\ < e" by simp - finally show "dist (X m) (X n) < e" . - qed -qed - -lemma convergent_Cauchy: "convergent X \ Cauchy X" -unfolding convergent_def -by (erule exE, erule LIMSEQ_imp_Cauchy) - -lemma Cauchy_convergent_iff: - fixes X :: "nat \ 'a::complete_space" - shows "Cauchy X = convergent X" -by (fast intro: Cauchy_convergent convergent_Cauchy) - -subsection {* Uniform Continuity *} - -definition - isUCont :: "['a::metric_space \ 'b::metric_space] \ bool" where - "isUCont f = (\r>0. \s>0. \x y. dist x y < s \ dist (f x) (f y) < r)" - -lemma isUCont_isCont: "isUCont f ==> isCont f x" -by (simp add: isUCont_def isCont_def LIM_def, force) - -lemma isUCont_Cauchy: - "\isUCont f; Cauchy X\ \ Cauchy (\n. f (X n))" -unfolding isUCont_def -apply (rule metric_CauchyI) -apply (drule_tac x=e in spec, safe) -apply (drule_tac e=s in metric_CauchyD, safe) -apply (rule_tac x=M in exI, simp) -done - -subsection {* The set of real numbers is a complete metric space *} - -instantiation real :: metric_space -begin - -definition dist_real_def: - "dist x y = \x - y\" - -definition open_real_def: - "open (S :: real set) \ (\x\S. \e>0. \y. dist y x < e \ y \ S)" - -instance - by default (auto simp: open_real_def dist_real_def) -end - -instance real :: linorder_topology -proof - show "(open :: real set \ bool) = generate_topology (range lessThan \ range greaterThan)" - proof (rule ext, safe) - fix S :: "real set" assume "open S" - then guess f unfolding open_real_def bchoice_iff .. - then have *: "S = (\x\S. {x - f x <..} \ {..< x + f x})" - by (fastforce simp: dist_real_def) - show "generate_topology (range lessThan \ range greaterThan) S" - apply (subst *) - apply (intro generate_topology_Union generate_topology.Int) - apply (auto intro: generate_topology.Basis) - done - next - fix S :: "real set" assume "generate_topology (range lessThan \ range greaterThan) S" - moreover have "\a::real. open {.. (\y. \y - x\ < a - x \ y \ {..e>0. \y. \y - x\ < e \ y \ {..a::real. open {a <..}" - unfolding open_real_def dist_real_def - proof clarify - fix x a :: real assume "a < x" - hence "0 < x - a \ (\y. \y - x\ < x - a \ y \ {a<..})" by auto - thus "\e>0. \y. \y - x\ < e \ y \ {a<..}" .. - qed - ultimately show "open S" - by induct auto - qed -qed - -lemmas open_real_greaterThan = open_greaterThan[where 'a=real] -lemmas open_real_lessThan = open_lessThan[where 'a=real] -lemmas open_real_greaterThanLessThan = open_greaterThanLessThan[where 'a=real] -lemmas closed_real_atMost = closed_atMost[where 'a=real] -lemmas closed_real_atLeast = closed_atLeast[where 'a=real] -lemmas closed_real_atLeastAtMost = closed_atLeastAtMost[where 'a=real] - -text {* -Proof that Cauchy sequences converge based on the one from -http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html -*} - -text {* - If sequence @{term "X"} is Cauchy, then its limit is the lub of - @{term "{r::real. \N. \n\N. r < X n}"} -*} - -lemma isUb_UNIV_I: "(\y. y \ S \ y \ u) \ isUb UNIV S u" -by (simp add: isUbI setleI) - -lemma increasing_LIMSEQ: - fixes f :: "nat \ real" - assumes inc: "\n. f n \ f (Suc n)" - and bdd: "\n. f n \ l" - and en: "\e. 0 < e \ \n. l \ f n + e" - shows "f ----> l" -proof (rule increasing_tendsto) - fix x assume "x < l" - with dense[of 0 "l - x"] obtain e where "0 < e" "e < l - x" - by auto - from en[OF `0 < e`] obtain n where "l - e \ f n" - by (auto simp: field_simps) - with `e < l - x` `0 < e` have "x < f n" by simp - with incseq_SucI[of f, OF inc] show "eventually (\n. x < f n) sequentially" - by (auto simp: eventually_sequentially incseq_def intro: less_le_trans) -qed (insert bdd, auto) - -lemma real_Cauchy_convergent: - fixes X :: "nat \ real" - assumes X: "Cauchy X" - shows "convergent X" -proof - - def S \ "{x::real. \N. \n\N. x < X n}" - then have mem_S: "\N x. \n\N. x < X n \ x \ S" by auto - - { fix N x assume N: "\n\N. X n < x" - have "isUb UNIV S x" - proof (rule isUb_UNIV_I) - fix y::real assume "y \ S" - hence "\M. \n\M. y < X n" - by (simp add: S_def) - then obtain M where "\n\M. y < X n" .. - hence "y < X (max M N)" by simp - also have "\ < x" using N by simp - finally show "y \ x" - by (rule order_less_imp_le) - qed } - note bound_isUb = this - - have "\u. isLub UNIV S u" - proof (rule reals_complete) - obtain N where "\m\N. \n\N. dist (X m) (X n) < 1" - using X[THEN metric_CauchyD, OF zero_less_one] by auto - hence N: "\n\N. dist (X n) (X N) < 1" by simp - show "\x. x \ S" - proof - from N have "\n\N. X N - 1 < X n" - by (simp add: abs_diff_less_iff dist_real_def) - thus "X N - 1 \ S" by (rule mem_S) - qed - show "\u. isUb UNIV S u" - proof - from N have "\n\N. X n < X N + 1" - by (simp add: abs_diff_less_iff dist_real_def) - thus "isUb UNIV S (X N + 1)" - by (rule bound_isUb) - qed - qed - then obtain x where x: "isLub UNIV S x" .. - have "X ----> x" - proof (rule metric_LIMSEQ_I) - fix r::real assume "0 < r" - hence r: "0 < r/2" by simp - obtain N where "\n\N. \m\N. dist (X n) (X m) < r/2" - using metric_CauchyD [OF X r] by auto - hence "\n\N. dist (X n) (X N) < r/2" by simp - hence N: "\n\N. X N - r/2 < X n \ X n < X N + r/2" - by (simp only: dist_real_def abs_diff_less_iff) - - from N have "\n\N. X N - r/2 < X n" by fast - hence "X N - r/2 \ S" by (rule mem_S) - hence 1: "X N - r/2 \ x" using x isLub_isUb isUbD by fast - - from N have "\n\N. X n < X N + r/2" by fast - hence "isUb UNIV S (X N + r/2)" by (rule bound_isUb) - hence 2: "x \ X N + r/2" using x isLub_le_isUb by fast - - show "\N. \n\N. dist (X n) x < r" - proof (intro exI allI impI) - fix n assume n: "N \ n" - from N n have "X n < X N + r/2" and "X N - r/2 < X n" by simp+ - thus "dist (X n) x < r" using 1 2 - by (simp add: abs_diff_less_iff dist_real_def) - qed - qed - then show ?thesis unfolding convergent_def by auto -qed - -instance real :: complete_space - by intro_classes (rule real_Cauchy_convergent) - -lemma tendsto_dist [tendsto_intros]: - fixes l m :: "'a :: metric_space" - assumes f: "(f ---> l) F" and g: "(g ---> m) F" - shows "((\x. dist (f x) (g x)) ---> dist l m) F" -proof (rule tendstoI) - fix e :: real assume "0 < e" - hence e2: "0 < e/2" by simp - from tendstoD [OF f e2] tendstoD [OF g e2] - show "eventually (\x. dist (dist (f x) (g x)) (dist l m) < e) F" - proof (eventually_elim) - case (elim x) - then show "dist (dist (f x) (g x)) (dist l m) < e" - unfolding dist_real_def - using dist_triangle2 [of "f x" "g x" "l"] - using dist_triangle2 [of "g x" "l" "m"] - using dist_triangle3 [of "l" "m" "f x"] - using dist_triangle [of "f x" "m" "g x"] - by arith - qed -qed - -lemma continuous_dist[continuous_intros]: - fixes f g :: "_ \ 'a :: metric_space" - shows "continuous F f \ continuous F g \ continuous F (\x. dist (f x) (g x))" - unfolding continuous_def by (rule tendsto_dist) - -lemma continuous_on_dist[continuous_on_intros]: - fixes f g :: "_ \ 'a :: metric_space" - shows "continuous_on s f \ continuous_on s g \ continuous_on s (\x. dist (f x) (g x))" - unfolding continuous_on_def by (auto intro: tendsto_dist) - -lemma tendsto_at_topI_sequentially: - fixes f :: "real \ real" - assumes mono: "mono f" - assumes limseq: "(\n. f (real n)) ----> y" - shows "(f ---> y) at_top" -proof (rule tendstoI) - fix e :: real assume "0 < e" - with limseq obtain N :: nat where N: "\n. N \ n \ \f (real n) - y\ < e" - by (auto simp: LIMSEQ_def dist_real_def) - { fix x :: real - from ex_le_of_nat[of x] guess n .. - note monoD[OF mono this] - also have "f (real_of_nat n) \ y" - by (rule LIMSEQ_le_const[OF limseq]) - (auto intro: exI[of _ n] monoD[OF mono] simp: real_eq_of_nat[symmetric]) - finally have "f x \ y" . } - note le = this - have "eventually (\x. real N \ x) at_top" - by (rule eventually_ge_at_top) - then show "eventually (\x. dist (f x) y < e) at_top" - proof eventually_elim - fix x assume N': "real N \ x" - with N[of N] le have "y - f (real N) < e" by auto - moreover note monoD[OF mono N'] - ultimately show "dist (f x) y < e" - using le[of x] by (auto simp: dist_real_def field_simps) - qed -qed - -lemma Cauchy_iff2: - "Cauchy X = (\j. (\M. \m \ M. \n \ M. \X m - X n\ < inverse(real (Suc j))))" - unfolding metric_Cauchy_iff2 dist_real_def .. - -end - diff -r 609914f0934a -r f415febf4234 src/HOL/Real_Vector_Spaces.thy --- a/src/HOL/Real_Vector_Spaces.thy Tue Mar 26 12:21:00 2013 +0100 +++ b/src/HOL/Real_Vector_Spaces.thy Tue Mar 26 12:21:01 2013 +0100 @@ -1,11 +1,12 @@ (* Title: HOL/Real_Vector_Spaces.thy Author: Brian Huffman + Author: Johannes Hölzl *) header {* Vector Spaces and Algebras over the Reals *} theory Real_Vector_Spaces -imports Metric_Spaces +imports Real Topological_Spaces begin subsection {* Locale for additive functions *} @@ -436,6 +437,9 @@ subsection {* Real normed vector spaces *} +class dist = + fixes dist :: "'a \ 'a \ real" + class norm = fixes norm :: "'a \ real" @@ -445,6 +449,9 @@ class dist_norm = dist + norm + minus + assumes dist_norm: "dist x y = norm (x - y)" +class open_dist = "open" + dist + + assumes open_dist: "open S \ (\x\S. \e>0. \y. dist y x < e \ y \ S)" + class real_normed_vector = real_vector + sgn_div_norm + dist_norm + open_dist + assumes norm_eq_zero [simp]: "norm x = 0 \ x = 0" and norm_triangle_ineq: "norm (x + y) \ norm x + norm y" @@ -653,6 +660,133 @@ shows "norm (x ^ n) = norm x ^ n" by (induct n) (simp_all add: norm_mult) + +subsection {* Metric spaces *} + +class metric_space = open_dist + + assumes dist_eq_0_iff [simp]: "dist x y = 0 \ x = y" + assumes dist_triangle2: "dist x y \ dist x z + dist y z" +begin + +lemma dist_self [simp]: "dist x x = 0" +by simp + +lemma zero_le_dist [simp]: "0 \ dist x y" +using dist_triangle2 [of x x y] by simp + +lemma zero_less_dist_iff: "0 < dist x y \ x \ y" +by (simp add: less_le) + +lemma dist_not_less_zero [simp]: "\ dist x y < 0" +by (simp add: not_less) + +lemma dist_le_zero_iff [simp]: "dist x y \ 0 \ x = y" +by (simp add: le_less) + +lemma dist_commute: "dist x y = dist y x" +proof (rule order_antisym) + show "dist x y \ dist y x" + using dist_triangle2 [of x y x] by simp + show "dist y x \ dist x y" + using dist_triangle2 [of y x y] by simp +qed + +lemma dist_triangle: "dist x z \ dist x y + dist y z" +using dist_triangle2 [of x z y] by (simp add: dist_commute) + +lemma dist_triangle3: "dist x y \ dist a x + dist a y" +using dist_triangle2 [of x y a] by (simp add: dist_commute) + +lemma dist_triangle_alt: + shows "dist y z <= dist x y + dist x z" +by (rule dist_triangle3) + +lemma dist_pos_lt: + shows "x \ y ==> 0 < dist x y" +by (simp add: zero_less_dist_iff) + +lemma dist_nz: + shows "x \ y \ 0 < dist x y" +by (simp add: zero_less_dist_iff) + +lemma dist_triangle_le: + shows "dist x z + dist y z <= e \ dist x y <= e" +by (rule order_trans [OF dist_triangle2]) + +lemma dist_triangle_lt: + shows "dist x z + dist y z < e ==> dist x y < e" +by (rule le_less_trans [OF dist_triangle2]) + +lemma dist_triangle_half_l: + shows "dist x1 y < e / 2 \ dist x2 y < e / 2 \ dist x1 x2 < e" +by (rule dist_triangle_lt [where z=y], simp) + +lemma dist_triangle_half_r: + shows "dist y x1 < e / 2 \ dist y x2 < e / 2 \ dist x1 x2 < e" +by (rule dist_triangle_half_l, simp_all add: dist_commute) + +subclass topological_space +proof + have "\e::real. 0 < e" + by (fast intro: zero_less_one) + then show "open UNIV" + unfolding open_dist by simp +next + fix S T assume "open S" "open T" + then show "open (S \ T)" + unfolding open_dist + apply clarify + apply (drule (1) bspec)+ + apply (clarify, rename_tac r s) + apply (rule_tac x="min r s" in exI, simp) + done +next + fix K assume "\S\K. open S" thus "open (\K)" + unfolding open_dist by fast +qed + +lemma open_ball: "open {y. dist x y < d}" +proof (unfold open_dist, intro ballI) + fix y assume *: "y \ {y. dist x y < d}" + then show "\e>0. \z. dist z y < e \ z \ {y. dist x y < d}" + by (auto intro!: exI[of _ "d - dist x y"] simp: field_simps dist_triangle_lt) +qed + +subclass first_countable_topology +proof + fix x + show "\A::nat \ 'a set. (\i. x \ A i \ open (A i)) \ (\S. open S \ x \ S \ (\i. A i \ S))" + proof (safe intro!: exI[of _ "\n. {y. dist x y < inverse (Suc n)}"]) + fix S assume "open S" "x \ S" + then obtain e where "0 < e" "{y. dist x y < e} \ S" + by (auto simp: open_dist subset_eq dist_commute) + moreover + then obtain i where "inverse (Suc i) < e" + by (auto dest!: reals_Archimedean) + then have "{y. dist x y < inverse (Suc i)} \ {y. dist x y < e}" + by auto + ultimately show "\i. {y. dist x y < inverse (Suc i)} \ S" + by blast + qed (auto intro: open_ball) +qed + +end + +instance metric_space \ t2_space +proof + fix x y :: "'a::metric_space" + assume xy: "x \ y" + let ?U = "{y'. dist x y' < dist x y / 2}" + let ?V = "{x'. dist y x' < 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] + using open_ball[of _ "dist x y / 2"] by auto + then show "\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" + by blast +qed + text {* Every normed vector space is a metric space. *} instance real_normed_vector < metric_space @@ -670,12 +804,19 @@ instantiation real :: real_normed_field begin +definition dist_real_def: + "dist x y = \x - y\" + +definition open_real_def: + "open (S :: real set) \ (\x\S. \e>0. \y. dist y x < e \ y \ S)" + definition real_norm_def [simp]: "norm r = \r\" instance apply (intro_classes, unfold real_norm_def real_scaleR_def) apply (rule dist_real_def) +apply (rule open_real_def) apply (simp add: sgn_real_def) apply (rule abs_eq_0) apply (rule abs_triangle_ineq) @@ -685,8 +826,49 @@ end +instance real :: linorder_topology +proof + show "(open :: real set \ bool) = generate_topology (range lessThan \ range greaterThan)" + proof (rule ext, safe) + fix S :: "real set" assume "open S" + then guess f unfolding open_real_def bchoice_iff .. + then have *: "S = (\x\S. {x - f x <..} \ {..< x + f x})" + by (fastforce simp: dist_real_def) + show "generate_topology (range lessThan \ range greaterThan) S" + apply (subst *) + apply (intro generate_topology_Union generate_topology.Int) + apply (auto intro: generate_topology.Basis) + done + next + fix S :: "real set" assume "generate_topology (range lessThan \ range greaterThan) S" + moreover have "\a::real. open {.. (\y. \y - x\ < a - x \ y \ {..e>0. \y. \y - x\ < e \ y \ {..a::real. open {a <..}" + unfolding open_real_def dist_real_def + proof clarify + fix x a :: real assume "a < x" + hence "0 < x - a \ (\y. \y - x\ < x - a \ y \ {a<..})" by auto + thus "\e>0. \y. \y - x\ < e \ y \ {a<..}" .. + qed + ultimately show "open S" + by induct auto + qed +qed + instance real :: linear_continuum_topology .. +lemmas open_real_greaterThan = open_greaterThan[where 'a=real] +lemmas open_real_lessThan = open_lessThan[where 'a=real] +lemmas open_real_greaterThanLessThan = open_greaterThanLessThan[where 'a=real] +lemmas closed_real_atMost = closed_atMost[where 'a=real] +lemmas closed_real_atLeast = closed_atLeast[where 'a=real] +lemmas closed_real_atLeastAtMost = closed_atLeastAtMost[where 'a=real] + subsection {* Extra type constraints *} text {* Only allow @{term "open"} in class @{text topological_space}. *} @@ -919,4 +1101,361 @@ by (clarsimp, rule_tac x="x + of_real (e/2)" in exI, simp) qed +subsection {* Filters and Limits on Metric Space *} + +lemma eventually_nhds_metric: + fixes a :: "'a :: metric_space" + shows "eventually P (nhds a) \ (\d>0. \x. dist x a < d \ P x)" +unfolding eventually_nhds open_dist +apply safe +apply fast +apply (rule_tac x="{x. dist x a < d}" in exI, simp) +apply clarsimp +apply (rule_tac x="d - dist x a" in exI, clarsimp) +apply (simp only: less_diff_eq) +apply (erule le_less_trans [OF dist_triangle]) +done + +lemma eventually_at: + fixes a :: "'a::metric_space" + shows "eventually P (at a) \ (\d>0. \x. x \ a \ dist x a < d \ P x)" +unfolding at_def eventually_within eventually_nhds_metric by auto + +lemma eventually_within_less: + fixes a :: "'a :: metric_space" + shows "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: + fixes a :: "'a :: metric_space" + shows "eventually P (at a within S) \ (\d>0. \x\S. 0 < dist x a \ dist x a \ d \ P x)" + unfolding eventually_within_less by auto (metis dense order_le_less_trans) + +lemma tendstoI: + fixes l :: "'a :: metric_space" + assumes "\e. 0 < e \ eventually (\x. dist (f x) l < e) F" + shows "(f ---> l) F" + apply (rule topological_tendstoI) + apply (simp add: open_dist) + apply (drule (1) bspec, clarify) + apply (drule assms) + apply (erule eventually_elim1, simp) + done + +lemma tendstoD: + fixes l :: "'a :: metric_space" + shows "(f ---> l) F \ 0 < e \ eventually (\x. dist (f x) l < e) F" + apply (drule_tac S="{x. dist x l < e}" in topological_tendstoD) + apply (clarsimp simp add: open_dist) + apply (rule_tac x="e - dist x l" in exI, clarsimp) + apply (simp only: less_diff_eq) + apply (erule le_less_trans [OF dist_triangle]) + apply simp + apply simp + done + +lemma tendsto_iff: + fixes l :: "'a :: metric_space" + shows "(f ---> l) F \ (\e>0. eventually (\x. dist (f x) l < e) F)" + using tendstoI tendstoD by fast + +lemma metric_tendsto_imp_tendsto: + fixes a :: "'a :: metric_space" and b :: "'b :: metric_space" + assumes f: "(f ---> a) F" + assumes le: "eventually (\x. dist (g x) b \ dist (f x) a) F" + shows "(g ---> b) F" +proof (rule tendstoI) + fix e :: real assume "0 < e" + with f have "eventually (\x. dist (f x) a < e) F" by (rule tendstoD) + with le show "eventually (\x. dist (g x) b < e) F" + using le_less_trans by (rule eventually_elim2) +qed + +lemma filterlim_real_sequentially: "LIM x sequentially. real x :> at_top" + unfolding filterlim_at_top + apply (intro allI) + apply (rule_tac c="natceiling (Z + 1)" in eventually_sequentiallyI) + apply (auto simp: natceiling_le_eq) + done + +subsubsection {* Limits of Sequences *} + +lemma LIMSEQ_def: "X ----> (L::'a::metric_space) \ (\r>0. \no. \n\no. dist (X n) L < r)" + unfolding tendsto_iff eventually_sequentially .. + +lemma LIMSEQ_iff_nz: "X ----> (L::'a::metric_space) = (\r>0. \no>0. \n\no. dist (X n) L < r)" + unfolding LIMSEQ_def by (metis Suc_leD zero_less_Suc) + +lemma metric_LIMSEQ_I: + "(\r. 0 < r \ \no. \n\no. dist (X n) L < r) \ X ----> (L::'a::metric_space)" +by (simp add: LIMSEQ_def) + +lemma metric_LIMSEQ_D: + "\X ----> (L::'a::metric_space); 0 < r\ \ \no. \n\no. dist (X n) L < r" +by (simp add: LIMSEQ_def) + + +subsubsection {* Limits of Functions *} + +lemma LIM_def: "f -- (a::'a::metric_space) --> (L::'b::metric_space) = + (\r > 0. \s > 0. \x. x \ a & dist x a < s + --> dist (f x) L < r)" +unfolding tendsto_iff eventually_at .. + +lemma metric_LIM_I: + "(\r. 0 < r \ \s>0. \x. x \ a \ dist x a < s \ dist (f x) L < r) + \ f -- (a::'a::metric_space) --> (L::'b::metric_space)" +by (simp add: LIM_def) + +lemma metric_LIM_D: + "\f -- (a::'a::metric_space) --> (L::'b::metric_space); 0 < r\ + \ \s>0. \x. x \ a \ dist x a < s \ dist (f x) L < r" +by (simp add: LIM_def) + +lemma metric_LIM_imp_LIM: + assumes f: "f -- a --> (l::'a::metric_space)" + assumes le: "\x. x \ a \ dist (g x) m \ dist (f x) l" + shows "g -- a --> (m::'b::metric_space)" + by (rule metric_tendsto_imp_tendsto [OF f]) (auto simp add: eventually_at_topological le) + +lemma metric_LIM_equal2: + assumes 1: "0 < R" + assumes 2: "\x. \x \ a; dist x a < R\ \ f x = g x" + shows "g -- a --> l \ f -- (a::'a::metric_space) --> l" +apply (rule topological_tendstoI) +apply (drule (2) topological_tendstoD) +apply (simp add: eventually_at, safe) +apply (rule_tac x="min d R" in exI, safe) +apply (simp add: 1) +apply (simp add: 2) +done + +lemma metric_LIM_compose2: + assumes f: "f -- (a::'a::metric_space) --> b" + assumes g: "g -- b --> c" + assumes inj: "\d>0. \x. x \ a \ dist x a < d \ f x \ b" + shows "(\x. g (f x)) -- a --> c" + using g f inj [folded eventually_at] + by (rule tendsto_compose_eventually) + +lemma metric_isCont_LIM_compose2: + fixes f :: "'a :: metric_space \ _" + assumes f [unfolded isCont_def]: "isCont f a" + assumes g: "g -- f a --> l" + assumes inj: "\d>0. \x. x \ a \ dist x a < d \ f x \ f a" + shows "(\x. g (f x)) -- a --> l" +by (rule metric_LIM_compose2 [OF f g inj]) + +subsection {* Complete metric spaces *} + +subsection {* Cauchy sequences *} + +definition (in metric_space) Cauchy :: "(nat \ 'a) \ bool" where + "Cauchy X = (\e>0. \M. \m \ M. \n \ M. dist (X m) (X n) < e)" + +subsection {* Cauchy Sequences *} + +lemma metric_CauchyI: + "(\e. 0 < e \ \M. \m\M. \n\M. dist (X m) (X n) < e) \ Cauchy X" + by (simp add: Cauchy_def) + +lemma metric_CauchyD: + "Cauchy X \ 0 < e \ \M. \m\M. \n\M. dist (X m) (X n) < e" + by (simp add: Cauchy_def) + +lemma metric_Cauchy_iff2: + "Cauchy X = (\j. (\M. \m \ M. \n \ M. dist (X m) (X n) < inverse(real (Suc j))))" +apply (simp add: Cauchy_def, auto) +apply (drule reals_Archimedean, safe) +apply (drule_tac x = n in spec, auto) +apply (rule_tac x = M in exI, auto) +apply (drule_tac x = m in spec, simp) +apply (drule_tac x = na in spec, auto) +done + +lemma Cauchy_iff2: + "Cauchy X = (\j. (\M. \m \ M. \n \ M. \X m - X n\ < inverse(real (Suc j))))" + unfolding metric_Cauchy_iff2 dist_real_def .. + +lemma Cauchy_subseq_Cauchy: + "\ Cauchy X; subseq f \ \ Cauchy (X o f)" +apply (auto simp add: Cauchy_def) +apply (drule_tac x=e in spec, clarify) +apply (rule_tac x=M in exI, clarify) +apply (blast intro: le_trans [OF _ seq_suble] dest!: spec) +done + +theorem LIMSEQ_imp_Cauchy: + assumes X: "X ----> a" shows "Cauchy X" +proof (rule metric_CauchyI) + fix e::real assume "0 < e" + hence "0 < e/2" by simp + with X have "\N. \n\N. dist (X n) a < e/2" by (rule metric_LIMSEQ_D) + then obtain N where N: "\n\N. dist (X n) a < e/2" .. + show "\N. \m\N. \n\N. dist (X m) (X n) < e" + proof (intro exI allI impI) + fix m assume "N \ m" + hence m: "dist (X m) a < e/2" using N by fast + fix n assume "N \ n" + hence n: "dist (X n) a < e/2" using N by fast + have "dist (X m) (X n) \ dist (X m) a + dist (X n) a" + by (rule dist_triangle2) + also from m n have "\ < e" by simp + finally show "dist (X m) (X n) < e" . + qed +qed + +lemma convergent_Cauchy: "convergent X \ Cauchy X" +unfolding convergent_def +by (erule exE, erule LIMSEQ_imp_Cauchy) + +subsubsection {* Cauchy Sequences are Convergent *} + +class complete_space = metric_space + + assumes Cauchy_convergent: "Cauchy X \ convergent X" + +lemma Cauchy_convergent_iff: + fixes X :: "nat \ 'a::complete_space" + shows "Cauchy X = convergent X" +by (fast intro: Cauchy_convergent convergent_Cauchy) + +subsection {* The set of real numbers is a complete metric space *} + +text {* +Proof that Cauchy sequences converge based on the one from +http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html +*} + +text {* + If sequence @{term "X"} is Cauchy, then its limit is the lub of + @{term "{r::real. \N. \n\N. r < X n}"} +*} + +lemma isUb_UNIV_I: "(\y. y \ S \ y \ u) \ isUb UNIV S u" +by (simp add: isUbI setleI) + +lemma increasing_LIMSEQ: + fixes f :: "nat \ real" + assumes inc: "\n. f n \ f (Suc n)" + and bdd: "\n. f n \ l" + and en: "\e. 0 < e \ \n. l \ f n + e" + shows "f ----> l" +proof (rule increasing_tendsto) + fix x assume "x < l" + with dense[of 0 "l - x"] obtain e where "0 < e" "e < l - x" + by auto + from en[OF `0 < e`] obtain n where "l - e \ f n" + by (auto simp: field_simps) + with `e < l - x` `0 < e` have "x < f n" by simp + with incseq_SucI[of f, OF inc] show "eventually (\n. x < f n) sequentially" + by (auto simp: eventually_sequentially incseq_def intro: less_le_trans) +qed (insert bdd, auto) + +lemma real_Cauchy_convergent: + fixes X :: "nat \ real" + assumes X: "Cauchy X" + shows "convergent X" +proof - + def S \ "{x::real. \N. \n\N. x < X n}" + then have mem_S: "\N x. \n\N. x < X n \ x \ S" by auto + + { fix N x assume N: "\n\N. X n < x" + have "isUb UNIV S x" + proof (rule isUb_UNIV_I) + fix y::real assume "y \ S" + hence "\M. \n\M. y < X n" + by (simp add: S_def) + then obtain M where "\n\M. y < X n" .. + hence "y < X (max M N)" by simp + also have "\ < x" using N by simp + finally show "y \ x" + by (rule order_less_imp_le) + qed } + note bound_isUb = this + + have "\u. isLub UNIV S u" + proof (rule reals_complete) + obtain N where "\m\N. \n\N. dist (X m) (X n) < 1" + using X[THEN metric_CauchyD, OF zero_less_one] by auto + hence N: "\n\N. dist (X n) (X N) < 1" by simp + show "\x. x \ S" + proof + from N have "\n\N. X N - 1 < X n" + by (simp add: abs_diff_less_iff dist_real_def) + thus "X N - 1 \ S" by (rule mem_S) + qed + show "\u. isUb UNIV S u" + proof + from N have "\n\N. X n < X N + 1" + by (simp add: abs_diff_less_iff dist_real_def) + thus "isUb UNIV S (X N + 1)" + by (rule bound_isUb) + qed + qed + then obtain x where x: "isLub UNIV S x" .. + have "X ----> x" + proof (rule metric_LIMSEQ_I) + fix r::real assume "0 < r" + hence r: "0 < r/2" by simp + obtain N where "\n\N. \m\N. dist (X n) (X m) < r/2" + using metric_CauchyD [OF X r] by auto + hence "\n\N. dist (X n) (X N) < r/2" by simp + hence N: "\n\N. X N - r/2 < X n \ X n < X N + r/2" + by (simp only: dist_real_def abs_diff_less_iff) + + from N have "\n\N. X N - r/2 < X n" by fast + hence "X N - r/2 \ S" by (rule mem_S) + hence 1: "X N - r/2 \ x" using x isLub_isUb isUbD by fast + + from N have "\n\N. X n < X N + r/2" by fast + hence "isUb UNIV S (X N + r/2)" by (rule bound_isUb) + hence 2: "x \ X N + r/2" using x isLub_le_isUb by fast + + show "\N. \n\N. dist (X n) x < r" + proof (intro exI allI impI) + fix n assume n: "N \ n" + from N n have "X n < X N + r/2" and "X N - r/2 < X n" by simp+ + thus "dist (X n) x < r" using 1 2 + by (simp add: abs_diff_less_iff dist_real_def) + qed + qed + then show ?thesis unfolding convergent_def by auto +qed + +instance real :: complete_space + by intro_classes (rule real_Cauchy_convergent) + +class banach = real_normed_vector + complete_space + +instance real :: banach by default + +lemma tendsto_at_topI_sequentially: + fixes f :: "real \ real" + assumes mono: "mono f" + assumes limseq: "(\n. f (real n)) ----> y" + shows "(f ---> y) at_top" +proof (rule tendstoI) + fix e :: real assume "0 < e" + with limseq obtain N :: nat where N: "\n. N \ n \ \f (real n) - y\ < e" + by (auto simp: LIMSEQ_def dist_real_def) + { fix x :: real + from ex_le_of_nat[of x] guess n .. + note monoD[OF mono this] + also have "f (real_of_nat n) \ y" + by (rule LIMSEQ_le_const[OF limseq]) + (auto intro: exI[of _ n] monoD[OF mono] simp: real_eq_of_nat[symmetric]) + finally have "f x \ y" . } + note le = this + have "eventually (\x. real N \ x) at_top" + by (rule eventually_ge_at_top) + then show "eventually (\x. dist (f x) y < e) at_top" + proof eventually_elim + fix x assume N': "real N \ x" + with N[of N] le have "y - f (real N) < e" by auto + moreover note monoD[OF mono N'] + ultimately show "dist (f x) y < e" + using le[of x] by (auto simp: dist_real_def field_simps) + qed +qed + end