src/HOL/MacLaurin.thy
author wenzelm
Mon, 22 May 2017 00:23:25 +0200
changeset 65897 94b0da1b242e
parent 65273 917ae0ba03a2
child 67091 1393c2340eec
permissions -rw-r--r--
back to scala-2.11.8 due to apparent non-termination of HOL-Codegenerator_Test;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
     1
(*  Title:      HOL/MacLaurin.thy
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
     2
    Author:     Jacques D. Fleuriot, 2001 University of Edinburgh
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
     3
    Author:     Lawrence C Paulson, 2004
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
     4
    Author:     Lukas Bulwahn and Bernhard Häupler, 2005
12224
02df7cbe7d25 even more theories from Jacques
paulson
parents:
diff changeset
     5
*)
02df7cbe7d25 even more theories from Jacques
paulson
parents:
diff changeset
     6
63570
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
     7
section \<open>MacLaurin and Taylor Series\<close>
15944
9b00875e21f7 from simplesubst to new subst
paulson
parents: 15561
diff changeset
     8
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15081
diff changeset
     9
theory MacLaurin
29811
026b0f9f579f fixed Proofs and dependencies ; Theory Dense_Linear_Order moved to Library
chaieb@chaieb-laptop
parents: 29803
diff changeset
    10
imports Transcendental
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15081
diff changeset
    11
begin
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    12
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    13
subsection \<open>Maclaurin's Theorem with Lagrange Form of Remainder\<close>
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    14
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    15
text \<open>This is a very long, messy proof even now that it's been broken down
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    16
  into lemmas.\<close>
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    17
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    18
lemma Maclaurin_lemma:
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    19
  "0 < h \<Longrightarrow>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    20
    \<exists>B::real. f h = (\<Sum>m<n. (j m / (fact m)) * (h^m)) + (B * ((h^n) /(fact n)))"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    21
  by (rule exI[where x = "(f h - (\<Sum>m<n. (j m / (fact m)) * h^m)) * (fact n) / (h^n)"]) simp
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    22
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    23
lemma eq_diff_eq': "x = y - z \<longleftrightarrow> y = x + z"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    24
  for x y z :: real
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    25
  by arith
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    26
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    27
lemma fact_diff_Suc: "n < Suc m \<Longrightarrow> fact (Suc m - n) = (Suc m - n) * fact (m - n)"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    28
  by (subst fact_reduce) auto
32038
4127b89f48ab Repaired uses of factorial.
avigad
parents: 31882
diff changeset
    29
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    30
lemma Maclaurin_lemma2:
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
    31
  fixes B
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    32
  assumes DERIV: "\<forall>m t. m < n \<and> 0\<le>t \<and> t\<le>h \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    33
    and INIT: "n = Suc k"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61284
diff changeset
    34
  defines "difg \<equiv>
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    35
    (\<lambda>m t::real. diff m t -
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    36
      ((\<Sum>p<n - m. diff (m + p) 0 / fact p * t ^ p) + B * (t ^ (n - m) / fact (n - m))))"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    37
    (is "difg \<equiv> (\<lambda>m t. diff m t - ?difg m t)")
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    38
  shows "\<forall>m t. m < n \<and> 0 \<le> t \<and> t \<le> h \<longrightarrow> DERIV (difg m) t :> difg (Suc m) t"
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
    39
proof (rule allI impI)+
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    40
  fix m t
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    41
  assume INIT2: "m < n \<and> 0 \<le> t \<and> t \<le> h"
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
    42
  have "DERIV (difg m) t :> diff (Suc m) t -
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    43
    ((\<Sum>x<n - m. real x * t ^ (x - Suc 0) * diff (m + x) 0 / fact x) +
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    44
     real (n - m) * t ^ (n - Suc m) * B / fact (n - m))"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    45
    by (auto simp: difg_def intro!: derivative_eq_intros DERIV[rule_format, OF INIT2])
56193
c726ecfb22b6 cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents: 56181
diff changeset
    46
  moreover
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
    47
  from INIT2 have intvl: "{..<n - m} = insert 0 (Suc ` {..<n - Suc m})" and "0 < n - m"
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
    48
    unfolding atLeast0LessThan[symmetric] by auto
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    49
  have "(\<Sum>x<n - m. real x * t ^ (x - Suc 0) * diff (m + x) 0 / fact x) =
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    50
      (\<Sum>x<n - Suc m. real (Suc x) * t ^ x * diff (Suc m + x) 0 / fact (Suc x))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63570
diff changeset
    51
    unfolding intvl by (subst sum.insert) (auto simp add: sum.reindex)
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
    52
  moreover
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    53
  have fact_neq_0: "\<And>x. (fact x) + real x * (fact x) \<noteq> 0"
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    54
    by (metis add_pos_pos fact_gt_zero less_add_same_cancel1 less_add_same_cancel2
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    55
        less_numeral_extra(3) mult_less_0_iff of_nat_less_0_iff)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    56
  have "\<And>x. (Suc x) * t ^ x * diff (Suc m + x) 0 / fact (Suc x) = diff (Suc m + x) 0 * t^x / fact x"
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    57
    by (rule nonzero_divide_eq_eq[THEN iffD2]) auto
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
    58
  moreover
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    59
  have "(n - m) * t ^ (n - Suc m) * B / fact (n - m) = B * (t ^ (n - Suc m) / fact (n - Suc m))"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    60
    using \<open>0 < n - m\<close> by (simp add: divide_simps fact_reduce)
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
    61
  ultimately show "DERIV (difg m) t :> difg (Suc m) t"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61284
diff changeset
    62
    unfolding difg_def  by (simp add: mult.commute)
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
    63
