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