diff -r 49996715bc6e -r 6a6d8004322f src/HOL/Hyperreal/Lim.thy --- a/src/HOL/Hyperreal/Lim.thy Sun Sep 17 16:42:38 2006 +0200 +++ b/src/HOL/Hyperreal/Lim.thy Sun Sep 17 16:44:05 2006 +0200 @@ -15,21 +15,21 @@ text{*Standard and Nonstandard Definitions*} definition - LIM :: "[real => 'a::real_normed_vector, real, 'a] => bool" + LIM :: "['a::real_normed_vector => 'b::real_normed_vector, 'a, 'b] => bool" ("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60) "f -- a --> L = - (\r > 0. \s > 0. \x. x \ a & \x + -a\ < s + (\r > 0. \s > 0. \x. x \ a & norm (x + -a) < s --> norm (f x + -L) < r)" - NSLIM :: "[real => 'a::real_normed_vector, real, 'a] => bool" + NSLIM :: "['a::real_normed_vector => 'b::real_normed_vector, 'a, 'b] => bool" ("((_)/ -- (_)/ --NS> (_))" [60, 0, 60] 60) "f -- a --NS> L = (\x. (x \ star_of a & x @= star_of a --> ( *f* f) x @= star_of L))" - isCont :: "[real => 'a::real_normed_vector, real] => bool" + isCont :: "['a::real_normed_vector => 'b::real_normed_vector, 'a] => bool" "isCont f a = (f -- a --> (f a))" - isNSCont :: "[real => 'a::real_normed_vector, real] => bool" + isNSCont :: "['a::real_normed_vector => 'b::real_normed_vector, 'a] => bool" --{*NS definition dispenses with limit notions*} "isNSCont f a = (\y. y @= star_of a --> ( *f* f) y @= star_of (f a))" @@ -78,23 +78,24 @@ lemma LIM_eq: "f -- a --> L = - (\r>0.\s>0.\x. x \ a & \x-a\ < s --> norm (f x - L) < r)" + (\r>0.\s>0.\x. x \ a & norm (x-a) < s --> norm (f x - L) < r)" by (simp add: LIM_def diff_minus) lemma LIM_I: - "(!!r. 0 \s>0.\x. x \ a & \x-a\ < s --> norm (f x - L) < r) + "(!!r. 0 \s>0.\x. x \ a & norm (x-a) < s --> norm (f x - L) < r) ==> f -- a --> L" by (simp add: LIM_eq) lemma LIM_D: "[| f -- a --> L; 0 \s>0.\x. x \ a & \x-a\ < s --> norm (f x - L) < r" + ==> \s>0.\x. x \ a & norm (x-a) < s --> norm (f x - L) < r" by (simp add: LIM_eq) lemma LIM_const [simp]: "(%x. k) -- x --> k" by (simp add: LIM_def) lemma LIM_add: + fixes f g :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes f: "f -- a --> L" and g: "g -- a --> M" shows "(%x. f x + g(x)) -- a --> (L + M)" proof (rule LIM_I) @@ -103,19 +104,19 @@ from LIM_D [OF f half_gt_zero [OF r]] obtain fs where fs: "0 < fs" - and fs_lt: "\x. x \ a & \x-a\ < fs --> norm (f x - L) < r/2" + and fs_lt: "\x. x \ a & norm (x-a) < fs --> norm (f x - L) < r/2" by blast from LIM_D [OF g half_gt_zero [OF r]] obtain gs where gs: "0 < gs" - and gs_lt: "\x. x \ a & \x-a\ < gs --> norm (g x - M) < r/2" + and gs_lt: "\x. x \ a & norm (x-a) < gs --> norm (g x - M) < r/2" by blast - show "\s>0.\x. x \ a \ \x-a\ < s \ norm (f x + g x - (L + M)) < r" + show "\s>0.\x. x \ a \ norm (x-a) < s \ norm (f x + g x - (L + M)) < r" proof (intro exI conjI strip) show "0 < min fs gs" by (simp add: fs gs) - fix x :: real - assume "x \ a \ \x-a\ < min fs gs" - hence "x \ a \ \x-a\ < fs \ \x-a\ < gs" by simp + fix x :: 'a + assume "x \ a \ norm (x-a) < min fs gs" + hence "x \ a \ norm (x-a) < fs \ norm (x-a) < gs" by simp with fs_lt gs_lt have "norm (f x - L) < r/2" and "norm (g x - M) < r/2" by blast+ hence "norm (f x - L) + norm (g x - M) < r" by arith @@ -141,24 +142,30 @@ by (simp only: diff_minus LIM_add LIM_minus) -lemma LIM_const_not_eq: "k \ L ==> ~ ((%x. k) -- a --> L)" +lemma LIM_const_not_eq: + fixes a :: "'a::real_normed_div_algebra" + shows "k \ L ==> ~ ((%x. k) -- a --> L)" apply (simp add: LIM_eq) apply (rule_tac x="norm (k - L)" in exI, simp, safe) -apply (rule_tac x="a + s/2" in exI, simp) +apply (rule_tac x="a + of_real (s/2)" in exI, simp add: norm_of_real) done -lemma LIM_const_eq: "(%x. k) -- x --> L ==> k = L" +lemma LIM_const_eq: + fixes a :: "'a::real_normed_div_algebra" + shows "(%x. k) -- a --> L ==> k = L" apply (rule ccontr) apply (blast dest: LIM_const_not_eq) done -lemma LIM_unique: "[| f -- a --> L; f -- a --> M |] ==> L = M" +lemma LIM_unique: + fixes a :: "'a::real_normed_div_algebra" + shows "[| f -- a --> L; f -- a --> M |] ==> L = M" apply (drule LIM_diff, assumption) apply (auto dest!: LIM_const_eq) done lemma LIM_mult_zero: - fixes f g :: "real \ 'a::real_normed_algebra" + fixes f g :: "'a::real_normed_vector \ 'b::real_normed_algebra" assumes f: "f -- a --> 0" and g: "g -- a --> 0" shows "(%x. f(x) * g(x)) -- a --> 0" proof (rule LIM_I, simp) @@ -167,19 +174,19 @@ from LIM_D [OF f zero_less_one] obtain fs where fs: "0 < fs" - and fs_lt: "\x. x \ a & \x-a\ < fs --> norm (f x) < 1" + and fs_lt: "\x. x \ a & norm (x-a) < fs --> norm (f x) < 1" by auto from LIM_D [OF g r] obtain gs where gs: "0 < gs" - and gs_lt: "\x. x \ a & \x-a\ < gs --> norm (g x) < r" + and gs_lt: "\x. x \ a & norm (x-a) < gs --> norm (g x) < r" by auto - show "\s. 0 < s \ (\x. x \ a \ \x-a\ < s \ norm (f x * g x) < r)" + show "\s. 0 < s \ (\x. x \ a \ norm (x-a) < s \ norm (f x * g x) < r)" proof (intro exI conjI strip) show "0 < min fs gs" by (simp add: fs gs) - fix x :: real - assume "x \ a \ \x-a\ < min fs gs" - hence "x \ a \ \x-a\ < fs \ \x-a\ < gs" by simp + fix x :: 'a + assume "x \ a \ norm (x-a) < min fs gs" + hence "x \ a \ norm (x-a) < fs \ norm (x-a) < gs" by simp with fs_lt gs_lt have "norm (f x) < 1" and "norm (g x) < r" by blast+ hence "norm (f x) * norm (g x) < 1*r" @@ -226,18 +233,18 @@ lemma lemma_LIM: fixes L :: "'a::real_normed_vector" - shows "\s>0. \x. x \ a \ \x + - a\ < s \ \ norm (f x + -L) < r + shows "\s>0. \x. x \ a \ norm (x + - a) < s \ \ norm (f x + -L) < r ==> \n::nat. \x. x \ a \ - \x + -a\ < inverse(real(Suc n)) \ \ norm (f x + -L) < r" + norm (x + -a) < inverse(real(Suc n)) \ \ norm (f x + -L) < r" apply clarify apply (cut_tac n1 = n in real_of_nat_Suc_gt_zero [THEN positive_imp_inverse_positive], auto) done lemma lemma_skolemize_LIM2: fixes L :: "'a::real_normed_vector" - shows "\s>0. \x. x \ a \ \x + - a\ < s \ \ norm (f x + -L) < r + shows "\s>0. \x. x \ a \ norm (x + - a) < s \ \ norm (f x + -L) < r ==> \X. \n::nat. X n \ a \ - \X n + -a\ < inverse(real(Suc n)) \ \ norm(f (X n) + -L) < r" + norm (X n + -a) < inverse(real(Suc n)) \ \ norm(f (X n) + -L) < r" apply (drule lemma_LIM) apply (drule choice, blast) done @@ -252,9 +259,9 @@ done lemma lemma_simp: "\n. X n \ a & - \X n + - a\ < inverse (real(Suc n)) & + norm (X n + - a) < inverse (real(Suc n)) & \ norm (f (X n) + - L) < r ==> - \n. \X n + - a\ < inverse (real(Suc n))" + \n. norm (X n + - a) < inverse (real(Suc n))" by auto @@ -359,35 +366,45 @@ apply (auto simp add: diff_minus add_assoc) done -lemma NSLIM_const_not_eq: "k \ L ==> ~ ((%x. k) -- x --NS> L)" +lemma NSLIM_const_not_eq: + fixes a :: real (* TODO: generalize to real_normed_div_algebra *) + shows "k \ L ==> ~ ((%x. k) -- a --NS> L)" apply (simp add: NSLIM_def) -apply (rule_tac x = "hypreal_of_real x + epsilon" in exI) +apply (rule_tac x="star_of a + epsilon" in exI) apply (auto intro: Infinitesimal_add_approx_self [THEN approx_sym] simp add: hypreal_epsilon_not_zero) done -lemma NSLIM_not_zero: "k \ 0 ==> ~ ((%x. k) -- x --NS> 0)" +lemma NSLIM_not_zero: + fixes a :: real + shows "k \ 0 ==> ~ ((%x. k) -- a --NS> 0)" by (rule NSLIM_const_not_eq) -lemma NSLIM_const_eq: "(%x. k) -- x --NS> L ==> k = L" +lemma NSLIM_const_eq: + fixes a :: real + shows "(%x. k) -- a --NS> L ==> k = L" apply (rule ccontr) apply (blast dest: NSLIM_const_not_eq) done text{* can actually be proved more easily by unfolding the definition!*} -lemma NSLIM_unique: "[| f -- x --NS> L; f -- x --NS> M |] ==> L = M" +lemma NSLIM_unique: + fixes a :: real + shows "[| f -- a --NS> L; f -- a --NS> M |] ==> L = M" apply (drule NSLIM_minus) apply (drule NSLIM_add, assumption) apply (auto dest!: NSLIM_const_eq [symmetric]) apply (simp add: diff_def [symmetric]) done -lemma LIM_unique2: "[| f -- x --> L; f -- x --> M |] ==> L = M" +lemma LIM_unique2: + fixes a :: real + shows "[| f -- a --> L; f -- a --> M |] ==> L = M" by (simp add: LIM_NSLIM_iff NSLIM_unique) lemma NSLIM_mult_zero: - fixes f g :: "real \ 'a::real_normed_algebra" + fixes f g :: "'a::real_normed_vector \ 'b::real_normed_algebra" shows "[| f -- x --NS> 0; g -- x --NS> 0 |] ==> (%x. f(x)*g(x)) -- x --NS> 0" by (drule NSLIM_mult, auto) @@ -395,7 +412,7 @@ (* for standard definition of limit *) lemma LIM_mult_zero2: - fixes f g :: "real \ 'a::real_normed_algebra" + fixes f g :: "'a::real_normed_vector \ 'b::real_normed_algebra" shows "[| f -- x --> 0; g -- x --> 0 |] ==> (%x. f(x)*g(x)) -- x --> 0" by (drule LIM_mult2, auto) @@ -416,7 +433,7 @@ lemma NSLIM_isNSCont: "f -- a --NS> (f a) ==> isNSCont f a" apply (simp add: isNSCont_def NSLIM_def, auto) -apply (rule_tac Q = "y = hypreal_of_real a" in excluded_middle [THEN disjE], auto) +apply (case_tac "y = star_of a", auto) done text{*NS continuity can be defined using NS Limit in @@ -449,13 +466,13 @@ (* seems easier than using standard def *) lemma NSLIM_h_iff: "(f -- a --NS> L) = ((%h. f(a + h)) -- 0 --NS> L)" apply (simp add: NSLIM_def, auto) -apply (drule_tac x = "hypreal_of_real a + x" in spec) -apply (drule_tac [2] x = "-hypreal_of_real a + x" in spec, safe, simp) -apply (rule mem_infmal_iff [THEN iffD2, THEN Infinitesimal_add_approx_self [THEN approx_sym]]) -apply (rule_tac [4] approx_minus_iff2 [THEN iffD1]) - prefer 3 apply (simp add: add_commute) +apply (drule_tac x = "star_of a + x" in spec) +apply (drule_tac [2] x = "- star_of a + x" in spec, safe, simp) +apply (erule mem_infmal_iff [THEN iffD2, THEN Infinitesimal_add_approx_self [THEN approx_sym]]) +apply (erule_tac [3] approx_minus_iff2 [THEN iffD1]) + prefer 2 apply (simp add: add_commute diff_def [symmetric]) +apply (rule_tac x = x in star_cases) apply (rule_tac [2] x = x in star_cases) -apply (rule_tac [4] x = x in star_cases) apply (auto simp add: starfun star_of_def star_n_minus star_n_add add_assoc approx_refl star_n_zero_num) done @@ -476,7 +493,7 @@ text{*mult continuous*} lemma isCont_mult: - fixes f g :: "real \ 'a::real_normed_algebra" + fixes f g :: "'a::real_normed_vector \ 'b::real_normed_algebra" shows "[| isCont f a; isCont g a |] ==> isCont (%x. f(x) * g(x)) a" by (auto intro!: starfun_mult_HFinite_approx simp del: starfun_mult [symmetric] @@ -497,14 +514,16 @@ by (auto simp add: isNSCont_isCont_iff [symmetric] isNSCont_minus) lemma isCont_inverse: - fixes f :: "real \ 'a::{real_normed_div_algebra,division_by_zero}" + fixes f :: "'a::real_normed_vector \ + 'b::{real_normed_div_algebra,division_by_zero}" shows "[| isCont f x; f x \ 0 |] ==> isCont (%x. inverse (f x)) x" apply (simp add: isCont_def) apply (blast intro: LIM_inverse) done lemma isNSCont_inverse: - fixes f :: "real \ 'a::{real_normed_div_algebra,division_by_zero}" + fixes f :: "'a::real_normed_vector \ + 'b::{real_normed_div_algebra,division_by_zero}" shows "[| isNSCont f x; f x \ 0 |] ==> isNSCont (%x. inverse (f x)) x" by (auto intro: isCont_inverse simp add: isNSCont_isCont_iff) @@ -520,12 +539,12 @@ lemma isNSCont_const [simp]: "isNSCont (%x. k) a" by (simp add: isNSCont_def) -lemma isNSCont_abs [simp]: "isNSCont abs a" +lemma isNSCont_abs [simp]: "isNSCont abs (a::real)" apply (simp add: isNSCont_def) apply (auto intro: approx_hrabs simp add: hypreal_of_real_hrabs [symmetric] starfun_rabs_hrabs) done -lemma isCont_abs [simp]: "isCont abs a" +lemma isCont_abs [simp]: "isCont abs (a::real)" by (auto simp add: isNSCont_isCont_iff [symmetric]) @@ -684,7 +703,7 @@ subsubsection{*Alternative definition for differentiability*} lemma DERIV_LIM_iff: - "((%h. (f(a + h) - f(a)) / h) -- 0 --> D) = + "((%h::real. (f(a + h) - f(a)) / h) -- 0 --> D) = ((%x. (f(x)-f(a)) / (x-a)) -- a --> D)" proof (intro iffI LIM_I) fix r::real @@ -696,12 +715,12 @@ and s_lt: "\x. x \ 0 & \x\ < s --> \(f (a + x) - f a) / x - D\ < r" by auto show "\s. 0 < s \ - (\x. x \ a \ \x-a\ < s \ norm ((f x - f a) / (x-a) - D) < r)" + (\x. x \ a \ norm (x-a) < s \ norm ((f x - f a) / (x-a) - D) < r)" proof (intro exI conjI strip) show "0 < s" by (rule s) next fix x::real - assume "x \ a \ \x-a\ < s" + assume "x \ a \ norm (x-a) < s" with s_lt [THEN spec [where x="x-a"]] show "norm ((f x - f a) / (x-a) - D) < r" by auto qed @@ -715,12 +734,12 @@ and s_lt: "\x. x \ a & \x-a\ < s --> \(f x - f a)/(x-a) - D\ < r" by auto show "\s. 0 < s \ - (\x. x \ 0 & \x - 0\ < s --> norm ((f (a + x) - f a) / x - D) < r)" + (\x. x \ 0 & norm (x - 0) < s --> norm ((f (a + x) - f a) / x - D) < r)" proof (intro exI conjI strip) show "0 < s" by (rule s) next fix x::real - assume "x \ 0 \ \x - 0\ < s" + assume "x \ 0 \ norm (x - 0) < s" with s_lt [THEN spec [where x="x+a"]] show "norm ((f (a + x) - f a) / x - D) < r" by (auto simp add: add_ac) qed @@ -1415,7 +1434,7 @@ subsection{*Intermediate Value Theorem: Prove Contrapositive by Bisection*} -lemma IVT: "[| f(a) \ (y::real); y \ f(b); +lemma IVT: "[| f(a::real) \ (y::real); y \ f(b); a \ b; (\x. a \ x & x \ b --> isCont f x) |] ==> \x. a \ x & x \ b & f(x) = y" @@ -1442,7 +1461,7 @@ apply (case_tac "x \ aa", simp_all) done -lemma IVT2: "[| f(b) \ (y::real); y \ f(a); +lemma IVT2: "[| f(b::real) \ (y::real); y \ f(a); a \ b; (\x. a \ x & x \ b --> isCont f x) |] ==> \x. a \ x & x \ b & f(x) = y" @@ -1452,13 +1471,13 @@ done (*HOL style here: object-level formulations*) -lemma IVT_objl: "(f(a) \ (y::real) & y \ f(b) & a \ b & +lemma IVT_objl: "(f(a::real) \ (y::real) & y \ f(b) & a \ b & (\x. a \ x & x \ b --> isCont f x)) --> (\x. a \ x & x \ b & f(x) = y)" apply (blast intro: IVT) done -lemma IVT2_objl: "(f(b) \ (y::real) & y \ f(a) & a \ b & +lemma IVT2_objl: "(f(b::real) \ (y::real) & y \ f(a) & a \ b & (\x. a \ x & x \ b --> isCont f x)) --> (\x. a \ x & x \ b & f(x) = y)" apply (blast intro: IVT2) @@ -1468,7 +1487,7 @@ lemma isCont_bounded: "[| a \ b; \x. a \ x & x \ b --> isCont f x |] - ==> \M::real. \x. a \ x & x \ b --> f(x) \ M" + ==> \M::real. \x::real. a \ x & x \ b --> f(x) \ M" apply (cut_tac P = "% (u,v) . a \ u & u \ v & v \ b --> (\M. \x. u \ x & x \ v --> f x \ M)" in lemma_BOLZANO2) apply safe apply simp_all @@ -1498,7 +1517,7 @@ by (blast intro: reals_complete) lemma isCont_has_Ub: "[| a \ b; \x. a \ x & x \ b --> isCont f x |] - ==> \M::real. (\x. a \ x & x \ b --> f(x) \ M) & + ==> \M::real. (\x::real. a \ x & x \ b --> f(x) \ M) & (\N. N < M --> (\x. a \ x & x \ b & N < f(x)))" apply (cut_tac S = "Collect (%y. \x. a \ x & x \ b & y = f x)" in lemma_reals_complete) @@ -1513,7 +1532,7 @@ lemma isCont_eq_Ub: assumes le: "a \ b" - and con: "\x. a \ x & x \ b --> isCont f x" + and con: "\x::real. a \ x & x \ b --> isCont f x" shows "\M::real. (\x. a \ x & x \ b --> f(x) \ M) & (\x. a \ x & x \ b & f(x) = M)" proof - @@ -1554,7 +1573,7 @@ text{*Same theorem for lower bound*} lemma isCont_eq_Lb: "[| a \ b; \x. a \ x & x \ b --> isCont f x |] - ==> \M::real. (\x. a \ x & x \ b --> M \ f(x)) & + ==> \M::real. (\x::real. a \ x & x \ b --> M \ f(x)) & (\x. a \ x & x \ b & f(x) = M)" apply (subgoal_tac "\x. a \ x & x \ b --> isCont (%x. - (f x)) x") prefer 2 apply (blast intro: isCont_minus) @@ -1567,7 +1586,7 @@ text{*Another version.*} lemma isCont_Lb_Ub: "[|a \ b; \x. a \ x & x \ b --> isCont f x |] - ==> \L M::real. (\x. a \ x & x \ b --> L \ f(x) & f(x) \ M) & + ==> \L M::real. (\x::real. a \ x & x \ b --> L \ f(x) & f(x) \ M) & (\y. L \ y & y \ M --> (\x. a \ x & x \ b & (f(x) = y)))" apply (frule isCont_eq_Lb) apply (frule_tac [2] isCont_eq_Ub) @@ -2032,6 +2051,7 @@ text{*Continuity of inverse function*} lemma isCont_inverse_function: + fixes f g :: "real \ real" assumes d: "0 < d" and inj: "\z. \z-x\ \ d --> g(f z) = z" and cont: "\z. \z-x\ \ d --> isCont f z" @@ -2205,11 +2225,12 @@ lemma LIMSEQ_SEQ_conv1: + fixes a :: real assumes "X -- a --> L" shows "\S. (\n. S n \ a) \ S ----> a \ (\n. X (S n)) ----> L" proof - { - from prems have Xdef: "\r > 0. \s > 0. \x. x \ a & \x + -a\ < s --> norm (X x + -L) < r" by (unfold LIM_def) + from prems have Xdef: "\r > 0. \s > 0. \x. x \ a & norm (x + -a) < s --> norm (X x + -L) < r" by (unfold LIM_def) fix S assume as: "(\n. S n \ a) \ S ----> a" @@ -2246,11 +2267,12 @@ ML {* fast_arith_split_limit := 0; *} (* FIXME *) lemma LIMSEQ_SEQ_conv2: + fixes a :: real assumes "\S. (\n. S n \ a) \ S ----> a \ (\n. X (S n)) ----> L" shows "X -- a --> L" proof (rule ccontr) assume "\ (X -- a --> L)" - hence "\ (\r > 0. \s > 0. \x. x \ a & \x + -a\ < s --> norm (X x + -L) < r)" by (unfold LIM_def) + hence "\ (\r > 0. \s > 0. \x. x \ a & norm (x + -a) < s --> norm (X x + -L) < r)" by (unfold LIM_def) hence "\r > 0. \s > 0. \x. \(x \ a \ \x + -a\ < s --> norm (X x + -L) < r)" by simp hence "\r > 0. \s > 0. \x. (x \ a \ \x + -a\ < s \ norm (X x + -L) \ r)" by (simp add: linorder_not_less) then obtain r where rdef: "r > 0 \ (\s > 0. \x. (x \ a \ \x + -a\ < s \ norm (X x + -L) \ r))" by auto @@ -2328,7 +2350,8 @@ ML {* fast_arith_split_limit := 9; *} (* FIXME *) lemma LIMSEQ_SEQ_conv: - "(\S. (\n. S n \ a) \ S ----> a \ (\n. X (S n)) ----> L) = (X -- a --> L)" + "(\S. (\n. S n \ a) \ S ----> (a::real) \ (\n. X (S n)) ----> L) = + (X -- a --> L)" proof assume "\S. (\n. S n \ a) \ S ----> a \ (\n. X (S n)) ----> L" show "X -- a --> L" by (rule LIMSEQ_SEQ_conv2) @@ -2356,26 +2379,26 @@ proof - have "f -- a --> L" . hence - fd: "\r > 0. \s > 0. \x. x \ a & \x + -a\ < s --> norm (f x + -L) < r" + fd: "\r > 0. \s > 0. \x. x \ a & norm (x + -a) < s --> norm (f x + -L) < r" by (unfold LIM_def) { fix r::real assume rgz: "0 < r" - with fd have "\s > 0. \x. x\a \ \x + -a\ < s --> norm (f x + -L) < r" by simp - then obtain s where sgz: "s > 0" and ax: "\x. x\a \ \x + -a\ < s \ norm (f x + -L) < r" by auto - from ax have ax2: "\x. (x+c)\a \ \(x+c) + -a\ < s \ norm (f (x+c) + -L) < r" by auto + with fd have "\s > 0. \x. x\a \ norm (x + -a) < s --> norm (f x + -L) < r" by simp + then obtain s where sgz: "s > 0" and ax: "\x. x\a \ norm (x + -a) < s \ norm (f x + -L) < r" by auto + from ax have ax2: "\x. (x+c)\a \ norm ((x+c) + -a) < s \ norm (f (x+c) + -L) < r" by auto { - fix x::real - from ax2 have nt: "(x+c)\a \ \(x+c) + -a\ < s \ norm (f (x+c) + -L) < r" .. + fix x + from ax2 have nt: "(x+c)\a \ norm ((x+c) + -a) < s \ norm (f (x+c) + -L) < r" .. moreover have "((x+c)\a) = (x\(a-c))" by auto moreover have "((x+c) + -a) = (x + -(a-c))" by simp - ultimately have "x\(a-c) \ \x + -(a-c)\ < s \ norm (f (x+c) + -L) < r" by simp + ultimately have "x\(a-c) \ norm (x + -(a-c)) < s \ norm (f (x+c) + -L) < r" by simp } - then have "\x. x\(a-c) \ \x + -(a-c)\ < s \ norm (f (x+c) + -L) < r" .. - with sgz have "\s > 0. \x. x\(a-c) \ \x + -(a-c)\ < s \ norm (f (x+c) + -L) < r" by auto + then have "\x. x\(a-c) \ norm (x + -(a-c)) < s \ norm (f (x+c) + -L) < r" .. + with sgz have "\s > 0. \x. x\(a-c) \ norm (x + -(a-c)) < s \ norm (f (x+c) + -L) < r" by auto } then have - "\r > 0. \s > 0. \x. x \ (a-c) & \x + -(a-c)\ < s --> norm (f (x+c) + -L) < r" by simp + "\r > 0. \s > 0. \x. x \ (a-c) & norm (x + -(a-c)) < s --> norm (f (x+c) + -L) < r" by simp thus ?thesis by (fold LIM_def) qed