src/HOL/Taylor.thy
author wenzelm
Wed, 10 Feb 2016 14:14:43 +0100
changeset 62278 c04e97be39d3
parent 61954 1d43f86f48be
child 63569 7e0b0db5e9ac
permissions -rw-r--r--
misc tuning and updates;
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: 25162
diff changeset
     1
(*  Title:      HOL/Taylor.thy
17634
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
     2
    Author:     Lukas Bulwahn, Bernhard Haeupler, Technische Universitaet Muenchen
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
     3
*)
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
     4
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59730
diff changeset
     5
section \<open>Taylor series\<close>
17634
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
     6
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
     7
theory Taylor
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
     8
imports MacLaurin
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
     9
begin
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    10
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59730
diff changeset
    11
text \<open>
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 60758
diff changeset
    12
We use MacLaurin and the translation of the expansion point \<open>c\<close> to \<open>0\<close>
17634
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    13
to prove Taylor's theorem.
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59730
diff changeset
    14
\<close>
17634
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    15
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    16
lemma taylor_up: 
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
    17
  assumes INIT: "n>0" "diff 0 = f"
17634
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    18
  and DERIV: "(\<forall> m t. m < n & a \<le> t & t \<le> b \<longrightarrow> DERIV (diff m) t :> (diff (Suc m) t))"
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    19
  and INTERV: "a \<le> c" "c < b" 
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    20
  shows "\<exists>t::real. c < t & t < b & 
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    21
    f b = (\<Sum>m<n. (diff m c / (fact m)) * (b - c)^m) + (diff n t / (fact n)) * (b - c)^n"
17634
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    22
proof -
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    23
  from INTERV have "0 < b-c" by arith
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    24
  moreover 
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
    25
  from INIT have "n>0" "((\<lambda>m x. diff m (x + c)) 0) = (\<lambda>x. f (x + c))" by auto
17634
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    26
  moreover
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
    27
  have "ALL m t. m < n & 0 <= t & t <= b - c --> DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c)"
17634
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    28
  proof (intro strip)
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    29
    fix m t
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    30
    assume "m < n & 0 <= t & t <= b - c"
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    31
    with DERIV and INTERV have "DERIV (diff m) (t + c) :> diff (Suc m) (t + c)" by auto
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    32
    moreover
23069
cdfff0241c12 rename lemmas LIM_ident, isCont_ident, DERIV_ident
huffman
parents: 17634
diff changeset
    33
    from DERIV_ident and DERIV_const have "DERIV (%x. x + c) t :> 1+0" by (rule DERIV_add)
17634
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    34
    ultimately have "DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c) * (1+0)"
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    35
      by (rule DERIV_chain2)
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    36
    thus "DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c)" by simp
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    37
  qed
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    38
  ultimately obtain x where 
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    39
        "0 < x & x < b - c & 
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    40
        f (b - c + c) = (\<Sum>m<n. diff m (0 + c) / (fact m) * (b - c) ^ m) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    41
          diff n (x + c) / (fact n) * (b - c) ^ n"
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    42
     by (rule Maclaurin [THEN exE])
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    43
  then have "c<x+c & x+c<b \<and> f b = (\<Sum>m<n. diff m c / (fact m) * (b - c) ^ m) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    44
    diff n (x+c) / (fact n) * (b - c) ^ n"
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    45
    by fastforce
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    46
  thus ?thesis by fastforce
17634
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    47
qed
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    48
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    49
lemma taylor_down:
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    50
  fixes a::real
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
    51
  assumes INIT: "n>0" "diff 0 = f"
17634
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    52
  and DERIV: "(\<forall> m t. m < n & a \<le> t & t \<le> b \<longrightarrow> DERIV (diff m) t :> (diff (Suc m) t))"
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    53
  and INTERV: "a < c" "c \<le> b"
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    54
  shows "\<exists> t. a < t & t < c & 
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    55
    f a = (\<Sum>m<n. (diff m c / (fact m)) * (a - c)^m) + (diff n t / (fact n)) * (a - c)^n" 
