diff -r 4a16f8e41879 -r 7338eb25226c src/HOL/Isar_Examples/Fibonacci.thy --- a/src/HOL/Isar_Examples/Fibonacci.thy Tue Oct 07 20:43:18 2014 +0200 +++ b/src/HOL/Isar_Examples/Fibonacci.thy Tue Oct 07 20:59:46 2014 +0200 @@ -12,21 +12,21 @@ (Addison-Wesley, 1989) *) -header {* Fib and Gcd commute *} +header \Fib and Gcd commute\ theory Fibonacci imports "../Number_Theory/Primes" begin -text_raw {* \footnote{Isar version by Gertrud Bauer. Original tactic +text_raw \\footnote{Isar version by Gertrud Bauer. Original tactic script by Larry Paulson. A few proofs of laws taken from - \cite{Concrete-Math}.} *} + @{cite "Concrete-Math"}.}\ declare One_nat_def [simp] -subsection {* Fibonacci numbers *} +subsection \Fibonacci numbers\ fun fib :: "nat \ nat" where "fib 0 = 0" @@ -37,7 +37,7 @@ by (induct n rule: fib.induct) simp_all -text {* Alternative induction rule. *} +text \Alternative induction rule.\ theorem fib_induct: fixes n :: nat @@ -45,14 +45,14 @@ by (induct rule: fib.induct) simp_all -subsection {* Fib and gcd commute *} +subsection \Fib and gcd commute\ -text {* A few laws taken from \cite{Concrete-Math}. *} +text \A few laws taken from @{cite "Concrete-Math"}.\ lemma fib_add: "fib (n + k + 1) = fib (k + 1) * fib (n + 1) + fib k * fib n" (is "?P n") - -- {* see \cite[page 280]{Concrete-Math} *} + -- \see @{cite \page 280\ "Concrete-Math"}\ proof (induct n rule: fib_induct) show "?P 0" by simp show "?P 1" by simp @@ -93,7 +93,7 @@ assume "0 < n" then have "gcd (n * k + m) n = gcd n (m mod n)" by (simp add: gcd_non_0_nat add.commute) - also from `0 < n` have "\ = gcd m n" + also from \0 < n\ have "\ = gcd m n" by (simp add: gcd_non_0_nat) finally show ?thesis . qed @@ -124,7 +124,7 @@ proof - have "gcd (fib m) (fib (n - m)) = gcd (fib m) (fib (n - m + m))" by (simp add: gcd_fib_add) - also from `m \ n` have "n - m + m = n" + also from \m \ n\ have "n - m + m = n" by simp finally show ?thesis . qed @@ -145,12 +145,12 @@ next case False then have "m \ n" by simp - from `0 < m` and False have "n - m < n" + from \0 < m\ and False have "n - m < n" by simp with hyp have "gcd (fib m) (fib ((n - m) mod m)) = gcd (fib m) (fib (n - m))" by simp also have "\ = gcd (fib m) (fib n)" - using `m \ n` by (rule gcd_fib_diff) + using \m \ n\ by (rule gcd_fib_diff) finally have "gcd (fib m) (fib ((n - m) mod m)) = gcd (fib m) (fib n)" . with False show ?thesis by simp