17456
|
1 |
(* Title: CCL/Trancl.thy
|
0
|
2 |
ID: $Id$
|
1474
|
3 |
Author: Martin Coen, Cambridge University Computer Laboratory
|
0
|
4 |
Copyright 1993 University of Cambridge
|
|
5 |
*)
|
|
6 |
|
17456
|
7 |
header {* Transitive closure of a relation *}
|
|
8 |
|
|
9 |
theory Trancl
|
|
10 |
imports CCL
|
|
11 |
begin
|
0
|
12 |
|
|
13 |
consts
|
17456
|
14 |
trans :: "i set => o" (*transitivity predicate*)
|
|
15 |
id :: "i set"
|
|
16 |
rtrancl :: "i set => i set" ("(_^*)" [100] 100)
|
|
17 |
trancl :: "i set => i set" ("(_^+)" [100] 100)
|
|
18 |
O :: "[i set,i set] => i set" (infixr 60)
|
0
|
19 |
|
17456
|
20 |
axioms
|
|
21 |
trans_def: "trans(r) == (ALL x y z. <x,y>:r --> <y,z>:r --> <x,z>:r)"
|
|
22 |
comp_def: (*composition of relations*)
|
|
23 |
"r O s == {xz. EX x y z. xz = <x,z> & <x,y>:s & <y,z>:r}"
|
|
24 |
id_def: (*the identity relation*)
|
|
25 |
"id == {p. EX x. p = <x,x>}"
|
|
26 |
rtrancl_def: "r^* == lfp(%s. id Un (r O s))"
|
|
27 |
trancl_def: "r^+ == r O rtrancl(r)"
|
|
28 |
|
|
29 |
ML {* use_legacy_bindings (the_context ()) *}
|
0
|
30 |
|
|
31 |
end
|