src/ZF/Nat.thy
author lcp
Tue, 29 Nov 1994 00:31:31 +0100
changeset 753 ec86863e87c8
parent 435 ca5356bd315a
child 1155 928a16e02f9f
permissions -rw-r--r--
replaced "rules" by "defs"

(*  Title: 	ZF/Nat.thy
    ID:         $Id$
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1994  University of Cambridge

Natural numbers in Zermelo-Fraenkel Set Theory 
*)

Nat = Ordinal + Bool + "mono" +
consts
    nat 	::      "i"
    nat_case    ::      "[i, i=>i, i]=>i"
    nat_rec     ::      "[i, i, [i,i]=>i]=>i"

defs

    nat_def     "nat == lfp(Inf, %X. {0} Un {succ(i). i:X})"

    nat_case_def
	"nat_case(a,b,k) == 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(a, %m. b(m, f`m), n))"

end