src/HOL/MacLaurin.thy
author huffman
Tue, 13 Jan 2009 14:08:24 -0800
changeset 29479 be8a15ffc511
parent 29187 7b09385234f9
child 29803 c56a5571f60a
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
28952
15a4b2cf8c34 made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents: 27239
diff changeset
     1
(*  Author      : Jacques D. Fleuriot
12224
02df7cbe7d25 even more theories from Jacques
paulson
parents:
diff changeset
     2
    Copyright   : 2001 University of Edinburgh
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
     3
    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
12224
02df7cbe7d25 even more theories from Jacques
paulson
parents:
diff changeset
     4
*)
02df7cbe7d25 even more theories from Jacques
paulson
parents:
diff changeset
     5
15944
9b00875e21f7 from simplesubst to new subst
paulson
parents: 15561
diff changeset
     6
header{*MacLaurin Series*}
9b00875e21f7 from simplesubst to new subst
paulson
parents: 15561
diff changeset
     7
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15081
diff changeset
     8
theory MacLaurin
26163
31e4ff2b9e5b Loads Dense_Linear_Order (needed dlo_simps)
chaieb
parents: 25162
diff changeset
     9
imports Dense_Linear_Order Transcendental
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15081
diff changeset
    10
begin
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    11
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    12
subsection{*Maclaurin's Theorem with Lagrange Form of Remainder*}
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    13
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    14
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: 14738
diff changeset
    15
into lemmas.*}
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    16
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    17
lemma Maclaurin_lemma:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    18
    "0 < h ==>
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
    19
     \<exists>B. f h = (\<Sum>m=0..<n. (j m / real (fact m)) * (h^m)) +
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    20
               (B * ((h^n) / real(fact n)))"
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
    21
apply (rule_tac x = "(f h - (\<Sum>m=0..<n. (j m / real (fact m)) * h^m)) *
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    22
                 real(fact n) / (h^n)"
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
    23
       in exI)
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
    24
apply (simp) 
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
    25
done
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    26
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    27
lemma eq_diff_eq': "(x = y - z) = (y = x + (z::real))"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    28
by arith
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    29
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    30
text{*A crude tactic to differentiate by proof.*}
24180
9f818139951b tuned ML setup;
wenzelm
parents: 23242
diff changeset
    31
9f818139951b tuned ML setup;
wenzelm
parents: 23242
diff changeset
    32
lemmas deriv_rulesI =
9f818139951b tuned ML setup;
wenzelm
parents: 23242
diff changeset
    33
  DERIV_ident DERIV_const DERIV_cos DERIV_cmult
9f818139951b tuned ML setup;
wenzelm
parents: 23242
diff changeset
    34
  DERIV_sin DERIV_exp DERIV_inverse DERIV_pow
9f818139951b tuned ML setup;
wenzelm
parents: 23242
diff changeset
    35
  DERIV_add DERIV_diff DERIV_mult DERIV_minus
9f818139951b tuned ML setup;
wenzelm
parents: 23242
diff changeset
    36
  DERIV_inverse_fun DERIV_quotient DERIV_fun_pow
9f818139951b tuned ML setup;
wenzelm
parents: 23242
diff changeset
    37
  DERIV_fun_exp DERIV_fun_sin DERIV_fun_cos
9f818139951b tuned ML setup;
wenzelm
parents: 23242
diff changeset
    38
  DERIV_ident DERIV_const DERIV_cos
