src/HOL/MacLaurin.thy
author hoelzl
Wed, 16 Mar 2016 11:49:56 +0100
changeset 62624 59ceeb6f3079
parent 61954 1d43f86f48be
child 63040 eb4ddd18d635
permissions -rw-r--r--
generalized some Borel measurable statements to support ennreal
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
28952
15a4b2cf8c34 made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents: 27239
diff changeset
     1
(*  Author      : Jacques D. Fleuriot
12224
02df7cbe7d25 even more theories from Jacques
paulson
parents:
diff changeset
     2
    Copyright   : 2001 University of Edinburgh
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
     3
    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
     4
    Conversion of Mac Laurin to Isar by 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
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
     7
section\<open>MacLaurin 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
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
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
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
    15
text\<open>This is a very long, messy proof even now that it's been broken down
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
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:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    19
    "0 < h ==>
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    20
     \<exists>B::real. f h = (\<Sum>m<n. (j m / (fact m)) * (h^m)) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    21
               (B * ((h^n) /(fact n)))"
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    22
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
    23
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    24
lemma eq_diff_eq': "(x = y - z) = (y = x + (z::real))"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    25
by arith
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    26
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
    27
lemma fact_diff_Suc:
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
    28
  "n < Suc m \<Longrightarrow> fact (Suc m - n) = (Suc m - n) * fact (m - n)"
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    29
  by (subst fact_reduce, auto)
32038
4127b89f48ab Repaired uses of factorial.
avigad
parents: 31882
diff changeset
    30
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    31
lemma Maclaurin_lemma2:
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
    32
  fixes B
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
    33
  assumes DERIV : "\<forall>m t. m < n \<and> 0\<le>t \<and> t\<le>h \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    34
      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
    35
  defines "difg \<equiv>
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
    36
      (\<lambda>m t::real. diff m t -
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
    37
         ((\<Sum>p<n - m. diff (m + p) 0 / (fact p) * t ^ p) + B * (t ^ (n - m) / (fact (n - m)))))"
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    38
        (is "difg \<equiv> (\<lambda>m t. diff m t - ?difg m t)")
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
    39
  shows "\<forall>m t. m < n & 0 \<le> t & t \<le> h --> DERIV (difg m) t :> difg (Suc m) t"
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
    40
proof (rule allI impI)+
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    41
  fix m and t::real
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    42
  assume INIT2: "m < n & 0 \<le> t & t \<le> h"
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
    43
  have "DERIV (difg m) t :> diff (Suc m) t -
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    44
    ((\<Sum>x<n - m. real x * t ^ (x - Suc 0) * diff (m + x) 0 / (fact x)) +
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
    45
     real (n - m) * t ^ (n - Suc m) * B / (fact (n - m)))"
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    46
    unfolding difg_def
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
    47
    by (auto 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
    48
  moreover
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
    49
  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
    50
    unfolding atLeast0LessThan[symmetric] by auto
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    51
  have "(\<Sum>x<n - m. real x * t ^ (x - Suc 0) * diff (m + x) 0 / (fact x)) =
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    52
      (\<Sum>x<n - Suc m. real (Suc x) * t ^ x * diff (Suc m + x) 0 / (fact (Suc x)))"
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
    53
    unfolding intvl atLeast0LessThan by (subst setsum.insert) (auto simp: setsum.reindex)
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
    54
  moreover
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    55
  have fact_neq_0: "\<And>x. (fact x) + real x * (fact x) \<noteq> 0"
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
    56
    by (metis add_pos_pos fact_gt_zero less_add_same_cancel1 less_add_same_cancel2 less_numeral_extra(3) mult_less_0_iff of_nat_less_0_iff)
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    57
  have "\<And>x. (Suc x) * t ^ x * diff (Suc m + x) 0 / (fact (Suc x)) =
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    58
            diff (Suc m + x) 0 * t^x / (fact x)"
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    59
    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
    60
  moreover
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    61
  have "(n - m) * t ^ (n - Suc m) * B / (fact (n - m)) =
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    62
        B * (t ^ (n - Suc m) / (fact (n - Suc m)))"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
    63
    using \<open>0 < n - m\<close>
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    64
    by (simp add: divide_simps fact_reduce)
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
    65
  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
    66
    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
    67
qed
32038
4127b89f48ab Repaired uses of factorial.
avigad
parents: 31882
diff changeset
    68
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    69
lemma Maclaurin:
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    70
  assumes h: "0 < h"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    71
  assumes n: "0 < n"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    72
  assumes diff_0: "diff 0 = f"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    73
  assumes diff_Suc:
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    74
    "\<forall>m t. m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    75
  shows
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    76
    "\<exists>t::real. 0 < t & t < h &
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    77
              f h =
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    78
              setsum (%m. (diff m 0 / (fact m)) * h ^ m) {..<n} +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    79
              (diff n t / (fact n)) * h ^ n"
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    80
proof -
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    81
  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
    82
    by (cases n) (simp add: n)
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    83
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    84
  obtain B where f_h: "f h =
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60758
diff changeset
    85
        (\<Sum>m<n. diff m (0::real) / (fact m) * h ^ m) + B * (h ^ n / (fact n))"
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    86
    using Maclaurin_lemma [OF h] ..
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    87
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
    88
  def g \<equiv> "(\<lambda>t. f t -
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    89
    (setsum (\<lambda>m. (diff m 0 / (fact m)) * t^m) {..<n} + (B * (t^n / (fact n)))))"
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    90
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    91
  have g2: "g 0 = 0 & g h = 0"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56536
diff changeset
    92
    by (simp add: m f_h g_def lessThan_Suc_eq_insert_0 image_iff diff_0 setsum.reindex)
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    93
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
    94
  def difg \<equiv> "(%m t. diff m t -
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    95
    (setsum (%p. (diff (m + p) 0 / (fact p)) * (t ^ p)) {..<n-m}
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    96
      + (B * ((t ^ (n - m)) / (fact (n - m))))))"
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    97
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    98
  have difg_0: "difg 0 = g"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
    99
    unfolding difg_def g_def by (simp add: diff_0)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   100
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60758
diff changeset
   101
  have difg_Suc: "\<forall>(m::nat) t::real.
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60758
diff changeset
   102
        m < n \<and> (0::real) \<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
   103
    using diff_Suc m unfolding difg_def by (rule Maclaurin_lemma2)
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   104
56193
c726ecfb22b6 cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents: 56181
diff changeset
   105
  have difg_eq_0: "\<forall>m<n. difg m 0 = 0"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56536
