| author | wenzelm |
| Fri, 03 Nov 2000 21:27:06 +0100 | |
| changeset 10379 | 93630e0c5ae9 |
| parent 10147 | 178deaacb244 |
| child 11049 | 7eef34adb852 |
| 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 = Primes + consts fib :: "nat => nat" recdef fib "less_than" zero "fib 0 = 0" one "fib 1 = 1" Suc_Suc "fib (Suc (Suc x)) = fib x + fib (Suc x)" end