qed
32038
4127b89f48ab Repaired uses of factorial.
avigad
parents: 31882
diff changeset
    64
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    65
lemma Maclaurin:
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    66
  assumes h: "0 < h"
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    67
    and n: "0 < n"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    68
    and diff_0: "diff 0 = f"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    69
    and diff_Suc: "\<forall>m t. m < n \<and> 0 \<le> t \<and> t \<le> h \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    70
  shows
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    71
    "\<exists>t::real. 0 < t \<and> t < h \<and>
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63570
diff changeset
    72
      f h = sum (\<lambda>m. (diff m 0 / fact m) * h ^ m) {..<n} + (diff n t / fact n) * h ^ n"
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    73
proof -
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    74
  from n obtain m where m: "n = Suc m"
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
    75
    by (cases n) (simp add: n)
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    76
  from m have "m < n" by simp
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    77
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    78
  obtain B where f_h: "f h = (\<Sum>m<n. diff m 0 / fact m * h ^ m) + B * (h ^ n / fact n)"
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    79
    using Maclaurin_lemma [OF h] ..
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    80
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 61954
diff changeset
    81
  define g where [abs_def]: "g t =
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63570
diff changeset
    82
    f t - (sum (\<lambda>m. (diff m 0 / fact m) * t^m) {..<n} + B * (t^n / fact n))" for t
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    83
  have g2: "g 0 = 0" "g h = 0"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63570
diff changeset
    84
    by (simp_all add: m f_h g_def lessThan_Suc_eq_insert_0 image_iff diff_0 sum.reindex)
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    85
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 61954
diff changeset
    86
  define difg where [abs_def]: "difg m t =
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63570
diff changeset
    87
    diff m t - (sum (\<lambda>p. (diff (m + p) 0 / fact p) * (t ^ p)) {..<n-m} +
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    88
      B * ((t ^ (n - m)) / fact (n - m)))" for m t
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    89
  have difg_0: "difg 0 = g"
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    90
    by (simp add: difg_def g_def diff_0)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    91
  have difg_Suc: "\<forall>m t. m < n \<and> 0 \<le> t \<and> t \<le> h \<longrightarrow> DERIV (difg m) t :> difg (Suc m) t"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 61954
diff changeset
    92
    using diff_Suc m unfolding difg_def [abs_def] by (rule Maclaurin_lemma2)
56193
c726ecfb22b6 cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents: 56181
diff changeset
    93
  have difg_eq_0: "\<forall>m<n. difg m 0 = 0"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63570
