src/Pure/General/graph.ML
author wenzelm
Tue, 26 Mar 2024 21:44:18 +0100
changeset 80018 ac4412562c7b
parent 79446 ec7a1603129a
permissions -rw-r--r--
more robust XML body: allow empty text, as well as arbitrary pro-forma markup (e.g. see XML.blob in Isabelle/ML);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/General/graph.ML
15759
wenzelm
parents: 15570
diff changeset
     2
    Author:     Markus Wenzel and Stefan Berghofer, TU Muenchen
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
     3
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
     4
Directed graphs.
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
     5
*)
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
     6
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
     7
signature GRAPH =
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
     8
sig
77731
48fbecc8fab1 tuned signature: more uniform structure Key;
wenzelm
parents: 77724
diff changeset
     9
  structure Key: KEY
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
    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: 44202
diff 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: 44202
diff 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: 44202
diff 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: 44202
diff 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: 44202
diff 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: 44202
diff 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: 44202
diff 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: 44202
diff changeset
    18
  end
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
    19
  type 'a T
9321
e0dda4bde88c tuned exceptions;
wenzelm
parents: 8806
diff changeset
    20
  exception DUP of key
19029
8635700e2c9c share exception UNDEF with Table;
wenzelm
parents: 18970
diff changeset
    21
  exception SAME
8635700e2c9c share exception UNDEF with Table;
wenzelm
parents: 18970
diff changeset
    22
  exception UNDEF of key
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
    23
  val empty: 'a T
28204
2d93b158ad99 added is_empty;
wenzelm
parents: 28183
diff changeset
    24
  val is_empty: 'a T -> bool
6659
7a056250899d removed get_nodes;
wenzelm
parents: 6152
diff changeset
    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: 44202
diff 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: 44202
diff changeset
    27
  val fold: (key * ('a * (Keys.T * Keys.T)) -> 'b -> 'b) -> 'a T -> 'b -> 'b
79446
ec7a1603129a clarified signature;
wenzelm
parents: 77731
diff changeset
    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: 44202
diff changeset
    29
  val get_entry: 'a T -> key -> key * ('a * (Keys.T * Keys.T))        (*exception UNDEF*)
15759
wenzelm
parents: 15570
diff changeset
    30
  val get_node: 'a T -> key -> 'a                                     (*exception UNDEF*)
6142
wenzelm
parents: 6134
diff changeset
    31
  val map_node: key -> ('a -> 'a) -> 'a T -> 'a T
46611
669601fa1a62 directed graphs (in Scala);
wenzelm
parents: 44338
diff changeset
    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: 44202
diff 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: 44202
diff 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: 44202
diff 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: 44202
diff changeset
    36
  val immediate_succs: 'a T -> key -> key list
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
    37
  val all_preds: 'a T -> key list -> key list
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
    38
  val all_succs: 'a T -> key list -> key list
46613
74331911375d further graph operations from ML;
wenzelm
parents: 46612
diff changeset
    39
  val strong_conn: 'a T -> key list list
55658
d696adf157e6 simultaneous mapping of strongly connected components
haftmann
parents: 55154
diff 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: 44202
diff 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: 44202
diff 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: 44202
diff 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: 44202
diff changeset
    44
  val is_maximal: 'a T -> key -> bool
66830
3b50269b90c2 more operations;
wenzelm
parents: 55658
diff changeset
    45
  val is_isolated: 'a T -> key -> bool
15759
wenzelm
parents: 15570
diff changeset
    46
  val new_node: key * 'a -> 'a T -> 'a T                              (*exception DUP*)
17179
28802c8a9816 canonical interface for 'default'
haftmann
parents: 17140
diff changeset
    47
  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
    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: 14161
diff changeset
    49
  val is_edge: 'a T -> key * key -> bool
44202
44c4ae5c5ce2 touch descendants of edited nodes;
wenzelm
parents: 44162
diff changeset
    50
  val add_edge: key * key -> 'a T -> 'a T                             (*exception UNDEF*)
44c4ae5c5ce2 touch descendants of edited nodes;
wenzelm
parents: 44162
diff changeset
    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: 46613
diff changeset
    52
  val restrict: (key -> bool) -> 'a T -> 'a T
49560
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
    53
  val dest: 'a T -> ((key * 'a) * key list) list
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
    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: 21935