9f818139951b tuned ML setup;
wenzelm
parents: 23242
diff changeset
    39
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    40
ML
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    41
{*
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16924
diff changeset
    42
local
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    43
exception DERIV_name;
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    44
fun get_fun_name (_ $ (Const ("Lim.deriv",_) $ Abs(_,_, Const (f,_) $ _) $ _ $ _)) = f
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    45
|   get_fun_name (_ $ (_ $ (Const ("Lim.deriv",_) $ Abs(_,_, Const (f,_) $ _) $ _ $ _))) = f
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    46
|   get_fun_name _ = raise DERIV_name;
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    47
24180
9f818139951b tuned ML setup;
wenzelm
parents: 23242
diff changeset
    48
in
9f818139951b tuned ML setup;
wenzelm
parents: 23242
diff changeset
    49
27227
184d7601c062 deriv_tac/DERIV_tac: proper context;
wenzelm
parents: 27153
diff changeset
    50
fun deriv_tac ctxt = SUBGOAL (fn (prem, i) =>
184d7601c062 deriv_tac/DERIV_tac: proper context;
wenzelm
parents: 27153
diff changeset
    51
  resolve_tac @{thms deriv_rulesI} i ORELSE
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27227
diff changeset
    52
    ((rtac (read_instantiate ctxt [(("f", 0), get_fun_name prem)]
27227
184d7601c062 deriv_tac/DERIV_tac: proper context;
wenzelm
parents: 27153
diff changeset
    53
                     @{thm DERIV_chain2}) i) handle DERIV_name => no_tac));
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    54
27227
184d7601c062 deriv_tac/DERIV_tac: proper context;
wenzelm
parents: 27153
diff changeset
    55
fun DERIV_tac ctxt = ALLGOALS (fn i => REPEAT (deriv_tac ctxt i));
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16924
diff changeset
    56
dfe940911617 misc cleanup;
wenzelm
parents: 16924
diff changeset
    57
end
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    58
*}
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    59
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    60
lemma Maclaurin_lemma2:
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    61
  assumes diff: "\<forall>m t. m < n \<and> 0\<le>t \<and> t\<le>h \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    62
  assumes n: "n = Suc k"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    63
  assumes difg: "difg =
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    64
        (\<lambda>m t. diff m t -
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    65
               ((\<Sum>p = 0..<n - m. diff (m + p) 0 / real (fact p) * t ^ p) +
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    66
                B * (t ^ (n - m) / real (fact (n - m)))))"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    67
  shows
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    68
      "\<forall>m t. m < n & 0 \<le> t & t \<le> h --> DERIV (difg m) t :> difg (Suc m) t"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    69
