src/HOL/MacLaurin.thy
author huffman
Fri, 19 Aug 2011 08:39:43 -0700
changeset 45166 33572a766836
parent 41412 4b2a457b17e8
child 45168 d2a6f9af02f4
permissions -rw-r--r--
fold definitions of sin_coeff and cos_coeff in Maclaurin lemmas
haftmann@28952
     1
(*  Author      : Jacques D. Fleuriot
paulson@12224
     2
    Copyright   : 2001 University of Edinburgh
paulson@15079
     3
    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
bulwahn@41368
     4
    Conversion of Mac Laurin to Isar by Lukas Bulwahn and Bernhard Häupler, 2005
paulson@12224
     5
*)
paulson@12224
     6
paulson@15944
     7
header{*MacLaurin Series*}
paulson@15944
     8
nipkow@15131
     9
theory MacLaurin
chaieb@29748
    10
imports Transcendental
nipkow@15131
    11
begin
obua@14738
    12
paulson@15079
    13
subsection{*Maclaurin's Theorem with Lagrange Form of Remainder*}
paulson@15079
    14
paulson@15079
    15
text{*This is a very long, messy proof even now that it's been broken down
paulson@15079
    16
into lemmas.*}
paulson@15079
    17
paulson@15079
    18
lemma Maclaurin_lemma:
paulson@15079
    19
    "0 < h ==>
nipkow@15539
    20
     \<exists>B. f h = (\<Sum>m=0..<n. (j m / real (fact m)) * (h^m)) +
paulson@15079
    21
               (B * ((h^n) / real(fact n)))"
hoelzl@41412
    22
by (rule exI[where x = "(f h - (\<Sum>m=0..<n. (j m / real (fact m)) * h^m)) *
hoelzl@41412
    23
                 real(fact n) / (h^n)"]) simp
paulson@15079
    24
paulson@15079
    25
lemma eq_diff_eq': "(x = y - z) = (y = x + (z::real))"
paulson@15079
    26
by arith
paulson@15079
    27
avigad@32031
    28
lemma fact_diff_Suc [rule_format]:
avigad@32031
    29
  "n < Suc m ==> fact (Suc m - n) = (Suc m - n) * fact (m - n)"
avigad@32031
    30
  by (subst fact_reduce_nat, auto)
avigad@32031
    31
paulson@15079
    32
lemma Maclaurin_lemma2:
hoelzl@41412
    33
  fixes B
bulwahn@41368
    34
  assumes DERIV : "\<forall>m t. m < n \<and> 0\<le>t \<and> t\<le>h \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
hoelzl@41412
    35
    and INIT : "n = Suc k"
hoelzl@41412
    36
  defines "difg \<equiv> (\<lambda>m t. diff m t - ((\<Sum>p = 0..<n - m. diff (m + p) 0 / real (fact p) * t ^ p) +
hoelzl@41412
    37
    B * (t ^ (n - m) / real (fact (n - m)))))" (is "difg \<equiv> (\<lambda>m t. diff m t - ?difg m t)")
bulwahn@41368
    38
  shows "\<forall>m t. m < n & 0 \<le> t & t \<le> h --> DERIV (difg m) t :> difg (Suc m) t"
hoelzl@41412
    39
proof (rule allI impI)+
hoelzl@41412
    40
  fix m t assume INIT2: "m < n & 0 \<le> t & t \<le> h"
hoelzl@41412
    41
  have "DERIV (difg m) t :> diff (Suc m) t -
hoelzl@41412
    42
    ((\<Sum>x = 0..<n - m. real x * t ^ (x - Suc 0) * diff (m + x) 0 / real (fact x)) +
hoelzl@41412
    43
     real (n - m) * t ^ (n - Suc m) * B / real (fact (n - m)))" unfolding difg_def
hoelzl@41412
    44
    by (auto intro!: DERIV_intros DERIV[rule_format, OF INIT2])
bulwahn@41368
    45
      moreover
hoelzl@41412
    46
  from INIT2 have intvl: "{..<n - m} = insert 0 (Suc ` {..<n - Suc m})" and "0 < n - m"
hoelzl@41412
    47
    unfolding atLeast0LessThan[symmetric] by auto
hoelzl@41412
    48
  have "(\<Sum>x = 0..<n - m. real x * t ^ (x - Suc 0) * diff (m + x) 0 / real (fact x)) =
hoelzl@41412
    49
      (\<Sum>x = 0..<n - Suc m. real (Suc x) * t ^ x * diff (Suc m + x) 0 / real (fact (Suc x)))"
hoelzl@41412
    50
    unfolding intvl atLeast0LessThan by (subst setsum.insert) (auto simp: setsum.reindex)
hoelzl@41412
    51
  moreover
hoelzl@41412
    52
  have fact_neq_0: "\<And>x::nat. real (fact x) + real x * real (fact x) \<noteq> 0"
hoelzl@41412
    53
    by (metis fact_gt_zero_nat not_add_less1 real_of_nat_add real_of_nat_mult real_of_nat_zero_iff)
hoelzl@41412
    54
  have "\<And>x. real (Suc x) * t ^ x * diff (Suc m + x) 0 / real (fact (Suc x)) =
hoelzl@41412
    55
      diff (Suc m + x) 0 * t^x / real (fact x)"
hoelzl@41412
    56
    by (auto simp: field_simps real_of_nat_Suc fact_neq_0 intro!: nonzero_divide_eq_eq[THEN iffD2])
hoelzl@41412
    57
  moreover
hoelzl@41412
    58
  have "real (n - m) * t ^ (n - Suc m) * B / real (fact (n - m)) =
hoelzl@41412
    59
      B * (t ^ (n - Suc m) / real (fact (n - Suc m)))"
hoelzl@41412
    60
    using `0 < n - m` by (simp add: fact_reduce_nat)
hoelzl@41412
    61
  ultimately show "DERIV (difg m) t :> difg (Suc m) t"
hoelzl@41412
    62
    unfolding difg_def by simp
bulwahn@41368
    63
qed
avigad@32031
    64
paulson@15079
    65
lemma Maclaurin:
huffman@29187
    66
  assumes h: "0 < h"
huffman@29187
    67
  assumes n: "0 < n"
huffman@29187
    68
  assumes diff_0: "diff 0 = f"
huffman@29187
    69
  assumes diff_Suc:
huffman@29187
    70
    "\<forall>m t. m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t"
huffman@29187
    71
  shows
huffman@29187
    72
    "\<exists>t. 0 < t & t < h &
paulson@15079
    73
              f h =
nipkow@15539
    74
              setsum (%m. (diff m 0 / real (fact m)) * h ^ m) {0..<n} +
hoelzl@41412
    75
              (diff n t / real (fact n)) * h ^ n"
huffman@29187
    76
proof -
huffman@29187
    77
  from n obtain m where m: "n = Suc m"
hoelzl@41412
    78
    by (cases n) (simp add: n)
huffman@29187
    79
huffman@29187
    80
  obtain B where f_h: "f h =
huffman@29187
    81
        (\<Sum>m = 0..<n. diff m (0\<Colon>real) / real (fact m) * h ^ m) +
huffman@29187
    82
        B * (h ^ n / real (fact n))"
huffman@29187
    83
    using Maclaurin_lemma [OF h] ..
huffman@29187
    84
hoelzl@41412
    85
  def g \<equiv> "(\<lambda>t. f t -
hoelzl@41412
    86
    (setsum (\<lambda>m. (diff m 0 / real(fact m)) * t^m) {0..<n}
hoelzl@41412
    87
      + (B * (t^n / real(fact n)))))"
