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