diff changeset
    55
  val merge: ('a * 'a -> bool) -> 'a T * 'a T -> 'a T                 (*exception DUP*)
19029
8635700e2c9c share exception UNDEF with Table;
wenzelm
parents: 18970
diff changeset
    56
  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
    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: 19482
diff changeset
    58
  val irreducible_paths: 'a T -> key * key -> key list list
6142
wenzelm
parents: 6134
diff changeset
    59
  exception CYCLES of key list list
44202
44c4ae5c5ce2 touch descendants of edited nodes;
wenzelm
parents: 44162
diff changeset
    60
  val add_edge_acyclic: key * key -> 'a T -> 'a T                     (*exception UNDEF | CYCLES*)
44c4ae5c5ce2 touch descendants of edited nodes;
wenzelm
parents: 44162
diff changeset
    61
  val add_deps_acyclic: key * key list -> 'a T -> 'a T                (*exception UNDEF | CYCLES*)
15759
wenzelm
parents: 15570
diff changeset
    62
  val merge_acyclic: ('a * 'a -> bool) -> 'a T * 'a T -> 'a T         (*exception CYCLES*)
23964
d2df2797519b added topological_order;
wenzelm
parents: 23655
diff changeset
    63
  val topological_order: 'a T -> key list
44202
44c4ae5c5ce2 touch descendants of edited nodes;
wenzelm
parents: 44162
diff changeset
    64
  val add_edge_trans_acyclic: key * key -> 'a T -> 'a T               (*exception UNDEF | CYCLES*)
15759
wenzelm
parents: 15570
diff changeset
    65
  val merge_trans_acyclic: ('a * 'a -> bool) -> 'a T * 'a T -> 'a T   (*exception CYCLES*)
44162
5434899d955c general Graph.schedule;
wenzelm
parents: 43792
diff changeset
    66
  exception DEP of key * key
