| author | paulson | 
| Fri, 20 Aug 2004 12:20:09 +0200 | |
| changeset 15150 | c7af682b9ee5 | 
| parent 15140 | 322485b816ac | 
| child 15229 | 1eb23f805c06 | 
| permissions | -rw-r--r-- | 
| 12224 | 1 | (* Title : MacLaurin.thy | 
| 2 | Author : Jacques D. Fleuriot | |
| 3 | Copyright : 2001 University of Edinburgh | |
| 4 | Description : MacLaurin series | |
| 15079 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 5 | Conversion to Isar and new proofs by Lawrence C Paulson, 2004 | 
| 12224 | 6 | *) | 
| 7 | ||
| 15131 | 8 | theory MacLaurin | 
| 15140 | 9 | imports Log | 
| 15131 | 10 | begin | 
| 15079 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 11 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 12 | lemma sumr_offset: "sumr 0 n (%m. f (m+k)) = sumr 0 (n+k) f - sumr 0 k f" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 13 | by (induct_tac "n", auto) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 14 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 15 | lemma sumr_offset2: "\<forall>f. sumr 0 n (%m. f (m+k)) = sumr 0 (n+k) f - sumr 0 k f" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 16 | by (induct_tac "n", auto) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 17 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 18 | lemma sumr_offset3: "sumr 0 (n+k) f = sumr 0 n (%m. f (m+k)) + sumr 0 k f" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 19 | by (simp add: sumr_offset) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 20 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 21 | lemma sumr_offset4: "\<forall>n f. sumr 0 (n+k) f = sumr 0 n (%m. f (m+k)) + sumr 0 k f" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 22 | by (simp add: sumr_offset) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 23 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 24 | lemma sumr_from_1_from_0: "0 < n ==> | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 25 | sumr (Suc 0) (Suc n) (%n. (if even(n) then 0 else | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 26 | ((- 1) ^ ((n - (Suc 0)) div 2))/(real (fact n))) * a ^ n) = | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 27 | sumr 0 (Suc n) (%n. (if even(n) then 0 else | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 28 | ((- 1) ^ ((n - (Suc 0)) div 2))/(real (fact n))) * a ^ n)" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 29 | by (rule_tac n1 = 1 in sumr_split_add [THEN subst], auto) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 30 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 31 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 32 | subsection{*Maclaurin's Theorem with Lagrange Form of Remainder*}
 | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 33 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 34 | text{*This is a very long, messy proof even now that it's been broken down
 | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 35 | into lemmas.*} | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 36 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 37 | lemma Maclaurin_lemma: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 38 | "0 < h ==> | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 39 | \<exists>B. f h = sumr 0 n (%m. (j m / real (fact m)) * (h^m)) + | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 40 | (B * ((h^n) / real(fact n)))" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 41 | by (rule_tac x = "(f h - sumr 0 n (%m. (j m / real (fact m)) * h^m)) * | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 42 | real(fact n) / (h^n)" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 43 | in exI, auto) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 44 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 45 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 46 | lemma eq_diff_eq': "(x = y - z) = (y = x + (z::real))" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 47 | by arith | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 48 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 49 | text{*A crude tactic to differentiate by proof.*}
 | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 50 | ML | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 51 | {*
 | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 52 | exception DERIV_name; | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 53 | fun get_fun_name (_ $ (Const ("Lim.deriv",_) $ Abs(_,_, Const (f,_) $ _) $ _ $ _)) = f
 | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 54 | |   get_fun_name (_ $ (_ $ (Const ("Lim.deriv",_) $ Abs(_,_, Const (f,_) $ _) $ _ $ _))) = f
 | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 55 | | get_fun_name _ = raise DERIV_name; | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 56 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 57 | val deriv_rulesI = [DERIV_Id,DERIV_const,DERIV_cos,DERIV_cmult, | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 58 | DERIV_sin, DERIV_exp, DERIV_inverse,DERIV_pow, | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 59 | DERIV_add, DERIV_diff, DERIV_mult, DERIV_minus, | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 60 | DERIV_inverse_fun,DERIV_quotient,DERIV_fun_pow, | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 61 | DERIV_fun_exp,DERIV_fun_sin,DERIV_fun_cos, | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 62 | DERIV_Id,DERIV_const,DERIV_cos]; | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 63 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 64 | val deriv_tac = | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 65 | SUBGOAL (fn (prem,i) => | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 66 | (resolve_tac deriv_rulesI i) ORELSE | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 67 |     ((rtac (read_instantiate [("f",get_fun_name prem)]
 | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 68 | DERIV_chain2) i) handle DERIV_name => no_tac));; | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 69 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 70 | val DERIV_tac = ALLGOALS(fn i => REPEAT(deriv_tac i)); | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 71 | *} | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 72 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 73 | lemma Maclaurin_lemma2: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 74 | "[| \<forall>m t. m < n \<and> 0\<le>t \<and> t\<le>h \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t; | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 75 | n = Suc k; | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 76 | difg = | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 77 | (\<lambda>m t. diff m t - | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 78 | ((\<Sum>p = 0..<n - m. diff (m + p) 0 / real (fact p) * t ^ p) + | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 79 | B * (t ^ (n - m) / real (fact (n - m)))))|] ==> | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 80 | \<forall>m t. m < n & 0 \<le> t & t \<le> h --> | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 81 | DERIV (difg m) t :> difg (Suc m) t" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 82 | apply clarify | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 83 | apply (rule DERIV_diff) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 84 | apply (simp (no_asm_simp)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 85 | apply (tactic DERIV_tac) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 86 | apply (tactic DERIV_tac) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 87 | apply (rule_tac [2] lemma_DERIV_subst) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 88 | apply (rule_tac [2] DERIV_quotient) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 89 | apply (rule_tac [3] DERIV_const) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 90 | apply (rule_tac [2] DERIV_pow) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 91 | prefer 3 apply (simp add: fact_diff_Suc) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 92 | prefer 2 apply simp | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 93 | apply (frule_tac m = m in less_add_one, clarify) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 94 | apply (simp del: sumr_Suc) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 95 | apply (insert sumr_offset4 [of 1]) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 96 | apply (simp del: sumr_Suc fact_Suc realpow_Suc) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 97 | apply (rule lemma_DERIV_subst) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 98 | apply (rule DERIV_add) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 99 | apply (rule_tac [2] DERIV_const) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 100 | apply (rule DERIV_sumr, clarify) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 101 | prefer 2 apply simp | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 102 | apply (simp (no_asm) add: divide_inverse mult_assoc del: fact_Suc realpow_Suc) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 103 | apply (rule DERIV_cmult) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 104 | apply (rule lemma_DERIV_subst) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 105 | apply (best intro: DERIV_chain2 intro!: DERIV_intros) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 106 | apply (subst fact_Suc) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 107 | apply (subst real_of_nat_mult) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 108 | apply (simp add: inverse_mult_distrib mult_ac) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 109 | done | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 110 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 111 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 112 | lemma Maclaurin_lemma3: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 113 | "[|\<forall>k t. k < Suc m \<and> 0\<le>t & t\<le>h \<longrightarrow> DERIV (difg k) t :> difg (Suc k) t; | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 114 | \<forall>k<Suc m. difg k 0 = 0; DERIV (difg n) t :> 0; n < m; 0 < t; | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 115 | t < h|] | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 116 | ==> \<exists>ta. 0 < ta & ta < t & DERIV (difg (Suc n)) ta :> 0" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 117 | apply (rule Rolle, assumption, simp) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 118 | apply (drule_tac x = n and P="%k. k<Suc m --> difg k 0 = 0" in spec) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 119 | apply (rule DERIV_unique) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 120 | prefer 2 apply assumption | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 121 | apply force | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 122 | apply (subgoal_tac "\<forall>ta. 0 \<le> ta & ta \<le> t --> (difg (Suc n)) differentiable ta") | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 123 | apply (simp add: differentiable_def) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 124 | apply (blast dest!: DERIV_isCont) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 125 | apply (simp add: differentiable_def, clarify) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 126 | apply (rule_tac x = "difg (Suc (Suc n)) ta" in exI) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 127 | apply force | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 128 | apply (simp add: differentiable_def, clarify) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 129 | apply (rule_tac x = "difg (Suc (Suc n)) x" in exI) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 130 | apply force | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 131 | done | 
| 14738 | 132 | |
| 15079 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 133 | lemma Maclaurin: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 134 | "[| 0 < h; 0 < n; diff 0 = f; | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 135 | \<forall>m t. m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t |] | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 136 | ==> \<exists>t. 0 < t & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 137 | t < h & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 138 | f h = | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 139 | sumr 0 n (%m. (diff m 0 / real (fact m)) * h ^ m) + | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 140 | (diff n t / real (fact n)) * h ^ n" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 141 | apply (case_tac "n = 0", force) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 142 | apply (drule not0_implies_Suc) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 143 | apply (erule exE) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 144 | apply (frule_tac f=f and n=n and j="%m. diff m 0" in Maclaurin_lemma) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 145 | apply (erule exE) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 146 | apply (subgoal_tac "\<exists>g. | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 147 | g = (%t. f t - (sumr 0 n (%m. (diff m 0 / real(fact m)) * t^m) + (B * (t^n / real(fact n)))))") | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 148 | prefer 2 apply blast | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 149 | apply (erule exE) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 150 | apply (subgoal_tac "g 0 = 0 & g h =0") | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 151 | prefer 2 | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 152 | apply (simp del: sumr_Suc) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 153 | apply (cut_tac n = m and k = 1 in sumr_offset2) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 154 | apply (simp add: eq_diff_eq' del: sumr_Suc) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 155 | apply (subgoal_tac "\<exists>difg. difg = (%m t. diff m t - (sumr 0 (n - m) (%p. (diff (m + p) 0 / real (fact p)) * (t ^ p)) + (B * ((t ^ (n - m)) / real (fact (n - m))))))") | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 156 | prefer 2 apply blast | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 157 | apply (erule exE) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 158 | apply (subgoal_tac "difg 0 = g") | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 159 | prefer 2 apply simp | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 160 | apply (frule Maclaurin_lemma2, assumption+) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 161 | apply (subgoal_tac "\<forall>ma. ma < n --> (\<exists>t. 0 < t & t < h & difg (Suc ma) t = 0) ") | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 162 | apply (drule_tac x = m and P="%m. m<n --> (\<exists>t. ?QQ m t)" in spec) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 163 | apply (erule impE) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 164 | apply (simp (no_asm_simp)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 165 | apply (erule exE) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 166 | apply (rule_tac x = t in exI) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 167 | apply (simp del: realpow_Suc fact_Suc) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 168 | apply (subgoal_tac "\<forall>m. m < n --> difg m 0 = 0") | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 169 | prefer 2 | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 170 | apply clarify | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 171 | apply simp | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 172 | apply (frule_tac m = ma in less_add_one, clarify) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 173 | apply (simp del: sumr_Suc) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 174 | apply (insert sumr_offset4 [of 1]) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 175 | apply (simp del: sumr_Suc fact_Suc realpow_Suc) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 176 | apply (subgoal_tac "\<forall>m. m < n --> (\<exists>t. 0 < t & t < h & DERIV (difg m) t :> 0) ") | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 177 | apply (rule allI, rule impI) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 178 | apply (drule_tac x = ma and P="%m. m<n --> (\<exists>t. ?QQ m t)" in spec) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 179 | apply (erule impE, assumption) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 180 | apply (erule exE) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 181 | apply (rule_tac x = t in exI) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 182 | (* do some tidying up *) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 183 | apply (erule_tac [!] V= "difg = (%m t. diff m t - (sumr 0 (n - m) (%p. diff (m + p) 0 / real (fact p) * t ^ p) + B * (t ^ (n - m) / real (fact (n - m)))))" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 184 | in thin_rl) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 185 | apply (erule_tac [!] V="g = (%t. f t - (sumr 0 n (%m. diff m 0 / real (fact m) * t ^ m) + B * (t ^ n / real (fact n))))" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 186 | in thin_rl) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 187 | apply (erule_tac [!] V="f h = sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + B * (h ^ n / real (fact n))" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 188 | in thin_rl) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 189 | (* back to business *) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 190 | apply (simp (no_asm_simp)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 191 | apply (rule DERIV_unique) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 192 | prefer 2 apply blast | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 193 | apply force | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 194 | apply (rule allI, induct_tac "ma") | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 195 | apply (rule impI, rule Rolle, assumption, simp, simp) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 196 | apply (subgoal_tac "\<forall>t. 0 \<le> t & t \<le> h --> g differentiable t") | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 197 | apply (simp add: differentiable_def) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 198 | apply (blast dest: DERIV_isCont) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 199 | apply (simp add: differentiable_def, clarify) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 200 | apply (rule_tac x = "difg (Suc 0) t" in exI) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 201 | apply force | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 202 | apply (simp add: differentiable_def, clarify) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 203 | apply (rule_tac x = "difg (Suc 0) x" in exI) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 204 | apply force | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 205 | apply safe | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 206 | apply force | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 207 | apply (frule Maclaurin_lemma3, assumption+, safe) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 208 | apply (rule_tac x = ta in exI, force) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 209 | done | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 210 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 211 | lemma Maclaurin_objl: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 212 | "0 < h & 0 < n & diff 0 = f & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 213 | (\<forall>m t. m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 214 | --> (\<exists>t. 0 < t & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 215 | t < h & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 216 | f h = | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 217 | sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 218 | diff n t / real (fact n) * h ^ n)" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 219 | by (blast intro: Maclaurin) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 220 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 221 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 222 | lemma Maclaurin2: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 223 | "[| 0 < h; diff 0 = f; | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 224 | \<forall>m t. | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 225 | m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t |] | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 226 | ==> \<exists>t. 0 < t & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 227 | t \<le> h & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 228 | f h = | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 229 | sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 230 | diff n t / real (fact n) * h ^ n" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 231 | apply (case_tac "n", auto) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 232 | apply (drule Maclaurin, auto) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 233 | done | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 234 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 235 | lemma Maclaurin2_objl: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 236 | "0 < h & diff 0 = f & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 237 | (\<forall>m t. | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 238 | m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 239 | --> (\<exists>t. 0 < t & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 240 | t \<le> h & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 241 | f h = | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 242 | sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 243 | diff n t / real (fact n) * h ^ n)" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 244 | by (blast intro: Maclaurin2) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 245 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 246 | lemma Maclaurin_minus: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 247 | "[| h < 0; 0 < n; diff 0 = f; | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 248 | \<forall>m t. m < n & h \<le> t & t \<le> 0 --> DERIV (diff m) t :> diff (Suc m) t |] | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 249 | ==> \<exists>t. h < t & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 250 | t < 0 & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 251 | f h = | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 252 | sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 253 | diff n t / real (fact n) * h ^ n" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 254 | apply (cut_tac f = "%x. f (-x)" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 255 | and diff = "%n x. ((- 1) ^ n) * diff n (-x)" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 256 | and h = "-h" and n = n in Maclaurin_objl) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 257 | apply simp | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 258 | apply safe | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 259 | apply (subst minus_mult_right) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 260 | apply (rule DERIV_cmult) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 261 | apply (rule lemma_DERIV_subst) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 262 | apply (rule DERIV_chain2 [where g=uminus]) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 263 | apply (rule_tac [2] DERIV_minus, rule_tac [2] DERIV_Id) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 264 | prefer 2 apply force | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 265 | apply force | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 266 | apply (rule_tac x = "-t" in exI, auto) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 267 | apply (subgoal_tac "(\<Sum>m = 0..<n. -1 ^ m * diff m 0 * (-h)^m / real(fact m)) = | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 268 | (\<Sum>m = 0..<n. diff m 0 * h ^ m / real(fact m))") | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 269 | apply (rule_tac [2] sumr_fun_eq) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 270 | apply (auto simp add: divide_inverse power_mult_distrib [symmetric]) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 271 | done | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 272 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 273 | lemma Maclaurin_minus_objl: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 274 | "(h < 0 & 0 < n & diff 0 = f & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 275 | (\<forall>m t. | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 276 | m < n & h \<le> t & t \<le> 0 --> DERIV (diff m) t :> diff (Suc m) t)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 277 | --> (\<exists>t. h < t & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 278 | t < 0 & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 279 | f h = | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 280 | sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 281 | diff n t / real (fact n) * h ^ n)" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 282 | by (blast intro: Maclaurin_minus) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 283 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 284 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 285 | subsection{*More Convenient "Bidirectional" Version.*}
 | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 286 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 287 | (* not good for PVS sin_approx, cos_approx *) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 288 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 289 | lemma Maclaurin_bi_le_lemma [rule_format]: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 290 | "0 < n \<longrightarrow> | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 291 | diff 0 0 = | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 292 | (\<Sum>m = 0..<n. diff m 0 * 0 ^ m / real (fact m)) + | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 293 | diff n 0 * 0 ^ n / real (fact n)" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 294 | by (induct_tac "n", auto) | 
| 14738 | 295 | |
| 15079 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 296 | lemma Maclaurin_bi_le: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 297 | "[| diff 0 = f; | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 298 | \<forall>m t. m < n & abs t \<le> abs x --> DERIV (diff m) t :> diff (Suc m) t |] | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 299 | ==> \<exists>t. abs t \<le> abs x & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 300 | f x = | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 301 | sumr 0 n (%m. diff m 0 / real (fact m) * x ^ m) + | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 302 | diff n t / real (fact n) * x ^ n" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 303 | apply (case_tac "n = 0", force) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 304 | apply (case_tac "x = 0") | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 305 | apply (rule_tac x = 0 in exI) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 306 | apply (force simp add: Maclaurin_bi_le_lemma) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 307 | apply (cut_tac x = x and y = 0 in linorder_less_linear, auto) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 308 | txt{*Case 1, where @{term "x < 0"}*}
 | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 309 | apply (cut_tac f = "diff 0" and diff = diff and h = x and n = n in Maclaurin_minus_objl, safe) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 310 | apply (simp add: abs_if) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 311 | apply (rule_tac x = t in exI) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 312 | apply (simp add: abs_if) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 313 | txt{*Case 2, where @{term "0 < x"}*}
 | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 314 | apply (cut_tac f = "diff 0" and diff = diff and h = x and n = n in Maclaurin_objl, safe) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 315 | apply (simp add: abs_if) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 316 | apply (rule_tac x = t in exI) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 317 | apply (simp add: abs_if) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 318 | done | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 319 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 320 | lemma Maclaurin_all_lt: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 321 | "[| diff 0 = f; | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 322 | \<forall>m x. DERIV (diff m) x :> diff(Suc m) x; | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 323 | x ~= 0; 0 < n | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 324 | |] ==> \<exists>t. 0 < abs t & abs t < abs x & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 325 | f x = sumr 0 n (%m. (diff m 0 / real (fact m)) * x ^ m) + | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 326 | (diff n t / real (fact n)) * x ^ n" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 327 | apply (rule_tac x = x and y = 0 in linorder_cases) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 328 | prefer 2 apply blast | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 329 | apply (drule_tac [2] diff=diff in Maclaurin) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 330 | apply (drule_tac diff=diff in Maclaurin_minus, simp_all, safe) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 331 | apply (rule_tac [!] x = t in exI, auto, arith+) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 332 | done | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 333 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 334 | lemma Maclaurin_all_lt_objl: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 335 | "diff 0 = f & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 336 | (\<forall>m x. DERIV (diff m) x :> diff(Suc m) x) & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 337 | x ~= 0 & 0 < n | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 338 | --> (\<exists>t. 0 < abs t & abs t < abs x & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 339 | f x = sumr 0 n (%m. (diff m 0 / real (fact m)) * x ^ m) + | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 340 | (diff n t / real (fact n)) * x ^ n)" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 341 | by (blast intro: Maclaurin_all_lt) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 342 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 343 | lemma Maclaurin_zero [rule_format]: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 344 | "x = (0::real) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 345 | ==> 0 < n --> | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 346 | sumr 0 n (%m. (diff m (0::real) / real (fact m)) * x ^ m) = | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 347 | diff 0 0" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 348 | by (induct n, auto) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 349 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 350 | lemma Maclaurin_all_le: "[| diff 0 = f; | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 351 | \<forall>m x. DERIV (diff m) x :> diff (Suc m) x | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 352 | |] ==> \<exists>t. abs t \<le> abs x & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 353 | f x = sumr 0 n (%m. (diff m 0 / real (fact m)) * x ^ m) + | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 354 | (diff n t / real (fact n)) * x ^ n" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 355 | apply (insert linorder_le_less_linear [of n 0]) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 356 | apply (erule disjE, force) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 357 | apply (case_tac "x = 0") | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 358 | apply (frule_tac diff = diff and n = n in Maclaurin_zero, assumption) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 359 | apply (drule gr_implies_not0 [THEN not0_implies_Suc]) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 360 | apply (rule_tac x = 0 in exI, force) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 361 | apply (frule_tac diff = diff and n = n in Maclaurin_all_lt, auto) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 362 | apply (rule_tac x = t in exI, auto) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 363 | done | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 364 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 365 | lemma Maclaurin_all_le_objl: "diff 0 = f & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 366 | (\<forall>m x. DERIV (diff m) x :> diff (Suc m) x) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 367 | --> (\<exists>t. abs t \<le> abs x & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 368 | f x = sumr 0 n (%m. (diff m 0 / real (fact m)) * x ^ m) + | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 369 | (diff n t / real (fact n)) * x ^ n)" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 370 | by (blast intro: Maclaurin_all_le) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 371 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 372 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 373 | subsection{*Version for Exponential Function*}
 | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 374 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 375 | lemma Maclaurin_exp_lt: "[| x ~= 0; 0 < n |] | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 376 | ==> (\<exists>t. 0 < abs t & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 377 | abs t < abs x & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 378 | exp x = sumr 0 n (%m. (x ^ m) / real (fact m)) + | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 379 | (exp t / real (fact n)) * x ^ n)" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 380 | by (cut_tac diff = "%n. exp" and f = exp and x = x and n = n in Maclaurin_all_lt_objl, auto) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 381 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 382 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 383 | lemma Maclaurin_exp_le: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 384 | "\<exists>t. abs t \<le> abs x & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 385 | exp x = sumr 0 n (%m. (x ^ m) / real (fact m)) + | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 386 | (exp t / real (fact n)) * x ^ n" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 387 | by (cut_tac diff = "%n. exp" and f = exp and x = x and n = n in Maclaurin_all_le_objl, auto) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 388 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 389 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 390 | subsection{*Version for Sine Function*}
 | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 391 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 392 | lemma MVT2: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 393 | "[| a < b; \<forall>x. a \<le> x & x \<le> b --> DERIV f x :> f'(x) |] | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 394 | ==> \<exists>z. a < z & z < b & (f b - f a = (b - a) * f'(z))" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 395 | apply (drule MVT) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 396 | apply (blast intro: DERIV_isCont) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 397 | apply (force dest: order_less_imp_le simp add: differentiable_def) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 398 | apply (blast dest: DERIV_unique order_less_imp_le) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 399 | done | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 400 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 401 | lemma mod_exhaust_less_4: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 402 | "m mod 4 = 0 | m mod 4 = 1 | m mod 4 = 2 | m mod 4 = (3::nat)" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 403 | by (case_tac "m mod 4", auto, arith) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 404 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 405 | lemma Suc_Suc_mult_two_diff_two [rule_format, simp]: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 406 | "0 < n --> Suc (Suc (2 * n - 2)) = 2*n" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 407 | by (induct_tac "n", auto) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 408 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 409 | lemma lemma_Suc_Suc_4n_diff_2 [rule_format, simp]: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 410 | "0 < n --> Suc (Suc (4*n - 2)) = 4*n" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 411 | by (induct_tac "n", auto) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 412 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 413 | lemma Suc_mult_two_diff_one [rule_format, simp]: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 414 | "0 < n --> Suc (2 * n - 1) = 2*n" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 415 | by (induct_tac "n", auto) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 416 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 417 | lemma Maclaurin_sin_expansion: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 418 | "\<exists>t. sin x = | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 419 | (sumr 0 n (%m. (if even m then 0 | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 420 | else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) * | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 421 | x ^ m)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 422 | + ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 423 | apply (cut_tac f = sin and n = n and x = x and diff = "%n x. sin (x + 1/2*real (n) *pi)" in Maclaurin_all_lt_objl) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 424 | apply safe | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 425 | apply (simp (no_asm)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 426 | apply (simp (no_asm)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 427 | apply (case_tac "n", clarify, simp) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 428 | apply (drule_tac x = 0 in spec, simp, simp) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 429 | apply (rule ccontr, simp) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 430 | apply (drule_tac x = x in spec, simp) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 431 | apply (erule ssubst) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 432 | apply (rule_tac x = t in exI, simp) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 433 | apply (rule sumr_fun_eq) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 434 | apply (auto simp add: odd_Suc_mult_two_ex) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 435 | apply (auto simp add: even_mult_two_ex simp del: fact_Suc realpow_Suc) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 436 | (*Could sin_zero_iff help?*) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 437 | done | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 438 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 439 | lemma Maclaurin_sin_expansion2: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 440 | "\<exists>t. abs t \<le> abs x & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 441 | sin x = | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 442 | (sumr 0 n (%m. (if even m then 0 | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 443 | else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) * | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 444 | x ^ m)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 445 | + ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 446 | apply (cut_tac f = sin and n = n and x = x | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 447 | and diff = "%n x. sin (x + 1/2*real n * pi)" in Maclaurin_all_lt_objl) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 448 | apply safe | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 449 | apply (simp (no_asm)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 450 | apply (simp (no_asm)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 451 | apply (case_tac "n", clarify, simp, simp) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 452 | apply (rule ccontr, simp) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 453 | apply (drule_tac x = x in spec, simp) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 454 | apply (erule ssubst) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 455 | apply (rule_tac x = t in exI, simp) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 456 | apply (rule sumr_fun_eq) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 457 | apply (auto simp add: odd_Suc_mult_two_ex) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 458 | apply (auto simp add: even_mult_two_ex simp del: fact_Suc realpow_Suc) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 459 | done | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 460 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 461 | lemma Maclaurin_sin_expansion3: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 462 | "[| 0 < n; 0 < x |] ==> | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 463 | \<exists>t. 0 < t & t < x & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 464 | sin x = | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 465 | (sumr 0 n (%m. (if even m then 0 | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 466 | else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) * | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 467 | x ^ m)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 468 | + ((sin(t + 1/2 * real(n) *pi) / real (fact n)) * x ^ n)" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 469 | apply (cut_tac f = sin and n = n and h = x and diff = "%n x. sin (x + 1/2*real (n) *pi)" in Maclaurin_objl) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 470 | apply safe | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 471 | apply simp | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 472 | apply (simp (no_asm)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 473 | apply (erule ssubst) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 474 | apply (rule_tac x = t in exI, simp) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 475 | apply (rule sumr_fun_eq) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 476 | apply (auto simp add: odd_Suc_mult_two_ex) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 477 | apply (auto simp add: even_mult_two_ex simp del: fact_Suc realpow_Suc) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 478 | done | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 479 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 480 | lemma Maclaurin_sin_expansion4: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 481 | "0 < x ==> | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 482 | \<exists>t. 0 < t & t \<le> x & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 483 | sin x = | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 484 | (sumr 0 n (%m. (if even m then 0 | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 485 | else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) * | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 486 | x ^ m)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 487 | + ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 488 | apply (cut_tac f = sin and n = n and h = x and diff = "%n x. sin (x + 1/2*real (n) *pi)" in Maclaurin2_objl) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 489 | apply safe | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 490 | apply simp | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 491 | apply (simp (no_asm)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 492 | apply (erule ssubst) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 493 | apply (rule_tac x = t in exI, simp) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 494 | apply (rule sumr_fun_eq) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 495 | apply (auto simp add: odd_Suc_mult_two_ex) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 496 | apply (auto simp add: even_mult_two_ex simp del: fact_Suc realpow_Suc) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 497 | done | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 498 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 499 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 500 | subsection{*Maclaurin Expansion for Cosine Function*}
 | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 501 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 502 | lemma sumr_cos_zero_one [simp]: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 503 | "sumr 0 (Suc n) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 504 | (%m. (if even m | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 505 | then (- 1) ^ (m div 2)/(real (fact m)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 506 | else 0) * | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 507 | 0 ^ m) = 1" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 508 | by (induct_tac "n", auto) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 509 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 510 | lemma Maclaurin_cos_expansion: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 511 | "\<exists>t. abs t \<le> abs x & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 512 | cos x = | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 513 | (sumr 0 n (%m. (if even m | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 514 | then (- 1) ^ (m div 2)/(real (fact m)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 515 | else 0) * | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 516 | x ^ m)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 517 | + ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 518 | apply (cut_tac f = cos and n = n and x = x and diff = "%n x. cos (x + 1/2*real (n) *pi)" in Maclaurin_all_lt_objl) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 519 | apply safe | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 520 | apply (simp (no_asm)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 521 | apply (simp (no_asm)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 522 | apply (case_tac "n", simp) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 523 | apply (simp del: sumr_Suc) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 524 | apply (rule ccontr, simp) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 525 | apply (drule_tac x = x in spec, simp) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 526 | apply (erule ssubst) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 527 | apply (rule_tac x = t in exI, simp) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 528 | apply (rule sumr_fun_eq) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 529 | apply (auto simp add: odd_Suc_mult_two_ex) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 530 | apply (auto simp add: even_mult_two_ex left_distrib cos_add simp del: fact_Suc realpow_Suc) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 531 | apply (simp add: mult_commute [of _ pi]) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 532 | done | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 533 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 534 | lemma Maclaurin_cos_expansion2: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 535 | "[| 0 < x; 0 < n |] ==> | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 536 | \<exists>t. 0 < t & t < x & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 537 | cos x = | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 538 | (sumr 0 n (%m. (if even m | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 539 | then (- 1) ^ (m div 2)/(real (fact m)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 540 | else 0) * | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 541 | x ^ m)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 542 | + ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 543 | apply (cut_tac f = cos and n = n and h = x and diff = "%n x. cos (x + 1/2*real (n) *pi)" in Maclaurin_objl) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 544 | apply safe | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 545 | apply simp | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 546 | apply (simp (no_asm)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 547 | apply (erule ssubst) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 548 | apply (rule_tac x = t in exI, simp) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 549 | apply (rule sumr_fun_eq) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 550 | apply (auto simp add: odd_Suc_mult_two_ex) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 551 | apply (auto simp add: even_mult_two_ex left_distrib cos_add simp del: fact_Suc realpow_Suc) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 552 | apply (simp add: mult_commute [of _ pi]) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 553 | done | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 554 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 555 | lemma Maclaurin_minus_cos_expansion: "[| x < 0; 0 < n |] ==> | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 556 | \<exists>t. x < t & t < 0 & | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 557 | cos x = | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 558 | (sumr 0 n (%m. (if even m | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 559 | then (- 1) ^ (m div 2)/(real (fact m)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 560 | else 0) * | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 561 | x ^ m)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 562 | + ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)" | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 563 | apply (cut_tac f = cos and n = n and h = x and diff = "%n x. cos (x + 1/2*real (n) *pi)" in Maclaurin_minus_objl) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 564 | apply safe | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 565 | apply simp | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 566 | apply (simp (no_asm)) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 567 | apply (erule ssubst) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 568 | apply (rule_tac x = t in exI, simp) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 569 | apply (rule sumr_fun_eq) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 570 | apply (auto simp add: odd_Suc_mult_two_ex) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 571 | apply (auto simp add: even_mult_two_ex left_distrib cos_add simp del: fact_Suc realpow_Suc) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 572 | apply (simp add: mult_commute [of _ pi]) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 573 | done | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 574 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 575 | (* ------------------------------------------------------------------------- *) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 576 | (* Version for ln(1 +/- x). Where is it?? *) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 577 | (* ------------------------------------------------------------------------- *) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 578 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 579 | lemma sin_bound_lemma: | 
| 15081 | 580 | "[|x = y; abs u \<le> (v::real) |] ==> \<bar>(x + u) - y\<bar> \<le> v" | 
| 15079 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 581 | by auto | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 582 | |
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 583 | lemma Maclaurin_sin_bound: | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 584 | "abs(sin x - sumr 0 n (%m. (if even m then 0 else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) * | 
| 15081 | 585 | x ^ m)) \<le> inverse(real (fact n)) * \<bar>x\<bar> ^ n" | 
| 14738 | 586 | proof - | 
| 15079 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 587 | have "!! x (y::real). x \<le> 1 \<Longrightarrow> 0 \<le> y \<Longrightarrow> x * y \<le> 1 * y" | 
| 14738 | 588 | by (rule_tac mult_right_mono,simp_all) | 
| 589 | note est = this[simplified] | |
| 590 | show ?thesis | |
| 15079 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 591 | apply (cut_tac f=sin and n=n and x=x and | 
| 14738 | 592 | 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)" | 
| 593 | in Maclaurin_all_le_objl) | |
| 15079 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 594 | apply safe | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 595 | apply simp | 
| 14738 | 596 | apply (subst mod_Suc_eq_Suc_mod) | 
| 15079 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 597 | apply (cut_tac m=m in mod_exhaust_less_4, safe, simp+) | 
| 14738 | 598 | apply (rule DERIV_minus, simp+) | 
| 599 | apply (rule lemma_DERIV_subst, rule DERIV_minus, rule DERIV_cos, simp) | |
| 15079 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 600 | apply (erule ssubst) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 601 | apply (rule sin_bound_lemma) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 602 | apply (rule sumr_fun_eq, safe) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 603 | apply (rule_tac f = "%u. u * (x^r)" in arg_cong) | 
| 14738 | 604 | apply (subst even_even_mod_4_iff) | 
| 15079 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 605 | apply (cut_tac m=r in mod_exhaust_less_4, simp, safe) | 
| 14738 | 606 | apply (simp_all add:even_num_iff) | 
| 607 | apply (drule lemma_even_mod_4_div_2[simplified]) | |
| 15079 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 608 | apply(simp add: numeral_2_eq_2 divide_inverse) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 609 | apply (drule lemma_odd_mod_4_div_2) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 610 | apply (simp add: numeral_2_eq_2 divide_inverse) | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 611 | apply (auto intro: mult_right_mono [where b=1, simplified] mult_right_mono | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 612 | simp add: est mult_pos_le mult_ac divide_inverse | 
| 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 613 | power_abs [symmetric]) | 
| 14738 | 614 | done | 
| 615 | qed | |
| 616 | ||
| 15079 
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
 paulson parents: 
14738diff
changeset | 617 | end |