| 1478 |      1 | (*  Title:      ZF/trancl.thy
 | 
| 0 |      2 |     ID:         $Id$
 | 
| 1478 |      3 |     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 | 
| 0 |      4 |     Copyright   1992  University of Cambridge
 | 
|  |      5 | 
 | 
|  |      6 | Transitive closure of a relation
 | 
|  |      7 | *)
 | 
|  |      8 | 
 | 
| 2469 |      9 | Trancl = Fixedpt + Perm + mono + Rel + 
 | 
| 0 |     10 | consts
 | 
| 1401 |     11 |     rtrancl :: i=>i  ("(_^*)" [100] 100)  (*refl/transitive closure*)
 | 
|  |     12 |     trancl  :: i=>i  ("(_^+)" [100] 100)  (*transitive closure*)
 | 
| 0 |     13 | 
 | 
| 753 |     14 | defs
 | 
| 1478 |     15 |     rtrancl_def "r^* == lfp(field(r)*field(r), %s. id(field(r)) Un (r O s))"
 | 
| 0 |     16 |     trancl_def  "r^+ == r O r^*"
 | 
|  |     17 | end
 |