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