paulson@15628: (* ID: $Id$ paulson@9944: Author: Lawrence C Paulson, Cambridge University Computer Laboratory paulson@9944: Copyright 1997 University of Cambridge paulson@9944: *) paulson@9944: wenzelm@11049: header {* The Fibonacci function *} wenzelm@11049: haftmann@16417: theory Fib imports Primes begin wenzelm@11049: wenzelm@11049: text {* wenzelm@11049: Fibonacci numbers: proofs of laws taken from: wenzelm@11049: R. L. Graham, D. E. Knuth, O. Patashnik. Concrete Mathematics. wenzelm@11049: (Addison-Wesley, 1989) wenzelm@11049: wenzelm@11049: \bigskip wenzelm@11049: *} wenzelm@11049: paulson@24549: fun fib :: "nat \ nat" paulson@24549: where paulson@24573: "fib 0 = 0" paulson@24573: | "fib (Suc 0) = 1" paulson@24549: | fib_2: "fib (Suc (Suc n)) = fib n + fib (Suc n)" wenzelm@11049: wenzelm@11049: text {* wenzelm@11049: \medskip The difficulty in these proofs is to ensure that the wenzelm@11049: induction hypotheses are applied before the definition of @{term wenzelm@11049: fib}. Towards this end, the @{term fib} equations are not declared wenzelm@11049: to the Simplifier and are applied very selectively at first. wenzelm@11049: *} wenzelm@11049: paulson@24549: text{*We disable @{text fib.fib_2fib_2} for simplification ...*} paulson@24549: declare fib_2 [simp del] wenzelm@11049: paulson@15628: text{*...then prove a version that has a more restrictive pattern.*} wenzelm@11049: lemma fib_Suc3: "fib (Suc (Suc (Suc n))) = fib (Suc n) + fib (Suc (Suc n))" paulson@24549: by (rule fib_2) wenzelm@11049: wenzelm@11049: text {* \medskip Concrete Mathematics, page 280 *} wenzelm@11049: wenzelm@11049: lemma fib_add: "fib (Suc (n + k)) = fib (Suc k) * fib (Suc n) + fib k * fib n" paulson@24573: proof (induct n rule: fib.induct) paulson@24573: case 1 show ?case by simp paulson@24573: next paulson@24573: case 2 show ?case by (simp add: fib_2) paulson@24573: next paulson@24573: case 3 thus ?case by (simp add: fib_2 add_mult_distrib2) paulson@24573: qed wenzelm@11049: paulson@15628: lemma fib_Suc_neq_0: "fib (Suc n) \ 0" wenzelm@11049: apply (induct n rule: fib.induct) paulson@24549: apply (simp_all add: fib_2) wenzelm@11049: done wenzelm@11049: paulson@15628: lemma fib_Suc_gr_0: "0 < fib (Suc n)" paulson@15628: by (insert fib_Suc_neq_0 [of n], simp) wenzelm@11049: wenzelm@11049: lemma fib_gr_0: "0 < n ==> 0 < fib n" paulson@15628: by (case_tac n, auto simp add: fib_Suc_gr_0) wenzelm@11049: paulson@9944: wenzelm@11049: text {* paulson@15628: \medskip Concrete Mathematics, page 278: Cassini's identity. The proof is paulson@15628: much easier using integers, not natural numbers! wenzelm@11049: *} wenzelm@11049: paulson@15628: lemma fib_Cassini_int: paulson@15003: "int (fib (Suc (Suc n)) * fib n) = paulson@11868: (if n mod 2 = 0 then int (fib (Suc n) * fib (Suc n)) - 1 paulson@11868: else int (fib (Suc n) * fib (Suc n)) + 1)" chaieb@23315: proof(induct n rule: fib.induct) paulson@24549: case 1 thus ?case by (simp add: fib_2) chaieb@23315: next paulson@24549: case 2 thus ?case by (simp add: fib_2 mod_Suc) chaieb@23315: next chaieb@23315: case (3 x) paulson@24573: have "Suc 0 \ x mod 2 \ x mod 2 = 0" by presburger paulson@24573: with "3.hyps" show ?case by (simp add: fib_2 add_mult_distrib add_mult_distrib2) chaieb@23315: qed paulson@15628: paulson@15628: text{*We now obtain a version for the natural numbers via the coercion paulson@15628: function @{term int}.*} paulson@15628: theorem fib_Cassini: paulson@15628: "fib (Suc (Suc n)) * fib n = paulson@15628: (if n mod 2 = 0 then fib (Suc n) * fib (Suc n) - 1 paulson@15628: else fib (Suc n) * fib (Suc n) + 1)" paulson@15628: apply (rule int_int_eq [THEN iffD1]) huffman@23431: apply (simp add: fib_Cassini_int) paulson@15628: apply (subst zdiff_int [symmetric]) paulson@15628: apply (insert fib_Suc_gr_0 [of n], simp_all) wenzelm@11049: done wenzelm@11049: wenzelm@11049: paulson@15628: text {* \medskip Toward Law 6.111 of Concrete Mathematics *} wenzelm@11049: wenzelm@11701: lemma gcd_fib_Suc_eq_1: "gcd (fib n, fib (Suc n)) = Suc 0" wenzelm@11049: apply (induct n rule: fib.induct) wenzelm@11049: prefer 3 wenzelm@11049: apply (simp add: gcd_commute fib_Suc3) paulson@24549: apply (simp_all add: fib_2) wenzelm@11049: done wenzelm@11049: wenzelm@11049: lemma gcd_fib_add: "gcd (fib m, fib (n + m)) = gcd (fib m, fib n)" paulson@15628: apply (simp add: gcd_commute [of "fib m"]) paulson@15628: apply (case_tac m) paulson@15628: apply simp wenzelm@11049: apply (simp add: fib_add) paulson@15628: apply (simp add: add_commute gcd_non_0 [OF fib_Suc_gr_0]) paulson@15628: apply (simp add: gcd_non_0 [OF fib_Suc_gr_0, symmetric]) wenzelm@11049: apply (simp add: gcd_fib_Suc_eq_1 gcd_mult_cancel) wenzelm@11049: done wenzelm@11049: wenzelm@11049: lemma gcd_fib_diff: "m \ n ==> gcd (fib m, fib (n - m)) = gcd (fib m, fib n)" paulson@15628: by (simp add: gcd_fib_add [symmetric, of _ "n-m"]) wenzelm@11049: wenzelm@11049: lemma gcd_fib_mod: "0 < m ==> gcd (fib m, fib (n mod m)) = gcd (fib m, fib n)" haftmann@23688: proof (induct n rule: less_induct) haftmann@23688: case (less n) haftmann@23688: from less.prems have pos_m: "0 < m" . haftmann@23688: show "gcd (fib m, fib (n mod m)) = gcd (fib m, fib n)" haftmann@23688: proof (cases "m < n") haftmann@23688: case True note m_n = True haftmann@23688: then have m_n': "m \ n" by auto haftmann@23688: with pos_m have pos_n: "0 < n" by auto haftmann@23688: with pos_m m_n have diff: "n - m < n" by auto haftmann@23688: have "gcd (fib m, fib (n mod m)) = gcd (fib m, fib ((n - m) mod m))" haftmann@23688: by (simp add: mod_if [of n]) (insert m_n, auto) haftmann@23688: also have "\ = gcd (fib m, fib (n - m))" by (simp add: less.hyps diff pos_m) haftmann@23688: also have "\ = gcd (fib m, fib n)" by (simp add: gcd_fib_diff m_n') haftmann@23688: finally show "gcd (fib m, fib (n mod m)) = gcd (fib m, fib n)" . haftmann@23688: next haftmann@23688: case False then show "gcd (fib m, fib (n mod m)) = gcd (fib m, fib n)" haftmann@23688: by (cases "m = n") auto haftmann@23688: qed haftmann@23688: qed wenzelm@11049: wenzelm@11049: lemma fib_gcd: "fib (gcd (m, n)) = gcd (fib m, fib n)" -- {* Law 6.111 *} wenzelm@11049: apply (induct m n rule: gcd_induct) paulson@15628: apply (simp_all add: gcd_non_0 gcd_commute gcd_fib_mod) wenzelm@11049: done wenzelm@11049: paulson@15628: theorem fib_mult_eq_setsum: wenzelm@11786: "fib (Suc n) * fib n = (\k \ {..n}. fib k * fib k)" wenzelm@11049: apply (induct n rule: fib.induct) paulson@24549: apply (auto simp add: atMost_Suc fib_2) wenzelm@11049: apply (simp add: add_mult_distrib add_mult_distrib2) wenzelm@11049: done paulson@9944: paulson@9944: end