src/HOL/Number_Theory/Fib.thy
changeset 64317 029e6247210e
parent 64267 b9a1486e79be
child 65393 079a6f850c02
equal deleted inserted replaced
64316:96fef7745c68 64317:029e6247210e
     1 (*  Title:      HOL/Number_Theory/Fib.thy
     1 (*  Title:      HOL/Number_Theory/Fib.thy
     2     Author:     Lawrence C. Paulson
     2     Author:     Lawrence C. Paulson
     3     Author:     Jeremy Avigad
     3     Author:     Jeremy Avigad
       
     4     Author:     Manuel Eberl
     4 *)
     5 *)
     5 
     6 
     6 section \<open>The fibonacci function\<close>
     7 section \<open>The fibonacci function\<close>
     7 
     8 
     8 theory Fib
     9 theory Fib
     9 imports Main GCD Binomial
    10   imports Complex_Main
    10 begin
    11 begin
    11 
    12 
    12 
    13 
    13 subsection \<open>Fibonacci numbers\<close>
    14 subsection \<open>Fibonacci numbers\<close>
    14 
    15 
    34 lemma fib_add: "fib (Suc (n+k)) = fib (Suc k) * fib (Suc n) + fib k * fib n"
    35 lemma fib_add: "fib (Suc (n+k)) = fib (Suc k) * fib (Suc n) + fib k * fib n"
    35   by (induct n rule: fib.induct) (auto simp add: field_simps)
    36   by (induct n rule: fib.induct) (auto simp add: field_simps)
    36 
    37 
    37 lemma fib_neq_0_nat: "n > 0 \<Longrightarrow> fib n > 0"
    38 lemma fib_neq_0_nat: "n > 0 \<Longrightarrow> fib n > 0"
    38   by (induct n rule: fib.induct) (auto simp add: )
    39   by (induct n rule: fib.induct) (auto simp add: )
       
    40 
       
    41 
       
    42 subsection \<open>More efficient code\<close>
       
    43 
       
    44 text \<open>
       
    45   The naive approach is very inefficient since the branching recursion leads to many
       
    46   values of @{term fib} being computed multiple times. We can avoid this by ``remembering''
       
    47   the last two values in the sequence, yielding a tail-recursive version.
       
    48   This is far from optimal (it takes roughly $O(n\cdot M(n))$ time where $M(n)$ is the 
       
    49   time required to multiply two $n$-bit integers), but much better than the naive version,
       
    50   which is exponential.
       
    51 \<close>
       
    52 
       
    53 fun gen_fib :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" where
       
    54   "gen_fib a b 0 = a"
       
    55 | "gen_fib a b (Suc 0) = b"
       
    56 | "gen_fib a b (Suc (Suc n)) = gen_fib b (a + b) (Suc n)"
       
    57 
       
    58 lemma gen_fib_recurrence: "gen_fib a b (Suc (Suc n)) = gen_fib a b n + gen_fib a b (Suc n)"
       
    59   by (induction a b n rule: gen_fib.induct) simp_all
       
    60   
       
    61 lemma gen_fib_fib: "gen_fib (fib n) (fib (Suc n)) m = fib (n + m)"
       
    62   by (induction m rule: fib.induct) (simp_all del: gen_fib.simps(3) add: gen_fib_recurrence)
       
    63 
       
    64 lemma fib_conv_gen_fib: "fib n = gen_fib 0 1 n"
       
    65   using gen_fib_fib[of 0 n] by simp
       
    66 
       
    67 declare fib_conv_gen_fib [code]
    39 
    68 
    40 
    69 
    41 subsection \<open>A Few Elementary Results\<close>
    70 subsection \<open>A Few Elementary Results\<close>
    42 
    71 
    43 text \<open>
    72 text \<open>
   102 
   131 
   103 theorem fib_mult_eq_sum_nat: "fib (Suc n) * fib n = (\<Sum>k \<in> {..n}. fib k * fib k)"
   132 theorem fib_mult_eq_sum_nat: "fib (Suc n) * fib n = (\<Sum>k \<in> {..n}. fib k * fib k)"
   104   by (induct n rule: nat.induct) (auto simp add:  field_simps)
   133   by (induct n rule: nat.induct) (auto simp add:  field_simps)
   105 
   134 
   106 
   135 
       
   136 subsection \<open>Closed form\<close>
       
   137 
       
   138 lemma fib_closed_form:
       
   139   defines "\<phi> \<equiv> (1 + sqrt 5) / (2::real)" and "\<psi> \<equiv> (1 - sqrt 5) / (2::real)"
       
   140   shows   "of_nat (fib n) = (\<phi> ^ n - \<psi> ^ n) / sqrt 5"
       
   141 proof (induction n rule: fib.induct)
       
   142   fix n :: nat
       
   143   assume IH1: "of_nat (fib n) = (\<phi> ^ n - \<psi> ^ n) / sqrt 5"
       
   144   assume IH2: "of_nat (fib (Suc n)) = (\<phi> ^ Suc n - \<psi> ^ Suc n) / sqrt 5"
       
   145   have "of_nat (fib (Suc (Suc n))) = of_nat (fib (Suc n)) + of_nat (fib n)" by simp
       
   146   also have "... = (\<phi>^n*(\<phi> + 1) - \<psi>^n*(\<psi> + 1)) / sqrt 5"
       
   147     by (simp add: IH1 IH2 field_simps)
       
   148   also have "\<phi> + 1 = \<phi>\<^sup>2" by (simp add: \<phi>_def field_simps power2_eq_square)
       
   149   also have "\<psi> + 1 = \<psi>\<^sup>2" by (simp add: \<psi>_def field_simps power2_eq_square)
       
   150   also have "\<phi>^n * \<phi>\<^sup>2 - \<psi>^n * \<psi>\<^sup>2 = \<phi> ^ Suc (Suc n) - \<psi> ^ Suc (Suc n)"  
       
   151     by (simp add: power2_eq_square)
       
   152   finally show "of_nat (fib (Suc (Suc n))) = (\<phi> ^ Suc (Suc n) - \<psi> ^ Suc (Suc n)) / sqrt 5" .
       
   153 qed (simp_all add: \<phi>_def \<psi>_def field_simps)
       
   154 
       
   155 lemma fib_closed_form':
       
   156   defines "\<phi> \<equiv> (1 + sqrt 5) / (2 :: real)" and "\<psi> \<equiv> (1 - sqrt 5) / (2 :: real)"
       
   157   assumes "n > 0"
       
   158   shows   "fib n = round (\<phi> ^ n / sqrt 5)"
       
   159 proof (rule sym, rule round_unique')
       
   160   have "\<bar>\<phi> ^ n / sqrt 5 - of_int (int (fib n))\<bar> = \<bar>\<psi>\<bar> ^ n / sqrt 5"
       
   161     by (simp add: fib_closed_form[folded \<phi>_def \<psi>_def] field_simps power_abs)
       
   162   also {
       
   163     from assms have "\<bar>\<psi>\<bar>^n \<le> \<bar>\<psi>\<bar>^1"
       
   164       by (intro power_decreasing) (simp_all add: algebra_simps real_le_lsqrt)
       
   165     also have "... < sqrt 5 / 2" by (simp add: \<psi>_def field_simps)
       
   166     finally have "\<bar>\<psi>\<bar>^n / sqrt 5 < 1/2" by (simp add: field_simps)
       
   167   }
       
   168   finally show "\<bar>\<phi> ^ n / sqrt 5 - of_int (int (fib n))\<bar> < 1/2" .
       
   169 qed
       
   170 
       
   171 lemma fib_asymptotics:
       
   172   defines "\<phi> \<equiv> (1 + sqrt 5) / (2 :: real)"
       
   173   shows   "(\<lambda>n. real (fib n) / (\<phi> ^ n / sqrt 5)) \<longlonglongrightarrow> 1"
       
   174 proof -
       
   175   define \<psi> where "\<psi> \<equiv> (1 - sqrt 5) / (2 :: real)"
       
   176   have "\<phi> > 1" by (simp add: \<phi>_def)
       
   177   hence A: "\<phi> \<noteq> 0" by auto
       
   178   have "(\<lambda>n. (\<psi> / \<phi>) ^ n) \<longlonglongrightarrow> 0"
       
   179     by (rule LIMSEQ_power_zero) (simp_all add: \<phi>_def \<psi>_def field_simps add_pos_pos)
       
   180   hence "(\<lambda>n. 1 - (\<psi> / \<phi>) ^ n) \<longlonglongrightarrow> 1 - 0" by (intro tendsto_diff tendsto_const)
       
   181   with A show ?thesis
       
   182     by (simp add: divide_simps fib_closed_form [folded \<phi>_def \<psi>_def])
       
   183 qed
       
   184 
       
   185 
       
   186 subsection \<open>Divide-and-Conquer recurrence\<close>
       
   187 
       
   188 text \<open>
       
   189   The following divide-and-conquer recurrence allows for a more efficient computation 
       
   190   of Fibonacci numbers; however, it requires memoisation of values to be reasonably 
       
   191   efficient, cutting the number of values to be computed to logarithmically many instead of
       
   192   linearly many. The vast majority of the computation time is then actually spent on the 
       
   193   multiplication, since the output number is exponential in the input number.
       
   194 \<close>
       
   195 
       
   196 lemma fib_rec_odd:
       
   197   defines "\<phi> \<equiv> (1 + sqrt 5) / (2 :: real)" and "\<psi> \<equiv> (1 - sqrt 5) / (2 :: real)"
       
   198   shows   "fib (Suc (2*n)) = fib n^2 + fib (Suc n)^2"
       
   199 proof -
       
   200   have "of_nat (fib n^2 + fib (Suc n)^2) = ((\<phi> ^ n - \<psi> ^ n)\<^sup>2 + (\<phi> * \<phi> ^ n - \<psi> * \<psi> ^ n)\<^sup>2)/5"
       
   201     by (simp add: fib_closed_form[folded \<phi>_def \<psi>_def] field_simps power2_eq_square)
       
   202   also have "(\<phi> ^ n - \<psi> ^ n)\<^sup>2 + (\<phi> * \<phi> ^ n - \<psi> * \<psi> ^ n)\<^sup>2 = 
       
   203     \<phi>^(2*n) + \<psi>^(2*n) - 2*(\<phi>*\<psi>)^n + \<phi>^(2*n+2) + \<psi>^(2*n+2) - 2*(\<phi>*\<psi>)^(n+1)" (is "_ = ?A")
       
   204       by (simp add: power2_eq_square algebra_simps power_mult power_mult_distrib)
       
   205   also have "\<phi> * \<psi> = -1" by (simp add: \<phi>_def \<psi>_def field_simps)
       
   206   hence "?A = \<phi>^(2*n+1) * (\<phi> + inverse \<phi>) + \<psi>^(2*n+1) * (\<psi> + inverse \<psi>)" 
       
   207     by (auto simp: field_simps power2_eq_square)
       
   208   also have "1 + sqrt 5 > 0" by (auto intro: add_pos_pos)
       
   209   hence "\<phi> + inverse \<phi> = sqrt 5" by (simp add: \<phi>_def field_simps)
       
   210   also have "\<psi> + inverse \<psi> = -sqrt 5" by (simp add: \<psi>_def field_simps)
       
   211   also have "(\<phi> ^ (2*n+1) * sqrt 5 + \<psi> ^ (2*n+1)* - sqrt 5) / 5 =
       
   212                (\<phi> ^ (2*n+1) - \<psi> ^ (2*n+1)) * (sqrt 5 / 5)" by (simp add: field_simps)
       
   213   also have "sqrt 5 / 5 = inverse (sqrt 5)" by (simp add: field_simps)
       
   214   also have "(\<phi> ^ (2*n+1) - \<psi> ^ (2*n+1)) * ... = of_nat (fib (Suc (2*n)))"
       
   215     by (simp add: fib_closed_form[folded \<phi>_def \<psi>_def] divide_inverse)
       
   216   finally show ?thesis by (simp only: of_nat_eq_iff)
       
   217 qed
       
   218 
       
   219 lemma fib_rec_even: "fib (2*n) = (fib (n - 1) + fib (n + 1)) * fib n"
       
   220 proof (induction n)
       
   221   case (Suc n)
       
   222   let ?rfib = "\<lambda>x. real (fib x)"
       
   223   have "2 * (Suc n) = Suc (Suc (2*n))" by simp
       
   224   also have "real (fib ...) = ?rfib n^2 + ?rfib (Suc n)^2 + (?rfib (n - 1) + ?rfib (n + 1)) * ?rfib n" 
       
   225     by (simp add: fib_rec_odd Suc)
       
   226   also have "(?rfib (n - 1) + ?rfib (n + 1)) * ?rfib n = (2 * ?rfib (n + 1) - ?rfib n) * ?rfib n"
       
   227     by (cases n) simp_all
       
   228   also have "?rfib n^2 + ?rfib (Suc n)^2 + ... = (?rfib (Suc n) + 2 * ?rfib n) * ?rfib (Suc n)"
       
   229     by (simp add: algebra_simps power2_eq_square)
       
   230   also have "... = real ((fib (Suc n - 1) + fib (Suc n + 1)) * fib (Suc n))" by simp
       
   231   finally show ?case by (simp only: of_nat_eq_iff)
       
   232 qed simp
       
   233 
       
   234 lemma fib_rec_even': "fib (2*n) = (2*fib (n - 1) + fib n) * fib n"
       
   235   by (subst fib_rec_even, cases n) simp_all
       
   236 
       
   237 lemma fib_rec:
       
   238   "fib n = (if n = 0 then 0 else if n = 1 then 1 else
       
   239             if even n then let n' = n div 2; fn = fib n' in (2 * fib (n' - 1) + fn) * fn
       
   240             else let n' = n div 2 in fib n' ^ 2 + fib (Suc n') ^ 2)"
       
   241   by (auto elim: evenE oddE simp: fib_rec_odd fib_rec_even' Let_def)
       
   242 
       
   243 
   107 subsection \<open>Fibonacci and Binomial Coefficients\<close>
   244 subsection \<open>Fibonacci and Binomial Coefficients\<close>
   108 
   245 
   109 lemma sum_drop_zero: "(\<Sum>k = 0..Suc n. if 0<k then (f (k - 1)) else 0) = (\<Sum>j = 0..n. f j)"
   246 lemma sum_drop_zero: "(\<Sum>k = 0..Suc n. if 0<k then (f (k - 1)) else 0) = (\<Sum>j = 0..n. f j)"
   110   by (induct n) auto
   247   by (induct n) auto
   111 
   248