17634
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    56
proof -
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    57
  from INTERV have "a-c < 0" by arith
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    58
  moreover 
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
    59
  from INIT have "n>0" "((\<lambda>m x. diff m (x + c)) 0) = (\<lambda>x. f (x + c))" by auto
17634
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    60
  moreover
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    61
  have "ALL m t. m < n & a-c <= t & t <= 0 --> DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c)"
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    62
  proof (rule allI impI)+
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    63
    fix m t
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    64
    assume "m < n & a-c <= t & t <= 0"
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    65
    with DERIV and INTERV have "DERIV (diff m) (t + c) :> diff (Suc m) (t + c)" by auto 
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    66
    moreover
23069
cdfff0241c12 rename lemmas LIM_ident, isCont_ident, DERIV_ident
huffman
parents: 17634
diff changeset
    67
    from DERIV_ident and DERIV_const have "DERIV (%x. x + c) t :> 1+0" by (rule DERIV_add)
17634
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    68
    ultimately have "DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c) * (1+0)" by (rule DERIV_chain2)
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    69
    thus "DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c)" by simp
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    70
  qed
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    71
  ultimately obtain x where 
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    72
         "a - c < x & x < 0 &
61954
1d43f86f48be more symbols;
wenzelm
parents: 61799
diff changeset
    73
      f (a - c + c) = (\<Sum>m<n. diff m (0 + c) / (fact m) * (a - c) ^ m) +
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    74
        diff n (x + c) / (fact n) * (a - c) ^ n"
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    75
     by (rule Maclaurin_minus [THEN exE])
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    76
  then have "a<x+c & x+c<c \<and> f a = (\<Sum>m<n. diff m c / (fact m) * (a - c) ^ m) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    77
      diff n (x+c) / (fact n) * (a - c) ^ n"
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    78
    by fastforce
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    79
  thus ?thesis by fastforce
17634
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    80
qed
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    81
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    82
lemma taylor:
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    83
  fixes a::real
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
    84
  assumes INIT: "n>0" "diff 0 = f"
17634
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    85
  and DERIV: "(\<forall> m t. m < n & a \<le> t & t \<le> b \<longrightarrow> DERIV (diff m) t :> (diff (Suc m) t))"
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    86
  and INTERV: "a \<le> c " "c \<le> b" "a \<le> x" "x \<le> b" "x \<noteq> c" 
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    87
  shows "\<exists> t. (if x<c then (x < t & t < c) else (c < t & t < x)) &
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    88
    f x = (\<Sum>m<n. (diff m c / (fact m)) * (x - c)^m) + (diff n t / (fact n)) * (x - c)^n" 
17634
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    89
proof (cases "x<c")
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    90
  case True
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    91
  note INIT
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    92
  moreover from DERIV and INTERV
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    93
  have "\<forall>m t. m < n \<and> x \<le> t \<and> t \<le> b \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 28952
diff changeset
    94
    by fastforce
17634
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    95
  moreover note True
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    96
  moreover from INTERV have "c \<le> b" by simp
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    97
  ultimately have "\<exists>t>x. t < c \<and> f x =
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    98
    (\<Sum>m<n. diff m c / (fact m) * (x - c) ^ m) + diff n t / (fact n) * (x - c) ^ n"
17634
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
    99
    by (rule taylor_down)
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
   100
  with True show ?thesis by simp
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
   101
next
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
   102
  case False
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
   103
  note INIT
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
   104
  moreover from DERIV and INTERV
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
   105
  have "\<forall>m t. m < n \<and> a \<le> t \<and> t \<le> x \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 28952
diff changeset
   106
    by fastforce
17634
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
   107
  moreover from INTERV have "a \<le> c" by arith
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
   108
  moreover from False and INTERV have "c < x" by arith
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   109
  ultimately have "\<exists>t>c. t < x \<and> f x =
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   110
    (\<Sum>m<n. diff m c / (fact m) * (x - c) ^ m) + diff n t / (fact n) * (x - c) ^ n" 
17634
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
   111
    by (rule taylor_up)
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
   112
  with False show ?thesis by simp
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
   113
qed
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
   114
e83ce8fe58fa Formalization of Taylor series by Lukas Bulwahn and
berghofe
parents:
diff changeset
   115
end