author | nipkow |
Fri, 28 Jul 1995 18:05:50 +0200 | |
changeset 1212 | 7059356e18e0 |
parent 753 | ec86863e87c8 |
child 1401 | 0c439768f45c |
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