# HG changeset patch # User huffman # Date 1162595471 -3600 # Node ID 0742fc979c674e437dae6f44f53cb691d83ff7b9 # Parent 6860f161111c3df95f9d01a9ac1b7d7d643b933d new Deriv.thy contains stuff from Lim.thy diff -r 6860f161111c -r 0742fc979c67 src/HOL/Hyperreal/Deriv.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hyperreal/Deriv.thy Sat Nov 04 00:11:11 2006 +0100 @@ -0,0 +1,1753 @@ +(* Title : Deriv.thy + ID : $Id$ + Author : Jacques D. Fleuriot + Copyright : 1998 University of Cambridge + Conversion to Isar and new proofs by Lawrence C Paulson, 2004 + GMVT by Benjamin Porter, 2005 +*) + +header{* Differentiation *} + +theory Deriv +imports Lim +begin + +text{*Standard and Nonstandard Definitions*} + +definition + + deriv :: "[real \ 'a::real_normed_vector, real, 'a] \ 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))" + + +consts + Bolzano_bisect :: "[real*real=>bool, real, real, nat] => (real*real)" +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))" + + +subsection {* Derivatives *} + +subsubsection {* Purely standard proofs *} + +lemma DERIV_iff: "(DERIV f x :> D) = ((%h. (f(x + h) - f(x))/#h) -- 0 --> D)" +by (simp add: deriv_def) + +lemma DERIV_D: "DERIV f x :> D ==> (%h. (f(x + h) - f(x))/#h) -- 0 --> D" +by (simp add: deriv_def) + +lemma DERIV_const [simp]: "DERIV (\x. k) x :> 0" +by (simp add: deriv_def) + +lemma DERIV_Id [simp]: "DERIV (\x. x) x :> 1" +by (simp add: deriv_def real_scaleR_def cong: LIM_cong) + +lemma add_diff_add: + fixes a b c d :: "'a::ab_group_add" + shows "(a + c) - (b + d) = (a - b) + (c - d)" +by simp + +lemma DERIV_add: + "\DERIV f x :> D; DERIV g x :> E\ \ DERIV (\x. f x + g x) x :> D + E" +by (simp only: deriv_def add_diff_add scaleR_right_distrib LIM_add) + +lemma DERIV_minus: + "DERIV f x :> D \ DERIV (\x. - f x) x :> - D" +by (simp only: deriv_def minus_diff_minus scaleR_minus_right LIM_minus) + +lemma DERIV_diff: + "\DERIV f x :> D; DERIV g x :> E\ \ DERIV (\x. f x - g x) x :> D - E" +by (simp only: diff_def DERIV_add DERIV_minus) + +lemma DERIV_add_minus: + "\DERIV f x :> D; DERIV g x :> E\ \ DERIV (\x. f x + - g x) x :> D + - E" +by (simp only: DERIV_add DERIV_minus) + +lemma DERIV_isCont: "DERIV f x :> D \ isCont f x" +proof (unfold isCont_iff) + assume "DERIV f x :> D" + hence "(\h. (f(x+h) - f(x)) /# h) -- 0 --> D" + by (rule DERIV_D) + hence "(\h. h *# ((f(x+h) - f(x)) /# h)) -- 0 --> 0 *# D" + by (intro LIM_scaleR LIM_self) + hence "(\h. (f(x+h) - f(x))) -- 0 --> 0" + by (simp cong: LIM_cong) + thus "(\h. f(x+h)) -- 0 --> f(x)" + by (simp add: LIM_def) +qed + +lemma DERIV_mult_lemma: + fixes a b c d :: "'a::real_algebra" + shows "(a * b - c * d) /# h = a * ((b - d) /# h) + ((a - c) /# h) * d" +by (simp add: diff_minus scaleR_right_distrib [symmetric] ring_distrib) + +lemma DERIV_mult': + fixes f g :: "real \ 'a::real_normed_algebra" + assumes f: "DERIV f x :> D" + assumes g: "DERIV g x :> E" + shows "DERIV (\x. f x * g x) x :> f x * E + D * g x" +proof (unfold deriv_def) + from f have "isCont f x" + by (rule DERIV_isCont) + hence "(\h. f(x+h)) -- 0 --> f x" + by (simp only: isCont_iff) + hence "(\h. f(x+h) * ((g(x+h) - g x) /# h) + + ((f(x+h) - f x) /# h) * g x) + -- 0 --> f x * E + D * g x" + by (intro LIM_add LIM_mult2 LIM_const DERIV_D f g) + thus "(\h. (f(x+h) * g(x+h) - f x * g x) /# h) + -- 0 --> f x * E + D * g x" + by (simp only: DERIV_mult_lemma) +qed + +lemma DERIV_mult: + fixes f g :: "real \ 'a::{real_normed_algebra,comm_ring}" shows + "[| DERIV f x :> Da; DERIV g x :> Db |] + ==> DERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))" +by (drule (1) DERIV_mult', simp only: mult_commute add_commute) + +lemma DERIV_unique: + "[| DERIV f x :> D; DERIV f x :> E |] ==> D = E" +apply (simp add: deriv_def) +apply (blast intro: LIM_unique) +done + +text{*Differentiation of finite sum*} + +lemma DERIV_sumr [rule_format (no_asm)]: + "(\r. m \ r & r < (m + n) --> DERIV (%x. f r x) x :> (f' r x)) + --> DERIV (%x. \n=m.. (\r=m.. D) = + ((%x. (f(x)-f(a)) / (x-a)) -- a --> D)" +apply (rule iffI) +apply (drule_tac k="- a" in LIM_offset) +apply (simp add: diff_minus) +apply (drule_tac k="a" in LIM_offset) +apply (simp add: add_commute) +done + +lemma DERIV_LIM_iff': + "((%h::real. (f(a + h) - f(a)) /# h) -- 0 --> D) = + ((%x. (f(x)-f(a)) /# (x-a)) -- a --> D)" +apply (rule iffI) +apply (drule_tac k="- a" in LIM_offset) +apply (simp add: diff_minus) +apply (drule_tac k="a" in LIM_offset) +apply (simp add: add_commute) +done + +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') + +lemma inverse_diff_inverse: + "\(a::'a::division_ring) \ 0; b \ 0\ + \ inverse a - inverse b = - (inverse a * (a - b) * inverse b)" +by (simp add: right_diff_distrib left_diff_distrib mult_assoc) + +lemma DERIV_inverse_lemma: + "\a \ 0; b \ (0::'a::real_normed_div_algebra)\ + \ (inverse a - inverse b) /# h + = - (inverse a * ((a - b) /# h) * inverse b)" +by (simp add: inverse_diff_inverse) + +lemma LIM_equal2: + assumes 1: "0 < R" + assumes 2: "\x. \x \ a; norm (x - a) < R\ \ f x = g x" + shows "g -- a --> l \ f -- a --> l" +apply (unfold LIM_def, safe) +apply (drule_tac x="r" in spec, safe) +apply (rule_tac x="min s R" in exI, safe) +apply (simp add: 1) +apply (simp add: 2) +done + +lemma DERIV_inverse': + fixes f :: "real \ 'a::real_normed_div_algebra" + assumes der: "DERIV f x :> D" + assumes neq: "f x \ 0" + shows "DERIV (\x. inverse (f x)) x :> - (inverse (f x) * D * inverse (f x))" + (is "DERIV _ _ :> ?E") +proof (unfold DERIV_iff2) + from der have lim_f: "f -- x --> f x" + by (rule DERIV_isCont [unfolded isCont_def]) + + from neq have "0 < norm (f x)" by simp + with LIM_D [OF lim_f] obtain s + where s: "0 < s" + and less_fx: "\z. \z \ x; norm (z - x) < s\ + \ norm (f z - f x) < norm (f x)" + by fast + + show "(\z. (inverse (f z) - inverse (f x)) /# (z - x)) -- x --> ?E" + proof (rule LIM_equal2 [OF s]) + fix z :: real + assume "z \ x" "norm (z - x) < s" + hence "norm (f z - f x) < norm (f x)" by (rule less_fx) + hence "f z \ 0" by auto + thus "(inverse (f z) - inverse (f x)) /# (z - x) = + - (inverse (f z) * ((f z - f x) /# (z - x)) * inverse (f x))" + using neq by (rule DERIV_inverse_lemma) + next + from der have "(\z. (f z - f x) /# (z - x)) -- x --> D" + by (unfold DERIV_iff2) + thus "(\z. - (inverse (f z) * ((f z - f x) /# (z - x)) * inverse (f x))) + -- x --> ?E" + by (intro LIM_mult2 LIM_inverse LIM_minus LIM_const lim_f neq) + qed +qed + +lemma DERIV_divide: + fixes D E :: "'a::{real_normed_div_algebra,field}" + shows "\DERIV f x :> D; DERIV g x :> E; g x \ 0\ + \ DERIV (\x. f x / g x) x :> (D * g x - f x * E) / (g x * g x)" +apply (subgoal_tac "f x * - (inverse (g x) * E * inverse (g x)) + + D * inverse (g x) = (D * g x - f x * E) / (g x * g x)") +apply (erule subst) +apply (unfold divide_inverse) +apply (erule DERIV_mult') +apply (erule (1) DERIV_inverse') +apply (simp add: left_diff_distrib nonzero_inverse_mult_distrib) +apply (simp add: mult_ac) +done + +lemma DERIV_power_Suc: + fixes f :: "real \ 'a::{real_normed_algebra,recpower}" + assumes f: "DERIV f x :> D" + shows "DERIV (\x. f x ^ Suc n) x :> (of_nat n + 1) *# (D * f x ^ n)" +proof (induct n) +case 0 + show ?case by (simp add: power_Suc f) +case (Suc k) + from DERIV_mult' [OF f Suc] show ?case + apply (simp only: of_nat_Suc scaleR_left_distrib scaleR_one) + apply (simp only: power_Suc right_distrib mult_scaleR_right mult_ac) + done +qed + +lemma DERIV_power: + fixes f :: "real \ 'a::{real_normed_algebra,recpower}" + assumes f: "DERIV f x :> D" + shows "DERIV (\x. f x ^ n) x :> of_nat n *# (D * f x ^ (n - Suc 0))" +by (cases "n", simp, simp add: DERIV_power_Suc f) + + +(* ------------------------------------------------------------------------ *) +(* Caratheodory formulation of derivative at a point: standard proof *) +(* ------------------------------------------------------------------------ *) + +lemma CARAT_DERIV: + "(DERIV f x :> l) = + (\g. (\z. f z - f x = (z-x) *# g z) & isCont g x & g x = l)" + (is "?lhs = ?rhs") +proof + assume der: "DERIV f x :> l" + show "\g. (\z. f z - f x = (z-x) *# g z) \ 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 = (z-x) *# ?g z" 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 = (z-x) *# g z)" 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 DERIV_chain': + assumes f: "DERIV f x :> D" + assumes g: "DERIV g (f x) :> E" + shows "DERIV (\x. g (f x)) x :> D *# E" +proof (unfold DERIV_iff2) + obtain d where d: "\y. g y - g (f x) = (y - f x) *# d y" + and cont_d: "isCont d (f x)" and dfx: "d (f x) = E" + using CARAT_DERIV [THEN iffD1, OF g] by fast + from f have "f -- x --> f x" + by (rule DERIV_isCont [unfolded isCont_def]) + with cont_d have "(\z. d (f z)) -- x --> d (f x)" + by (rule LIM_compose) + hence "(\z. ((f z - f x) /# (z - x)) *# d (f z)) + -- x --> D *# d (f x)" + by (rule LIM_scaleR [OF f [unfolded DERIV_iff2]]) + thus "(\z. (g (f z) - g (f x)) /# (z - x)) -- x --> D *# E" + by (simp add: d dfx real_scaleR_def) +qed + + +subsubsection {* Nonstandard proofs *} + +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 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) + +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 + +text {*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: + 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{*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 = "xa - hypreal_of_real x" 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{*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) + +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 diff_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: diff_def add_ac) +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_diff_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 diff_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, assumption+) +apply (drule_tac + 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 + +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_diff_distrib [symmetric]) +apply (erule NSLIM_const [THEN NSLIM_mult]) +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 + +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 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 + +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 + +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) + +lemma NSDERIV_cmult_Id [simp]: "NSDERIV (op * c) x :> c" +by (cut_tac c = c and x = x in NSDERIV_Id [THEN NSDERIV_cmult], simp) + +(*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 (1) Infinitesimal_add_not_zero) +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 diff_def + 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 + +subsubsection {* Equivalence of NS and Standard definitions *} + +lemma divideR_eq_divide [simp]: "x /# y = x / y" +by (simp add: real_scaleR_def divide_inverse mult_commute) + +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) + +(* let's do the standard proof though theorem *) +(* LIM_mult2 follows from a NS proof *) + +lemma DERIV_cmult: + fixes f :: "real \ 'a::real_normed_algebra" shows + "DERIV f x :> D ==> DERIV (%x. c * f x) x :> c*D" +by (drule DERIV_mult' [OF DERIV_const], simp) + +(* standard version *) +lemma DERIV_chain: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (f o g) x :> Da * Db" +by (drule (1) DERIV_chain', simp add: o_def real_scaleR_def mult_commute) + +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) + +(*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 DERIV_pow: "DERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))" +apply (cut_tac DERIV_power [OF DERIV_Id]) +apply (simp add: real_scaleR_def real_of_nat_def) +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*} + +lemma DERIV_inverse: "x \ 0 ==> DERIV (%x. inverse(x)) x :> (-(inverse x ^ Suc (Suc 0)))" +by (drule DERIV_inverse' [OF DERIV_Id], simp) + +text{*Derivative of inverse*} +lemma DERIV_inverse_fun: "[| DERIV f x :> d; f(x) \ 0 |] + ==> DERIV (%x. inverse(f x)::real) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))" +by (drule (1) DERIV_inverse', simp add: mult_ac) + +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) :: real) x :> (d*g(x) - (e*f(x))) / (g(x) ^ Suc (Suc 0))" +by (drule (2) DERIV_divide, simp add: mult_commute) + +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) + +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 + real_scaleR_def mult_commute) + +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 + +subsubsection {* Differentiability predicate *} + +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) + +lemma differentiable_const: "(\z. a) differentiable x" + apply (unfold differentiable_def) + apply (rule_tac x=0 in exI) + apply simp + done + +lemma differentiable_sum: + assumes "f differentiable x" + and "g differentiable x" + shows "(\x. f x + g x) differentiable x" +proof - + from prems have "\D. DERIV f x :> D" by (unfold differentiable_def) + then obtain df where "DERIV f x :> df" .. + moreover from prems have "\D. DERIV g x :> D" by (unfold differentiable_def) + then obtain dg where "DERIV g x :> dg" .. + ultimately have "DERIV (\x. f x + g x) x :> df + dg" by (rule DERIV_add) + hence "\D. DERIV (\x. f x + g x) x :> D" by auto + thus ?thesis by (fold differentiable_def) +qed + +lemma differentiable_diff: + assumes "f differentiable x" + and "g differentiable x" + shows "(\x. f x - g x) differentiable x" +proof - + from prems have "f differentiable x" by simp + moreover + from prems have "\D. DERIV g x :> D" by (unfold differentiable_def) + then obtain dg where "DERIV g x :> dg" .. + then have "DERIV (\x. - g x) x :> -dg" by (rule DERIV_minus) + hence "\D. DERIV (\x. - g x) x :> D" by auto + hence "(\x. - g x) differentiable x" by (fold differentiable_def) + ultimately + show ?thesis + by (auto simp: real_diff_def dest: differentiable_sum) +qed + +lemma differentiable_mult: + assumes "f differentiable x" + and "g differentiable x" + shows "(\x. f x * g x) differentiable x" +proof - + from prems have "\D. DERIV f x :> D" by (unfold differentiable_def) + then obtain df where "DERIV f x :> df" .. + moreover from prems have "\D. DERIV g x :> D" by (unfold differentiable_def) + then obtain dg where "DERIV g x :> dg" .. + ultimately have "DERIV (\x. f x * g x) x :> df * g x + dg * f x" by (simp add: DERIV_mult) + hence "\D. DERIV (\x. f x * g x) x :> D" by auto + thus ?thesis by (fold differentiable_def) +qed + +subsection {*(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 + +subsection {* Nested Intervals and Bisection *} + +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 :: nat \ real)" +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 :: nat \ real)" +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: + fixes f :: "nat \ real" + shows "\\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 (drule_tac no=no and m=n in lemma_f_mono_add) +apply (auto simp add: add_commute) +apply (induct_tac "no") +apply simp +apply (auto intro: order_trans simp add: diff_minus abs_if) +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: + fixes g :: "nat \ real" + shows "\\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 :: real. 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::real. ((\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 ?Q r" and x = "d/2" in spec) +apply (drule_tac P = "%r. 0 ?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 *} + +text {*Prove Contrapositive by Bisection*} + +lemma IVT: "[| f(a::real) \ (y::real); y \ f(b); + a \ b; + (\x. a \ x & x \ b --> isCont f x) |] + ==> \x. a \ x & x \ b & f(x) = y" +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 ?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::real) \ (y::real); y \ f(a); + a \ b; + (\x. a \ x & x \ b --> isCont f x) + |] ==> \x. a \ x & x \ b & f(x) = y" +apply (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::real) \ (y::real) & y \ f(b) & a \ b & + (\x. a \ x & x \ b --> isCont f x)) + --> (\x. a \ x & x \ b & f(x) = y)" +apply (blast intro: IVT) +done + +lemma IVT2_objl: "(f(b::real) \ (y::real) & y \ f(a) & a \ b & + (\x. a \ x & x \ b --> isCont f x)) + --> (\x. a \ x & x \ b & f(x) = y)" +apply (blast intro: IVT2) +done + +text{*By bisection, function continuous on closed interval is bounded above*} + +lemma isCont_bounded: + "[| a \ b; \x. a \ x & x \ b --> isCont f x |] + ==> \M::real. \x::real. a \ x & x \ b --> f(x) \ M" +apply (cut_tac P = "% (u,v) . a \ u & u \ v & v \ b --> (\M. \x. u \ x & x \ v --> f x \ M)" in lemma_BOLZANO2) +apply safe +apply simp_all +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) +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::real. (\x::real. a \ x & x \ b --> f(x) \ M) & + (\N. N < M --> (\x. a \ x & x \ b & N < f(x)))" +apply (cut_tac S = "Collect (%y. \x. a \ x & x \ b & y = f x)" + in lemma_reals_complete) +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::real. a \ x & x \ b --> isCont f x" + shows "\M::real. (\x. a \ x & x \ b --> f(x) \ M) & + (\x. a \ x & x \ b & f(x) = M)" +proof - + from isCont_has_Ub [OF le con] + obtain M where M1: "\x. a \ x \ x \ b \ f x \ M" + and M2: "!!N. N \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::real. (\x::real. a \ x & x \ b --> M \ f(x)) & + (\x. a \ x & x \ b & f(x) = M)" +apply (subgoal_tac "\x. a \ x & x \ b --> isCont (%x. - (f x)) x") +prefer 2 apply (blast intro: isCont_minus) +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::real. (\x::real. a \ x & x \ b --> L \ f(x) & f(x) \ M) & + (\y. L \ y & y \ M --> (\x. a \ x & x \ b & (f(x) = y)))" +apply (frule isCont_eq_Lb) +apply (frule_tac [2] isCont_eq_Ub) +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 + + +text{*If @{term "0 < f'(x)"} then @{term x} is Locally Strictly Increasing At The Right*} + +lemma DERIV_left_inc: + fixes f :: "real => real" + 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 real" + 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 real" + 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: + fixes f :: "real => real" + shows "[| 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: + fixes f :: "real => real" + shows "[| 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: "ay. \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': "ay. \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: "0y. \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: + fixes f :: "real => real" + shows "[| 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: + fixes f :: "real => real" + shows "[| 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: + fixes f :: "real => real" + shows "[| 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: + fixes f :: "real => real" + shows "\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: + fixes v :: "real => real" + 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: + fixes f :: "real \ real" + 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: + fixes f g :: "real \ real" + shows "[|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: + fixes f :: "real \ real" + assumes d: "0 < d" + and inj: "\z. \z-x\ \ d --> g(f z) = z" + and cont: "\z. \z-x\ \ d --> isCont f z" + 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 "0y. \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: + fixes f g :: "real \ real" + assumes d: "0 < d" + and inj: "\z. \z-x\ \ d --> g(f z) = z" + and cont: "\z. \z-x\ \ d --> isCont f z" + 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 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 "0z. 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 + +theorem GMVT: + assumes alb: "a < b" + and fc: "\x. a \ x \ x \ b \ isCont f x" + and fd: "\x. a < x \ x < b \ f differentiable x" + and gc: "\x. a \ x \ x \ b \ isCont g x" + and gd: "\x. a < x \ x < b \ g differentiable x" + shows "\g'c f'c c. DERIV g c :> g'c \ DERIV f c :> f'c \ a < c \ c < b \ ((f b - f a) * g'c) = ((g b - g a) * f'c)" +proof - + let ?h = "\x. (f b - f a)*(g x) - (g b - g a)*(f x)" + from prems have "a < b" by simp + moreover have "\x. a \ x \ x \ b \ isCont ?h x" + proof - + have "\x. a <= x \ x <= b \ isCont (\x. f b - f a) x" by simp + with gc have "\x. a <= x \ x <= b \ isCont (\x. (f b - f a) * g x) x" + by (auto intro: isCont_mult) + moreover + have "\x. a <= x \ x <= b \ isCont (\x. g b - g a) x" by simp + with fc have "\x. a <= x \ x <= b \ isCont (\x. (g b - g a) * f x) x" + by (auto intro: isCont_mult) + ultimately show ?thesis + by (fastsimp intro: isCont_diff) + qed + moreover + have "\x. a < x \ x < b \ ?h differentiable x" + proof - + have "\x. a < x \ x < b \ (\x. f b - f a) differentiable x" by (simp add: differentiable_const) + with gd have "\x. a < x \ x < b \ (\x. (f b - f a) * g x) differentiable x" by (simp add: differentiable_mult) + moreover + have "\x. a < x \ x < b \ (\x. g b - g a) differentiable x" by (simp add: differentiable_const) + with fd have "\x. a < x \ x < b \ (\x. (g b - g a) * f x) differentiable x" by (simp add: differentiable_mult) + ultimately show ?thesis by (simp add: differentiable_diff) + qed + ultimately have "\l z. a < z \ z < b \ DERIV ?h z :> l \ ?h b - ?h a = (b - a) * l" by (rule MVT) + then obtain l where ldef: "\z. a < z \ z < b \ DERIV ?h z :> l \ ?h b - ?h a = (b - a) * l" .. + then obtain c where cdef: "a < c \ c < b \ DERIV ?h c :> l \ ?h b - ?h a = (b - a) * l" .. + + from cdef have cint: "a < c \ c < b" by auto + with gd have "g differentiable c" by simp + hence "\D. DERIV g c :> D" by (rule differentiableD) + then obtain g'c where g'cdef: "DERIV g c :> g'c" .. + + from cdef have "a < c \ c < b" by auto + with fd have "f differentiable c" by simp + hence "\D. DERIV f c :> D" by (rule differentiableD) + then obtain f'c where f'cdef: "DERIV f c :> f'c" .. + + from cdef have "DERIV ?h c :> l" by auto + moreover + { + from g'cdef have "DERIV (\x. (f b - f a) * g x) c :> g'c * (f b - f a)" + apply (insert DERIV_const [where k="f b - f a"]) + apply (drule meta_spec [of _ c]) + apply (drule DERIV_mult [where f="(\x. f b - f a)" and g=g]) + by simp_all + moreover from f'cdef have "DERIV (\x. (g b - g a) * f x) c :> f'c * (g b - g a)" + apply (insert DERIV_const [where k="g b - g a"]) + apply (drule meta_spec [of _ c]) + apply (drule DERIV_mult [where f="(\x. g b - g a)" and g=f]) + by simp_all + ultimately have "DERIV ?h c :> g'c * (f b - f a) - f'c * (g b - g a)" + by (simp add: DERIV_diff) + } + ultimately have leq: "l = g'c * (f b - f a) - f'c * (g b - g a)" by (rule DERIV_unique) + + { + from cdef have "?h b - ?h a = (b - a) * l" by auto + also with leq have "\ = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp + finally have "?h b - ?h a = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp + } + moreover + { + have "?h b - ?h a = + ((f b)*(g b) - (f a)*(g b) - (g b)*(f b) + (g a)*(f b)) - + ((f b)*(g a) - (f a)*(g a) - (g b)*(f a) + (g a)*(f a))" + by (simp add: mult_ac add_ac real_diff_mult_distrib) + hence "?h b - ?h a = 0" by auto + } + ultimately have "(b - a) * (g'c * (f b - f a) - f'c * (g b - g a)) = 0" by auto + with alb have "g'c * (f b - f a) - f'c * (g b - g a) = 0" by simp + hence "g'c * (f b - f a) = f'c * (g b - g a)" by simp + hence "(f b - f a) * g'c = (g b - g a) * f'c" by (simp add: mult_ac) + + with g'cdef f'cdef cint show ?thesis by auto +qed + +end diff -r 6860f161111c -r 0742fc979c67 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Fri Nov 03 15:28:13 2006 +0100 +++ b/src/HOL/IsaMakefile Sat Nov 04 00:11:11 2006 +0100 @@ -177,7 +177,7 @@ Hyperreal/Ln.thy Hyperreal/MacLaurin.thy Hyperreal/NatStar.thy \ Hyperreal/NSA.thy Hyperreal/NthRoot.thy Hyperreal/Poly.thy \ Hyperreal/SEQ.thy Hyperreal/Series.thy Hyperreal/Star.thy \ - Hyperreal/Taylor.thy \ + Hyperreal/Taylor.thy Hyperreal/Deriv.thy \ Hyperreal/Transcendental.thy Hyperreal/fuf.ML Hyperreal/hypreal_arith.ML \ Complex/Complex_Main.thy Complex/CLim.thy Complex/CSeries.thy \ Complex/CStar.thy Complex/Complex.thy Complex/ComplexBin.thy \