author | wenzelm |
Mon, 29 Nov 1999 15:52:49 +0100 | |
changeset 8039 | a901bafe4578 |
parent 6481 | dbf2d9b3d6c8 |
child 8658 | 3cf533397c5a |
permissions | -rw-r--r-- |
(* Title: ex/Fib ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1997 University of Cambridge The Fibonacci function. Demonstrates the use of recdef. *) Fib = Divides + Primes + consts fib :: "nat => nat" recdef fib "less_than" "fib 0 = 0" "fib 1 = 1" "fib (Suc(Suc x)) = (fib x + fib (Suc x))" end