diff changeset
   106
    by (auto simp: difg_def m Suc_diff_le lessThan_Suc_eq_insert_0 image_iff setsum.reindex)
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   107
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   108
  have isCont_difg: "\<And>m x. \<lbrakk>m < n; 0 \<le> x; x \<le> h\<rbrakk> \<Longrightarrow> isCont (difg m) x"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   109
    by (rule DERIV_isCont [OF difg_Suc [rule_format]]) simp
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   110
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   111
  have differentiable_difg:
56181
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 51489
diff changeset
   112
    "\<And>m x. \<lbrakk>m < n; 0 \<le> x; x \<le> h\<rbrakk> \<Longrightarrow> difg m differentiable (at x)"
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   113
    by (rule differentiableI [OF difg_Suc [rule_format]]) simp
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   114
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   115
  have difg_Suc_eq_0: "\<And>m t. \<lbrakk>m < n; 0 \<le> t; t \<le> h; DERIV (difg m) t :> 0\<rbrakk>
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   116
        \<Longrightarrow> difg (Suc m) t = 0"
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   117
    by (rule DERIV_unique [OF difg_Suc [rule_format]]) simp
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   118
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   119
  have "m < n" using m by simp
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   120
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   121
  have "\<exists>t. 0 < t \<and> t < h \<and> DERIV (difg m) t :> 0"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   122
  using \<open>m < n\<close>
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   123
  proof (induct m)
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   124
    case 0
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   125
    show ?case
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   126
    proof (rule Rolle)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   127
      show "0 < h" by fact
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   128
      show "difg 0 0 = difg 0 h" by (simp add: difg_0 g2)
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60758
diff changeset
   129
      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
   130
        by (simp add: isCont_difg n)
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60758
diff changeset
   131
      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
   132
        by (simp add: differentiable_difg n)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   133
    qed
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   134
  next
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   135
    case (Suc m')
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   136
    hence "\<exists>t. 0 < t \<and> t < h \<and> DERIV (difg m') t :> 0" by simp
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   137
    then obtain t where t: "0 < t" "t < h" "DERIV (difg m') t :> 0" by fast
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   138
    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
   139
    proof (rule Rolle)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   140
      show "0 < t" by fact
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   141
      show "difg (Suc m') 0 = difg (Suc m') t"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   142
        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
   143
      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
   144
        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
   145
      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
   146
        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
   147
    qed
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   148
    thus ?case
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   149
      using \<open>t < h\<close> by auto
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   150
  qed
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   151
  then obtain t where "0 < t" "t < h" "DERIV (difg m) t :> 0" by fast
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   152
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   153
  hence "difg (Suc m) t = 0"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   154
    using \<open>m < n\<close> by (simp add: difg_Suc_eq_0)
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   155
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   156
  show ?thesis
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   157
  proof (intro exI conjI)
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   158
    show "0 < t" by fact
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   159
    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
   160
    show "f h = (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) + diff n t / (fact n) * h ^ n"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   161
      using \<open>difg (Suc m) t = 0\<close>
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   162
      by (simp add: m f_h difg_def)
29187
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   163
  qed
7b09385234f9 clean up proofs of lemma Maclaurin
huffman
parents: 29168
diff changeset
   164
qed
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   165
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   166
lemma Maclaurin_objl:
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   167
  "0 < h & n>0 & diff 0 = f &
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   168
  (\<forall>m t. m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t)
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   169
   --> (\<exists>t::real. 0 < t & t < h &
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   170
            f h = (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   171
                  diff n t / (fact n) * h ^ n)"
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   172
by (blast intro: Maclaurin)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   173
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:
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   176
  assumes INIT1: "0 < h " and INIT2: "diff 0 = f"
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   177
  and DERIV: "\<forall>m t::real.
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   178
  m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t"
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   179
  shows "\<exists>t. 0 < t \<and> t \<le> h \<and> f h =
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   180
  (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   181
  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
   182
proof (cases "n")
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44319
diff changeset
   183
  case 0 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
   184
next
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   185
  case Suc
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   186
  hence "n > 0" by simp
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   187
  from INIT1 this INIT2 DERIV have "\<exists>t>0. t < h \<and>
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   188
    f h =
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   189
    (\<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
   190
    by (rule Maclaurin)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44319
diff changeset
   191
  thus ?thesis by fastforce
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   192
qed
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   193
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   194
lemma Maclaurin2_objl:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   195
     "0 < h & diff 0 = f &
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   196
       (\<forall>m t. m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t)
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   197
    --> (\<exists>t::real. 0 < t &
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   198
              t \<le> h &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   199
              f h =
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   200
              (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   201
              diff n t / (fact n) * h ^ n)"
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   202
by (blast intro: Maclaurin2)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   203
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   204
lemma Maclaurin_minus:
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   205
  fixes h::real
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   206
  assumes "h < 0" "0 < n" "diff 0 = f"
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   207
  and DERIV: "\<forall>m t. m < n & h \<le> t & t \<le> 0 --> DERIV (diff m) t :> diff (Suc m) t"
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   208
  shows "\<exists>t. h < t & t < 0 &
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   209
         f h = (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   210
         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
   211
proof -
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61609
diff changeset
   212
  txt "Transform \<open>ABL'\<close> into \<open>derivative_intros\<close> format."
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   213
  note DERIV' = DERIV_chain'[OF _ DERIV[rule_format], THEN DERIV_cong]
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   214
  from assms
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   215
  have "\<exists>t>0. t < - h \<and>
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   216
    f (- (- h)) =
56193
c726ecfb22b6 cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents: 56181
diff changeset
   217
    (\<Sum>m<n.
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   218
    (- 1) ^ m * diff m (- 0) / (fact m) * (- h) ^ m) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   219
    (- 1) ^ n * diff n (- t) / (fact n) * (- h) ^ n"
56381
0556204bc230 merged DERIV_intros, has_derivative_intros into derivative_intros
hoelzl
parents: 56238
diff changeset
   220
    by (intro Maclaurin) (auto intro!: derivative_eq_intros DERIV')
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   221
  then guess t ..
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   222
  moreover
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   223
  have "(- 1) ^ n * diff n (- t) * (- h) ^ n / (fact n) = diff n (- t) * h ^ n / (fact n)"
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   224
    by (auto simp add: power_mult_distrib[symmetric])
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   225
  moreover
61954
1d43f86f48be more symbols;
wenzelm
parents: 61944
diff changeset
   226
  have "(\<Sum>m<n. (- 1) ^ m * diff m 0 * (- h) ^ m / (fact m)) = (\<Sum>m<n. diff m 0 * h ^ m / (fact m))"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56536
diff changeset
   227
    by (auto intro: setsum.cong simp add: 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
   228
  ultimately have " h < - t \<and>
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   229
    - t < 0 \<and>
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   230
    f h =
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   231
    (\<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
   232
    by auto
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   233
  thus ?thesis ..
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   234
qed
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   235
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   236
lemma Maclaurin_minus_objl:
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   237
  fixes h::real
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   238
  shows
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   239
     "(h < 0 & n > 0 & diff 0 = f &
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   240
       (\<forall>m t.
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   241
          m < n & h \<le> t & t \<le> 0 --> DERIV (diff m) t :> diff (Suc m) t))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   242
    --> (\<exists>t. h < t &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   243
              t < 0 &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   244
              f h =
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   245
              (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   246
              diff n t / (fact n) * h ^ n)"
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   247
by (blast intro: Maclaurin_minus)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   248
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   249
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   250
subsection\<open>More Convenient "Bidirectional" Version.\<close>
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   251
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   252
(* not good for PVS sin_approx, cos_approx *)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   253
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   254
lemma Maclaurin_bi_le_lemma:
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   255
  "n>0 \<Longrightarrow>
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   256
   diff 0 0 =
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   257
   (\<Sum>m<n. diff m 0 * 0 ^ m / (fact m)) + diff n 0 * 0 ^ n / (fact n :: real)"
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   258
by (induct "n") auto
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   259
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   260
lemma Maclaurin_bi_le:
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   261
   assumes "diff 0 = f"
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61799
diff changeset
   262
   and DERIV : "\<forall>m t::real. m < n & \<bar>t\<bar> \<le> \<bar>x\<bar> --> DERIV (diff m) t :> diff (Suc m) t"
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61799
diff changeset
   263
   shows "\<exists>t. \<bar>t\<bar> \<le> \<bar>x\<bar> &
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   264
              f x =
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   265
              (\<Sum>m<n. diff m 0 / (fact m) * x ^ m) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   266
     diff n t / (fact n) * x ^ n" (is "\<exists>t. _ \<and> f x = ?f x t")
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   267
proof cases
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   268
  assume "n = 0" 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
   269
next
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   270
  assume "n \<noteq> 0"
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   271
  show ?thesis
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   272
  proof (cases rule: linorder_cases)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   273
    assume "x = 0" with \<open>n \<noteq> 0\<close> \<open>diff 0 = f\<close> DERIV
56238
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56193
diff changeset
   274
    have "\<bar>0\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x 0" by (auto simp add: Maclaurin_bi_le_lemma)
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   275
    thus ?thesis ..
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   276
  next
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   277
    assume "x < 0"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   278
    with \<open>n \<noteq> 0\<close> DERIV
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   279
    have "\<exists>t>x. t < 0 \<and> diff 0 x = ?f x t" by (intro Maclaurin_minus) auto
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   280
    then guess t ..
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   281
    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
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   282
    thus ?thesis ..
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   283
  next
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   284
    assume "x > 0"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   285
    with \<open>n \<noteq> 0\<close> \<open>diff 0 = f\<close> DERIV
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   286
    have "\<exists>t>0. t < x \<and> diff 0 x = ?f x t" by (intro Maclaurin) auto
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   287
    then guess t ..
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   288
    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
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   289
    thus ?thesis ..
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   290
  qed
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   291
qed
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   292
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   293
lemma Maclaurin_all_lt:
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   294
  fixes x::real
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   295
  assumes INIT1: "diff 0 = f" and INIT2: "0 < n" and INIT3: "x \<noteq> 0"
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   296
  and DERIV: "\<forall>m x. DERIV (diff m) x :> diff(Suc m) x"
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61799
diff changeset
   297
  shows "\<exists>t. 0 < \<bar>t\<bar> & \<bar>t\<bar> < \<bar>x\<bar> & f x =
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   298
    (\<Sum>m<n. (diff m 0 / (fact m)) * x ^ m) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   299
                (diff n t / (fact n)) * x ^ n" (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
   300
proof (cases rule: linorder_cases)
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   301
  assume "x = 0" with INIT3 show "?thesis"..
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   302
next
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   303
  assume "x < 0"
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   304
  with assms have "\<exists>t>x. t < 0 \<and> f x = ?f x t" by (intro Maclaurin_minus) auto
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   305
  then guess t ..
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   306
  with \<open>x < 0\<close> have "0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t" by simp
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   307
  thus ?thesis ..
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   308
next
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   309
  assume "x > 0"
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   310
  with assms have "\<exists>t>0. t < x \<and> f x = ?f x t " by (intro Maclaurin) auto
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   311
  then guess t ..
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   312
  with \<open>x > 0\<close> have "0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t" by simp
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   313
  thus ?thesis ..
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   314
qed
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   315
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   316
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   317
lemma Maclaurin_all_lt_objl:
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   318
  fixes x::real
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   319
  shows
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   320
     "diff 0 = f &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   321
      (\<forall>m x. DERIV (diff m) x :> diff(Suc m) x) &
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   322
      x ~= 0 & n > 0
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61799
diff changeset
   323
      --> (\<exists>t. 0 < \<bar>t\<bar> & \<bar>t\<bar> < \<bar>x\<bar> &
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   324
               f x = (\<Sum>m<n. (diff m 0 / (fact m)) * x ^ m) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   325
                     (diff n t / (fact n)) * x ^ n)"
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   326
by (blast intro: Maclaurin_all_lt)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   327
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   328
lemma Maclaurin_zero [rule_format]:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   329
     "x = (0::real)
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   330
      ==> n \<noteq> 0 -->
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   331
          (\<Sum>m<n. (diff m (0::real) / (fact m)) * x ^ m) =
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   332
          diff 0 0"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   333
by (induct n, auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   334
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   335
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   336
lemma Maclaurin_all_le:
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   337
  assumes INIT: "diff 0 = f"
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   338
  and DERIV: "\<forall>m x::real. DERIV (diff m) x :> diff (Suc m) x"
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61799
diff changeset
   339
  shows "\<exists>t. \<bar>t\<bar> \<le> \<bar>x\<bar> & f x =
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   340
    (\<Sum>m<n. (diff m 0 / (fact m)) * x ^ m) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   341
    (diff n t / (fact n)) * x ^ n" (is "\<exists>t. _ \<and> f x = ?f x t")
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   342
proof cases
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   343
  assume "n = 0" with INIT 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
   344
  next
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   345
  assume "n \<noteq> 0"
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   346
  show ?thesis
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   347
  proof cases
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   348
    assume "x = 0"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   349
    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
   350
      by (intro Maclaurin_zero) auto
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   351
    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" by force
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   352
    thus ?thesis ..
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   353
  next
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   354
    assume "x \<noteq> 0"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   355
    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
   356
      by (intro Maclaurin_all_lt) auto
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   357
    then guess t ..
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   358
    hence "\<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x t" by simp
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   359
    thus ?thesis ..
41120
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   360
  qed
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   361
qed
74e41b2d48ea adding an Isar version of the MacLaurin theorem from some students' work in 2005
bulwahn
parents: 36974
diff changeset
   362
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   363
lemma Maclaurin_all_le_objl:
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   364
  "diff 0 = f &
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   365
      (\<forall>m x. DERIV (diff m) x :> diff (Suc m) x)
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61799
diff changeset
   366
      --> (\<exists>t::real. \<bar>t\<bar> \<le> \<bar>x\<bar> &
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   367
              f x = (\<Sum>m<n. (diff m 0 / (fact m)) * x ^ m) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   368
                    (diff n t / (fact n)) * x ^ n)"
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   369
by (blast intro: Maclaurin_all_le)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   370
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   371
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   372
subsection\<open>Version for Exponential Function\<close>
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   373
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   374
lemma Maclaurin_exp_lt:
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   375
  fixes x::real
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   376
  shows
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   377
  "[| x ~= 0; n > 0 |]
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61799
diff changeset
   378
      ==> (\<exists>t. 0 < \<bar>t\<bar> &
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61799
diff changeset
   379
                \<bar>t\<bar> < \<bar>x\<bar> &
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   380
                exp x = (\<Sum>m<n. (x ^ m) / (fact m)) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   381
                        (exp t / (fact n)) * x ^ n)"
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   382
by (cut_tac diff = "%n. exp" and f = exp and x = x and n = n in Maclaurin_all_lt_objl, auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   383
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   384
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   385
lemma Maclaurin_exp_le:
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61799
diff changeset
   386
     "\<exists>t::real. \<bar>t\<bar> \<le> \<bar>x\<bar> &
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   387
            exp x = (\<Sum>m<n. (x ^ m) / (fact m)) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   388
                       (exp t / (fact n)) * x ^ n"
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   389
by (cut_tac diff = "%n. exp" and f = exp and x = x and n = n in Maclaurin_all_le_objl, auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   390
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
   391
lemma exp_lower_taylor_quadratic:
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
   392
  fixes x::real
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
   393
  shows "0 \<le> x \<Longrightarrow> 1 + x + x\<^sup>2 / 2 \<le> exp x"
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
   394
  using Maclaurin_exp_le [of x 3]
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
   395
  by (auto simp: numeral_3_eq_3 power2_eq_square power_Suc)
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
   396
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   397
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   398
subsection\<open>Version for Sine Function\<close>
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   399
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   400
lemma mod_exhaust_less_4:
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   401
  "m mod 4 = 0 | m mod 4 = 1 | m mod 4 = 2 | m mod 4 = (3::nat)"
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19765
diff changeset
   402
by auto
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   403
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   404
lemma Suc_Suc_mult_two_diff_two [rule_format, simp]:
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   405
  "n\<noteq>0 --> Suc (Suc (2 * n - 2)) = 2*n"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15234
diff changeset
   406
by (induct "n", auto)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   407
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   408
lemma lemma_Suc_Suc_4n_diff_2 [rule_format, simp]:
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   409
  "n\<noteq>0 --> Suc (Suc (4*n - 2)) = 4*n"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15234
diff changeset
   410
by (induct "n", auto)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   411
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   412
lemma Suc_mult_two_diff_one [rule_format, simp]:
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   413
  "n\<noteq>0 --> Suc (2 * n - 1) = 2*n"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15234
diff changeset
   414
by (induct "n", auto)
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
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   417
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
   418
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
   419
lemma sin_expansion_lemma:
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   420
     "sin (x + real (Suc m) * pi / 2) =
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
   421
      cos (x + real (m) * pi / 2)"
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
   422
by (simp only: cos_add sin_add of_nat_Suc add_divide_distrib distrib_right, auto)
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
   423
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   424
lemma Maclaurin_sin_expansion2:
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61799
diff changeset
   425
     "\<exists>t. \<bar>t\<bar> \<le> \<bar>x\<bar> &
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   426
       sin x =
56193
c726ecfb22b6 cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents: 56181
diff changeset
   427
       (\<Sum>m<n. sin_coeff m * x ^ m)
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   428
      + ((sin(t + 1/2 * real (n) *pi) / (fact n)) * x ^ n)"
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   429
apply (cut_tac f = sin and n = n and x = x
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   430
        and diff = "%n x. sin (x + 1/2*real n * pi)" in Maclaurin_all_lt_objl)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   431
apply safe
61284
2314c2f62eb1 real_of_nat_Suc is now a simprule
paulson <lp15@cam.ac.uk>
parents: 61076
diff changeset
   432
    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
   433
   apply (simp 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
   434
   apply (force intro!: derivative_eq_intros)
2314c2f62eb1 real_of_nat_Suc is now a simprule
paulson <lp15@cam.ac.uk>
parents: 61076
diff changeset
   435
  apply (subst (asm) setsum.neutral, auto)[1]
2314c2f62eb1 real_of_nat_Suc is now a simprule
paulson <lp15@cam.ac.uk>
parents: 61076
diff changeset
   436
 apply (rule ccontr, simp)
2314c2f62eb1 real_of_nat_Suc is now a simprule
paulson <lp15@cam.ac.uk>
parents: 61076
diff changeset
   437
 apply (drule_tac x = x in spec, simp)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   438
apply (erule ssubst)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   439
apply (rule_tac x = t in exI, simp)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56536
diff changeset
   440
apply (rule setsum.cong[OF refl])
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
   441
apply (auto simp add: sin_coeff_def sin_zero_iff elim: oddE simp del: of_nat_Suc)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   442
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   443
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   444
lemma Maclaurin_sin_expansion:
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   445
     "\<exists>t. sin x =
56193
c726ecfb22b6 cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents: 56181
diff changeset
   446
       (\<Sum>m<n. sin_coeff m * x ^ m)
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   447
      + ((sin(t + 1/2 * real (n) *pi) / (fact n)) * x ^ n)"
41166
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   448
apply (insert Maclaurin_sin_expansion2 [of x n])
4b2a457b17e8 beautify MacLaurin proofs; make better use of DERIV_intros
hoelzl
parents: 41120
diff changeset
   449
apply (blast intro: elim:)
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   450
done
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   451
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   452
lemma Maclaurin_sin_expansion3:
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   453
     "[| n > 0; 0 < x |] ==>
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   454
       \<exists>t. 0 < t & t < x &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   455
       sin x =
56193
c726ecfb22b6 cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents: 56181
diff changeset
   456
       (\<Sum>m<n. sin_coeff m * x ^ m)
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   457
      + ((sin(t + 1/2 * real(n) *pi) / (fact n)) * x ^ n)"
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   458
apply (cut_tac f = sin and n = n and h = x and diff = "%n x. sin (x + 1/2*real (n) *pi)" in Maclaurin_objl)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   459
apply safe
61284
2314c2f62eb1 real_of_nat_Suc is now a simprule
paulson <lp15@cam.ac.uk>
parents: 61076
diff changeset
   460
    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
   461
   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
   462
   apply (force intro!: derivative_eq_intros)
2314c2f62eb1 real_of_nat_Suc is now a simprule
paulson <lp15@cam.ac.uk>
parents: 61076
diff changeset
   463
  apply (erule ssubst)
2314c2f62eb1 real_of_nat_Suc is now a simprule
paulson <lp15@cam.ac.uk>
parents: 61076
diff changeset
   464
  apply (rule_tac x = t in exI, simp)
2314c2f62eb1 real_of_nat_Suc is now a simprule
paulson <lp15@cam.ac.uk>
parents: 61076
diff changeset
   465
 apply (rule setsum.cong[OF refl])
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
   466
 apply (auto simp add: sin_coeff_def sin_zero_iff elim: oddE simp del: of_nat_Suc)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   467
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   468
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   469
lemma Maclaurin_sin_expansion4:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   470
     "0 < x ==>
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   471
       \<exists>t. 0 < t & t \<le> x &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   472
       sin x =
56193
c726ecfb22b6 cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents: 56181
diff changeset
   473
       (\<Sum>m<n. sin_coeff m * x ^ m)
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   474
      + ((sin(t + 1/2 * real (n) *pi) / (fact n)) * x ^ n)"
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   475
apply (cut_tac f = sin and n = n and h = x and diff = "%n x. sin (x + 1/2*real (n) *pi)" in Maclaurin2_objl)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   476
apply safe
61284
2314c2f62eb1 real_of_nat_Suc is now a simprule
paulson <lp15@cam.ac.uk>
parents: 61076
diff changeset
   477
    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
   478
   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
   479
   apply (force intro!: derivative_eq_intros)
2314c2f62eb1 real_of_nat_Suc is now a simprule
paulson <lp15@cam.ac.uk>
parents: 61076
diff changeset
   480
  apply (erule ssubst)
2314c2f62eb1 real_of_nat_Suc is now a simprule
paulson <lp15@cam.ac.uk>
parents: 61076
diff changeset
   481
  apply (rule_tac x = t in exI, simp)
2314c2f62eb1 real_of_nat_Suc is now a simprule
paulson <lp15@cam.ac.uk>
parents: 61076
diff changeset
   482
 apply (rule setsum.cong[OF refl])
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
   483
 apply (auto simp add: sin_coeff_def sin_zero_iff elim: oddE simp del: of_nat_Suc)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   484
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   485
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   486
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   487
subsection\<open>Maclaurin Expansion for Cosine Function\<close>
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   488
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   489
lemma sumr_cos_zero_one [simp]:
56193
c726ecfb22b6 cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents: 56181
diff changeset
   490
  "(\<Sum>m<(Suc n). cos_coeff m * 0 ^ m) = 1"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15234
diff changeset
   491
by (induct "n", auto)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   492
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
   493
lemma cos_expansion_lemma:
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
   494
  "cos (x + real(Suc m) * pi / 2) = -sin (x + real m * pi / 2)"
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
   495
by (simp only: cos_add sin_add of_nat_Suc distrib_right add_divide_distrib, auto)
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
   496
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   497
lemma Maclaurin_cos_expansion:
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61799
diff changeset
   498
     "\<exists>t::real. \<bar>t\<bar> \<le> \<bar>x\<bar> &
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   499
       cos x =
56193
c726ecfb22b6 cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents: 56181
diff changeset
   500
       (\<Sum>m<n. cos_coeff m * x ^ m)
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   501
      + ((cos(t + 1/2 * real (n) *pi) / (fact n)) * x ^ n)"
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   502
apply (cut_tac f = cos and n = n and x = x and diff = "%n x. cos (x + 1/2*real (n) *pi)" in Maclaurin_all_lt_objl)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   503
apply safe
61284
2314c2f62eb1 real_of_nat_Suc is now a simprule
paulson <lp15@cam.ac.uk>
parents: 61076
diff changeset
   504
    apply (simp (no_asm))
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
   505
   apply (simp (no_asm) add: cos_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
   506
  apply (case_tac "n", simp)
2314c2f62eb1 real_of_nat_Suc is now a simprule
paulson <lp15@cam.ac.uk>
parents: 61076
diff changeset
   507
  apply (simp del: setsum_lessThan_Suc)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   508
apply (rule ccontr, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   509
apply (drule_tac x = x in spec, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   510
apply (erule ssubst)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   511
apply (rule_tac x = t in exI, simp)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56536
diff changeset
   512
apply (rule setsum.cong[OF refl])
58709
efdc6c533bd3 prefer generic elimination rules for even/odd over specialized unfold rules for nat
haftmann
parents: 58410
diff changeset
   513
apply (auto simp add: cos_coeff_def cos_zero_iff elim: evenE)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   514
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   515
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   516
lemma Maclaurin_cos_expansion2:
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   517
     "[| 0 < x; n > 0 |] ==>
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   518
       \<exists>t. 0 < t & t < x &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   519
       cos x =
56193
c726ecfb22b6 cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents: 56181
diff changeset
   520
       (\<Sum>m<n. cos_coeff m * x ^ m)
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   521
      + ((cos(t + 1/2 * real (n) *pi) / (fact n)) * x ^ n)"
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   522
apply (cut_tac f = cos and n = n and h = x and diff = "%n x. cos (x + 1/2*real (n) *pi)" in Maclaurin_objl)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   523
apply safe
61284
2314c2f62eb1 real_of_nat_Suc is now a simprule
paulson <lp15@cam.ac.uk>
parents: 61076
diff changeset
   524
  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
   525
  apply (simp (no_asm) add: cos_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
   526
 apply (erule ssubst)
2314c2f62eb1 real_of_nat_Suc is now a simprule
paulson <lp15@cam.ac.uk>
parents: 61076
diff changeset
   527
 apply (rule_tac x = t in exI, simp)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56536
diff changeset
   528
apply (rule setsum.cong[OF refl])
58709
efdc6c533bd3 prefer generic elimination rules for even/odd over specialized unfold rules for nat
haftmann
parents: 58410
diff changeset
   529
apply (auto simp add: cos_coeff_def cos_zero_iff elim: evenE)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   530
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   531
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   532
lemma Maclaurin_minus_cos_expansion:
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   533
     "[| x < 0; n > 0 |] ==>
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   534
       \<exists>t. x < t & t < 0 &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   535
       cos x =
56193
c726ecfb22b6 cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents: 56181
diff changeset
   536
       (\<Sum>m<n. cos_coeff m * x ^ m)
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   537
      + ((cos(t + 1/2 * real (n) *pi) / (fact n)) * x ^ n)"
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   538
apply (cut_tac f = cos and n = n and h = x and diff = "%n x. cos (x + 1/2*real (n) *pi)" in Maclaurin_minus_objl)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   539
apply safe
61284
2314c2f62eb1 real_of_nat_Suc is now a simprule
paulson <lp15@cam.ac.uk>
parents: 61076
diff changeset
   540
  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
   541
 apply (simp (no_asm) add: cos_expansion_lemma del: of_nat_Suc)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   542
apply (erule ssubst)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   543
apply (rule_tac x = t in exI, simp)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56536
diff changeset
   544
apply (rule setsum.cong[OF refl])
58709
efdc6c533bd3 prefer generic elimination rules for even/odd over specialized unfold rules for nat
haftmann
parents: 58410
diff changeset
   545
apply (auto simp add: cos_coeff_def cos_zero_iff elim: evenE)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   546
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   547
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   548
(* ------------------------------------------------------------------------- *)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   549
(* Version for ln(1 +/- x). Where is it??                                    *)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   550
(* ------------------------------------------------------------------------- *)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   551
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   552
lemma sin_bound_lemma:
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61799
diff changeset
   553
    "[|x = y; \<bar>u\<bar> \<le> (v::real) |] ==> \<bar>(x + u) - y\<bar> \<le> v"
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   554
by auto
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   555
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   556
lemma Maclaurin_sin_bound:
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61799
diff changeset
   557
  "\<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
   558
proof -
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   559
  have "!! x (y::real). x \<le> 1 \<Longrightarrow> 0 \<le> y \<Longrightarrow> x * y \<le> 1 * y"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   560
    by (rule_tac mult_right_mono,simp_all)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   561
  note est = this[simplified]
22985
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   562
  let ?diff = "\<lambda>(n::nat) x. if n mod 4 = 0 then sin(x) else if n mod 4 = 1 then cos(x) else if n mod 4 = 2 then -sin(x) else -cos(x)"
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   563
  have diff_0: "?diff 0 = sin" by simp
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   564
  have DERIV_diff: "\<forall>m x. DERIV (?diff m) x :> ?diff (Suc m) x"
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   565
    apply (clarify)
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   566
    apply (subst (1 2 3) mod_Suc_eq_Suc_mod)
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   567
    apply (cut_tac m=m in mod_exhaust_less_4)
56381
0556204bc230 merged DERIV_intros, has_derivative_intros into derivative_intros
hoelzl
parents: 56238
diff changeset
   568
    apply (safe, auto intro!: derivative_eq_intros)
22985
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   569
    done
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   570
  from Maclaurin_all_le [OF diff_0 DERIV_diff]
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   571
  obtain t where t1: "\<bar>t\<bar> \<le> \<bar>x\<bar>" and
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   572
    t2: "sin x = (\<Sum>m<n. ?diff m 0 / (fact m) * x ^ m) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   573
      ?diff n t / (fact n) * x ^ n" by fast
22985
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   574
  have diff_m_0:
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   575
    "\<And>m. ?diff m 0 = (if even m then 0
58410
6d46ad54a2ab explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents: 57514
diff changeset
   576
         else (- 1) ^ ((m - Suc 0) div 2))"
22985
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   577
    apply (subst even_even_mod_4_iff)
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   578
    apply (cut_tac m=m in mod_exhaust_less_4)
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   579
    apply (elim disjE, simp_all)
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   580
    apply (safe dest!: mod_eqD, simp_all)
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   581
    done
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   582
  show ?thesis
44306
33572a766836 fold definitions of sin_coeff and cos_coeff in Maclaurin lemmas
huffman
parents: 41166
diff changeset
   583
    unfolding sin_coeff_def
22985
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   584
    apply (subst t2)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   585
    apply (rule sin_bound_lemma)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56536
diff changeset
   586
    apply (rule setsum.cong[OF refl])
22985
501e6dfe4e5a cleaned up proof of Maclaurin_sin_bound
huffman
parents: 22983
diff changeset
   587
    apply (subst diff_m_0, simp)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   588
    apply (auto intro: mult_right_mono [where b=1, simplified] mult_right_mono
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57492
diff changeset
   589
                simp add: est 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
   590
    done
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   591
qed
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   592
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   593
end