5434899d955c general Graph.schedule;
wenzelm
parents: 43792
diff changeset
    67
  val schedule: ((key * 'b) list -> key * 'a -> 'b) -> 'a T -> 'b list  (*exception DEP*)
49560
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
    68
  val encode: key XML.Encode.T -> 'a XML.Encode.T -> 'a T XML.Encode.T
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
    69
  val decode: key XML.Decode.T -> 'a XML.Decode.T -> 'a T XML.Decode.T
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
    70
end;
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
    71
31971
8c1b845ed105 renamed functor TableFun to Table, and GraphFun to Graph;
wenzelm
parents: 31616
diff changeset
    72
functor Graph(Key: KEY): GRAPH =
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
    73
struct
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
    74
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
    75
(* keys *)
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
    76
77731
48fbecc8fab1 tuned signature: more uniform structure Key;
wenzelm
parents: 77724
diff changeset
    77
structure Key = Key;
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
    78
type key = Key.key;
77731
48fbecc8fab1 tuned signature: more uniform structure Key;
wenzelm
parents: 77724
diff changeset
    79
18970
d055a29ddd23 Library.is_equal;
wenzelm
parents: 18921
diff changeset
    80
val eq_key = is_equal o Key.ord;
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
    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
structure Table = Table(Key);
77723
b761c91c2447 performance tuning: prefer functor Set() over Table();
wenzelm
parents: 70773
diff 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: 44202
diff 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: 44202
diff 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: 44202
diff changeset
    86
struct
6152
wenzelm
parents: 6142
diff changeset
    87
77723
b761c91c2447 performance tuning: prefer functor Set() over Table();
wenzelm
parents: 70773
diff 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: 44202
diff changeset
    89
with
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
    90
77723
b761c91c2447 performance tuning: prefer functor Set() over Table();
wenzelm
parents: 70773
diff changeset
    91
val empty = Keys Set.empty;
b761c91c2447 performance tuning: prefer functor Set() over Table();
wenzelm
parents: 70773
diff changeset
    92
fun is_empty (Keys set) = Set.is_empty set;
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
    93
77723
b761c91c2447 performance tuning: prefer functor Set() over Table();
wenzelm
parents: 70773
diff changeset
    94
fun member (Keys set) = Set.member set;
b761c91c2447 performance tuning: prefer functor Set() over Table();
wenzelm
parents: 70773
diff changeset
    95
fun insert x (Keys set) = Keys (Set.insert x set);
b761c91c2447 performance tuning: prefer functor Set() over Table();
wenzelm
parents: 70773
diff 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: 44202
diff changeset
    97
77723
b761c91c2447 performance tuning: prefer functor Set() over Table();
wenzelm
parents: 70773
diff changeset
    98
fun fold f (Keys set) = Set.fold f set;
b761c91c2447 performance tuning: prefer functor Set() over Table();
wenzelm
parents: 70773
diff changeset
    99
fun fold_rev f (Keys set) = Set.fold_rev f set;
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
   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: 44202
diff changeset
   101
fun dest keys = fold_rev cons keys [];
6142
wenzelm
parents: 6134
diff changeset
   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: 44202
diff 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: 44202
diff 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: 44202
diff 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: 44202
diff changeset
   106
end;
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
   107
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
   108
6142
wenzelm
parents: 6134
diff changeset
   109
(* graphs *)
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
   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: 44202
diff changeset
   111
datatype 'a T = Graph of ('a * (Keys.T * Keys.T)) Table.table;
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
   112
9321
e0dda4bde88c tuned exceptions;
wenzelm
parents: 8806
diff changeset
   113
exception DUP = Table.DUP;
19029
8635700e2c9c share exception UNDEF with Table;
wenzelm
parents: 18970
diff changeset
   114
exception UNDEF = Table.UNDEF;
8635700e2c9c share exception UNDEF with Table;
wenzelm
parents: 18970
diff changeset
   115
exception SAME = Table.SAME;
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
   116
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
   117
val empty = Graph Table.empty;
28204
2d93b158ad99 added is_empty;
wenzelm
parents: 28183
diff changeset
   118
fun is_empty (Graph tab) = Table.is_empty tab;
6659
7a056250899d removed get_nodes;
wenzelm
parents: 6152
diff changeset
   119
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
   120
35012
c3e3ac3ca091 removed unused "boundary" of Table/Graph.get_first;
wenzelm
parents: 32710
diff changeset
   121
fun get_first f (Graph tab) = Table.get_first f tab;
19615
e3ab6cd838a4 added fold;
wenzelm
parents: 19580
diff changeset
   122
fun fold_graph f (Graph tab) = Table.fold f tab;
e3ab6cd838a4 added fold;
wenzelm
parents: 19580
diff changeset
   123
79446
ec7a1603129a clarified signature;
wenzelm
parents: 77731
diff changeset
   124
fun defined (Graph tab) = Table.defined tab;
ec7a1603129a clarified signature;
wenzelm
parents: 77731
diff changeset
   125
6142
wenzelm
parents: 6134
diff changeset
   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: 39021
diff changeset
   127
  (case Table.lookup_key tab x of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15160
diff changeset
   128
    SOME entry => entry
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15160
diff changeset
   129
  | NONE => raise UNDEF x);
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
   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: 39021
diff changeset
   131
fun map_entry x f (G as Graph tab) = Graph (Table.update (x, f (#2 (get_entry G x))) tab);
19290
wenzelm
parents: 19029
diff changeset
   132
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
   133
6142
wenzelm
parents: 6134
diff changeset
   134
(* nodes *)
wenzelm
parents: 6134
diff changeset
   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: 39021
diff changeset
   136
fun get_node G = #1 o #2 o get_entry G;
18133
1d403623dabc avoid code redundancy;
wenzelm
parents: 18126
diff changeset
   137
6142
wenzelm
parents: 6134
diff changeset
   138
fun map_node x f = map_entry x (fn (i, ps) => (f i, ps));
19290
wenzelm
parents: 19029
diff changeset
   139
46611
669601fa1a62 directed graphs (in Scala);
wenzelm
parents: 44338
diff changeset
   140
fun map_nodes f (Graph tab) = Graph (Table.map (apfst o f) tab);
669601fa1a62 directed graphs (in Scala);
wenzelm
parents: 44338
diff changeset
   141
18133
1d403623dabc avoid code redundancy;
wenzelm
parents: 18126
diff changeset
   142
6142
wenzelm
parents: 6134
diff changeset
   143
(* reachability *)
wenzelm
parents: 6134
diff changeset
   144
6659
7a056250899d removed get_nodes;
wenzelm
parents: 6152
diff changeset
   145
(*nodes reachable from xs -- topologically sorted for acyclic graphs*)
6142
wenzelm
parents: 6134
diff changeset
   146
fun reachable next xs =
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
   147
  let
18006
535de280c812 reachable - abandoned foldl_map in favor of fold_map
haftmann
parents: 17912
diff 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: 44202
diff changeset
   149
      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
   150
      else Keys.fold_rev reach (next x) (rs, Keys.insert x R) |>> cons x;
32710
wenzelm
parents: 32709
diff changeset
   151
    fun reachs x (rss, R) =
wenzelm
parents: 32709
diff changeset
   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: 44202
diff changeset
   153
  in fold reachs xs ([], Keys.empty) end;
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
   154
6142
wenzelm
parents: 6134
diff changeset
   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: 39021
diff 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: 39021
diff changeset
   157
fun imm_succs G = #2 o #2 o #2 o get_entry G;
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
   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: 44202
diff 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: 44202
diff 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: 44202
diff changeset
   161
6142
wenzelm
parents: 6134
diff changeset
   162
(*transitive*)
32710
wenzelm
parents: 32709
diff changeset
   163
fun all_preds G = flat o #1 o reachable (imm_preds G);
wenzelm
parents: 32709
diff changeset
   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: 12451
diff changeset
   165
46613
74331911375d further graph operations from ML;
wenzelm
parents: 46612
diff changeset
   166
(*strongly connected components; see: David King and John Launchbury,
74331911375d further graph operations from ML;
wenzelm
parents: 46612
diff changeset
   167
  "Structuring Depth First Search Algorithms in Haskell"*)
74331911375d further graph operations from ML;
wenzelm
parents: 46612
diff changeset
   168
fun strong_conn G =
74331911375d further graph operations from ML;
wenzelm
parents: 46612
diff changeset
   169
  rev (filter_out null (#1 (reachable (imm_preds G) (all_succs G (keys G)))));
74331911375d further graph operations from ML;
wenzelm
parents: 46612
diff changeset
   170
55658
d696adf157e6 simultaneous mapping of strongly connected components
haftmann
parents: 55154
diff changeset
   171
fun map_strong_conn f G =
d696adf157e6 simultaneous mapping of strongly connected components
haftmann
parents: 55154
diff changeset
   172
  let
d696adf157e6 simultaneous mapping of strongly connected components
haftmann
parents: 55154
diff changeset
   173
    val xss = strong_conn G;
d696adf157e6 simultaneous mapping of strongly connected components
haftmann
parents: 55154
diff changeset
   174
    fun map' xs =
d696adf157e6 simultaneous mapping of strongly connected components
haftmann
parents: 55154
diff changeset
   175
      fold2 (curry Table.update) xs (f (AList.make (get_node G) xs));
d696adf157e6 simultaneous mapping of strongly connected components
haftmann
parents: 55154
diff changeset
   176
    val tab' = Table.empty
d696adf157e6 simultaneous mapping of strongly connected components
haftmann
parents: 55154
diff changeset
   177
      |> fold map' xss;
d696adf157e6 simultaneous mapping of strongly connected components
haftmann
parents: 55154
diff changeset
   178
  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
   179
46613
74331911375d further graph operations from ML;
wenzelm
parents: 46612
diff changeset
   180
74331911375d further graph operations from ML;
wenzelm
parents: 46612
diff changeset
   181
(* minimal and maximal elements *)
74331911375d further graph operations from ML;
wenzelm
parents: 46612
diff changeset
   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: 44202
diff 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: 44202
diff 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: 44202
diff 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: 44202
diff 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: 44202
diff changeset
   187
66830
3b50269b90c2 more operations;
wenzelm
parents: 55658
diff changeset
   188
fun is_isolated G x = is_minimal G x andalso is_maximal G x;
3b50269b90c2 more operations;
wenzelm
parents: 55658
diff changeset
   189
18133
1d403623dabc avoid code redundancy;
wenzelm
parents: 18126
diff changeset
   190
46668
9034b44844bd tuned comments;
wenzelm
parents: 46667
diff changeset
   191
(* node operations *)
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
   192
6152
wenzelm
parents: 6142
diff changeset
   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: 44202
diff changeset
   194
  Graph (Table.update_new (x, (info, (Keys.empty, Keys.empty))) tab);
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
   195
17179
28802c8a9816 canonical interface for 'default'
haftmann
parents: 17140
diff changeset
   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: 44202
diff changeset
   197
  Graph (Table.default (x, (info, (Keys.empty, Keys.empty))) tab);
17140
5be3a21ec949 added 'default' function
haftmann
parents: 16894
diff changeset
   198
28333
507b64f4cd2a added del_node, which is more efficient for sparse graphs;
wenzelm
parents: 28204
diff changeset
   199
fun del_node x (G as Graph tab) =
507b64f4cd2a added del_node, which is more efficient for sparse graphs;
wenzelm
parents: 28204
diff 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: 44202
diff 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: 44202
diff 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: 39021
diff 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: 28204
diff changeset
   204
  in
507b64f4cd2a added del_node, which is more efficient for sparse graphs;
wenzelm
parents: 28204
diff changeset
   205
    Graph (tab
507b64f4cd2a added del_node, which is more efficient for sparse graphs;
wenzelm
parents: 28204
diff 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: 44202
diff 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: 44202
diff changeset
   208
      |> Keys.fold (del_adjacent apfst) succs)
28333
507b64f4cd2a added del_node, which is more efficient for sparse graphs;
wenzelm
parents: 28204
diff changeset
   209
  end;
507b64f4cd2a added del_node, which is more efficient for sparse graphs;
wenzelm
parents: 28204
diff changeset
   210
46614
165886a4fe64 clarified Graph.restrict (formerly Graph.subgraph) based on public graph operations;
wenzelm
parents: 46613
diff changeset
   211
fun restrict pred G =
165886a4fe64 clarified Graph.restrict (formerly Graph.subgraph) based on public graph operations;
wenzelm
parents: 46613
diff 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: 46613
diff changeset
   213
6152
wenzelm
parents: 6142
diff changeset
   214
46668
9034b44844bd tuned comments;
wenzelm
parents: 46667
diff changeset
   215
(* edge operations *)
9321
e0dda4bde88c tuned exceptions;
wenzelm
parents: 8806
diff changeset
   216
77724
b4032c468d74 performanc tuning: avoid exception overhead, potentially relevant for Sorts.class_less;
wenzelm
parents: 77723
diff changeset
   217
fun is_edge (Graph tab) (x, y) =
b4032c468d74 performanc tuning: avoid exception overhead, potentially relevant for Sorts.class_less;
wenzelm
parents: 77723
diff changeset
   218
  (case Table.lookup tab x of
b4032c468d74 performanc tuning: avoid exception overhead, potentially relevant for Sorts.class_less;
wenzelm
parents: 77723
diff changeset
   219
    SOME (_, (_, succs)) => Keys.member succs y
b4032c468d74 performanc tuning: avoid exception overhead, potentially relevant for Sorts.class_less;
wenzelm
parents: 77723
diff changeset
   220
  | NONE => false);
14793
32d94d1e4842 added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents: 14161
diff changeset
   221
32d94d1e4842 added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents: 14161
diff changeset
   222
fun add_edge (x, y) G =
32d94d1e4842 added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents: 14161
diff changeset
   223
  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
   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: 44202
diff 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: 44202
diff 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: 14161
diff changeset
   227
32d94d1e4842 added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents: 14161
diff changeset
   228
fun del_edge (x, y) G =
32d94d1e4842 added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents: 14161
diff 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: 44202
diff 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: 44202
diff 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: 14161
diff changeset
   232
  else G;
9321
e0dda4bde88c tuned exceptions;
wenzelm
parents: 8806
diff changeset
   233
14793
32d94d1e4842 added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents: 14161
diff changeset
   234
fun diff_edges G1 G2 =
49560
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   235
  fold_graph (fn (x, (_, (_, succs))) =>
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   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: 14161
diff changeset
   237
32d94d1e4842 added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents: 14161
diff changeset
   238
fun edges G = diff_edges G empty;
32d94d1e4842 added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents: 14161
diff changeset
   239
32d94d1e4842 added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents: 14161
diff changeset
   240
49560
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   241
(* dest and make *)
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   242
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   243
fun dest G = fold_graph (fn (x, (i, (_, succs))) => cons ((x, i), Keys.dest succs)) G [];
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   244
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   245
fun make entries =
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   246
  empty
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   247
  |> fold (new_node o fst) entries
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   248
  |> fold (fn ((x, _), ys) => fold (fn y => add_edge (x, y)) ys) entries;
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   249
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   250
18126
b74145e46e0d added join function
haftmann
parents: 18006
diff changeset
   251
(* join and merge *)
b74145e46e0d added join function
haftmann
parents: 18006
diff changeset
   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: 44202
diff changeset
   253
fun no_edges (i, _) = (i, (Keys.empty, Keys.empty));
18133
1d403623dabc avoid code redundancy;
wenzelm
parents: 18126
diff changeset
   254
35974
3a588b344749 low-level tuning for join/merge: ignore identical versions (SUBTLE CHANGE IN SEMANTICS);
wenzelm
parents: 35405
diff 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: 35405
diff 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: 35405
diff changeset
   257
    if pointer_eq (G1, G2) then G1
39020
ac0f24f850c9 replaced Table.map' by Table.map
haftmann
parents: 37853
diff changeset
   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: 35405
diff changeset
   259
  end;
6152
wenzelm
parents: 6142
diff changeset
   260
35974
3a588b344749 low-level tuning for join/merge: ignore identical versions (SUBTLE CHANGE IN SEMANTICS);
wenzelm
parents: 35405
diff 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: 35405
diff 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: 35405
diff changeset
   263
    if pointer_eq (G1, G2) then G1
39020
ac0f24f850c9 replaced Table.map' by Table.map
haftmann
parents: 37853
diff changeset
   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: 35405
diff changeset
   265
  end;
6152
wenzelm
parents: 6142
diff changeset
   266
14793
32d94d1e4842 added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents: 14161
diff 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: 14161
diff changeset
   268
18133
1d403623dabc avoid code redundancy;
wenzelm
parents: 18126
diff changeset
   269
19580
c878a09fb849 replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents: 19482
diff changeset
   270
(* irreducible paths -- Hasse diagram *)
c878a09fb849 replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents: 19482
diff changeset
   271
c878a09fb849 replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents: 19482
diff 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: 19482
diff changeset
   273
  let
c878a09fb849 replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents: 19482
diff 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: 19482
diff changeset
   275
    fun irreds [] xs' = xs'
c878a09fb849 replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents: 19482
diff 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: 44202
diff 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: 19482
diff 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: 19482
diff changeset
   279
          then irreds xs xs'
c878a09fb849 replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents: 19482
diff 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: 44202
diff 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: 19482
diff changeset
   282
c878a09fb849 replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents: 19482
diff 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: 19482
diff changeset
   284
  let
c878a09fb849 replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents: 19482
diff 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: 19482
diff changeset
   286
    fun paths path z =
c878a09fb849 replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents: 19482
diff 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: 19482
diff 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: 19482
diff 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: 19482
diff changeset
   290
c878a09fb849 replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents: 19482
diff changeset
   291
14793
32d94d1e4842 added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents: 14161
diff changeset
   292
(* maintain acyclic graphs *)
6142
wenzelm
parents: 6134
diff changeset
   293
wenzelm
parents: 6134
diff changeset
   294
exception CYCLES of key list list;
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
   295
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
   296
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
   297
  if is_edge G (x, y) then G
9347
1791a62b33e7 improved add_edges_cyclic;
wenzelm
parents: 9321
diff changeset
   298
  else
19580
c878a09fb849 replaced find_paths by irreducible_paths, i.e. produce paths within a Hasse diagram;
wenzelm
parents: 19482
diff changeset
   299
    (case irreducible_paths G (y, x) of
9347
1791a62b33e7 improved add_edges_cyclic;
wenzelm
parents: 9321
diff changeset
   300
      [] => add_edge (x, y) G
1791a62b33e7 improved add_edges_cyclic;
wenzelm
parents: 9321
diff changeset
   301
    | cycles => raise CYCLES (map (cons x) cycles));
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
   302
15759
wenzelm
parents: 15570
diff changeset
   303
fun add_deps_acyclic (y, xs) = fold (fn x => add_edge_acyclic (x, y)) xs;
9321
e0dda4bde88c tuned exceptions;
wenzelm
parents: 8806
diff changeset
   304
14793
32d94d1e4842 added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents: 14161
diff changeset
   305
fun merge_acyclic eq GG = gen_merge add_edge_acyclic eq GG;
9321
e0dda4bde88c tuned exceptions;
wenzelm
parents: 8806
diff changeset
   306
23964
d2df2797519b added topological_order;
wenzelm
parents: 23655
diff changeset
   307
fun topological_order G = minimals G |> all_succs G;
d2df2797519b added topological_order;
wenzelm
parents: 23655
diff changeset
   308
14793
32d94d1e4842 added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents: 14161
diff changeset
   309
32d94d1e4842 added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents: 14161
diff changeset
   310
(* maintain transitive acyclic graphs *)
9321
e0dda4bde88c tuned exceptions;
wenzelm
parents: 8806
diff changeset
   311
14793
32d94d1e4842 added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents: 14161
diff changeset
   312
fun add_edge_trans_acyclic (x, y) G =
19290
wenzelm
parents: 19029
diff changeset
   313
  add_edge_acyclic (x, y) G
25538
58e8ba3b792b map_product and fold_product
haftmann
parents: 23964
diff changeset
   314
  |> fold_product (curry add_edge) (all_preds G [x]) (all_succs G [y]);
9321
e0dda4bde88c tuned exceptions;
wenzelm
parents: 8806
diff changeset
   315
14793
32d94d1e4842 added dest, minimals, maximals, is_edge, add_edge/merge_trans_acyclic;
wenzelm
parents: 14161
diff 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: 35405
diff 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: 35405
diff changeset
   318
  else
3a588b344749 low-level tuning for join/merge: ignore identical versions (SUBTLE CHANGE IN SEMANTICS);
wenzelm
parents: 35405
diff changeset
   319
    merge_acyclic eq (G1, G2)
3a588b344749 low-level tuning for join/merge: ignore identical versions (SUBTLE CHANGE IN SEMANTICS);
wenzelm
parents: 35405
diff 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: 35405
diff changeset
   321
    |> fold add_edge_trans_acyclic (diff_edges G2 G1);
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
   322
31540
4cfe0f756feb removed unused make;
wenzelm
parents: 31516
diff changeset
   323
44162
5434899d955c general Graph.schedule;
wenzelm
parents: 43792
diff changeset
   324
(* schedule acyclic graph *)
5434899d955c general Graph.schedule;
wenzelm
parents: 43792
diff changeset
   325
5434899d955c general Graph.schedule;
wenzelm
parents: 43792
diff changeset
   326
exception DEP of key * key;
5434899d955c general Graph.schedule;
wenzelm
parents: 43792
diff changeset
   327
5434899d955c general Graph.schedule;
wenzelm
parents: 43792
diff changeset
   328
fun schedule f G =
5434899d955c general Graph.schedule;
wenzelm
parents: 43792
diff changeset
   329
  let
5434899d955c general Graph.schedule;
wenzelm
parents: 43792
diff changeset
   330
    val xs = topological_order G;
5434899d955c general Graph.schedule;
wenzelm
parents: 43792
diff changeset
   331
    val results = (xs, Table.empty) |-> fold (fn x => fn tab =>
5434899d955c general Graph.schedule;
wenzelm
parents: 43792
diff changeset
   332
      let
5434899d955c general Graph.schedule;
wenzelm
parents: 43792
diff changeset
   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: 44202
diff changeset
   334
        val deps = immediate_preds G x |> map (fn y =>
44162
5434899d955c general Graph.schedule;
wenzelm
parents: 43792
diff changeset
   335
          (case Table.lookup tab y of
5434899d955c general Graph.schedule;
wenzelm
parents: 43792
diff changeset
   336
            SOME b => (y, b)
5434899d955c general Graph.schedule;
wenzelm
parents: 43792
diff changeset
   337
          | NONE => raise DEP (x, y)));
5434899d955c general Graph.schedule;
wenzelm
parents: 43792
diff changeset
   338
      in Table.update (x, f deps (x, a)) tab end);
5434899d955c general Graph.schedule;
wenzelm
parents: 43792
diff changeset
   339
  in map (the o Table.lookup results) xs end;
5434899d955c general Graph.schedule;
wenzelm
parents: 43792
diff changeset
   340
5434899d955c general Graph.schedule;
wenzelm
parents: 43792
diff changeset
   341
49560
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   342
(* XML data representation *)
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   343
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   344
fun encode key info G =
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   345
  dest G |>
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   346
    let open XML.Encode
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   347
    in list (pair (pair key info) (list key)) end;
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   348
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   349
fun decode key info body =
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   350
  body |>
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   351
    let open XML.Decode
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   352
    in list (pair (pair key info) (list key)) end |> make;
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   353
11430dd89e35 added graph encode/decode operations;
wenzelm
parents: 48350
diff changeset
   354
19615
e3ab6cd838a4 added fold;
wenzelm
parents: 19580
diff changeset
   355
(*final declarations of this structure!*)
39021
139aada5caf8 Graph.map, in analogy to Table.map
haftmann
parents: 39020
diff changeset
   356
val map = map_nodes;
19615
e3ab6cd838a4 added fold;
wenzelm
parents: 19580
diff changeset
   357
val fold = fold_graph;
e3ab6cd838a4 added fold;
wenzelm
parents: 19580
diff changeset
   358
6134
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
   359
end;
ec6092b0599d added General/graph.ML: generic direct graphs;
wenzelm
parents:
diff changeset
   360
77731
48fbecc8fab1 tuned signature: more uniform structure Key;
wenzelm
parents: 77724
diff changeset
   361
structure Graph = Graph(Symtab.Key);
46667
318b669fe660 standard Graph instances;
wenzelm
parents: 46665
diff changeset
   362
structure String_Graph = Graph(type key = string val ord = string_ord);
77731
48fbecc8fab1 tuned signature: more uniform structure Key;
wenzelm
parents: 77724
diff changeset
   363
structure Int_Graph = Graph(Inttab.Key);