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