author | wenzelm |
Fri, 13 Dec 1996 17:50:04 +0100 | |
changeset 2390 | 4e183a4d9cd0 |
parent 1478 | 2b8c2a7547ab |
child 2469 | b50b8c0eec01 |
permissions | -rw-r--r-- |
(* Title: ZF/epsilon.thy ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge Epsilon induction and recursion *) Epsilon = Nat + "mono" + consts eclose,rank :: i=>i transrec :: [i, [i,i]=>i] =>i defs eclose_def "eclose(A) == UN n:nat. nat_rec(n, A, %m r. Union(r))" transrec_def "transrec(a,H) == wfrec(Memrel(eclose({a})), a, H)" rank_def "rank(a) == transrec(a, %x f. UN y:x. succ(f`y))" end