diff -r 000000000000 -r a5a9c433f639 src/ZF/nat.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/nat.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,26 @@ +(* Title: ZF/nat.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Natural numbers in Zermelo-Fraenkel Set Theory +*) + +Nat = Ord + Bool + +consts + nat :: "i" + nat_case :: "[i, i, i=>i]=>i" + nat_rec :: "[i, i, [i,i]=>i]=>i" + +rules + + nat_def "nat == lfp(Inf, %X. {0} Un {succ(i). i:X})" + + nat_case_def + "nat_case(k,a,b) == THE y. k=0 & y=a | (EX x. k=succ(x) & y=b(x))" + + nat_rec_def + "nat_rec(k,a,b) == \ +\ wfrec(Memrel(nat), k, %n f. nat_case(n, a, %m. b(m, f`m)))" + +end