10213
|
1 |
(* Title: HOL/Transitive_Closure.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
|
4 |
Copyright 1992 University of Cambridge
|
|
5 |
|
|
6 |
Relfexive and Transitive closure of a relation
|
|
7 |
|
|
8 |
rtrancl is reflexive/transitive closure;
|
|
9 |
trancl is transitive closure
|
|
10 |
reflcl is reflexive closure
|
|
11 |
|
10331
|
12 |
These postfix operators have MAXIMUM PRIORITY, forcing their operands
|
|
13 |
to be atomic.
|
10213
|
14 |
*)
|
|
15 |
|
|
16 |
Transitive_Closure = Lfp + Relation +
|
|
17 |
|
|
18 |
constdefs
|
10331
|
19 |
rtrancl :: "('a * 'a) set => ('a * 'a) set" ("(_^*)" [1000] 999)
|
|
20 |
"r^* == lfp (%s. Id Un (r O s))"
|
10213
|
21 |
|
10331
|
22 |
trancl :: "('a * 'a) set => ('a * 'a) set" ("(_^+)" [1000] 999)
|
|
23 |
"r^+ == r O rtrancl r"
|
10213
|
24 |
|
|
25 |
syntax
|
10331
|
26 |
"_reflcl" :: "('a * 'a) set => ('a * 'a) set" ("(_^=)" [1000] 999)
|
10213
|
27 |
translations
|
|
28 |
"r^=" == "r Un Id"
|
|
29 |
|
10331
|
30 |
syntax (xsymbols)
|
|
31 |
rtrancl :: "('a * 'a) set => ('a * 'a) set" ("(_\\<^sup>*)" [1000] 999)
|
|
32 |
trancl :: "('a * 'a) set => ('a * 'a) set" ("(_\\<^sup>+)" [1000] 999)
|
|
33 |
"_reflcl" :: "('a * 'a) set => ('a * 'a) set" ("(_\\<^sup>=)" [1000] 999)
|
|
34 |
|
10213
|
35 |
end
|