(* Title : Lim.thy ID : $Id: Lim.thy,v 1.24 2005/09/09 17:34:22 huffman Exp $ Author : Jacques D. Fleuriot Copyright : 1998 University of Cambridge Conversion to Isar and new proofs by Lawrence C Paulson, 2004 *) header{*Limits, Continuity and Differentiation*} theory Lim imports SEQ begin text{*Standard and Nonstandard Definitions*} constdefs LIM :: "[real=>real,real,real] => bool" ("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60) "f -- a --> L == ∀r > 0. ∃s > 0. ∀x. x ≠ a & ¦x + -a¦ < s --> ¦f x + -L¦ < r" NSLIM :: "[real=>real,real,real] => bool" ("((_)/ -- (_)/ --NS> (_))" [60, 0, 60] 60) "f -- a --NS> L == (∀x. (x ≠ hypreal_of_real a & x @= hypreal_of_real a --> ( *f* f) x @= hypreal_of_real L))" isCont :: "[real=>real,real] => bool" "isCont f a == (f -- a --> (f a))" isNSCont :: "[real=>real,real] => bool" --{*NS definition dispenses with limit notions*} "isNSCont f a == (∀y. y @= hypreal_of_real a --> ( *f* f) y @= hypreal_of_real (f a))" deriv:: "[real=>real,real,real] => bool" --{*Differentiation: D is derivative of function f at x*} ("(DERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) "DERIV f x :> D == ((%h. (f(x + h) + -f x)/h) -- 0 --> D)" nsderiv :: "[real=>real,real,real] => bool" ("(NSDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) "NSDERIV f x :> D == (∀h ∈ Infinitesimal - {0}. (( *f* f)(hypreal_of_real x + h) + - hypreal_of_real (f x))/h @= hypreal_of_real D)" differentiable :: "[real=>real,real] => bool" (infixl "differentiable" 60) "f differentiable x == (∃D. DERIV f x :> D)" NSdifferentiable :: "[real=>real,real] => bool" (infixl "NSdifferentiable" 60) "f NSdifferentiable x == (∃D. NSDERIV f x :> D)" increment :: "[real=>real,real,hypreal] => hypreal" "increment f x h == (@inc. f NSdifferentiable x & inc = ( *f* f)(hypreal_of_real x + h) + -hypreal_of_real (f x))" isUCont :: "(real=>real) => bool" "isUCont f == ∀r > 0. ∃s > 0. ∀x y. ¦x + -y¦ < s --> ¦f x + -f y¦ < r" isNSUCont :: "(real=>real) => bool" "isNSUCont f == (∀x y. x @= y --> ( *f* f) x @= ( *f* f) y)" consts Bolzano_bisect :: "[real*real=>bool, real, real, nat] => (real*real)" --{*Used in the proof of the Bolzano theorem*} primrec "Bolzano_bisect P a b 0 = (a,b)" "Bolzano_bisect P a b (Suc n) = (let (x,y) = Bolzano_bisect P a b n in if P(x, (x+y)/2) then ((x+y)/2, y) else (x, (x+y)/2))" section{*Some Purely Standard Proofs*} lemma LIM_eq: "f -- a --> L = (∀r>0.∃s>0.∀x. x ≠ a & ¦x-a¦ < s --> ¦f x - L¦ < r)" by (simp add: LIM_def diff_minus) lemma LIM_D: "[| f -- a --> L; 0<r |] ==> ∃s>0.∀x. x ≠ a & ¦x-a¦ < s --> ¦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: assumes f: "f -- a --> L" and g: "g -- a --> M" shows "(%x. f x + g(x)) -- a --> (L + M)" proof (simp add: LIM_eq, clarify) fix r :: real assume r: "0<r" 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 --> ¦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 --> ¦g x - M¦ < r/2" by blast show "∃s>0.∀x. x ≠ a ∧ ¦x-a¦ < s --> ¦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" with fs_lt gs_lt have "¦f x - L¦ < r/2" and "¦g x - M¦ < r/2" by (auto simp add: fs_lt) hence "¦f x - L¦ + ¦g x - M¦ < r" by arith thus "¦f x + g x - (L + M)¦ < r" by (blast intro: abs_diff_triangle_ineq order_le_less_trans) qed qed lemma LIM_minus: "f -- a --> L ==> (%x. -f(x)) -- a --> -L" apply (simp add: LIM_eq) apply (subgoal_tac "∀x. ¦- f x + L¦ = ¦f x - L¦") apply (simp_all add: abs_if) done lemma LIM_add_minus: "[| f -- x --> l; g -- x --> m |] ==> (%x. f(x) + -g(x)) -- x --> (l + -m)" by (blast dest: LIM_add LIM_minus) lemma LIM_diff: "[| f -- x --> l; g -- x --> m |] ==> (%x. f(x) - g(x)) -- x --> l-m" by (simp add: diff_minus LIM_add_minus) lemma LIM_const_not_eq: "k ≠ L ==> ~ ((%x. k) -- a --> L)" proof (simp add: linorder_neq_iff LIM_eq, elim disjE) assume k: "k < L" show "∃r>0. ∀s>0. (∃x. (x < a ∨ a < x) ∧ ¦x-a¦ < s) ∧ ¬ ¦k-L¦ < r" proof (intro exI conjI strip) show "0 < L-k" by (simp add: k compare_rls) fix s :: real assume s: "0<s" { from s show "s/2 + a < a ∨ a < s/2 + a" by arith next from s show "¦s / 2 + a - a¦ < s" by (simp add: abs_if) next from s show "~ ¦k-L¦ < L-k" by (simp add: abs_if) } qed next assume k: "L < k" show "∃r>0.∀s>0. (∃x. (x < a ∨ a < x) ∧ ¦x-a¦ < s) ∧ ¬ ¦k-L¦ < r" proof (intro exI conjI strip) show "0 < k-L" by (simp add: k compare_rls) fix s :: real assume s: "0<s" { from s show "s/2 + a < a ∨ a < s/2 + a" by arith next from s show "¦s / 2 + a - a¦ < s" by (simp add: abs_if) next from s show "~ ¦k-L¦ < k-L" by (simp add: abs_if) } qed qed lemma LIM_const_eq: "(%x. k) -- x --> 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" apply (drule LIM_diff, assumption) apply (auto dest!: LIM_const_eq) done lemma LIM_mult_zero: assumes f: "f -- a --> 0" and g: "g -- a --> 0" shows "(%x. f(x) * g(x)) -- a --> 0" proof (simp add: LIM_eq abs_mult, clarify) fix r :: real assume r: "0<r" from LIM_D [OF f zero_less_one] obtain fs where fs: "0 < fs" and fs_lt: "∀x. x ≠ a & ¦x-a¦ < fs --> ¦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 --> ¦g x¦ < r" by auto show "∃s. 0 < s ∧ (∀x. x ≠ a ∧ ¦x-a¦ < s --> ¦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" with fs_lt gs_lt have "¦f x¦ < 1" and "¦g x¦ < r" by (auto simp add: fs_lt) hence "¦f x¦ * ¦g x¦ < 1*r" by (rule abs_mult_less) thus "¦f x¦ * ¦g x¦ < r" by simp qed qed lemma LIM_self: "(%x. x) -- a --> a" by (auto simp add: LIM_def) text{*Limits are equal for functions equal except at limit point*} lemma LIM_equal: "[| ∀x. x ≠ a --> (f x = g x) |] ==> (f -- a --> l) = (g -- a --> l)" by (simp add: LIM_def) text{*Two uses in Hyperreal/Transcendental.ML*} lemma LIM_trans: "[| (%x. f(x) + -g(x)) -- a --> 0; g -- a --> l |] ==> f -- a --> l" apply (drule LIM_add, assumption) apply (auto simp add: add_assoc) done subsection{*Relationships Between Standard and Nonstandard Concepts*} text{*Standard and NS definitions of Limit*} (*NEEDS STRUCTURING*) lemma LIM_NSLIM: "f -- x --> L ==> f -- x --NS> L" apply (simp add: LIM_def NSLIM_def approx_def) apply (simp add: Infinitesimal_FreeUltrafilterNat_iff, safe) apply (rule_tac x = xa in star_cases) apply (auto simp add: real_add_minus_iff starfun star_n_minus star_of_def star_n_add star_n_eq_iff) apply (rule bexI [OF _ Rep_star_star_n], clarify) apply (drule_tac x = u in spec, clarify) apply (drule_tac x = s in spec, clarify) apply (subgoal_tac "∀n::nat. (Xa n) ≠ x & ¦(Xa n) + - x¦ < s --> ¦f (Xa n) + - L¦ < u") prefer 2 apply blast apply (drule FreeUltrafilterNat_all, ultra) done subsubsection{*Limit: The NS definition implies the standard definition.*} lemma lemma_LIM: "∀s>0.∃xa. xa ≠ x & ¦xa + - x¦ < s & r ≤ ¦f xa + -L¦ ==> ∀n::nat. ∃xa. xa ≠ x & ¦xa + -x¦ < inverse(real(Suc n)) & r ≤ ¦f xa + -L¦" 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: "∀s>0.∃xa. xa ≠ x & ¦xa + - x¦ < s & r ≤ ¦f xa + -L¦ ==> ∃X. ∀n::nat. X n ≠ x & ¦X n + -x¦ < inverse(real(Suc n)) & r ≤ abs(f (X n) + -L)" apply (drule lemma_LIM) apply (drule choice, blast) done lemma lemma_simp: "∀n. X n ≠ x & ¦X n + - x¦ < inverse (real(Suc n)) & r ≤ abs (f (X n) + - L) ==> ∀n. ¦X n + - x¦ < inverse (real(Suc n))" by auto text{*NSLIM => LIM*} lemma NSLIM_LIM: "f -- x --NS> L ==> f -- x --> L" apply (simp add: LIM_def NSLIM_def approx_def) apply (simp add: Infinitesimal_FreeUltrafilterNat_iff, clarify) apply (rule ccontr, simp) apply (simp add: linorder_not_less) apply (drule lemma_skolemize_LIM2, safe) apply (drule_tac x = "star_n X" in spec) apply (auto simp add: starfun star_n_minus star_of_def star_n_add star_n_eq_iff) apply (drule lemma_simp [THEN real_seq_to_hypreal_Infinitesimal]) apply (simp add: Infinitesimal_FreeUltrafilterNat_iff star_of_def star_n_minus star_n_add star_n_eq_iff, blast) apply (drule spec, drule mp, assumption) apply (drule FreeUltrafilterNat_all, ultra) done theorem LIM_NSLIM_iff: "(f -- x --> L) = (f -- x --NS> L)" by (blast intro: LIM_NSLIM NSLIM_LIM) text{*Proving properties of limits using nonstandard definition. The properties hold for standard limits as well!*} lemma NSLIM_mult: "[| 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 LIM_mult2: "[| f -- x --> l; g -- x --> m |] ==> (%x. f(x) * g(x)) -- x --> (l * m)" by (simp add: LIM_NSLIM_iff NSLIM_mult) 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 LIM_add2: "[| f -- x --> l; g -- x --> m |] ==> (%x. f(x) + g(x)) -- x --> (l + m)" by (simp add: LIM_NSLIM_iff NSLIM_add) lemma NSLIM_const [simp]: "(%x. k) -- x --NS> k" by (simp add: NSLIM_def) lemma LIM_const2: "(%x. k) -- x --> k" by (simp add: LIM_NSLIM_iff) lemma NSLIM_minus: "f -- a --NS> L ==> (%x. -f(x)) -- a --NS> -L" by (simp add: NSLIM_def) lemma LIM_minus2: "f -- a --> L ==> (%x. -f(x)) -- a --> -L" by (simp add: LIM_NSLIM_iff NSLIM_minus) lemma NSLIM_add_minus: "[| f -- x --NS> l; g -- x --NS> m |] ==> (%x. f(x) + -g(x)) -- x --NS> (l + -m)" by (blast dest: NSLIM_add NSLIM_minus) lemma LIM_add_minus2: "[| f -- x --> l; g -- x --> m |] ==> (%x. f(x) + -g(x)) -- x --> (l + -m)" by (simp add: LIM_NSLIM_iff NSLIM_add_minus) lemma NSLIM_inverse: "[| 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: hypreal_of_real_approx_inverse) done lemma LIM_inverse: "[| f -- a --> L; L ≠ 0 |] ==> (%x. inverse(f(x))) -- a --> (inverse L)" by (simp add: LIM_NSLIM_iff NSLIM_inverse) 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_add_minus [OF f NSLIM_const]) thus ?thesis by simp qed lemma LIM_zero2: "f -- a --> l ==> (%x. f(x) + -l) -- a --> 0" by (simp add: LIM_NSLIM_iff NSLIM_zero) 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 LIM_zero_cancel: "(%x. f(x) - l) -- x --> 0 ==> f -- x --> l" apply (drule_tac g = "%x. l" and M = l in LIM_add) apply (auto simp add: diff_minus add_assoc) done lemma NSLIM_not_zero: "k ≠ 0 ==> ~ ((%x. k) -- x --NS> 0)" apply (simp add: NSLIM_def) apply (rule_tac x = "hypreal_of_real x + epsilon" in exI) apply (auto intro: Infinitesimal_add_approx_self [THEN approx_sym] simp add: hypreal_epsilon_not_zero) done lemma NSLIM_const_not_eq: "k ≠ L ==> ~ ((%x. k) -- x --NS> L)" apply (simp add: NSLIM_def) apply (rule_tac x = "hypreal_of_real x + epsilon" in exI) apply (auto intro: Infinitesimal_add_approx_self [THEN approx_sym] simp add: hypreal_epsilon_not_zero) done lemma NSLIM_const_eq: "(%x. k) -- x --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" apply (drule NSLIM_minus) apply (drule NSLIM_add, assumption) apply (auto dest!: NSLIM_const_eq [symmetric]) done lemma LIM_unique2: "[| f -- x --> L; f -- x --> M |] ==> L = M" by (simp add: LIM_NSLIM_iff NSLIM_unique) lemma NSLIM_mult_zero: "[| f -- x --NS> 0; g -- x --NS> 0 |] ==> (%x. f(x)*g(x)) -- x --NS> 0" by (drule NSLIM_mult, auto) (* we can use the corresponding thm LIM_mult2 *) (* for standard definition of limit *) lemma LIM_mult_zero2: "[| f -- x --> 0; g -- x --> 0 |] ==> (%x. f(x)*g(x)) -- x --> 0" by (drule LIM_mult2, auto) lemma NSLIM_self: "(%x. x) -- a --NS> a" by (simp add: NSLIM_def) subsection{* Derivatives and Continuity: NS and Standard properties*} subsubsection{*Continuity*} lemma isNSContD: "[| isNSCont f a; y ≈ hypreal_of_real a |] ==> ( *f* f) y ≈ hypreal_of_real (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 (rule_tac Q = "y = hypreal_of_real a" in excluded_middle [THEN disjE], 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 = "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 (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 lemma NSLIM_isCont_iff: "(f -- a --NS> f a) = ((%h. f(a + h)) -- 0 --NS> f a)" by (rule NSLIM_h_iff) lemma LIM_isCont_iff: "(f -- a --> f a) = ((%h. f(a + h)) -- 0 --> f(a))" by (simp add: LIM_NSLIM_iff NSLIM_isCont_iff) lemma isCont_iff: "(isCont f x) = ((%h. f(x + h)) -- 0 --> f(x))" by (simp add: isCont_def LIM_isCont_iff) text{*Immediate application of nonstandard criterion for continuity can offer very simple proofs of some standard property of continuous functions*} text{*sum continuous*} lemma isCont_add: "[| isCont f a; isCont g a |] ==> isCont (%x. f(x) + g(x)) a" by (auto intro: approx_add simp add: isNSCont_isCont_iff [symmetric] isNSCont_def) text{*mult continuous*} lemma isCont_mult: "[| 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] simp add: isNSCont_isCont_iff [symmetric] isNSCont_def) text{*composition of continuous functions Note very short straightforard proof!*} lemma isCont_o: "[| isCont f a; isCont g (f a) |] ==> isCont (g o f) a" by (auto simp add: isNSCont_isCont_iff [symmetric] isNSCont_def starfun_o [symmetric]) lemma isCont_o2: "[| isCont f a; isCont g (f a) |] ==> isCont (%x. g (f x)) a" by (auto dest: isCont_o simp add: o_def) lemma isNSCont_minus: "isNSCont f a ==> isNSCont (%x. - f x) a" by (simp add: isNSCont_def) lemma isCont_minus: "isCont f a ==> isCont (%x. - f x) a" by (auto simp add: isNSCont_isCont_iff [symmetric] isNSCont_minus) lemma isCont_inverse: "[| 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: "[| isNSCont f x; f x ≠ 0 |] ==> isNSCont (%x. inverse (f x)) x" by (auto intro: isCont_inverse simp add: isNSCont_isCont_iff) lemma isCont_diff: "[| isCont f a; isCont g a |] ==> isCont (%x. f(x) - g(x)) a" apply (simp add: diff_minus) apply (auto intro: isCont_add isCont_minus) done lemma isCont_const [simp]: "isCont (%x. k) a" by (simp add: isCont_def) lemma isNSCont_const [simp]: "isNSCont (%x. k) a" by (simp add: isNSCont_def) lemma isNSCont_abs [simp]: "isNSCont abs a" 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" by (auto simp add: isNSCont_isCont_iff [symmetric]) (**************************************************************** (%* 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"; *******************************************************************) text{*Uniform continuity*} lemma isNSUContD: "[| isNSUCont f; x ≈ y|] ==> ( *f* f) x ≈ ( *f* f) y" by (simp add: isNSUCont_def) lemma isUCont_isCont: "isUCont f ==> isCont f x" by (simp add: isUCont_def isCont_def LIM_def, meson) lemma isUCont_isNSUCont: "isUCont f ==> isNSUCont f" apply (simp add: isNSUCont_def isUCont_def approx_def) apply (simp add: Infinitesimal_FreeUltrafilterNat_iff, safe) apply (rule_tac x = x in star_cases) apply (rule_tac x = y in star_cases) apply (auto simp add: starfun star_n_minus star_n_add) apply (rule bexI [OF _ Rep_star_star_n], safe) apply (drule_tac x = u in spec, clarify) apply (drule_tac x = s in spec, clarify) apply (subgoal_tac "∀n::nat. abs ((Xa n) + - (Xb n)) < s --> abs (f (Xa n) + - f (Xb n)) < u") prefer 2 apply blast apply (erule_tac V = "∀x y. ¦x + - y¦ < s --> ¦f x + - f y¦ < u" in thin_rl) apply (drule FreeUltrafilterNat_all, ultra) done lemma lemma_LIMu: "∀s>0.∃z y. ¦z + - y¦ < s & r ≤ ¦f z + -f y¦ ==> ∀n::nat. ∃z y. ¦z + -y¦ < inverse(real(Suc n)) & r ≤ ¦f z + -f y¦" apply clarify apply (cut_tac n1 = n in real_of_nat_Suc_gt_zero [THEN positive_imp_inverse_positive], auto) done lemma lemma_skolemize_LIM2u: "∀s>0.∃z y. ¦z + - y¦ < s & r ≤ ¦f z + -f y¦ ==> ∃X Y. ∀n::nat. abs(X n + -(Y n)) < inverse(real(Suc n)) & r ≤ abs(f (X n) + -f (Y n))" apply (drule lemma_LIMu) apply (drule choice, safe) apply (drule choice, blast) done lemma lemma_simpu: "∀n. ¦X n + -Y n¦ < inverse (real(Suc n)) & r ≤ abs (f (X n) + - f(Y n)) ==> ∀n. ¦X n + - Y n¦ < inverse (real(Suc n))" by auto lemma isNSUCont_isUCont: "isNSUCont f ==> isUCont f" apply (simp add: isNSUCont_def isUCont_def approx_def) apply (simp add: Infinitesimal_FreeUltrafilterNat_iff, safe) apply (rule ccontr, simp) apply (simp add: linorder_not_less) apply (drule lemma_skolemize_LIM2u, safe) apply (drule_tac x = "star_n X" in spec) apply (drule_tac x = "star_n Y" in spec) apply (simp add: starfun star_n_minus star_n_add, auto) apply (drule lemma_simpu [THEN real_seq_to_hypreal_Infinitesimal2]) apply (simp add: Infinitesimal_FreeUltrafilterNat_iff star_n_minus star_n_add, blast) apply (drule_tac x = r in spec, clarify) apply (drule FreeUltrafilterNat_all, ultra) done text{*Derivatives*} lemma DERIV_iff: "(DERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- 0 --> D)" by (simp add: deriv_def) lemma DERIV_NS_iff: "(DERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- 0 --NS> D)" by (simp add: deriv_def LIM_NSLIM_iff) lemma DERIV_D: "DERIV f x :> D ==> (%h. (f(x + h) + - f(x))/h) -- 0 --> D" by (simp add: deriv_def) lemma NS_DERIV_D: "DERIV f x :> D ==> (%h. (f(x + h) + - f(x))/h) -- 0 --NS> D" by (simp add: deriv_def LIM_NSLIM_iff) subsubsection{*Uniqueness*} lemma DERIV_unique: "[| DERIV f x :> D; DERIV f x :> E |] ==> D = E" apply (simp add: deriv_def) apply (blast intro: LIM_unique) done lemma NSDeriv_unique: "[| NSDERIV f x :> D; NSDERIV f x :> E |] ==> D = E" apply (simp add: nsderiv_def) apply (cut_tac Infinitesimal_epsilon hypreal_epsilon_not_zero) apply (auto dest!: bspec [where x=epsilon] intro!: inj_hypreal_of_real [THEN injD] dest: approx_trans3) done subsubsection{*Differentiable*} lemma differentiableD: "f differentiable x ==> ∃D. DERIV f x :> D" by (simp add: differentiable_def) lemma differentiableI: "DERIV f x :> D ==> f differentiable x" by (force simp add: differentiable_def) lemma NSdifferentiableD: "f NSdifferentiable x ==> ∃D. NSDERIV f x :> D" by (simp add: NSdifferentiable_def) lemma NSdifferentiableI: "NSDERIV f x :> D ==> f NSdifferentiable x" by (force simp add: NSdifferentiable_def) subsubsection{*Alternative definition for differentiability*} lemma LIM_I: "(!!r. 0<r ==> ∃s>0.∀x. x ≠ a & ¦x-a¦ < s --> ¦f x - L¦ < r) ==> f -- a --> L" by (simp add: LIM_eq) lemma DERIV_LIM_iff: "((%h. (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 assume r: "0<r" assume "(λh. (f (a + h) - f a) / h) -- 0 --> D" from LIM_D [OF this r] obtain s where s: "0 < s" 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 --> ¦(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" with s_lt [THEN spec [where x="x-a"]] show "¦(f x - f a) / (x-a) - D¦ < r" by auto qed next fix r::real assume r: "0<r" assume "(λx. (f x - f a) / (x-a)) -- a --> D" from LIM_D [OF this r] obtain s where s: "0 < s" 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 --> ¦(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" with s_lt [THEN spec [where x="x+a"]] show "¦(f (a + x) - f a) / x - D¦ < r" by (auto simp add: add_ac) qed qed lemma DERIV_iff2: "(DERIV f x :> D) = ((%z. (f(z) - f(x)) / (z-x)) -- x --> D)" by (simp add: deriv_def diff_minus [symmetric] DERIV_LIM_iff) subsection{*Equivalence of NS and standard definitions of differentiation*} subsubsection{*First NSDERIV in terms of NSLIM*} text{*first equivalence *} lemma NSDERIV_NSLIM_iff: "(NSDERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- 0 --NS> D)" apply (simp add: nsderiv_def NSLIM_def, auto) apply (drule_tac x = xa in bspec) apply (rule_tac [3] ccontr) apply (drule_tac [3] x = h in spec) apply (auto simp add: mem_infmal_iff starfun_lambda_cancel) done text{*second equivalence *} lemma NSDERIV_NSLIM_iff2: "(NSDERIV f x :> D) = ((%z. (f(z) - f(x)) / (z-x)) -- x --NS> D)" by (simp add: NSDERIV_NSLIM_iff DERIV_LIM_iff diff_minus [symmetric] LIM_NSLIM_iff [symmetric]) (* while we're at it! *) lemma NSDERIV_iff2: "(NSDERIV f x :> D) = (∀w. w ≠ hypreal_of_real x & w ≈ hypreal_of_real x --> ( *f* (%z. (f z - f x) / (z-x))) w ≈ hypreal_of_real D)" by (simp add: NSDERIV_NSLIM_iff2 NSLIM_def) (*FIXME DELETE*) lemma hypreal_not_eq_minus_iff: "(x ≠ a) = (x + -a ≠ (0::hypreal))" by (auto dest: hypreal_eq_minus_iff [THEN iffD2]) lemma NSDERIVD5: "(NSDERIV f x :> D) ==> (∀u. u ≈ hypreal_of_real x --> ( *f* (%z. f z - f x)) u ≈ hypreal_of_real D * (u - hypreal_of_real x))" apply (auto simp add: NSDERIV_iff2) apply (case_tac "u = hypreal_of_real x", auto) apply (drule_tac x = u in spec, auto) apply (drule_tac c = "u - hypreal_of_real x" and b = "hypreal_of_real D" in approx_mult1) apply (drule_tac [!] hypreal_not_eq_minus_iff [THEN iffD1]) apply (subgoal_tac [2] "( *f* (%z. z-x)) u ≠ (0::hypreal) ") apply (auto simp add: diff_minus approx_minus_iff [THEN iffD1, THEN mem_infmal_iff [THEN iffD2]] Infinitesimal_subset_HFinite [THEN subsetD]) done lemma NSDERIVD4: "(NSDERIV f x :> D) ==> (∀h ∈ Infinitesimal. (( *f* f)(hypreal_of_real x + h) - hypreal_of_real (f x))≈ (hypreal_of_real D) * h)" apply (auto simp add: nsderiv_def) apply (case_tac "h = (0::hypreal) ") apply (auto simp add: diff_minus) apply (drule_tac x = h in bspec) apply (drule_tac [2] c = h in approx_mult1) apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] simp add: diff_minus) done lemma NSDERIVD3: "(NSDERIV f x :> D) ==> (∀h ∈ Infinitesimal - {0}. (( *f* f)(hypreal_of_real x + h) - hypreal_of_real (f x))≈ (hypreal_of_real D) * h)" apply (auto simp add: nsderiv_def) apply (rule ccontr, drule_tac x = h in bspec) apply (drule_tac [2] c = h in approx_mult1) apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] simp add: mult_assoc diff_minus) done text{*Now equivalence between NSDERIV and DERIV*} lemma NSDERIV_DERIV_iff: "(NSDERIV f x :> D) = (DERIV f x :> D)" by (simp add: deriv_def NSDERIV_NSLIM_iff LIM_NSLIM_iff) text{*Differentiability implies continuity nice and simple "algebraic" proof*} lemma NSDERIV_isNSCont: "NSDERIV f x :> D ==> isNSCont f x" apply (auto simp add: nsderiv_def isNSCont_NSLIM_iff NSLIM_def) apply (drule approx_minus_iff [THEN iffD1]) apply (drule hypreal_not_eq_minus_iff [THEN iffD1]) apply (drule_tac x = "-hypreal_of_real x + xa" in bspec) prefer 2 apply (simp add: add_assoc [symmetric]) apply (auto simp add: mem_infmal_iff [symmetric] add_commute) apply (drule_tac c = "xa + -hypreal_of_real x" in approx_mult1) apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] simp add: mult_assoc) apply (drule_tac x3=D in HFinite_hypreal_of_real [THEN [2] Infinitesimal_HFinite_mult, THEN mem_infmal_iff [THEN iffD1]]) apply (auto simp add: mult_commute intro: approx_trans approx_minus_iff [THEN iffD2]) done text{*Now Sandard proof*} lemma DERIV_isCont: "DERIV f x :> D ==> isCont f x" by (simp add: NSDERIV_DERIV_iff [symmetric] isNSCont_isCont_iff [symmetric] NSDERIV_isNSCont) text{*Differentiation rules for combinations of functions follow from clear, straightforard, algebraic manipulations*} text{*Constant function*} (* use simple constant nslimit theorem *) lemma NSDERIV_const [simp]: "(NSDERIV (%x. k) x :> 0)" by (simp add: NSDERIV_NSLIM_iff) lemma DERIV_const [simp]: "(DERIV (%x. k) x :> 0)" by (simp add: NSDERIV_DERIV_iff [symmetric]) text{*Sum of functions- proved easily*} lemma NSDERIV_add: "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] ==> NSDERIV (%x. f x + g x) x :> Da + Db" apply (auto simp add: NSDERIV_NSLIM_iff NSLIM_def) apply (auto simp add: add_divide_distrib dest!: spec) apply (drule_tac b = "hypreal_of_real Da" and d = "hypreal_of_real Db" in approx_add) apply (auto simp add: add_ac) done (* Standard theorem *) lemma DERIV_add: "[| DERIV f x :> Da; DERIV g x :> Db |] ==> DERIV (%x. f x + g x) x :> Da + Db" apply (simp add: NSDERIV_add NSDERIV_DERIV_iff [symmetric]) done text{*Product of functions - Proof is trivial but tedious and long due to rearrangement of terms*} lemma lemma_nsderiv1: "((a::hypreal)*b) + -(c*d) = (b*(a + -c)) + (c*(b + -d))" by (simp add: right_distrib) lemma lemma_nsderiv2: "[| (x + y) / z = hypreal_of_real D + yb; z ≠ 0; z ∈ Infinitesimal; yb ∈ Infinitesimal |] ==> x + y ≈ 0" apply (frule_tac c1 = z in hypreal_mult_right_cancel [THEN iffD2], assumption) apply (erule_tac V = "(x + y) / z = hypreal_of_real D + yb" in thin_rl) apply (auto intro!: Infinitesimal_HFinite_mult2 HFinite_add simp add: mult_assoc mem_infmal_iff [symmetric]) apply (erule Infinitesimal_subset_HFinite [THEN subsetD]) done lemma NSDERIV_mult: "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] ==> NSDERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))" apply (auto simp add: NSDERIV_NSLIM_iff NSLIM_def) apply (auto dest!: spec simp add: starfun_lambda_cancel lemma_nsderiv1) apply (simp (no_asm) add: add_divide_distrib) apply (drule bex_Infinitesimal_iff2 [THEN iffD2])+ apply (auto simp add: times_divide_eq_right [symmetric] simp del: times_divide_eq) apply (drule_tac D = Db in lemma_nsderiv2) apply (drule_tac [4] approx_minus_iff [THEN iffD2, THEN bex_Infinitesimal_iff2 [THEN iffD2]]) apply (auto intro!: approx_add_mono1 simp add: left_distrib right_distrib mult_commute add_assoc) apply (rule_tac b1 = "hypreal_of_real Db * hypreal_of_real (f x)" in add_commute [THEN subst]) apply (auto intro!: Infinitesimal_add_approx_self2 [THEN approx_sym] Infinitesimal_add Infinitesimal_mult Infinitesimal_hypreal_of_real_mult Infinitesimal_hypreal_of_real_mult2 simp add: add_assoc [symmetric]) done lemma DERIV_mult: "[| DERIV f x :> Da; DERIV g x :> Db |] ==> DERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))" by (simp add: NSDERIV_mult NSDERIV_DERIV_iff [symmetric]) text{*Multiplying by a constant*} lemma NSDERIV_cmult: "NSDERIV f x :> D ==> NSDERIV (%x. c * f x) x :> c*D" apply (simp only: times_divide_eq_right [symmetric] NSDERIV_NSLIM_iff minus_mult_right right_distrib [symmetric]) apply (erule NSLIM_const [THEN NSLIM_mult]) done (* let's do the standard proof though theorem *) (* LIM_mult2 follows from a NS proof *) lemma DERIV_cmult: "DERIV f x :> D ==> DERIV (%x. c * f x) x :> c*D" apply (simp only: deriv_def times_divide_eq_right [symmetric] NSDERIV_NSLIM_iff minus_mult_right right_distrib [symmetric]) apply (erule LIM_const [THEN LIM_mult2]) done text{*Negation of function*} lemma NSDERIV_minus: "NSDERIV f x :> D ==> NSDERIV (%x. -(f x)) x :> -D" proof (simp add: NSDERIV_NSLIM_iff) assume "(λh. (f (x + h) + - f x) / h) -- 0 --NS> D" hence deriv: "(λh. - ((f(x+h) + - f x) / h)) -- 0 --NS> - D" by (rule NSLIM_minus) have "∀h. - ((f (x + h) + - f x) / h) = (- f (x + h) + f x) / h" by (simp add: minus_divide_left) with deriv show "(λh. (- f (x + h) + f x) / h) -- 0 --NS> - D" by simp qed lemma DERIV_minus: "DERIV f x :> D ==> DERIV (%x. -(f x)) x :> -D" by (simp add: NSDERIV_minus NSDERIV_DERIV_iff [symmetric]) text{*Subtraction*} lemma NSDERIV_add_minus: "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] ==> NSDERIV (%x. f x + -g x) x :> Da + -Db" by (blast dest: NSDERIV_add NSDERIV_minus) lemma DERIV_add_minus: "[| DERIV f x :> Da; DERIV g x :> Db |] ==> DERIV (%x. f x + -g x) x :> Da + -Db" by (blast dest: DERIV_add DERIV_minus) lemma NSDERIV_diff: "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] ==> NSDERIV (%x. f x - g x) x :> Da-Db" apply (simp add: diff_minus) apply (blast intro: NSDERIV_add_minus) done lemma DERIV_diff: "[| DERIV f x :> Da; DERIV g x :> Db |] ==> DERIV (%x. f x - g x) x :> Da-Db" apply (simp add: diff_minus) apply (blast intro: DERIV_add_minus) done text{*(NS) Increment*} lemma incrementI: "f NSdifferentiable x ==> increment f x h = ( *f* f) (hypreal_of_real(x) + h) + -hypreal_of_real (f x)" by (simp add: increment_def) lemma incrementI2: "NSDERIV f x :> D ==> increment f x h = ( *f* f) (hypreal_of_real(x) + h) + -hypreal_of_real (f x)" apply (erule NSdifferentiableI [THEN incrementI]) done (* The Increment theorem -- Keisler p. 65 *) lemma increment_thm: "[| NSDERIV f x :> D; h ∈ Infinitesimal; h ≠ 0 |] ==> ∃e ∈ Infinitesimal. increment f x h = hypreal_of_real(D)*h + e*h" apply (frule_tac h = h in incrementI2, simp add: nsderiv_def) apply (drule bspec, auto) apply (drule bex_Infinitesimal_iff2 [THEN iffD2], clarify) apply (frule_tac b1 = "hypreal_of_real (D) + y" in hypreal_mult_right_cancel [THEN iffD2]) apply (erule_tac [2] V = "(( *f* f) (hypreal_of_real (x) + h) + - hypreal_of_real (f x)) / h = hypreal_of_real (D) + y" in thin_rl) apply assumption apply (simp add: times_divide_eq_right [symmetric]) apply (auto simp add: left_distrib) done lemma increment_thm2: "[| NSDERIV f x :> D; h ≈ 0; h ≠ 0 |] ==> ∃e ∈ Infinitesimal. increment f x h = hypreal_of_real(D)*h + e*h" by (blast dest!: mem_infmal_iff [THEN iffD2] intro!: increment_thm) lemma increment_approx_zero: "[| NSDERIV f x :> D; h ≈ 0; h ≠ 0 |] ==> increment f x h ≈ 0" apply (drule increment_thm2, auto intro!: Infinitesimal_HFinite_mult2 HFinite_add simp add: left_distrib [symmetric] mem_infmal_iff [symmetric]) apply (erule Infinitesimal_subset_HFinite [THEN subsetD]) done text{* Similarly to the above, the chain rule admits an entirely straightforward derivation. Compare this with Harrison's HOL proof of the chain rule, which proved to be trickier and required an alternative characterisation of differentiability- the so-called Carathedory derivative. Our main problem is manipulation of terms.*} (* lemmas *) lemma NSDERIV_zero: "[| NSDERIV g x :> D; ( *f* g) (hypreal_of_real(x) + xa) = hypreal_of_real(g x); xa ∈ Infinitesimal; xa ≠ 0 |] ==> D = 0" apply (simp add: nsderiv_def) apply (drule bspec, auto) done (* can be proved differently using NSLIM_isCont_iff *) lemma NSDERIV_approx: "[| NSDERIV f x :> D; h ∈ Infinitesimal; h ≠ 0 |] ==> ( *f* f) (hypreal_of_real(x) + h) + -hypreal_of_real(f x) ≈ 0" apply (simp add: nsderiv_def) apply (simp add: mem_infmal_iff [symmetric]) apply (rule Infinitesimal_ratio) apply (rule_tac [3] approx_hypreal_of_real_HFinite, auto) done (*--------------------------------------------------------------- from one version of differentiability f(x) - f(a) --------------- ≈ Db x - a ---------------------------------------------------------------*) lemma NSDERIVD1: "[| NSDERIV f (g x) :> Da; ( *f* g) (hypreal_of_real(x) + xa) ≠ hypreal_of_real (g x); ( *f* g) (hypreal_of_real(x) + xa) ≈ hypreal_of_real (g x) |] ==> (( *f* f) (( *f* g) (hypreal_of_real(x) + xa)) + - hypreal_of_real (f (g x))) / (( *f* g) (hypreal_of_real(x) + xa) + - hypreal_of_real (g x)) ≈ hypreal_of_real(Da)" by (auto simp add: NSDERIV_NSLIM_iff2 NSLIM_def diff_minus [symmetric]) (*-------------------------------------------------------------- from other version of differentiability f(x + h) - f(x) ----------------- ≈ Db h --------------------------------------------------------------*) lemma NSDERIVD2: "[| NSDERIV g x :> Db; xa ∈ Infinitesimal; xa ≠ 0 |] ==> (( *f* g) (hypreal_of_real(x) + xa) + - hypreal_of_real(g x)) / xa ≈ hypreal_of_real(Db)" by (auto simp add: NSDERIV_NSLIM_iff NSLIM_def mem_infmal_iff starfun_lambda_cancel) lemma lemma_chain: "(z::hypreal) ≠ 0 ==> x*y = (x*inverse(z))*(z*y)" by auto text{*This proof uses both definitions of differentiability.*} lemma NSDERIV_chain: "[| NSDERIV f (g x) :> Da; NSDERIV g x :> Db |] ==> NSDERIV (f o g) x :> Da * Db" apply (simp (no_asm_simp) add: NSDERIV_NSLIM_iff NSLIM_def mem_infmal_iff [symmetric]) apply clarify apply (frule_tac f = g in NSDERIV_approx) apply (auto simp add: starfun_lambda_cancel2 starfun_o [symmetric]) apply (case_tac "( *f* g) (hypreal_of_real (x) + xa) = hypreal_of_real (g x) ") apply (drule_tac g = g in NSDERIV_zero) apply (auto simp add: divide_inverse) apply (rule_tac z1 = "( *f* g) (hypreal_of_real (x) + xa) + -hypreal_of_real (g x) " and y1 = "inverse xa" in lemma_chain [THEN ssubst]) apply (erule hypreal_not_eq_minus_iff [THEN iffD1]) apply (rule approx_mult_hypreal_of_real) apply (simp_all add: divide_inverse [symmetric]) apply (blast intro: NSDERIVD1 approx_minus_iff [THEN iffD2]) apply (blast intro: NSDERIVD2) done (* standard version *) lemma DERIV_chain: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (f o g) x :> Da * Db" by (simp add: NSDERIV_DERIV_iff [symmetric] NSDERIV_chain) lemma DERIV_chain2: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (%x. f (g x)) x :> Da * Db" by (auto dest: DERIV_chain simp add: o_def) text{*Differentiation of natural number powers*} lemma NSDERIV_Id [simp]: "NSDERIV (%x. x) x :> 1" by (simp add: NSDERIV_NSLIM_iff NSLIM_def divide_self del: divide_self_if) (*derivative of the identity function*) lemma DERIV_Id [simp]: "DERIV (%x. x) x :> 1" by (simp add: NSDERIV_DERIV_iff [symmetric]) lemmas isCont_Id = DERIV_Id [THEN DERIV_isCont, standard] (*derivative of linear multiplication*) lemma DERIV_cmult_Id [simp]: "DERIV (op * c) x :> c" by (cut_tac c = c and x = x in DERIV_Id [THEN DERIV_cmult], simp) lemma NSDERIV_cmult_Id [simp]: "NSDERIV (op * c) x :> c" by (simp add: NSDERIV_DERIV_iff) lemma DERIV_pow: "DERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))" apply (induct "n") apply (drule_tac [2] DERIV_Id [THEN DERIV_mult]) apply (auto simp add: real_of_nat_Suc left_distrib) apply (case_tac "0 < n") apply (drule_tac x = x in realpow_minus_mult) apply (auto simp add: mult_assoc add_commute) done (* NS version *) lemma NSDERIV_pow: "NSDERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))" by (simp add: NSDERIV_DERIV_iff DERIV_pow) text{*Power of -1*} (*Can't get rid of x ≠ 0 because it isn't continuous at zero*) lemma NSDERIV_inverse: "x ≠ 0 ==> NSDERIV (%x. inverse(x)) x :> (- (inverse x ^ Suc (Suc 0)))" apply (simp add: nsderiv_def) apply (rule ballI, simp, clarify) apply (frule Infinitesimal_add_not_zero) prefer 2 apply (simp add: add_commute) apply (auto simp add: starfun_inverse_inverse realpow_two simp del: minus_mult_left [symmetric] minus_mult_right [symmetric]) apply (simp add: inverse_add inverse_mult_distrib [symmetric] inverse_minus_eq [symmetric] add_ac mult_ac del: inverse_mult_distrib inverse_minus_eq minus_mult_left [symmetric] minus_mult_right [symmetric]) apply (simp (no_asm_simp) add: mult_assoc [symmetric] right_distrib del: minus_mult_left [symmetric] minus_mult_right [symmetric]) apply (rule_tac y = "inverse (- hypreal_of_real x * hypreal_of_real x)" in approx_trans) apply (rule inverse_add_Infinitesimal_approx2) apply (auto dest!: hypreal_of_real_HFinite_diff_Infinitesimal simp add: inverse_minus_eq [symmetric] HFinite_minus_iff) apply (rule Infinitesimal_HFinite_mult2, auto) done lemma DERIV_inverse: "x ≠ 0 ==> DERIV (%x. inverse(x)) x :> (-(inverse x ^ Suc (Suc 0)))" by (simp add: NSDERIV_inverse NSDERIV_DERIV_iff [symmetric] del: realpow_Suc) text{*Derivative of inverse*} lemma DERIV_inverse_fun: "[| DERIV f x :> d; f(x) ≠ 0 |] ==> DERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))" apply (simp only: mult_commute [of d] minus_mult_left power_inverse) apply (fold o_def) apply (blast intro!: DERIV_chain DERIV_inverse) done lemma NSDERIV_inverse_fun: "[| NSDERIV f x :> d; f(x) ≠ 0 |] ==> NSDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))" by (simp add: NSDERIV_DERIV_iff DERIV_inverse_fun del: realpow_Suc) text{*Derivative of quotient*} lemma DERIV_quotient: "[| DERIV f x :> d; DERIV g x :> e; g(x) ≠ 0 |] ==> DERIV (%y. f(y) / (g y)) x :> (d*g(x) + -(e*f(x))) / (g(x) ^ Suc (Suc 0))" apply (drule_tac f = g in DERIV_inverse_fun) apply (drule_tac [2] DERIV_mult) apply (assumption+) apply (simp add: divide_inverse right_distrib power_inverse minus_mult_left mult_ac del: realpow_Suc minus_mult_right [symmetric] minus_mult_left [symmetric]) done lemma NSDERIV_quotient: "[| NSDERIV f x :> d; DERIV g x :> e; g(x) ≠ 0 |] ==> NSDERIV (%y. f(y) / (g y)) x :> (d*g(x) + -(e*f(x))) / (g(x) ^ Suc (Suc 0))" by (simp add: NSDERIV_DERIV_iff DERIV_quotient del: realpow_Suc) (* ------------------------------------------------------------------------ *) (* Caratheodory formulation of derivative at a point: standard proof *) (* ------------------------------------------------------------------------ *) lemma CARAT_DERIV: "(DERIV f x :> l) = (∃g. (∀z. f z - f x = g z * (z-x)) & isCont g x & g x = l)" (is "?lhs = ?rhs") proof assume der: "DERIV f x :> l" show "∃g. (∀z. f z - f x = g z * (z-x)) ∧ isCont g x ∧ g x = l" proof (intro exI conjI) let ?g = "(%z. if z = x then l else (f z - f x) / (z-x))" show "∀z. f z - f x = ?g z * (z-x)" by (simp) show "isCont ?g x" using der by (simp add: isCont_iff DERIV_iff diff_minus cong: LIM_equal [rule_format]) show "?g x = l" by simp qed next assume "?rhs" then obtain g where "(∀z. f z - f x = g z * (z-x))" and "isCont g x" and "g x = l" by blast thus "(DERIV f x :> l)" by (auto simp add: isCont_iff DERIV_iff diff_minus cong: LIM_equal [rule_format]) qed lemma CARAT_NSDERIV: "NSDERIV f x :> l ==> ∃g. (∀z. f z - f x = g z * (z-x)) & isNSCont g x & g x = l" by (auto simp add: NSDERIV_DERIV_iff isNSCont_isCont_iff CARAT_DERIV) lemma hypreal_eq_minus_iff3: "(x = y + z) = (x + -z = (y::hypreal))" by auto lemma CARAT_DERIVD: assumes all: "∀z. f z - f x = g z * (z-x)" and nsc: "isNSCont g x" shows "NSDERIV f x :> g x" proof - from nsc have "∀w. w ≠ hypreal_of_real x ∧ w ≈ hypreal_of_real x --> ( *f* g) w * (w - hypreal_of_real x) / (w - hypreal_of_real x) ≈ hypreal_of_real (g x)" by (simp add: diff_minus isNSCont_def) thus ?thesis using all by (simp add: NSDERIV_iff2 starfun_if_eq cong: if_cong) qed text{*Lemmas about nested intervals and proof by bisection (cf.Harrison). All considerably tidied by lcp.*} lemma lemma_f_mono_add [rule_format (no_asm)]: "(∀n. (f::nat=>real) n ≤ f (Suc n)) --> f m ≤ f(m + no)" apply (induct "no") apply (auto intro: order_trans) done lemma f_inc_g_dec_Beq_f: "[| ∀n. f(n) ≤ f(Suc n); ∀n. g(Suc n) ≤ g(n); ∀n. f(n) ≤ g(n) |] ==> Bseq f" apply (rule_tac k = "f 0" and K = "g 0" in BseqI2, rule allI) apply (induct_tac "n") apply (auto intro: order_trans) apply (rule_tac y = "g (Suc na)" in order_trans) apply (induct_tac [2] "na") apply (auto intro: order_trans) done lemma f_inc_g_dec_Beq_g: "[| ∀n. f(n) ≤ f(Suc n); ∀n. g(Suc n) ≤ g(n); ∀n. f(n) ≤ g(n) |] ==> Bseq g" apply (subst Bseq_minus_iff [symmetric]) apply (rule_tac g = "%x. - (f x)" in f_inc_g_dec_Beq_f) apply auto done lemma f_inc_imp_le_lim: "[| ∀n. f n ≤ f (Suc n); convergent f |] ==> f n ≤ lim f" apply (rule linorder_not_less [THEN iffD1]) apply (auto simp add: convergent_LIMSEQ_iff LIMSEQ_iff monoseq_Suc) apply (drule real_less_sum_gt_zero) apply (drule_tac x = "f n + - lim f" in spec, safe) apply (drule_tac P = "%na. no≤na --> ?Q na" and x = "no + n" in spec, auto) apply (subgoal_tac "lim f ≤ f (no + n) ") apply (induct_tac [2] "no") apply (auto intro: order_trans simp add: diff_minus abs_if) apply (drule_tac no=no and m=n in lemma_f_mono_add) apply (auto simp add: add_commute) done lemma lim_uminus: "convergent g ==> lim (%x. - g x) = - (lim g)" apply (rule LIMSEQ_minus [THEN limI]) apply (simp add: convergent_LIMSEQ_iff) done lemma g_dec_imp_lim_le: "[| ∀n. g(Suc n) ≤ g(n); convergent g |] ==> lim g ≤ g n" apply (subgoal_tac "- (g n) ≤ - (lim g) ") apply (cut_tac [2] f = "%x. - (g x)" in f_inc_imp_le_lim) apply (auto simp add: lim_uminus convergent_minus_iff [symmetric]) done lemma lemma_nest: "[| ∀n. f(n) ≤ f(Suc n); ∀n. g(Suc n) ≤ g(n); ∀n. f(n) ≤ g(n) |] ==> ∃l m. l ≤ m & ((∀n. f(n) ≤ l) & f ----> l) & ((∀n. m ≤ g(n)) & g ----> m)" apply (subgoal_tac "monoseq f & monoseq g") prefer 2 apply (force simp add: LIMSEQ_iff monoseq_Suc) apply (subgoal_tac "Bseq f & Bseq g") prefer 2 apply (blast intro: f_inc_g_dec_Beq_f f_inc_g_dec_Beq_g) apply (auto dest!: Bseq_monoseq_convergent simp add: convergent_LIMSEQ_iff) apply (rule_tac x = "lim f" in exI) apply (rule_tac x = "lim g" in exI) apply (auto intro: LIMSEQ_le) apply (auto simp add: f_inc_imp_le_lim g_dec_imp_lim_le convergent_LIMSEQ_iff) done lemma lemma_nest_unique: "[| ∀n. f(n) ≤ f(Suc n); ∀n. g(Suc n) ≤ g(n); ∀n. f(n) ≤ g(n); (%n. f(n) - g(n)) ----> 0 |] ==> ∃l. ((∀n. f(n) ≤ l) & f ----> l) & ((∀n. l ≤ g(n)) & g ----> l)" apply (drule lemma_nest, auto) apply (subgoal_tac "l = m") apply (drule_tac [2] X = f in LIMSEQ_diff) apply (auto intro: LIMSEQ_unique) done text{*The universal quantifiers below are required for the declaration of @{text Bolzano_nest_unique} below.*} lemma Bolzano_bisect_le: "a ≤ b ==> ∀n. fst (Bolzano_bisect P a b n) ≤ snd (Bolzano_bisect P a b n)" apply (rule allI) apply (induct_tac "n") apply (auto simp add: Let_def split_def) done lemma Bolzano_bisect_fst_le_Suc: "a ≤ b ==> ∀n. fst(Bolzano_bisect P a b n) ≤ fst (Bolzano_bisect P a b (Suc n))" apply (rule allI) apply (induct_tac "n") apply (auto simp add: Bolzano_bisect_le Let_def split_def) done lemma Bolzano_bisect_Suc_le_snd: "a ≤ b ==> ∀n. snd(Bolzano_bisect P a b (Suc n)) ≤ snd (Bolzano_bisect P a b n)" apply (rule allI) apply (induct_tac "n") apply (auto simp add: Bolzano_bisect_le Let_def split_def) done lemma eq_divide_2_times_iff: "((x::real) = y / (2 * z)) = (2 * x = y/z)" apply (auto) apply (drule_tac f = "%u. (1/2) *u" in arg_cong) apply (simp) done lemma Bolzano_bisect_diff: "a ≤ b ==> snd(Bolzano_bisect P a b n) - fst(Bolzano_bisect P a b n) = (b-a) / (2 ^ n)" apply (induct "n") apply (auto simp add: eq_divide_2_times_iff add_divide_distrib Let_def split_def) done lemmas Bolzano_nest_unique = lemma_nest_unique [OF Bolzano_bisect_fst_le_Suc Bolzano_bisect_Suc_le_snd Bolzano_bisect_le] lemma not_P_Bolzano_bisect: assumes P: "!!a b c. [| P(a,b); P(b,c); a ≤ b; b ≤ c|] ==> P(a,c)" and notP: "~ P(a,b)" and le: "a ≤ b" shows "~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))" proof (induct n) case 0 thus ?case by simp next case (Suc n) thus ?case by (auto simp del: surjective_pairing [symmetric] simp add: Let_def split_def Bolzano_bisect_le [OF le] P [of "fst (Bolzano_bisect P a b n)" _ "snd (Bolzano_bisect P a b n)"]) qed (*Now we re-package P_prem as a formula*) lemma not_P_Bolzano_bisect': "[| ∀a b c. P(a,b) & P(b,c) & a ≤ b & b ≤ c --> P(a,c); ~ P(a,b); a ≤ b |] ==> ∀n. ~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))" by (blast elim!: not_P_Bolzano_bisect [THEN [2] rev_notE]) lemma lemma_BOLZANO: "[| ∀a b c. P(a,b) & P(b,c) & a ≤ b & b ≤ c --> P(a,c); ∀x. ∃d::real. 0 < d & (∀a b. a ≤ x & x ≤ b & (b-a) < d --> P(a,b)); a ≤ b |] ==> P(a,b)" apply (rule Bolzano_nest_unique [where P1=P, THEN exE], assumption+) apply (rule LIMSEQ_minus_cancel) apply (simp (no_asm_simp) add: Bolzano_bisect_diff LIMSEQ_divide_realpow_zero) apply (rule ccontr) apply (drule not_P_Bolzano_bisect', assumption+) apply (rename_tac "l") apply (drule_tac x = l in spec, clarify) apply (simp add: LIMSEQ_def) apply (drule_tac P = "%r. 0<r --> ?Q r" and x = "d/2" in spec) apply (drule_tac P = "%r. 0<r --> ?Q r" and x = "d/2" in spec) apply (drule real_less_half_sum, auto) apply (drule_tac x = "fst (Bolzano_bisect P a b (no + noa))" in spec) apply (drule_tac x = "snd (Bolzano_bisect P a b (no + noa))" in spec) apply safe apply (simp_all (no_asm_simp)) apply (rule_tac y = "abs (fst (Bolzano_bisect P a b (no + noa)) - l) + abs (snd (Bolzano_bisect P a b (no + noa)) - l)" in order_le_less_trans) apply (simp (no_asm_simp) add: abs_if) apply (rule real_sum_of_halves [THEN subst]) apply (rule add_strict_mono) apply (simp_all add: diff_minus [symmetric]) done lemma lemma_BOLZANO2: "((∀a b c. (a ≤ b & b ≤ c & P(a,b) & P(b,c)) --> P(a,c)) & (∀x. ∃d::real. 0 < d & (∀a b. a ≤ x & x ≤ b & (b-a) < d --> P(a,b)))) --> (∀a b. a ≤ b --> P(a,b))" apply clarify apply (blast intro: lemma_BOLZANO) done subsection{*Intermediate Value Theorem: Prove Contrapositive by Bisection*} lemma IVT: "[| f(a) ≤ y; y ≤ f(b); a ≤ b; (∀x. a ≤ x & x ≤ b --> isCont f x) |] ==> ∃x. a ≤ x & x ≤ b & f(x) = y" apply (rule contrapos_pp, assumption) apply (cut_tac P = "% (u,v) . a ≤ u & u ≤ v & v ≤ b --> ~ (f (u) ≤ y & y ≤ f (v))" in lemma_BOLZANO2) apply safe apply simp_all apply (simp add: isCont_iff LIM_def) apply (rule ccontr) apply (subgoal_tac "a ≤ x & x ≤ b") prefer 2 apply simp apply (drule_tac P = "%d. 0<d --> ?P d" and x = 1 in spec, arith) apply (drule_tac x = x in spec)+ apply simp apply (drule_tac P = "%r. ?P r --> (∃s>0. ?Q r s) " and x = "¦y - f x¦" in spec) apply safe apply simp apply (drule_tac x = s in spec, clarify) apply (cut_tac x = "f x" and y = y in linorder_less_linear, safe) apply (drule_tac x = "ba-x" in spec) apply (simp_all add: abs_if) apply (drule_tac x = "aa-x" in spec) apply (case_tac "x ≤ aa", simp_all) done lemma IVT2: "[| f(b) ≤ y; y ≤ f(a); a ≤ b; (∀x. a ≤ x & x ≤ b --> isCont f x) |] ==> ∃x. a ≤ x & x ≤ b & f(x) = y" apply (subgoal_tac "- f a ≤ -y & -y ≤ - f b", clarify) apply (drule IVT [where f = "%x. - f x"], assumption) apply (auto intro: isCont_minus) done (*HOL style here: object-level formulations*) lemma IVT_objl: "(f(a) ≤ y & 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 & 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) done subsection{*By bisection, function continuous on closed interval is bounded above*} lemma isCont_bounded: "[| a ≤ b; ∀x. a ≤ x & x ≤ b --> isCont f x |] ==> ∃M. ∀x. 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 apply (rename_tac x xa ya M Ma) apply (cut_tac x = M and y = Ma in linorder_linear, safe) apply (rule_tac x = Ma in exI, clarify) apply (cut_tac x = xb and y = xa in linorder_linear, force) apply (rule_tac x = M in exI, clarify) apply (cut_tac x = xb and y = xa in linorder_linear, force) apply (case_tac "a ≤ x & x ≤ b") apply (rule_tac [2] x = 1 in exI) prefer 2 apply force apply (simp add: LIM_def isCont_iff) apply (drule_tac x = x in spec, auto) apply (erule_tac V = "∀M. ∃x. a ≤ x & x ≤ b & ~ f x ≤ M" in thin_rl) apply (drule_tac x = 1 in spec, auto) apply (rule_tac x = s in exI, clarify) apply (rule_tac x = "¦f x¦ + 1" in exI, clarify) apply (drule_tac x = "xa-x" in spec) apply (auto simp add: abs_ge_self, arith+) done text{*Refine the above to existence of least upper bound*} lemma lemma_reals_complete: "((∃x. x ∈ S) & (∃y. isUb UNIV S (y::real))) --> (∃t. isLub UNIV S t)" by (blast intro: reals_complete) lemma isCont_has_Ub: "[| a ≤ b; ∀x. a ≤ x & x ≤ b --> isCont f x |] ==> ∃M. (∀x. 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) apply auto apply (drule isCont_bounded, assumption) apply (auto simp add: isUb_def leastP_def isLub_def setge_def setle_def) apply (rule exI, auto) apply (auto dest!: spec simp add: linorder_not_less) done text{*Now show that it attains its upper bound*} lemma isCont_eq_Ub: assumes le: "a ≤ b" and con: "∀x. a ≤ x & x ≤ b --> isCont f x" shows "∃M. (∀x. a ≤ x & x ≤ b --> f(x) ≤ M) & (∃x. a ≤ x & x ≤ b & f(x) = M)" proof - from isCont_has_Ub [OF le con] obtain M where M1: "∀x. a ≤ x ∧ x ≤ b --> f x ≤ M" and M2: "!!N. N<M ==> ∃x. a ≤ x ∧ x ≤ b ∧ N < f x" by blast show ?thesis proof (intro exI, intro conjI) show " ∀x. a ≤ x ∧ x ≤ b --> f x ≤ M" by (rule M1) show "∃x. a ≤ x ∧ x ≤ b ∧ f x = M" proof (rule ccontr) assume "¬ (∃x. a ≤ x ∧ x ≤ b ∧ f x = M)" with M1 have M3: "∀x. a ≤ x & x ≤ b --> f x < M" by (fastsimp simp add: linorder_not_le [symmetric]) hence "∀x. a ≤ x & x ≤ b --> isCont (%x. inverse (M - f x)) x" by (auto simp add: isCont_inverse isCont_diff con) from isCont_bounded [OF le this] obtain k where k: "!!x. a ≤ x & x ≤ b --> inverse (M - f x) ≤ k" by auto have Minv: "!!x. a ≤ x & x ≤ b --> 0 < inverse (M - f (x))" by (simp add: M3 compare_rls) have "!!x. a ≤ x & x ≤ b --> inverse (M - f x) < k+1" using k by (auto intro: order_le_less_trans [of _ k]) with Minv have "!!x. a ≤ x & x ≤ b --> inverse(k+1) < inverse(inverse(M - f x))" by (intro strip less_imp_inverse_less, simp_all) hence invlt: "!!x. a ≤ x & x ≤ b --> inverse(k+1) < M - f x" by simp have "M - inverse (k+1) < M" using k [of a] Minv [of a] le by (simp, arith) from M2 [OF this] obtain x where ax: "a ≤ x & x ≤ b & M - inverse(k+1) < f x" .. thus False using invlt [of x] by force qed qed qed text{*Same theorem for lower bound*} lemma isCont_eq_Lb: "[| a ≤ b; ∀x. a ≤ x & x ≤ b --> isCont f x |] ==> ∃M. (∀x. 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) apply (drule_tac f = "(%x. - (f x))" in isCont_eq_Ub) apply safe apply auto done text{*Another version.*} lemma isCont_Lb_Ub: "[|a ≤ b; ∀x. a ≤ x & x ≤ b --> isCont f x |] ==> ∃L M. (∀x. 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) apply (assumption+, safe) apply (rule_tac x = "f x" in exI) apply (rule_tac x = "f xa" in exI, simp, safe) apply (cut_tac x = x and y = xa in linorder_linear, safe) apply (cut_tac f = f and a = x and b = xa and y = y in IVT_objl) apply (cut_tac [2] f = f and a = xa and b = x and y = y in IVT2_objl, safe) apply (rule_tac [2] x = xb in exI) apply (rule_tac [4] x = xb in exI, simp_all) done subsection{*If @{term "0 < f'(x)"} then @{term x} is Locally Strictly Increasing At The Right*} lemma DERIV_left_inc: assumes der: "DERIV f x :> l" and l: "0 < l" shows "∃d > 0. ∀h > 0. h < d --> f(x) < f(x + h)" proof - from l der [THEN DERIV_D, THEN LIM_D [where r = "l"]] have "∃s > 0. (∀z. z ≠ 0 ∧ ¦z¦ < s --> ¦(f(x+z) - f x) / z - l¦ < l)" by (simp add: diff_minus) then obtain s where s: "0 < s" and all: "!!z. z ≠ 0 ∧ ¦z¦ < s --> ¦(f(x+z) - f x) / z - l¦ < l" by auto thus ?thesis proof (intro exI conjI strip) show "0<s" . fix h::real assume "0 < h" "h < s" with all [of h] show "f x < f (x+h)" proof (simp add: abs_if pos_less_divide_eq diff_minus [symmetric] split add: split_if_asm) assume "~ (f (x+h) - f x) / h < l" and h: "0 < h" with l have "0 < (f (x+h) - f x) / h" by arith thus "f x < f (x+h)" by (simp add: pos_less_divide_eq h) qed qed qed lemma DERIV_left_dec: assumes der: "DERIV f x :> l" and l: "l < 0" shows "∃d > 0. ∀h > 0. h < d --> f(x) < f(x-h)" proof - from l der [THEN DERIV_D, THEN LIM_D [where r = "-l"]] have "∃s > 0. (∀z. z ≠ 0 ∧ ¦z¦ < s --> ¦(f(x+z) - f x) / z - l¦ < -l)" by (simp add: diff_minus) then obtain s where s: "0 < s" and all: "!!z. z ≠ 0 ∧ ¦z¦ < s --> ¦(f(x+z) - f x) / z - l¦ < -l" by auto thus ?thesis proof (intro exI conjI strip) show "0<s" . fix h::real assume "0 < h" "h < s" with all [of "-h"] show "f x < f (x-h)" proof (simp add: abs_if pos_less_divide_eq diff_minus [symmetric] split add: split_if_asm) assume " - ((f (x-h) - f x) / h) < l" and h: "0 < h" with l have "0 < (f (x-h) - f x) / h" by arith thus "f x < f (x-h)" by (simp add: pos_less_divide_eq h) qed qed qed lemma DERIV_local_max: assumes der: "DERIV f x :> l" and d: "0 < d" and le: "∀y. ¦x-y¦ < d --> f(y) ≤ f(x)" shows "l = 0" proof (cases rule: linorder_cases [of l 0]) case equal show ?thesis . next case less from DERIV_left_dec [OF der less] obtain d' where d': "0 < d'" and lt: "∀h > 0. h < d' --> f x < f (x-h)" by blast from real_lbound_gt_zero [OF d d'] obtain e where "0 < e ∧ e < d ∧ e < d'" .. with lt le [THEN spec [where x="x-e"]] show ?thesis by (auto simp add: abs_if) next case greater from DERIV_left_inc [OF der greater] obtain d' where d': "0 < d'" and lt: "∀h > 0. h < d' --> f x < f (x + h)" by blast from real_lbound_gt_zero [OF d d'] obtain e where "0 < e ∧ e < d ∧ e < d'" .. with lt le [THEN spec [where x="x+e"]] show ?thesis by (auto simp add: abs_if) qed text{*Similar theorem for a local minimum*} lemma DERIV_local_min: "[| DERIV f x :> l; 0 < d; ∀y. ¦x-y¦ < d --> f(x) ≤ f(y) |] ==> l = 0" by (drule DERIV_minus [THEN DERIV_local_max], auto) text{*In particular, if a function is locally flat*} lemma DERIV_local_const: "[| DERIV f x :> l; 0 < d; ∀y. ¦x-y¦ < d --> f(x) = f(y) |] ==> l = 0" by (auto dest!: DERIV_local_max) text{*Lemma about introducing open ball in open interval*} lemma lemma_interval_lt: "[| a < x; x < b |] ==> ∃d::real. 0 < d & (∀y. ¦x-y¦ < d --> a < y & y < b)" apply (simp add: abs_interval_iff) apply (insert linorder_linear [of "x-a" "b-x"], safe) apply (rule_tac x = "x-a" in exI) apply (rule_tac [2] x = "b-x" in exI, auto) done lemma lemma_interval: "[| a < x; x < b |] ==> ∃d::real. 0 < d & (∀y. ¦x-y¦ < d --> a ≤ y & y ≤ b)" apply (drule lemma_interval_lt, auto) apply (auto intro!: exI) done text{*Rolle's Theorem. If @{term f} is defined and continuous on the closed interval @{text "[a,b]"} and differentiable on the open interval @{text "(a,b)"}, and @{term "f(a) = f(b)"}, then there exists @{text "x0 ∈ (a,b)"} such that @{term "f'(x0) = 0"}*} theorem Rolle: assumes lt: "a < b" and eq: "f(a) = f(b)" and con: "∀x. a ≤ x & x ≤ b --> isCont f x" and dif [rule_format]: "∀x. a < x & x < b --> f differentiable x" shows "∃z. a < z & z < b & DERIV f z :> 0" proof - have le: "a ≤ b" using lt by simp from isCont_eq_Ub [OF le con] obtain x where x_max: "∀z. a ≤ z ∧ z ≤ b --> f z ≤ f x" and alex: "a ≤ x" and xleb: "x ≤ b" by blast from isCont_eq_Lb [OF le con] obtain x' where x'_min: "∀z. a ≤ z ∧ z ≤ b --> f x' ≤ f z" and alex': "a ≤ x'" and x'leb: "x' ≤ b" by blast show ?thesis proof cases assume axb: "a < x & x < b" --{*@{term f} attains its maximum within the interval*} hence ax: "a<x" and xb: "x<b" by auto from lemma_interval [OF ax xb] obtain d where d: "0<d" and bound: "∀y. ¦x-y¦ < d --> a ≤ y ∧ y ≤ b" by blast hence bound': "∀y. ¦x-y¦ < d --> f y ≤ f x" using x_max by blast from differentiableD [OF dif [OF axb]] obtain l where der: "DERIV f x :> l" .. have "l=0" by (rule DERIV_local_max [OF der d bound']) --{*the derivative at a local maximum is zero*} thus ?thesis using ax xb der by auto next assume notaxb: "~ (a < x & x < b)" hence xeqab: "x=a | x=b" using alex xleb by arith hence fb_eq_fx: "f b = f x" by (auto simp add: eq) show ?thesis proof cases assume ax'b: "a < x' & x' < b" --{*@{term f} attains its minimum within the interval*} hence ax': "a<x'" and x'b: "x'<b" by auto from lemma_interval [OF ax' x'b] obtain d where d: "0<d" and bound: "∀y. ¦x'-y¦ < d --> a ≤ y ∧ y ≤ b" by blast hence bound': "∀y. ¦x'-y¦ < d --> f x' ≤ f y" using x'_min by blast from differentiableD [OF dif [OF ax'b]] obtain l where der: "DERIV f x' :> l" .. have "l=0" by (rule DERIV_local_min [OF der d bound']) --{*the derivative at a local minimum is zero*} thus ?thesis using ax' x'b der by auto next assume notax'b: "~ (a < x' & x' < b)" --{*@{term f} is constant througout the interval*} hence x'eqab: "x'=a | x'=b" using alex' x'leb by arith hence fb_eq_fx': "f b = f x'" by (auto simp add: eq) from dense [OF lt] obtain r where ar: "a < r" and rb: "r < b" by blast from lemma_interval [OF ar rb] obtain d where d: "0<d" and bound: "∀y. ¦r-y¦ < d --> a ≤ y ∧ y ≤ b" by blast have eq_fb: "∀z. a ≤ z --> z ≤ b --> f z = f b" proof (clarify) fix z::real assume az: "a ≤ z" and zb: "z ≤ b" show "f z = f b" proof (rule order_antisym) show "f z ≤ f b" by (simp add: fb_eq_fx x_max az zb) show "f b ≤ f z" by (simp add: fb_eq_fx' x'_min az zb) qed qed have bound': "∀y. ¦r-y¦ < d --> f r = f y" proof (intro strip) fix y::real assume lt: "¦r-y¦ < d" hence "f y = f b" by (simp add: eq_fb bound) thus "f r = f y" by (simp add: eq_fb ar rb order_less_imp_le) qed from differentiableD [OF dif [OF conjI [OF ar rb]]] obtain l where der: "DERIV f r :> l" .. have "l=0" by (rule DERIV_local_const [OF der d bound']) --{*the derivative of a constant function is zero*} thus ?thesis using ar rb der by auto qed qed qed subsection{*Mean Value Theorem*} lemma lemma_MVT: "f a - (f b - f a)/(b-a) * a = f b - (f b - f a)/(b-a) * (b::real)" proof cases assume "a=b" thus ?thesis by simp next assume "a≠b" hence ba: "b-a ≠ 0" by arith show ?thesis by (rule real_mult_left_cancel [OF ba, THEN iffD1], simp add: right_diff_distrib, simp add: left_diff_distrib) qed theorem MVT: assumes lt: "a < b" and con: "∀x. a ≤ x & x ≤ b --> isCont f x" and dif [rule_format]: "∀x. a < x & x < b --> f differentiable x" shows "∃l z. a < z & z < b & DERIV f z :> l & (f(b) - f(a) = (b-a) * l)" proof - let ?F = "%x. f x - ((f b - f a) / (b-a)) * x" have contF: "∀x. a ≤ x ∧ x ≤ b --> isCont ?F x" using con by (fast intro: isCont_diff isCont_const isCont_mult isCont_Id) have difF: "∀x. a < x ∧ x < b --> ?F differentiable x" proof (clarify) fix x::real assume ax: "a < x" and xb: "x < b" from differentiableD [OF dif [OF conjI [OF ax xb]]] obtain l where der: "DERIV f x :> l" .. show "?F differentiable x" by (rule differentiableI [where D = "l - (f b - f a)/(b-a)"], blast intro: DERIV_diff DERIV_cmult_Id der) qed from Rolle [where f = ?F, OF lt lemma_MVT contF difF] obtain z where az: "a < z" and zb: "z < b" and der: "DERIV ?F z :> 0" by blast have "DERIV (%x. ((f b - f a)/(b-a)) * x) z :> (f b - f a)/(b-a)" by (rule DERIV_cmult_Id) hence derF: "DERIV (λx. ?F x + (f b - f a) / (b - a) * x) z :> 0 + (f b - f a) / (b - a)" by (rule DERIV_add [OF der]) show ?thesis proof (intro exI conjI) show "a < z" . show "z < b" . show "f b - f a = (b - a) * ((f b - f a)/(b-a))" by (simp) show "DERIV f z :> ((f b - f a)/(b-a))" using derF by simp qed qed text{*A function is constant if its derivative is 0 over an interval.*} lemma DERIV_isconst_end: "[| a < b; ∀x. a ≤ x & x ≤ b --> isCont f x; ∀x. a < x & x < b --> DERIV f x :> 0 |] ==> f b = f a" apply (drule MVT, assumption) apply (blast intro: differentiableI) apply (auto dest!: DERIV_unique simp add: diff_eq_eq) done lemma DERIV_isconst1: "[| a < b; ∀x. a ≤ x & x ≤ b --> isCont f x; ∀x. a < x & x < b --> DERIV f x :> 0 |] ==> ∀x. a ≤ x & x ≤ b --> f x = f a" apply safe apply (drule_tac x = a in order_le_imp_less_or_eq, safe) apply (drule_tac b = x in DERIV_isconst_end, auto) done lemma DERIV_isconst2: "[| a < b; ∀x. a ≤ x & x ≤ b --> isCont f x; ∀x. a < x & x < b --> DERIV f x :> 0; a ≤ x; x ≤ b |] ==> f x = f a" apply (blast dest: DERIV_isconst1) done lemma DERIV_isconst_all: "∀x. DERIV f x :> 0 ==> f(x) = f(y)" apply (rule linorder_cases [of x y]) apply (blast intro: sym DERIV_isCont DERIV_isconst_end)+ done lemma DERIV_const_ratio_const: "[|a ≠ b; ∀x. DERIV f x :> k |] ==> (f(b) - f(a)) = (b-a) * k" apply (rule linorder_cases [of a b], auto) apply (drule_tac [!] f = f in MVT) apply (auto dest: DERIV_isCont DERIV_unique simp add: differentiable_def) apply (auto dest: DERIV_unique simp add: left_distrib diff_minus) done lemma DERIV_const_ratio_const2: "[|a ≠ b; ∀x. DERIV f x :> k |] ==> (f(b) - f(a))/(b-a) = k" apply (rule_tac c1 = "b-a" in real_mult_right_cancel [THEN iffD1]) apply (auto dest!: DERIV_const_ratio_const simp add: mult_assoc) done lemma real_average_minus_first [simp]: "((a + b) /2 - a) = (b-a)/(2::real)" by (simp) lemma real_average_minus_second [simp]: "((b + a)/2 - a) = (b-a)/(2::real)" by (simp) text{*Gallileo's "trick": average velocity = av. of end velocities*} lemma DERIV_const_average: assumes neq: "a ≠ (b::real)" and der: "∀x. DERIV v x :> k" shows "v ((a + b)/2) = (v a + v b)/2" proof (cases rule: linorder_cases [of a b]) case equal with neq show ?thesis by simp next case less have "(v b - v a) / (b - a) = k" by (rule DERIV_const_ratio_const2 [OF neq der]) hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp moreover have "(v ((a + b) / 2) - v a) / ((a + b) / 2 - a) = k" by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq) ultimately show ?thesis using neq by force next case greater have "(v b - v a) / (b - a) = k" by (rule DERIV_const_ratio_const2 [OF neq der]) hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp moreover have " (v ((b + a) / 2) - v a) / ((b + a) / 2 - a) = k" by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq) ultimately show ?thesis using neq by (force simp add: add_commute) qed text{*Dull lemma: an continuous injection on an interval must have a strict maximum at an end point, not in the middle.*} lemma lemma_isCont_inj: assumes d: "0 < d" and inj [rule_format]: "∀z. ¦z-x¦ ≤ d --> g(f z) = z" and cont: "∀z. ¦z-x¦ ≤ d --> isCont f z" shows "∃z. ¦z-x¦ ≤ d & f x < f z" proof (rule ccontr) assume "~ (∃z. ¦z-x¦ ≤ d & f x < f z)" hence all [rule_format]: "∀z. ¦z - x¦ ≤ d --> f z ≤ f x" by auto show False proof (cases rule: linorder_le_cases [of "f(x-d)" "f(x+d)"]) case le from d cont all [of "x+d"] have flef: "f(x+d) ≤ f x" and xlex: "x - d ≤ x" and cont': "∀z. x - d ≤ z ∧ z ≤ x --> isCont f z" by (auto simp add: abs_if) from IVT [OF le flef xlex cont'] obtain x' where "x-d ≤ x'" "x' ≤ x" "f x' = f(x+d)" by blast moreover hence "g(f x') = g (f(x+d))" by simp ultimately show False using d inj [of x'] inj [of "x+d"] by (simp add: abs_le_interval_iff) next case ge from d cont all [of "x-d"] have flef: "f(x-d) ≤ f x" and xlex: "x ≤ x+d" and cont': "∀z. x ≤ z ∧ z ≤ x+d --> isCont f z" by (auto simp add: abs_if) from IVT2 [OF ge flef xlex cont'] obtain x' where "x ≤ x'" "x' ≤ x+d" "f x' = f(x-d)" by blast moreover hence "g(f x') = g (f(x-d))" by simp ultimately show False using d inj [of x'] inj [of "x-d"] by (simp add: abs_le_interval_iff) qed qed text{*Similar version for lower bound.*} lemma lemma_isCont_inj2: "[|0 < d; ∀z. ¦z-x¦ ≤ d --> g(f z) = z; ∀z. ¦z-x¦ ≤ d --> isCont f z |] ==> ∃z. ¦z-x¦ ≤ d & f z < f x" apply (insert lemma_isCont_inj [where f = "%x. - f x" and g = "%y. g(-y)" and x = x and d = d]) apply (simp add: isCont_minus linorder_not_le) done text{*Show there's an interval surrounding @{term "f(x)"} in @{text "f[[x - d, x + d]]"} .*} lemma isCont_inj_range: assumes d: "0 < d" and inj: "∀z. ¦z-x¦ ≤ d --> g(f z) = z" and cont: "∀z. ¦z-x¦ ≤ d --> isCont f z" shows "∃e>0. ∀y. ¦y - f x¦ ≤ e --> (∃z. ¦z-x¦ ≤ d & f z = y)" proof - have "x-d ≤ x+d" "∀z. x-d ≤ z ∧ z ≤ x+d --> isCont f z" using cont d by (auto simp add: abs_le_interval_iff) from isCont_Lb_Ub [OF this] obtain L M where all1 [rule_format]: "∀z. x-d ≤ z ∧ z ≤ x+d --> L ≤ f z ∧ f z ≤ M" and all2 [rule_format]: "∀y. L ≤ y ∧ y ≤ M --> (∃z. x-d ≤ z ∧ z ≤ x+d ∧ f z = y)" by auto with d have "L ≤ f x & f x ≤ M" by simp moreover have "L ≠ f x" proof - from lemma_isCont_inj2 [OF d inj cont] obtain u where "¦u - x¦ ≤ d" "f u < f x" by auto thus ?thesis using all1 [of u] by arith qed moreover have "f x ≠ M" proof - from lemma_isCont_inj [OF d inj cont] obtain u where "¦u - x¦ ≤ d" "f x < f u" by auto thus ?thesis using all1 [of u] by arith qed ultimately have "L < f x & f x < M" by arith hence "0 < f x - L" "0 < M - f x" by arith+ from real_lbound_gt_zero [OF this] obtain e where e: "0 < e" "e < f x - L" "e < M - f x" by auto thus ?thesis proof (intro exI conjI) show "0<e" . show "∀y. ¦y - f x¦ ≤ e --> (∃z. ¦z - x¦ ≤ d ∧ f z = y)" proof (intro strip) fix y::real assume "¦y - f x¦ ≤ e" with e have "L ≤ y ∧ y ≤ M" by arith from all2 [OF this] obtain z where "x - d ≤ z" "z ≤ x + d" "f z = y" by blast thus "∃z. ¦z - x¦ ≤ d ∧ f z = y" by (force simp add: abs_le_interval_iff) qed qed qed text{*Continuity of inverse function*} lemma isCont_inverse_function: assumes d: "0 < d" and inj: "∀z. ¦z-x¦ ≤ d --> g(f z) = z" and cont: "∀z. ¦z-x¦ ≤ d --> isCont f z" shows "isCont g (f x)" proof (simp add: isCont_iff LIM_eq) show "∀r. 0 < r --> (∃s>0. ∀z. z≠0 ∧ ¦z¦ < s --> ¦g(f x + z) - g(f x)¦ < r)" proof (intro strip) fix r::real assume r: "0<r" from real_lbound_gt_zero [OF r d] obtain e where e: "0 < e" and e_lt: "e < r ∧ e < d" by blast with inj cont have e_simps: "∀z. ¦z-x¦ ≤ e --> g (f z) = z" "∀z. ¦z-x¦ ≤ e --> isCont f z" by auto from isCont_inj_range [OF e this] obtain e' where e': "0 < e'" and all: "∀y. ¦y - f x¦ ≤ e' --> (∃z. ¦z - x¦ ≤ e ∧ f z = y)" by blast show "∃s>0. ∀z. z≠0 ∧ ¦z¦ < s --> ¦g(f x + z) - g(f x)¦ < r" proof (intro exI conjI) show "0<e'" . show "∀z. z ≠ 0 ∧ ¦z¦ < e' --> ¦g (f x + z) - g (f x)¦ < r" proof (intro strip) fix z::real assume z: "z ≠ 0 ∧ ¦z¦ < e'" with e e_lt e_simps all [rule_format, of "f x + z"] show "¦g (f x + z) - g (f x)¦ < r" by force qed qed qed qed ML {* val LIM_def = thm"LIM_def"; val NSLIM_def = thm"NSLIM_def"; val isCont_def = thm"isCont_def"; val isNSCont_def = thm"isNSCont_def"; val deriv_def = thm"deriv_def"; val nsderiv_def = thm"nsderiv_def"; val differentiable_def = thm"differentiable_def"; val NSdifferentiable_def = thm"NSdifferentiable_def"; val increment_def = thm"increment_def"; val isUCont_def = thm"isUCont_def"; val isNSUCont_def = thm"isNSUCont_def"; val half_gt_zero_iff = thm "half_gt_zero_iff"; val half_gt_zero = thms "half_gt_zero"; val abs_diff_triangle_ineq = thm "abs_diff_triangle_ineq"; val LIM_eq = thm "LIM_eq"; val LIM_D = thm "LIM_D"; val LIM_const = thm "LIM_const"; val LIM_add = thm "LIM_add"; val LIM_minus = thm "LIM_minus"; val LIM_add_minus = thm "LIM_add_minus"; val LIM_diff = thm "LIM_diff"; val LIM_const_not_eq = thm "LIM_const_not_eq"; val LIM_const_eq = thm "LIM_const_eq"; val LIM_unique = thm "LIM_unique"; val LIM_mult_zero = thm "LIM_mult_zero"; val LIM_self = thm "LIM_self"; val LIM_equal = thm "LIM_equal"; val LIM_trans = thm "LIM_trans"; val LIM_NSLIM = thm "LIM_NSLIM"; val NSLIM_LIM = thm "NSLIM_LIM"; val LIM_NSLIM_iff = thm "LIM_NSLIM_iff"; val NSLIM_mult = thm "NSLIM_mult"; val LIM_mult2 = thm "LIM_mult2"; val NSLIM_add = thm "NSLIM_add"; val LIM_add2 = thm "LIM_add2"; val NSLIM_const = thm "NSLIM_const"; val LIM_const2 = thm "LIM_const2"; val NSLIM_minus = thm "NSLIM_minus"; val LIM_minus2 = thm "LIM_minus2"; val NSLIM_add_minus = thm "NSLIM_add_minus"; val LIM_add_minus2 = thm "LIM_add_minus2"; val NSLIM_inverse = thm "NSLIM_inverse"; val LIM_inverse = thm "LIM_inverse"; val NSLIM_zero = thm "NSLIM_zero"; val LIM_zero2 = thm "LIM_zero2"; val NSLIM_zero_cancel = thm "NSLIM_zero_cancel"; val LIM_zero_cancel = thm "LIM_zero_cancel"; val NSLIM_not_zero = thm "NSLIM_not_zero"; val NSLIM_const_not_eq = thm "NSLIM_const_not_eq"; val NSLIM_const_eq = thm "NSLIM_const_eq"; val NSLIM_unique = thm "NSLIM_unique"; val LIM_unique2 = thm "LIM_unique2"; val NSLIM_mult_zero = thm "NSLIM_mult_zero"; val LIM_mult_zero2 = thm "LIM_mult_zero2"; val NSLIM_self = thm "NSLIM_self"; val isNSContD = thm "isNSContD"; val isNSCont_NSLIM = thm "isNSCont_NSLIM"; val NSLIM_isNSCont = thm "NSLIM_isNSCont"; val isNSCont_NSLIM_iff = thm "isNSCont_NSLIM_iff"; val isNSCont_LIM_iff = thm "isNSCont_LIM_iff"; val isNSCont_isCont_iff = thm "isNSCont_isCont_iff"; val isCont_isNSCont = thm "isCont_isNSCont"; val isNSCont_isCont = thm "isNSCont_isCont"; val NSLIM_h_iff = thm "NSLIM_h_iff"; val NSLIM_isCont_iff = thm "NSLIM_isCont_iff"; val LIM_isCont_iff = thm "LIM_isCont_iff"; val isCont_iff = thm "isCont_iff"; val isCont_add = thm "isCont_add"; val isCont_mult = thm "isCont_mult"; val isCont_o = thm "isCont_o"; val isCont_o2 = thm "isCont_o2"; val isNSCont_minus = thm "isNSCont_minus"; val isCont_minus = thm "isCont_minus"; val isCont_inverse = thm "isCont_inverse"; val isNSCont_inverse = thm "isNSCont_inverse"; val isCont_diff = thm "isCont_diff"; val isCont_const = thm "isCont_const"; val isNSCont_const = thm "isNSCont_const"; val isNSUContD = thm "isNSUContD"; val isUCont_isCont = thm "isUCont_isCont"; val isUCont_isNSUCont = thm "isUCont_isNSUCont"; val isNSUCont_isUCont = thm "isNSUCont_isUCont"; val DERIV_iff = thm "DERIV_iff"; val DERIV_NS_iff = thm "DERIV_NS_iff"; val DERIV_D = thm "DERIV_D"; val NS_DERIV_D = thm "NS_DERIV_D"; val DERIV_unique = thm "DERIV_unique"; val NSDeriv_unique = thm "NSDeriv_unique"; val differentiableD = thm "differentiableD"; val differentiableI = thm "differentiableI"; val NSdifferentiableD = thm "NSdifferentiableD"; val NSdifferentiableI = thm "NSdifferentiableI"; val LIM_I = thm "LIM_I"; val DERIV_LIM_iff = thm "DERIV_LIM_iff"; val DERIV_iff2 = thm "DERIV_iff2"; val NSDERIV_NSLIM_iff = thm "NSDERIV_NSLIM_iff"; val NSDERIV_NSLIM_iff2 = thm "NSDERIV_NSLIM_iff2"; val NSDERIV_iff2 = thm "NSDERIV_iff2"; val hypreal_not_eq_minus_iff = thm "hypreal_not_eq_minus_iff"; val NSDERIVD5 = thm "NSDERIVD5"; val NSDERIVD4 = thm "NSDERIVD4"; val NSDERIVD3 = thm "NSDERIVD3"; val NSDERIV_DERIV_iff = thm "NSDERIV_DERIV_iff"; val NSDERIV_isNSCont = thm "NSDERIV_isNSCont"; val DERIV_isCont = thm "DERIV_isCont"; val NSDERIV_const = thm "NSDERIV_const"; val DERIV_const = thm "DERIV_const"; val NSDERIV_add = thm "NSDERIV_add"; val DERIV_add = thm "DERIV_add"; val NSDERIV_mult = thm "NSDERIV_mult"; val DERIV_mult = thm "DERIV_mult"; val NSDERIV_cmult = thm "NSDERIV_cmult"; val DERIV_cmult = thm "DERIV_cmult"; val NSDERIV_minus = thm "NSDERIV_minus"; val DERIV_minus = thm "DERIV_minus"; val NSDERIV_add_minus = thm "NSDERIV_add_minus"; val DERIV_add_minus = thm "DERIV_add_minus"; val NSDERIV_diff = thm "NSDERIV_diff"; val DERIV_diff = thm "DERIV_diff"; val incrementI = thm "incrementI"; val incrementI2 = thm "incrementI2"; val increment_thm = thm "increment_thm"; val increment_thm2 = thm "increment_thm2"; val increment_approx_zero = thm "increment_approx_zero"; val NSDERIV_zero = thm "NSDERIV_zero"; val NSDERIV_approx = thm "NSDERIV_approx"; val NSDERIVD1 = thm "NSDERIVD1"; val NSDERIVD2 = thm "NSDERIVD2"; val NSDERIV_chain = thm "NSDERIV_chain"; val DERIV_chain = thm "DERIV_chain"; val DERIV_chain2 = thm "DERIV_chain2"; val NSDERIV_Id = thm "NSDERIV_Id"; val DERIV_Id = thm "DERIV_Id"; val isCont_Id = thms "isCont_Id"; val DERIV_cmult_Id = thm "DERIV_cmult_Id"; val NSDERIV_cmult_Id = thm "NSDERIV_cmult_Id"; val DERIV_pow = thm "DERIV_pow"; val NSDERIV_pow = thm "NSDERIV_pow"; val NSDERIV_inverse = thm "NSDERIV_inverse"; val DERIV_inverse = thm "DERIV_inverse"; val DERIV_inverse_fun = thm "DERIV_inverse_fun"; val NSDERIV_inverse_fun = thm "NSDERIV_inverse_fun"; val DERIV_quotient = thm "DERIV_quotient"; val NSDERIV_quotient = thm "NSDERIV_quotient"; val CARAT_DERIV = thm "CARAT_DERIV"; val CARAT_NSDERIV = thm "CARAT_NSDERIV"; val hypreal_eq_minus_iff3 = thm "hypreal_eq_minus_iff3"; val starfun_if_eq = thm "starfun_if_eq"; val CARAT_DERIVD = thm "CARAT_DERIVD"; val f_inc_g_dec_Beq_f = thm "f_inc_g_dec_Beq_f"; val f_inc_g_dec_Beq_g = thm "f_inc_g_dec_Beq_g"; val f_inc_imp_le_lim = thm "f_inc_imp_le_lim"; val lim_uminus = thm "lim_uminus"; val g_dec_imp_lim_le = thm "g_dec_imp_lim_le"; val Bolzano_bisect_le = thm "Bolzano_bisect_le"; val Bolzano_bisect_fst_le_Suc = thm "Bolzano_bisect_fst_le_Suc"; val Bolzano_bisect_Suc_le_snd = thm "Bolzano_bisect_Suc_le_snd"; val eq_divide_2_times_iff = thm "eq_divide_2_times_iff"; val Bolzano_bisect_diff = thm "Bolzano_bisect_diff"; val Bolzano_nest_unique = thms "Bolzano_nest_unique"; val not_P_Bolzano_bisect = thm "not_P_Bolzano_bisect"; val not_P_Bolzano_bisect = thm "not_P_Bolzano_bisect"; val lemma_BOLZANO2 = thm "lemma_BOLZANO2"; val IVT = thm "IVT"; val IVT2 = thm "IVT2"; val IVT_objl = thm "IVT_objl"; val IVT2_objl = thm "IVT2_objl"; val isCont_bounded = thm "isCont_bounded"; val isCont_has_Ub = thm "isCont_has_Ub"; val isCont_eq_Ub = thm "isCont_eq_Ub"; val isCont_eq_Lb = thm "isCont_eq_Lb"; val isCont_Lb_Ub = thm "isCont_Lb_Ub"; val DERIV_left_inc = thm "DERIV_left_inc"; val DERIV_left_dec = thm "DERIV_left_dec"; val DERIV_local_max = thm "DERIV_local_max"; val DERIV_local_min = thm "DERIV_local_min"; val DERIV_local_const = thm "DERIV_local_const"; val Rolle = thm "Rolle"; val MVT = thm "MVT"; val DERIV_isconst_end = thm "DERIV_isconst_end"; val DERIV_isconst1 = thm "DERIV_isconst1"; val DERIV_isconst2 = thm "DERIV_isconst2"; val DERIV_isconst_all = thm "DERIV_isconst_all"; val DERIV_const_ratio_const = thm "DERIV_const_ratio_const"; val DERIV_const_ratio_const2 = thm "DERIV_const_ratio_const2"; val real_average_minus_first = thm "real_average_minus_first"; val real_average_minus_second = thm "real_average_minus_second"; val DERIV_const_average = thm "DERIV_const_average"; val isCont_inj_range = thm "isCont_inj_range"; val isCont_inverse_function = thm "isCont_inverse_function"; *} end
lemma LIM_eq:
f -- a --> L = (∀r>0. ∃s>0. ∀x. x ≠ a ∧ ¦x - a¦ < s --> ¦f x - L¦ < r)
lemma LIM_D:
[| f -- a --> L; 0 < r |] ==> ∃s>0. ∀x. x ≠ a ∧ ¦x - a¦ < s --> ¦f x - L¦ < r
lemma LIM_const:
(%x. k) -- x --> k
lemma LIM_add:
[| f -- a --> L; g -- a --> M |] ==> (%x. f x + g x) -- a --> L + M
lemma LIM_minus:
f -- a --> L ==> (%x. - f x) -- a --> - L
lemma LIM_add_minus:
[| f -- x --> l; g -- x --> m |] ==> (%x. f x + - g x) -- x --> l + - m
lemma LIM_diff:
[| f -- x --> l; g -- x --> m |] ==> (%x. f x - g x) -- x --> l - m
lemma LIM_const_not_eq:
k ≠ L ==> ¬ (%x. k) -- a --> L
lemma LIM_const_eq:
(%x. k) -- x --> L ==> k = L
lemma LIM_unique:
[| f -- a --> L; f -- a --> M |] ==> L = M
lemma LIM_mult_zero:
[| f -- a --> 0; g -- a --> 0 |] ==> (%x. f x * g x) -- a --> 0
lemma LIM_self:
(%x. x) -- a --> a
lemma LIM_equal:
∀x. x ≠ a --> f x = g x ==> f -- a --> l = g -- a --> l
lemma LIM_trans:
[| (%x. f x + - g x) -- a --> 0; g -- a --> l |] ==> f -- a --> l
lemma LIM_NSLIM:
f -- x --> L ==> f -- x --NS> L
lemma lemma_LIM:
∀s>0. ∃xa. xa ≠ x ∧ ¦xa + - x¦ < s ∧ r ≤ ¦f xa + - L¦ ==> ∀n. ∃xa. xa ≠ x ∧ ¦xa + - x¦ < inverse (real (Suc n)) ∧ r ≤ ¦f xa + - L¦
lemma lemma_skolemize_LIM2:
∀s>0. ∃xa. xa ≠ x ∧ ¦xa + - x¦ < s ∧ r ≤ ¦f xa + - L¦ ==> ∃X. ∀n. X n ≠ x ∧ ¦X n + - x¦ < inverse (real (Suc n)) ∧ r ≤ ¦f (X n) + - L¦
lemma lemma_simp:
∀n. X n ≠ x ∧ ¦X n + - x¦ < inverse (real (Suc n)) ∧ r ≤ ¦f (X n) + - L¦ ==> ∀n. ¦X n + - x¦ < inverse (real (Suc n))
lemma NSLIM_LIM:
f -- x --NS> L ==> f -- x --> L
theorem LIM_NSLIM_iff:
f -- x --> L = f -- x --NS> L
lemma NSLIM_mult:
[| f -- x --NS> l; g -- x --NS> m |] ==> (%x. f x * g x) -- x --NS> l * m
lemma LIM_mult2:
[| f -- x --> l; g -- x --> m |] ==> (%x. f x * g x) -- x --> l * m
lemma NSLIM_add:
[| f -- x --NS> l; g -- x --NS> m |] ==> (%x. f x + g x) -- x --NS> l + m
lemma LIM_add2:
[| f -- x --> l; g -- x --> m |] ==> (%x. f x + g x) -- x --> l + m
lemma NSLIM_const:
(%x. k) -- x --NS> k
lemma LIM_const2:
(%x. k) -- x --> k
lemma NSLIM_minus:
f -- a --NS> L ==> (%x. - f x) -- a --NS> - L
lemma LIM_minus2:
f -- a --> L ==> (%x. - f x) -- a --> - L
lemma NSLIM_add_minus:
[| f -- x --NS> l; g -- x --NS> m |] ==> (%x. f x + - g x) -- x --NS> l + - m
lemma LIM_add_minus2:
[| f -- x --> l; g -- x --> m |] ==> (%x. f x + - g x) -- x --> l + - m
lemma NSLIM_inverse:
[| f -- a --NS> L; L ≠ 0 |] ==> (%x. inverse (f x)) -- a --NS> inverse L
lemma LIM_inverse:
[| f -- a --> L; L ≠ 0 |] ==> (%x. inverse (f x)) -- a --> inverse L
lemma NSLIM_zero:
f -- a --NS> l ==> (%x. f x + - l) -- a --NS> 0
lemma LIM_zero2:
f -- a --> l ==> (%x. f x + - l) -- a --> 0
lemma NSLIM_zero_cancel:
(%x. f x - l) -- x --NS> 0 ==> f -- x --NS> l
lemma LIM_zero_cancel:
(%x. f x - l) -- x --> 0 ==> f -- x --> l
lemma NSLIM_not_zero:
k ≠ 0 ==> ¬ (%x. k) -- x --NS> 0
lemma NSLIM_const_not_eq:
k ≠ L ==> ¬ (%x. k) -- x --NS> L
lemma NSLIM_const_eq:
(%x. k) -- x --NS> L ==> k = L
lemma NSLIM_unique:
[| f -- x --NS> L; f -- x --NS> M |] ==> L = M
lemma LIM_unique2:
[| f -- x --> L; f -- x --> M |] ==> L = M
lemma NSLIM_mult_zero:
[| f -- x --NS> 0; g -- x --NS> 0 |] ==> (%x. f x * g x) -- x --NS> 0
lemma LIM_mult_zero2:
[| f -- x --> 0; g -- x --> 0 |] ==> (%x. f x * g x) -- x --> 0
lemma NSLIM_self:
(%x. x) -- a --NS> a
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 --NS> L
lemma NSLIM_isCont_iff:
f -- a --NS> f a = (%h. f (a + h)) -- 0 --NS> f a
lemma LIM_isCont_iff:
f -- a --> f a = (%h. f (a + h)) -- 0 --> f a
lemma isCont_iff:
isCont f x = (%h. f (x + h)) -- 0 --> f x
lemma isCont_add:
[| isCont f a; isCont g a |] ==> isCont (%x. f x + g x) a
lemma isCont_mult:
[| isCont f a; isCont g a |] ==> isCont (%x. f x * g x) a
lemma isCont_o:
[| isCont f a; isCont g (f a) |] ==> isCont (g o f) a
lemma isCont_o2:
[| isCont f a; isCont g (f a) |] ==> isCont (%x. g (f x)) a
lemma isNSCont_minus:
isNSCont f a ==> isNSCont (%x. - f x) a
lemma isCont_minus:
isCont f a ==> isCont (%x. - f x) a
lemma isCont_inverse:
[| isCont f x; f x ≠ 0 |] ==> isCont (%x. inverse (f x)) x
lemma isNSCont_inverse:
[| isNSCont f x; f x ≠ 0 |] ==> isNSCont (%x. inverse (f x)) x
lemma isCont_diff:
[| isCont f a; isCont g a |] ==> isCont (%x. f x - g x) a
lemma isCont_const:
isCont (%x. k) a
lemma isNSCont_const:
isNSCont (%x. k) a
lemma isNSCont_abs:
isNSCont abs a
lemma isCont_abs:
isCont abs a
lemma isNSUContD:
[| isNSUCont f; x ≈ y |] ==> (*f* f) x ≈ (*f* f) y
lemma isUCont_isCont:
isUCont f ==> isCont f x
lemma isUCont_isNSUCont:
isUCont f ==> isNSUCont f
lemma lemma_LIMu:
∀s>0. ∃z y. ¦z + - y¦ < s ∧ r ≤ ¦f z + - f y¦ ==> ∀n. ∃z y. ¦z + - y¦ < inverse (real (Suc n)) ∧ r ≤ ¦f z + - f y¦
lemma lemma_skolemize_LIM2u:
∀s>0. ∃z y. ¦z + - y¦ < s ∧ r ≤ ¦f z + - f y¦ ==> ∃X Y. ∀n. ¦X n + - Y n¦ < inverse (real (Suc n)) ∧ r ≤ ¦f (X n) + - f (Y n)¦
lemma lemma_simpu:
∀n. ¦X n + - Y n¦ < inverse (real (Suc n)) ∧ r ≤ ¦f (X n) + - f (Y n)¦ ==> ∀n. ¦X n + - Y n¦ < inverse (real (Suc n))
lemma isNSUCont_isUCont:
isNSUCont f ==> isUCont f
lemma DERIV_iff:
DERIV f x :> D = (%h. (f (x + h) + - f x) / h) -- 0 --> D
lemma DERIV_NS_iff:
DERIV f x :> D = (%h. (f (x + h) + - f x) / h) -- 0 --NS> D
lemma DERIV_D:
DERIV f x :> D ==> (%h. (f (x + h) + - f x) / h) -- 0 --> D
lemma NS_DERIV_D:
DERIV f x :> D ==> (%h. (f (x + h) + - f x) / h) -- 0 --NS> D
lemma DERIV_unique:
[| DERIV f x :> D; DERIV f x :> E |] ==> D = E
lemma NSDeriv_unique:
[| NSDERIV f x :> D; NSDERIV f x :> E |] ==> D = E
lemma differentiableD:
f differentiable x ==> ∃D. DERIV f x :> D
lemma differentiableI:
DERIV f x :> D ==> f differentiable x
lemma NSdifferentiableD:
f NSdifferentiable x ==> ∃D. NSDERIV f x :> D
lemma NSdifferentiableI:
NSDERIV f x :> D ==> f NSdifferentiable x
lemma LIM_I:
(!!r. 0 < r ==> ∃s>0. ∀x. x ≠ a ∧ ¦x - a¦ < s --> ¦f x - L¦ < r) ==> f -- a --> L
lemma DERIV_LIM_iff:
(%h. (f (a + h) - f a) / h) -- 0 --> D = (%x. (f x - f a) / (x - a)) -- a --> D
lemma DERIV_iff2:
DERIV f x :> D = (%z. (f z - f x) / (z - x)) -- x --> D
lemma NSDERIV_NSLIM_iff:
NSDERIV f x :> D = (%h. (f (x + h) + - f x) / h) -- 0 --NS> D
lemma NSDERIV_NSLIM_iff2:
NSDERIV f x :> D = (%z. (f z - f x) / (z - x)) -- x --NS> D
lemma NSDERIV_iff2:
NSDERIV f x :> D = (∀w. w ≠ star_of x ∧ w ≈ star_of x --> (*f* (%z. (f z - f x) / (z - x))) w ≈ star_of D)
lemma hypreal_not_eq_minus_iff:
(x ≠ a) = (x + - a ≠ 0)
lemma NSDERIVD5:
NSDERIV f x :> D ==> ∀u. u ≈ star_of x --> (*f* (%z. f z - f x)) u ≈ star_of D * (u - star_of x)
lemma NSDERIVD4:
NSDERIV f x :> D ==> ∀h∈Infinitesimal. (*f* f) (star_of x + h) - star_of (f x) ≈ star_of D * h
lemma NSDERIVD3:
NSDERIV f x :> D ==> ∀h∈Infinitesimal - {0}. (*f* f) (star_of x + h) - star_of (f x) ≈ star_of D * h
lemma NSDERIV_DERIV_iff:
NSDERIV f x :> D = DERIV f x :> D
lemma NSDERIV_isNSCont:
NSDERIV f x :> D ==> isNSCont f x
lemma DERIV_isCont:
DERIV f x :> D ==> isCont f x
lemma NSDERIV_const:
NSDERIV (%x. k) x :> 0
lemma DERIV_const:
DERIV (%x. k) x :> 0
lemma NSDERIV_add:
[| NSDERIV f x :> Da; NSDERIV g x :> Db |] ==> NSDERIV (%x. f x + g x) x :> Da + Db
lemma DERIV_add:
[| DERIV f x :> Da; DERIV g x :> Db |] ==> DERIV (%x. f x + g x) x :> Da + Db
lemma lemma_nsderiv1:
a * b + - (c * d) = b * (a + - c) + c * (b + - d)
lemma lemma_nsderiv2:
[| (x + y) / z = star_of D + yb; z ≠ 0; z ∈ Infinitesimal; yb ∈ Infinitesimal |] ==> x + y ≈ 0
lemma NSDERIV_mult:
[| NSDERIV f x :> Da; NSDERIV g x :> Db |] ==> NSDERIV (%x. f x * g x) x :> Da * g x + Db * f x
lemma DERIV_mult:
[| DERIV f x :> Da; DERIV g x :> Db |] ==> DERIV (%x. f x * g x) x :> Da * g x + Db * f x
lemma NSDERIV_cmult:
NSDERIV f x :> D ==> NSDERIV (%x. c * f x) x :> c * D
lemma DERIV_cmult:
DERIV f x :> D ==> DERIV (%x. c * f x) x :> c * D
lemma NSDERIV_minus:
NSDERIV f x :> D ==> NSDERIV (%x. - f x) x :> - D
lemma DERIV_minus:
DERIV f x :> D ==> DERIV (%x. - f x) x :> - D
lemma NSDERIV_add_minus:
[| NSDERIV f x :> Da; NSDERIV g x :> Db |] ==> NSDERIV (%x. f x + - g x) x :> Da + - Db
lemma DERIV_add_minus:
[| DERIV f x :> Da; DERIV g x :> Db |] ==> DERIV (%x. f x + - g x) x :> Da + - Db
lemma NSDERIV_diff:
[| NSDERIV f x :> Da; NSDERIV g x :> Db |] ==> NSDERIV (%x. f x - g x) x :> Da - Db
lemma DERIV_diff:
[| DERIV f x :> Da; DERIV g x :> Db |] ==> DERIV (%x. f x - g x) x :> Da - Db
lemma incrementI:
f NSdifferentiable x ==> increment f x h = (*f* f) (star_of x + h) + - star_of (f x)
lemma incrementI2:
NSDERIV f x :> D ==> increment f x h = (*f* f) (star_of x + h) + - star_of (f x)
lemma increment_thm:
[| NSDERIV f x :> D; h ∈ Infinitesimal; h ≠ 0 |] ==> ∃e∈Infinitesimal. increment f x h = star_of D * h + e * h
lemma increment_thm2:
[| NSDERIV f x :> D; h ≈ 0; h ≠ 0 |] ==> ∃e∈Infinitesimal. increment f x h = star_of D * h + e * h
lemma increment_approx_zero:
[| NSDERIV f x :> D; h ≈ 0; h ≠ 0 |] ==> increment f x h ≈ 0
lemma NSDERIV_zero:
[| NSDERIV g x :> D; (*f* g) (star_of x + xa) = star_of (g x); xa ∈ Infinitesimal; xa ≠ 0 |] ==> D = 0
lemma NSDERIV_approx:
[| NSDERIV f x :> D; h ∈ Infinitesimal; h ≠ 0 |] ==> (*f* f) (star_of x + h) + - star_of (f x) ≈ 0
lemma NSDERIVD1:
[| NSDERIV f (g x) :> Da; (*f* g) (star_of x + xa) ≠ star_of (g x); (*f* g) (star_of x + xa) ≈ star_of (g x) |] ==> ((*f* f) ((*f* g) (star_of x + xa)) + - star_of (f (g x))) / ((*f* g) (star_of x + xa) + - star_of (g x)) ≈ star_of Da
lemma NSDERIVD2:
[| NSDERIV g x :> Db; xa ∈ Infinitesimal; xa ≠ 0 |] ==> ((*f* g) (star_of x + xa) + - star_of (g x)) / xa ≈ star_of Db
lemma lemma_chain:
z ≠ 0 ==> x * y = x * inverse z * (z * y)
lemma NSDERIV_chain:
[| NSDERIV f (g x) :> Da; NSDERIV g x :> Db |] ==> NSDERIV (f o g) x :> Da * Db
lemma DERIV_chain:
[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (f o g) x :> Da * Db
lemma DERIV_chain2:
[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (%x. f (g x)) x :> Da * Db
lemma NSDERIV_Id:
NSDERIV (%x. x) x :> 1
lemma DERIV_Id:
DERIV (%x. x) x :> 1
lemmas isCont_Id:
isCont (%x. x) x
lemmas isCont_Id:
isCont (%x. x) x
lemma DERIV_cmult_Id:
DERIV (op * c) x :> c
lemma NSDERIV_cmult_Id:
NSDERIV (op * c) x :> c
lemma DERIV_pow:
DERIV (%x. x ^ n) x :> real n * x ^ (n - Suc 0)
lemma NSDERIV_pow:
NSDERIV (%x. x ^ n) x :> real n * x ^ (n - Suc 0)
lemma NSDERIV_inverse:
x ≠ 0 ==> NSDERIV inverse x :> - (inverse x ^ Suc (Suc 0))
lemma DERIV_inverse:
x ≠ 0 ==> DERIV inverse x :> - (inverse x ^ Suc (Suc 0))
lemma DERIV_inverse_fun:
[| DERIV f x :> d; f x ≠ 0 |] ==> DERIV (%x. inverse (f x)) x :> - (d * inverse (f x ^ Suc (Suc 0)))
lemma NSDERIV_inverse_fun:
[| NSDERIV f x :> d; f x ≠ 0 |] ==> NSDERIV (%x. inverse (f x)) x :> - (d * inverse (f x ^ Suc (Suc 0)))
lemma DERIV_quotient:
[| DERIV f x :> d; DERIV g x :> e; g x ≠ 0 |] ==> DERIV (%y. f y / g y) x :> (d * g x + - (e * f x)) / g x ^ Suc (Suc 0)
lemma NSDERIV_quotient:
[| NSDERIV f x :> d; DERIV g x :> e; g x ≠ 0 |] ==> NSDERIV (%y. f y / g y) x :> (d * g x + - (e * f x)) / g x ^ Suc (Suc 0)
lemma CARAT_DERIV:
DERIV f x :> l = (∃g. (∀z. f z - f x = g z * (z - x)) ∧ isCont g x ∧ g x = l)
lemma CARAT_NSDERIV:
NSDERIV f x :> l ==> ∃g. (∀z. f z - f x = g z * (z - x)) ∧ isNSCont g x ∧ g x = l
lemma hypreal_eq_minus_iff3:
(x = y + z) = (x + - z = y)
lemma CARAT_DERIVD:
[| ∀z. f z - f x = g z * (z - x); isNSCont g x |] ==> NSDERIV f x :> g x
lemma lemma_f_mono_add:
∀n. f n ≤ f (Suc n) ==> f m ≤ f (m + no)
lemma f_inc_g_dec_Beq_f:
[| ∀n. f n ≤ f (Suc n); ∀n. g (Suc n) ≤ g n; ∀n. f n ≤ g n |] ==> Bseq f
lemma f_inc_g_dec_Beq_g:
[| ∀n. f n ≤ f (Suc n); ∀n. g (Suc n) ≤ g n; ∀n. f n ≤ g n |] ==> Bseq g
lemma f_inc_imp_le_lim:
[| ∀n. f n ≤ f (Suc n); convergent f |] ==> f n ≤ lim f
lemma lim_uminus:
convergent g ==> lim (%x. - g x) = - lim g
lemma g_dec_imp_lim_le:
[| ∀n. g (Suc n) ≤ g n; convergent g |] ==> lim g ≤ g n
lemma lemma_nest:
[| ∀n. f n ≤ f (Suc n); ∀n. g (Suc n) ≤ g n; ∀n. f n ≤ g n |] ==> ∃l m. l ≤ m ∧ ((∀n. f n ≤ l) ∧ f ----> l) ∧ (∀n. m ≤ g n) ∧ g ----> m
lemma lemma_nest_unique:
[| ∀n. f n ≤ f (Suc n); ∀n. g (Suc n) ≤ g n; ∀n. f n ≤ g n; (%n. f n - g n) ----> 0 |] ==> ∃l. ((∀n. f n ≤ l) ∧ f ----> l) ∧ (∀n. l ≤ g n) ∧ g ----> l
lemma Bolzano_bisect_le:
a ≤ b ==> ∀n. fst (Bolzano_bisect P a b n) ≤ snd (Bolzano_bisect P a b n)
lemma Bolzano_bisect_fst_le_Suc:
a ≤ b ==> ∀n. fst (Bolzano_bisect P a b n) ≤ fst (Bolzano_bisect P a b (Suc n))
lemma Bolzano_bisect_Suc_le_snd:
a ≤ b ==> ∀n. snd (Bolzano_bisect P a b (Suc n)) ≤ snd (Bolzano_bisect P a b n)
lemma eq_divide_2_times_iff:
(x = y / (2 * z)) = (2 * x = y / z)
lemma Bolzano_bisect_diff:
a ≤ b ==> snd (Bolzano_bisect P a b n) - fst (Bolzano_bisect P a b n) = (b - a) / 2 ^ n
lemmas Bolzano_nest_unique:
[| a1 ≤ b1; a1 ≤ b1; a1 ≤ b1; (%n. fst (Bolzano_bisect P1 a1 b1 n) - snd (Bolzano_bisect P1 a1 b1 n)) ----> 0 |] ==> ∃l. ((∀n. fst (Bolzano_bisect P1 a1 b1 n) ≤ l) ∧ (%n. fst (Bolzano_bisect P1 a1 b1 n)) ----> l) ∧ (∀n. l ≤ snd (Bolzano_bisect P1 a1 b1 n)) ∧ (%n. snd (Bolzano_bisect P1 a1 b1 n)) ----> l
lemmas Bolzano_nest_unique:
[| a1 ≤ b1; a1 ≤ b1; a1 ≤ b1; (%n. fst (Bolzano_bisect P1 a1 b1 n) - snd (Bolzano_bisect P1 a1 b1 n)) ----> 0 |] ==> ∃l. ((∀n. fst (Bolzano_bisect P1 a1 b1 n) ≤ l) ∧ (%n. fst (Bolzano_bisect P1 a1 b1 n)) ----> l) ∧ (∀n. l ≤ snd (Bolzano_bisect P1 a1 b1 n)) ∧ (%n. snd (Bolzano_bisect P1 a1 b1 n)) ----> l
lemma not_P_Bolzano_bisect:
[| !!a b c. [| P (a, b); P (b, c); a ≤ b; b ≤ c |] ==> P (a, c); ¬ P (a, b); a ≤ b |] ==> ¬ P (fst (Bolzano_bisect P a b n), snd (Bolzano_bisect P a b n))
lemma not_P_Bolzano_bisect':
[| ∀a b c. P (a, b) ∧ P (b, c) ∧ a ≤ b ∧ b ≤ c --> P (a, c); ¬ P (a, b); a ≤ b |] ==> ∀n. ¬ P (fst (Bolzano_bisect P a b n), snd (Bolzano_bisect P a b n))
lemma lemma_BOLZANO:
[| ∀a b c. P (a, b) ∧ P (b, c) ∧ a ≤ b ∧ b ≤ c --> P (a, c); ∀x. ∃d>0. ∀a b. a ≤ x ∧ x ≤ b ∧ b - a < d --> P (a, b); a ≤ b |] ==> P (a, b)
lemma lemma_BOLZANO2:
(∀a b c. a ≤ b ∧ b ≤ c ∧ P (a, b) ∧ P (b, c) --> P (a, c)) ∧ (∀x. ∃d>0. ∀a b. a ≤ x ∧ x ≤ b ∧ b - a < d --> P (a, b)) --> (∀a b. a ≤ b --> P (a, b))
lemma IVT:
[| f a ≤ y; y ≤ f b; a ≤ b; ∀x. a ≤ x ∧ x ≤ b --> isCont f x |] ==> ∃x≥a. x ≤ b ∧ f x = y
lemma IVT2:
[| f b ≤ y; y ≤ f a; a ≤ b; ∀x. a ≤ x ∧ x ≤ b --> isCont f x |] ==> ∃x≥a. x ≤ b ∧ f x = y
lemma IVT_objl:
f a ≤ y ∧ y ≤ f b ∧ a ≤ b ∧ (∀x. a ≤ x ∧ x ≤ b --> isCont f x) --> (∃x≥a. x ≤ b ∧ f x = y)
lemma IVT2_objl:
f b ≤ y ∧ y ≤ f a ∧ a ≤ b ∧ (∀x. a ≤ x ∧ x ≤ b --> isCont f x) --> (∃x≥a. x ≤ b ∧ f x = y)
lemma isCont_bounded:
[| a ≤ b; ∀x. a ≤ x ∧ x ≤ b --> isCont f x |] ==> ∃M. ∀x. a ≤ x ∧ x ≤ b --> f x ≤ M
lemma lemma_reals_complete:
(∃x. x ∈ S) ∧ (∃y. isUb UNIV S y) --> (∃t. isLub UNIV S t)
lemma isCont_has_Ub:
[| a ≤ b; ∀x. a ≤ x ∧ x ≤ b --> isCont f x |] ==> ∃M. (∀x. a ≤ x ∧ x ≤ b --> f x ≤ M) ∧ (∀N<M. ∃x≥a. x ≤ b ∧ N < f x)
lemma isCont_eq_Ub:
[| a ≤ b; ∀x. a ≤ x ∧ x ≤ b --> isCont f x |] ==> ∃M. (∀x. a ≤ x ∧ x ≤ b --> f x ≤ M) ∧ (∃x≥a. x ≤ b ∧ f x = M)
lemma isCont_eq_Lb:
[| a ≤ b; ∀x. a ≤ x ∧ x ≤ b --> isCont f x |] ==> ∃M. (∀x. a ≤ x ∧ x ≤ b --> M ≤ f x) ∧ (∃x≥a. x ≤ b ∧ f x = M)
lemma isCont_Lb_Ub:
[| a ≤ b; ∀x. a ≤ x ∧ x ≤ b --> isCont f x |] ==> ∃L M. (∀x. a ≤ x ∧ x ≤ b --> L ≤ f x ∧ f x ≤ M) ∧ (∀y. L ≤ y ∧ y ≤ M --> (∃x≥a. x ≤ b ∧ f x = y))
lemma DERIV_left_inc:
[| DERIV f x :> l; 0 < l |] ==> ∃d>0. ∀h>0. h < d --> f x < f (x + h)
lemma DERIV_left_dec:
[| DERIV f x :> l; l < 0 |] ==> ∃d>0. ∀h>0. h < d --> f x < f (x - h)
lemma DERIV_local_max:
[| DERIV f x :> l; 0 < d; ∀y. ¦x - y¦ < d --> f y ≤ f x |] ==> l = 0
lemma DERIV_local_min:
[| DERIV f x :> l; 0 < d; ∀y. ¦x - y¦ < d --> f x ≤ f y |] ==> l = 0
lemma DERIV_local_const:
[| DERIV f x :> l; 0 < d; ∀y. ¦x - y¦ < d --> f x = f y |] ==> l = 0
lemma lemma_interval_lt:
[| a < x; x < b |] ==> ∃d>0. ∀y. ¦x - y¦ < d --> a < y ∧ y < b
lemma lemma_interval:
[| a < x; x < b |] ==> ∃d>0. ∀y. ¦x - y¦ < d --> a ≤ y ∧ y ≤ b
theorem Rolle:
[| a < b; f a = f b; ∀x. a ≤ x ∧ x ≤ b --> isCont f x; ∀x. a < x ∧ x < b --> f differentiable x |] ==> ∃z>a. z < b ∧ DERIV f z :> 0
lemma lemma_MVT:
f a - (f b - f a) / (b - a) * a = f b - (f b - f a) / (b - a) * b
theorem MVT:
[| a < b; ∀x. a ≤ x ∧ x ≤ b --> isCont f x; ∀x. a < x ∧ x < b --> f differentiable x |] ==> ∃l z. a < z ∧ z < b ∧ DERIV f z :> l ∧ f b - f a = (b - a) * l
lemma DERIV_isconst_end:
[| a < b; ∀x. a ≤ x ∧ x ≤ b --> isCont f x; ∀x. a < x ∧ x < b --> DERIV f x :> 0 |] ==> f b = f a
lemma DERIV_isconst1:
[| a < b; ∀x. a ≤ x ∧ x ≤ b --> isCont f x; ∀x. a < x ∧ x < b --> DERIV f x :> 0 |] ==> ∀x. a ≤ x ∧ x ≤ b --> f x = f a
lemma DERIV_isconst2:
[| a < b; ∀x. a ≤ x ∧ x ≤ b --> isCont f x; ∀x. a < x ∧ x < b --> DERIV f x :> 0; a ≤ x; x ≤ b |] ==> f x = f a
lemma DERIV_isconst_all:
∀x. DERIV f x :> 0 ==> f x = f y
lemma DERIV_const_ratio_const:
[| a ≠ b; ∀x. DERIV f x :> k |] ==> f b - f a = (b - a) * k
lemma DERIV_const_ratio_const2:
[| a ≠ b; ∀x. DERIV f x :> k |] ==> (f b - f a) / (b - a) = k
lemma real_average_minus_first:
(a + b) / 2 - a = (b - a) / 2
lemma real_average_minus_second:
(b + a) / 2 - a = (b - a) / 2
lemma DERIV_const_average:
[| a ≠ b; ∀x. DERIV v x :> k |] ==> v ((a + b) / 2) = (v a + v b) / 2
lemma lemma_isCont_inj:
[| 0 < d; ∀z. ¦z - x¦ ≤ d --> g (f z) = z; ∀z. ¦z - x¦ ≤ d --> isCont f z |] ==> ∃z. ¦z - x¦ ≤ d ∧ f x < f z
lemma lemma_isCont_inj2:
[| 0 < d; ∀z. ¦z - x¦ ≤ d --> g (f z) = z; ∀z. ¦z - x¦ ≤ d --> isCont f z |] ==> ∃z. ¦z - x¦ ≤ d ∧ f z < f x
lemma isCont_inj_range:
[| 0 < d; ∀z. ¦z - x¦ ≤ d --> g (f z) = z; ∀z. ¦z - x¦ ≤ d --> isCont f z |] ==> ∃e>0. ∀y. ¦y - f x¦ ≤ e --> (∃z. ¦z - x¦ ≤ d ∧ f z = y)
lemma isCont_inverse_function:
[| 0 < d; ∀z. ¦z - x¦ ≤ d --> g (f z) = z; ∀z. ¦z - x¦ ≤ d --> isCont f z |] ==> isCont g (f x)