huffman@29187
    88
huffman@29187
    89
  have g2: "g 0 = 0 & g h = 0"
huffman@29187
    90
    apply (simp add: m f_h g_def del: setsum_op_ivl_Suc)
huffman@30019
    91
    apply (cut_tac n = m and k = "Suc 0" in sumr_offset2)
huffman@29187
    92
    apply (simp add: eq_diff_eq' diff_0 del: setsum_op_ivl_Suc)
huffman@29187
    93
    done
huffman@29187
    94
hoelzl@41412
    95
  def difg \<equiv> "(%m t. diff m t -
huffman@29187
    96
    (setsum (%p. (diff (m + p) 0 / real (fact p)) * (t ^ p)) {0..<n-m}
hoelzl@41412
    97
      + (B * ((t ^ (n - m)) / real (fact (n - m))))))"
huffman@29187
    98
huffman@29187
    99
  have difg_0: "difg 0 = g"
huffman@29187
   100
    unfolding difg_def g_def by (simp add: diff_0)
huffman@29187
   101
huffman@29187
   102
  have difg_Suc: "\<forall>(m\<Colon>nat) t\<Colon>real.
huffman@29187
   103
        m < n \<and> (0\<Colon>real) \<le> t \<and> t \<le> h \<longrightarrow> DERIV (difg m) t :> difg (Suc m) t"
hoelzl@41412
   104
    using diff_Suc m unfolding difg_def by (rule Maclaurin_lemma2)
huffman@29187
   105
huffman@29187
   106
  have difg_eq_0: "\<forall>m. m < n --> difg m 0 = 0"
huffman@29187
   107
    apply clarify
huffman@29187
   108
    apply (simp add: m difg_def)
huffman@29187
   109
    apply (frule less_iff_Suc_add [THEN iffD1], clarify)
huffman@29187
   110
    apply (simp del: setsum_op_ivl_Suc)
huffman@30019
   111
    apply (insert sumr_offset4 [of "Suc 0"])
avigad@32039
   112
    apply (simp del: setsum_op_ivl_Suc fact_Suc)
huffman@29187
   113
    done
huffman@29187
   114
huffman@29187
   115
  have isCont_difg: "\<And>m x. \<lbrakk>m < n; 0 \<le> x; x \<le> h\<rbrakk> \<Longrightarrow> isCont (difg m) x"
huffman@29187
   116
    by (rule DERIV_isCont [OF difg_Suc [rule_format]]) simp
huffman@29187
   117
huffman@29187
   118
  have differentiable_difg:
huffman@29187
   119
    "\<And>m x. \<lbrakk>m < n; 0 \<le> x; x \<le> h\<rbrakk> \<Longrightarrow> difg m differentiable x"
huffman@29187
   120
    by (rule differentiableI [OF difg_Suc [rule_format]]) simp
huffman@29187
   121
huffman@29187
   122
  have difg_Suc_eq_0: "\<And>m t. \<lbrakk>m < n; 0 \<le> t; t \<le> h; DERIV (difg m) t :> 0\<rbrakk>
huffman@29187
   123
        \<Longrightarrow> difg (Suc m) t = 0"
huffman@29187
   124
    by (rule DERIV_unique [OF difg_Suc [rule_format]]) simp
huffman@29187
   125
huffman@29187
   126
  have "m < n" using m by simp
huffman@29187
   127
