author | wenzelm |
Mon, 28 Jun 1999 21:38:50 +0200 | |
changeset 6845 | 598d2f32d452 |
parent 6481 | dbf2d9b3d6c8 |
child 8658 | 3cf533397c5a |
permissions | -rw-r--r-- |
3300 | 1 |
(* Title: ex/Fib |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1997 University of Cambridge |
|
5 |
||
3494 | 6 |
The Fibonacci function. Demonstrates the use of recdef. |
3300 | 7 |
*) |
8 |
||
6481 | 9 |
Fib = Divides + Primes + |
3300 | 10 |
|
11 |
consts fib :: "nat => nat" |
|
12 |
recdef fib "less_than" |
|
13 |
"fib 0 = 0" |
|
14 |
"fib 1 = 1" |
|
15 |
"fib (Suc(Suc x)) = (fib x + fib (Suc x))" |
|
16 |
||
17 |
end |