| 1478 |      1 | (*  Title:      ZF/wf.thy
 | 
| 0 |      2 |     ID:         $Id$
 | 
| 1478 |      3 |     Author:     Tobias Nipkow and Lawrence C Paulson
 | 
| 435 |      4 |     Copyright   1994  University of Cambridge
 | 
| 0 |      5 | 
 | 
|  |      6 | Well-founded Recursion
 | 
|  |      7 | *)
 | 
|  |      8 | 
 | 
| 2469 |      9 | WF = Trancl + mono + equalities +
 | 
| 0 |     10 | consts
 | 
| 1401 |     11 |   wf           :: i=>o
 | 
| 1478 |     12 |   wf_on        :: [i,i]=>o                      ("wf[_]'(_')")
 | 
| 435 |     13 | 
 | 
| 1401 |     14 |   wftrec,wfrec :: [i, i, [i,i]=>i] =>i
 | 
| 1478 |     15 |   wfrec_on     :: [i, i, i, [i,i]=>i] =>i       ("wfrec[_]'(_,_,_')")
 | 
| 1401 |     16 |   is_recfun    :: [i, i, [i,i]=>i, i] =>o
 | 
|  |     17 |   the_recfun   :: [i, i, [i,i]=>i] =>i
 | 
| 0 |     18 | 
 | 
| 930 |     19 | defs
 | 
| 0 |     20 |   (*r is a well-founded relation*)
 | 
| 1478 |     21 |   wf_def         "wf(r) == ALL Z. Z=0 | (EX x:Z. ALL y. <y,x>:r --> ~ y:Z)"
 | 
| 0 |     22 | 
 | 
| 435 |     23 |   (*r is well-founded relation over A*)
 | 
|  |     24 |   wf_on_def      "wf_on(A,r) == wf(r Int A*A)"
 | 
|  |     25 | 
 | 
| 1155 |     26 |   is_recfun_def  "is_recfun(r,a,H,f) == 
 | 
| 1478 |     27 |                         (f = (lam x: r-``{a}. H(x, restrict(f, r-``{x}))))"
 | 
| 0 |     28 | 
 | 
|  |     29 |   the_recfun_def "the_recfun(r,a,H) == (THE f.is_recfun(r,a,H,f))"
 | 
|  |     30 | 
 | 
| 1478 |     31 |   wftrec_def     "wftrec(r,a,H) == H(a, the_recfun(r,a,H))"
 | 
| 0 |     32 | 
 | 
|  |     33 |   (*public version.  Does not require r to be transitive*)
 | 
|  |     34 |   wfrec_def "wfrec(r,a,H) == wftrec(r^+, a, %x f. H(x, restrict(f,r-``{x})))"
 | 
|  |     35 | 
 | 
| 435 |     36 |   wfrec_on_def   "wfrec[A](r,a,H) == wfrec(r Int A*A, a, H)"
 | 
|  |     37 | 
 | 
| 0 |     38 | end
 |