| author | wenzelm | 
| Tue, 10 Mar 2009 22:27:32 +0100 | |
| changeset 30423 | 6baef860dfa6 | 
| parent 30290 | f49d70426690 | 
| child 31516 | 9801a92d52d7 | 
| 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 | |
| 10 | type 'a T | |
| 9321 | 11 | exception DUP of key | 
| 19029 | 12 | exception SAME | 
| 13 | exception UNDEF of key | |
| 6134 | 14 | val empty: 'a T | 
| 28204 | 15 | val is_empty: 'a T -> bool | 
| 6659 | 16 | val keys: 'a T -> key list | 
| 14793 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 17 | val dest: 'a T -> (key * key list) list | 
| 28183 | 18 |   val get_first: (key * ('a * (key list * key list)) -> 'b option) -> 'a T -> 'b option
 | 
| 19615 | 19 |   val fold: (key * ('a * (key list * key list)) -> 'b -> 'b) -> 'a T -> 'b -> 'b
 | 
| 14793 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 20 | val minimals: 'a T -> key list | 
| 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 21 | val maximals: 'a T -> key list | 
| 21935 
4e20a5397b57
renamed project to subgraph, improved presentation, avoided unnecessary evaluation of predicate;
 wenzelm parents: 
21565diff
changeset | 22 | val subgraph: (key -> bool) -> 'a T -> 'a T | 
| 6142 | 23 |   val map_nodes: ('a -> 'b) -> 'a T -> 'b T
 | 
| 15759 | 24 | val get_node: 'a T -> key -> 'a (*exception UNDEF*) | 
| 6142 | 25 |   val map_node: key -> ('a -> 'a) -> 'a T -> 'a T
 | 
| 17767 | 26 |   val map_node_yield: key -> ('a -> 'b * 'a) -> 'a T -> 'b * 'a T
 | 
| 6142 | 27 | val imm_preds: 'a T -> key -> key list | 
| 28 | val imm_succs: 'a T -> key -> key list | |
| 6134 | 29 | val all_preds: 'a T -> key list -> key list | 
| 30 | val all_succs: 'a T -> key list -> key list | |
| 14161 
73ad4884441f
Added function strong_conn for computing the strongly connected components
 berghofe parents: 
12451diff
changeset | 31 | val strong_conn: 'a T -> key list list | 
| 15759 | 32 | val new_node: key * 'a -> 'a T -> 'a T (*exception DUP*) | 
| 17179 | 33 | val default_node: key * 'a -> 'a T -> 'a T | 
| 15759 | 34 | val del_nodes: key list -> 'a T -> 'a T (*exception UNDEF*) | 
| 28333 
507b64f4cd2a
added del_node, which is more efficient for sparse graphs;
 wenzelm parents: 
28204diff
changeset | 35 | 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 | 36 | val is_edge: 'a T -> key * key -> bool | 
| 6134 | 37 | val add_edge: key * key -> 'a T -> 'a T | 
| 6152 | 38 | val del_edge: key * key -> 'a T -> 'a T | 
| 23655 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 wenzelm parents: 
21935diff
changeset | 39 |   val merge: ('a * 'a -> bool) -> 'a T * 'a T -> 'a T                 (*exception DUP*)
 | 
| 19029 | 40 | val join: (key -> 'a * 'a -> 'a) (*exception DUP/SAME*) -> | 
| 23655 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 wenzelm parents: 
21935diff
changeset | 41 | '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 | 42 | val irreducible_paths: 'a T -> key * key -> key list list | 
| 20679 | 43 | val all_paths: 'a T -> key * key -> key list list | 
| 6142 | 44 | exception CYCLES of key list list | 
| 15759 | 45 | val add_edge_acyclic: key * key -> 'a T -> 'a T (*exception CYCLES*) | 
| 46 | val add_deps_acyclic: key * key list -> 'a T -> 'a T (*exception CYCLES*) | |
| 47 |   val merge_acyclic: ('a * 'a -> bool) -> 'a T * 'a T -> 'a T         (*exception CYCLES*)
 | |
| 23964 | 48 | val topological_order: 'a T -> key list | 
| 15759 | 49 | val add_edge_trans_acyclic: key * key -> 'a T -> 'a T (*exception CYCLES*) | 
| 50 |   val merge_trans_acyclic: ('a * 'a -> bool) -> 'a T * 'a T -> 'a T   (*exception CYCLES*)
 | |
| 6134 | 51 | end; | 
| 52 | ||
| 53 | functor GraphFun(Key: KEY): GRAPH = | |
| 54 | struct | |
| 55 | ||
| 56 | (* keys *) | |
| 57 | ||
| 58 | type key = Key.key; | |
| 59 | ||
| 18970 | 60 | val eq_key = is_equal o Key.ord; | 
| 6134 | 61 | |
| 18921 | 62 | val member_key = member eq_key; | 
| 15759 | 63 | val remove_key = remove eq_key; | 
| 6152 | 64 | |
| 6134 | 65 | |
| 66 | (* tables and sets of keys *) | |
| 67 | ||
| 68 | structure Table = TableFun(Key); | |
| 69 | type keys = unit Table.table; | |
| 70 | ||
| 6142 | 71 | val empty_keys = Table.empty: keys; | 
| 72 | ||
| 18921 | 73 | fun member_keys tab = Table.defined (tab: keys); | 
| 74 | fun insert_keys x tab = Table.insert (K true) (x, ()) (tab: keys); | |
| 6134 | 75 | |
| 76 | ||
| 6142 | 77 | (* graphs *) | 
| 6134 | 78 | |
| 79 | datatype 'a T = Graph of ('a * (key list * key list)) Table.table;
 | |
| 80 | ||
| 9321 | 81 | exception DUP = Table.DUP; | 
| 19029 | 82 | exception UNDEF = Table.UNDEF; | 
| 83 | exception SAME = Table.SAME; | |
| 6134 | 84 | |
| 85 | val empty = Graph Table.empty; | |
| 28204 | 86 | fun is_empty (Graph tab) = Table.is_empty tab; | 
| 6659 | 87 | fun keys (Graph tab) = Table.keys tab; | 
| 14793 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 88 | fun dest (Graph tab) = map (fn (x, (_, (_, succs))) => (x, succs)) (Table.dest tab); | 
| 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 89 | |
| 28183 | 90 | fun get_first f (Graph tab) = Table.get_first f tab; | 
| 19615 | 91 | fun fold_graph f (Graph tab) = Table.fold f tab; | 
| 92 | ||
| 93 | fun minimals G = fold_graph (fn (m, (_, ([], _))) => cons m | _ => I) G []; | |
| 94 | fun maximals G = fold_graph (fn (m, (_, (_, []))) => cons m | _ => I) G []; | |
| 6134 | 95 | |
| 21935 
4e20a5397b57
renamed project to subgraph, improved presentation, avoided unnecessary evaluation of predicate;
 wenzelm parents: 
21565diff
changeset | 96 | fun subgraph P G = | 
| 
4e20a5397b57
renamed project to subgraph, improved presentation, avoided unnecessary evaluation of predicate;
 wenzelm parents: 
21565diff
changeset | 97 | let | 
| 
4e20a5397b57
renamed project to subgraph, improved presentation, avoided unnecessary evaluation of predicate;
 wenzelm parents: 
21565diff
changeset | 98 | fun subg (k, (i, (preds, succs))) = | 
| 
4e20a5397b57
renamed project to subgraph, improved presentation, avoided unnecessary evaluation of predicate;
 wenzelm parents: 
21565diff
changeset | 99 | if P k then Table.update (k, (i, (filter P preds, filter P succs))) else I; | 
| 
4e20a5397b57
renamed project to subgraph, improved presentation, avoided unnecessary evaluation of predicate;
 wenzelm parents: 
21565diff
changeset | 100 | in Graph (fold_graph subg G Table.empty) end; | 
| 
4e20a5397b57
renamed project to subgraph, improved presentation, avoided unnecessary evaluation of predicate;
 wenzelm parents: 
21565diff
changeset | 101 | |
| 6142 | 102 | fun get_entry (Graph tab) x = | 
| 17412 | 103 | (case Table.lookup tab x of | 
| 15531 | 104 | SOME entry => entry | 
| 105 | | NONE => raise UNDEF x); | |
| 6134 | 106 | |
| 17412 | 107 | fun map_entry x f (G as Graph tab) = Graph (Table.update (x, f (get_entry G x)) tab); | 
| 19290 | 108 | |
| 17767 | 109 | fun map_entry_yield x f (G as Graph tab) = | 
| 110 | let val (a, node') = f (get_entry G x) | |
| 111 | in (a, Graph (Table.update (x, node') tab)) end; | |
| 6134 | 112 | |
| 113 | ||
| 6142 | 114 | (* nodes *) | 
| 115 | ||
| 116 | fun map_nodes f (Graph tab) = Graph (Table.map (fn (i, ps) => (f i, ps)) tab); | |
| 6134 | 117 | |
| 6142 | 118 | fun get_node G = #1 o get_entry G; | 
| 18133 | 119 | |
| 6142 | 120 | fun map_node x f = map_entry x (fn (i, ps) => (f i, ps)); | 
| 19290 | 121 | |
| 17767 | 122 | fun map_node_yield x f = map_entry_yield x (fn (i, ps) => | 
| 123 | let val (a, i') = f i in (a, (i', ps)) end); | |
| 6142 | 124 | |
| 18133 | 125 | |
| 6142 | 126 | (* reachability *) | 
| 127 | ||
| 6659 | 128 | (*nodes reachable from xs -- topologically sorted for acyclic graphs*) | 
| 6142 | 129 | fun reachable next xs = | 
| 6134 | 130 | let | 
| 18006 
535de280c812
reachable - abandoned foldl_map in favor of fold_map
 haftmann parents: 
17912diff
changeset | 131 | fun reach x (rs, R) = | 
| 18921 | 132 | if member_keys R x then (rs, R) | 
| 133 | else apfst (cons x) (fold reach (next x) (rs, insert_keys x R)) | |
| 23964 | 134 | in fold_map (fn x => fn X => reach x ([], X)) xs empty_keys end; | 
| 6134 | 135 | |
| 6142 | 136 | (*immediate*) | 
| 137 | fun imm_preds G = #1 o #2 o get_entry G; | |
| 138 | fun imm_succs G = #2 o #2 o get_entry G; | |
| 6134 | 139 | |
| 6142 | 140 | (*transitive*) | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19409diff
changeset | 141 | fun all_preds G = flat o fst o reachable (imm_preds G); | 
| 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19409diff
changeset | 142 | fun all_succs G = flat o fst o reachable (imm_succs G); | 
| 14161 
73ad4884441f
Added function strong_conn for computing the strongly connected components
 berghofe parents: 
12451diff
changeset | 143 | |
| 14793 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 144 | (*strongly connected components; see: David King and John Launchbury, | 
| 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 145 | "Structuring Depth First Search Algorithms in Haskell"*) | 
| 18006 
535de280c812
reachable - abandoned foldl_map in favor of fold_map
 haftmann parents: 
17912diff
changeset | 146 | fun strong_conn G = filter_out null (fst (reachable (imm_preds G) | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19409diff
changeset | 147 | (flat (rev (fst (reachable (imm_succs G) (keys G))))))); | 
| 6134 | 148 | |
| 18133 | 149 | |
| 9321 | 150 | (* nodes *) | 
| 6134 | 151 | |
| 6152 | 152 | fun new_node (x, info) (Graph tab) = | 
| 17412 | 153 | Graph (Table.update_new (x, (info, ([], []))) tab); | 
| 6134 | 154 | |
| 17179 | 155 | fun default_node (x, info) (Graph tab) = | 
| 156 | Graph (Table.default (x, (info, ([], []))) tab); | |
| 17140 | 157 | |
| 6659 | 158 | fun del_nodes xs (Graph tab) = | 
| 15759 | 159 | Graph (tab | 
| 160 | |> fold Table.delete xs | |
| 161 | |> Table.map (fn (i, (preds, succs)) => | |
| 162 | (i, (fold remove_key xs preds, fold remove_key xs succs)))); | |
| 6659 | 163 | |
| 28333 
507b64f4cd2a
added del_node, which is more efficient for sparse graphs;
 wenzelm parents: 
28204diff
changeset | 164 | fun del_node x (G as Graph tab) = | 
| 
507b64f4cd2a
added del_node, which is more efficient for sparse graphs;
 wenzelm parents: 
28204diff
changeset | 165 | let | 
| 
507b64f4cd2a
added del_node, which is more efficient for sparse graphs;
 wenzelm parents: 
28204diff
changeset | 166 | fun del_adjacent which y = Table.map_entry y (fn (i, ps) => (i, (which (remove_key x) ps))); | 
| 
507b64f4cd2a
added del_node, which is more efficient for sparse graphs;
 wenzelm parents: 
28204diff
changeset | 167 | val (preds, succs) = #2 (get_entry G x); | 
| 
507b64f4cd2a
added del_node, which is more efficient for sparse graphs;
 wenzelm parents: 
28204diff
changeset | 168 | in | 
| 
507b64f4cd2a
added del_node, which is more efficient for sparse graphs;
 wenzelm parents: 
28204diff
changeset | 169 | Graph (tab | 
| 
507b64f4cd2a
added del_node, which is more efficient for sparse graphs;
 wenzelm parents: 
28204diff
changeset | 170 | |> Table.delete x | 
| 
507b64f4cd2a
added del_node, which is more efficient for sparse graphs;
 wenzelm parents: 
28204diff
changeset | 171 | |> fold (del_adjacent apsnd) preds | 
| 
507b64f4cd2a
added del_node, which is more efficient for sparse graphs;
 wenzelm parents: 
28204diff
changeset | 172 | |> fold (del_adjacent apfst) succs) | 
| 
507b64f4cd2a
added del_node, which is more efficient for sparse graphs;
 wenzelm parents: 
28204diff
changeset | 173 | end; | 
| 
507b64f4cd2a
added del_node, which is more efficient for sparse graphs;
 wenzelm parents: 
28204diff
changeset | 174 | |
| 6152 | 175 | |
| 9321 | 176 | (* edges *) | 
| 177 | ||
| 18921 | 178 | fun is_edge G (x, y) = member_key (imm_succs G x) y handle UNDEF _ => false; | 
| 14793 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 179 | |
| 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 180 | fun add_edge (x, y) G = | 
| 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 181 | if is_edge G (x, y) then G | 
| 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 182 | else | 
| 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 183 | G |> map_entry y (fn (i, (preds, succs)) => (i, (x :: preds, succs))) | 
| 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 184 | |> map_entry x (fn (i, (preds, succs)) => (i, (preds, y :: succs))); | 
| 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 185 | |
| 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 186 | fun del_edge (x, y) G = | 
| 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 187 | if is_edge G (x, y) then | 
| 15759 | 188 | G |> map_entry y (fn (i, (preds, succs)) => (i, (remove_key x preds, succs))) | 
| 189 | |> map_entry x (fn (i, (preds, succs)) => (i, (preds, remove_key y succs))) | |
| 14793 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 190 | else G; | 
| 9321 | 191 | |
| 14793 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 192 | fun diff_edges G1 G2 = | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19409diff
changeset | 193 | flat (dest G1 |> map (fn (x, ys) => ys |> map_filter (fn y => | 
| 15531 | 194 | if is_edge G2 (x, y) then NONE else SOME (x, y)))); | 
| 14793 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 195 | |
| 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 196 | fun edges G = diff_edges G empty; | 
| 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 197 | |
| 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 198 | |
| 18126 | 199 | (* join and merge *) | 
| 200 | ||
| 18133 | 201 | fun no_edges (i, _) = (i, ([], [])); | 
| 202 | ||
| 203 | fun join f (Graph tab1, G2 as Graph tab2) = | |
| 19029 | 204 | let fun join_node key ((i1, edges1), (i2, _)) = (f key (i1, i2), edges1) | 
| 18133 | 205 | in fold add_edge (edges G2) (Graph (Table.join join_node (tab1, Table.map no_edges tab2))) end; | 
| 6152 | 206 | |
| 14793 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 207 | fun gen_merge add eq (Graph tab1, G2 as Graph tab2) = | 
| 18133 | 208 | let fun eq_node ((i1, _), (i2, _)) = eq (i1, i2) | 
| 14793 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 209 | in fold add (edges G2) (Graph (Table.merge eq_node (tab1, Table.map no_edges tab2))) end; | 
| 6152 | 210 | |
| 14793 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 211 | 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 | 212 | |
| 18133 | 213 | |
| 19580 
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
 wenzelm parents: 
19482diff
changeset | 214 | (* irreducible paths -- Hasse diagram *) | 
| 
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
 wenzelm parents: 
19482diff
changeset | 215 | |
| 
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
 wenzelm parents: 
19482diff
changeset | 216 | 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 | 217 | let | 
| 
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
 wenzelm parents: 
19482diff
changeset | 218 | 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 | 219 | fun irreds [] xs' = xs' | 
| 
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
 wenzelm parents: 
19482diff
changeset | 220 | | irreds (x :: xs) xs' = | 
| 
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
 wenzelm parents: 
19482diff
changeset | 221 | if not (member_keys X x) orelse eq_key (x, z) orelse member_key path x orelse | 
| 
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
 wenzelm parents: 
19482diff
changeset | 222 | 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 | 223 | then irreds xs xs' | 
| 
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
 wenzelm parents: 
19482diff
changeset | 224 | else irreds xs (x :: xs'); | 
| 
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
 wenzelm parents: 
19482diff
changeset | 225 | in irreds (imm_preds G z) [] end; | 
| 
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
 wenzelm parents: 
19482diff
changeset | 226 | |
| 
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
 wenzelm parents: 
19482diff
changeset | 227 | 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 | 228 | let | 
| 
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
 wenzelm parents: 
19482diff
changeset | 229 | 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 | 230 | fun paths path z = | 
| 
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
 wenzelm parents: 
19482diff
changeset | 231 | 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 | 232 | 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 | 233 | 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 | 234 | |
| 
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
 wenzelm parents: 
19482diff
changeset | 235 | |
| 20736 | 236 | (* all paths *) | 
| 20679 | 237 | |
| 238 | fun all_paths G (x, y) = | |
| 239 | let | |
| 240 | val (_, X) = reachable (imm_succs G) [x]; | |
| 20736 | 241 | fun paths path z = | 
| 242 | if not (null path) andalso eq_key (x, z) then [z :: path] | |
| 243 | else if member_keys X z andalso not (member_key path z) | |
| 244 | then maps (paths (z :: path)) (imm_preds G z) | |
| 20679 | 245 | else []; | 
| 246 | in paths [] y end; | |
| 247 | ||
| 248 | ||
| 14793 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 249 | (* maintain acyclic graphs *) | 
| 6142 | 250 | |
| 251 | exception CYCLES of key list list; | |
| 6134 | 252 | |
| 253 | fun add_edge_acyclic (x, y) G = | |
| 14793 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 254 | if is_edge G (x, y) then G | 
| 9347 | 255 | else | 
| 19580 
c878a09fb849
replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
 wenzelm parents: 
19482diff
changeset | 256 | (case irreducible_paths G (y, x) of | 
| 9347 | 257 | [] => add_edge (x, y) G | 
| 258 | | cycles => raise CYCLES (map (cons x) cycles)); | |
| 6134 | 259 | |
| 15759 | 260 | fun add_deps_acyclic (y, xs) = fold (fn x => add_edge_acyclic (x, y)) xs; | 
| 9321 | 261 | |
| 14793 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 262 | fun merge_acyclic eq GG = gen_merge add_edge_acyclic eq GG; | 
| 9321 | 263 | |
| 23964 | 264 | fun topological_order G = minimals G |> all_succs G; | 
| 265 | ||
| 14793 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 266 | |
| 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 267 | (* maintain transitive acyclic graphs *) | 
| 9321 | 268 | |
| 14793 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 269 | fun add_edge_trans_acyclic (x, y) G = | 
| 19290 | 270 | add_edge_acyclic (x, y) G | 
| 25538 | 271 | |> fold_product (curry add_edge) (all_preds G [x]) (all_succs G [y]); | 
| 9321 | 272 | |
| 14793 
32d94d1e4842
added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
 wenzelm parents: 
14161diff
changeset | 273 | fun merge_trans_acyclic eq (G1, G2) = | 
| 19290 | 274 | merge_acyclic eq (G1, G2) | 
| 275 | |> fold add_edge_trans_acyclic (diff_edges G1 G2) | |
| 276 | |> fold add_edge_trans_acyclic (diff_edges G2 G1); | |
| 6134 | 277 | |
| 19615 | 278 | |
| 279 | (*final declarations of this structure!*) | |
| 280 | val fold = fold_graph; | |
| 281 | ||
| 6134 | 282 | end; | 
| 283 | ||
| 16810 | 284 | structure Graph = GraphFun(type key = string val ord = fast_string_ord); | 
| 19615 | 285 | structure IntGraph = GraphFun(type key = int val ord = int_ord); |