| author | paulson |
| Fri, 17 Oct 1997 11:00:00 +0200 | |
| changeset 3912 | 4ed64ad7fb42 |
| parent 3494 | f7ac2d1e2051 |
| child 4809 | 595f905cc348 |
| 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 = WF_Rel + Divides + consts fib :: "nat => nat" recdef fib "less_than" "fib 0 = 0" "fib 1 = 1" "fib (Suc(Suc x)) = (fib x + fib (Suc x))" end