huffman@29187
   128
  have "\<exists>t. 0 < t \<and> t < h \<and> DERIV (difg m) t :> 0"
huffman@29187
   129
  using `m < n`
huffman@29187
   130
  proof (induct m)
hoelzl@41412
   131
    case 0
huffman@29187
   132
    show ?case
huffman@29187
   133
    proof (rule Rolle)
huffman@29187
   134
      show "0 < h" by fact
huffman@29187
   135
      show "difg 0 0 = difg 0 h" by (simp add: difg_0 g2)
huffman@29187
   136
      show "\<forall>x. 0 \<le> x \<and> x \<le> h \<longrightarrow> isCont (difg (0\<Colon>nat)) x"
huffman@29187
   137
        by (simp add: isCont_difg n)
huffman@29187
   138
      show "\<forall>x. 0 < x \<and> x < h \<longrightarrow> difg (0\<Colon>nat) differentiable x"
huffman@29187
   139
        by (simp add: differentiable_difg n)
huffman@29187
   140
    qed
huffman@29187
   141
  next
hoelzl@41412
   142
    case (Suc m')
huffman@29187
   143
    hence "\<exists>t. 0 < t \<and> t < h \<and> DERIV (difg m') t :> 0" by simp
huffman@29187
   144
    then obtain t where t: "0 < t" "t < h" "DERIV (difg m') t :> 0" by fast
huffman@29187
   145
    have "\<exists>t'. 0 < t' \<and> t' < t \<and> DERIV (difg (Suc m')) t' :> 0"
huffman@29187
   146
    proof (rule Rolle)
huffman@29187
   147
      show "0 < t" by fact
huffman@29187
   148
      show "difg (Suc m') 0 = difg (Suc m') t"
huffman@29187
   149
        using t `Suc m' < n` by (simp add: difg_Suc_eq_0 difg_eq_0)
huffman@29187
   150
      show "\<forall>x. 0 \<le> x \<and> x \<le> t \<longrightarrow> isCont (difg (Suc m')) x"
huffman@29187
   151
        using `t < h` `Suc m' < n` by (simp add: isCont_difg)
huffman@29187
   152
      show "\<forall>x. 0 < x \<and> x < t \<longrightarrow> difg (Suc m') differentiable x"
huffman@29187
   153
        using `t < h` `Suc m' < n` by (simp add: differentiable_difg)
huffman@29187
   154
    qed
huffman@29187
   155
    thus ?case
huffman@29187
   156
      using `t < h` by auto
huffman@29187
   157
  qed
huffman@29187
   158
huffman@29187
   159
  then obtain t where "0 < t" "t < h" "DERIV (difg m) t :> 0" by fast
huffman@29187
   160
huffman@29187
   161
  hence "difg (Suc m) t = 0"
huffman@29187
   162
    using `m < n` by (simp add: difg_Suc_eq_0)
huffman@29187
   163
huffman@29187
   164
  show ?thesis
huffman@29187
   165
  proof (intro exI conjI)
huffman@29187
   166
    show "0 < t" by fact
huffman@29187
   167
    show "t < h" by fact
huffman@29187
   168
    show "f h =
huffman@29187
   169
      (\<Sum>m = 0..<n. diff m 0 / real (fact m) * h ^ m) +
huffman@29187
   170
      diff n t / real (fact n) * h ^ n"
huffman@29187
   171
      using `difg (Suc m) t = 0`
avigad@32039
   172
      by (simp add: m f_h difg_def del: fact_Suc)
huffman@29187
   173
  qed
huffman@29187
   174
qed
paulson@15079
   175
paulson@15079
   176
lemma Maclaurin_objl:
nipkow@25162
   177
  "0 < h & n>0 & diff 0 = f &
nipkow@25134
   178
  (\<forall>m t. m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t)
nipkow@25134
   179
   --> (\<exists>t. 0 < t & t < h &
nipkow@25134
   180
            f h = (\<Sum>m=0..<n. diff m 0 / real (fact m) * h ^ m) +
nipkow@25134
   181
                  diff n t / real (fact n) * h ^ n)"
paulson@15079
   182
by (blast intro: Maclaurin)
paulson@15079
   183
paulson@15079
   184
paulson@15079
   185
lemma Maclaurin2:
bulwahn@41368
   186
  assumes INIT1: "0 < h " and INIT2: "diff 0 = f"
bulwahn@41368
   187
  and DERIV: "\<forall>m t.
bulwahn@41368
   188
  m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t"
bulwahn@41368
   189
  shows "\<exists>t. 0 < t \<and> t \<le> h \<and> f h =
bulwahn@41368
   190
  (\<Sum>m=0..<n. diff m 0 / real (fact m) * h ^ m) +
bulwahn@41368
   191
  diff n t / real (fact n) * h ^ n"
bulwahn@41368
   192
proof (cases "n")
bulwahn@41368
   193
  case 0 with INIT1 INIT2 show ?thesis by fastsimp
bulwahn@41368
   194
next
hoelzl@41412
   195
  case Suc
bulwahn@41368
   196
  hence "n > 0" by simp
bulwahn@41368
   197
  from INIT1 this INIT2 DERIV have "\<exists>t>0. t < h \<and>
bulwahn@41368
   198
    f h =
hoelzl@41412
   199
    (\<Sum>m = 0..<n. diff m 0 / real (fact m) * h ^ m) + diff n t / real (fact n) * h ^ n"
bulwahn@41368
   200
    by (rule Maclaurin)
bulwahn@41368
   201
  thus ?thesis by fastsimp
bulwahn@41368
   202
qed
paulson@15079
   203
paulson@15079
   204
lemma Maclaurin2_objl:
paulson@15079
   205
     "0 < h & diff 0 = f &
paulson@15079
   206
       (\<forall>m t.
paulson@15079
   207
          m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t)
paulson@15079
   208
    --> (\<exists>t. 0 < t &
paulson@15079
   209
              t \<le> h &
paulson@15079
   210
              f h =
nipkow@15539
   211
              (\<Sum>m=0..<n. diff m 0 / real (fact m) * h ^ m) +
paulson@15079
   212
              diff n t / real (fact n) * h ^ n)"
paulson@15079
   213
by (blast intro: Maclaurin2)
paulson@15079
   214
paulson@15079
   215
lemma Maclaurin_minus:
hoelzl@41412
   216
  assumes "h < 0" "0 < n" "diff 0 = f"
hoelzl@41412
   217
  and DERIV: "\<forall>m t. m < n & h \<le> t & t \<le> 0 --> DERIV (diff m) t :> diff (Suc m) t"
bulwahn@41368
   218
  shows "\<exists>t. h < t & t < 0 &
bulwahn@41368
   219
         f h = (\<Sum>m=0..<n. diff m 0 / real (fact m) * h ^ m) +
bulwahn@41368
   220
         diff n t / real (fact n) * h ^ n"
bulwahn@41368
   221
proof -
hoelzl@41412
   222
  txt "Transform @{text ABL'} into @{text DERIV_intros} format."
hoelzl@41412
   223
  note DERIV' = DERIV_chain'[OF _ DERIV[rule_format], THEN DERIV_cong]
hoelzl@41412
   224
  from assms
hoelzl@41412
   225
  have "\<exists>t>0. t < - h \<and>
bulwahn@41368
   226
    f (- (- h)) =
bulwahn@41368
   227
    (\<Sum>m = 0..<n.
bulwahn@41368
   228
    (- 1) ^ m * diff m (- 0) / real (fact m) * (- h) ^ m) +
hoelzl@41412
   229
    (- 1) ^ n * diff n (- t) / real (fact n) * (- h) ^ n"
hoelzl@41412
   230
    by (intro Maclaurin) (auto intro!: DERIV_intros DERIV')
hoelzl@41412
   231
  then guess t ..
bulwahn@41368
   232
  moreover
bulwahn@41368
   233
  have "-1 ^ n * diff n (- t) * (- h) ^ n / real (fact n) = diff n (- t) * h ^ n / real (fact n)"
bulwahn@41368
   234
    by (auto simp add: power_mult_distrib[symmetric])
bulwahn@41368
   235
  moreover
bulwahn@41368
   236
  have "(SUM m = 0..<n. -1 ^ m * diff m 0 * (- h) ^ m / real (fact m)) = (SUM m = 0..<n. diff m 0 * h ^ m / real (fact m))"
bulwahn@41368
   237
    by (auto intro: setsum_cong simp add: power_mult_distrib[symmetric])
bulwahn@41368
   238
  ultimately have " h < - t \<and>
bulwahn@41368
   239
    - t < 0 \<and>
bulwahn@41368
   240
    f h =
bulwahn@41368
   241
    (\<Sum>m = 0..<n. diff m 0 / real (fact m) * h ^ m) + diff n (- t) / real (fact n) * h ^ n"
bulwahn@41368
   242
    by auto
bulwahn@41368
   243
  thus ?thesis ..
bulwahn@41368
   244
qed
paulson@15079
   245
paulson@15079
   246
lemma Maclaurin_minus_objl:
nipkow@25162
   247
     "(h < 0 & n > 0 & diff 0 = f &
paulson@15079
   248
       (\<forall>m t.
paulson@15079
   249
          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 =
nipkow@15539
   253
              (\<Sum>m=0..<n. diff m 0 / real (fact m) * h ^ m) +
paulson@15079
   254
              diff n t / real (fact n) * h ^ n)"
paulson@15079
   255
by (blast intro: Maclaurin_minus)
paulson@15079
   256
paulson@15079
   257
paulson@15079
   258
subsection{*More Convenient "Bidirectional" Version.*}
paulson@15079
   259
paulson@15079
   260
(* not good for PVS sin_approx, cos_approx *)
paulson@15079
   261
paulson@15079
   262
lemma Maclaurin_bi_le_lemma [rule_format]:
nipkow@25162
   263
  "n>0 \<longrightarrow>
nipkow@25134
   264
   diff 0 0 =
nipkow@25134
   265
   (\<Sum>m = 0..<n. diff m 0 * 0 ^ m / real (fact m)) +
nipkow@25134
   266
   diff n 0 * 0 ^ n / real (fact n)"
hoelzl@41412
   267
by (induct "n") auto
paulson@15079
   268
paulson@15079
   269
lemma Maclaurin_bi_le:
hoelzl@41412
   270
   assumes "diff 0 = f"
bulwahn@41368
   271
   and DERIV : "\<forall>m t. m < n & abs t \<le> abs x --> DERIV (diff m) t :> diff (Suc m) t"
bulwahn@41368
   272
   shows "\<exists>t. abs t \<le> abs x &
paulson@15079
   273
              f x =
nipkow@15539
   274
              (\<Sum>m=0..<n. diff m 0 / real (fact m) * x ^ m) +
hoelzl@41412
   275
     diff n t / real (fact n) * x ^ n" (is "\<exists>t. _ \<and> f x = ?f x t")
hoelzl@41412
   276
proof cases
hoelzl@41412
   277
  assume "n = 0" with `diff 0 = f` show ?thesis by force
bulwahn@41368
   278
next
hoelzl@41412
   279
  assume "n \<noteq> 0"
hoelzl@41412
   280
  show ?thesis
hoelzl@41412
   281
  proof (cases rule: linorder_cases)
hoelzl@41412
   282
    assume "x = 0" with `n \<noteq> 0` `diff 0 = f` DERIV
hoelzl@41412
   283
    have "\<bar>0\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x 0" by (force simp add: Maclaurin_bi_le_lemma)
hoelzl@41412
   284
    thus ?thesis ..
bulwahn@41368
   285
  next
hoelzl@41412
   286
    assume "x < 0"
hoelzl@41412
   287
    with `n \<noteq> 0` DERIV
hoelzl@41412
   288
    have "\<exists>t>x. t < 0 \<and> diff 0 x = ?f x t" by (intro Maclaurin_minus) auto
hoelzl@41412
   289
    then guess t ..
hoelzl@41412
   290
    with `x < 0` `diff 0 = f` have "\<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x t" by simp
hoelzl@41412
   291
    thus ?thesis ..
hoelzl@41412
   292
  next
hoelzl@41412
   293
    assume "x > 0"
hoelzl@41412
   294
    with `n \<noteq> 0` `diff 0 = f` DERIV
hoelzl@41412
   295
    have "\<exists>t>0. t < x \<and> diff 0 x = ?f x t" by (intro Maclaurin) auto
hoelzl@41412
   296
    then guess t ..
hoelzl@41412
   297
    with `x > 0` `diff 0 = f` have "\<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x t" by simp
hoelzl@41412
   298
    thus ?thesis ..
bulwahn@41368
   299
  qed
bulwahn@41368
   300
qed
bulwahn@41368
   301
paulson@15079
   302
lemma Maclaurin_all_lt:
bulwahn@41368
   303
  assumes INIT1: "diff 0 = f" and INIT2: "0 < n" and INIT3: "x \<noteq> 0"
bulwahn@41368
   304
  and DERIV: "\<forall>m x. DERIV (diff m) x :> diff(Suc m) x"
hoelzl@41412
   305
  shows "\<exists>t. 0 < abs t & abs t < abs x & f x =
hoelzl@41412
   306
    (\<Sum>m=0..<n. (diff m 0 / real (fact m)) * x ^ m) +
hoelzl@41412
   307
                (diff n t / real (fact n)) * x ^ n" (is "\<exists>t. _ \<and> _ \<and> f x = ?f x t")
hoelzl@41412
   308
proof (cases rule: linorder_cases)
hoelzl@41412
   309
  assume "x = 0" with INIT3 show "?thesis"..
hoelzl@41412
   310
next
hoelzl@41412
   311
  assume "x < 0"
hoelzl@41412
   312
  with assms have "\<exists>t>x. t < 0 \<and> f x = ?f x t" by (intro Maclaurin_minus) auto
hoelzl@41412
   313
  then guess t ..
hoelzl@41412
   314
  with `x < 0` have "0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t" by simp
hoelzl@41412
   315
  thus ?thesis ..
hoelzl@41412
   316
next
hoelzl@41412
   317
  assume "x > 0"
hoelzl@41412
   318
  with assms have "\<exists>t>0. t < x \<and> f x = ?f x t " by (intro Maclaurin) auto
hoelzl@41412
   319
  then guess t ..
hoelzl@41412
   320
  with `x > 0` have "0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t" by simp
hoelzl@41412
   321
  thus ?thesis ..
bulwahn@41368
   322
qed
bulwahn@41368
   323
paulson@15079
   324
paulson@15079
   325
lemma Maclaurin_all_lt_objl:
paulson@15079
   326
     "diff 0 = f &
paulson@15079
   327
      (\<forall>m x. DERIV (diff m) x :> diff(Suc m) x) &
nipkow@25162
   328
      x ~= 0 & n > 0
paulson@15079
   329
      --> (\<exists>t. 0 < abs t & abs t < abs x &
nipkow@15539
   330
               f x = (\<Sum>m=0..<n. (diff m 0 / real (fact m)) * x ^ m) +
paulson@15079
   331
                     (diff n t / real (fact n)) * x ^ n)"
paulson@15079
   332
by (blast intro: Maclaurin_all_lt)
paulson@15079
   333
paulson@15079
   334
lemma Maclaurin_zero [rule_format]:
paulson@15079
   335
     "x = (0::real)
nipkow@25134
   336
      ==> n \<noteq> 0 -->
nipkow@15539
   337
          (\<Sum>m=0..<n. (diff m (0::real) / real (fact m)) * x ^ m) =
paulson@15079
   338
          diff 0 0"
paulson@15079
   339
by (induct n, auto)
paulson@15079
   340
bulwahn@41368
   341
bulwahn@41368
   342
lemma Maclaurin_all_le:
bulwahn@41368
   343
  assumes INIT: "diff 0 = f"
bulwahn@41368
   344
  and DERIV: "\<forall>m x. DERIV (diff m) x :> diff (Suc m) x"
hoelzl@41412
   345
  shows "\<exists>t. abs t \<le> abs x & f x =
hoelzl@41412
   346
    (\<Sum>m=0..<n. (diff m 0 / real (fact m)) * x ^ m) +
hoelzl@41412
   347
    (diff n t / real (fact n)) * x ^ n" (is "\<exists>t. _ \<and> f x = ?f x t")
hoelzl@41412
   348
proof cases
hoelzl@41412
   349
  assume "n = 0" with INIT show ?thesis by force
bulwahn@41368
   350
  next
hoelzl@41412
   351
  assume "n \<noteq> 0"
hoelzl@41412
   352
  show ?thesis
hoelzl@41412
   353
  proof cases
hoelzl@41412
   354
    assume "x = 0"
hoelzl@41412
   355
    with `n \<noteq> 0` have "(\<Sum>m = 0..<n. diff m 0 / real (fact m) * x ^ m) = diff 0 0"
hoelzl@41412
   356
      by (intro Maclaurin_zero) auto
hoelzl@41412
   357
    with INIT `x = 0` `n \<noteq> 0` have " \<bar>0\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x 0" by force
hoelzl@41412
   358
    thus ?thesis ..
hoelzl@41412
   359
  next
hoelzl@41412
   360
    assume "x \<noteq> 0"
hoelzl@41412
   361
    with INIT `n \<noteq> 0` DERIV have "\<exists>t. 0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t"
hoelzl@41412
   362
      by (intro Maclaurin_all_lt) auto
hoelzl@41412
   363
    then guess t ..
hoelzl@41412
   364
    hence "\<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x t" by simp
hoelzl@41412
   365
    thus ?thesis ..
bulwahn@41368
   366
  qed
bulwahn@41368
   367
qed
bulwahn@41368
   368
paulson@15079
   369
lemma Maclaurin_all_le_objl: "diff 0 = f &
paulson@15079
   370
      (\<forall>m x. DERIV (diff m) x :> diff (Suc m) x)
paulson@15079
   371
      --> (\<exists>t. abs t \<le> abs x &
nipkow@15539
   372
              f x = (\<Sum>m=0..<n. (diff m 0 / real (fact m)) * x ^ m) +
paulson@15079
   373
                    (diff n t / real (fact n)) * x ^ n)"
paulson@15079
   374
by (blast intro: Maclaurin_all_le)
paulson@15079
   375
paulson@15079
   376
paulson@15079
   377
subsection{*Version for Exponential Function*}
paulson@15079
   378
nipkow@25162
   379
lemma Maclaurin_exp_lt: "[| x ~= 0; n > 0 |]
paulson@15079
   380
      ==> (\<exists>t. 0 < abs t &
paulson@15079
   381
                abs t < abs x &
nipkow@15539
   382
                exp x = (\<Sum>m=0..<n. (x ^ m) / real (fact m)) +
paulson@15079
   383
                        (exp t / real (fact n)) * x ^ n)"
paulson@15079
   384
by (cut_tac diff = "%n. exp" and f = exp and x = x and n = n in Maclaurin_all_lt_objl, auto)
paulson@15079
   385
paulson@15079
   386
paulson@15079
   387
lemma Maclaurin_exp_le:
paulson@15079
   388
     "\<exists>t. abs t \<le> abs x &
nipkow@15539
   389
            exp x = (\<Sum>m=0..<n. (x ^ m) / real (fact m)) +
paulson@15079
   390
                       (exp t / real (fact n)) * x ^ n"
paulson@15079
   391
by (cut_tac diff = "%n. exp" and f = exp and x = x and n = n in Maclaurin_all_le_objl, auto)
paulson@15079
   392
paulson@15079
   393
paulson@15079
   394
subsection{*Version for Sine Function*}
paulson@15079
   395
paulson@15079
   396
lemma mod_exhaust_less_4:
nipkow@25134
   397
  "m mod 4 = 0 | m mod 4 = 1 | m mod 4 = 2 | m mod 4 = (3::nat)"
webertj@20217
   398
by auto
paulson@15079
   399
paulson@15079
   400
lemma Suc_Suc_mult_two_diff_two [rule_format, simp]:
nipkow@25134
   401
  "n\<noteq>0 --> Suc (Suc (2 * n - 2)) = 2*n"
paulson@15251
   402
by (induct "n", auto)
paulson@15079
   403
paulson@15079
   404
lemma lemma_Suc_Suc_4n_diff_2 [rule_format, simp]:
nipkow@25134
   405
  "n\<noteq>0 --> Suc (Suc (4*n - 2)) = 4*n"
paulson@15251
   406
by (induct "n", auto)
paulson@15079
   407
paulson@15079
   408
lemma Suc_mult_two_diff_one [rule_format, simp]:
nipkow@25134
   409
  "n\<noteq>0 --> Suc (2 * n - 1) = 2*n"
paulson@15251
   410
by (induct "n", auto)
paulson@15079
   411
paulson@15234
   412
paulson@15234
   413
text{*It is unclear why so many variant results are needed.*}
paulson@15079
   414
huffman@36974
   415
lemma sin_expansion_lemma:
hoelzl@41412
   416
     "sin (x + real (Suc m) * pi / 2) =
huffman@36974
   417
      cos (x + real (m) * pi / 2)"
huffman@36974
   418
by (simp only: cos_add sin_add real_of_nat_Suc add_divide_distrib left_distrib, auto)
huffman@36974
   419
huffman@45166
   420
lemma sin_coeff_0 [simp]: "sin_coeff 0 = 0"
huffman@45166
   421
  unfolding sin_coeff_def by simp (* TODO: move *)
huffman@45166
   422
paulson@15079
   423
lemma Maclaurin_sin_expansion2:
paulson@15079
   424
     "\<exists>t. abs t \<le> abs x &
paulson@15079
   425
       sin x =
huffman@45166
   426
       (\<Sum>m=0..<n. sin_coeff m * 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))
huffman@36974
   432
apply (simp (no_asm) add: sin_expansion_lemma)
huffman@45166
   433
apply (subst (asm) setsum_0', clarify, case_tac "a", simp, simp)
huffman@45166
   434
apply (cases n, simp, simp)
paulson@15079
   435
apply (rule ccontr, simp)
paulson@15079
   436
apply (drule_tac x = x in spec, simp)
paulson@15079
   437
apply (erule ssubst)
paulson@15079
   438
apply (rule_tac x = t in exI, simp)
nipkow@15536
   439
apply (rule setsum_cong[OF refl])
huffman@45166
   440
apply (auto simp add: sin_coeff_def sin_zero_iff odd_Suc_mult_two_ex)
paulson@15079
   441
done
paulson@15079
   442
paulson@15234
   443
lemma Maclaurin_sin_expansion:
paulson@15234
   444
     "\<exists>t. sin x =
huffman@45166
   445
       (\<Sum>m=0..<n. sin_coeff m * x ^ m)
paulson@15234
   446
      + ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
hoelzl@41412
   447
apply (insert Maclaurin_sin_expansion2 [of x n])
hoelzl@41412
   448
apply (blast intro: elim:)
paulson@15234
   449
done
paulson@15234
   450
paulson@15079
   451
lemma Maclaurin_sin_expansion3:
nipkow@25162
   452
     "[| n > 0; 0 < x |] ==>
paulson@15079
   453
       \<exists>t. 0 < t & t < x &
paulson@15079
   454
       sin x =
huffman@45166
   455
       (\<Sum>m=0..<n. sin_coeff m * x ^ m)
paulson@15079
   456
      + ((sin(t + 1/2 * real(n) *pi) / real (fact n)) * x ^ n)"
paulson@15079
   457
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
   458
apply safe
paulson@15079
   459
apply simp
huffman@36974
   460
apply (simp (no_asm) add: sin_expansion_lemma)
paulson@15079
   461
apply (erule ssubst)
paulson@15079
   462
apply (rule_tac x = t in exI, simp)
nipkow@15536
   463
apply (rule setsum_cong[OF refl])
huffman@45166
   464
apply (auto simp add: sin_coeff_def sin_zero_iff odd_Suc_mult_two_ex)
paulson@15079
   465
done
paulson@15079
   466
paulson@15079
   467
lemma Maclaurin_sin_expansion4:
paulson@15079
   468
     "0 < x ==>
paulson@15079
   469
       \<exists>t. 0 < t & t \<le> x &
paulson@15079
   470
       sin x =
huffman@45166
   471
       (\<Sum>m=0..<n. sin_coeff m * x ^ m)
paulson@15079
   472
      + ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
paulson@15079
   473
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
   474
apply safe
paulson@15079
   475
apply simp
huffman@36974
   476
apply (simp (no_asm) add: sin_expansion_lemma)
paulson@15079
   477
apply (erule ssubst)
paulson@15079
   478
apply (rule_tac x = t in exI, simp)
nipkow@15536
   479
apply (rule setsum_cong[OF refl])
huffman@45166
   480
apply (auto simp add: sin_coeff_def sin_zero_iff odd_Suc_mult_two_ex)
paulson@15079
   481
done
paulson@15079
   482
paulson@15079
   483
paulson@15079
   484
subsection{*Maclaurin Expansion for Cosine Function*}
paulson@15079
   485
huffman@45166
   486
lemma cos_coeff_0 [simp]: "cos_coeff 0 = 1"
huffman@45166
   487
  unfolding cos_coeff_def by simp (* TODO: move *)
huffman@45166
   488
paulson@15079
   489
lemma sumr_cos_zero_one [simp]:
huffman@45166
   490
  "(\<Sum>m=0..<(Suc n). cos_coeff m * 0 ^ m) = 1"
paulson@15251
   491
by (induct "n", auto)
paulson@15079
   492
huffman@36974
   493
lemma cos_expansion_lemma:
huffman@36974
   494
  "cos (x + real(Suc m) * pi / 2) = -sin (x + real m * pi / 2)"
huffman@36974
   495
by (simp only: cos_add sin_add real_of_nat_Suc left_distrib add_divide_distrib, auto)
huffman@36974
   496
paulson@15079
   497
lemma Maclaurin_cos_expansion:
paulson@15079
   498
     "\<exists>t. abs t \<le> abs x &
paulson@15079
   499
       cos x =
huffman@45166
   500
       (\<Sum>m=0..<n. cos_coeff m * x ^ m)
paulson@15079
   501
      + ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
paulson@15079
   502
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
   503
apply safe
paulson@15079
   504
apply (simp (no_asm))
huffman@36974
   505
apply (simp (no_asm) add: cos_expansion_lemma)
paulson@15079
   506
apply (case_tac "n", simp)
nipkow@15561
   507
apply (simp del: setsum_op_ivl_Suc)
paulson@15079
   508
apply (rule ccontr, simp)
paulson@15079
   509
apply (drule_tac x = x in spec, simp)
paulson@15079
   510
apply (erule ssubst)
paulson@15079
   511
apply (rule_tac x = t in exI, simp)
nipkow@15536
   512
apply (rule setsum_cong[OF refl])
huffman@45166
   513
apply (auto simp add: cos_coeff_def cos_zero_iff even_mult_two_ex)
paulson@15079
   514
done
paulson@15079
   515
paulson@15079
   516
lemma Maclaurin_cos_expansion2:
nipkow@25162
   517
     "[| 0 < x; n > 0 |] ==>
paulson@15079
   518
       \<exists>t. 0 < t & t < x &
paulson@15079
   519
       cos x =
huffman@45166
   520
       (\<Sum>m=0..<n. cos_coeff m * x ^ m)
paulson@15079
   521
      + ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
paulson@15079
   522
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
   523
apply safe
paulson@15079
   524
apply simp
huffman@36974
   525
apply (simp (no_asm) add: cos_expansion_lemma)
paulson@15079
   526
apply (erule ssubst)
paulson@15079
   527
apply (rule_tac x = t in exI, simp)
nipkow@15536
   528
apply (rule setsum_cong[OF refl])
huffman@45166
   529
apply (auto simp add: cos_coeff_def cos_zero_iff even_mult_two_ex)
paulson@15079
   530
done
paulson@15079
   531
paulson@15234
   532
lemma Maclaurin_minus_cos_expansion:
nipkow@25162
   533
     "[| x < 0; n > 0 |] ==>
paulson@15079
   534
       \<exists>t. x < t & t < 0 &
paulson@15079
   535
       cos x =
huffman@45166
   536
       (\<Sum>m=0..<n. cos_coeff m * x ^ m)
paulson@15079
   537
      + ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
paulson@15079
   538
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
   539
apply safe
paulson@15079
   540
apply simp
huffman@36974
   541
apply (simp (no_asm) add: cos_expansion_lemma)
paulson@15079
   542
apply (erule ssubst)
paulson@15079
   543
apply (rule_tac x = t in exI, simp)
nipkow@15536
   544
apply (rule setsum_cong[OF refl])
huffman@45166
   545
apply (auto simp add: cos_coeff_def cos_zero_iff even_mult_two_ex)
paulson@15079
   546
done
paulson@15079
   547
paulson@15079
   548
(* ------------------------------------------------------------------------- *)
paulson@15079
   549
(* Version for ln(1 +/- x). Where is it??                                    *)
paulson@15079
   550
(* ------------------------------------------------------------------------- *)
paulson@15079
   551
paulson@15079
   552
lemma sin_bound_lemma:
paulson@15081
   553
    "[|x = y; abs u \<le> (v::real) |] ==> \<bar>(x + u) - y\<bar> \<le> v"
paulson@15079
   554
by auto
paulson@15079
   555
paulson@15079
   556
lemma Maclaurin_sin_bound:
huffman@45166
   557
  "abs(sin x - (\<Sum>m=0..<n. sin_coeff m * x ^ m))
huffman@45166
   558
  \<le> inverse(real (fact n)) * \<bar>x\<bar> ^ n"
obua@14738
   559
proof -
paulson@15079
   560
  have "!! x (y::real). x \<le> 1 \<Longrightarrow> 0 \<le> y \<Longrightarrow> x * y \<le> 1 * y"
obua@14738
   561
    by (rule_tac mult_right_mono,simp_all)
obua@14738
   562
  note est = this[simplified]
huffman@22985
   563
  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)"
huffman@22985
   564
  have diff_0: "?diff 0 = sin" by simp
huffman@22985
   565
  have DERIV_diff: "\<forall>m x. DERIV (?diff m) x :> ?diff (Suc m) x"
huffman@22985
   566
    apply (clarify)
huffman@22985
   567
    apply (subst (1 2 3) mod_Suc_eq_Suc_mod)
huffman@22985
   568
    apply (cut_tac m=m in mod_exhaust_less_4)
hoelzl@31880
   569
    apply (safe, auto intro!: DERIV_intros)
huffman@22985
   570
    done
huffman@22985
   571
  from Maclaurin_all_le [OF diff_0 DERIV_diff]
huffman@22985
   572
  obtain t where t1: "\<bar>t\<bar> \<le> \<bar>x\<bar>" and
huffman@22985
   573
    t2: "sin x = (\<Sum>m = 0..<n. ?diff m 0 / real (fact m) * x ^ m) +
huffman@22985
   574
      ?diff n t / real (fact n) * x ^ n" by fast
huffman@22985
   575
  have diff_m_0:
huffman@22985
   576
    "\<And>m. ?diff m 0 = (if even m then 0
huffman@23177
   577
         else -1 ^ ((m - Suc 0) div 2))"
huffman@22985
   578
    apply (subst even_even_mod_4_iff)
huffman@22985
   579
    apply (cut_tac m=m in mod_exhaust_less_4)
huffman@22985
   580
    apply (elim disjE, simp_all)
huffman@22985
   581
    apply (safe dest!: mod_eqD, simp_all)
huffman@22985
   582
    done
obua@14738
   583
  show ?thesis
huffman@45166
   584
    unfolding sin_coeff_def
huffman@22985
   585
    apply (subst t2)
paulson@15079
   586
    apply (rule sin_bound_lemma)
nipkow@15536
   587
    apply (rule setsum_cong[OF refl])
huffman@22985
   588
    apply (subst diff_m_0, simp)
paulson@15079
   589
    apply (auto intro: mult_right_mono [where b=1, simplified] mult_right_mono
hoelzl@41412
   590
                simp add: est mult_nonneg_nonneg mult_ac divide_inverse
paulson@16924
   591
                          power_abs [symmetric] abs_mult)
obua@14738
   592
    done
obua@14738
   593
qed
obua@14738
   594
paulson@15079
   595
end