src/ZF/Epsilon.thy
author paulson
Mon, 23 Sep 1996 18:18:18 +0200
changeset 2010 0a22b9d63a18
parent 1478 2b8c2a7547ab
child 2469 b50b8c0eec01
permissions -rw-r--r--
Simplification of definition of synth
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/epsilon.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
Epsilon induction and recursion
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
124
858ab9a9b047 made pseudo theories for all ML files;
clasohm
parents: 0
diff changeset
     9
Epsilon = Nat + "mono" +
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
consts
1401
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 753
diff changeset
    11
    eclose,rank ::      i=>i
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 753
diff changeset
    12
    transrec    ::      [i, [i,i]=>i] =>i
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
753
ec86863e87c8 replaced "rules" by "defs"
lcp
parents: 124
diff changeset
    14
defs
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    15
  eclose_def    "eclose(A) == UN n:nat. nat_rec(n, A, %m r. Union(r))"
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    16
  transrec_def  "transrec(a,H) == wfrec(Memrel(eclose({a})), a, H)"
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    17
  rank_def      "rank(a) == transrec(a, %x f. UN y:x. succ(f`y))"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    18
end