(* Title : HLim.thy ID : $Id: HLim.thy,v 1.5 2007/10/16 21:12:56 haftmann Exp $ Author : Jacques D. Fleuriot Copyright : 1998 University of Cambridge Conversion to Isar and new proofs by Lawrence C Paulson, 2004 *) header{* Limits and Continuity (Nonstandard) *} theory HLim imports Star Lim begin text{*Nonstandard Definitions*} definition NSLIM :: "['a::real_normed_vector => 'b::real_normed_vector, 'a, 'b] => bool" ("((_)/ -- (_)/ --NS> (_))" [60, 0, 60] 60) where "f -- a --NS> L = (∀x. (x ≠ star_of a & x @= star_of a --> ( *f* f) x @= star_of L))" definition isNSCont :: "['a::real_normed_vector => 'b::real_normed_vector, 'a] => bool" where --{*NS definition dispenses with limit notions*} "isNSCont f a = (∀y. y @= star_of a --> ( *f* f) y @= star_of (f a))" definition isNSUCont :: "['a::real_normed_vector => 'b::real_normed_vector] => bool" where "isNSUCont f = (∀x y. x @= y --> ( *f* f) x @= ( *f* f) y)" subsection {* Limits of Functions *} lemma NSLIM_I: "(!!x. [|x ≠ star_of a; x ≈ star_of a|] ==> starfun f x ≈ star_of L) ==> f -- a --NS> L" by (simp add: NSLIM_def) lemma NSLIM_D: "[|f -- a --NS> L; x ≠ star_of a; x ≈ star_of a|] ==> starfun f x ≈ star_of L" by (simp add: NSLIM_def) text{*Proving properties of limits using nonstandard definition. The properties hold for standard limits as well!*} lemma NSLIM_mult: fixes l m :: "'a::real_normed_algebra" shows "[| f -- x --NS> l; g -- x --NS> m |] ==> (%x. f(x) * g(x)) -- x --NS> (l * m)" by (auto simp add: NSLIM_def intro!: approx_mult_HFinite) lemma starfun_scaleR [simp]: "starfun (λx. f x *R g x) = (λx. scaleHR (starfun f x) (starfun g x))" by transfer (rule refl) lemma NSLIM_scaleR: "[| f -- x --NS> l; g -- x --NS> m |] ==> (%x. f(x) *R g(x)) -- x --NS> (l *R m)" by (auto simp add: NSLIM_def intro!: approx_scaleR_HFinite) lemma NSLIM_add: "[| f -- x --NS> l; g -- x --NS> m |] ==> (%x. f(x) + g(x)) -- x --NS> (l + m)" by (auto simp add: NSLIM_def intro!: approx_add) lemma NSLIM_const [simp]: "(%x. k) -- x --NS> k" by (simp add: NSLIM_def) lemma NSLIM_minus: "f -- a --NS> L ==> (%x. -f(x)) -- a --NS> -L" by (simp add: NSLIM_def) lemma NSLIM_diff: "[|f -- x --NS> l; g -- x --NS> m|] ==> (λx. f x - g x) -- x --NS> (l - m)" by (simp only: diff_def NSLIM_add NSLIM_minus) lemma NSLIM_add_minus: "[| f -- x --NS> l; g -- x --NS> m |] ==> (%x. f(x) + -g(x)) -- x --NS> (l + -m)" by (simp only: NSLIM_add NSLIM_minus) lemma NSLIM_inverse: fixes L :: "'a::real_normed_div_algebra" shows "[| f -- a --NS> L; L ≠ 0 |] ==> (%x. inverse(f(x))) -- a --NS> (inverse L)" apply (simp add: NSLIM_def, clarify) apply (drule spec) apply (auto simp add: star_of_approx_inverse) done lemma NSLIM_zero: assumes f: "f -- a --NS> l" shows "(%x. f(x) - l) -- a --NS> 0" proof - have "(λx. f x - l) -- a --NS> l - l" by (rule NSLIM_diff [OF f NSLIM_const]) thus ?thesis by simp qed lemma NSLIM_zero_cancel: "(%x. f(x) - l) -- x --NS> 0 ==> f -- x --NS> l" apply (drule_tac g = "%x. l" and m = l in NSLIM_add) apply (auto simp add: diff_minus add_assoc) done lemma NSLIM_const_not_eq: fixes a :: "'a::real_normed_algebra_1" shows "k ≠ L ==> ¬ (λx. k) -- a --NS> L" apply (simp add: NSLIM_def) apply (rule_tac x="star_of a + of_hypreal epsilon" in exI) apply (simp add: hypreal_epsilon_not_zero approx_def) done lemma NSLIM_not_zero: fixes a :: "'a::real_normed_algebra_1" shows "k ≠ 0 ==> ¬ (λx. k) -- a --NS> 0" by (rule NSLIM_const_not_eq) lemma NSLIM_const_eq: fixes a :: "'a::real_normed_algebra_1" shows "(λx. k) -- a --NS> L ==> k = L" apply (rule ccontr) apply (blast dest: NSLIM_const_not_eq) done lemma NSLIM_unique: fixes a :: "'a::real_normed_algebra_1" shows "[|f -- a --NS> L; f -- a --NS> M|] ==> L = M" apply (drule (1) NSLIM_diff) apply (auto dest!: NSLIM_const_eq) done lemma NSLIM_mult_zero: 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) lemma NSLIM_self: "(%x. x) -- a --NS> a" by (simp add: NSLIM_def) subsubsection {* Equivalence of @{term LIM} and @{term NSLIM} *} lemma LIM_NSLIM: assumes f: "f -- a --> L" shows "f -- a --NS> L" proof (rule NSLIM_I) fix x assume neq: "x ≠ star_of a" assume approx: "x ≈ star_of a" have "starfun f x - star_of L ∈ Infinitesimal" proof (rule InfinitesimalI2) fix r::real assume r: "0 < r" from LIM_D [OF f r] obtain s where s: "0 < s" and less_r: "!!x. [|x ≠ a; norm (x - a) < s|] ==> norm (f x - L) < r" by fast from less_r have less_r': "!!x. [|x ≠ star_of a; hnorm (x - star_of a) < star_of s|] ==> hnorm (starfun f x - star_of L) < star_of r" by transfer from approx have "x - star_of a ∈ Infinitesimal" by (unfold approx_def) hence "hnorm (x - star_of a) < star_of s" using s by (rule InfinitesimalD2) with neq show "hnorm (starfun f x - star_of L) < star_of r" by (rule less_r') qed thus "starfun f x ≈ star_of L" by (unfold approx_def) qed lemma NSLIM_LIM: assumes f: "f -- a --NS> L" shows "f -- a --> L" proof (rule LIM_I) fix r::real assume r: "0 < r" have "∃s>0. ∀x. x ≠ star_of a ∧ hnorm (x - star_of a) < s --> hnorm (starfun f x - star_of L) < star_of r" proof (rule exI, safe) show "0 < epsilon" by (rule hypreal_epsilon_gt_zero) next fix x assume neq: "x ≠ star_of a" assume "hnorm (x - star_of a) < epsilon" with Infinitesimal_epsilon have "x - star_of a ∈ Infinitesimal" by (rule hnorm_less_Infinitesimal) hence "x ≈ star_of a" by (unfold approx_def) with f neq have "starfun f x ≈ star_of L" by (rule NSLIM_D) hence "starfun f x - star_of L ∈ Infinitesimal" by (unfold approx_def) thus "hnorm (starfun f x - star_of L) < star_of r" using r by (rule InfinitesimalD2) qed thus "∃s>0. ∀x. x ≠ a ∧ norm (x - a) < s --> norm (f x - L) < r" by transfer qed theorem LIM_NSLIM_iff: "(f -- x --> L) = (f -- x --NS> L)" by (blast intro: LIM_NSLIM NSLIM_LIM) subsection {* Continuity *} lemma isNSContD: "[|isNSCont f a; y ≈ star_of a|] ==> ( *f* f) y ≈ star_of (f a)" by (simp add: isNSCont_def) lemma isNSCont_NSLIM: "isNSCont f a ==> f -- a --NS> (f a) " by (simp add: isNSCont_def NSLIM_def) lemma NSLIM_isNSCont: "f -- a --NS> (f a) ==> isNSCont f a" apply (simp add: isNSCont_def NSLIM_def, auto) apply (case_tac "y = star_of a", auto) done text{*NS continuity can be defined using NS Limit in similar fashion to standard def of continuity*} lemma isNSCont_NSLIM_iff: "(isNSCont f a) = (f -- a --NS> (f a))" by (blast intro: isNSCont_NSLIM NSLIM_isNSCont) text{*Hence, NS continuity can be given in terms of standard limit*} lemma isNSCont_LIM_iff: "(isNSCont f a) = (f -- a --> (f a))" by (simp add: LIM_NSLIM_iff isNSCont_NSLIM_iff) text{*Moreover, it's trivial now that NS continuity is equivalent to standard continuity*} lemma isNSCont_isCont_iff: "(isNSCont f a) = (isCont f a)" apply (simp add: isCont_def) apply (rule isNSCont_LIM_iff) done text{*Standard continuity ==> NS continuity*} lemma isCont_isNSCont: "isCont f a ==> isNSCont f a" by (erule isNSCont_isCont_iff [THEN iffD2]) text{*NS continuity ==> Standard continuity*} lemma isNSCont_isCont: "isNSCont f a ==> isCont f a" by (erule isNSCont_isCont_iff [THEN iffD1]) text{*Alternative definition of continuity*} (* Prove equivalence between NS limits - *) (* 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 = "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 (auto simp add: starfun star_of_def star_n_minus star_n_add add_assoc approx_refl star_n_zero_num) done lemma NSLIM_isCont_iff: "(f -- a --NS> f a) = ((%h. f(a + h)) -- 0 --NS> f a)" by (rule NSLIM_h_iff) lemma isNSCont_minus: "isNSCont f a ==> isNSCont (%x. - f x) a" by (simp add: isNSCont_def) lemma isNSCont_inverse: fixes f :: "'a::real_normed_vector => 'b::real_normed_div_algebra" shows "[| isNSCont f x; f x ≠ 0 |] ==> isNSCont (%x. inverse (f x)) x" by (auto intro: isCont_inverse simp add: isNSCont_isCont_iff) lemma isNSCont_const [simp]: "isNSCont (%x. k) a" by (simp add: isNSCont_def) lemma isNSCont_abs [simp]: "isNSCont abs (a::real)" apply (simp add: isNSCont_def) apply (auto intro: approx_hrabs simp add: starfun_rabs_hrabs) done (**************************************************************** (%* Leave as commented until I add topology theory or remove? *%) (%*------------------------------------------------------------ Elementary topology proof for a characterisation of continuity now: a function f is continuous if and only if the inverse image, {x. f(x) ∈ A}, of any open set A is always an open set ------------------------------------------------------------*%) Goal "[| isNSopen A; ∀x. isNSCont f x |] ==> isNSopen {x. f x ∈ A}" by (auto_tac (claset(),simpset() addsimps [isNSopen_iff1])); by (dtac (mem_monad_approx RS approx_sym); by (dres_inst_tac [("x","a")] spec 1); by (dtac isNSContD 1 THEN assume_tac 1) by (dtac bspec 1 THEN assume_tac 1) by (dres_inst_tac [("x","( *f* f) x")] approx_mem_monad2 1); by (blast_tac (claset() addIs [starfun_mem_starset]); qed "isNSCont_isNSopen"; Goalw [isNSCont_def] "∀A. isNSopen A --> isNSopen {x. f x ∈ A} \ \ ==> isNSCont f x"; by (auto_tac (claset() addSIs [(mem_infmal_iff RS iffD1) RS (approx_minus_iff RS iffD2)],simpset() addsimps [Infinitesimal_def,SReal_iff])); by (dres_inst_tac [("x","{z. abs(z + -f(x)) < ya}")] spec 1); by (etac (isNSopen_open_interval RSN (2,impE)); by (auto_tac (claset(),simpset() addsimps [isNSopen_def,isNSnbhd_def])); by (dres_inst_tac [("x","x")] spec 1); by (auto_tac (claset() addDs [approx_sym RS approx_mem_monad], simpset() addsimps [hypreal_of_real_zero RS sym,STAR_starfun_rabs_add_minus])); qed "isNSopen_isNSCont"; Goal "(∀x. isNSCont f x) = \ \ (∀A. isNSopen A --> isNSopen {x. f(x) ∈ A})"; by (blast_tac (claset() addIs [isNSCont_isNSopen, isNSopen_isNSCont]); qed "isNSCont_isNSopen_iff"; (%*------- Standard version of same theorem --------*%) Goal "(∀x. isCont f x) = \ \ (∀A. isopen A --> isopen {x. f(x) ∈ A})"; by (auto_tac (claset() addSIs [isNSCont_isNSopen_iff], simpset() addsimps [isNSopen_isopen_iff RS sym, isNSCont_isCont_iff RS sym])); qed "isCont_isopen_iff"; *******************************************************************) subsection {* Uniform Continuity *} lemma isNSUContD: "[| isNSUCont f; x ≈ y|] ==> ( *f* f) x ≈ ( *f* f) y" by (simp add: isNSUCont_def) lemma isUCont_isNSUCont: fixes f :: "'a::real_normed_vector => 'b::real_normed_vector" assumes f: "isUCont f" shows "isNSUCont f" proof (unfold isNSUCont_def, safe) fix x y :: "'a star" assume approx: "x ≈ y" have "starfun f x - starfun f y ∈ Infinitesimal" proof (rule InfinitesimalI2) fix r::real assume r: "0 < r" with f obtain s where s: "0 < s" and less_r: "!!x y. norm (x - y) < s ==> norm (f x - f y) < r" by (auto simp add: isUCont_def) from less_r have less_r': "!!x y. hnorm (x - y) < star_of s ==> hnorm (starfun f x - starfun f y) < star_of r" by transfer from approx have "x - y ∈ Infinitesimal" by (unfold approx_def) hence "hnorm (x - y) < star_of s" using s by (rule InfinitesimalD2) thus "hnorm (starfun f x - starfun f y) < star_of r" by (rule less_r') qed thus "starfun f x ≈ starfun f y" by (unfold approx_def) qed lemma isNSUCont_isUCont: fixes f :: "'a::real_normed_vector => 'b::real_normed_vector" assumes f: "isNSUCont f" shows "isUCont f" proof (unfold isUCont_def, safe) fix r::real assume r: "0 < r" have "∃s>0. ∀x y. hnorm (x - y) < s --> hnorm (starfun f x - starfun f y) < star_of r" proof (rule exI, safe) show "0 < epsilon" by (rule hypreal_epsilon_gt_zero) next fix x y :: "'a star" assume "hnorm (x - y) < epsilon" with Infinitesimal_epsilon have "x - y ∈ Infinitesimal" by (rule hnorm_less_Infinitesimal) hence "x ≈ y" by (unfold approx_def) with f have "starfun f x ≈ starfun f y" by (simp add: isNSUCont_def) hence "starfun f x - starfun f y ∈ Infinitesimal" by (unfold approx_def) thus "hnorm (starfun f x - starfun f y) < star_of r" using r by (rule InfinitesimalD2) qed thus "∃s>0. ∀x y. norm (x - y) < s --> norm (f x - f y) < r" by transfer qed end
lemma NSLIM_I:
(!!x. [| x ≠ star_of a; x ≈ star_of a |] ==> (*f* f) x ≈ star_of L)
==> f -- a --NS> L
lemma NSLIM_D:
[| f -- a --NS> L; x ≠ star_of a; x ≈ star_of a |] ==> (*f* f) x ≈ star_of L
lemma NSLIM_mult:
[| f -- x --NS> l; g -- x --NS> m |] ==> (λx. f x * g x) -- x --NS> l * m
lemma starfun_scaleR:
*f* (λx. f x *R g x) = (λx. scaleHR ((*f* f) x) ((*f* g) x))
lemma NSLIM_scaleR:
[| f -- x --NS> l; g -- x --NS> m |] ==> (λx. f x *R g x) -- x --NS> l *R m
lemma NSLIM_add:
[| f -- x --NS> l; g -- x --NS> m |] ==> (λx. f x + g x) -- x --NS> l + m
lemma NSLIM_const:
(λx. k) -- x --NS> k
lemma NSLIM_minus:
f -- a --NS> L ==> (λx. - f x) -- a --NS> - L
lemma NSLIM_diff:
[| f -- x --NS> l; g -- x --NS> m |] ==> (λx. f x - g x) -- x --NS> l - m
lemma NSLIM_add_minus:
[| f -- x --NS> l; g -- x --NS> m |] ==> (λx. f x + - g x) -- x --NS> l + - m
lemma NSLIM_inverse:
[| f -- a --NS> L; L ≠ (0::'a) |] ==> (λx. inverse (f x)) -- a --NS> inverse L
lemma NSLIM_zero:
f -- a --NS> l ==> (λx. f x - l) -- a --NS> (0::'b)
lemma NSLIM_zero_cancel:
(λx. f x - l) -- x --NS> (0::'b) ==> f -- x --NS> l
lemma NSLIM_const_not_eq:
k ≠ L ==> ¬ (λx. k) -- a --NS> L
lemma NSLIM_not_zero:
k ≠ (0::'b) ==> ¬ (λx. k) -- a --NS> (0::'b)
lemma NSLIM_const_eq:
(λx. k) -- a --NS> L ==> k = L
lemma NSLIM_unique:
[| f -- a --NS> L; f -- a --NS> M |] ==> L = M
lemma NSLIM_mult_zero:
[| f -- x --NS> (0::'b); g -- x --NS> (0::'b) |]
==> (λx. f x * g x) -- x --NS> (0::'b)
lemma NSLIM_self:
(λx. x) -- a --NS> a
lemma LIM_NSLIM:
f -- a --> L ==> f -- a --NS> L
lemma NSLIM_LIM:
f -- a --NS> L ==> f -- a --> L
theorem LIM_NSLIM_iff:
f -- x --> L = f -- x --NS> L
lemma isNSContD:
[| isNSCont f a; y ≈ star_of a |] ==> (*f* f) y ≈ star_of (f a)
lemma isNSCont_NSLIM:
isNSCont f a ==> f -- a --NS> f a
lemma NSLIM_isNSCont:
f -- a --NS> f a ==> isNSCont f a
lemma isNSCont_NSLIM_iff:
isNSCont f a = f -- a --NS> f a
lemma isNSCont_LIM_iff:
isNSCont f a = f -- a --> f a
lemma isNSCont_isCont_iff:
isNSCont f a = isCont f a
lemma isCont_isNSCont:
isCont f a ==> isNSCont f a
lemma isNSCont_isCont:
isNSCont f a ==> isCont f a
lemma NSLIM_h_iff:
f -- a --NS> L = (λh. f (a + h)) -- 0::'a --NS> L
lemma NSLIM_isCont_iff:
f -- a --NS> f a = (λh. f (a + h)) -- 0::'a --NS> f a
lemma isNSCont_minus:
isNSCont f a ==> isNSCont (λx. - f x) a
lemma isNSCont_inverse:
[| isNSCont f x; f x ≠ (0::'b) |] ==> isNSCont (λx. inverse (f x)) x
lemma isNSCont_const:
isNSCont (λx. k) a
lemma isNSCont_abs:
isNSCont abs a
lemma isNSUContD:
[| isNSUCont f; x ≈ y |] ==> (*f* f) x ≈ (*f* f) y
lemma isUCont_isNSUCont:
isUCont f ==> isNSUCont f
lemma isNSUCont_isUCont:
isNSUCont f ==> isUCont f