huffman@21164: (* Title : Deriv.thy huffman@21164: ID : $Id$ huffman@21164: Author : Jacques D. Fleuriot huffman@21164: Copyright : 1998 University of Cambridge huffman@21164: Conversion to Isar and new proofs by Lawrence C Paulson, 2004 huffman@21164: GMVT by Benjamin Porter, 2005 huffman@21164: *) huffman@21164: huffman@21164: header{* Differentiation *} huffman@21164: huffman@21164: theory Deriv huffman@21164: imports Lim huffman@21164: begin huffman@21164: huffman@21164: text{*Standard and Nonstandard Definitions*} huffman@21164: huffman@21164: definition huffman@21164: deriv :: "[real \ 'a::real_normed_vector, real, 'a] \ bool" huffman@21164: --{*Differentiation: D is derivative of function f at x*} wenzelm@21404: ("(DERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where huffman@21164: "DERIV f x :> D = ((%h. (f(x + h) - f x) /# h) -- 0 --> D)" huffman@21164: wenzelm@21404: definition huffman@21164: nsderiv :: "[real=>real,real,real] => bool" wenzelm@21404: ("(NSDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where huffman@21164: "NSDERIV f x :> D = (\h \ Infinitesimal - {0}. huffman@21164: (( *f* f)(hypreal_of_real x + h) huffman@21164: - hypreal_of_real (f x))/h @= hypreal_of_real D)" huffman@21164: wenzelm@21404: definition wenzelm@21404: differentiable :: "[real=>real,real] => bool" (infixl "differentiable" 60) where huffman@21164: "f differentiable x = (\D. DERIV f x :> D)" huffman@21164: wenzelm@21404: definition huffman@21164: NSdifferentiable :: "[real=>real,real] => bool" wenzelm@21404: (infixl "NSdifferentiable" 60) where huffman@21164: "f NSdifferentiable x = (\D. NSDERIV f x :> D)" huffman@21164: wenzelm@21404: definition wenzelm@21404: increment :: "[real=>real,real,hypreal] => hypreal" where huffman@21164: "increment f x h = (@inc. f NSdifferentiable x & huffman@21164: inc = ( *f* f)(hypreal_of_real x + h) - hypreal_of_real (f x))" huffman@21164: huffman@21164: huffman@21164: consts huffman@21164: Bolzano_bisect :: "[real*real=>bool, real, real, nat] => (real*real)" huffman@21164: primrec huffman@21164: "Bolzano_bisect P a b 0 = (a,b)" huffman@21164: "Bolzano_bisect P a b (Suc n) = huffman@21164: (let (x,y) = Bolzano_bisect P a b n huffman@21164: in if P(x, (x+y)/2) then ((x+y)/2, y) huffman@21164: else (x, (x+y)/2))" huffman@21164: huffman@21164: huffman@21164: subsection {* Derivatives *} huffman@21164: huffman@21164: subsubsection {* Purely standard proofs *} huffman@21164: huffman@21164: lemma DERIV_iff: "(DERIV f x :> D) = ((%h. (f(x + h) - f(x))/#h) -- 0 --> D)" huffman@21164: by (simp add: deriv_def) huffman@21164: huffman@21164: lemma DERIV_D: "DERIV f x :> D ==> (%h. (f(x + h) - f(x))/#h) -- 0 --> D" huffman@21164: by (simp add: deriv_def) huffman@21164: huffman@21164: lemma DERIV_const [simp]: "DERIV (\x. k) x :> 0" huffman@21164: by (simp add: deriv_def) huffman@21164: huffman@21164: lemma DERIV_Id [simp]: "DERIV (\x. x) x :> 1" huffman@21164: by (simp add: deriv_def real_scaleR_def cong: LIM_cong) huffman@21164: huffman@21164: lemma add_diff_add: huffman@21164: fixes a b c d :: "'a::ab_group_add" huffman@21164: shows "(a + c) - (b + d) = (a - b) + (c - d)" huffman@21164: by simp huffman@21164: huffman@21164: lemma DERIV_add: huffman@21164: "\DERIV f x :> D; DERIV g x :> E\ \ DERIV (\x. f x + g x) x :> D + E" huffman@21164: by (simp only: deriv_def add_diff_add scaleR_right_distrib LIM_add) huffman@21164: huffman@21164: lemma DERIV_minus: huffman@21164: "DERIV f x :> D \ DERIV (\x. - f x) x :> - D" huffman@21164: by (simp only: deriv_def minus_diff_minus scaleR_minus_right LIM_minus) huffman@21164: huffman@21164: lemma DERIV_diff: huffman@21164: "\DERIV f x :> D; DERIV g x :> E\ \ DERIV (\x. f x - g x) x :> D - E" huffman@21164: by (simp only: diff_def DERIV_add DERIV_minus) huffman@21164: huffman@21164: lemma DERIV_add_minus: huffman@21164: "\DERIV f x :> D; DERIV g x :> E\ \ DERIV (\x. f x + - g x) x :> D + - E" huffman@21164: by (simp only: DERIV_add DERIV_minus) huffman@21164: huffman@21164: lemma DERIV_isCont: "DERIV f x :> D \ isCont f x" huffman@21164: proof (unfold isCont_iff) huffman@21164: assume "DERIV f x :> D" huffman@21164: hence "(\h. (f(x+h) - f(x)) /# h) -- 0 --> D" huffman@21164: by (rule DERIV_D) huffman@21164: hence "(\h. h *# ((f(x+h) - f(x)) /# h)) -- 0 --> 0 *# D" huffman@21164: by (intro LIM_scaleR LIM_self) huffman@21164: hence "(\h. (f(x+h) - f(x))) -- 0 --> 0" huffman@21164: by (simp cong: LIM_cong) huffman@21164: thus "(\h. f(x+h)) -- 0 --> f(x)" huffman@21164: by (simp add: LIM_def) huffman@21164: qed huffman@21164: huffman@21164: lemma DERIV_mult_lemma: huffman@21164: fixes a b c d :: "'a::real_algebra" huffman@21164: shows "(a * b - c * d) /# h = a * ((b - d) /# h) + ((a - c) /# h) * d" huffman@21164: by (simp add: diff_minus scaleR_right_distrib [symmetric] ring_distrib) huffman@21164: huffman@21164: lemma DERIV_mult': huffman@21164: fixes f g :: "real \ 'a::real_normed_algebra" huffman@21164: assumes f: "DERIV f x :> D" huffman@21164: assumes g: "DERIV g x :> E" huffman@21164: shows "DERIV (\x. f x * g x) x :> f x * E + D * g x" huffman@21164: proof (unfold deriv_def) huffman@21164: from f have "isCont f x" huffman@21164: by (rule DERIV_isCont) huffman@21164: hence "(\h. f(x+h)) -- 0 --> f x" huffman@21164: by (simp only: isCont_iff) huffman@21164: hence "(\h. f(x+h) * ((g(x+h) - g x) /# h) + huffman@21164: ((f(x+h) - f x) /# h) * g x) huffman@21164: -- 0 --> f x * E + D * g x" huffman@21164: by (intro LIM_add LIM_mult2 LIM_const DERIV_D f g) huffman@21164: thus "(\h. (f(x+h) * g(x+h) - f x * g x) /# h) huffman@21164: -- 0 --> f x * E + D * g x" huffman@21164: by (simp only: DERIV_mult_lemma) huffman@21164: qed huffman@21164: huffman@21164: lemma DERIV_mult: huffman@21164: fixes f g :: "real \ 'a::{real_normed_algebra,comm_ring}" shows huffman@21164: "[| DERIV f x :> Da; DERIV g x :> Db |] huffman@21164: ==> DERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))" huffman@21164: by (drule (1) DERIV_mult', simp only: mult_commute add_commute) huffman@21164: huffman@21164: lemma DERIV_unique: huffman@21164: "[| DERIV f x :> D; DERIV f x :> E |] ==> D = E" huffman@21164: apply (simp add: deriv_def) huffman@21164: apply (blast intro: LIM_unique) huffman@21164: done huffman@21164: huffman@21164: text{*Differentiation of finite sum*} huffman@21164: huffman@21164: lemma DERIV_sumr [rule_format (no_asm)]: huffman@21164: "(\r. m \ r & r < (m + n) --> DERIV (%x. f r x) x :> (f' r x)) huffman@21164: --> DERIV (%x. \n=m.. (\r=m.. D) = huffman@21164: ((%x. (f(x)-f(a)) / (x-a)) -- a --> D)" huffman@21164: apply (rule iffI) huffman@21164: apply (drule_tac k="- a" in LIM_offset) huffman@21164: apply (simp add: diff_minus) huffman@21164: apply (drule_tac k="a" in LIM_offset) huffman@21164: apply (simp add: add_commute) huffman@21164: done huffman@21164: huffman@21164: lemma DERIV_LIM_iff': huffman@21164: "((%h::real. (f(a + h) - f(a)) /# h) -- 0 --> D) = huffman@21164: ((%x. (f(x)-f(a)) /# (x-a)) -- a --> D)" huffman@21164: apply (rule iffI) huffman@21164: apply (drule_tac k="- a" in LIM_offset) huffman@21164: apply (simp add: diff_minus) huffman@21164: apply (drule_tac k="a" in LIM_offset) huffman@21164: apply (simp add: add_commute) huffman@21164: done huffman@21164: huffman@21164: lemma DERIV_iff2: "(DERIV f x :> D) = ((%z. (f(z) - f(x)) /# (z-x)) -- x --> D)" huffman@21164: by (simp add: deriv_def diff_minus [symmetric] DERIV_LIM_iff') huffman@21164: huffman@21164: lemma inverse_diff_inverse: huffman@21164: "\(a::'a::division_ring) \ 0; b \ 0\ huffman@21164: \ inverse a - inverse b = - (inverse a * (a - b) * inverse b)" huffman@21164: by (simp add: right_diff_distrib left_diff_distrib mult_assoc) huffman@21164: huffman@21164: lemma DERIV_inverse_lemma: huffman@21164: "\a \ 0; b \ (0::'a::real_normed_div_algebra)\ huffman@21164: \ (inverse a - inverse b) /# h huffman@21164: = - (inverse a * ((a - b) /# h) * inverse b)" huffman@21164: by (simp add: inverse_diff_inverse) huffman@21164: huffman@21164: lemma LIM_equal2: huffman@21164: assumes 1: "0 < R" huffman@21164: assumes 2: "\x. \x \ a; norm (x - a) < R\ \ f x = g x" huffman@21164: shows "g -- a --> l \ f -- a --> l" huffman@21164: apply (unfold LIM_def, safe) huffman@21164: apply (drule_tac x="r" in spec, safe) huffman@21164: apply (rule_tac x="min s R" in exI, safe) huffman@21164: apply (simp add: 1) huffman@21164: apply (simp add: 2) huffman@21164: done huffman@21164: huffman@21164: lemma DERIV_inverse': huffman@21164: fixes f :: "real \ 'a::real_normed_div_algebra" huffman@21164: assumes der: "DERIV f x :> D" huffman@21164: assumes neq: "f x \ 0" huffman@21164: shows "DERIV (\x. inverse (f x)) x :> - (inverse (f x) * D * inverse (f x))" huffman@21164: (is "DERIV _ _ :> ?E") huffman@21164: proof (unfold DERIV_iff2) huffman@21164: from der have lim_f: "f -- x --> f x" huffman@21164: by (rule DERIV_isCont [unfolded isCont_def]) huffman@21164: huffman@21164: from neq have "0 < norm (f x)" by simp huffman@21164: with LIM_D [OF lim_f] obtain s huffman@21164: where s: "0 < s" huffman@21164: and less_fx: "\z. \z \ x; norm (z - x) < s\ huffman@21164: \ norm (f z - f x) < norm (f x)" huffman@21164: by fast huffman@21164: huffman@21164: show "(\z. (inverse (f z) - inverse (f x)) /# (z - x)) -- x --> ?E" huffman@21164: proof (rule LIM_equal2 [OF s]) huffman@21164: fix z :: real huffman@21164: assume "z \ x" "norm (z - x) < s" huffman@21164: hence "norm (f z - f x) < norm (f x)" by (rule less_fx) huffman@21164: hence "f z \ 0" by auto huffman@21164: thus "(inverse (f z) - inverse (f x)) /# (z - x) = huffman@21164: - (inverse (f z) * ((f z - f x) /# (z - x)) * inverse (f x))" huffman@21164: using neq by (rule DERIV_inverse_lemma) huffman@21164: next huffman@21164: from der have "(\z. (f z - f x) /# (z - x)) -- x --> D" huffman@21164: by (unfold DERIV_iff2) huffman@21164: thus "(\z. - (inverse (f z) * ((f z - f x) /# (z - x)) * inverse (f x))) huffman@21164: -- x --> ?E" huffman@21164: by (intro LIM_mult2 LIM_inverse LIM_minus LIM_const lim_f neq) huffman@21164: qed huffman@21164: qed huffman@21164: huffman@21164: lemma DERIV_divide: huffman@21164: fixes D E :: "'a::{real_normed_div_algebra,field}" huffman@21164: shows "\DERIV f x :> D; DERIV g x :> E; g x \ 0\ huffman@21164: \ DERIV (\x. f x / g x) x :> (D * g x - f x * E) / (g x * g x)" huffman@21164: apply (subgoal_tac "f x * - (inverse (g x) * E * inverse (g x)) + huffman@21164: D * inverse (g x) = (D * g x - f x * E) / (g x * g x)") huffman@21164: apply (erule subst) huffman@21164: apply (unfold divide_inverse) huffman@21164: apply (erule DERIV_mult') huffman@21164: apply (erule (1) DERIV_inverse') huffman@21164: apply (simp add: left_diff_distrib nonzero_inverse_mult_distrib) huffman@21164: apply (simp add: mult_ac) huffman@21164: done huffman@21164: huffman@21164: lemma DERIV_power_Suc: krauss@21199: fixes f :: "real \ 'a::{real_normed_algebra,recpower,comm_monoid_mult}" huffman@21164: assumes f: "DERIV f x :> D" huffman@21164: shows "DERIV (\x. f x ^ Suc n) x :> (of_nat n + 1) *# (D * f x ^ n)" huffman@21164: proof (induct n) huffman@21164: case 0 huffman@21164: show ?case by (simp add: power_Suc f) huffman@21164: case (Suc k) huffman@21164: from DERIV_mult' [OF f Suc] show ?case huffman@21164: apply (simp only: of_nat_Suc scaleR_left_distrib scaleR_one) huffman@21164: apply (simp only: power_Suc right_distrib mult_scaleR_right mult_ac) huffman@21164: done huffman@21164: qed huffman@21164: huffman@21164: lemma DERIV_power: krauss@21199: fixes f :: "real \ 'a::{real_normed_algebra,recpower,comm_monoid_mult}" huffman@21164: assumes f: "DERIV f x :> D" huffman@21164: shows "DERIV (\x. f x ^ n) x :> of_nat n *# (D * f x ^ (n - Suc 0))" huffman@21164: by (cases "n", simp, simp add: DERIV_power_Suc f) huffman@21164: huffman@21164: huffman@21164: (* ------------------------------------------------------------------------ *) huffman@21164: (* Caratheodory formulation of derivative at a point: standard proof *) huffman@21164: (* ------------------------------------------------------------------------ *) huffman@21164: huffman@21164: lemma CARAT_DERIV: huffman@21164: "(DERIV f x :> l) = huffman@21164: (\g. (\z. f z - f x = (z-x) *# g z) & isCont g x & g x = l)" huffman@21164: (is "?lhs = ?rhs") huffman@21164: proof huffman@21164: assume der: "DERIV f x :> l" huffman@21164: show "\g. (\z. f z - f x = (z-x) *# g z) \ isCont g x \ g x = l" huffman@21164: proof (intro exI conjI) huffman@21164: let ?g = "(%z. if z = x then l else (f z - f x) /# (z-x))" huffman@21164: show "\z. f z - f x = (z-x) *# ?g z" by (simp) huffman@21164: show "isCont ?g x" using der huffman@21164: by (simp add: isCont_iff DERIV_iff diff_minus huffman@21164: cong: LIM_equal [rule_format]) huffman@21164: show "?g x = l" by simp huffman@21164: qed huffman@21164: next huffman@21164: assume "?rhs" huffman@21164: then obtain g where huffman@21164: "(\z. f z - f x = (z-x) *# g z)" and "isCont g x" and "g x = l" by blast huffman@21164: thus "(DERIV f x :> l)" huffman@21164: by (auto simp add: isCont_iff DERIV_iff diff_minus huffman@21164: cong: LIM_equal [rule_format]) huffman@21164: qed huffman@21164: huffman@21164: lemma DERIV_chain': huffman@21164: assumes f: "DERIV f x :> D" huffman@21164: assumes g: "DERIV g (f x) :> E" huffman@21164: shows "DERIV (\x. g (f x)) x :> D *# E" huffman@21164: proof (unfold DERIV_iff2) huffman@21164: obtain d where d: "\y. g y - g (f x) = (y - f x) *# d y" huffman@21164: and cont_d: "isCont d (f x)" and dfx: "d (f x) = E" huffman@21164: using CARAT_DERIV [THEN iffD1, OF g] by fast huffman@21164: from f have "f -- x --> f x" huffman@21164: by (rule DERIV_isCont [unfolded isCont_def]) huffman@21164: with cont_d have "(\z. d (f z)) -- x --> d (f x)" huffman@21239: by (rule isCont_LIM_compose) huffman@21164: hence "(\z. ((f z - f x) /# (z - x)) *# d (f z)) huffman@21164: -- x --> D *# d (f x)" huffman@21164: by (rule LIM_scaleR [OF f [unfolded DERIV_iff2]]) huffman@21164: thus "(\z. (g (f z) - g (f x)) /# (z - x)) -- x --> D *# E" huffman@21164: by (simp add: d dfx real_scaleR_def) huffman@21164: qed huffman@21164: huffman@21164: huffman@21164: subsubsection {* Nonstandard proofs *} huffman@21164: huffman@21164: lemma DERIV_NS_iff: huffman@21164: "(DERIV f x :> D) = ((%h. (f(x + h) - f(x))/#h) -- 0 --NS> D)" huffman@21164: by (simp add: deriv_def LIM_NSLIM_iff) huffman@21164: huffman@21164: lemma NS_DERIV_D: "DERIV f x :> D ==> (%h. (f(x + h) - f(x))/#h) -- 0 --NS> D" huffman@21164: by (simp add: deriv_def LIM_NSLIM_iff) huffman@21164: huffman@21164: lemma NSDeriv_unique: huffman@21164: "[| NSDERIV f x :> D; NSDERIV f x :> E |] ==> D = E" huffman@21164: apply (simp add: nsderiv_def) huffman@21164: apply (cut_tac Infinitesimal_epsilon hypreal_epsilon_not_zero) huffman@21164: apply (auto dest!: bspec [where x=epsilon] huffman@21164: intro!: inj_hypreal_of_real [THEN injD] huffman@21164: dest: approx_trans3) huffman@21164: done huffman@21164: huffman@21164: text {*First NSDERIV in terms of NSLIM*} huffman@21164: huffman@21164: text{*first equivalence *} huffman@21164: lemma NSDERIV_NSLIM_iff: huffman@21164: "(NSDERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --NS> D)" huffman@21164: apply (simp add: nsderiv_def NSLIM_def, auto) huffman@21164: apply (drule_tac x = xa in bspec) huffman@21164: apply (rule_tac [3] ccontr) huffman@21164: apply (drule_tac [3] x = h in spec) huffman@21164: apply (auto simp add: mem_infmal_iff starfun_lambda_cancel) huffman@21164: done huffman@21164: huffman@21164: text{*second equivalence *} huffman@21164: lemma NSDERIV_NSLIM_iff2: huffman@21164: "(NSDERIV f x :> D) = ((%z. (f(z) - f(x)) / (z-x)) -- x --NS> D)" huffman@21164: by (simp add: NSDERIV_NSLIM_iff DERIV_LIM_iff diff_minus [symmetric] huffman@21164: LIM_NSLIM_iff [symmetric]) huffman@21164: huffman@21164: (* while we're at it! *) huffman@21164: lemma NSDERIV_iff2: huffman@21164: "(NSDERIV f x :> D) = huffman@21164: (\w. huffman@21164: w \ hypreal_of_real x & w \ hypreal_of_real x --> huffman@21164: ( *f* (%z. (f z - f x) / (z-x))) w \ hypreal_of_real D)" huffman@21164: by (simp add: NSDERIV_NSLIM_iff2 NSLIM_def) huffman@21164: huffman@21164: (*FIXME DELETE*) huffman@21164: lemma hypreal_not_eq_minus_iff: "(x \ a) = (x - a \ (0::hypreal))" huffman@21164: by (auto dest: hypreal_eq_minus_iff [THEN iffD2]) huffman@21164: huffman@21164: lemma NSDERIVD5: huffman@21164: "(NSDERIV f x :> D) ==> huffman@21164: (\u. u \ hypreal_of_real x --> huffman@21164: ( *f* (%z. f z - f x)) u \ hypreal_of_real D * (u - hypreal_of_real x))" huffman@21164: apply (auto simp add: NSDERIV_iff2) huffman@21164: apply (case_tac "u = hypreal_of_real x", auto) huffman@21164: apply (drule_tac x = u in spec, auto) huffman@21164: apply (drule_tac c = "u - hypreal_of_real x" and b = "hypreal_of_real D" in approx_mult1) huffman@21164: apply (drule_tac [!] hypreal_not_eq_minus_iff [THEN iffD1]) huffman@21164: apply (subgoal_tac [2] "( *f* (%z. z-x)) u \ (0::hypreal) ") huffman@21164: apply (auto simp add: huffman@21164: approx_minus_iff [THEN iffD1, THEN mem_infmal_iff [THEN iffD2]] huffman@21164: Infinitesimal_subset_HFinite [THEN subsetD]) huffman@21164: done huffman@21164: huffman@21164: lemma NSDERIVD4: huffman@21164: "(NSDERIV f x :> D) ==> huffman@21164: (\h \ Infinitesimal. huffman@21164: (( *f* f)(hypreal_of_real x + h) - huffman@21164: hypreal_of_real (f x))\ (hypreal_of_real D) * h)" huffman@21164: apply (auto simp add: nsderiv_def) huffman@21164: apply (case_tac "h = (0::hypreal) ") huffman@21164: apply (auto simp add: diff_minus) huffman@21164: apply (drule_tac x = h in bspec) huffman@21164: apply (drule_tac [2] c = h in approx_mult1) huffman@21164: apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] huffman@21164: simp add: diff_minus) huffman@21164: done huffman@21164: huffman@21164: lemma NSDERIVD3: huffman@21164: "(NSDERIV f x :> D) ==> huffman@21164: (\h \ Infinitesimal - {0}. huffman@21164: (( *f* f)(hypreal_of_real x + h) - huffman@21164: hypreal_of_real (f x))\ (hypreal_of_real D) * h)" huffman@21164: apply (auto simp add: nsderiv_def) huffman@21164: apply (rule ccontr, drule_tac x = h in bspec) huffman@21164: apply (drule_tac [2] c = h in approx_mult1) huffman@21164: apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] huffman@21164: simp add: mult_assoc diff_minus) huffman@21164: done huffman@21164: huffman@21164: text{*Differentiability implies continuity huffman@21164: nice and simple "algebraic" proof*} huffman@21164: lemma NSDERIV_isNSCont: "NSDERIV f x :> D ==> isNSCont f x" huffman@21164: apply (auto simp add: nsderiv_def isNSCont_NSLIM_iff NSLIM_def) huffman@21164: apply (drule approx_minus_iff [THEN iffD1]) huffman@21164: apply (drule hypreal_not_eq_minus_iff [THEN iffD1]) huffman@21164: apply (drule_tac x = "xa - hypreal_of_real x" in bspec) huffman@21164: prefer 2 apply (simp add: add_assoc [symmetric]) huffman@21164: apply (auto simp add: mem_infmal_iff [symmetric] add_commute) huffman@21164: apply (drule_tac c = "xa - hypreal_of_real x" in approx_mult1) huffman@21164: apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] huffman@21164: simp add: mult_assoc) huffman@21164: apply (drule_tac x3=D in huffman@21164: HFinite_hypreal_of_real [THEN [2] Infinitesimal_HFinite_mult, huffman@21164: THEN mem_infmal_iff [THEN iffD1]]) huffman@21164: apply (auto simp add: mult_commute huffman@21164: intro: approx_trans approx_minus_iff [THEN iffD2]) huffman@21164: done huffman@21164: huffman@21164: text{*Differentiation rules for combinations of functions huffman@21164: follow from clear, straightforard, algebraic huffman@21164: manipulations*} huffman@21164: text{*Constant function*} huffman@21164: huffman@21164: (* use simple constant nslimit theorem *) huffman@21164: lemma NSDERIV_const [simp]: "(NSDERIV (%x. k) x :> 0)" huffman@21164: by (simp add: NSDERIV_NSLIM_iff) huffman@21164: huffman@21164: text{*Sum of functions- proved easily*} huffman@21164: huffman@21164: lemma NSDERIV_add: "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] huffman@21164: ==> NSDERIV (%x. f x + g x) x :> Da + Db" huffman@21164: apply (auto simp add: NSDERIV_NSLIM_iff NSLIM_def) huffman@21164: apply (auto simp add: add_divide_distrib diff_divide_distrib dest!: spec) huffman@21164: apply (drule_tac b = "hypreal_of_real Da" and d = "hypreal_of_real Db" in approx_add) huffman@21164: apply (auto simp add: diff_def add_ac) huffman@21164: done huffman@21164: huffman@21164: text{*Product of functions - Proof is trivial but tedious huffman@21164: and long due to rearrangement of terms*} huffman@21164: huffman@21164: lemma lemma_nsderiv1: "((a::hypreal)*b) - (c*d) = (b*(a - c)) + (c*(b - d))" huffman@21164: by (simp add: right_diff_distrib) huffman@21164: huffman@21164: lemma lemma_nsderiv2: "[| (x - y) / z = hypreal_of_real D + yb; z \ 0; huffman@21164: z \ Infinitesimal; yb \ Infinitesimal |] huffman@21164: ==> x - y \ 0" huffman@21164: apply (frule_tac c1 = z in hypreal_mult_right_cancel [THEN iffD2], assumption) huffman@21164: apply (erule_tac V = "(x - y) / z = hypreal_of_real D + yb" in thin_rl) huffman@21164: apply (auto intro!: Infinitesimal_HFinite_mult2 HFinite_add huffman@21164: simp add: mult_assoc mem_infmal_iff [symmetric]) huffman@21164: apply (erule Infinitesimal_subset_HFinite [THEN subsetD]) huffman@21164: done huffman@21164: huffman@21164: lemma NSDERIV_mult: "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] huffman@21164: ==> NSDERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))" huffman@21164: apply (auto simp add: NSDERIV_NSLIM_iff NSLIM_def) huffman@21164: apply (auto dest!: spec huffman@21164: simp add: starfun_lambda_cancel lemma_nsderiv1) huffman@21164: apply (simp (no_asm) add: add_divide_distrib diff_divide_distrib) huffman@21164: apply (drule bex_Infinitesimal_iff2 [THEN iffD2])+ huffman@21164: apply (auto simp add: times_divide_eq_right [symmetric] huffman@21164: simp del: times_divide_eq) huffman@21164: apply (drule_tac D = Db in lemma_nsderiv2, assumption+) huffman@21164: apply (drule_tac huffman@21164: approx_minus_iff [THEN iffD2, THEN bex_Infinitesimal_iff2 [THEN iffD2]]) huffman@21164: apply (auto intro!: approx_add_mono1 huffman@21164: simp add: left_distrib right_distrib mult_commute add_assoc) huffman@21164: apply (rule_tac b1 = "hypreal_of_real Db * hypreal_of_real (f x)" huffman@21164: in add_commute [THEN subst]) huffman@21164: apply (auto intro!: Infinitesimal_add_approx_self2 [THEN approx_sym] huffman@21164: Infinitesimal_add Infinitesimal_mult huffman@21164: Infinitesimal_hypreal_of_real_mult huffman@21164: Infinitesimal_hypreal_of_real_mult2 huffman@21164: simp add: add_assoc [symmetric]) huffman@21164: done huffman@21164: huffman@21164: text{*Multiplying by a constant*} huffman@21164: lemma NSDERIV_cmult: "NSDERIV f x :> D huffman@21164: ==> NSDERIV (%x. c * f x) x :> c*D" huffman@21164: apply (simp only: times_divide_eq_right [symmetric] NSDERIV_NSLIM_iff huffman@21164: minus_mult_right right_diff_distrib [symmetric]) huffman@21164: apply (erule NSLIM_const [THEN NSLIM_mult]) huffman@21164: done huffman@21164: huffman@21164: text{*Negation of function*} huffman@21164: lemma NSDERIV_minus: "NSDERIV f x :> D ==> NSDERIV (%x. -(f x)) x :> -D" huffman@21164: proof (simp add: NSDERIV_NSLIM_iff) huffman@21164: assume "(\h. (f (x + h) - f x) / h) -- 0 --NS> D" huffman@21164: hence deriv: "(\h. - ((f(x+h) - f x) / h)) -- 0 --NS> - D" huffman@21164: by (rule NSLIM_minus) huffman@21164: have "\h. - ((f (x + h) - f x) / h) = (- f (x + h) + f x) / h" huffman@21164: by (simp add: minus_divide_left) huffman@21164: with deriv huffman@21164: show "(\h. (- f (x + h) + f x) / h) -- 0 --NS> - D" by simp huffman@21164: qed huffman@21164: huffman@21164: text{*Subtraction*} huffman@21164: lemma NSDERIV_add_minus: "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] ==> NSDERIV (%x. f x + -g x) x :> Da + -Db" huffman@21164: by (blast dest: NSDERIV_add NSDERIV_minus) huffman@21164: huffman@21164: lemma NSDERIV_diff: huffman@21164: "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] huffman@21164: ==> NSDERIV (%x. f x - g x) x :> Da-Db" huffman@21164: apply (simp add: diff_minus) huffman@21164: apply (blast intro: NSDERIV_add_minus) huffman@21164: done huffman@21164: huffman@21164: text{* Similarly to the above, the chain rule admits an entirely huffman@21164: straightforward derivation. Compare this with Harrison's huffman@21164: HOL proof of the chain rule, which proved to be trickier and huffman@21164: required an alternative characterisation of differentiability- huffman@21164: the so-called Carathedory derivative. Our main problem is huffman@21164: manipulation of terms.*} huffman@21164: huffman@21164: huffman@21164: (* lemmas *) huffman@21164: lemma NSDERIV_zero: huffman@21164: "[| NSDERIV g x :> D; huffman@21164: ( *f* g) (hypreal_of_real(x) + xa) = hypreal_of_real(g x); huffman@21164: xa \ Infinitesimal; huffman@21164: xa \ 0 huffman@21164: |] ==> D = 0" huffman@21164: apply (simp add: nsderiv_def) huffman@21164: apply (drule bspec, auto) huffman@21164: done huffman@21164: huffman@21164: (* can be proved differently using NSLIM_isCont_iff *) huffman@21164: lemma NSDERIV_approx: huffman@21164: "[| NSDERIV f x :> D; h \ Infinitesimal; h \ 0 |] huffman@21164: ==> ( *f* f) (hypreal_of_real(x) + h) - hypreal_of_real(f x) \ 0" huffman@21164: apply (simp add: nsderiv_def) huffman@21164: apply (simp add: mem_infmal_iff [symmetric]) huffman@21164: apply (rule Infinitesimal_ratio) huffman@21164: apply (rule_tac [3] approx_hypreal_of_real_HFinite, auto) huffman@21164: done huffman@21164: huffman@21164: (*--------------------------------------------------------------- huffman@21164: from one version of differentiability huffman@21164: huffman@21164: f(x) - f(a) huffman@21164: --------------- \ Db huffman@21164: x - a huffman@21164: ---------------------------------------------------------------*) huffman@21164: lemma NSDERIVD1: "[| NSDERIV f (g x) :> Da; huffman@21164: ( *f* g) (hypreal_of_real(x) + xa) \ hypreal_of_real (g x); huffman@21164: ( *f* g) (hypreal_of_real(x) + xa) \ hypreal_of_real (g x) huffman@21164: |] ==> (( *f* f) (( *f* g) (hypreal_of_real(x) + xa)) huffman@21164: - hypreal_of_real (f (g x))) huffman@21164: / (( *f* g) (hypreal_of_real(x) + xa) - hypreal_of_real (g x)) huffman@21164: \ hypreal_of_real(Da)" huffman@21164: by (auto simp add: NSDERIV_NSLIM_iff2 NSLIM_def diff_minus [symmetric]) huffman@21164: huffman@21164: (*-------------------------------------------------------------- huffman@21164: from other version of differentiability huffman@21164: huffman@21164: f(x + h) - f(x) huffman@21164: ----------------- \ Db huffman@21164: h huffman@21164: --------------------------------------------------------------*) huffman@21164: lemma NSDERIVD2: "[| NSDERIV g x :> Db; xa \ Infinitesimal; xa \ 0 |] huffman@21164: ==> (( *f* g) (hypreal_of_real(x) + xa) - hypreal_of_real(g x)) / xa huffman@21164: \ hypreal_of_real(Db)" huffman@21164: by (auto simp add: NSDERIV_NSLIM_iff NSLIM_def mem_infmal_iff starfun_lambda_cancel) huffman@21164: huffman@21164: lemma lemma_chain: "(z::hypreal) \ 0 ==> x*y = (x*inverse(z))*(z*y)" huffman@21164: by auto huffman@21164: huffman@21164: text{*This proof uses both definitions of differentiability.*} huffman@21164: lemma NSDERIV_chain: "[| NSDERIV f (g x) :> Da; NSDERIV g x :> Db |] huffman@21164: ==> NSDERIV (f o g) x :> Da * Db" huffman@21164: apply (simp (no_asm_simp) add: NSDERIV_NSLIM_iff NSLIM_def huffman@21164: mem_infmal_iff [symmetric]) huffman@21164: apply clarify huffman@21164: apply (frule_tac f = g in NSDERIV_approx) huffman@21164: apply (auto simp add: starfun_lambda_cancel2 starfun_o [symmetric]) huffman@21164: apply (case_tac "( *f* g) (hypreal_of_real (x) + xa) = hypreal_of_real (g x) ") huffman@21164: apply (drule_tac g = g in NSDERIV_zero) huffman@21164: apply (auto simp add: divide_inverse) huffman@21164: 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]) huffman@21164: apply (erule hypreal_not_eq_minus_iff [THEN iffD1]) huffman@21164: apply (rule approx_mult_hypreal_of_real) huffman@21164: apply (simp_all add: divide_inverse [symmetric]) huffman@21164: apply (blast intro: NSDERIVD1 approx_minus_iff [THEN iffD2]) huffman@21164: apply (blast intro: NSDERIVD2) huffman@21164: done huffman@21164: huffman@21164: text{*Differentiation of natural number powers*} huffman@21164: lemma NSDERIV_Id [simp]: "NSDERIV (%x. x) x :> 1" huffman@21164: by (simp add: NSDERIV_NSLIM_iff NSLIM_def divide_self del: divide_self_if) huffman@21164: huffman@21164: lemma NSDERIV_cmult_Id [simp]: "NSDERIV (op * c) x :> c" huffman@21164: by (cut_tac c = c and x = x in NSDERIV_Id [THEN NSDERIV_cmult], simp) huffman@21164: huffman@21164: (*Can't get rid of x \ 0 because it isn't continuous at zero*) huffman@21164: lemma NSDERIV_inverse: huffman@21164: "x \ 0 ==> NSDERIV (%x. inverse(x)) x :> (- (inverse x ^ Suc (Suc 0)))" huffman@21164: apply (simp add: nsderiv_def) huffman@21164: apply (rule ballI, simp, clarify) huffman@21164: apply (frule (1) Infinitesimal_add_not_zero) huffman@21164: apply (simp add: add_commute) huffman@21164: (*apply (auto simp add: starfun_inverse_inverse realpow_two huffman@21164: simp del: minus_mult_left [symmetric] minus_mult_right [symmetric])*) huffman@21164: apply (simp add: inverse_add inverse_mult_distrib [symmetric] huffman@21164: inverse_minus_eq [symmetric] add_ac mult_ac diff_def huffman@21164: del: inverse_mult_distrib inverse_minus_eq huffman@21164: minus_mult_left [symmetric] minus_mult_right [symmetric]) huffman@21164: apply (simp (no_asm_simp) add: mult_assoc [symmetric] right_distrib huffman@21164: del: minus_mult_left [symmetric] minus_mult_right [symmetric]) huffman@21164: apply (rule_tac y = "inverse (- hypreal_of_real x * hypreal_of_real x)" in approx_trans) huffman@21164: apply (rule inverse_add_Infinitesimal_approx2) huffman@21164: apply (auto dest!: hypreal_of_real_HFinite_diff_Infinitesimal huffman@21164: simp add: inverse_minus_eq [symmetric] HFinite_minus_iff) huffman@21164: apply (rule Infinitesimal_HFinite_mult2, auto) huffman@21164: done huffman@21164: huffman@21164: subsubsection {* Equivalence of NS and Standard definitions *} huffman@21164: huffman@21164: lemma divideR_eq_divide [simp]: "x /# y = x / y" huffman@21164: by (simp add: real_scaleR_def divide_inverse mult_commute) huffman@21164: huffman@21164: text{*Now equivalence between NSDERIV and DERIV*} huffman@21164: lemma NSDERIV_DERIV_iff: "(NSDERIV f x :> D) = (DERIV f x :> D)" huffman@21164: by (simp add: deriv_def NSDERIV_NSLIM_iff LIM_NSLIM_iff) huffman@21164: huffman@21164: (* let's do the standard proof though theorem *) huffman@21164: (* LIM_mult2 follows from a NS proof *) huffman@21164: huffman@21164: lemma DERIV_cmult: huffman@21164: fixes f :: "real \ 'a::real_normed_algebra" shows huffman@21164: "DERIV f x :> D ==> DERIV (%x. c * f x) x :> c*D" huffman@21164: by (drule DERIV_mult' [OF DERIV_const], simp) huffman@21164: huffman@21164: (* standard version *) huffman@21164: lemma DERIV_chain: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (f o g) x :> Da * Db" huffman@21164: by (drule (1) DERIV_chain', simp add: o_def real_scaleR_def mult_commute) huffman@21164: huffman@21164: lemma DERIV_chain2: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (%x. f (g x)) x :> Da * Db" huffman@21164: by (auto dest: DERIV_chain simp add: o_def) huffman@21164: huffman@21164: (*derivative of linear multiplication*) huffman@21164: lemma DERIV_cmult_Id [simp]: "DERIV (op * c) x :> c" huffman@21164: by (cut_tac c = c and x = x in DERIV_Id [THEN DERIV_cmult], simp) huffman@21164: huffman@21164: lemma DERIV_pow: "DERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))" huffman@21164: apply (cut_tac DERIV_power [OF DERIV_Id]) huffman@21164: apply (simp add: real_scaleR_def real_of_nat_def) huffman@21164: done huffman@21164: huffman@21164: (* NS version *) huffman@21164: lemma NSDERIV_pow: "NSDERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))" huffman@21164: by (simp add: NSDERIV_DERIV_iff DERIV_pow) huffman@21164: huffman@21164: text{*Power of -1*} huffman@21164: huffman@21164: lemma DERIV_inverse: "x \ 0 ==> DERIV (%x. inverse(x)) x :> (-(inverse x ^ Suc (Suc 0)))" huffman@21164: by (drule DERIV_inverse' [OF DERIV_Id], simp) huffman@21164: huffman@21164: text{*Derivative of inverse*} huffman@21164: lemma DERIV_inverse_fun: "[| DERIV f x :> d; f(x) \ 0 |] huffman@21164: ==> DERIV (%x. inverse(f x)::real) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))" huffman@21164: by (drule (1) DERIV_inverse', simp add: mult_ac) huffman@21164: huffman@21164: lemma NSDERIV_inverse_fun: "[| NSDERIV f x :> d; f(x) \ 0 |] huffman@21164: ==> NSDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))" huffman@21164: by (simp add: NSDERIV_DERIV_iff DERIV_inverse_fun del: realpow_Suc) huffman@21164: huffman@21164: text{*Derivative of quotient*} huffman@21164: lemma DERIV_quotient: "[| DERIV f x :> d; DERIV g x :> e; g(x) \ 0 |] huffman@21164: ==> DERIV (%y. f(y) / (g y) :: real) x :> (d*g(x) - (e*f(x))) / (g(x) ^ Suc (Suc 0))" huffman@21164: by (drule (2) DERIV_divide, simp add: mult_commute) huffman@21164: huffman@21164: lemma NSDERIV_quotient: "[| NSDERIV f x :> d; DERIV g x :> e; g(x) \ 0 |] huffman@21164: ==> NSDERIV (%y. f(y) / (g y)) x :> (d*g(x) huffman@21164: - (e*f(x))) / (g(x) ^ Suc (Suc 0))" huffman@21164: by (simp add: NSDERIV_DERIV_iff DERIV_quotient del: realpow_Suc) huffman@21164: huffman@21164: lemma CARAT_NSDERIV: "NSDERIV f x :> l ==> huffman@21164: \g. (\z. f z - f x = g z * (z-x)) & isNSCont g x & g x = l" huffman@21164: by (auto simp add: NSDERIV_DERIV_iff isNSCont_isCont_iff CARAT_DERIV huffman@21164: real_scaleR_def mult_commute) huffman@21164: huffman@21164: lemma hypreal_eq_minus_iff3: "(x = y + z) = (x + -z = (y::hypreal))" huffman@21164: by auto huffman@21164: huffman@21164: lemma CARAT_DERIVD: huffman@21164: assumes all: "\z. f z - f x = g z * (z-x)" huffman@21164: and nsc: "isNSCont g x" huffman@21164: shows "NSDERIV f x :> g x" huffman@21164: proof - huffman@21164: from nsc huffman@21164: have "\w. w \ hypreal_of_real x \ w \ hypreal_of_real x \ huffman@21164: ( *f* g) w * (w - hypreal_of_real x) / (w - hypreal_of_real x) \ huffman@21164: hypreal_of_real (g x)" huffman@21164: by (simp add: diff_minus isNSCont_def) huffman@21164: thus ?thesis using all huffman@21164: by (simp add: NSDERIV_iff2 starfun_if_eq cong: if_cong) huffman@21164: qed huffman@21164: huffman@21164: subsubsection {* Differentiability predicate *} huffman@21164: huffman@21164: lemma differentiableD: "f differentiable x ==> \D. DERIV f x :> D" huffman@21164: by (simp add: differentiable_def) huffman@21164: huffman@21164: lemma differentiableI: "DERIV f x :> D ==> f differentiable x" huffman@21164: by (force simp add: differentiable_def) huffman@21164: huffman@21164: lemma NSdifferentiableD: "f NSdifferentiable x ==> \D. NSDERIV f x :> D" huffman@21164: by (simp add: NSdifferentiable_def) huffman@21164: huffman@21164: lemma NSdifferentiableI: "NSDERIV f x :> D ==> f NSdifferentiable x" huffman@21164: by (force simp add: NSdifferentiable_def) huffman@21164: huffman@21164: lemma differentiable_const: "(\z. a) differentiable x" huffman@21164: apply (unfold differentiable_def) huffman@21164: apply (rule_tac x=0 in exI) huffman@21164: apply simp huffman@21164: done huffman@21164: huffman@21164: lemma differentiable_sum: huffman@21164: assumes "f differentiable x" huffman@21164: and "g differentiable x" huffman@21164: shows "(\x. f x + g x) differentiable x" huffman@21164: proof - huffman@21164: from prems have "\D. DERIV f x :> D" by (unfold differentiable_def) huffman@21164: then obtain df where "DERIV f x :> df" .. huffman@21164: moreover from prems have "\D. DERIV g x :> D" by (unfold differentiable_def) huffman@21164: then obtain dg where "DERIV g x :> dg" .. huffman@21164: ultimately have "DERIV (\x. f x + g x) x :> df + dg" by (rule DERIV_add) huffman@21164: hence "\D. DERIV (\x. f x + g x) x :> D" by auto huffman@21164: thus ?thesis by (fold differentiable_def) huffman@21164: qed huffman@21164: huffman@21164: lemma differentiable_diff: huffman@21164: assumes "f differentiable x" huffman@21164: and "g differentiable x" huffman@21164: shows "(\x. f x - g x) differentiable x" huffman@21164: proof - huffman@21164: from prems have "f differentiable x" by simp huffman@21164: moreover huffman@21164: from prems have "\D. DERIV g x :> D" by (unfold differentiable_def) huffman@21164: then obtain dg where "DERIV g x :> dg" .. huffman@21164: then have "DERIV (\x. - g x) x :> -dg" by (rule DERIV_minus) huffman@21164: hence "\D. DERIV (\x. - g x) x :> D" by auto huffman@21164: hence "(\x. - g x) differentiable x" by (fold differentiable_def) huffman@21164: ultimately huffman@21164: show ?thesis huffman@21164: by (auto simp: real_diff_def dest: differentiable_sum) huffman@21164: qed huffman@21164: huffman@21164: lemma differentiable_mult: huffman@21164: assumes "f differentiable x" huffman@21164: and "g differentiable x" huffman@21164: shows "(\x. f x * g x) differentiable x" huffman@21164: proof - huffman@21164: from prems have "\D. DERIV f x :> D" by (unfold differentiable_def) huffman@21164: then obtain df where "DERIV f x :> df" .. huffman@21164: moreover from prems have "\D. DERIV g x :> D" by (unfold differentiable_def) huffman@21164: then obtain dg where "DERIV g x :> dg" .. huffman@21164: ultimately have "DERIV (\x. f x * g x) x :> df * g x + dg * f x" by (simp add: DERIV_mult) huffman@21164: hence "\D. DERIV (\x. f x * g x) x :> D" by auto huffman@21164: thus ?thesis by (fold differentiable_def) huffman@21164: qed huffman@21164: huffman@21164: subsection {*(NS) Increment*} huffman@21164: lemma incrementI: huffman@21164: "f NSdifferentiable x ==> huffman@21164: increment f x h = ( *f* f) (hypreal_of_real(x) + h) - huffman@21164: hypreal_of_real (f x)" huffman@21164: by (simp add: increment_def) huffman@21164: huffman@21164: lemma incrementI2: "NSDERIV f x :> D ==> huffman@21164: increment f x h = ( *f* f) (hypreal_of_real(x) + h) - huffman@21164: hypreal_of_real (f x)" huffman@21164: apply (erule NSdifferentiableI [THEN incrementI]) huffman@21164: done huffman@21164: huffman@21164: (* The Increment theorem -- Keisler p. 65 *) huffman@21164: lemma increment_thm: "[| NSDERIV f x :> D; h \ Infinitesimal; h \ 0 |] huffman@21164: ==> \e \ Infinitesimal. increment f x h = hypreal_of_real(D)*h + e*h" huffman@21164: apply (frule_tac h = h in incrementI2, simp add: nsderiv_def) huffman@21164: apply (drule bspec, auto) huffman@21164: apply (drule bex_Infinitesimal_iff2 [THEN iffD2], clarify) huffman@21164: apply (frule_tac b1 = "hypreal_of_real (D) + y" huffman@21164: in hypreal_mult_right_cancel [THEN iffD2]) huffman@21164: 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) huffman@21164: apply assumption huffman@21164: apply (simp add: times_divide_eq_right [symmetric]) huffman@21164: apply (auto simp add: left_distrib) huffman@21164: done huffman@21164: huffman@21164: lemma increment_thm2: huffman@21164: "[| NSDERIV f x :> D; h \ 0; h \ 0 |] huffman@21164: ==> \e \ Infinitesimal. increment f x h = huffman@21164: hypreal_of_real(D)*h + e*h" huffman@21164: by (blast dest!: mem_infmal_iff [THEN iffD2] intro!: increment_thm) huffman@21164: huffman@21164: huffman@21164: lemma increment_approx_zero: "[| NSDERIV f x :> D; h \ 0; h \ 0 |] huffman@21164: ==> increment f x h \ 0" huffman@21164: apply (drule increment_thm2, huffman@21164: auto intro!: Infinitesimal_HFinite_mult2 HFinite_add simp add: left_distrib [symmetric] mem_infmal_iff [symmetric]) huffman@21164: apply (erule Infinitesimal_subset_HFinite [THEN subsetD]) huffman@21164: done huffman@21164: huffman@21164: subsection {* Nested Intervals and Bisection *} huffman@21164: huffman@21164: text{*Lemmas about nested intervals and proof by bisection (cf.Harrison). huffman@21164: All considerably tidied by lcp.*} huffman@21164: huffman@21164: lemma lemma_f_mono_add [rule_format (no_asm)]: "(\n. (f::nat=>real) n \ f (Suc n)) --> f m \ f(m + no)" huffman@21164: apply (induct "no") huffman@21164: apply (auto intro: order_trans) huffman@21164: done huffman@21164: huffman@21164: lemma f_inc_g_dec_Beq_f: "[| \n. f(n) \ f(Suc n); huffman@21164: \n. g(Suc n) \ g(n); huffman@21164: \n. f(n) \ g(n) |] huffman@21164: ==> Bseq (f :: nat \ real)" huffman@21164: apply (rule_tac k = "f 0" and K = "g 0" in BseqI2, rule allI) huffman@21164: apply (induct_tac "n") huffman@21164: apply (auto intro: order_trans) huffman@21164: apply (rule_tac y = "g (Suc na)" in order_trans) huffman@21164: apply (induct_tac [2] "na") huffman@21164: apply (auto intro: order_trans) huffman@21164: done huffman@21164: huffman@21164: lemma f_inc_g_dec_Beq_g: "[| \n. f(n) \ f(Suc n); huffman@21164: \n. g(Suc n) \ g(n); huffman@21164: \n. f(n) \ g(n) |] huffman@21164: ==> Bseq (g :: nat \ real)" huffman@21164: apply (subst Bseq_minus_iff [symmetric]) huffman@21164: apply (rule_tac g = "%x. - (f x)" in f_inc_g_dec_Beq_f) huffman@21164: apply auto huffman@21164: done huffman@21164: huffman@21164: lemma f_inc_imp_le_lim: huffman@21164: fixes f :: "nat \ real" huffman@21164: shows "\\n. f n \ f (Suc n); convergent f\ \ f n \ lim f" huffman@21164: apply (rule linorder_not_less [THEN iffD1]) huffman@21164: apply (auto simp add: convergent_LIMSEQ_iff LIMSEQ_iff monoseq_Suc) huffman@21164: apply (drule real_less_sum_gt_zero) huffman@21164: apply (drule_tac x = "f n + - lim f" in spec, safe) huffman@21164: apply (drule_tac P = "%na. no\na --> ?Q na" and x = "no + n" in spec, auto) huffman@21164: apply (subgoal_tac "lim f \ f (no + n) ") huffman@21164: apply (drule_tac no=no and m=n in lemma_f_mono_add) huffman@21164: apply (auto simp add: add_commute) huffman@21164: apply (induct_tac "no") huffman@21164: apply simp huffman@21164: apply (auto intro: order_trans simp add: diff_minus abs_if) huffman@21164: done huffman@21164: huffman@21164: lemma lim_uminus: "convergent g ==> lim (%x. - g x) = - (lim g)" huffman@21164: apply (rule LIMSEQ_minus [THEN limI]) huffman@21164: apply (simp add: convergent_LIMSEQ_iff) huffman@21164: done huffman@21164: huffman@21164: lemma g_dec_imp_lim_le: huffman@21164: fixes g :: "nat \ real" huffman@21164: shows "\\n. g (Suc n) \ g(n); convergent g\ \ lim g \ g n" huffman@21164: apply (subgoal_tac "- (g n) \ - (lim g) ") huffman@21164: apply (cut_tac [2] f = "%x. - (g x)" in f_inc_imp_le_lim) huffman@21164: apply (auto simp add: lim_uminus convergent_minus_iff [symmetric]) huffman@21164: done huffman@21164: huffman@21164: lemma lemma_nest: "[| \n. f(n) \ f(Suc n); huffman@21164: \n. g(Suc n) \ g(n); huffman@21164: \n. f(n) \ g(n) |] huffman@21164: ==> \l m :: real. l \ m & ((\n. f(n) \ l) & f ----> l) & huffman@21164: ((\n. m \ g(n)) & g ----> m)" huffman@21164: apply (subgoal_tac "monoseq f & monoseq g") huffman@21164: prefer 2 apply (force simp add: LIMSEQ_iff monoseq_Suc) huffman@21164: apply (subgoal_tac "Bseq f & Bseq g") huffman@21164: prefer 2 apply (blast intro: f_inc_g_dec_Beq_f f_inc_g_dec_Beq_g) huffman@21164: apply (auto dest!: Bseq_monoseq_convergent simp add: convergent_LIMSEQ_iff) huffman@21164: apply (rule_tac x = "lim f" in exI) huffman@21164: apply (rule_tac x = "lim g" in exI) huffman@21164: apply (auto intro: LIMSEQ_le) huffman@21164: apply (auto simp add: f_inc_imp_le_lim g_dec_imp_lim_le convergent_LIMSEQ_iff) huffman@21164: done huffman@21164: huffman@21164: lemma lemma_nest_unique: "[| \n. f(n) \ f(Suc n); huffman@21164: \n. g(Suc n) \ g(n); huffman@21164: \n. f(n) \ g(n); huffman@21164: (%n. f(n) - g(n)) ----> 0 |] huffman@21164: ==> \l::real. ((\n. f(n) \ l) & f ----> l) & huffman@21164: ((\n. l \ g(n)) & g ----> l)" huffman@21164: apply (drule lemma_nest, auto) huffman@21164: apply (subgoal_tac "l = m") huffman@21164: apply (drule_tac [2] X = f in LIMSEQ_diff) huffman@21164: apply (auto intro: LIMSEQ_unique) huffman@21164: done huffman@21164: huffman@21164: text{*The universal quantifiers below are required for the declaration huffman@21164: of @{text Bolzano_nest_unique} below.*} huffman@21164: huffman@21164: lemma Bolzano_bisect_le: huffman@21164: "a \ b ==> \n. fst (Bolzano_bisect P a b n) \ snd (Bolzano_bisect P a b n)" huffman@21164: apply (rule allI) huffman@21164: apply (induct_tac "n") huffman@21164: apply (auto simp add: Let_def split_def) huffman@21164: done huffman@21164: huffman@21164: lemma Bolzano_bisect_fst_le_Suc: "a \ b ==> huffman@21164: \n. fst(Bolzano_bisect P a b n) \ fst (Bolzano_bisect P a b (Suc n))" huffman@21164: apply (rule allI) huffman@21164: apply (induct_tac "n") huffman@21164: apply (auto simp add: Bolzano_bisect_le Let_def split_def) huffman@21164: done huffman@21164: huffman@21164: lemma Bolzano_bisect_Suc_le_snd: "a \ b ==> huffman@21164: \n. snd(Bolzano_bisect P a b (Suc n)) \ snd (Bolzano_bisect P a b n)" huffman@21164: apply (rule allI) huffman@21164: apply (induct_tac "n") huffman@21164: apply (auto simp add: Bolzano_bisect_le Let_def split_def) huffman@21164: done huffman@21164: huffman@21164: lemma eq_divide_2_times_iff: "((x::real) = y / (2 * z)) = (2 * x = y/z)" huffman@21164: apply (auto) huffman@21164: apply (drule_tac f = "%u. (1/2) *u" in arg_cong) huffman@21164: apply (simp) huffman@21164: done huffman@21164: huffman@21164: lemma Bolzano_bisect_diff: huffman@21164: "a \ b ==> huffman@21164: snd(Bolzano_bisect P a b n) - fst(Bolzano_bisect P a b n) = huffman@21164: (b-a) / (2 ^ n)" huffman@21164: apply (induct "n") huffman@21164: apply (auto simp add: eq_divide_2_times_iff add_divide_distrib Let_def split_def) huffman@21164: done huffman@21164: huffman@21164: lemmas Bolzano_nest_unique = huffman@21164: lemma_nest_unique huffman@21164: [OF Bolzano_bisect_fst_le_Suc Bolzano_bisect_Suc_le_snd Bolzano_bisect_le] huffman@21164: huffman@21164: huffman@21164: lemma not_P_Bolzano_bisect: huffman@21164: assumes P: "!!a b c. [| P(a,b); P(b,c); a \ b; b \ c|] ==> P(a,c)" huffman@21164: and notP: "~ P(a,b)" huffman@21164: and le: "a \ b" huffman@21164: shows "~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))" huffman@21164: proof (induct n) huffman@21164: case 0 thus ?case by simp huffman@21164: next huffman@21164: case (Suc n) huffman@21164: thus ?case huffman@21164: by (auto simp del: surjective_pairing [symmetric] huffman@21164: simp add: Let_def split_def Bolzano_bisect_le [OF le] huffman@21164: P [of "fst (Bolzano_bisect P a b n)" _ "snd (Bolzano_bisect P a b n)"]) huffman@21164: qed huffman@21164: huffman@21164: (*Now we re-package P_prem as a formula*) huffman@21164: lemma not_P_Bolzano_bisect': huffman@21164: "[| \a b c. P(a,b) & P(b,c) & a \ b & b \ c --> P(a,c); huffman@21164: ~ P(a,b); a \ b |] ==> huffman@21164: \n. ~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))" huffman@21164: by (blast elim!: not_P_Bolzano_bisect [THEN [2] rev_notE]) huffman@21164: huffman@21164: huffman@21164: huffman@21164: lemma lemma_BOLZANO: huffman@21164: "[| \a b c. P(a,b) & P(b,c) & a \ b & b \ c --> P(a,c); huffman@21164: \x. \d::real. 0 < d & huffman@21164: (\a b. a \ x & x \ b & (b-a) < d --> P(a,b)); huffman@21164: a \ b |] huffman@21164: ==> P(a,b)" huffman@21164: apply (rule Bolzano_nest_unique [where P1=P, THEN exE], assumption+) huffman@21164: apply (rule LIMSEQ_minus_cancel) huffman@21164: apply (simp (no_asm_simp) add: Bolzano_bisect_diff LIMSEQ_divide_realpow_zero) huffman@21164: apply (rule ccontr) huffman@21164: apply (drule not_P_Bolzano_bisect', assumption+) huffman@21164: apply (rename_tac "l") huffman@21164: apply (drule_tac x = l in spec, clarify) huffman@21164: apply (simp add: LIMSEQ_def) huffman@21164: apply (drule_tac P = "%r. 0 ?Q r" and x = "d/2" in spec) huffman@21164: apply (drule_tac P = "%r. 0 ?Q r" and x = "d/2" in spec) huffman@21164: apply (drule real_less_half_sum, auto) huffman@21164: apply (drule_tac x = "fst (Bolzano_bisect P a b (no + noa))" in spec) huffman@21164: apply (drule_tac x = "snd (Bolzano_bisect P a b (no + noa))" in spec) huffman@21164: apply safe huffman@21164: apply (simp_all (no_asm_simp)) huffman@21164: 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) huffman@21164: apply (simp (no_asm_simp) add: abs_if) huffman@21164: apply (rule real_sum_of_halves [THEN subst]) huffman@21164: apply (rule add_strict_mono) huffman@21164: apply (simp_all add: diff_minus [symmetric]) huffman@21164: done huffman@21164: huffman@21164: huffman@21164: lemma lemma_BOLZANO2: "((\a b c. (a \ b & b \ c & P(a,b) & P(b,c)) --> P(a,c)) & huffman@21164: (\x. \d::real. 0 < d & huffman@21164: (\a b. a \ x & x \ b & (b-a) < d --> P(a,b)))) huffman@21164: --> (\a b. a \ b --> P(a,b))" huffman@21164: apply clarify huffman@21164: apply (blast intro: lemma_BOLZANO) huffman@21164: done huffman@21164: huffman@21164: huffman@21164: subsection {* Intermediate Value Theorem *} huffman@21164: huffman@21164: text {*Prove Contrapositive by Bisection*} huffman@21164: huffman@21164: lemma IVT: "[| f(a::real) \ (y::real); y \ f(b); huffman@21164: a \ b; huffman@21164: (\x. a \ x & x \ b --> isCont f x) |] huffman@21164: ==> \x. a \ x & x \ b & f(x) = y" huffman@21164: apply (rule contrapos_pp, assumption) huffman@21164: apply (cut_tac P = "% (u,v) . a \ u & u \ v & v \ b --> ~ (f (u) \ y & y \ f (v))" in lemma_BOLZANO2) huffman@21164: apply safe huffman@21164: apply simp_all huffman@21164: apply (simp add: isCont_iff LIM_def) huffman@21164: apply (rule ccontr) huffman@21164: apply (subgoal_tac "a \ x & x \ b") huffman@21164: prefer 2 huffman@21164: apply simp huffman@21164: apply (drule_tac P = "%d. 0 ?P d" and x = 1 in spec, arith) huffman@21164: apply (drule_tac x = x in spec)+ huffman@21164: apply simp huffman@21164: apply (drule_tac P = "%r. ?P r --> (\s>0. ?Q r s) " and x = "\y - f x\" in spec) huffman@21164: apply safe huffman@21164: apply simp huffman@21164: apply (drule_tac x = s in spec, clarify) huffman@21164: apply (cut_tac x = "f x" and y = y in linorder_less_linear, safe) huffman@21164: apply (drule_tac x = "ba-x" in spec) huffman@21164: apply (simp_all add: abs_if) huffman@21164: apply (drule_tac x = "aa-x" in spec) huffman@21164: apply (case_tac "x \ aa", simp_all) huffman@21164: done huffman@21164: huffman@21164: lemma IVT2: "[| f(b::real) \ (y::real); y \ f(a); huffman@21164: a \ b; huffman@21164: (\x. a \ x & x \ b --> isCont f x) huffman@21164: |] ==> \x. a \ x & x \ b & f(x) = y" huffman@21164: apply (subgoal_tac "- f a \ -y & -y \ - f b", clarify) huffman@21164: apply (drule IVT [where f = "%x. - f x"], assumption) huffman@21164: apply (auto intro: isCont_minus) huffman@21164: done huffman@21164: huffman@21164: (*HOL style here: object-level formulations*) huffman@21164: lemma IVT_objl: "(f(a::real) \ (y::real) & y \ f(b) & a \ b & huffman@21164: (\x. a \ x & x \ b --> isCont f x)) huffman@21164: --> (\x. a \ x & x \ b & f(x) = y)" huffman@21164: apply (blast intro: IVT) huffman@21164: done huffman@21164: huffman@21164: lemma IVT2_objl: "(f(b::real) \ (y::real) & y \ f(a) & a \ b & huffman@21164: (\x. a \ x & x \ b --> isCont f x)) huffman@21164: --> (\x. a \ x & x \ b & f(x) = y)" huffman@21164: apply (blast intro: IVT2) huffman@21164: done huffman@21164: huffman@21164: text{*By bisection, function continuous on closed interval is bounded above*} huffman@21164: huffman@21164: lemma isCont_bounded: huffman@21164: "[| a \ b; \x. a \ x & x \ b --> isCont f x |] huffman@21164: ==> \M::real. \x::real. a \ x & x \ b --> f(x) \ M" huffman@21164: apply (cut_tac P = "% (u,v) . a \ u & u \ v & v \ b --> (\M. \x. u \ x & x \ v --> f x \ M)" in lemma_BOLZANO2) huffman@21164: apply safe huffman@21164: apply simp_all huffman@21164: apply (rename_tac x xa ya M Ma) huffman@21164: apply (cut_tac x = M and y = Ma in linorder_linear, safe) huffman@21164: apply (rule_tac x = Ma in exI, clarify) huffman@21164: apply (cut_tac x = xb and y = xa in linorder_linear, force) huffman@21164: apply (rule_tac x = M in exI, clarify) huffman@21164: apply (cut_tac x = xb and y = xa in linorder_linear, force) huffman@21164: apply (case_tac "a \ x & x \ b") huffman@21164: apply (rule_tac [2] x = 1 in exI) huffman@21164: prefer 2 apply force huffman@21164: apply (simp add: LIM_def isCont_iff) huffman@21164: apply (drule_tac x = x in spec, auto) huffman@21164: apply (erule_tac V = "\M. \x. a \ x & x \ b & ~ f x \ M" in thin_rl) huffman@21164: apply (drule_tac x = 1 in spec, auto) huffman@21164: apply (rule_tac x = s in exI, clarify) huffman@21164: apply (rule_tac x = "\f x\ + 1" in exI, clarify) huffman@21164: apply (drule_tac x = "xa-x" in spec) huffman@21164: apply (auto simp add: abs_ge_self) huffman@21164: done huffman@21164: huffman@21164: text{*Refine the above to existence of least upper bound*} huffman@21164: huffman@21164: lemma lemma_reals_complete: "((\x. x \ S) & (\y. isUb UNIV S (y::real))) --> huffman@21164: (\t. isLub UNIV S t)" huffman@21164: by (blast intro: reals_complete) huffman@21164: huffman@21164: lemma isCont_has_Ub: "[| a \ b; \x. a \ x & x \ b --> isCont f x |] huffman@21164: ==> \M::real. (\x::real. a \ x & x \ b --> f(x) \ M) & huffman@21164: (\N. N < M --> (\x. a \ x & x \ b & N < f(x)))" huffman@21164: apply (cut_tac S = "Collect (%y. \x. a \ x & x \ b & y = f x)" huffman@21164: in lemma_reals_complete) huffman@21164: apply auto huffman@21164: apply (drule isCont_bounded, assumption) huffman@21164: apply (auto simp add: isUb_def leastP_def isLub_def setge_def setle_def) huffman@21164: apply (rule exI, auto) huffman@21164: apply (auto dest!: spec simp add: linorder_not_less) huffman@21164: done huffman@21164: huffman@21164: text{*Now show that it attains its upper bound*} huffman@21164: huffman@21164: lemma isCont_eq_Ub: huffman@21164: assumes le: "a \ b" huffman@21164: and con: "\x::real. a \ x & x \ b --> isCont f x" huffman@21164: shows "\M::real. (\x. a \ x & x \ b --> f(x) \ M) & huffman@21164: (\x. a \ x & x \ b & f(x) = M)" huffman@21164: proof - huffman@21164: from isCont_has_Ub [OF le con] huffman@21164: obtain M where M1: "\x. a \ x \ x \ b \ f x \ M" huffman@21164: and M2: "!!N. N \x. a \ x \ x \ b \ N < f x" by blast huffman@21164: show ?thesis huffman@21164: proof (intro exI, intro conjI) huffman@21164: show " \x. a \ x \ x \ b \ f x \ M" by (rule M1) huffman@21164: show "\x. a \ x \ x \ b \ f x = M" huffman@21164: proof (rule ccontr) huffman@21164: assume "\ (\x. a \ x \ x \ b \ f x = M)" huffman@21164: with M1 have M3: "\x. a \ x & x \ b --> f x < M" huffman@21164: by (fastsimp simp add: linorder_not_le [symmetric]) huffman@21164: hence "\x. a \ x & x \ b --> isCont (%x. inverse (M - f x)) x" huffman@21164: by (auto simp add: isCont_inverse isCont_diff con) huffman@21164: from isCont_bounded [OF le this] huffman@21164: obtain k where k: "!!x. a \ x & x \ b --> inverse (M - f x) \ k" by auto huffman@21164: have Minv: "!!x. a \ x & x \ b --> 0 < inverse (M - f (x))" huffman@21164: by (simp add: M3 compare_rls) huffman@21164: have "!!x. a \ x & x \ b --> inverse (M - f x) < k+1" using k huffman@21164: by (auto intro: order_le_less_trans [of _ k]) huffman@21164: with Minv huffman@21164: have "!!x. a \ x & x \ b --> inverse(k+1) < inverse(inverse(M - f x))" huffman@21164: by (intro strip less_imp_inverse_less, simp_all) huffman@21164: hence invlt: "!!x. a \ x & x \ b --> inverse(k+1) < M - f x" huffman@21164: by simp huffman@21164: have "M - inverse (k+1) < M" using k [of a] Minv [of a] le huffman@21164: by (simp, arith) huffman@21164: from M2 [OF this] huffman@21164: obtain x where ax: "a \ x & x \ b & M - inverse(k+1) < f x" .. huffman@21164: thus False using invlt [of x] by force huffman@21164: qed huffman@21164: qed huffman@21164: qed huffman@21164: huffman@21164: huffman@21164: text{*Same theorem for lower bound*} huffman@21164: huffman@21164: lemma isCont_eq_Lb: "[| a \ b; \x. a \ x & x \ b --> isCont f x |] huffman@21164: ==> \M::real. (\x::real. a \ x & x \ b --> M \ f(x)) & huffman@21164: (\x. a \ x & x \ b & f(x) = M)" huffman@21164: apply (subgoal_tac "\x. a \ x & x \ b --> isCont (%x. - (f x)) x") huffman@21164: prefer 2 apply (blast intro: isCont_minus) huffman@21164: apply (drule_tac f = "(%x. - (f x))" in isCont_eq_Ub) huffman@21164: apply safe huffman@21164: apply auto huffman@21164: done huffman@21164: huffman@21164: huffman@21164: text{*Another version.*} huffman@21164: huffman@21164: lemma isCont_Lb_Ub: "[|a \ b; \x. a \ x & x \ b --> isCont f x |] huffman@21164: ==> \L M::real. (\x::real. a \ x & x \ b --> L \ f(x) & f(x) \ M) & huffman@21164: (\y. L \ y & y \ M --> (\x. a \ x & x \ b & (f(x) = y)))" huffman@21164: apply (frule isCont_eq_Lb) huffman@21164: apply (frule_tac [2] isCont_eq_Ub) huffman@21164: apply (assumption+, safe) huffman@21164: apply (rule_tac x = "f x" in exI) huffman@21164: apply (rule_tac x = "f xa" in exI, simp, safe) huffman@21164: apply (cut_tac x = x and y = xa in linorder_linear, safe) huffman@21164: apply (cut_tac f = f and a = x and b = xa and y = y in IVT_objl) huffman@21164: apply (cut_tac [2] f = f and a = xa and b = x and y = y in IVT2_objl, safe) huffman@21164: apply (rule_tac [2] x = xb in exI) huffman@21164: apply (rule_tac [4] x = xb in exI, simp_all) huffman@21164: done huffman@21164: huffman@21164: huffman@21164: text{*If @{term "0 < f'(x)"} then @{term x} is Locally Strictly Increasing At The Right*} huffman@21164: huffman@21164: lemma DERIV_left_inc: huffman@21164: fixes f :: "real => real" huffman@21164: assumes der: "DERIV f x :> l" huffman@21164: and l: "0 < l" huffman@21164: shows "\d > 0. \h > 0. h < d --> f(x) < f(x + h)" huffman@21164: proof - huffman@21164: from l der [THEN DERIV_D, THEN LIM_D [where r = "l"]] huffman@21164: have "\s > 0. (\z. z \ 0 \ \z\ < s \ \(f(x+z) - f x) / z - l\ < l)" huffman@21164: by (simp add: diff_minus) huffman@21164: then obtain s huffman@21164: where s: "0 < s" huffman@21164: and all: "!!z. z \ 0 \ \z\ < s \ \(f(x+z) - f x) / z - l\ < l" huffman@21164: by auto huffman@21164: thus ?thesis huffman@21164: proof (intro exI conjI strip) huffman@21164: show "0 real" huffman@21164: assumes der: "DERIV f x :> l" huffman@21164: and l: "l < 0" huffman@21164: shows "\d > 0. \h > 0. h < d --> f(x) < f(x-h)" huffman@21164: proof - huffman@21164: from l der [THEN DERIV_D, THEN LIM_D [where r = "-l"]] huffman@21164: have "\s > 0. (\z. z \ 0 \ \z\ < s \ \(f(x+z) - f x) / z - l\ < -l)" huffman@21164: by (simp add: diff_minus) huffman@21164: then obtain s huffman@21164: where s: "0 < s" huffman@21164: and all: "!!z. z \ 0 \ \z\ < s \ \(f(x+z) - f x) / z - l\ < -l" huffman@21164: by auto huffman@21164: thus ?thesis huffman@21164: proof (intro exI conjI strip) huffman@21164: show "0 real" huffman@21164: assumes der: "DERIV f x :> l" huffman@21164: and d: "0 < d" huffman@21164: and le: "\y. \x-y\ < d --> f(y) \ f(x)" huffman@21164: shows "l = 0" huffman@21164: proof (cases rule: linorder_cases [of l 0]) huffman@21164: case equal show ?thesis . huffman@21164: next huffman@21164: case less huffman@21164: from DERIV_left_dec [OF der less] huffman@21164: obtain d' where d': "0 < d'" huffman@21164: and lt: "\h > 0. h < d' \ f x < f (x-h)" by blast huffman@21164: from real_lbound_gt_zero [OF d d'] huffman@21164: obtain e where "0 < e \ e < d \ e < d'" .. huffman@21164: with lt le [THEN spec [where x="x-e"]] huffman@21164: show ?thesis by (auto simp add: abs_if) huffman@21164: next huffman@21164: case greater huffman@21164: from DERIV_left_inc [OF der greater] huffman@21164: obtain d' where d': "0 < d'" huffman@21164: and lt: "\h > 0. h < d' \ f x < f (x + h)" by blast huffman@21164: from real_lbound_gt_zero [OF d d'] huffman@21164: obtain e where "0 < e \ e < d \ e < d'" .. huffman@21164: with lt le [THEN spec [where x="x+e"]] huffman@21164: show ?thesis by (auto simp add: abs_if) huffman@21164: qed huffman@21164: huffman@21164: huffman@21164: text{*Similar theorem for a local minimum*} huffman@21164: lemma DERIV_local_min: huffman@21164: fixes f :: "real => real" huffman@21164: shows "[| DERIV f x :> l; 0 < d; \y. \x-y\ < d --> f(x) \ f(y) |] ==> l = 0" huffman@21164: by (drule DERIV_minus [THEN DERIV_local_max], auto) huffman@21164: huffman@21164: huffman@21164: text{*In particular, if a function is locally flat*} huffman@21164: lemma DERIV_local_const: huffman@21164: fixes f :: "real => real" huffman@21164: shows "[| DERIV f x :> l; 0 < d; \y. \x-y\ < d --> f(x) = f(y) |] ==> l = 0" huffman@21164: by (auto dest!: DERIV_local_max) huffman@21164: huffman@21164: text{*Lemma about introducing open ball in open interval*} huffman@21164: lemma lemma_interval_lt: huffman@21164: "[| a < x; x < b |] huffman@21164: ==> \d::real. 0 < d & (\y. \x-y\ < d --> a < y & y < b)" huffman@21164: apply (simp add: abs_interval_iff) huffman@21164: apply (insert linorder_linear [of "x-a" "b-x"], safe) huffman@21164: apply (rule_tac x = "x-a" in exI) huffman@21164: apply (rule_tac [2] x = "b-x" in exI, auto) huffman@21164: done huffman@21164: huffman@21164: lemma lemma_interval: "[| a < x; x < b |] ==> huffman@21164: \d::real. 0 < d & (\y. \x-y\ < d --> a \ y & y \ b)" huffman@21164: apply (drule lemma_interval_lt, auto) huffman@21164: apply (auto intro!: exI) huffman@21164: done huffman@21164: huffman@21164: text{*Rolle's Theorem. huffman@21164: If @{term f} is defined and continuous on the closed interval huffman@21164: @{text "[a,b]"} and differentiable on the open interval @{text "(a,b)"}, huffman@21164: and @{term "f(a) = f(b)"}, huffman@21164: then there exists @{text "x0 \ (a,b)"} such that @{term "f'(x0) = 0"}*} huffman@21164: theorem Rolle: huffman@21164: assumes lt: "a < b" huffman@21164: and eq: "f(a) = f(b)" huffman@21164: and con: "\x. a \ x & x \ b --> isCont f x" huffman@21164: and dif [rule_format]: "\x. a < x & x < b --> f differentiable x" huffman@21164: shows "\z. a < z & z < b & DERIV f z :> 0" huffman@21164: proof - huffman@21164: have le: "a \ b" using lt by simp huffman@21164: from isCont_eq_Ub [OF le con] huffman@21164: obtain x where x_max: "\z. a \ z \ z \ b \ f z \ f x" huffman@21164: and alex: "a \ x" and xleb: "x \ b" huffman@21164: by blast huffman@21164: from isCont_eq_Lb [OF le con] huffman@21164: obtain x' where x'_min: "\z. a \ z \ z \ b \ f x' \ f z" huffman@21164: and alex': "a \ x'" and x'leb: "x' \ b" huffman@21164: by blast huffman@21164: show ?thesis huffman@21164: proof cases huffman@21164: assume axb: "a < x & x < b" huffman@21164: --{*@{term f} attains its maximum within the interval*} huffman@21164: hence ax: "ay. \x-y\ < d \ a \ y \ y \ b" huffman@21164: by blast huffman@21164: hence bound': "\y. \x-y\ < d \ f y \ f x" using x_max huffman@21164: by blast huffman@21164: from differentiableD [OF dif [OF axb]] huffman@21164: obtain l where der: "DERIV f x :> l" .. huffman@21164: have "l=0" by (rule DERIV_local_max [OF der d bound']) huffman@21164: --{*the derivative at a local maximum is zero*} huffman@21164: thus ?thesis using ax xb der by auto huffman@21164: next huffman@21164: assume notaxb: "~ (a < x & x < b)" huffman@21164: hence xeqab: "x=a | x=b" using alex xleb by arith huffman@21164: hence fb_eq_fx: "f b = f x" by (auto simp add: eq) huffman@21164: show ?thesis huffman@21164: proof cases huffman@21164: assume ax'b: "a < x' & x' < b" huffman@21164: --{*@{term f} attains its minimum within the interval*} huffman@21164: hence ax': "ay. \x'-y\ < d \ a \ y \ y \ b" huffman@21164: by blast huffman@21164: hence bound': "\y. \x'-y\ < d \ f x' \ f y" using x'_min huffman@21164: by blast huffman@21164: from differentiableD [OF dif [OF ax'b]] huffman@21164: obtain l where der: "DERIV f x' :> l" .. huffman@21164: have "l=0" by (rule DERIV_local_min [OF der d bound']) huffman@21164: --{*the derivative at a local minimum is zero*} huffman@21164: thus ?thesis using ax' x'b der by auto huffman@21164: next huffman@21164: assume notax'b: "~ (a < x' & x' < b)" huffman@21164: --{*@{term f} is constant througout the interval*} huffman@21164: hence x'eqab: "x'=a | x'=b" using alex' x'leb by arith huffman@21164: hence fb_eq_fx': "f b = f x'" by (auto simp add: eq) huffman@21164: from dense [OF lt] huffman@21164: obtain r where ar: "a < r" and rb: "r < b" by blast huffman@21164: from lemma_interval [OF ar rb] huffman@21164: obtain d where d: "0y. \r-y\ < d \ a \ y \ y \ b" huffman@21164: by blast huffman@21164: have eq_fb: "\z. a \ z --> z \ b --> f z = f b" huffman@21164: proof (clarify) huffman@21164: fix z::real huffman@21164: assume az: "a \ z" and zb: "z \ b" huffman@21164: show "f z = f b" huffman@21164: proof (rule order_antisym) huffman@21164: show "f z \ f b" by (simp add: fb_eq_fx x_max az zb) huffman@21164: show "f b \ f z" by (simp add: fb_eq_fx' x'_min az zb) huffman@21164: qed huffman@21164: qed huffman@21164: have bound': "\y. \r-y\ < d \ f r = f y" huffman@21164: proof (intro strip) huffman@21164: fix y::real huffman@21164: assume lt: "\r-y\ < d" huffman@21164: hence "f y = f b" by (simp add: eq_fb bound) huffman@21164: thus "f r = f y" by (simp add: eq_fb ar rb order_less_imp_le) huffman@21164: qed huffman@21164: from differentiableD [OF dif [OF conjI [OF ar rb]]] huffman@21164: obtain l where der: "DERIV f r :> l" .. huffman@21164: have "l=0" by (rule DERIV_local_const [OF der d bound']) huffman@21164: --{*the derivative of a constant function is zero*} huffman@21164: thus ?thesis using ar rb der by auto huffman@21164: qed huffman@21164: qed huffman@21164: qed huffman@21164: huffman@21164: huffman@21164: subsection{*Mean Value Theorem*} huffman@21164: huffman@21164: lemma lemma_MVT: huffman@21164: "f a - (f b - f a)/(b-a) * a = f b - (f b - f a)/(b-a) * (b::real)" huffman@21164: proof cases huffman@21164: assume "a=b" thus ?thesis by simp huffman@21164: next huffman@21164: assume "a\b" huffman@21164: hence ba: "b-a \ 0" by arith huffman@21164: show ?thesis huffman@21164: by (rule real_mult_left_cancel [OF ba, THEN iffD1], huffman@21164: simp add: right_diff_distrib, huffman@21164: simp add: left_diff_distrib) huffman@21164: qed huffman@21164: huffman@21164: theorem MVT: huffman@21164: assumes lt: "a < b" huffman@21164: and con: "\x. a \ x & x \ b --> isCont f x" huffman@21164: and dif [rule_format]: "\x. a < x & x < b --> f differentiable x" huffman@21164: shows "\l z. a < z & z < b & DERIV f z :> l & huffman@21164: (f(b) - f(a) = (b-a) * l)" huffman@21164: proof - huffman@21164: let ?F = "%x. f x - ((f b - f a) / (b-a)) * x" huffman@21164: have contF: "\x. a \ x \ x \ b \ isCont ?F x" using con huffman@21164: by (fast intro: isCont_diff isCont_const isCont_mult isCont_Id) huffman@21164: have difF: "\x. a < x \ x < b \ ?F differentiable x" huffman@21164: proof (clarify) huffman@21164: fix x::real huffman@21164: assume ax: "a < x" and xb: "x < b" huffman@21164: from differentiableD [OF dif [OF conjI [OF ax xb]]] huffman@21164: obtain l where der: "DERIV f x :> l" .. huffman@21164: show "?F differentiable x" huffman@21164: by (rule differentiableI [where D = "l - (f b - f a)/(b-a)"], huffman@21164: blast intro: DERIV_diff DERIV_cmult_Id der) huffman@21164: qed huffman@21164: from Rolle [where f = ?F, OF lt lemma_MVT contF difF] huffman@21164: obtain z where az: "a < z" and zb: "z < b" and der: "DERIV ?F z :> 0" huffman@21164: by blast huffman@21164: have "DERIV (%x. ((f b - f a)/(b-a)) * x) z :> (f b - f a)/(b-a)" huffman@21164: by (rule DERIV_cmult_Id) huffman@21164: hence derF: "DERIV (\x. ?F x + (f b - f a) / (b - a) * x) z huffman@21164: :> 0 + (f b - f a) / (b - a)" huffman@21164: by (rule DERIV_add [OF der]) huffman@21164: show ?thesis huffman@21164: proof (intro exI conjI) huffman@21164: show "a < z" . huffman@21164: show "z < b" . huffman@21164: show "f b - f a = (b - a) * ((f b - f a)/(b-a))" by (simp) huffman@21164: show "DERIV f z :> ((f b - f a)/(b-a))" using derF by simp huffman@21164: qed huffman@21164: qed huffman@21164: huffman@21164: huffman@21164: text{*A function is constant if its derivative is 0 over an interval.*} huffman@21164: huffman@21164: lemma DERIV_isconst_end: huffman@21164: fixes f :: "real => real" huffman@21164: shows "[| a < b; huffman@21164: \x. a \ x & x \ b --> isCont f x; huffman@21164: \x. a < x & x < b --> DERIV f x :> 0 |] huffman@21164: ==> f b = f a" huffman@21164: apply (drule MVT, assumption) huffman@21164: apply (blast intro: differentiableI) huffman@21164: apply (auto dest!: DERIV_unique simp add: diff_eq_eq) huffman@21164: done huffman@21164: huffman@21164: lemma DERIV_isconst1: huffman@21164: fixes f :: "real => real" huffman@21164: shows "[| a < b; huffman@21164: \x. a \ x & x \ b --> isCont f x; huffman@21164: \x. a < x & x < b --> DERIV f x :> 0 |] huffman@21164: ==> \x. a \ x & x \ b --> f x = f a" huffman@21164: apply safe huffman@21164: apply (drule_tac x = a in order_le_imp_less_or_eq, safe) huffman@21164: apply (drule_tac b = x in DERIV_isconst_end, auto) huffman@21164: done huffman@21164: huffman@21164: lemma DERIV_isconst2: huffman@21164: fixes f :: "real => real" huffman@21164: shows "[| a < b; huffman@21164: \x. a \ x & x \ b --> isCont f x; huffman@21164: \x. a < x & x < b --> DERIV f x :> 0; huffman@21164: a \ x; x \ b |] huffman@21164: ==> f x = f a" huffman@21164: apply (blast dest: DERIV_isconst1) huffman@21164: done huffman@21164: huffman@21164: lemma DERIV_isconst_all: huffman@21164: fixes f :: "real => real" huffman@21164: shows "\x. DERIV f x :> 0 ==> f(x) = f(y)" huffman@21164: apply (rule linorder_cases [of x y]) huffman@21164: apply (blast intro: sym DERIV_isCont DERIV_isconst_end)+ huffman@21164: done huffman@21164: huffman@21164: lemma DERIV_const_ratio_const: huffman@21164: "[|a \ b; \x. DERIV f x :> k |] ==> (f(b) - f(a)) = (b-a) * k" huffman@21164: apply (rule linorder_cases [of a b], auto) huffman@21164: apply (drule_tac [!] f = f in MVT) huffman@21164: apply (auto dest: DERIV_isCont DERIV_unique simp add: differentiable_def) huffman@21164: apply (auto dest: DERIV_unique simp add: left_distrib diff_minus) huffman@21164: done huffman@21164: huffman@21164: lemma DERIV_const_ratio_const2: huffman@21164: "[|a \ b; \x. DERIV f x :> k |] ==> (f(b) - f(a))/(b-a) = k" huffman@21164: apply (rule_tac c1 = "b-a" in real_mult_right_cancel [THEN iffD1]) huffman@21164: apply (auto dest!: DERIV_const_ratio_const simp add: mult_assoc) huffman@21164: done huffman@21164: huffman@21164: lemma real_average_minus_first [simp]: "((a + b) /2 - a) = (b-a)/(2::real)" huffman@21164: by (simp) huffman@21164: huffman@21164: lemma real_average_minus_second [simp]: "((b + a)/2 - a) = (b-a)/(2::real)" huffman@21164: by (simp) huffman@21164: huffman@21164: text{*Gallileo's "trick": average velocity = av. of end velocities*} huffman@21164: huffman@21164: lemma DERIV_const_average: huffman@21164: fixes v :: "real => real" huffman@21164: assumes neq: "a \ (b::real)" huffman@21164: and der: "\x. DERIV v x :> k" huffman@21164: shows "v ((a + b)/2) = (v a + v b)/2" huffman@21164: proof (cases rule: linorder_cases [of a b]) huffman@21164: case equal with neq show ?thesis by simp huffman@21164: next huffman@21164: case less huffman@21164: have "(v b - v a) / (b - a) = k" huffman@21164: by (rule DERIV_const_ratio_const2 [OF neq der]) huffman@21164: hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp huffman@21164: moreover have "(v ((a + b) / 2) - v a) / ((a + b) / 2 - a) = k" huffman@21164: by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq) huffman@21164: ultimately show ?thesis using neq by force huffman@21164: next huffman@21164: case greater huffman@21164: have "(v b - v a) / (b - a) = k" huffman@21164: by (rule DERIV_const_ratio_const2 [OF neq der]) huffman@21164: hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp huffman@21164: moreover have " (v ((b + a) / 2) - v a) / ((b + a) / 2 - a) = k" huffman@21164: by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq) huffman@21164: ultimately show ?thesis using neq by (force simp add: add_commute) huffman@21164: qed huffman@21164: huffman@21164: huffman@21164: text{*Dull lemma: an continuous injection on an interval must have a huffman@21164: strict maximum at an end point, not in the middle.*} huffman@21164: huffman@21164: lemma lemma_isCont_inj: huffman@21164: fixes f :: "real \ real" huffman@21164: assumes d: "0 < d" huffman@21164: and inj [rule_format]: "\z. \z-x\ \ d --> g(f z) = z" huffman@21164: and cont: "\z. \z-x\ \ d --> isCont f z" huffman@21164: shows "\z. \z-x\ \ d & f x < f z" huffman@21164: proof (rule ccontr) huffman@21164: assume "~ (\z. \z-x\ \ d & f x < f z)" huffman@21164: hence all [rule_format]: "\z. \z - x\ \ d --> f z \ f x" by auto huffman@21164: show False huffman@21164: proof (cases rule: linorder_le_cases [of "f(x-d)" "f(x+d)"]) huffman@21164: case le huffman@21164: from d cont all [of "x+d"] huffman@21164: have flef: "f(x+d) \ f x" huffman@21164: and xlex: "x - d \ x" huffman@21164: and cont': "\z. x - d \ z \ z \ x \ isCont f z" huffman@21164: by (auto simp add: abs_if) huffman@21164: from IVT [OF le flef xlex cont'] huffman@21164: obtain x' where "x-d \ x'" "x' \ x" "f x' = f(x+d)" by blast huffman@21164: moreover huffman@21164: hence "g(f x') = g (f(x+d))" by simp huffman@21164: ultimately show False using d inj [of x'] inj [of "x+d"] huffman@21164: by (simp add: abs_le_interval_iff) huffman@21164: next huffman@21164: case ge huffman@21164: from d cont all [of "x-d"] huffman@21164: have flef: "f(x-d) \ f x" huffman@21164: and xlex: "x \ x+d" huffman@21164: and cont': "\z. x \ z \ z \ x+d \ isCont f z" huffman@21164: by (auto simp add: abs_if) huffman@21164: from IVT2 [OF ge flef xlex cont'] huffman@21164: obtain x' where "x \ x'" "x' \ x+d" "f x' = f(x-d)" by blast huffman@21164: moreover huffman@21164: hence "g(f x') = g (f(x-d))" by simp huffman@21164: ultimately show False using d inj [of x'] inj [of "x-d"] huffman@21164: by (simp add: abs_le_interval_iff) huffman@21164: qed huffman@21164: qed huffman@21164: huffman@21164: huffman@21164: text{*Similar version for lower bound.*} huffman@21164: huffman@21164: lemma lemma_isCont_inj2: huffman@21164: fixes f g :: "real \ real" huffman@21164: shows "[|0 < d; \z. \z-x\ \ d --> g(f z) = z; huffman@21164: \z. \z-x\ \ d --> isCont f z |] huffman@21164: ==> \z. \z-x\ \ d & f z < f x" huffman@21164: apply (insert lemma_isCont_inj huffman@21164: [where f = "%x. - f x" and g = "%y. g(-y)" and x = x and d = d]) huffman@21164: apply (simp add: isCont_minus linorder_not_le) huffman@21164: done huffman@21164: huffman@21164: text{*Show there's an interval surrounding @{term "f(x)"} in huffman@21164: @{text "f[[x - d, x + d]]"} .*} huffman@21164: huffman@21164: lemma isCont_inj_range: huffman@21164: fixes f :: "real \ real" huffman@21164: assumes d: "0 < d" huffman@21164: and inj: "\z. \z-x\ \ d --> g(f z) = z" huffman@21164: and cont: "\z. \z-x\ \ d --> isCont f z" huffman@21164: shows "\e>0. \y. \y - f x\ \ e --> (\z. \z-x\ \ d & f z = y)" huffman@21164: proof - huffman@21164: have "x-d \ x+d" "\z. x-d \ z \ z \ x+d \ isCont f z" using cont d huffman@21164: by (auto simp add: abs_le_interval_iff) huffman@21164: from isCont_Lb_Ub [OF this] huffman@21164: obtain L M huffman@21164: where all1 [rule_format]: "\z. x-d \ z \ z \ x+d \ L \ f z \ f z \ M" huffman@21164: and all2 [rule_format]: huffman@21164: "\y. L \ y \ y \ M \ (\z. x-d \ z \ z \ x+d \ f z = y)" huffman@21164: by auto huffman@21164: with d have "L \ f x & f x \ M" by simp huffman@21164: moreover have "L \ f x" huffman@21164: proof - huffman@21164: from lemma_isCont_inj2 [OF d inj cont] huffman@21164: obtain u where "\u - x\ \ d" "f u < f x" by auto huffman@21164: thus ?thesis using all1 [of u] by arith huffman@21164: qed huffman@21164: moreover have "f x \ M" huffman@21164: proof - huffman@21164: from lemma_isCont_inj [OF d inj cont] huffman@21164: obtain u where "\u - x\ \ d" "f x < f u" by auto huffman@21164: thus ?thesis using all1 [of u] by arith huffman@21164: qed huffman@21164: ultimately have "L < f x & f x < M" by arith huffman@21164: hence "0 < f x - L" "0 < M - f x" by arith+ huffman@21164: from real_lbound_gt_zero [OF this] huffman@21164: obtain e where e: "0 < e" "e < f x - L" "e < M - f x" by auto huffman@21164: thus ?thesis huffman@21164: proof (intro exI conjI) huffman@21164: show "0y. \y - f x\ \ e \ (\z. \z - x\ \ d \ f z = y)" huffman@21164: proof (intro strip) huffman@21164: fix y::real huffman@21164: assume "\y - f x\ \ e" huffman@21164: with e have "L \ y \ y \ M" by arith huffman@21164: from all2 [OF this] huffman@21164: obtain z where "x - d \ z" "z \ x + d" "f z = y" by blast huffman@21164: thus "\z. \z - x\ \ d \ f z = y" huffman@21164: by (force simp add: abs_le_interval_iff) huffman@21164: qed huffman@21164: qed huffman@21164: qed huffman@21164: huffman@21164: huffman@21164: text{*Continuity of inverse function*} huffman@21164: huffman@21164: lemma isCont_inverse_function: huffman@21164: fixes f g :: "real \ real" huffman@21164: assumes d: "0 < d" huffman@21164: and inj: "\z. \z-x\ \ d --> g(f z) = z" huffman@21164: and cont: "\z. \z-x\ \ d --> isCont f z" huffman@21164: shows "isCont g (f x)" huffman@21164: proof (simp add: isCont_iff LIM_eq) huffman@21164: show "\r. 0 < r \ huffman@21164: (\s>0. \z. z\0 \ \z\ < s \ \g(f x + z) - g(f x)\ < r)" huffman@21164: proof (intro strip) huffman@21164: fix r::real huffman@21164: assume r: "0 e < d" by blast huffman@21164: with inj cont huffman@21164: have e_simps: "\z. \z-x\ \ e --> g (f z) = z" huffman@21164: "\z. \z-x\ \ e --> isCont f z" by auto huffman@21164: from isCont_inj_range [OF e this] huffman@21164: obtain e' where e': "0 < e'" huffman@21164: and all: "\y. \y - f x\ \ e' \ (\z. \z - x\ \ e \ f z = y)" huffman@21164: by blast huffman@21164: show "\s>0. \z. z\0 \ \z\ < s \ \g(f x + z) - g(f x)\ < r" huffman@21164: proof (intro exI conjI) huffman@21164: show "0z. z \ 0 \ \z\ < e' \ \g (f x + z) - g (f x)\ < r" huffman@21164: proof (intro strip) huffman@21164: fix z::real huffman@21164: assume z: "z \ 0 \ \z\ < e'" huffman@21164: with e e_lt e_simps all [rule_format, of "f x + z"] huffman@21164: show "\g (f x + z) - g (f x)\ < r" by force huffman@21164: qed huffman@21164: qed huffman@21164: qed huffman@21164: qed huffman@21164: huffman@21164: theorem GMVT: huffman@21164: assumes alb: "a < b" huffman@21164: and fc: "\x. a \ x \ x \ b \ isCont f x" huffman@21164: and fd: "\x. a < x \ x < b \ f differentiable x" huffman@21164: and gc: "\x. a \ x \ x \ b \ isCont g x" huffman@21164: and gd: "\x. a < x \ x < b \ g differentiable x" huffman@21164: 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)" huffman@21164: proof - huffman@21164: let ?h = "\x. (f b - f a)*(g x) - (g b - g a)*(f x)" huffman@21164: from prems have "a < b" by simp huffman@21164: moreover have "\x. a \ x \ x \ b \ isCont ?h x" huffman@21164: proof - huffman@21164: have "\x. a <= x \ x <= b \ isCont (\x. f b - f a) x" by simp huffman@21164: with gc have "\x. a <= x \ x <= b \ isCont (\x. (f b - f a) * g x) x" huffman@21164: by (auto intro: isCont_mult) huffman@21164: moreover huffman@21164: have "\x. a <= x \ x <= b \ isCont (\x. g b - g a) x" by simp huffman@21164: with fc have "\x. a <= x \ x <= b \ isCont (\x. (g b - g a) * f x) x" huffman@21164: by (auto intro: isCont_mult) huffman@21164: ultimately show ?thesis huffman@21164: by (fastsimp intro: isCont_diff) huffman@21164: qed huffman@21164: moreover huffman@21164: have "\x. a < x \ x < b \ ?h differentiable x" huffman@21164: proof - huffman@21164: have "\x. a < x \ x < b \ (\x. f b - f a) differentiable x" by (simp add: differentiable_const) huffman@21164: with gd have "\x. a < x \ x < b \ (\x. (f b - f a) * g x) differentiable x" by (simp add: differentiable_mult) huffman@21164: moreover huffman@21164: have "\x. a < x \ x < b \ (\x. g b - g a) differentiable x" by (simp add: differentiable_const) huffman@21164: with fd have "\x. a < x \ x < b \ (\x. (g b - g a) * f x) differentiable x" by (simp add: differentiable_mult) huffman@21164: ultimately show ?thesis by (simp add: differentiable_diff) huffman@21164: qed huffman@21164: ultimately have "\l z. a < z \ z < b \ DERIV ?h z :> l \ ?h b - ?h a = (b - a) * l" by (rule MVT) huffman@21164: then obtain l where ldef: "\z. a < z \ z < b \ DERIV ?h z :> l \ ?h b - ?h a = (b - a) * l" .. huffman@21164: then obtain c where cdef: "a < c \ c < b \ DERIV ?h c :> l \ ?h b - ?h a = (b - a) * l" .. huffman@21164: huffman@21164: from cdef have cint: "a < c \ c < b" by auto huffman@21164: with gd have "g differentiable c" by simp huffman@21164: hence "\D. DERIV g c :> D" by (rule differentiableD) huffman@21164: then obtain g'c where g'cdef: "DERIV g c :> g'c" .. huffman@21164: huffman@21164: from cdef have "a < c \ c < b" by auto huffman@21164: with fd have "f differentiable c" by simp huffman@21164: hence "\D. DERIV f c :> D" by (rule differentiableD) huffman@21164: then obtain f'c where f'cdef: "DERIV f c :> f'c" .. huffman@21164: huffman@21164: from cdef have "DERIV ?h c :> l" by auto huffman@21164: moreover huffman@21164: { huffman@21164: from g'cdef have "DERIV (\x. (f b - f a) * g x) c :> g'c * (f b - f a)" huffman@21164: apply (insert DERIV_const [where k="f b - f a"]) huffman@21164: apply (drule meta_spec [of _ c]) huffman@21164: apply (drule DERIV_mult [where f="(\x. f b - f a)" and g=g]) huffman@21164: by simp_all huffman@21164: moreover from f'cdef have "DERIV (\x. (g b - g a) * f x) c :> f'c * (g b - g a)" huffman@21164: apply (insert DERIV_const [where k="g b - g a"]) huffman@21164: apply (drule meta_spec [of _ c]) huffman@21164: apply (drule DERIV_mult [where f="(\x. g b - g a)" and g=f]) huffman@21164: by simp_all huffman@21164: ultimately have "DERIV ?h c :> g'c * (f b - f a) - f'c * (g b - g a)" huffman@21164: by (simp add: DERIV_diff) huffman@21164: } huffman@21164: ultimately have leq: "l = g'c * (f b - f a) - f'c * (g b - g a)" by (rule DERIV_unique) huffman@21164: huffman@21164: { huffman@21164: from cdef have "?h b - ?h a = (b - a) * l" by auto huffman@21164: also with leq have "\ = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp huffman@21164: finally have "?h b - ?h a = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp huffman@21164: } huffman@21164: moreover huffman@21164: { huffman@21164: have "?h b - ?h a = huffman@21164: ((f b)*(g b) - (f a)*(g b) - (g b)*(f b) + (g a)*(f b)) - huffman@21164: ((f b)*(g a) - (f a)*(g a) - (g b)*(f a) + (g a)*(f a))" huffman@21164: by (simp add: mult_ac add_ac real_diff_mult_distrib) huffman@21164: hence "?h b - ?h a = 0" by auto huffman@21164: } huffman@21164: ultimately have "(b - a) * (g'c * (f b - f a) - f'c * (g b - g a)) = 0" by auto huffman@21164: with alb have "g'c * (f b - f a) - f'c * (g b - g a) = 0" by simp huffman@21164: hence "g'c * (f b - f a) = f'c * (g b - g a)" by simp huffman@21164: hence "(f b - f a) * g'c = (g b - g a) * f'c" by (simp add: mult_ac) huffman@21164: huffman@21164: with g'cdef f'cdef cint show ?thesis by auto huffman@21164: qed huffman@21164: huffman@21164: end