| author | hoelzl | 
| Mon, 14 Mar 2011 14:37:35 +0100 | |
| changeset 41970 | 47d6e13d1710 | 
| parent 38159 | e9b4835a54ee | 
| child 44821 | a92f65e174cf | 
| permissions | -rw-r--r-- | 
| 38159 | 1 | (* Title: HOL/Old_Number_Theory/Fib.thy | 
| 9944 | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 3 | Copyright 1997 University of Cambridge | |
| 4 | *) | |
| 5 | ||
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 6 | header {* The Fibonacci function *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 7 | |
| 27556 | 8 | theory Fib | 
| 9 | imports Primes | |
| 10 | begin | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 11 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 12 | text {*
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 13 | Fibonacci numbers: proofs of laws taken from: | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 14 | R. L. Graham, D. E. Knuth, O. Patashnik. Concrete Mathematics. | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 15 | (Addison-Wesley, 1989) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 16 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 17 | \bigskip | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 18 | *} | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 19 | |
| 24549 | 20 | fun fib :: "nat \<Rightarrow> nat" | 
| 21 | where | |
| 38159 | 22 | "fib 0 = 0" | 
| 23 | | "fib (Suc 0) = 1" | |
| 24549 | 24 | | fib_2: "fib (Suc (Suc n)) = fib n + fib (Suc n)" | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 25 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 26 | text {*
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 27 | \medskip The difficulty in these proofs is to ensure that the | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 28 |   induction hypotheses are applied before the definition of @{term
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 29 |   fib}.  Towards this end, the @{term fib} equations are not declared
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 30 | to the Simplifier and are applied very selectively at first. | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 31 | *} | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 32 | |
| 24549 | 33 | text{*We disable @{text fib.fib_2fib_2} for simplification ...*}
 | 
| 34 | declare fib_2 [simp del] | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 35 | |
| 15628 | 36 | text{*...then prove a version that has a more restrictive pattern.*}
 | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 37 | lemma fib_Suc3: "fib (Suc (Suc (Suc n))) = fib (Suc n) + fib (Suc (Suc n))" | 
| 24549 | 38 | by (rule fib_2) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 39 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 40 | text {* \medskip Concrete Mathematics, page 280 *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 41 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 42 | lemma fib_add: "fib (Suc (n + k)) = fib (Suc k) * fib (Suc n) + fib k * fib n" | 
| 24573 | 43 | proof (induct n rule: fib.induct) | 
| 44 | case 1 show ?case by simp | |
| 45 | next | |
| 46 | case 2 show ?case by (simp add: fib_2) | |
| 47 | next | |
| 25361 
1aa441e48496
function package: using the names of the equations as case names turned out to be impractical => disabled
 krauss parents: 
25222diff
changeset | 48 | case 3 thus ?case by (simp add: fib_2 add_mult_distrib2) | 
| 24573 | 49 | qed | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 50 | |
| 15628 | 51 | lemma fib_Suc_neq_0: "fib (Suc n) \<noteq> 0" | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 52 | apply (induct n rule: fib.induct) | 
| 24549 | 53 | apply (simp_all add: fib_2) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 54 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 55 | |
| 15628 | 56 | lemma fib_Suc_gr_0: "0 < fib (Suc n)" | 
| 57 | by (insert fib_Suc_neq_0 [of n], simp) | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 58 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 59 | lemma fib_gr_0: "0 < n ==> 0 < fib n" | 
| 15628 | 60 | by (case_tac n, auto simp add: fib_Suc_gr_0) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 61 | |
| 9944 | 62 | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 63 | text {*
 | 
| 15628 | 64 | \medskip Concrete Mathematics, page 278: Cassini's identity. The proof is | 
| 65 | much easier using integers, not natural numbers! | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 66 | *} | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 67 | |
| 15628 | 68 | lemma fib_Cassini_int: | 
| 15003 | 69 | "int (fib (Suc (Suc n)) * fib n) = | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11786diff
changeset | 70 | (if n mod 2 = 0 then int (fib (Suc n) * fib (Suc n)) - 1 | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11786diff
changeset | 71 | else int (fib (Suc n) * fib (Suc n)) + 1)" | 
| 23315 | 72 | proof(induct n rule: fib.induct) | 
| 24549 | 73 | case 1 thus ?case by (simp add: fib_2) | 
| 23315 | 74 | next | 
| 24549 | 75 | case 2 thus ?case by (simp add: fib_2 mod_Suc) | 
| 23315 | 76 | next | 
| 25361 
1aa441e48496
function package: using the names of the equations as case names turned out to be impractical => disabled
 krauss parents: 
25222diff
changeset | 77 | case (3 x) | 
| 24573 | 78 | have "Suc 0 \<noteq> x mod 2 \<longrightarrow> x mod 2 = 0" by presburger | 
| 25361 
1aa441e48496
function package: using the names of the equations as case names turned out to be impractical => disabled
 krauss parents: 
25222diff
changeset | 79 | with "3.hyps" show ?case by (simp add: fib.simps add_mult_distrib add_mult_distrib2) | 
| 23315 | 80 | qed | 
| 15628 | 81 | |
| 82 | text{*We now obtain a version for the natural numbers via the coercion 
 | |
| 83 |    function @{term int}.*}
 | |
| 84 | theorem fib_Cassini: | |
| 85 | "fib (Suc (Suc n)) * fib n = | |
| 86 | (if n mod 2 = 0 then fib (Suc n) * fib (Suc n) - 1 | |
| 87 | else fib (Suc n) * fib (Suc n) + 1)" | |
| 88 | apply (rule int_int_eq [THEN iffD1]) | |
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23365diff
changeset | 89 | apply (simp add: fib_Cassini_int) | 
| 15628 | 90 | apply (subst zdiff_int [symmetric]) | 
| 91 | apply (insert fib_Suc_gr_0 [of n], simp_all) | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 92 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 93 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 94 | |
| 15628 | 95 | text {* \medskip Toward Law 6.111 of Concrete Mathematics *}
 | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 96 | |
| 27556 | 97 | lemma gcd_fib_Suc_eq_1: "gcd (fib n) (fib (Suc n)) = Suc 0" | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 98 | apply (induct n rule: fib.induct) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 99 | prefer 3 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 100 | apply (simp add: gcd_commute fib_Suc3) | 
| 24549 | 101 | apply (simp_all add: fib_2) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 102 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 103 | |
| 27556 | 104 | lemma gcd_fib_add: "gcd (fib m) (fib (n + m)) = gcd (fib m) (fib n)" | 
| 15628 | 105 | apply (simp add: gcd_commute [of "fib m"]) | 
| 106 | apply (case_tac m) | |
| 107 | apply simp | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 108 | apply (simp add: fib_add) | 
| 15628 | 109 | apply (simp add: add_commute gcd_non_0 [OF fib_Suc_gr_0]) | 
| 110 | apply (simp add: gcd_non_0 [OF fib_Suc_gr_0, symmetric]) | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 111 | apply (simp add: gcd_fib_Suc_eq_1 gcd_mult_cancel) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 112 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 113 | |
| 27556 | 114 | lemma gcd_fib_diff: "m \<le> n ==> gcd (fib m) (fib (n - m)) = gcd (fib m) (fib n)" | 
| 15628 | 115 | by (simp add: gcd_fib_add [symmetric, of _ "n-m"]) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 116 | |
| 27556 | 117 | lemma gcd_fib_mod: "0 < m ==> gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)" | 
| 23688 | 118 | proof (induct n rule: less_induct) | 
| 119 | case (less n) | |
| 120 | from less.prems have pos_m: "0 < m" . | |
| 27556 | 121 | show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)" | 
| 23688 | 122 | proof (cases "m < n") | 
| 123 | case True note m_n = True | |
| 124 | then have m_n': "m \<le> n" by auto | |
| 125 | with pos_m have pos_n: "0 < n" by auto | |
| 126 | with pos_m m_n have diff: "n - m < n" by auto | |
| 27556 | 127 | have "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib ((n - m) mod m))" | 
| 23688 | 128 | by (simp add: mod_if [of n]) (insert m_n, auto) | 
| 27556 | 129 | also have "\<dots> = gcd (fib m) (fib (n - m))" by (simp add: less.hyps diff pos_m) | 
| 130 | also have "\<dots> = gcd (fib m) (fib n)" by (simp add: gcd_fib_diff m_n') | |
| 131 | finally show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)" . | |
| 23688 | 132 | next | 
| 27556 | 133 | case False then show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)" | 
| 23688 | 134 | by (cases "m = n") auto | 
| 135 | qed | |
| 136 | qed | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 137 | |
| 27556 | 138 | lemma fib_gcd: "fib (gcd m n) = gcd (fib m) (fib n)"  -- {* Law 6.111 *}
 | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 139 | apply (induct m n rule: gcd_induct) | 
| 15628 | 140 | apply (simp_all add: gcd_non_0 gcd_commute gcd_fib_mod) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 141 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 142 | |
| 15628 | 143 | theorem fib_mult_eq_setsum: | 
| 11786 | 144 |     "fib (Suc n) * fib n = (\<Sum>k \<in> {..n}. fib k * fib k)"
 | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 145 | apply (induct n rule: fib.induct) | 
| 24549 | 146 | apply (auto simp add: atMost_Suc fib_2) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 147 | apply (simp add: add_mult_distrib add_mult_distrib2) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 148 | done | 
| 9944 | 149 | |
| 150 | end |