unfolding difg
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    70
 apply clarify
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    71
 apply (rule DERIV_diff)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    72
  apply (simp add: diff)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    73
 apply (simp only: n)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    74
 apply (rule DERIV_add)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    75
  apply (rule_tac [2] DERIV_cmult)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    76
  apply (rule_tac [2] lemma_DERIV_subst)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    77
   apply (rule_tac [2] DERIV_quotient)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    78
     apply (rule_tac [3] DERIV_const)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    79
    apply (rule_tac [2] DERIV_pow)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    80
   prefer 3 apply (simp add: fact_diff_Suc)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    81
  prefer 2 apply simp
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    82
 apply (frule less_iff_Suc_add [THEN iffD1], clarify)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    83
 apply (simp del: setsum_op_ivl_Suc)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    84
 apply (insert sumr_offset4 [of 1])
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    85
 apply (simp del: setsum_op_ivl_Suc fact_Suc realpow_Suc)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    86
 apply (rule lemma_DERIV_subst)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    87
  apply (rule DERIV_add)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    88
   apply (rule_tac [2] DERIV_const)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    89
  apply (rule DERIV_sumr, clarify)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    90
  prefer 2 apply simp
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    91
 apply (simp (no_asm) add: divide_inverse mult_assoc del: fact_Suc realpow_Suc)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    92
 apply (rule DERIV_cmult)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    93
 apply (rule lemma_DERIV_subst)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    94
  apply (best intro: DERIV_chain2 intro!: DERIV_intros)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    95
 apply (subst fact_Suc)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    96
 apply (subst real_of_nat_mult)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    97
 apply (simp add: mult_ac)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    98
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    99
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   100
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   101
lemma Maclaurin:
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   102
  assumes h: "0 < h"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   103
  assumes n: "0 < n"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   104
  assumes diff_0: "diff 0 = f"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   105
  assumes diff_Suc:
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   106
    "\<forall>m t. m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   107
  shows
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   108
    "\<exists>t. 0 < t & t < h &
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   109
              f h =
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   110
              setsum (%m. (diff m 0 / real (fact m)) * h ^ m) {0..<n} +
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   111
              (diff n t / real (fact n)) * h ^ n"
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   112
proof -
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   113
  from n obtain m where m: "n = Suc m"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   114
    by (cases n, simp add: n)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   115
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   116
  obtain B where f_h: "f h =
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   117
        (\<Sum>m = 0..<n. diff m (0\<Colon>real) / real (fact m) * h ^ m) +
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   118
        B * (h ^ n / real (fact n))"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   119
    using Maclaurin_lemma [OF h] ..
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   120
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   121
  obtain g where g_def: "g = (%t. f t -
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   122
    (setsum (%m. (diff m 0 / real(fact m)) * t^m) {0..<n}
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   123
      + (B * (t^n / real(fact n)))))" by blast
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   124
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   125
  have g2: "g 0 = 0 & g h = 0"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   126
    apply (simp add: m f_h g_def del: setsum_op_ivl_Suc)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   127
    apply (cut_tac n = m and k = 1 in sumr_offset2)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   128
    apply (simp add: eq_diff_eq' diff_0 del: setsum_op_ivl_Suc)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   129
    done
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   130
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   131
  obtain difg where difg_def: "difg = (%m t. diff m t -
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   132
    (setsum (%p. (diff (m + p) 0 / real (fact p)) * (t ^ p)) {0..<n-m}
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   133
      + (B * ((t ^ (n - m)) / real (fact (n - m))))))" by blast
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   134
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   135
  have difg_0: "difg 0 = g"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   136
    unfolding difg_def g_def by (simp add: diff_0)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   137
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   138
  have difg_Suc: "\<forall>(m\<Colon>nat) t\<Colon>real.
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   139
        m < n \<and> (0\<Colon>real) \<le> t \<and> t \<le> h \<longrightarrow> DERIV (difg m) t :> difg (Suc m) t"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   140
    using diff_Suc m difg_def by (rule Maclaurin_lemma2)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   141
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   142
  have difg_eq_0: "\<forall>m. m < n --> difg m 0 = 0"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   143
    apply clarify
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   144
    apply (simp add: m difg_def)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   145
    apply (frule less_iff_Suc_add [THEN iffD1], clarify)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   146
    apply (simp del: setsum_op_ivl_Suc)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   147
    apply (insert sumr_offset4 [of 1])
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   148
    apply (simp del: setsum_op_ivl_Suc fact_Suc realpow_Suc)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   149
    done
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   150
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   151
  have isCont_difg: "\<And>m x. \<lbrakk>m < n; 0 \<le> x; x \<le> h\<rbrakk> \<Longrightarrow> isCont (difg m) x"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   152
    by (rule DERIV_isCont [OF difg_Suc [rule_format]]) simp
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   153
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   154
  have differentiable_difg:
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   155
    "\<And>m x. \<lbrakk>m < n; 0 \<le> x; x \<le> h\<rbrakk> \<Longrightarrow> difg m differentiable x"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   156
    by (rule differentiableI [OF difg_Suc [rule_format]]) simp
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   157
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   158
  have difg_Suc_eq_0: "\<And>m t. \<lbrakk>m < n; 0 \<le> t; t \<le> h; DERIV (difg m) t :> 0\<rbrakk>
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   159
        \<Longrightarrow> difg (Suc m) t = 0"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   160
    by (rule DERIV_unique [OF difg_Suc [rule_format]]) simp
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   161
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   162
  have "m < n" using m by simp
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   163
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   164
  have "\<exists>t. 0 < t \<and> t < h \<and> DERIV (difg m) t :> 0"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   165
  using `m < n`
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   166
  proof (induct m)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   167
  case 0
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   168
    show ?case
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   169
    proof (rule Rolle)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   170
      show "0 < h" by fact
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   171
      show "difg 0 0 = difg 0 h" by (simp add: difg_0 g2)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   172
      show "\<forall>x. 0 \<le> x \<and> x \<le> h \<longrightarrow> isCont (difg (0\<Colon>nat)) x"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   173
        by (simp add: isCont_difg n)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   174
      show "\<forall>x. 0 < x \<and> x < h \<longrightarrow> difg (0\<Colon>nat) differentiable x"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   175
        by (simp add: differentiable_difg n)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   176
    qed
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   177
  next
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   178
  case (Suc m')
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   179
    hence "\<exists>t. 0 < t \<and> t < h \<and> DERIV (difg m') t :> 0" by simp
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   180
    then obtain t where t: "0 < t" "t < h" "DERIV (difg m') t :> 0" by fast
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   181
    have "\<exists>t'. 0 < t' \<and> t' < t \<and> DERIV (difg (Suc m')) t' :> 0"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   182
    proof (rule Rolle)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   183
      show "0 < t" by fact
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   184
      show "difg (Suc m') 0 = difg (Suc m') t"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   185
        using t `Suc m' < n` by (simp add: difg_Suc_eq_0 difg_eq_0)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   186
      show "\<forall>x. 0 \<le> x \<and> x \<le> t \<longrightarrow> isCont (difg (Suc m')) x"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   187
        using `t < h` `Suc m' < n` by (simp add: isCont_difg)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   188
      show "\<forall>x. 0 < x \<and> x < t \<longrightarrow> difg (Suc m') differentiable x"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   189
        using `t < h` `Suc m' < n` by (simp add: differentiable_difg)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   190
    qed
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   191
    thus ?case
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   192
      using `t < h` by auto
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   193
  qed
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   194
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   195
  then obtain t where "0 < t" "t < h" "DERIV (difg m) t :> 0" by fast
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   196
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   197
  hence "difg (Suc m) t = 0"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   198
    using `m < n` by (simp add: difg_Suc_eq_0)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   199
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   200
  show ?thesis
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   201
  proof (intro exI conjI)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   202
    show "0 < t" by fact
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   203
    show "t < h" by fact
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   204
    show "f h =
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   205
      (\<Sum>m = 0..<n. diff m 0 / real (fact m) * h ^ m) +
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   206
      diff n t / real (fact n) * h ^ n"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   207
      using `difg (Suc m) t = 0`
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   208
      by (simp add: m f_h difg_def del: realpow_Suc fact_Suc)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   209
  qed
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   210
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   211
qed
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   212
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   213
lemma Maclaurin_objl:
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   214
  "0 < h & n>0 & diff 0 = f &
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   215
  (\<forall>m t. m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t)
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   216
   --> (\<exists>t. 0 < t & t < h &
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   217
            f h = (\<Sum>m=0..<n. diff m 0 / real (fact m) * h ^ m) +
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   218
                  diff n t / real (fact n) * h ^ n)"
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   219
by (blast intro: Maclaurin)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   220
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   221
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   222
lemma Maclaurin2:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   223
   "[| 0 < h; diff 0 = f;
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   224
       \<forall>m t.
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff 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: 14738
diff changeset
   226
    ==> \<exists>t. 0 < t &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   227
              t \<le> h &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   228
              f h =
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   229
              (\<Sum>m=0..<n. diff m 0 / real (fact m) * h ^ m) +
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   230
              diff n t / real (fact n) * h ^ n"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   231
apply (case_tac "n", auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   232
apply (drule Maclaurin, auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   233
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   234
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   235
lemma Maclaurin2_objl:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   236
     "0 < h & diff 0 = f &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   237
       (\<forall>m t.
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff 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: 14738
diff changeset
   239
    --> (\<exists>t. 0 < t &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   240
              t \<le> h &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   241
              f h =
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   242
              (\<Sum>m=0..<n. diff m 0 / real (fact m) * h ^ m) +
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   243
              diff n t / real (fact n) * h ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   244
by (blast intro: Maclaurin2)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   245
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   246
lemma Maclaurin_minus:
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   247
   "[| h < 0; n > 0; diff 0 = f;
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff 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: 14738
diff changeset
   249
    ==> \<exists>t. h < t &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   250
              t < 0 &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   251
              f h =
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   252
              (\<Sum>m=0..<n. diff m 0 / real (fact m) * h ^ m) +
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   253
              diff n t / real (fact n) * h ^ n"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   254
apply (cut_tac f = "%x. f (-x)"
23177
3004310c95b1 replace (- 1) with -1
huffman
parents: 23069
diff changeset
   255
        and diff = "%n x. (-1 ^ n) * diff n (-x)"
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   256
        and h = "-h" and n = n in Maclaurin_objl)
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   257
apply (simp)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   258
apply safe
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   259
apply (subst minus_mult_right)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   260
apply (rule DERIV_cmult)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   261
apply (rule lemma_DERIV_subst)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   262
apply (rule DERIV_chain2 [where g=uminus])
23069
cdfff0241c12 rename lemmas LIM_ident, isCont_ident, DERIV_ident
huffman
parents: 22985
diff changeset
   263
apply (rule_tac [2] DERIV_minus, rule_tac [2] DERIV_ident)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   264
prefer 2 apply force
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   265
apply force
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   266
apply (rule_tac x = "-t" in exI, auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff 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: 14738
diff changeset
   268
                    (\<Sum>m = 0..<n. diff m 0 * h ^ m / real(fact m))")
15536
3ce1cb7a24f0 starting to get rid of sumr
nipkow
parents: 15481
diff changeset
   269
apply (rule_tac [2] setsum_cong[OF refl])
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   270
apply (auto simp add: divide_inverse power_mult_distrib [symmetric])
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   271
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   272
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   273
lemma Maclaurin_minus_objl:
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   274
     "(h < 0 & n > 0 & diff 0 = f &
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   275
       (\<forall>m t.
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff 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: 14738
diff changeset
   277
    --> (\<exists>t. h < t &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   278
              t < 0 &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   279
              f h =
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   280
              (\<Sum>m=0..<n. diff m 0 / real (fact m) * h ^ m) +
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   281
              diff n t / real (fact n) * h ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   282
by (blast intro: Maclaurin_minus)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   283
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   284
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   285
subsection{*More Convenient "Bidirectional" Version.*}
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   286
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   287
(* not good for PVS sin_approx, cos_approx *)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   288
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   289
lemma Maclaurin_bi_le_lemma [rule_format]:
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   290
  "n>0 \<longrightarrow>
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   291
   diff 0 0 =
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   292
   (\<Sum>m = 0..<n. diff m 0 * 0 ^ m / real (fact m)) +
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   293
   diff n 0 * 0 ^ n / real (fact n)"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15234
diff changeset
   294
by (induct "n", auto)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   295
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   296
lemma Maclaurin_bi_le:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   297
   "[| diff 0 = f;
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff 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: 14738
diff changeset
   299
    ==> \<exists>t. abs t \<le> abs x &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   300
              f x =
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   301
              (\<Sum>m=0..<n. diff m 0 / real (fact m) * x ^ m) +
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   302
              diff n t / real (fact n) * x ^ n"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   303
apply (case_tac "n = 0", force)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   304
apply (case_tac "x = 0")
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   305
 apply (rule_tac x = 0 in exI)
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   306
 apply (force simp add: Maclaurin_bi_le_lemma)
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   307
apply (cut_tac x = x and y = 0 in linorder_less_linear, auto)
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   308
 txt{*Case 1, where @{term "x < 0"}*}
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   309
 apply (cut_tac f = "diff 0" and diff = diff and h = x and n = n in Maclaurin_minus_objl, safe)
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   310
  apply (simp add: abs_if)
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   311
 apply (rule_tac x = t in exI)
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   312
 apply (simp add: abs_if)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   313
txt{*Case 2, where @{term "0 < x"}*}
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   314
apply (cut_tac f = "diff 0" and diff = diff and h = x and n = n in Maclaurin_objl, safe)
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   315
 apply (simp add: abs_if)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   316
apply (rule_tac x = t in exI)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   317
apply (simp add: abs_if)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   318
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   319
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   320
lemma Maclaurin_all_lt:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   321
     "[| diff 0 = f;
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   322
         \<forall>m x. DERIV (diff m) x :> diff(Suc m) x;
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   323
        x ~= 0; n > 0
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   324
      |] ==> \<exists>t. 0 < abs t & abs t < abs x &
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   325
               f x = (\<Sum>m=0..<n. (diff m 0 / real (fact m)) * x ^ m) +
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   326
                     (diff n t / real (fact n)) * x ^ n"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   327
apply (rule_tac x = x and y = 0 in linorder_cases)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   328
prefer 2 apply blast
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   329
apply (drule_tac [2] diff=diff in Maclaurin)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   330
apply (drule_tac diff=diff in Maclaurin_minus, simp_all, safe)
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   331
apply (rule_tac [!] x = t in exI, auto)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   332
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   333
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   334
lemma Maclaurin_all_lt_objl:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   335
     "diff 0 = f &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   336
      (\<forall>m x. DERIV (diff m) x :> diff(Suc m) x) &
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   337
      x ~= 0 & n > 0
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   338
      --> (\<exists>t. 0 < abs t & abs t < abs x &
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   339
               f x = (\<Sum>m=0..<n. (diff m 0 / real (fact m)) * x ^ m) +
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   340
                     (diff n t / real (fact n)) * x ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   341
by (blast intro: Maclaurin_all_lt)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   342
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   343
lemma Maclaurin_zero [rule_format]:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   344
     "x = (0::real)
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   345
      ==> n \<noteq> 0 -->
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   346
          (\<Sum>m=0..<n. (diff m (0::real) / real (fact m)) * x ^ m) =
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   347
          diff 0 0"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   348
by (induct n, auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   349
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   350
lemma Maclaurin_all_le: "[| diff 0 = f;
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   351
        \<forall>m x. DERIV (diff m) x :> diff (Suc m) x
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   352
      |] ==> \<exists>t. abs t \<le> abs x &
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   353
              f x = (\<Sum>m=0..<n. (diff m 0 / real (fact m)) * x ^ m) +
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   354
                    (diff n t / real (fact n)) * x ^ n"
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   355
apply(cases "n=0")
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   356
apply (force)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   357
apply (case_tac "x = 0")
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   358
apply (frule_tac diff = diff and n = n in Maclaurin_zero, assumption)
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   359
apply (drule not0_implies_Suc)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   360
apply (rule_tac x = 0 in exI, force)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff 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: 14738
diff changeset
   362
apply (rule_tac x = t in exI, auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   363
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   364
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   365
lemma Maclaurin_all_le_objl: "diff 0 = f &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   366
      (\<forall>m x. DERIV (diff m) x :> diff (Suc m) x)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   367
      --> (\<exists>t. abs t \<le> abs x &
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   368
              f x = (\<Sum>m=0..<n. (diff m 0 / real (fact m)) * x ^ m) +
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   369
                    (diff n t / real (fact n)) * x ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   370
by (blast intro: Maclaurin_all_le)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   371
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   372
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   373
subsection{*Version for Exponential Function*}
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   374
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   375
lemma Maclaurin_exp_lt: "[| x ~= 0; n > 0 |]
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   376
      ==> (\<exists>t. 0 < abs t &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   377
                abs t < abs x &
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   378
                exp x = (\<Sum>m=0..<n. (x ^ m) / real (fact m)) +
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   379
                        (exp t / real (fact n)) * x ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff 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: 14738
diff changeset
   381
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   382
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   383
lemma Maclaurin_exp_le:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   384
     "\<exists>t. abs t \<le> abs x &
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   385
            exp x = (\<Sum>m=0..<n. (x ^ m) / real (fact m)) +
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   386
                       (exp t / real (fact n)) * x ^ n"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff 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: 14738
diff changeset
   388
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   389
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   390
subsection{*Version for Sine Function*}
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   391
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   392
lemma MVT2:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   393
     "[| a < b; \<forall>x. a \<le> x & x \<le> b --> DERIV f x :> f'(x) |]
21782
bf055d30d3ad add type annotation
huffman
parents: 20792
diff changeset
   394
      ==> \<exists>z::real. a < z & z < b & (f b - f a = (b - a) * f'(z))"
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   395
apply (drule MVT)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   396
apply (blast intro: DERIV_isCont)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   397
apply (force dest: order_less_imp_le simp add: differentiable_def)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   398
apply (blast dest: DERIV_unique order_less_imp_le)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   399
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   400
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   401
lemma mod_exhaust_less_4:
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   402
  "m mod 4 = 0 | m mod 4 = 1 | m mod 4 = 2 | m mod 4 = (3::nat)"
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19765
diff changeset
   403
by auto
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   404
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   405
lemma Suc_Suc_mult_two_diff_two [rule_format, simp]:
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   406
  "n\<noteq>0 --> Suc (Suc (2 * n - 2)) = 2*n"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15234
diff changeset
   407
by (induct "n", auto)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   408
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   409
lemma lemma_Suc_Suc_4n_diff_2 [rule_format, simp]:
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   410
  "n\<noteq>0 --> Suc (Suc (4*n - 2)) = 4*n"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15234
diff changeset
   411
by (induct "n", auto)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   412
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   413
lemma Suc_mult_two_diff_one [rule_format, simp]:
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   414
  "n\<noteq>0 --> Suc (2 * n - 1) = 2*n"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15234
diff changeset
   415
by (induct "n", auto)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   416
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   417
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   418
text{*It is unclear why so many variant results are needed.*}
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   419
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   420
lemma Maclaurin_sin_expansion2:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   421
     "\<exists>t. abs t \<le> abs x &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   422
       sin x =
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   423
       (\<Sum>m=0..<n. (if even m then 0
23177
3004310c95b1 replace (- 1) with -1
huffman
parents: 23069
diff changeset
   424
                       else (-1 ^ ((m - Suc 0) div 2)) / real (fact m)) *
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   425
                       x ^ m)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   426
      + ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   427
apply (cut_tac f = sin and n = n and x = x
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   428
        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: 14738
diff changeset
   429
apply safe
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   430
apply (simp (no_asm))
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   431
apply (simp (no_asm))
23242
e1526d5fa80d remove simp attribute from lemma_STAR theorems
huffman
parents: 23177
diff changeset
   432
apply (case_tac "n", clarify, simp, simp add: lemma_STAR_sin)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   433
apply (rule ccontr, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   434
apply (drule_tac x = x in spec, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   435
apply (erule ssubst)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   436
apply (rule_tac x = t in exI, simp)
15536
3ce1cb7a24f0 starting to get rid of sumr
nipkow
parents: 15481
diff changeset
   437
apply (rule setsum_cong[OF refl])
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   438
apply (auto simp add: sin_zero_iff odd_Suc_mult_two_ex)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   439
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   440
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   441
lemma Maclaurin_sin_expansion:
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   442
     "\<exists>t. sin x =
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   443
       (\<Sum>m=0..<n. (if even m then 0
23177
3004310c95b1 replace (- 1) with -1
huffman
parents: 23069
diff changeset
   444
                       else (-1 ^ ((m - Suc 0) div 2)) / real (fact m)) *
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   445
                       x ^ m)
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   446
      + ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   447
apply (insert Maclaurin_sin_expansion2 [of x n]) 
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   448
apply (blast intro: elim:); 
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   449
done
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   450
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   451
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   452
lemma Maclaurin_sin_expansion3:
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   453
     "[| n > 0; 0 < x |] ==>
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   454
       \<exists>t. 0 < t & t < x &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   455
       sin x =
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   456
       (\<Sum>m=0..<n. (if even m then 0
23177
3004310c95b1 replace (- 1) with -1
huffman
parents: 23069
diff changeset
   457
                       else (-1 ^ ((m - Suc 0) div 2)) / real (fact m)) *
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   458
                       x ^ m)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   459
      + ((sin(t + 1/2 * real(n) *pi) / real (fact n)) * x ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   460
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: 14738
diff changeset
   461
apply safe
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   462
apply simp
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   463
apply (simp (no_asm))
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   464
apply (erule ssubst)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   465
apply (rule_tac x = t in exI, simp)
15536
3ce1cb7a24f0 starting to get rid of sumr
nipkow
parents: 15481
diff changeset
   466
apply (rule setsum_cong[OF refl])
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   467
apply (auto simp add: sin_zero_iff odd_Suc_mult_two_ex)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   468
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   469
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   470
lemma Maclaurin_sin_expansion4:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   471
     "0 < x ==>
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   472
       \<exists>t. 0 < t & t \<le> x &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   473
       sin x =
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   474
       (\<Sum>m=0..<n. (if even m then 0
23177
3004310c95b1 replace (- 1) with -1
huffman
parents: 23069
diff changeset
   475
                       else (-1 ^ ((m - Suc 0) div 2)) / real (fact m)) *
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   476
                       x ^ m)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   477
      + ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   478
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: 14738
diff changeset
   479
apply safe
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   480
apply simp
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   481
apply (simp (no_asm))
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   482
apply (erule ssubst)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   483
apply (rule_tac x = t in exI, simp)
15536
3ce1cb7a24f0 starting to get rid of sumr
nipkow
parents: 15481
diff changeset
   484
apply (rule setsum_cong[OF refl])
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   485
apply (auto simp add: sin_zero_iff odd_Suc_mult_two_ex)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   486
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   487
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   488
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   489
subsection{*Maclaurin Expansion for Cosine Function*}
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   490
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   491
lemma sumr_cos_zero_one [simp]:
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   492
 "(\<Sum>m=0..<(Suc n).
23177
3004310c95b1 replace (- 1) with -1
huffman
parents: 23069
diff changeset
   493
     (if even m then -1 ^ (m div 2)/(real  (fact m)) else 0) * 0 ^ m) = 1"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15234
diff changeset
   494
by (induct "n", auto)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   495
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   496
lemma Maclaurin_cos_expansion:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   497
     "\<exists>t. abs t \<le> abs x &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   498
       cos x =
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   499
       (\<Sum>m=0..<n. (if even m
23177
3004310c95b1 replace (- 1) with -1
huffman
parents: 23069
diff changeset
   500
                       then -1 ^ (m div 2)/(real (fact m))
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   501
                       else 0) *
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   502
                       x ^ m)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   503
      + ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   504
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: 14738
diff changeset
   505
apply safe
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   506
apply (simp (no_asm))
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   507
apply (simp (no_asm))
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   508
apply (case_tac "n", simp)
15561
045a07ac35a7 another reorganization of setsums and intervals
nipkow
parents: 15539
diff changeset
   509
apply (simp del: setsum_op_ivl_Suc)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   510
apply (rule ccontr, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   511
apply (drule_tac x = x in spec, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   512
apply (erule ssubst)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   513
apply (rule_tac x = t in exI, simp)
15536
3ce1cb7a24f0 starting to get rid of sumr
nipkow
parents: 15481
diff changeset
   514
apply (rule setsum_cong[OF refl])
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   515
apply (auto simp add: cos_zero_iff even_mult_two_ex)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   516
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   517
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   518
lemma Maclaurin_cos_expansion2:
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   519
     "[| 0 < x; n > 0 |] ==>
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   520
       \<exists>t. 0 < t & t < x &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   521
       cos x =
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   522
       (\<Sum>m=0..<n. (if even m
23177
3004310c95b1 replace (- 1) with -1
huffman
parents: 23069
diff changeset
   523
                       then -1 ^ (m div 2)/(real (fact m))
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   524
                       else 0) *
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   525
                       x ^ m)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   526
      + ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   527
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: 14738
diff changeset
   528
apply safe
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   529
apply simp
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   530
apply (simp (no_asm))
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   531
apply (erule ssubst)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   532
apply (rule_tac x = t in exI, simp)
15536
3ce1cb7a24f0 starting to get rid of sumr
nipkow
parents: 15481
diff changeset
   533
apply (rule setsum_cong[OF refl])
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   534
apply (auto simp add: cos_zero_iff even_mult_two_ex)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   535
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   536
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   537
lemma Maclaurin_minus_cos_expansion:
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   538
     "[| x < 0; n > 0 |] ==>
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   539
       \<exists>t. x < t & t < 0 &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   540
       cos x =
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   541
       (\<Sum>m=0..<n. (if even m
23177
3004310c95b1 replace (- 1) with -1
huffman
parents: 23069
diff changeset
   542
                       then -1 ^ (m div 2)/(real (fact m))
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   543
                       else 0) *
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   544
                       x ^ m)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   545
      + ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   546
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: 14738
diff changeset
   547
apply safe
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   548
apply simp
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   549
apply (simp (no_asm))
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   550
apply (erule ssubst)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   551
apply (rule_tac x = t in exI, simp)
15536
3ce1cb7a24f0 starting to get rid of sumr
nipkow
parents: 15481
diff changeset
   552
apply (rule setsum_cong[OF refl])
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   553
apply (auto simp add: cos_zero_iff even_mult_two_ex)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   554
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   555
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   556
(* ------------------------------------------------------------------------- *)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   557
(* Version for ln(1 +/- x). Where is it??                                    *)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   558
(* ------------------------------------------------------------------------- *)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   559
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   560
lemma sin_bound_lemma:
15081
32402f5624d1 abs notation
paulson
parents: 15079
diff changeset
   561
    "[|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: 14738
diff changeset
   562
by auto
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   563
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   564
lemma Maclaurin_sin_bound:
23177
3004310c95b1 replace (- 1) with -1
huffman
parents: 23069
diff changeset
   565
  "abs(sin x - (\<Sum>m=0..<n. (if even m then 0 else (-1 ^ ((m - Suc 0) div 2)) / real (fact m)) *
15081
32402f5624d1 abs notation
paulson
parents: 15079
diff changeset
   566
  x ^ m))  \<le> inverse(real (fact n)) * \<bar>x\<bar> ^ n"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   567
proof -
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   568
  have "!! x (y::real). x \<le> 1 \<Longrightarrow> 0 \<le> y \<Longrightarrow> x * y \<le> 1 * y"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   569
    by (rule_tac mult_right_mono,simp_all)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   570
  note est = this[simplified]
22985
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   571
  let ?diff = "\<lambda>(n::nat) 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)"
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   572
  have diff_0: "?diff 0 = sin" by simp
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   573
  have DERIV_diff: "\<forall>m x. DERIV (?diff m) x :> ?diff (Suc m) x"
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   574
    apply (clarify)
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   575
    apply (subst (1 2 3) mod_Suc_eq_Suc_mod)
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   576
    apply (cut_tac m=m in mod_exhaust_less_4)
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   577
    apply (safe, simp_all)
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   578
    apply (rule DERIV_minus, simp)
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   579
    apply (rule lemma_DERIV_subst, rule DERIV_minus, rule DERIV_cos, simp)
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   580
    done
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   581
  from Maclaurin_all_le [OF diff_0 DERIV_diff]
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   582
  obtain t where t1: "\<bar>t\<bar> \<le> \<bar>x\<bar>" and
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   583
    t2: "sin x = (\<Sum>m = 0..<n. ?diff m 0 / real (fact m) * x ^ m) +
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   584
      ?diff n t / real (fact n) * x ^ n" by fast
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   585
  have diff_m_0:
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   586
    "\<And>m. ?diff m 0 = (if even m then 0
23177
3004310c95b1 replace (- 1) with -1
huffman
parents: 23069
diff changeset
   587
         else -1 ^ ((m - Suc 0) div 2))"
22985
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   588
    apply (subst even_even_mod_4_iff)
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   589
    apply (cut_tac m=m in mod_exhaust_less_4)
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   590
    apply (elim disjE, simp_all)
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   591
    apply (safe dest!: mod_eqD, simp_all)
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   592
    done
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   593
  show ?thesis
22985
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   594
    apply (subst t2)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   595
    apply (rule sin_bound_lemma)
15536
3ce1cb7a24f0 starting to get rid of sumr
nipkow
parents: 15481
diff changeset
   596
    apply (rule setsum_cong[OF refl])
22985
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   597
    apply (subst diff_m_0, simp)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   598
    apply (auto intro: mult_right_mono [where b=1, simplified] mult_right_mono
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 15944
diff changeset
   599
                   simp add: est mult_nonneg_nonneg mult_ac divide_inverse
16924
04246269386e removed the dependence on abs_mult
paulson
parents: 16819
diff changeset
   600
                          power_abs [symmetric] abs_mult)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   601
    done
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   602
qed
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   603
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   604
end