clasohm@0: (* Title: ZF/nat.thy clasohm@0: ID: $Id$ clasohm@0: Author: Lawrence C Paulson, Cambridge University Computer Laboratory clasohm@0: Copyright 1992 University of Cambridge clasohm@0: clasohm@0: Natural numbers in Zermelo-Fraenkel Set Theory clasohm@0: *) clasohm@0: clasohm@0: Nat = Ord + Bool + clasohm@0: consts clasohm@0: nat :: "i" clasohm@0: nat_case :: "[i, i, i=>i]=>i" clasohm@0: nat_rec :: "[i, i, [i,i]=>i]=>i" clasohm@0: clasohm@0: rules clasohm@0: clasohm@0: nat_def "nat == lfp(Inf, %X. {0} Un {succ(i). i:X})" clasohm@0: clasohm@0: nat_case_def clasohm@0: "nat_case(k,a,b) == THE y. k=0 & y=a | (EX x. k=succ(x) & y=b(x))" clasohm@0: clasohm@0: nat_rec_def clasohm@0: "nat_rec(k,a,b) == \ clasohm@0: \ wfrec(Memrel(nat), k, %n f. nat_case(n, a, %m. b(m, f`m)))" clasohm@0: clasohm@0: end