paulson@15944: (* ID : $Id$ paulson@12224: Author : Jacques D. Fleuriot paulson@12224: Copyright : 2001 University of Edinburgh paulson@15079: Conversion to Isar and new proofs by Lawrence C Paulson, 2004 paulson@12224: *) paulson@12224: paulson@15944: header{*MacLaurin Series*} paulson@15944: nipkow@15131: theory MacLaurin nipkow@15140: imports Log nipkow@15131: begin paulson@15079: paulson@15079: subsection{*Maclaurin's Theorem with Lagrange Form of Remainder*} paulson@15079: paulson@15079: text{*This is a very long, messy proof even now that it's been broken down paulson@15079: into lemmas.*} paulson@15079: paulson@15079: lemma Maclaurin_lemma: paulson@15079: "0 < h ==> nipkow@15539: \B. f h = (\m=0.. paulson@15079: (resolve_tac deriv_rulesI i) ORELSE paulson@15079: ((rtac (read_instantiate [("f",get_fun_name prem)] paulson@15079: DERIV_chain2) i) handle DERIV_name => no_tac));; paulson@15079: paulson@15079: val DERIV_tac = ALLGOALS(fn i => REPEAT(deriv_tac i)); wenzelm@19765: wenzelm@19765: end paulson@15079: *} paulson@15079: paulson@15079: lemma Maclaurin_lemma2: paulson@15079: "[| \m t. m < n \ 0\t \ t\h \ DERIV (diff m) t :> diff (Suc m) t; paulson@15079: n = Suc k; paulson@15079: difg = paulson@15079: (\m t. diff m t - paulson@15079: ((\p = 0.. paulson@15079: \m t. m < n & 0 \ t & t \ h --> paulson@15079: DERIV (difg m) t :> difg (Suc m) t" paulson@15079: apply clarify paulson@15079: apply (rule DERIV_diff) paulson@15079: apply (simp (no_asm_simp)) paulson@15079: apply (tactic DERIV_tac) paulson@15079: apply (tactic DERIV_tac) paulson@15079: apply (rule_tac [2] lemma_DERIV_subst) paulson@15079: apply (rule_tac [2] DERIV_quotient) paulson@15079: apply (rule_tac [3] DERIV_const) paulson@15079: apply (rule_tac [2] DERIV_pow) paulson@15079: prefer 3 apply (simp add: fact_diff_Suc) paulson@15079: prefer 2 apply simp paulson@15079: apply (frule_tac m = m in less_add_one, clarify) nipkow@15561: apply (simp del: setsum_op_ivl_Suc) paulson@15079: apply (insert sumr_offset4 [of 1]) nipkow@15561: apply (simp del: setsum_op_ivl_Suc fact_Suc realpow_Suc) paulson@15079: apply (rule lemma_DERIV_subst) paulson@15079: apply (rule DERIV_add) paulson@15079: apply (rule_tac [2] DERIV_const) paulson@15079: apply (rule DERIV_sumr, clarify) paulson@15079: prefer 2 apply simp paulson@15079: apply (simp (no_asm) add: divide_inverse mult_assoc del: fact_Suc realpow_Suc) paulson@15079: apply (rule DERIV_cmult) paulson@15079: apply (rule lemma_DERIV_subst) paulson@15079: apply (best intro: DERIV_chain2 intro!: DERIV_intros) paulson@15079: apply (subst fact_Suc) paulson@15079: apply (subst real_of_nat_mult) nipkow@15539: apply (simp add: mult_ac) paulson@15079: done paulson@15079: paulson@15079: paulson@15079: lemma Maclaurin_lemma3: huffman@20792: fixes difg :: "nat => real => real" shows paulson@15079: "[|\k t. k < Suc m \ 0\t & t\h \ DERIV (difg k) t :> difg (Suc k) t; paulson@15079: \k 0; n < m; 0 < t; paulson@15079: t < h|] paulson@15079: ==> \ta. 0 < ta & ta < t & DERIV (difg (Suc n)) ta :> 0" paulson@15079: apply (rule Rolle, assumption, simp) paulson@15079: apply (drule_tac x = n and P="%k. k difg k 0 = 0" in spec) paulson@15079: apply (rule DERIV_unique) paulson@15079: prefer 2 apply assumption paulson@15079: apply force paulson@15079: apply (subgoal_tac "\ta. 0 \ ta & ta \ t --> (difg (Suc n)) differentiable ta") paulson@15079: apply (simp add: differentiable_def) paulson@15079: apply (blast dest!: DERIV_isCont) paulson@15079: apply (simp add: differentiable_def, clarify) paulson@15079: apply (rule_tac x = "difg (Suc (Suc n)) ta" in exI) paulson@15079: apply force paulson@15079: apply (simp add: differentiable_def, clarify) paulson@15079: apply (rule_tac x = "difg (Suc (Suc n)) x" in exI) paulson@15079: apply force paulson@15079: done obua@14738: paulson@15079: lemma Maclaurin: paulson@15079: "[| 0 < h; 0 < n; diff 0 = f; paulson@15079: \m t. m < n & 0 \ t & t \ h --> DERIV (diff m) t :> diff (Suc m) t |] paulson@15079: ==> \t. 0 < t & paulson@15079: t < h & paulson@15079: f h = nipkow@15539: setsum (%m. (diff m 0 / real (fact m)) * h ^ m) {0..g. nipkow@15539: g = (%t. f t - (setsum (%m. (diff m 0 / real(fact m)) * t^m) {0..difg. difg = (%m t. diff m t - (setsum (%p. (diff (m + p) 0 / real (fact p)) * (t ^ p)) {0..ma. ma < n --> (\t. 0 < t & t < h & difg (Suc ma) t = 0) ") paulson@15234: apply (drule_tac x = m and P="%m. m (\t. ?QQ m t)" in spec) paulson@15234: apply (erule impE) paulson@15234: apply (simp (no_asm_simp)) paulson@15234: apply (erule exE) paulson@15234: apply (rule_tac x = t in exI) nipkow@15539: apply (simp del: realpow_Suc fact_Suc) paulson@15079: apply (subgoal_tac "\m. m < n --> difg m 0 = 0") paulson@15079: prefer 2 paulson@15079: apply clarify paulson@15079: apply simp paulson@15079: apply (frule_tac m = ma in less_add_one, clarify) nipkow@15561: apply (simp del: setsum_op_ivl_Suc) paulson@15079: apply (insert sumr_offset4 [of 1]) nipkow@15561: apply (simp del: setsum_op_ivl_Suc fact_Suc realpow_Suc) paulson@15079: apply (subgoal_tac "\m. m < n --> (\t. 0 < t & t < h & DERIV (difg m) t :> 0) ") paulson@15079: apply (rule allI, rule impI) paulson@15079: apply (drule_tac x = ma and P="%m. m (\t. ?QQ m t)" in spec) paulson@15079: apply (erule impE, assumption) paulson@15079: apply (erule exE) paulson@15079: apply (rule_tac x = t in exI) paulson@15079: (* do some tidying up *) nipkow@15539: apply (erule_tac [!] V= "difg = (%m t. diff m t - (setsum (%p. diff (m + p) 0 / real (fact p) * t ^ p) {0..t. 0 \ t & t \ h --> g differentiable t") paulson@15079: apply (simp add: differentiable_def) paulson@15079: apply (blast dest: DERIV_isCont) paulson@15079: apply (simp add: differentiable_def, clarify) paulson@15079: apply (rule_tac x = "difg (Suc 0) t" in exI) paulson@15079: apply force paulson@15079: apply (simp add: differentiable_def, clarify) paulson@15079: apply (rule_tac x = "difg (Suc 0) x" in exI) paulson@15079: apply force paulson@15079: apply safe paulson@15079: apply force paulson@15079: apply (frule Maclaurin_lemma3, assumption+, safe) paulson@15079: apply (rule_tac x = ta in exI, force) paulson@15079: done paulson@15079: paulson@15079: lemma Maclaurin_objl: paulson@15079: "0 < h & 0 < n & diff 0 = f & paulson@15079: (\m t. m < n & 0 \ t & t \ h --> DERIV (diff m) t :> diff (Suc m) t) paulson@15079: --> (\t. 0 < t & paulson@15079: t < h & paulson@15079: f h = nipkow@15539: (\m=0..m t. paulson@15079: m < n & 0 \ t & t \ h --> DERIV (diff m) t :> diff (Suc m) t |] paulson@15079: ==> \t. 0 < t & paulson@15079: t \ h & paulson@15079: f h = nipkow@15539: (\m=0..m t. paulson@15079: m < n & 0 \ t & t \ h --> DERIV (diff m) t :> diff (Suc m) t) paulson@15079: --> (\t. 0 < t & paulson@15079: t \ h & paulson@15079: f h = nipkow@15539: (\m=0..m t. m < n & h \ t & t \ 0 --> DERIV (diff m) t :> diff (Suc m) t |] paulson@15079: ==> \t. h < t & paulson@15079: t < 0 & paulson@15079: f h = nipkow@15539: (\m=0..m = 0..m = 0..m t. paulson@15079: m < n & h \ t & t \ 0 --> DERIV (diff m) t :> diff (Suc m) t)) paulson@15079: --> (\t. h < t & paulson@15079: t < 0 & paulson@15079: f h = nipkow@15539: (\m=0.. paulson@15079: diff 0 0 = paulson@15079: (\m = 0..m t. m < n & abs t \ abs x --> DERIV (diff m) t :> diff (Suc m) t |] paulson@15079: ==> \t. abs t \ abs x & paulson@15079: f x = nipkow@15539: (\m=0..m x. DERIV (diff m) x :> diff(Suc m) x; paulson@15079: x ~= 0; 0 < n paulson@15079: |] ==> \t. 0 < abs t & abs t < abs x & nipkow@15539: f x = (\m=0..m x. DERIV (diff m) x :> diff(Suc m) x) & paulson@15079: x ~= 0 & 0 < n paulson@15079: --> (\t. 0 < abs t & abs t < abs x & nipkow@15539: f x = (\m=0.. 0 < n --> nipkow@15539: (\m=0..m x. DERIV (diff m) x :> diff (Suc m) x paulson@15079: |] ==> \t. abs t \ abs x & nipkow@15539: f x = (\m=0..m x. DERIV (diff m) x :> diff (Suc m) x) paulson@15079: --> (\t. abs t \ abs x & nipkow@15539: f x = (\m=0.. (\t. 0 < abs t & paulson@15079: abs t < abs x & nipkow@15539: exp x = (\m=0..t. abs t \ abs x & nipkow@15539: exp x = (\m=0..x. a \ x & x \ b --> DERIV f x :> f'(x) |] paulson@15079: ==> \z. a < z & z < b & (f b - f a = (b - a) * f'(z))" paulson@15079: apply (drule MVT) paulson@15079: apply (blast intro: DERIV_isCont) paulson@15079: apply (force dest: order_less_imp_le simp add: differentiable_def) paulson@15079: apply (blast dest: DERIV_unique order_less_imp_le) paulson@15079: done paulson@15079: paulson@15079: lemma mod_exhaust_less_4: paulson@15079: "m mod 4 = 0 | m mod 4 = 1 | m mod 4 = 2 | m mod 4 = (3::nat)" webertj@20217: by auto paulson@15079: paulson@15079: lemma Suc_Suc_mult_two_diff_two [rule_format, simp]: paulson@15079: "0 < n --> Suc (Suc (2 * n - 2)) = 2*n" paulson@15251: by (induct "n", auto) paulson@15079: paulson@15079: lemma lemma_Suc_Suc_4n_diff_2 [rule_format, simp]: paulson@15079: "0 < n --> Suc (Suc (4*n - 2)) = 4*n" paulson@15251: by (induct "n", auto) paulson@15079: paulson@15079: lemma Suc_mult_two_diff_one [rule_format, simp]: paulson@15079: "0 < n --> Suc (2 * n - 1) = 2*n" paulson@15251: by (induct "n", auto) paulson@15079: paulson@15234: paulson@15234: text{*It is unclear why so many variant results are needed.*} paulson@15079: paulson@15079: lemma Maclaurin_sin_expansion2: paulson@15079: "\t. abs t \ abs x & paulson@15079: sin x = nipkow@15539: (\m=0..t. sin x = nipkow@15539: (\m=0.. paulson@15079: \t. 0 < t & t < x & paulson@15079: sin x = nipkow@15539: (\m=0.. paulson@15079: \t. 0 < t & t \ x & paulson@15079: sin x = nipkow@15539: (\m=0..m=0..<(Suc n). nipkow@15539: (if even m then (- 1) ^ (m div 2)/(real (fact m)) else 0) * 0 ^ m) = 1" paulson@15251: by (induct "n", auto) paulson@15079: paulson@15079: lemma Maclaurin_cos_expansion: paulson@15079: "\t. abs t \ abs x & paulson@15079: cos x = nipkow@15539: (\m=0.. paulson@15079: \t. 0 < t & t < x & paulson@15079: cos x = nipkow@15539: (\m=0.. paulson@15079: \t. x < t & t < 0 & paulson@15079: cos x = nipkow@15539: (\m=0.. (v::real) |] ==> \(x + u) - y\ \ v" paulson@15079: by auto paulson@15079: paulson@15079: lemma Maclaurin_sin_bound: nipkow@15539: "abs(sin x - (\m=0.. inverse(real (fact n)) * \x\ ^ n" obua@14738: proof - paulson@15079: have "!! x (y::real). x \ 1 \ 0 \ y \ x * y \ 1 * y" obua@14738: by (rule_tac mult_right_mono,simp_all) obua@14738: note est = this[simplified] obua@14738: show ?thesis paulson@15079: apply (cut_tac f=sin and n=n and x=x and obua@14738: diff = "%n x. if n mod 4 = 0 then sin(x) else if n mod 4 = 1 then cos(x) else if n mod 4 = 2 then -sin(x) else -cos(x)" obua@14738: in Maclaurin_all_le_objl) paulson@15079: apply safe paulson@15079: apply simp paulson@15944: apply (subst (1 2 3) mod_Suc_eq_Suc_mod) paulson@15079: apply (cut_tac m=m in mod_exhaust_less_4, safe, simp+) obua@14738: apply (rule DERIV_minus, simp+) obua@14738: apply (rule lemma_DERIV_subst, rule DERIV_minus, rule DERIV_cos, simp) paulson@15079: apply (erule ssubst) paulson@15079: apply (rule sin_bound_lemma) nipkow@15536: apply (rule setsum_cong[OF refl]) nipkow@15536: apply (rule_tac f = "%u. u * (x^xa)" in arg_cong) obua@14738: apply (subst even_even_mod_4_iff) nipkow@15536: apply (cut_tac m=xa in mod_exhaust_less_4, simp, safe) obua@14738: apply (simp_all add:even_num_iff) obua@14738: apply (drule lemma_even_mod_4_div_2[simplified]) paulson@15079: apply(simp add: numeral_2_eq_2 divide_inverse) paulson@15079: apply (drule lemma_odd_mod_4_div_2) paulson@15079: apply (simp add: numeral_2_eq_2 divide_inverse) paulson@15079: apply (auto intro: mult_right_mono [where b=1, simplified] mult_right_mono avigad@16775: simp add: est mult_nonneg_nonneg mult_ac divide_inverse paulson@16924: power_abs [symmetric] abs_mult) obua@14738: done obua@14738: qed obua@14738: paulson@15079: end