| author | haftmann | 
| Sat, 27 Jun 2015 20:20:33 +0200 | |
| changeset 60597 | 2da9b632069b | 
| parent 59730 | b7c394c7a619 | 
| child 60758 | d8d85a8172b5 | 
| permissions | -rw-r--r-- | 
| 28952 
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
 haftmann parents: 
25162diff
changeset | 1 | (* Title: HOL/Taylor.thy | 
| 17634 | 2 | Author: Lukas Bulwahn, Bernhard Haeupler, Technische Universitaet Muenchen | 
| 3 | *) | |
| 4 | ||
| 58889 | 5 | section {* Taylor series *}
 | 
| 17634 | 6 | |
| 7 | theory Taylor | |
| 8 | imports MacLaurin | |
| 9 | begin | |
| 10 | ||
| 11 | text {*
 | |
| 12 | We use MacLaurin and the translation of the expansion point @{text c} to @{text 0}
 | |
| 13 | to prove Taylor's theorem. | |
| 14 | *} | |
| 15 | ||
| 16 | lemma taylor_up: | |
| 25162 | 17 | assumes INIT: "n>0" "diff 0 = f" | 
| 17634 | 18 | and DERIV: "(\<forall> m t. m < n & a \<le> t & t \<le> b \<longrightarrow> DERIV (diff m) t :> (diff (Suc m) t))" | 
| 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: 
58889diff
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: 
58889diff
changeset | 21 | f b = (\<Sum>m<n. (diff m c / (fact m)) * (b - c)^m) + (diff n t / (fact n)) * (b - c)^n" | 
| 17634 | 22 | proof - | 
| 23 | from INTERV have "0 < b-c" by arith | |
| 24 | moreover | |
| 25162 | 25 | from INIT have "n>0" "((\<lambda>m x. diff m (x + c)) 0) = (\<lambda>x. f (x + c))" by auto | 
| 17634 | 26 | moreover | 
| 25162 | 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 | 28 | proof (intro strip) | 
| 29 | fix m t | |
| 30 | assume "m < n & 0 <= t & t <= b - c" | |
| 31 | with DERIV and INTERV have "DERIV (diff m) (t + c) :> diff (Suc m) (t + c)" by auto | |
| 32 | moreover | |
| 23069 
cdfff0241c12
rename lemmas LIM_ident, isCont_ident, DERIV_ident
 huffman parents: 
17634diff
changeset | 33 | from DERIV_ident and DERIV_const have "DERIV (%x. x + c) t :> 1+0" by (rule DERIV_add) | 
| 17634 | 34 | ultimately have "DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c) * (1+0)" | 
| 35 | by (rule DERIV_chain2) | |
| 36 | thus "DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c)" by simp | |
| 37 | qed | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 38 | ultimately obtain x where | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 39 | "0 < x & x < b - c & | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
58889diff
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: 
58889diff
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: 
58889diff
changeset | 42 | by (rule Maclaurin [THEN exE]) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
58889diff
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: 
58889diff
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: 
58889diff
changeset | 45 | by fastforce | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 46 | thus ?thesis by fastforce | 
| 17634 | 47 | qed | 
| 48 | ||
| 49 | lemma taylor_down: | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 50 | fixes a::real | 
| 25162 | 51 | assumes INIT: "n>0" "diff 0 = f" | 
| 17634 | 52 | and DERIV: "(\<forall> m t. m < n & a \<le> t & t \<le> b \<longrightarrow> DERIV (diff m) t :> (diff (Suc m) t))" | 
| 53 | and INTERV: "a < c" "c \<le> b" | |
| 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: 
58889diff
changeset | 55 | f a = (\<Sum>m<n. (diff m c / (fact m)) * (a - c)^m) + (diff n t / (fact n)) * (a - c)^n" | 
| 17634 | 56 | proof - | 
| 57 | from INTERV have "a-c < 0" by arith | |
| 58 | moreover | |
| 25162 | 59 | from INIT have "n>0" "((\<lambda>m x. diff m (x + c)) 0) = (\<lambda>x. f (x + c))" by auto | 
| 17634 | 60 | moreover | 
| 61 | have "ALL m t. m < n & a-c <= t & t <= 0 --> DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c)" | |
| 62 | proof (rule allI impI)+ | |
| 63 | fix m t | |
| 64 | assume "m < n & a-c <= t & t <= 0" | |
| 65 | with DERIV and INTERV have "DERIV (diff m) (t + c) :> diff (Suc m) (t + c)" by auto | |
| 66 | moreover | |
| 23069 
cdfff0241c12
rename lemmas LIM_ident, isCont_ident, DERIV_ident
 huffman parents: 
17634diff
changeset | 67 | from DERIV_ident and DERIV_const have "DERIV (%x. x + c) t :> 1+0" by (rule DERIV_add) | 
| 17634 | 68 | ultimately have "DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c) * (1+0)" by (rule DERIV_chain2) | 
| 69 | thus "DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c)" by simp | |
| 70 | qed | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 71 | ultimately obtain x where | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 72 | "a - c < x & x < 0 & | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 73 | f (a - c + c) = (SUM m<n. diff m (0 + c) / (fact m) * (a - c) ^ m) + | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
58889diff
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: 
58889diff
changeset | 75 | by (rule Maclaurin_minus [THEN exE]) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
58889diff
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: 
58889diff
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: 
58889diff
changeset | 78 | by fastforce | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 79 | thus ?thesis by fastforce | 
| 17634 | 80 | qed | 
| 81 | ||
| 82 | lemma taylor: | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 83 | fixes a::real | 
| 25162 | 84 | assumes INIT: "n>0" "diff 0 = f" | 
| 17634 | 85 | and DERIV: "(\<forall> m t. m < n & a \<le> t & t \<le> b \<longrightarrow> DERIV (diff m) t :> (diff (Suc m) t))" | 
| 86 | and INTERV: "a \<le> c " "c \<le> b" "a \<le> x" "x \<le> b" "x \<noteq> c" | |
| 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: 
58889diff
changeset | 88 | f x = (\<Sum>m<n. (diff m c / (fact m)) * (x - c)^m) + (diff n t / (fact n)) * (x - c)^n" | 
| 17634 | 89 | proof (cases "x<c") | 
| 90 | case True | |
| 91 | note INIT | |
| 92 | moreover from DERIV and INTERV | |
| 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: 
28952diff
changeset | 94 | by fastforce | 
| 17634 | 95 | moreover note True | 
| 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: 
58889diff
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: 
58889diff
changeset | 98 | (\<Sum>m<n. diff m c / (fact m) * (x - c) ^ m) + diff n t / (fact n) * (x - c) ^ n" | 
| 17634 | 99 | by (rule taylor_down) | 
| 100 | with True show ?thesis by simp | |
| 101 | next | |
| 102 | case False | |
| 103 | note INIT | |
| 104 | moreover from DERIV and INTERV | |
| 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: 
28952diff
changeset | 106 | by fastforce | 
| 17634 | 107 | moreover from INTERV have "a \<le> c" by arith | 
| 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: 
58889diff
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: 
58889diff
changeset | 110 | (\<Sum>m<n. diff m c / (fact m) * (x - c) ^ m) + diff n t / (fact n) * (x - c) ^ n" | 
| 17634 | 111 | by (rule taylor_up) | 
| 112 | with False show ?thesis by simp | |
| 113 | qed | |
| 114 | ||
| 115 | end |