author | wenzelm |
Tue, 21 Jun 2022 14:51:17 +0200 | |
changeset 75571 | ac5e633ad9b3 |
parent 70773 | 60abd1e94168 |
child 77723 | b761c91c2447 |
permissions | -rw-r--r-- |
6134 | 1 |
(* Title: Pure/General/graph.ML |
15759 | 2 |
Author: Markus Wenzel and Stefan Berghofer, TU Muenchen |
6134 | 3 |
|
4 |
Directed graphs. |
|
5 |
*) |
|
6 |
||
7 |
signature GRAPH = |
|
8 |
sig |
|
9 |
type key |
|
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
10 |
structure Keys: |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
11 |
sig |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
12 |
type T |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
13 |
val is_empty: T -> bool |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
14 |
val fold: (key -> 'a -> 'a) -> T -> 'a -> 'a |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
15 |
val fold_rev: (key -> 'a -> 'a) -> T -> 'a -> 'a |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
16 |
val dest: T -> key list |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
17 |
end |
6134 | 18 |
type 'a T |
9321 | 19 |
exception DUP of key |
19029 | 20 |
exception SAME |
21 |
exception UNDEF of key |
|
6134 | 22 |
val empty: 'a T |
28204 | 23 |
val is_empty: 'a T -> bool |
6659 | 24 |
val keys: 'a T -> key list |
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
25 |
val get_first: (key * ('a * (Keys.T * Keys.T)) -> 'b option) -> 'a T -> 'b option |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
26 |
val fold: (key * ('a * (Keys.T * Keys.T)) -> 'b -> 'b) -> 'a T -> 'b -> 'b |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
27 |
val get_entry: 'a T -> key -> key * ('a * (Keys.T * Keys.T)) (*exception UNDEF*) |
15759 | 28 |
val get_node: 'a T -> key -> 'a (*exception UNDEF*) |
6142 | 29 |
val map_node: key -> ('a -> 'a) -> 'a T -> 'a T |
46611 | 30 |
val map: (key -> 'a -> 'b) -> 'a T -> 'b T |
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
31 |
val imm_preds: 'a T -> key -> Keys.T |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
32 |
val imm_succs: 'a T -> key -> Keys.T |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
33 |
val immediate_preds: 'a T -> key -> key list |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
34 |
val immediate_succs: 'a T -> key -> key list |
6134 | 35 |
val all_preds: 'a T -> key list -> key list |
36 |
val all_succs: 'a T -> key list -> key list |
|
46613 | 37 |
val strong_conn: 'a T -> key list list |
55658
d696adf157e6
simultaneous mapping of strongly connected components
haftmann
parents:
55154
diff
changeset
|
38 |
val map_strong_conn: ((key * 'a) list -> 'b list) -> 'a T -> 'b T |
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
39 |
val minimals: 'a T -> key list |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
40 |
val maximals: 'a T -> key list |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
41 |
val is_minimal: 'a T -> key -> bool |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
42 |
val is_maximal: 'a T -> key -> bool |
66830 | 43 |
val is_isolated: 'a T -> key -> bool |
15759 | 44 |
val new_node: key * 'a -> 'a T -> 'a T (*exception DUP*) |
17179 | 45 |
val default_node: key * 'a -> 'a T -> 'a T |
28333
507b64f4cd2a
added del_node, which is more efficient for sparse graphs;
wenzelm
parents:
28204
diff
changeset
|
46 |
val del_node: key -> 'a T -> 'a T (*exception UNDEF*) |
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
47 |
val is_edge: 'a T -> key * key -> bool |
44202 | 48 |
val add_edge: key * key -> 'a T -> 'a T (*exception UNDEF*) |
49 |
val del_edge: key * key -> 'a T -> 'a T (*exception UNDEF*) |
|
46614
165886a4fe64
clarified Graph.restrict (formerly Graph.subgraph) based on public graph operations;
wenzelm
parents:
46613
diff
changeset
|
50 |
val restrict: (key -> bool) -> 'a T -> 'a T |
49560 | 51 |
val dest: 'a T -> ((key * 'a) * key list) list |
52 |
val make: ((key * 'a) * key list) list -> 'a T (*exception DUP | UNDEF*) |
|
23655
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents:
21935
diff
changeset
|
53 |
val merge: ('a * 'a -> bool) -> 'a T * 'a T -> 'a T (*exception DUP*) |
19029 | 54 |
val join: (key -> 'a * 'a -> 'a) (*exception DUP/SAME*) -> |
23655
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents:
21935
diff
changeset
|
55 |
'a T * 'a T -> 'a T (*exception DUP*) |
19580
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents:
19482
diff
changeset
|
56 |
val irreducible_paths: 'a T -> key * key -> key list list |
6142 | 57 |
exception CYCLES of key list list |
44202 | 58 |
val add_edge_acyclic: key * key -> 'a T -> 'a T (*exception UNDEF | CYCLES*) |
59 |
val add_deps_acyclic: key * key list -> 'a T -> 'a T (*exception UNDEF | CYCLES*) |
|
15759 | 60 |
val merge_acyclic: ('a * 'a -> bool) -> 'a T * 'a T -> 'a T (*exception CYCLES*) |
23964 | 61 |
val topological_order: 'a T -> key list |
44202 | 62 |
val add_edge_trans_acyclic: key * key -> 'a T -> 'a T (*exception UNDEF | CYCLES*) |
15759 | 63 |
val merge_trans_acyclic: ('a * 'a -> bool) -> 'a T * 'a T -> 'a T (*exception CYCLES*) |
44162 | 64 |
exception DEP of key * key |
65 |
val schedule: ((key * 'b) list -> key * 'a -> 'b) -> 'a T -> 'b list (*exception DEP*) |
|
49560 | 66 |
val encode: key XML.Encode.T -> 'a XML.Encode.T -> 'a T XML.Encode.T |
67 |
val decode: key XML.Decode.T -> 'a XML.Decode.T -> 'a T XML.Decode.T |
|
6134 | 68 |
end; |
69 |
||
31971
8c1b845ed105
renamed functor TableFun to Table, and GraphFun to Graph;
wenzelm
parents:
31616
diff
changeset
|
70 |
functor Graph(Key: KEY): GRAPH = |
6134 | 71 |
struct |
72 |
||
73 |
(* keys *) |
|
74 |
||
75 |
type key = Key.key; |
|
18970 | 76 |
val eq_key = is_equal o Key.ord; |
6134 | 77 |
|
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
78 |
structure Table = Table(Key); |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
79 |
|
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
80 |
structure Keys = |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
81 |
struct |
6152 | 82 |
|
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
83 |
abstype T = Keys of unit Table.table |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
84 |
with |
6134 | 85 |
|
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
86 |
val empty = Keys Table.empty; |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
87 |
fun is_empty (Keys tab) = Table.is_empty tab; |
6134 | 88 |
|
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
89 |
fun member (Keys tab) = Table.defined tab; |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
90 |
fun insert x (Keys tab) = Keys (Table.insert (K true) (x, ()) tab); |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
91 |
fun remove x (Keys tab) = Keys (Table.delete_safe x tab); |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
92 |
|
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
93 |
fun fold f (Keys tab) = Table.fold (f o #1) tab; |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
94 |
fun fold_rev f (Keys tab) = Table.fold_rev (f o #1) tab; |
6134 | 95 |
|
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
96 |
fun dest keys = fold_rev cons keys []; |
6142 | 97 |
|
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
98 |
fun filter P keys = fold (fn x => P x ? insert x) keys empty; |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
99 |
|
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
100 |
end; |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
101 |
end; |
6134 | 102 |
|
103 |
||
6142 | 104 |
(* graphs *) |
6134 | 105 |
|
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
106 |
datatype 'a T = Graph of ('a * (Keys.T * Keys.T)) Table.table; |
6134 | 107 |
|
9321 | 108 |
exception DUP = Table.DUP; |
19029 | 109 |
exception UNDEF = Table.UNDEF; |
110 |
exception SAME = Table.SAME; |
|
6134 | 111 |
|
112 |
val empty = Graph Table.empty; |
|
28204 | 113 |
fun is_empty (Graph tab) = Table.is_empty tab; |
6659 | 114 |
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
|
115 |
|
35012
c3e3ac3ca091
removed unused "boundary" of Table/Graph.get_first;
wenzelm
parents:
32710
diff
changeset
|
116 |
fun get_first f (Graph tab) = Table.get_first f tab; |
19615 | 117 |
fun fold_graph f (Graph tab) = Table.fold f tab; |
118 |
||
6142 | 119 |
fun get_entry (Graph tab) x = |
43792
d5803c3d537a
Table.lookup_key and Graph.get_entry allow to retrieve the original key, which is not necessarily identical to the given one;
wenzelm
parents:
39021
diff
changeset
|
120 |
(case Table.lookup_key tab x of |
15531 | 121 |
SOME entry => entry |
122 |
| NONE => raise UNDEF x); |
|
6134 | 123 |
|
43792
d5803c3d537a
Table.lookup_key and Graph.get_entry allow to retrieve the original key, which is not necessarily identical to the given one;
wenzelm
parents:
39021
diff
changeset
|
124 |
fun map_entry x f (G as Graph tab) = Graph (Table.update (x, f (#2 (get_entry G x))) tab); |
19290 | 125 |
|
6134 | 126 |
|
6142 | 127 |
(* nodes *) |
128 |
||
43792
d5803c3d537a
Table.lookup_key and Graph.get_entry allow to retrieve the original key, which is not necessarily identical to the given one;
wenzelm
parents:
39021
diff
changeset
|
129 |
fun get_node G = #1 o #2 o get_entry G; |
18133 | 130 |
|
6142 | 131 |
fun map_node x f = map_entry x (fn (i, ps) => (f i, ps)); |
19290 | 132 |
|
46611 | 133 |
fun map_nodes f (Graph tab) = Graph (Table.map (apfst o f) tab); |
134 |
||
18133 | 135 |
|
6142 | 136 |
(* reachability *) |
137 |
||
6659 | 138 |
(*nodes reachable from xs -- topologically sorted for acyclic graphs*) |
6142 | 139 |
fun reachable next xs = |
6134 | 140 |
let |
18006
535de280c812
reachable - abandoned foldl_map in favor of fold_map
haftmann
parents:
17912
diff
changeset
|
141 |
fun reach x (rs, R) = |
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
142 |
if Keys.member R x then (rs, R) |
48350
09bf3b73e446
clarified topological ordering: preserve order of adjacency via reverse fold;
wenzelm
parents:
46668
diff
changeset
|
143 |
else Keys.fold_rev reach (next x) (rs, Keys.insert x R) |>> cons x; |
32710 | 144 |
fun reachs x (rss, R) = |
145 |
reach x ([], R) |>> (fn rs => rs :: rss); |
|
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
146 |
in fold reachs xs ([], Keys.empty) end; |
6134 | 147 |
|
6142 | 148 |
(*immediate*) |
43792
d5803c3d537a
Table.lookup_key and Graph.get_entry allow to retrieve the original key, which is not necessarily identical to the given one;
wenzelm
parents:
39021
diff
changeset
|
149 |
fun imm_preds G = #1 o #2 o #2 o get_entry G; |
d5803c3d537a
Table.lookup_key and Graph.get_entry allow to retrieve the original key, which is not necessarily identical to the given one;
wenzelm
parents:
39021
diff
changeset
|
150 |
fun imm_succs G = #2 o #2 o #2 o get_entry G; |
6134 | 151 |
|
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
152 |
fun immediate_preds G = Keys.dest o imm_preds G; |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
153 |
fun immediate_succs G = Keys.dest o imm_succs G; |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
154 |
|
6142 | 155 |
(*transitive*) |
32710 | 156 |
fun all_preds G = flat o #1 o reachable (imm_preds G); |
157 |
fun all_succs G = flat o #1 o reachable (imm_succs G); |
|
14161
73ad4884441f
Added function strong_conn for computing the strongly connected components
berghofe
parents:
12451
diff
changeset
|
158 |
|
46613 | 159 |
(*strongly connected components; see: David King and John Launchbury, |
160 |
"Structuring Depth First Search Algorithms in Haskell"*) |
|
161 |
fun strong_conn G = |
|
162 |
rev (filter_out null (#1 (reachable (imm_preds G) (all_succs G (keys G))))); |
|
163 |
||
55658
d696adf157e6
simultaneous mapping of strongly connected components
haftmann
parents:
55154
diff
changeset
|
164 |
fun map_strong_conn f G = |
d696adf157e6
simultaneous mapping of strongly connected components
haftmann
parents:
55154
diff
changeset
|
165 |
let |
d696adf157e6
simultaneous mapping of strongly connected components
haftmann
parents:
55154
diff
changeset
|
166 |
val xss = strong_conn G; |
d696adf157e6
simultaneous mapping of strongly connected components
haftmann
parents:
55154
diff
changeset
|
167 |
fun map' xs = |
d696adf157e6
simultaneous mapping of strongly connected components
haftmann
parents:
55154
diff
changeset
|
168 |
fold2 (curry Table.update) xs (f (AList.make (get_node G) xs)); |
d696adf157e6
simultaneous mapping of strongly connected components
haftmann
parents:
55154
diff
changeset
|
169 |
val tab' = Table.empty |
d696adf157e6
simultaneous mapping of strongly connected components
haftmann
parents:
55154
diff
changeset
|
170 |
|> fold map' xss; |
d696adf157e6
simultaneous mapping of strongly connected components
haftmann
parents:
55154
diff
changeset
|
171 |
in map_nodes (fn x => fn _ => the (Table.lookup tab' x)) G end; |
d696adf157e6
simultaneous mapping of strongly connected components
haftmann
parents:
55154
diff
changeset
|
172 |
|
46613 | 173 |
|
174 |
(* minimal and maximal elements *) |
|
175 |
||
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
176 |
fun minimals G = fold_graph (fn (m, (_, (preds, _))) => Keys.is_empty preds ? cons m) G []; |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
177 |
fun maximals G = fold_graph (fn (m, (_, (_, succs))) => Keys.is_empty succs ? cons m) G []; |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
178 |
fun is_minimal G x = Keys.is_empty (imm_preds G x); |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
179 |
fun is_maximal G x = Keys.is_empty (imm_succs G x); |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
180 |
|
66830 | 181 |
fun is_isolated G x = is_minimal G x andalso is_maximal G x; |
182 |
||
18133 | 183 |
|
46668 | 184 |
(* node operations *) |
6134 | 185 |
|
6152 | 186 |
fun new_node (x, info) (Graph tab) = |
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
187 |
Graph (Table.update_new (x, (info, (Keys.empty, Keys.empty))) tab); |
6134 | 188 |
|
17179 | 189 |
fun default_node (x, info) (Graph tab) = |
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
190 |
Graph (Table.default (x, (info, (Keys.empty, Keys.empty))) tab); |
17140 | 191 |
|
28333
507b64f4cd2a
added del_node, which is more efficient for sparse graphs;
wenzelm
parents:
28204
diff
changeset
|
192 |
fun del_node x (G as Graph tab) = |
507b64f4cd2a
added del_node, which is more efficient for sparse graphs;
wenzelm
parents:
28204
diff
changeset
|
193 |
let |
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
194 |
fun del_adjacent which y = |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
195 |
Table.map_entry y (fn (i, ps) => (i, (which (Keys.remove x) ps))); |
43792
d5803c3d537a
Table.lookup_key and Graph.get_entry allow to retrieve the original key, which is not necessarily identical to the given one;
wenzelm
parents:
39021
diff
changeset
|
196 |
val (preds, succs) = #2 (#2 (get_entry G x)); |
28333
507b64f4cd2a
added del_node, which is more efficient for sparse graphs;
wenzelm
parents:
28204
diff
changeset
|
197 |
in |
507b64f4cd2a
added del_node, which is more efficient for sparse graphs;
wenzelm
parents:
28204
diff
changeset
|
198 |
Graph (tab |
507b64f4cd2a
added del_node, which is more efficient for sparse graphs;
wenzelm
parents:
28204
diff
changeset
|
199 |
|> Table.delete x |
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
200 |
|> Keys.fold (del_adjacent apsnd) preds |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
201 |
|> Keys.fold (del_adjacent apfst) succs) |
28333
507b64f4cd2a
added del_node, which is more efficient for sparse graphs;
wenzelm
parents:
28204
diff
changeset
|
202 |
end; |
507b64f4cd2a
added del_node, which is more efficient for sparse graphs;
wenzelm
parents:
28204
diff
changeset
|
203 |
|
46614
165886a4fe64
clarified Graph.restrict (formerly Graph.subgraph) based on public graph operations;
wenzelm
parents:
46613
diff
changeset
|
204 |
fun restrict pred G = |
165886a4fe64
clarified Graph.restrict (formerly Graph.subgraph) based on public graph operations;
wenzelm
parents:
46613
diff
changeset
|
205 |
fold_graph (fn (x, _) => not (pred x) ? del_node x) G G; |
165886a4fe64
clarified Graph.restrict (formerly Graph.subgraph) based on public graph operations;
wenzelm
parents:
46613
diff
changeset
|
206 |
|
6152 | 207 |
|
46668 | 208 |
(* edge operations *) |
9321 | 209 |
|
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
210 |
fun is_edge G (x, y) = Keys.member (imm_succs G x) y handle UNDEF _ => false; |
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
211 |
|
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
212 |
fun add_edge (x, y) G = |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
213 |
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
|
214 |
else |
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
215 |
G |> map_entry y (fn (i, (preds, succs)) => (i, (Keys.insert x preds, succs))) |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
216 |
|> map_entry x (fn (i, (preds, succs)) => (i, (preds, Keys.insert y succs))); |
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
217 |
|
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
218 |
fun del_edge (x, y) G = |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
219 |
if is_edge G (x, y) then |
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
220 |
G |> map_entry y (fn (i, (preds, succs)) => (i, (Keys.remove x preds, succs))) |
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
221 |
|> map_entry x (fn (i, (preds, succs)) => (i, (preds, Keys.remove y succs))) |
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
222 |
else G; |
9321 | 223 |
|
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
224 |
fun diff_edges G1 G2 = |
49560 | 225 |
fold_graph (fn (x, (_, (_, succs))) => |
226 |
Keys.fold (fn y => not (is_edge G2 (x, y)) ? cons (x, y)) succs) G1 []; |
|
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
227 |
|
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
228 |
fun edges G = diff_edges G empty; |
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
229 |
|
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
230 |
|
49560 | 231 |
(* dest and make *) |
232 |
||
233 |
fun dest G = fold_graph (fn (x, (i, (_, succs))) => cons ((x, i), Keys.dest succs)) G []; |
|
234 |
||
235 |
fun make entries = |
|
236 |
empty |
|
237 |
|> fold (new_node o fst) entries |
|
238 |
|> fold (fn ((x, _), ys) => fold (fn y => add_edge (x, y)) ys) entries; |
|
239 |
||
240 |
||
18126 | 241 |
(* join and merge *) |
242 |
||
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
243 |
fun no_edges (i, _) = (i, (Keys.empty, Keys.empty)); |
18133 | 244 |
|
35974
3a588b344749
low-level tuning for join/merge: ignore identical versions (SUBTLE CHANGE IN SEMANTICS);
wenzelm
parents:
35405
diff
changeset
|
245 |
fun join f (G1 as Graph tab1, G2 as Graph tab2) = |
3a588b344749
low-level tuning for join/merge: ignore identical versions (SUBTLE CHANGE IN SEMANTICS);
wenzelm
parents:
35405
diff
changeset
|
246 |
let fun join_node key ((i1, edges1), (i2, _)) = (f key (i1, i2), edges1) in |
3a588b344749
low-level tuning for join/merge: ignore identical versions (SUBTLE CHANGE IN SEMANTICS);
wenzelm
parents:
35405
diff
changeset
|
247 |
if pointer_eq (G1, G2) then G1 |
39020 | 248 |
else fold add_edge (edges G2) (Graph (Table.join join_node (tab1, Table.map (K no_edges) tab2))) |
35974
3a588b344749
low-level tuning for join/merge: ignore identical versions (SUBTLE CHANGE IN SEMANTICS);
wenzelm
parents:
35405
diff
changeset
|
249 |
end; |
6152 | 250 |
|
35974
3a588b344749
low-level tuning for join/merge: ignore identical versions (SUBTLE CHANGE IN SEMANTICS);
wenzelm
parents:
35405
diff
changeset
|
251 |
fun gen_merge add eq (G1 as Graph tab1, G2 as Graph tab2) = |
3a588b344749
low-level tuning for join/merge: ignore identical versions (SUBTLE CHANGE IN SEMANTICS);
wenzelm
parents:
35405
diff
changeset
|
252 |
let fun eq_node ((i1, _), (i2, _)) = eq (i1, i2) in |
3a588b344749
low-level tuning for join/merge: ignore identical versions (SUBTLE CHANGE IN SEMANTICS);
wenzelm
parents:
35405
diff
changeset
|
253 |
if pointer_eq (G1, G2) then G1 |
39020 | 254 |
else fold add (edges G2) (Graph (Table.merge eq_node (tab1, Table.map (K no_edges) tab2))) |
35974
3a588b344749
low-level tuning for join/merge: ignore identical versions (SUBTLE CHANGE IN SEMANTICS);
wenzelm
parents:
35405
diff
changeset
|
255 |
end; |
6152 | 256 |
|
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
257 |
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
|
258 |
|
18133 | 259 |
|
19580
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents:
19482
diff
changeset
|
260 |
(* irreducible paths -- Hasse diagram *) |
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents:
19482
diff
changeset
|
261 |
|
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents:
19482
diff
changeset
|
262 |
fun irreducible_preds G X path z = |
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents:
19482
diff
changeset
|
263 |
let |
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents:
19482
diff
changeset
|
264 |
fun red x x' = is_edge G (x, x') andalso not (eq_key (x', z)); |
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents:
19482
diff
changeset
|
265 |
fun irreds [] xs' = xs' |
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents:
19482
diff
changeset
|
266 |
| irreds (x :: xs) xs' = |
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
267 |
if not (Keys.member X x) orelse eq_key (x, z) orelse member eq_key path x orelse |
19580
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents:
19482
diff
changeset
|
268 |
exists (red x) xs orelse exists (red x) xs' |
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents:
19482
diff
changeset
|
269 |
then irreds xs xs' |
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents:
19482
diff
changeset
|
270 |
else irreds xs (x :: xs'); |
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
271 |
in irreds (immediate_preds G z) [] end; |
19580
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents:
19482
diff
changeset
|
272 |
|
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents:
19482
diff
changeset
|
273 |
fun irreducible_paths G (x, y) = |
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents:
19482
diff
changeset
|
274 |
let |
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents:
19482
diff
changeset
|
275 |
val (_, X) = reachable (imm_succs G) [x]; |
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents:
19482
diff
changeset
|
276 |
fun paths path z = |
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents:
19482
diff
changeset
|
277 |
if eq_key (x, z) then cons (z :: path) |
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents:
19482
diff
changeset
|
278 |
else fold (paths (z :: path)) (irreducible_preds G X path z); |
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents:
19482
diff
changeset
|
279 |
in if eq_key (x, y) andalso not (is_edge G (x, x)) then [[]] else paths [] y [] end; |
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents:
19482
diff
changeset
|
280 |
|
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents:
19482
diff
changeset
|
281 |
|
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
282 |
(* maintain acyclic graphs *) |
6142 | 283 |
|
284 |
exception CYCLES of key list list; |
|
6134 | 285 |
|
286 |
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
|
287 |
if is_edge G (x, y) then G |
9347 | 288 |
else |
19580
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents:
19482
diff
changeset
|
289 |
(case irreducible_paths G (y, x) of |
9347 | 290 |
[] => add_edge (x, y) G |
291 |
| cycles => raise CYCLES (map (cons x) cycles)); |
|
6134 | 292 |
|
15759 | 293 |
fun add_deps_acyclic (y, xs) = fold (fn x => add_edge_acyclic (x, y)) xs; |
9321 | 294 |
|
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
295 |
fun merge_acyclic eq GG = gen_merge add_edge_acyclic eq GG; |
9321 | 296 |
|
23964 | 297 |
fun topological_order G = minimals G |> all_succs G; |
298 |
||
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
299 |
|
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
300 |
(* maintain transitive acyclic graphs *) |
9321 | 301 |
|
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
302 |
fun add_edge_trans_acyclic (x, y) G = |
19290 | 303 |
add_edge_acyclic (x, y) G |
25538 | 304 |
|> fold_product (curry add_edge) (all_preds G [x]) (all_succs G [y]); |
9321 | 305 |
|
14793
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents:
14161
diff
changeset
|
306 |
fun merge_trans_acyclic eq (G1, G2) = |
35974
3a588b344749
low-level tuning for join/merge: ignore identical versions (SUBTLE CHANGE IN SEMANTICS);
wenzelm
parents:
35405
diff
changeset
|
307 |
if pointer_eq (G1, G2) then G1 |
3a588b344749
low-level tuning for join/merge: ignore identical versions (SUBTLE CHANGE IN SEMANTICS);
wenzelm
parents:
35405
diff
changeset
|
308 |
else |
3a588b344749
low-level tuning for join/merge: ignore identical versions (SUBTLE CHANGE IN SEMANTICS);
wenzelm
parents:
35405
diff
changeset
|
309 |
merge_acyclic eq (G1, G2) |
3a588b344749
low-level tuning for join/merge: ignore identical versions (SUBTLE CHANGE IN SEMANTICS);
wenzelm
parents:
35405
diff
changeset
|
310 |
|> fold add_edge_trans_acyclic (diff_edges G1 G2) |
3a588b344749
low-level tuning for join/merge: ignore identical versions (SUBTLE CHANGE IN SEMANTICS);
wenzelm
parents:
35405
diff
changeset
|
311 |
|> fold add_edge_trans_acyclic (diff_edges G2 G1); |
6134 | 312 |
|
31540 | 313 |
|
44162 | 314 |
(* schedule acyclic graph *) |
315 |
||
316 |
exception DEP of key * key; |
|
317 |
||
318 |
fun schedule f G = |
|
319 |
let |
|
320 |
val xs = topological_order G; |
|
321 |
val results = (xs, Table.empty) |-> fold (fn x => fn tab => |
|
322 |
let |
|
323 |
val a = get_node G x; |
|
44338
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents:
44202
diff
changeset
|
324 |
val deps = immediate_preds G x |> map (fn y => |
44162 | 325 |
(case Table.lookup tab y of |
326 |
SOME b => (y, b) |
|
327 |
| NONE => raise DEP (x, y))); |
|
328 |
in Table.update (x, f deps (x, a)) tab end); |
|
329 |
in map (the o Table.lookup results) xs end; |
|
330 |
||
331 |
||
49560 | 332 |
(* XML data representation *) |
333 |
||
334 |
fun encode key info G = |
|
335 |
dest G |> |
|
336 |
let open XML.Encode |
|
337 |
in list (pair (pair key info) (list key)) end; |
|
338 |
||
339 |
fun decode key info body = |
|
340 |
body |> |
|
341 |
let open XML.Decode |
|
342 |
in list (pair (pair key info) (list key)) end |> make; |
|
343 |
||
344 |
||
19615 | 345 |
(*final declarations of this structure!*) |
39021 | 346 |
val map = map_nodes; |
19615 | 347 |
val fold = fold_graph; |
348 |
||
6134 | 349 |
end; |
350 |
||
31971
8c1b845ed105
renamed functor TableFun to Table, and GraphFun to Graph;
wenzelm
parents:
31616
diff
changeset
|
351 |
structure Graph = Graph(type key = string val ord = fast_string_ord); |
46667 | 352 |
structure String_Graph = Graph(type key = string val ord = string_ord); |
35403 | 353 |
structure Int_Graph = Graph(type key = int val ord = int_ord); |