| author | wenzelm | 
| Mon, 26 Jun 2000 00:00:40 +0200 | |
| changeset 9143 | 6180c29d2db6 | 
| parent 6906 | 46652582f831 | 
| permissions | -rw-r--r-- | 
| 1475 | 1  | 
(* Title: HOL/Trancl.thy  | 
| 923 | 2  | 
ID: $Id$  | 
| 1475 | 3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 923 | 4  | 
Copyright 1992 University of Cambridge  | 
5  | 
||
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents: 
972 
diff
changeset
 | 
6  | 
Relfexive and Transitive closure of a relation  | 
| 923 | 7  | 
|
| 1642 | 8  | 
rtrancl is reflexive/transitive closure;  | 
9  | 
trancl is transitive closure  | 
|
10  | 
reflcl is reflexive closure  | 
|
| 
1916
 
43521f79f016
Changed precedences to remove ambiguities in r^+ notation
 
paulson 
parents: 
1642 
diff
changeset
 | 
11  | 
|
| 
 
43521f79f016
Changed precedences to remove ambiguities in r^+ notation
 
paulson 
parents: 
1642 
diff
changeset
 | 
12  | 
These postfix operators have MAXIMUM PRIORITY, forcing their operands to be  | 
| 
 
43521f79f016
Changed precedences to remove ambiguities in r^+ notation
 
paulson 
parents: 
1642 
diff
changeset
 | 
13  | 
atomic.  | 
| 923 | 14  | 
*)  | 
15  | 
||
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents: 
972 
diff
changeset
 | 
16  | 
Trancl = Lfp + Relation +  | 
| 1558 | 17  | 
|
18  | 
constdefs  | 
|
| 
3499
 
ce1664057431
Reduced priority of postfix ^* etc operators such that they are the same as
 
nipkow 
parents: 
1916 
diff
changeset
 | 
19  | 
  rtrancl :: "('a * 'a)set => ('a * 'a)set"   ("(_^*)" [1000] 999)
 | 
| 5608 | 20  | 
"r^* == lfp(%s. Id Un (r O s))"  | 
| 1558 | 21  | 
|
| 
3499
 
ce1664057431
Reduced priority of postfix ^* etc operators such that they are the same as
 
nipkow 
parents: 
1916 
diff
changeset
 | 
22  | 
  trancl  :: "('a * 'a)set => ('a * 'a)set"   ("(_^+)" [1000] 999)
 | 
| 1558 | 23  | 
"r^+ == r O rtrancl(r)"  | 
24  | 
||
| 1301 | 25  | 
syntax  | 
| 6906 | 26  | 
  "_reflcl"  :: "('a*'a)set => ('a*'a)set"       ("(_^=)" [1000] 999)
 | 
| 1558 | 27  | 
|
| 1301 | 28  | 
translations  | 
| 5608 | 29  | 
"r^=" == "r Un Id"  | 
| 1558 | 30  | 
|
| 923 | 31  | 
end  |