author | paulson |
Tue, 16 Jul 1996 15:49:46 +0200 | |
changeset 1868 | 836950047d85 |
parent 1642 | 21db0cf9a1a4 |
child 1916 | 43521f79f016 |
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 |
|
923 | 11 |
*) |
12 |
||
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
972
diff
changeset
|
13 |
Trancl = Lfp + Relation + |
1558 | 14 |
|
15 |
constdefs |
|
16 |
rtrancl :: "('a * 'a)set => ('a * 'a)set" ("(_^*)" [100] 100) |
|
17 |
"r^* == lfp(%s. id Un (r O s))" |
|
18 |
||
19 |
trancl :: "('a * 'a)set => ('a * 'a)set" ("(_^+)" [100] 100) |
|
20 |
"r^+ == r O rtrancl(r)" |
|
21 |
||
1301 | 22 |
syntax |
1558 | 23 |
reflcl :: "('a*'a)set => ('a*'a)set" ("(_^=)" [100] 100) |
24 |
||
1301 | 25 |
translations |
1558 | 26 |
"r^=" == "r Un id" |
27 |
||
923 | 28 |
end |