diff changeset
    94
    by (auto simp: difg_def m Suc_diff_le lessThan_Suc_eq_insert_0 image_iff sum.reindex)
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    95
  have isCont_difg: "\<And>m x. m < n \<Longrightarrow> 0 \<le> x \<Longrightarrow> x \<le> h \<Longrightarrow> isCont (difg m) x"
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    96
    by (rule DERIV_isCont [OF difg_Suc [rule_format]]) simp
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    97
  have differentiable_difg: "\<And>m x. m < n \<Longrightarrow> 0 \<le> x \<Longrightarrow> x \<le> h \<Longrightarrow> difg m differentiable (at x)"
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    98
    by (rule differentiableI [OF difg_Suc [rule_format]]) simp
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
    99
  have difg_Suc_eq_0:
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   100
    "\<And>m t. m < n \<Longrightarrow> 0 \<le> t \<Longrightarrow> t \<le> h \<Longrightarrow> DERIV (difg m) t :> 0 \<Longrightarrow> difg (Suc m) t = 0"
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   101
    by (rule DERIV_unique [OF difg_Suc [rule_format]]) simp
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   102
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   103
  have "\<exists>t. 0 < t \<and> t < h \<and> DERIV (difg m) t :> 0"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   104
  using \<open>m < n\<close>
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   105
  proof (induct m)
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   106
    case 0
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   107
    show ?case
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   108
    proof (rule Rolle)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   109
      show "0 < h" by fact
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   110
      show "difg 0 0 = difg 0 h"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   111
        by (simp add: difg_0 g2)
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60758
diff changeset
   112
      show "\<forall>x. 0 \<le> x \<and> x \<le> h \<longrightarrow> isCont (difg (0::nat)) x"
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   113
        by (simp add: isCont_difg n)
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60758
diff changeset
   114
      show "\<forall>x. 0 < x \<and> x < h \<longrightarrow> difg (0::nat) differentiable (at x)"
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   115
        by (simp add: differentiable_difg n)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   116
    qed
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   117
  next
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   118
    case (Suc m')
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   119
    then have "\<exists>t. 0 < t \<and> t < h \<and> DERIV (difg m') t :> 0"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   120
      by simp
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   121
    then obtain t where t: "0 < t" "t < h" "DERIV (difg m') t :> 0"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   122
      by fast
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   123
    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
   124
    proof (rule Rolle)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   125
      show "0 < t" by fact
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   126
      show "difg (Suc m') 0 = difg (Suc m') t"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   127
        using t \<open>Suc m' < n\<close> by (simp add: difg_Suc_eq_0 difg_eq_0)
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   128
      show "\<forall>x. 0 \<le> x \<and> x \<le> t \<longrightarrow> isCont (difg (Suc m')) x"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   129
        using \<open>t < h\<close> \<open>Suc m' < n\<close> by (simp add: isCont_difg)
56181
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 51489
diff changeset
   130
      show "\<forall>x. 0 < x \<and> x < t \<longrightarrow> difg (Suc m') differentiable (at x)"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   131
        using \<open>t < h\<close> \<open>Suc m' < n\<close> by (simp add: differentiable_difg)
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   132
    qed
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   133
    with \<open>t < h\<close> show ?case
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   134
      by auto
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   135
  qed
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   136
  then obtain t where "0 < t" "t < h" "DERIV (difg m) t :> 0"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   137
    by fast
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   138
  with \<open>m < n\<close> have "difg (Suc m) t = 0"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   139
    by (simp add: difg_Suc_eq_0)
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   140
  show ?thesis
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   141
  proof (intro exI conjI)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   142
    show "0 < t" by fact
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   143
    show "t < h" by fact
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   144
    show "f h = (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) + diff n t / (fact n) * h ^ n"
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   145
      using \<open>difg (Suc m) t = 0\<close> by (simp add: m f_h difg_def)
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   146
  qed
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   147
qed
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   148
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   149
lemma Maclaurin_objl:
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   150
  "0 < h \<and> n > 0 \<and> diff 0 = f \<and>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   151
    (\<forall>m t. m < n \<and> 0 \<le> t \<and> t \<le> h \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t) \<longrightarrow>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   152
    (\<exists>t. 0 < t \<and> t < h \<and> f h = (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) + diff n t / fact n * h ^ n)"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   153
  for n :: nat and h :: real
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   154
  by (blast intro: Maclaurin)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   155
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   156
lemma Maclaurin2:
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   157
  fixes n :: nat
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   158
    and h :: real
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   159
  assumes INIT1: "0 < h"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   160
    and INIT2: "diff 0 = f"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   161
    and DERIV: "\<forall>m t. m < n \<and> 0 \<le> t \<and> t \<le> h \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   162
  shows "\<exists>t. 0 < t \<and> t \<le> h \<and> f h = (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) + diff n t / fact n * h ^ n"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   163
proof (cases n)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   164
  case 0
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   165
  with INIT1 INIT2 show ?thesis by fastforce
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   166
next
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   167
  case Suc
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   168
  then have "n > 0" by simp
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   169
  from INIT1 this INIT2 DERIV
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   170
  have "\<exists>t>0. t < h \<and> f h = (\<Sum>m<n. diff m 0 / fact m * h ^ m) + diff n t / fact n * h ^ n"
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   171
    by (rule Maclaurin)
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   172
  then show ?thesis by fastforce
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   173
qed
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   174
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   175
lemma Maclaurin2_objl:
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   176
  "0 < h \<and> diff 0 = f \<and>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   177
    (\<forall>m t. m < n \<and> 0 \<le> t \<and> t \<le> h \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t) \<longrightarrow>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   178
    (\<exists>t. 0 < t \<and> t \<le> h \<and> f h = (\<Sum>m<n. diff m 0 / fact m * h ^ m) + diff n t / fact n * h ^ n)"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   179
  for n :: nat and h :: real
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   180
  by (blast intro: Maclaurin2)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   181
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   182
lemma Maclaurin_minus:
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   183
  fixes n :: nat and h :: real
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   184
  assumes "h < 0" "0 < n" "diff 0 = f"
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   185
    and DERIV: "\<forall>m t. m < n \<and> h \<le> t \<and> t \<le> 0 \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   186
  shows "\<exists>t. h < t \<and> t < 0 \<and> f h = (\<Sum>m<n. diff m 0 / fact m * h ^ m) + diff n t / fact n * h ^ n"
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   187
proof -
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   188
  txt \<open>Transform \<open>ABL'\<close> into \<open>derivative_intros\<close> format.\<close>
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   189
  note DERIV' = DERIV_chain'[OF _ DERIV[rule_format], THEN DERIV_cong]
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   190
  let ?sum = "\<lambda>t.
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   191
    (\<Sum>m<n. (- 1) ^ m * diff m (- 0) / (fact m) * (- h) ^ m) +
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   192
    (- 1) ^ n * diff n (- t) / (fact n) * (- h) ^ n"
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   193
  from assms have "\<exists>t>0. t < - h \<and> f (- (- h)) = ?sum t"
56381
0556204bc230 merged DERIV_intros, has_derivative_intros into derivative_intros
hoelzl
parents: 56238
diff changeset
   194
    by (intro Maclaurin) (auto intro!: derivative_eq_intros DERIV')
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   195
  then obtain t where "0 < t" "t < - h" "f (- (- h)) = ?sum t"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   196
    by blast
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   197
  moreover have "(- 1) ^ n * diff n (- t) * (- h) ^ n / fact n = diff n (- t) * h ^ n / fact n"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   198
    by (auto simp: power_mult_distrib[symmetric])
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   199
  moreover
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   200
    have "(\<Sum>m<n. (- 1) ^ m * diff m 0 * (- h) ^ m / fact m) = (\<Sum>m<n. diff m 0 * h ^ m / fact m)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63570
diff changeset
   201
    by (auto intro: sum.cong simp add: power_mult_distrib[symmetric])
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   202
  ultimately have "h < - t \<and> - t < 0 \<and>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   203
    f h = (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) + diff n (- t) / (fact n) * h ^ n"
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   204
    by auto
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   205
  then show ?thesis ..
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   206
qed
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   207
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   208
lemma Maclaurin_minus_objl:
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   209
  fixes n :: nat and h :: real
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   210
  shows
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   211
    "h < 0 \<and> n > 0 \<and> diff 0 = f \<and>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   212
      (\<forall>m t. m < n & h \<le> t & t \<le> 0 --> DERIV (diff m) t :> diff (Suc m) t) \<longrightarrow>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   213
    (\<exists>t. h < t \<and> t < 0 \<and> f h = (\<Sum>m<n. diff m 0 / fact m * h ^ m) + diff n t / fact n * h ^ n)"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   214
  by (blast intro: Maclaurin_minus)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   215
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   216
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   217
subsection \<open>More Convenient "Bidirectional" Version.\<close>
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   218
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   219
(* not good for PVS sin_approx, cos_approx *)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   220
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   221
lemma Maclaurin_bi_le_lemma:
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   222
  "n > 0 \<Longrightarrow>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   223
    diff 0 0 = (\<Sum>m<n. diff m 0 * 0 ^ m / (fact m)) + diff n 0 * 0 ^ n / (fact n :: real)"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   224
  by (induct n) auto
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   225
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   226
lemma Maclaurin_bi_le:
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   227
  fixes n :: nat and x :: real
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   228
  assumes "diff 0 = f"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   229
    and DERIV : "\<forall>m t. m < n \<and> \<bar>t\<bar> \<le> \<bar>x\<bar> \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   230
  shows "\<exists>t. \<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = (\<Sum>m<n. diff m 0 / (fact m) * x ^ m) + diff n t / (fact n) * x ^ n"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   231
    (is "\<exists>t. _ \<and> f x = ?f x t")
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   232
proof (cases "n = 0")
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   233
  case True
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   234
  with \<open>diff 0 = f\<close> show ?thesis by force
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   235
next
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   236
  case False
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   237
  show ?thesis
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   238
  proof (cases rule: linorder_cases)
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   239
    assume "x = 0"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   240
    with \<open>n \<noteq> 0\<close> \<open>diff 0 = f\<close> DERIV have "\<bar>0\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x 0"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   241
      by (auto simp add: Maclaurin_bi_le_lemma)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   242
    then show ?thesis ..
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   243
  next
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   244
    assume "x < 0"
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   245
    with \<open>n \<noteq> 0\<close> DERIV have "\<exists>t>x. t < 0 \<and> diff 0 x = ?f x t"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   246
      by (intro Maclaurin_minus) auto
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   247
    then obtain t where "x < t" "t < 0"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   248
      "diff 0 x = (\<Sum>m<n. diff m 0 / fact m * x ^ m) + diff n t / fact n * x ^ n"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   249
      by blast
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   250
    with \<open>x < 0\<close> \<open>diff 0 = f\<close> have "\<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x t"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   251
      by simp
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   252
    then show ?thesis ..
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   253
  next
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   254
    assume "x > 0"
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   255
    with \<open>n \<noteq> 0\<close> \<open>diff 0 = f\<close> DERIV have "\<exists>t>0. t < x \<and> diff 0 x = ?f x t"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   256
      by (intro Maclaurin) auto
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   257
    then obtain t where "0 < t" "t < x"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   258
      "diff 0 x = (\<Sum>m<n. diff m 0 / fact m * x ^ m) + diff n t / fact n * x ^ n"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   259
      by blast
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   260
    with \<open>x > 0\<close> \<open>diff 0 = f\<close> have "\<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x t" by simp
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   261
    then show ?thesis ..
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   262
  qed
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   263
qed
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   264
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   265
lemma Maclaurin_all_lt:
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   266
  fixes x :: real
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   267
  assumes INIT1: "diff 0 = f"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   268
    and INIT2: "0 < n"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   269
    and INIT3: "x \<noteq> 0"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   270
    and DERIV: "\<forall>m x. DERIV (diff m) x :> diff(Suc m) x"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   271
  shows "\<exists>t. 0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x =
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   272
      (\<Sum>m<n. (diff m 0 / fact m) * x ^ m) + (diff n t / fact n) * x ^ n"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   273
    (is "\<exists>t. _ \<and> _ \<and> f x = ?f x t")
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   274
proof (cases rule: linorder_cases)
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   275
  assume "x = 0"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   276
  with INIT3 show ?thesis ..
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   277
next
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   278
  assume "x < 0"
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   279
  with assms have "\<exists>t>x. t < 0 \<and> f x = ?f x t"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   280
    by (intro Maclaurin_minus) auto
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   281
  then obtain t where "t > x" "t < 0" "f x = ?f x t"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   282
    by blast
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   283
  with \<open>x < 0\<close> have "0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   284
    by simp
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   285
  then show ?thesis ..
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   286
next
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   287
  assume "x > 0"
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   288
  with assms have "\<exists>t>0. t < x \<and> f x = ?f x t"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   289
    by (intro Maclaurin) auto
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   290
  then obtain t where "t > 0" "t < x" "f x = ?f x t"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   291
    by blast
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   292
  with \<open>x > 0\<close> have "0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   293
    by simp
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   294
  then show ?thesis ..
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   295
qed
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   296
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   297
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   298
lemma Maclaurin_all_lt_objl:
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   299
  fixes x :: real
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   300
  shows
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   301
    "diff 0 = f \<and> (\<forall>m x. DERIV (diff m) x :> diff (Suc m) x) \<and> x \<noteq> 0 \<and> n > 0 \<longrightarrow>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   302
    (\<exists>t. 0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   303
      f x = (\<Sum>m<n. (diff m 0 / fact m) * x ^ m) + (diff n t / fact n) * x ^ n)"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   304
  by (blast intro: Maclaurin_all_lt)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   305
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   306
lemma Maclaurin_zero: "x = 0 \<Longrightarrow> n \<noteq> 0 \<Longrightarrow> (\<Sum>m<n. (diff m 0 / fact m) * x ^ m) = diff 0 0"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   307
  for x :: real and n :: nat
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   308
  by (induct n) auto
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   309
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   310
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   311
lemma Maclaurin_all_le:
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   312
  fixes x :: real and n :: nat
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   313
  assumes INIT: "diff 0 = f"
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   314
    and DERIV: "\<forall>m x. DERIV (diff m) x :> diff (Suc m) x"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   315
  shows "\<exists>t. \<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = (\<Sum>m<n. (diff m 0 / fact m) * x ^ m) + (diff n t / fact n) * x ^ n"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   316
    (is "\<exists>t. _ \<and> f x = ?f x t")
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   317
proof (cases "n = 0")
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   318
  case True
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   319
  with INIT show ?thesis by force
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   320
next
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   321
  case False
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   322
  show ?thesis
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   323
  proof (cases "x = 0")
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   324
    case True
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   325
    with \<open>n \<noteq> 0\<close> have "(\<Sum>m<n. diff m 0 / (fact m) * x ^ m) = diff 0 0"
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   326
      by (intro Maclaurin_zero) auto
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   327
    with INIT \<open>x = 0\<close> \<open>n \<noteq> 0\<close> have " \<bar>0\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x 0"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   328
      by force
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   329
    then show ?thesis ..
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   330
  next
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   331
    case False
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   332
    with INIT \<open>n \<noteq> 0\<close> DERIV have "\<exists>t. 0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t"
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   333
      by (intro Maclaurin_all_lt) auto
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   334
    then obtain t where "0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t" ..
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   335
    then have "\<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x t"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   336
      by simp
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   337
    then show ?thesis ..
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   338
  qed
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   339
qed
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   340
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   341
lemma Maclaurin_all_le_objl:
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   342
  "diff 0 = f \<and> (\<forall>m x. DERIV (diff m) x :> diff (Suc m) x) \<longrightarrow>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   343
    (\<exists>t::real. \<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = (\<Sum>m<n. (diff m 0 / fact m) * x ^ m) + (diff n t / fact n) * x ^ n)"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   344
  for x :: real and n :: nat
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   345
  by (blast intro: Maclaurin_all_le)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   346
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   347
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   348
subsection \<open>Version for Exponential Function\<close>
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   349
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   350
lemma Maclaurin_exp_lt:
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   351
  fixes x :: real and n :: nat
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   352
  shows
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   353
    "x \<noteq> 0 \<Longrightarrow> n > 0 \<Longrightarrow>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   354
      (\<exists>t. 0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> exp x = (\<Sum>m<n. (x ^ m) / fact m) + (exp t / fact n) * x ^ n)"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   355
 using Maclaurin_all_lt_objl [where diff = "\<lambda>n. exp" and f = exp and x = x and n = n] by auto
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   356
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   357
lemma Maclaurin_exp_le:
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   358
  fixes x :: real and n :: nat
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   359
  shows "\<exists>t. \<bar>t\<bar> \<le> \<bar>x\<bar> \<and> exp x = (\<Sum>m<n. (x ^ m) / fact m) + (exp t / fact n) * x ^ n"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   360
  using Maclaurin_all_le_objl [where diff = "\<lambda>n. exp" and f = exp and x = x and n = n] by auto
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   361
65273
917ae0ba03a2 Removal of [simp] status for greaterThan_0. Moved two theorems into main HOL.
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   362
corollary exp_lower_taylor_quadratic: "0 \<le> x \<Longrightarrow> 1 + x + x\<^sup>2 / 2 \<le> exp x"
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   363
  for x :: real
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   364
  using Maclaurin_exp_le [of x 3] by (auto simp: numeral_3_eq_3 power2_eq_square)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   365
65273
917ae0ba03a2 Removal of [simp] status for greaterThan_0. Moved two theorems into main HOL.
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   366
corollary ln_2_less_1: "ln 2 < (1::real)"
917ae0ba03a2 Removal of [simp] status for greaterThan_0. Moved two theorems into main HOL.
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   367
proof -
917ae0ba03a2 Removal of [simp] status for greaterThan_0. Moved two theorems into main HOL.
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   368
  have "2 < 5/(2::real)" by simp
917ae0ba03a2 Removal of [simp] status for greaterThan_0. Moved two theorems into main HOL.
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   369
  also have "5/2 \<le> exp (1::real)" using exp_lower_taylor_quadratic[of 1, simplified] by simp
917ae0ba03a2 Removal of [simp] status for greaterThan_0. Moved two theorems into main HOL.
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   370
  finally have "exp (ln 2) < exp (1::real)" by simp
917ae0ba03a2 Removal of [simp] status for greaterThan_0. Moved two theorems into main HOL.
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   371
  thus "ln 2 < (1::real)" by (subst (asm) exp_less_cancel_iff) simp
917ae0ba03a2 Removal of [simp] status for greaterThan_0. Moved two theorems into main HOL.
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   372
qed
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   373
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   374
subsection \<open>Version for Sine Function\<close>
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   375
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   376
lemma mod_exhaust_less_4: "m mod 4 = 0 | m mod 4 = 1 | m mod 4 = 2 | m mod 4 = 3"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   377
  for m :: nat
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   378
  by auto
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   379
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   380
lemma Suc_Suc_mult_two_diff_two [simp]: "n \<noteq> 0 \<Longrightarrow> Suc (Suc (2 * n - 2)) = 2 * n"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   381
  by (induct n) auto
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   382
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   383
lemma lemma_Suc_Suc_4n_diff_2 [simp]: "n \<noteq> 0 \<Longrightarrow> Suc (Suc (4 * n - 2)) = 4 * n"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   384
  by (induct n) auto
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   385
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   386
lemma Suc_mult_two_diff_one [simp]: "n \<noteq> 0 \<Longrightarrow> Suc (2 * n - 1) = 2 * n"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   387
  by (induct n) auto
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59730
diff changeset
   388
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   389
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   390
text \<open>It is unclear why so many variant results are needed.\<close>
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   391
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   392
lemma sin_expansion_lemma: "sin (x + real (Suc m) * pi / 2) = cos (x + real m * pi / 2)"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   393
  by (auto simp: cos_add sin_add add_divide_distrib distrib_right)
36974
b877866b5b00 remove some unnamed simp rules from Transcendental.thy; move the needed ones to MacLaurin.thy where they are used
huffman
parents: 32047
diff changeset
   394
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   395
lemma Maclaurin_sin_expansion2:
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   396
  "\<exists>t. \<bar>t\<bar> \<le> \<bar>x\<bar> \<and>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   397
    sin x = (\<Sum>m<n. sin_coeff m * x ^ m) + (sin (t + 1/2 * real n * pi) / fact n) * x ^ n"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   398
  using Maclaurin_all_lt_objl
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   399
    [where f = sin and n = n and x = x and diff = "\<lambda>n x. sin (x + 1/2 * real n * pi)"]
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   400
  apply safe
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   401
      apply simp
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   402
     apply (simp add: sin_expansion_lemma del: of_nat_Suc)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   403
     apply (force intro!: derivative_eq_intros)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63570
diff changeset
   404
    apply (subst (asm) sum.neutral; auto)
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   405
   apply (rule ccontr)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   406
   apply simp
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   407
   apply (drule_tac x = x in spec)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   408
   apply simp
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   409
  apply (erule ssubst)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   410
  apply (rule_tac x = t in exI)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   411
  apply simp
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63570
diff changeset
   412
  apply (rule sum.cong[OF refl])
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   413
  apply (auto simp add: sin_coeff_def sin_zero_iff elim: oddE simp del: of_nat_Suc)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   414
  done
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   415
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   416
lemma Maclaurin_sin_expansion:
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   417
  "\<exists>t. sin x = (\<Sum>m<n. sin_coeff m * x ^ m) + (sin (t + 1/2 * real n * pi) / fact n) * x ^ n"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   418
  using Maclaurin_sin_expansion2 [of x n] by blast
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   419
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   420
lemma Maclaurin_sin_expansion3:
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   421
  "n > 0 \<Longrightarrow> 0 < x \<Longrightarrow>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   422
    \<exists>t. 0 < t \<and> t < x \<and>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   423
       sin x = (\<Sum>m<n. sin_coeff m * x ^ m) + (sin (t + 1/2 * real n * pi) / fact n) * x ^ n"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   424
  using Maclaurin_objl
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   425
    [where f = sin and n = n and h = x and diff = "\<lambda>n x. sin (x + 1/2 * real n * pi)"]
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   426
  apply safe
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   427
    apply simp
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   428
   apply (simp (no_asm) add: sin_expansion_lemma del: of_nat_Suc)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   429
   apply (force intro!: derivative_eq_intros)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   430
  apply (erule ssubst)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   431
  apply (rule_tac x = t in exI)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   432
  apply simp
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63570
diff changeset
   433
  apply (rule sum.cong[OF refl])
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   434
  apply (auto simp add: sin_coeff_def sin_zero_iff elim: oddE simp del: of_nat_Suc)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   435
  done
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   436
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   437
lemma Maclaurin_sin_expansion4:
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   438
  "0 < x \<Longrightarrow>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   439
    \<exists>t. 0 < t \<and> t \<le> x \<and>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   440
      sin x = (\<Sum>m<n. sin_coeff m * x ^ m) + (sin (t + 1/2 * real n * pi) / fact n) * x ^ n"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   441
  using Maclaurin2_objl
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   442
    [where f = sin and n = n and h = x and diff = "\<lambda>n x. sin (x + 1/2 * real n * pi)"]
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   443
  apply safe
61284
2314c2f62eb1 real_of_nat_Suc is now a simprule
paulson <lp15@cam.ac.uk>
parents: 61076
diff changeset
   444
    apply simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61284
diff changeset
   445
   apply (simp (no_asm) add: sin_expansion_lemma del: of_nat_Suc)
61284
2314c2f62eb1 real_of_nat_Suc is now a simprule
paulson <lp15@cam.ac.uk>
parents: 61076
diff changeset
   446
   apply (force intro!: derivative_eq_intros)
2314c2f62eb1 real_of_nat_Suc is now a simprule
paulson <lp15@cam.ac.uk>
parents: 61076
diff changeset
   447
  apply (erule ssubst)
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   448
  apply (rule_tac x = t in exI)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   449
  apply simp
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63570
diff changeset
   450
  apply (rule sum.cong[OF refl])
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   451
  apply (auto simp add: sin_coeff_def sin_zero_iff elim: oddE simp del: of_nat_Suc)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   452
  done
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   453
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   454
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   455
subsection \<open>Maclaurin Expansion for Cosine Function\<close>
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   456
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   457
lemma sumr_cos_zero_one [simp]: "(\<Sum>m<Suc n. cos_coeff m * 0 ^ m) = 1"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   458
  by (induct n) auto
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   459
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   460
lemma cos_expansion_lemma: "cos (x + real (Suc m) * pi / 2) = - sin (x + real m * pi / 2)"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   461
  by (auto simp: cos_add sin_add distrib_right add_divide_distrib)
36974
b877866b5b00 remove some unnamed simp rules from Transcendental.thy; move the needed ones to MacLaurin.thy where they are used
huffman
parents: 32047
diff changeset
   462
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   463
lemma Maclaurin_cos_expansion:
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   464
  "\<exists>t::real. \<bar>t\<bar> \<le> \<bar>x\<bar> \<and>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   465
    cos x = (\<Sum>m<n. cos_coeff m * x ^ m) + (cos(t + 1/2 * real n * pi) / fact n) * x ^ n"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   466
  using Maclaurin_all_lt_objl
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   467
    [where f = cos and n = n and x = x and diff = "\<lambda>n x. cos (x + 1/2 * real n * pi)"]
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   468
  apply safe
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   469
      apply simp
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   470
     apply (simp add: cos_expansion_lemma del: of_nat_Suc)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   471
    apply (cases n)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   472
     apply simp
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63570
diff changeset
   473
    apply (simp del: sum_lessThan_Suc)
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   474
   apply (rule ccontr)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   475
   apply simp
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   476
   apply (drule_tac x = x in spec)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   477
   apply simp
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   478
  apply (erule ssubst)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   479
  apply (rule_tac x = t in exI)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   480
  apply simp
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63570
diff changeset
   481
  apply (rule sum.cong[OF refl])
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   482
  apply (auto simp add: cos_coeff_def cos_zero_iff elim: evenE)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   483
  done
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   484
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   485
lemma Maclaurin_cos_expansion2:
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   486
  "0 < x \<Longrightarrow> n > 0 \<Longrightarrow>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   487
    \<exists>t. 0 < t \<and> t < x \<and>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   488
      cos x = (\<Sum>m<n. cos_coeff m * x ^ m) + (cos (t + 1/2 * real n * pi) / fact n) * x ^ n"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   489
  using Maclaurin_objl
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   490
    [where f = cos and n = n and h = x and diff = "\<lambda>n x. cos (x + 1/2 * real n * pi)"]
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   491
  apply safe
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   492
    apply simp
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   493
   apply (simp add: cos_expansion_lemma del: of_nat_Suc)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   494
  apply (erule ssubst)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   495
  apply (rule_tac x = t in exI)
61284
2314c2f62eb1 real_of_nat_Suc is now a simprule
paulson <lp15@cam.ac.uk>
parents: 61076
diff changeset
   496
  apply simp
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63570
diff changeset
   497
  apply (rule sum.cong[OF refl])
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   498
  apply (auto simp: cos_coeff_def cos_zero_iff elim: evenE)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   499
  done
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   500
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   501
lemma Maclaurin_minus_cos_expansion:
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   502
  "x < 0 \<Longrightarrow> n > 0 \<Longrightarrow>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   503
    \<exists>t. x < t \<and> t < 0 \<and>
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   504
      cos x = (\<Sum>m<n. cos_coeff m * x ^ m) + ((cos (t + 1/2 * real n * pi) / fact n) * x ^ n)"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   505
  using Maclaurin_minus_objl
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   506
    [where f = cos and n = n and h = x and diff = "\<lambda>n x. cos (x + 1/2 * real n *pi)"]
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   507
  apply safe
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   508
    apply simp
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   509
   apply (simp add: cos_expansion_lemma del: of_nat_Suc)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   510
  apply (erule ssubst)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   511
  apply (rule_tac x = t in exI)
61284
2314c2f62eb1 real_of_nat_Suc is now a simprule
paulson <lp15@cam.ac.uk>
parents: 61076
diff changeset
   512
  apply simp
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63570
diff changeset
   513
  apply (rule sum.cong[OF refl])
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   514
  apply (auto simp: cos_coeff_def cos_zero_iff elim: evenE)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   515
  done
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   516
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   517
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   518
(* Version for ln(1 +/- x). Where is it?? *)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   519
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   520
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   521
lemma sin_bound_lemma: "x = y \<Longrightarrow> \<bar>u\<bar> \<le> v \<Longrightarrow> \<bar>(x + u) - y\<bar> \<le> v"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   522
  for x y u v :: real
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   523
  by auto
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   524
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   525
lemma Maclaurin_sin_bound: "\<bar>sin x - (\<Sum>m<n. sin_coeff m * x ^ m)\<bar> \<le> inverse (fact n) * \<bar>x\<bar> ^ n"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   526
proof -
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   527
  have est: "x \<le> 1 \<Longrightarrow> 0 \<le> y \<Longrightarrow> x * y \<le> 1 * y" for x y :: real
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   528
    by (rule mult_right_mono) simp_all
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   529
  let ?diff = "\<lambda>(n::nat) x.
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   530
    if n mod 4 = 0 then sin x
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   531
    else if n mod 4 = 1 then cos x
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   532
    else if n mod 4 = 2 then - sin x
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   533
    else - cos x"
22985
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   534
  have diff_0: "?diff 0 = sin" by simp
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   535
  have DERIV_diff: "\<forall>m x. DERIV (?diff m) x :> ?diff (Suc m) x"
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   536
    apply clarify
22985
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   537
    apply (subst (1 2 3) mod_Suc_eq_Suc_mod)
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   538
    apply (cut_tac m=m in mod_exhaust_less_4)
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   539
    apply safe
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   540
       apply (auto intro!: derivative_eq_intros)
22985
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   541
    done
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   542
  from Maclaurin_all_le [OF diff_0 DERIV_diff]
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   543
  obtain t where t1: "\<bar>t\<bar> \<le> \<bar>x\<bar>"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   544
    and t2: "sin x = (\<Sum>m<n. ?diff m 0 / (fact m) * x ^ m) + ?diff n t / (fact n) * x ^ n"
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   545
    by fast
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   546
  have diff_m_0: "\<And>m. ?diff m 0 = (if even m then 0 else (- 1) ^ ((m - Suc 0) div 2))"
22985
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   547
    apply (subst even_even_mod_4_iff)
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   548
    apply (cut_tac m=m in mod_exhaust_less_4)
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   549
    apply (elim disjE)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   550
       apply simp_all
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   551
     apply (safe dest!: mod_eqD)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   552
     apply simp_all
22985
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   553
    done
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   554
  show ?thesis
44306
33572a766836 fold definitions of sin_coeff and cos_coeff in Maclaurin lemmas
huffman
parents: 41166
diff changeset
   555
    unfolding sin_coeff_def
22985
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   556
    apply (subst t2)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   557
    apply (rule sin_bound_lemma)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63570
diff changeset
   558
     apply (rule sum.cong[OF refl])
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   559
     apply (subst diff_m_0, simp)
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   560
    using est
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   561
    apply (auto intro: mult_right_mono [where b=1, simplified] mult_right_mono
63569
7e0b0db5e9ac misc tuning and modernization;
wenzelm
parents: 63365
diff changeset
   562
        simp: ac_simps divide_inverse power_abs [symmetric] abs_mult)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   563
    done
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   564
qed
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   565
63570
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   566
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   567
section \<open>Taylor series\<close>
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   568
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   569
text \<open>
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   570
  We use MacLaurin and the translation of the expansion point \<open>c\<close> to \<open>0\<close>
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   571
  to prove Taylor's theorem.
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   572
\<close>
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   573
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   574
lemma taylor_up:
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   575
  assumes INIT: "n > 0" "diff 0 = f"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   576
    and DERIV: "\<forall>m t. m < n \<and> a \<le> t \<and> t \<le> b \<longrightarrow> DERIV (diff m) t :> (diff (Suc m) t)"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   577
    and INTERV: "a \<le> c" "c < b"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   578
  shows "\<exists>t::real. c < t \<and> t < b \<and>
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   579
    f b = (\<Sum>m<n. (diff m c / fact m) * (b - c)^m) + (diff n t / fact n) * (b - c)^n"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   580
proof -
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   581
  from INTERV have "0 < b - c" by arith
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   582
  moreover from INIT have "n > 0" "(\<lambda>m x. diff m (x + c)) 0 = (\<lambda>x. f (x + c))"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   583
    by auto
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   584
  moreover
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   585
  have "\<forall>m t. m < n \<and> 0 \<le> t \<and> t \<le> b - c \<longrightarrow> DERIV (\<lambda>x. diff m (x + c)) t :> diff (Suc m) (t + c)"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   586
  proof (intro strip)
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   587
    fix m t
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   588
    assume "m < n \<and> 0 \<le> t \<and> t \<le> b - c"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   589
    with DERIV and INTERV have "DERIV (diff m) (t + c) :> diff (Suc m) (t + c)"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   590
      by auto
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   591
    moreover from DERIV_ident and DERIV_const have "DERIV (\<lambda>x. x + c) t :> 1 + 0"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   592
      by (rule DERIV_add)
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   593
    ultimately have "DERIV (\<lambda>x. diff m (x + c)) t :> diff (Suc m) (t + c) * (1 + 0)"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   594
      by (rule DERIV_chain2)
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   595
    then show "DERIV (\<lambda>x. diff m (x + c)) t :> diff (Suc m) (t + c)"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   596
      by simp
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   597
  qed
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   598
  ultimately obtain x where
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   599
    "0 < x \<and> x < b - c \<and>
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   600
      f (b - c + c) =
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   601
        (\<Sum>m<n. diff m (0 + c) / fact m * (b - c) ^ m) + diff n (x + c) / fact n * (b - c) ^ n"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   602
     by (rule Maclaurin [THEN exE])
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   603
   then have "c < x + c \<and> x + c < b \<and> f b =
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   604
     (\<Sum>m<n. diff m c / fact m * (b - c) ^ m) + diff n (x + c) / fact n * (b - c) ^ n"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   605
    by fastforce
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   606
  then show ?thesis by fastforce
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   607
qed
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   608
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   609
lemma taylor_down:
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   610
  fixes a :: real and n :: nat
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   611
  assumes INIT: "n > 0" "diff 0 = f"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   612
    and DERIV: "(\<forall>m t. m < n \<and> a \<le> t \<and> t \<le> b \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t)"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   613
    and INTERV: "a < c" "c \<le> b"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   614
  shows "\<exists>t. a < t \<and> t < c \<and>
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   615
    f a = (\<Sum>m<n. (diff m c / fact m) * (a - c)^m) + (diff n t / fact n) * (a - c)^n"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   616
proof -
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   617
  from INTERV have "a-c < 0" by arith
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   618
  moreover from INIT have "n > 0" "(\<lambda>m x. diff m (x + c)) 0 = (\<lambda>x. f (x + c))"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   619
    by auto
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   620
  moreover
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   621
  have "\<forall>m t. m < n \<and> a - c \<le> t \<and> t \<le> 0 \<longrightarrow> DERIV (\<lambda>x. diff m (x + c)) t :> diff (Suc m) (t + c)"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   622
  proof (rule allI impI)+
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   623
    fix m t
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   624
    assume "m < n \<and> a - c \<le> t \<and> t \<le> 0"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   625
    with DERIV and INTERV have "DERIV (diff m) (t + c) :> diff (Suc m) (t + c)"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   626
      by auto
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   627
    moreover from DERIV_ident and DERIV_const have "DERIV (\<lambda>x. x + c) t :> 1 + 0"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   628
      by (rule DERIV_add)
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   629
    ultimately have "DERIV (\<lambda>x. diff m (x + c)) t :> diff (Suc m) (t + c) * (1 + 0)"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   630
      by (rule DERIV_chain2)
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   631
    then show "DERIV (\<lambda>x. diff m (x + c)) t :> diff (Suc m) (t + c)"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   632
      by simp
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   633
  qed
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   634
  ultimately obtain x where
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   635
    "a - c < x \<and> x < 0 \<and>
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   636
      f (a - c + c) =
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   637
        (\<Sum>m<n. diff m (0 + c) / fact m * (a - c) ^ m) + diff n (x + c) / fact n * (a - c) ^ n"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   638
    by (rule Maclaurin_minus [THEN exE])
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   639
  then have "a < x + c \<and> x + c < c \<and>
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   640
    f a = (\<Sum>m<n. diff m c / fact m * (a - c) ^ m) + diff n (x + c) / fact n * (a - c) ^ n"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   641
    by fastforce
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   642
  then show ?thesis by fastforce
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   643
qed
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   644
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   645
theorem taylor:
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   646
  fixes a :: real and n :: nat
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   647
  assumes INIT: "n > 0" "diff 0 = f"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   648
    and DERIV: "\<forall>m t. m < n \<and> a \<le> t \<and> t \<le> b \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   649
    and INTERV: "a \<le> c " "c \<le> b" "a \<le> x" "x \<le> b" "x \<noteq> c"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   650
  shows "\<exists>t.
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   651
    (if x < c then x < t \<and> t < c else c < t \<and> t < x) \<and>
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   652
    f x = (\<Sum>m<n. (diff m c / fact m) * (x - c)^m) + (diff n t / fact n) * (x - c)^n"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   653
proof (cases "x < c")
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   654
  case True
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   655
  note INIT
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   656
  moreover have "\<forall>m t. m < n \<and> x \<le> t \<and> t \<le> b \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   657
    using DERIV and INTERV by fastforce
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   658
  moreover note True
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   659
  moreover from INTERV have "c \<le> b"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   660
    by simp
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   661
  ultimately have "\<exists>t>x. t < c \<and> f x =
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   662
    (\<Sum>m<n. diff m c / (fact m) * (x - c) ^ m) + diff n t / (fact n) * (x - c) ^ n"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   663
    by (rule taylor_down)
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   664
  with True show ?thesis by simp
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   665
next
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   666
  case False
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   667
  note INIT
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   668
  moreover have "\<forall>m t. m < n \<and> a \<le> t \<and> t \<le> x \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   669
    using DERIV and INTERV by fastforce
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   670
  moreover from INTERV have "a \<le> c"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   671
    by arith
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   672
  moreover from False and INTERV have "c < x"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   673
    by arith
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   674
  ultimately have "\<exists>t>c. t < x \<and> f x =
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   675
    (\<Sum>m<n. diff m c / (fact m) * (x - c) ^ m) + diff n t / (fact n) * (x - c) ^ n"
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   676
    by (rule taylor_up)
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   677
  with False show ?thesis by simp
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   678
qed
1826a90b9cbc simplified theory structure;
wenzelm
parents: 63569
diff changeset
   679
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   680
end