author | wenzelm |
Wed, 18 May 2005 11:31:00 +0200 | |
changeset 16003 | 48ae07a95c70 |
parent 15927 | db77bed00211 |
child 16445 | bc90e58bb6ac |
permissions | -rw-r--r-- |
6134 | 1 |
(* Title: Pure/General/graph.ML |
2 |
ID: $Id$ |
|
15759 | 3 |
Author: Markus Wenzel and Stefan Berghofer, TU Muenchen |
6134 | 4 |
|
5 |
Directed graphs. |
|
6 |
*) |
|
7 |
||
8 |
signature GRAPH = |
|
9 |
sig |
|
10 |
type key |
|
11 |
type 'a T |
|
9321 | 12 |
exception UNDEF of key |
13 |
exception DUP of key |
|
14 |
exception DUPS of key list |
|
6134 | 15 |
val empty: 'a T |
6659 | 16 |
val keys: 'a T -> key list |
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
17 |
val dest: 'a T -> (key * key list) list |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
18 |
val minimals: 'a T -> key list |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
19 |
val maximals: 'a T -> key list |
6142 | 20 |
val map_nodes: ('a -> 'b) -> 'a T -> 'b T |
15759 | 21 |
val get_node: 'a T -> key -> 'a (*exception UNDEF*) |
6142 | 22 |
val map_node: key -> ('a -> 'a) -> 'a T -> 'a T |
23 |
val imm_preds: 'a T -> key -> key list |
|
24 |
val imm_succs: 'a T -> key -> key list |
|
6134 | 25 |
val all_preds: 'a T -> key list -> key list |
26 |
val all_succs: 'a T -> key list -> key list |
|
14161
73ad4884441f
Added function strong_conn for computing the strongly connected components
berghofe
parents:
12451
diff
changeset
|
27 |
val strong_conn: 'a T -> key list list |
6134 | 28 |
val find_paths: 'a T -> key * key -> key list list |
15759 | 29 |
val new_node: key * 'a -> 'a T -> 'a T (*exception DUP*) |
30 |
val del_nodes: key list -> 'a T -> 'a T (*exception UNDEF*) |
|
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
31 |
val is_edge: 'a T -> key * key -> bool |
6134 | 32 |
val add_edge: key * key -> 'a T -> 'a T |
6152 | 33 |
val del_edge: key * key -> 'a T -> 'a T |
15759 | 34 |
val merge: ('a * 'a -> bool) -> 'a T * 'a T -> 'a T (*exception DUPS*) |
6142 | 35 |
exception CYCLES of key list list |
15759 | 36 |
val add_edge_acyclic: key * key -> 'a T -> 'a T (*exception CYCLES*) |
37 |
val add_deps_acyclic: key * key list -> 'a T -> 'a T (*exception CYCLES*) |
|
38 |
val merge_acyclic: ('a * 'a -> bool) -> 'a T * 'a T -> 'a T (*exception CYCLES*) |
|
39 |
val add_edge_trans_acyclic: key * key -> 'a T -> 'a T (*exception CYCLES*) |
|
40 |
val merge_trans_acyclic: ('a * 'a -> bool) -> 'a T * 'a T -> 'a T (*exception CYCLES*) |
|
6134 | 41 |
end; |
42 |
||
43 |
functor GraphFun(Key: KEY): GRAPH = |
|
44 |
struct |
|
45 |
||
46 |
(* keys *) |
|
47 |
||
48 |
type key = Key.key; |
|
49 |
||
50 |
val eq_key = equal EQUAL o Key.ord; |
|
51 |
||
52 |
infix mem_key; |
|
53 |
val op mem_key = gen_mem eq_key; |
|
54 |
||
15759 | 55 |
val remove_key = remove eq_key; |
6152 | 56 |
|
6134 | 57 |
|
58 |
(* tables and sets of keys *) |
|
59 |
||
60 |
structure Table = TableFun(Key); |
|
61 |
type keys = unit Table.table; |
|
62 |
||
6142 | 63 |
val empty_keys = Table.empty: keys; |
64 |
||
6134 | 65 |
infix mem_keys; |
16003
48ae07a95c70
removed update_node, which is just an instance of map_node;
wenzelm
parents:
15927
diff
changeset
|
66 |
fun x mem_keys tab = is_some (Table.lookup (tab: keys, x)); |
6134 | 67 |
|
68 |
infix ins_keys; |
|
15759 | 69 |
fun x ins_keys tab = Table.insert (K true) (x, ()) (tab: keys); |
6134 | 70 |
|
71 |
||
6142 | 72 |
(* graphs *) |
6134 | 73 |
|
74 |
datatype 'a T = Graph of ('a * (key list * key list)) Table.table; |
|
75 |
||
9321 | 76 |
exception UNDEF of key; |
77 |
exception DUP = Table.DUP; |
|
78 |
exception DUPS = Table.DUPS; |
|
6134 | 79 |
|
80 |
val empty = Graph Table.empty; |
|
6659 | 81 |
fun keys (Graph tab) = Table.keys tab; |
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
82 |
fun dest (Graph tab) = map (fn (x, (_, (_, succs))) => (x, succs)) (Table.dest tab); |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
83 |
|
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
84 |
fun minimals (Graph tab) = |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
85 |
Table.foldl (fn (ms, (m, (_, ([], _)))) => m :: ms | (ms, _) => ms) ([], tab); |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
86 |
fun maximals (Graph tab) = |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
87 |
Table.foldl (fn (ms, (m, (_, (_, [])))) => m :: ms | (ms, _) => ms) ([], tab); |
6134 | 88 |
|
6142 | 89 |
fun get_entry (Graph tab) x = |
6134 | 90 |
(case Table.lookup (tab, x) of |
15531 | 91 |
SOME entry => entry |
92 |
| NONE => raise UNDEF x); |
|
6134 | 93 |
|
6142 | 94 |
fun map_entry x f (G as Graph tab) = Graph (Table.update ((x, f (get_entry G x)), tab)); |
6134 | 95 |
|
96 |
||
6142 | 97 |
(* nodes *) |
98 |
||
99 |
fun map_nodes f (Graph tab) = Graph (Table.map (fn (i, ps) => (f i, ps)) tab); |
|
6134 | 100 |
|
6142 | 101 |
fun get_node G = #1 o get_entry G; |
102 |
fun map_node x f = map_entry x (fn (i, ps) => (f i, ps)); |
|
103 |
||
104 |
||
105 |
(* reachability *) |
|
106 |
||
6659 | 107 |
(*nodes reachable from xs -- topologically sorted for acyclic graphs*) |
6142 | 108 |
fun reachable next xs = |
6134 | 109 |
let |
14161
73ad4884441f
Added function strong_conn for computing the strongly connected components
berghofe
parents:
12451
diff
changeset
|
110 |
fun reach ((R, rs), x) = |
73ad4884441f
Added function strong_conn for computing the strongly connected components
berghofe
parents:
12451
diff
changeset
|
111 |
if x mem_keys R then (R, rs) |
73ad4884441f
Added function strong_conn for computing the strongly connected components
berghofe
parents:
12451
diff
changeset
|
112 |
else apsnd (cons x) (reachs ((x ins_keys R, rs), next x)) |
15570 | 113 |
and reachs R_xs = Library.foldl reach R_xs; |
14161
73ad4884441f
Added function strong_conn for computing the strongly connected components
berghofe
parents:
12451
diff
changeset
|
114 |
in foldl_map (reach o apfst (rpair [])) (empty_keys, xs) end; |
6134 | 115 |
|
6142 | 116 |
(*immediate*) |
117 |
fun imm_preds G = #1 o #2 o get_entry G; |
|
118 |
fun imm_succs G = #2 o #2 o get_entry G; |
|
6134 | 119 |
|
6142 | 120 |
(*transitive*) |
15570 | 121 |
fun all_preds G = List.concat o snd o reachable (imm_preds G); |
122 |
fun all_succs G = List.concat o snd o reachable (imm_succs G); |
|
14161
73ad4884441f
Added function strong_conn for computing the strongly connected components
berghofe
parents:
12451
diff
changeset
|
123 |
|
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
124 |
(*strongly connected components; see: David King and John Launchbury, |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
125 |
"Structuring Depth First Search Algorithms in Haskell"*) |
14161
73ad4884441f
Added function strong_conn for computing the strongly connected components
berghofe
parents:
12451
diff
changeset
|
126 |
fun strong_conn G = filter_out null (snd (reachable (imm_preds G) |
15570 | 127 |
(List.concat (rev (snd (reachable (imm_succs G) (keys G))))))); |
6134 | 128 |
|
129 |
||
6142 | 130 |
(* paths *) |
6134 | 131 |
|
132 |
fun find_paths G (x, y) = |
|
133 |
let |
|
14161
73ad4884441f
Added function strong_conn for computing the strongly connected components
berghofe
parents:
12451
diff
changeset
|
134 |
val (X, _) = reachable (imm_succs G) [x]; |
6134 | 135 |
fun paths ps p = |
12451 | 136 |
if not (null ps) andalso eq_key (p, x) then [p :: ps] |
137 |
else if p mem_keys X andalso not (p mem_key ps) |
|
15570 | 138 |
then List.concat (map (paths (p :: ps)) (imm_preds G p)) |
12451 | 139 |
else []; |
140 |
in paths [] y end; |
|
6134 | 141 |
|
142 |
||
9321 | 143 |
(* nodes *) |
6134 | 144 |
|
6152 | 145 |
fun new_node (x, info) (Graph tab) = |
9321 | 146 |
Graph (Table.update_new ((x, (info, ([], []))), tab)); |
6134 | 147 |
|
6659 | 148 |
fun del_nodes xs (Graph tab) = |
15759 | 149 |
Graph (tab |
150 |
|> fold Table.delete xs |
|
151 |
|> Table.map (fn (i, (preds, succs)) => |
|
152 |
(i, (fold remove_key xs preds, fold remove_key xs succs)))); |
|
6659 | 153 |
|
6152 | 154 |
|
9321 | 155 |
(* edges *) |
156 |
||
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
157 |
fun is_edge G (x, y) = y mem_key imm_succs G x handle UNDEF _ => false; |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
158 |
|
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
159 |
fun add_edge (x, y) G = |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
160 |
if is_edge G (x, y) then G |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
161 |
else |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
162 |
G |> map_entry y (fn (i, (preds, succs)) => (i, (x :: preds, succs))) |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
163 |
|> map_entry x (fn (i, (preds, succs)) => (i, (preds, y :: succs))); |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
164 |
|
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
165 |
fun del_edge (x, y) G = |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
166 |
if is_edge G (x, y) then |
15759 | 167 |
G |> map_entry y (fn (i, (preds, succs)) => (i, (remove_key x preds, succs))) |
168 |
|> map_entry x (fn (i, (preds, succs)) => (i, (preds, remove_key y succs))) |
|
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
169 |
else G; |
9321 | 170 |
|
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
171 |
fun diff_edges G1 G2 = |
15570 | 172 |
List.concat (dest G1 |> map (fn (x, ys) => ys |> List.mapPartial (fn y => |
15531 | 173 |
if is_edge G2 (x, y) then NONE else SOME (x, y)))); |
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
174 |
|
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
175 |
fun edges G = diff_edges G empty; |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
176 |
|
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
177 |
|
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
178 |
(* merge *) |
6152 | 179 |
|
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
180 |
fun gen_merge add eq (Graph tab1, G2 as Graph tab2) = |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
181 |
let |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
182 |
fun eq_node ((i1, _), (i2, _)) = eq (i1, i2); |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
183 |
fun no_edges (i, _) = (i, ([], [])); |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
184 |
in fold add (edges G2) (Graph (Table.merge eq_node (tab1, Table.map no_edges tab2))) end; |
6152 | 185 |
|
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
186 |
fun merge eq GG = gen_merge add_edge eq GG; |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
187 |
|
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
188 |
|
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
189 |
(* maintain acyclic graphs *) |
6142 | 190 |
|
191 |
exception CYCLES of key list list; |
|
6134 | 192 |
|
193 |
fun add_edge_acyclic (x, y) G = |
|
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
194 |
if is_edge G (x, y) then G |
9347 | 195 |
else |
196 |
(case find_paths G (y, x) of |
|
197 |
[] => add_edge (x, y) G |
|
198 |
| cycles => raise CYCLES (map (cons x) cycles)); |
|
6134 | 199 |
|
15759 | 200 |
fun add_deps_acyclic (y, xs) = fold (fn x => add_edge_acyclic (x, y)) xs; |
9321 | 201 |
|
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
202 |
fun merge_acyclic eq GG = gen_merge add_edge_acyclic eq GG; |
9321 | 203 |
|
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
204 |
|
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
205 |
(* maintain transitive acyclic graphs *) |
9321 | 206 |
|
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
207 |
fun add_edge_trans_acyclic (x, y) G = |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
208 |
add_edge_acyclic (x, y) G |> |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
209 |
fold add_edge (Library.product (all_preds G [x]) (all_succs G [y])); |
9321 | 210 |
|
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
211 |
fun merge_trans_acyclic eq (G1, G2) = |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
212 |
merge_acyclic eq (G1, G2) |> |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
213 |
fold add_edge_trans_acyclic (diff_edges G1 G2 @ diff_edges G2 G1); |
6134 | 214 |
|
215 |
end; |
|
216 |
||
217 |
||
218 |
(*graphs indexed by strings*) |
|
219 |
structure Graph = GraphFun(type key = string val ord = string_ord); |