| author | Fabian Huch <huch@in.tum.de> | 
| Fri, 15 Mar 2024 17:57:03 +0100 | |
| changeset 79899 | c73a36081b1c | 
| parent 78235 | aece9a063271 | 
| permissions | -rw-r--r-- | 
| 46611 | 1 | /* Title: Pure/General/graph.scala | 
| 2 | Author: Makarius | |
| 3 | ||
| 4 | Directed graphs. | |
| 5 | */ | |
| 6 | ||
| 7 | package isabelle | |
| 8 | ||
| 9 | ||
| 46661 
d2ac78ba805e
prefer sorted Map/Set for canonical order of results -- pass ordering via fresh copy of empty;
 wenzelm parents: 
46659diff
changeset | 10 | import scala.collection.immutable.{SortedMap, SortedSet}
 | 
| 46611 | 11 | import scala.annotation.tailrec | 
| 12 | ||
| 13 | ||
| 75393 | 14 | object Graph {
 | 
| 15 |   class Duplicate[Key](val key: Key) extends Exception {
 | |
| 70651 | 16 |     override def toString: String = "Graph.Duplicate(" + key.toString + ")"
 | 
| 17 | } | |
| 75393 | 18 |   class Undefined[Key](val key: Key) extends Exception {
 | 
| 70651 | 19 |     override def toString: String = "Graph.Undefined(" + key.toString + ")"
 | 
| 20 | } | |
| 75393 | 21 |   class Cycles[Key](val cycles: List[List[Key]]) extends Exception {
 | 
| 70651 | 22 |     override def toString: String = cycles.mkString("Graph.Cycles(", ",\n", ")")
 | 
| 23 | } | |
| 46611 | 24 | |
| 46661 
d2ac78ba805e
prefer sorted Map/Set for canonical order of results -- pass ordering via fresh copy of empty;
 wenzelm parents: 
46659diff
changeset | 25 | def empty[Key, A](implicit ord: Ordering[Key]): Graph[Key, A] = | 
| 
d2ac78ba805e
prefer sorted Map/Set for canonical order of results -- pass ordering via fresh copy of empty;
 wenzelm parents: 
46659diff
changeset | 26 | new Graph[Key, A](SortedMap.empty(ord)) | 
| 46667 | 27 | |
| 75393 | 28 | def make[Key, A]( | 
| 29 | entries: List[((Key, A), List[Key])], | |
| 76840 | 30 | converse: Boolean = false)( | 
| 75393 | 31 | implicit ord: Ordering[Key] | 
| 32 |   ): Graph[Key, A] = {
 | |
| 49560 | 33 | val graph1 = | 
| 73359 | 34 |       entries.foldLeft(empty[Key, A](ord)) {
 | 
| 35 | case (graph, ((x, info), _)) => graph.new_node(x, info) | |
| 36 | } | |
| 49560 | 37 | val graph2 = | 
| 73359 | 38 |       entries.foldLeft(graph1) {
 | 
| 39 | case (graph, ((x, _), ys)) => | |
| 40 |           ys.foldLeft(graph) {
 | |
| 76840 | 41 | case (g, y) => if (converse) g.add_edge(y, x) else g.add_edge(x, y) | 
| 73359 | 42 | } | 
| 43 | } | |
| 49560 | 44 | graph2 | 
| 45 | } | |
| 46 | ||
| 46667 | 47 | def string[A]: Graph[String, A] = empty(Ordering.String) | 
| 48 | def int[A]: Graph[Int, A] = empty(Ordering.Int) | |
| 49 | def long[A]: Graph[Long, A] = empty(Ordering.Long) | |
| 49560 | 50 | |
| 51 | ||
| 52 | /* XML data representation */ | |
| 53 | ||
| 54 | def encode[Key, A](key: XML.Encode.T[Key], info: XML.Encode.T[A]): XML.Encode.T[Graph[Key, A]] = | |
| 75394 | 55 |     { (graph: Graph[Key, A]) =>
 | 
| 49560 | 56 | import XML.Encode._ | 
| 57 | list(pair(pair(key, info), list(key)))(graph.dest) | |
| 60215 | 58 | } | 
| 49560 | 59 | |
| 59245 
be4180f3c236
more formal Graph_Display.Node (with ordering) and Graph_Display.Edge;
 wenzelm parents: 
59212diff
changeset | 60 | def decode[Key, A](key: XML.Decode.T[Key], info: XML.Decode.T[A])( | 
| 59212 | 61 | implicit ord: Ordering[Key]): XML.Decode.T[Graph[Key, A]] = | 
| 75394 | 62 |     { (body: XML.Body) =>
 | 
| 49560 | 63 | import XML.Decode._ | 
| 59245 
be4180f3c236
more formal Graph_Display.Node (with ordering) and Graph_Display.Edge;
 wenzelm parents: 
59212diff
changeset | 64 | make(list(pair(pair(key, info), list(key)))(body))(ord) | 
| 60215 | 65 | } | 
| 46611 | 66 | } | 
| 67 | ||
| 68 | ||
| 78235 | 69 | final class Graph[Key, A] private( | 
| 70 | protected val rep: SortedMap[Key, (A, (SortedSet[Key], SortedSet[Key]))] | |
| 71 | ) {
 | |
| 72 | override def hashCode: Int = rep.hashCode | |
| 73 | override def equals(that: Any): Boolean = | |
| 74 |     that match {
 | |
| 75 | case other: Graph[_, _] => rep == other.rep | |
| 76 | case _ => false | |
| 77 | } | |
| 78 | ||
| 46661 
d2ac78ba805e
prefer sorted Map/Set for canonical order of results -- pass ordering via fresh copy of empty;
 wenzelm parents: 
46659diff
changeset | 79 | type Keys = SortedSet[Key] | 
| 46611 | 80 | type Entry = (A, (Keys, Keys)) | 
| 81 | ||
| 46661 
d2ac78ba805e
prefer sorted Map/Set for canonical order of results -- pass ordering via fresh copy of empty;
 wenzelm parents: 
46659diff
changeset | 82 | def ordering: Ordering[Key] = rep.ordering | 
| 
d2ac78ba805e
prefer sorted Map/Set for canonical order of results -- pass ordering via fresh copy of empty;
 wenzelm parents: 
46659diff
changeset | 83 | def empty_keys: Keys = SortedSet.empty[Key](ordering) | 
| 
d2ac78ba805e
prefer sorted Map/Set for canonical order of results -- pass ordering via fresh copy of empty;
 wenzelm parents: 
46659diff
changeset | 84 | |
| 46666 
b01b6977a5e8
clarified signature -- avoid oddities of Iterable like Iterator.map;
 wenzelm parents: 
46661diff
changeset | 85 | |
| 
b01b6977a5e8
clarified signature -- avoid oddities of Iterable like Iterator.map;
 wenzelm parents: 
46661diff
changeset | 86 | /* graphs */ | 
| 46611 | 87 | |
| 88 | def is_empty: Boolean = rep.isEmpty | |
| 48507 | 89 | def defined(x: Key): Boolean = rep.isDefinedAt(x) | 
| 68321 | 90 | def domain: Set[Key] = rep.keySet | 
| 46611 | 91 | |
| 68496 | 92 | def size: Int = rep.size | 
| 56372 
fadb0fef09d7
more explicit iterator terminology, in accordance to Scala 2.8 library;
 wenzelm parents: 
50452diff
changeset | 93 | def iterator: Iterator[(Key, Entry)] = rep.iterator | 
| 
fadb0fef09d7
more explicit iterator terminology, in accordance to Scala 2.8 library;
 wenzelm parents: 
50452diff
changeset | 94 | |
| 
fadb0fef09d7
more explicit iterator terminology, in accordance to Scala 2.8 library;
 wenzelm parents: 
50452diff
changeset | 95 | def keys_iterator: Iterator[Key] = iterator.map(_._1) | 
| 
fadb0fef09d7
more explicit iterator terminology, in accordance to Scala 2.8 library;
 wenzelm parents: 
50452diff
changeset | 96 | def keys: List[Key] = keys_iterator.toList | 
| 46611 | 97 | |
| 49560 | 98 | def dest: List[((Key, A), List[Key])] = | 
| 56372 
fadb0fef09d7
more explicit iterator terminology, in accordance to Scala 2.8 library;
 wenzelm parents: 
50452diff
changeset | 99 | (for ((x, (i, (_, succs))) <- iterator) yield ((x, i), succs.toList)).toList | 
| 46611 | 100 | |
| 46666 
b01b6977a5e8
clarified signature -- avoid oddities of Iterable like Iterator.map;
 wenzelm parents: 
46661diff
changeset | 101 | override def toString: String = | 
| 49560 | 102 |     dest.map({ case ((x, _), ys) =>
 | 
| 103 |         x.toString + " -> " + ys.iterator.map(_.toString).mkString("{", ", ", "}") })
 | |
| 46666 
b01b6977a5e8
clarified signature -- avoid oddities of Iterable like Iterator.map;
 wenzelm parents: 
46661diff
changeset | 104 |       .mkString("Graph(", ", ", ")")
 | 
| 46611 | 105 | |
| 106 | private def get_entry(x: Key): Entry = | |
| 107 |     rep.get(x) match {
 | |
| 108 | case Some(entry) => entry | |
| 109 | case None => throw new Graph.Undefined(x) | |
| 110 | } | |
| 111 | ||
| 112 | private def map_entry(x: Key, f: Entry => Entry): Graph[Key, A] = | |
| 113 | new Graph[Key, A](rep + (x -> f(get_entry(x)))) | |
| 114 | ||
| 115 | ||
| 116 | /* nodes */ | |
| 117 | ||
| 118 | def get_node(x: Key): A = get_entry(x)._1 | |
| 119 | ||
| 120 | def map_node(x: Key, f: A => A): Graph[Key, A] = | |
| 121 |     map_entry(x, { case (i, ps) => (f(i), ps) })
 | |
| 122 | ||
| 123 | ||
| 124 | /* reachability */ | |
| 125 | ||
| 70764 
6d36b30fdd67
more operations -- incremental exploration of reachable nodes;
 wenzelm parents: 
70699diff
changeset | 126 | /*reachable nodes with length of longest path*/ | 
| 70794 | 127 | def reachable_length( | 
| 70800 | 128 | count: Key => Long, | 
| 70794 | 129 | next: Key => Keys, | 
| 75393 | 130 | xs: List[Key] | 
| 131 |   ): Map[Key, Long] = {
 | |
| 70800 | 132 | def reach(length: Long)(rs: Map[Key, Long], x: Key): Map[Key, Long] = | 
| 70764 
6d36b30fdd67
more operations -- incremental exploration of reachable nodes;
 wenzelm parents: 
70699diff
changeset | 133 |       if (rs.get(x) match { case Some(n) => n >= length case None => false }) rs
 | 
| 73359 | 134 | else next(x).foldLeft(rs + (x -> length))(reach(length + count(x))) | 
| 135 | xs.foldLeft(Map.empty[Key, Long])(reach(0)) | |
| 70764 
6d36b30fdd67
more operations -- incremental exploration of reachable nodes;
 wenzelm parents: 
70699diff
changeset | 136 | } | 
| 70800 | 137 | def node_height(count: Key => Long): Map[Key, Long] = | 
| 71601 | 138 | reachable_length(count, imm_preds, maximals) | 
| 70800 | 139 | def node_depth(count: Key => Long): Map[Key, Long] = | 
| 71601 | 140 | reachable_length(count, imm_succs, minimals) | 
| 70764 
6d36b30fdd67
more operations -- incremental exploration of reachable nodes;
 wenzelm parents: 
70699diff
changeset | 141 | |
| 
6d36b30fdd67
more operations -- incremental exploration of reachable nodes;
 wenzelm parents: 
70699diff
changeset | 142 | /*reachable nodes with size limit*/ | 
| 70772 
030a6baa5cb2
support headless_load_limit for more scalable load process;
 wenzelm parents: 
70764diff
changeset | 143 | def reachable_limit( | 
| 70800 | 144 | limit: Long, | 
| 145 | count: Key => Long, | |
| 70772 
030a6baa5cb2
support headless_load_limit for more scalable load process;
 wenzelm parents: 
70764diff
changeset | 146 | next: Key => Keys, | 
| 75393 | 147 | xs: List[Key] | 
| 148 |   ): Keys = {
 | |
| 149 |     def reach(res: (Long, Keys), x: Key): (Long, Keys) = {
 | |
| 70764 
6d36b30fdd67
more operations -- incremental exploration of reachable nodes;
 wenzelm parents: 
70699diff
changeset | 150 | val (n, rs) = res | 
| 70772 
030a6baa5cb2
support headless_load_limit for more scalable load process;
 wenzelm parents: 
70764diff
changeset | 151 | if (rs.contains(x)) res | 
| 73359 | 152 | else next(x).foldLeft((n + count(x), rs + x))(reach) | 
| 70764 
6d36b30fdd67
more operations -- incremental exploration of reachable nodes;
 wenzelm parents: 
70699diff
changeset | 153 | } | 
| 
6d36b30fdd67
more operations -- incremental exploration of reachable nodes;
 wenzelm parents: 
70699diff
changeset | 154 | |
| 70800 | 155 | @tailrec def reachs(size: Long, rs: Keys, rest: List[Key]): Keys = | 
| 70764 
6d36b30fdd67
more operations -- incremental exploration of reachable nodes;
 wenzelm parents: 
70699diff
changeset | 156 |       rest match {
 | 
| 70772 
030a6baa5cb2
support headless_load_limit for more scalable load process;
 wenzelm parents: 
70764diff
changeset | 157 | case Nil => rs | 
| 70764 
6d36b30fdd67
more operations -- incremental exploration of reachable nodes;
 wenzelm parents: 
70699diff
changeset | 158 | case head :: tail => | 
| 
6d36b30fdd67
more operations -- incremental exploration of reachable nodes;
 wenzelm parents: 
70699diff
changeset | 159 | val (size1, rs1) = reach((size, rs), head) | 
| 70772 
030a6baa5cb2
support headless_load_limit for more scalable load process;
 wenzelm parents: 
70764diff
changeset | 160 | if (size > 0 && size1 > limit) rs | 
| 70764 
6d36b30fdd67
more operations -- incremental exploration of reachable nodes;
 wenzelm parents: 
70699diff
changeset | 161 | else reachs(size1, rs1, tail) | 
| 
6d36b30fdd67
more operations -- incremental exploration of reachable nodes;
 wenzelm parents: 
70699diff
changeset | 162 | } | 
| 
6d36b30fdd67
more operations -- incremental exploration of reachable nodes;
 wenzelm parents: 
70699diff
changeset | 163 | |
| 
6d36b30fdd67
more operations -- incremental exploration of reachable nodes;
 wenzelm parents: 
70699diff
changeset | 164 | reachs(0, empty_keys, xs) | 
| 
6d36b30fdd67
more operations -- incremental exploration of reachable nodes;
 wenzelm parents: 
70699diff
changeset | 165 | } | 
| 
6d36b30fdd67
more operations -- incremental exploration of reachable nodes;
 wenzelm parents: 
70699diff
changeset | 166 | |
| 
6d36b30fdd67
more operations -- incremental exploration of reachable nodes;
 wenzelm parents: 
70699diff
changeset | 167 | /*reachable nodes with result in topological order (for acyclic graphs)*/ | 
| 75393 | 168 | private def reachable( | 
| 169 | next: Key => Keys, | |
| 170 | xs: List[Key], | |
| 171 | rev: Boolean = false | |
| 172 |   ): (List[List[Key]], Keys) = {
 | |
| 173 |     def reach(x: Key, reached: (List[Key], Keys)): (List[Key], Keys) = {
 | |
| 46611 | 174 | val (rs, r_set) = reached | 
| 175 | if (r_set(x)) reached | |
| 176 |       else {
 | |
| 72065 
11dc8929832d
clarified order --- proper sorting of requirements;
 wenzelm parents: 
71601diff
changeset | 177 | val (rs1, r_set1) = | 
| 73359 | 178 |           if (rev) next(x).foldLeft((rs, r_set + x)) { case (b, a) => reach(a, b) }
 | 
| 73360 | 179 | else next(x).foldRight((rs, r_set + x))(reach) | 
| 46611 | 180 | (x :: rs1, r_set1) | 
| 181 | } | |
| 182 | } | |
| 75393 | 183 |     def reachs(reached: (List[List[Key]], Keys), x: Key): (List[List[Key]], Keys) = {
 | 
| 46611 | 184 | val (rss, r_set) = reached | 
| 48350 
09bf3b73e446
clarified topological ordering: preserve order of adjacency via reverse fold;
 wenzelm parents: 
48348diff
changeset | 185 | val (rs, r_set1) = reach(x, (Nil, r_set)) | 
| 46611 | 186 | (rs :: rss, r_set1) | 
| 187 | } | |
| 73359 | 188 | xs.foldLeft((List.empty[List[Key]], empty_keys))(reachs) | 
| 46611 | 189 | } | 
| 190 | ||
| 191 | /*immediate*/ | |
| 192 | def imm_preds(x: Key): Keys = get_entry(x)._2._1 | |
| 193 | def imm_succs(x: Key): Keys = get_entry(x)._2._2 | |
| 194 | ||
| 195 | /*transitive*/ | |
| 72065 
11dc8929832d
clarified order --- proper sorting of requirements;
 wenzelm parents: 
71601diff
changeset | 196 | def all_preds_rev(xs: List[Key]): List[Key] = | 
| 
11dc8929832d
clarified order --- proper sorting of requirements;
 wenzelm parents: 
71601diff
changeset | 197 | reachable(imm_preds, xs, rev = true)._1.flatten.reverse | 
| 46611 | 198 | def all_preds(xs: List[Key]): List[Key] = reachable(imm_preds, xs)._1.flatten | 
| 199 | def all_succs(xs: List[Key]): List[Key] = reachable(imm_succs, xs)._1.flatten | |
| 200 | ||
| 46613 | 201 | /*strongly connected components; see: David King and John Launchbury, | 
| 202 | "Structuring Depth First Search Algorithms in Haskell"*/ | |
| 203 | def strong_conn: List[List[Key]] = | |
| 56372 
fadb0fef09d7
more explicit iterator terminology, in accordance to Scala 2.8 library;
 wenzelm parents: 
50452diff
changeset | 204 | reachable(imm_preds, all_succs(keys))._1.filterNot(_.isEmpty).reverse | 
| 46613 | 205 | |
| 46611 | 206 | |
| 207 | /* minimal and maximal elements */ | |
| 208 | ||
| 209 | def minimals: List[Key] = | |
| 73359 | 210 |     rep.foldLeft(List.empty[Key]) {
 | 
| 211 | case (ms, (m, (_, (preds, _)))) => if (preds.isEmpty) m :: ms else ms | |
| 212 | } | |
| 46611 | 213 | |
| 214 | def maximals: List[Key] = | |
| 73359 | 215 |     rep.foldLeft(List.empty[Key]) {
 | 
| 216 | case (ms, (m, (_, (_, succs)))) => if (succs.isEmpty) m :: ms else ms | |
| 217 | } | |
| 46611 | 218 | |
| 219 | def is_minimal(x: Key): Boolean = imm_preds(x).isEmpty | |
| 220 | def is_maximal(x: Key): Boolean = imm_succs(x).isEmpty | |
| 221 | ||
| 66830 | 222 | def is_isolated(x: Key): Boolean = is_minimal(x) && is_maximal(x) | 
| 223 | ||
| 46611 | 224 | |
| 46668 | 225 | /* node operations */ | 
| 46611 | 226 | |
| 75393 | 227 |   def new_node(x: Key, info: A): Graph[Key, A] = {
 | 
| 48648 | 228 | if (defined(x)) throw new Graph.Duplicate(x) | 
| 46661 
d2ac78ba805e
prefer sorted Map/Set for canonical order of results -- pass ordering via fresh copy of empty;
 wenzelm parents: 
46659diff
changeset | 229 | else new Graph[Key, A](rep + (x -> (info, (empty_keys, empty_keys)))) | 
| 46611 | 230 | } | 
| 231 | ||
| 46613 | 232 | def default_node(x: Key, info: A): Graph[Key, A] = | 
| 48648 | 233 | if (defined(x)) this else new_node(x, info) | 
| 46613 | 234 | |
| 46661 
d2ac78ba805e
prefer sorted Map/Set for canonical order of results -- pass ordering via fresh copy of empty;
 wenzelm parents: 
46659diff
changeset | 235 | private def del_adjacent(fst: Boolean, x: Key)(map: SortedMap[Key, Entry], y: Key) | 
| 
d2ac78ba805e
prefer sorted Map/Set for canonical order of results -- pass ordering via fresh copy of empty;
 wenzelm parents: 
46659diff
changeset | 236 | : SortedMap[Key, Entry] = | 
| 46611 | 237 |     map.get(y) match {
 | 
| 238 | case None => map | |
| 239 | case Some((i, (preds, succs))) => | |
| 240 | map + (y -> (i, if (fst) (preds - x, succs) else (preds, succs - x))) | |
| 241 | } | |
| 242 | ||
| 75393 | 243 |   def del_node(x: Key): Graph[Key, A] = {
 | 
| 46611 | 244 | val (preds, succs) = get_entry(x)._2 | 
| 245 | new Graph[Key, A]( | |
| 73359 | 246 | succs.foldLeft(preds.foldLeft(rep - x)(del_adjacent(false, x)))(del_adjacent(true, x))) | 
| 46611 | 247 | } | 
| 248 | ||
| 46614 
165886a4fe64
clarified Graph.restrict (formerly Graph.subgraph) based on public graph operations;
 wenzelm parents: 
46613diff
changeset | 249 | def restrict(pred: Key => Boolean): Graph[Key, A] = | 
| 73359 | 250 |     iterator.foldLeft(this) { case (graph, (x, _)) => if (!pred(x)) graph.del_node(x) else graph }
 | 
| 46614 
165886a4fe64
clarified Graph.restrict (formerly Graph.subgraph) based on public graph operations;
 wenzelm parents: 
46613diff
changeset | 251 | |
| 70699 | 252 | def exclude(pred: Key => Boolean): Graph[Key, A] = restrict(name => !pred(name)) | 
| 253 | ||
| 46611 | 254 | |
| 46668 | 255 | /* edge operations */ | 
| 46611 | 256 | |
| 59259 
399506ee38a5
clarified static full_graph vs. dynamic visible_graph;
 wenzelm parents: 
59245diff
changeset | 257 | def edges_iterator: Iterator[(Key, Key)] = | 
| 
399506ee38a5
clarified static full_graph vs. dynamic visible_graph;
 wenzelm parents: 
59245diff
changeset | 258 |     for { x <- keys_iterator; y <- imm_succs(x).iterator } yield (x, y)
 | 
| 
399506ee38a5
clarified static full_graph vs. dynamic visible_graph;
 wenzelm parents: 
59245diff
changeset | 259 | |
| 46611 | 260 | def is_edge(x: Key, y: Key): Boolean = | 
| 48507 | 261 | defined(x) && defined(y) && imm_succs(x)(y) | 
| 46611 | 262 | |
| 263 | def add_edge(x: Key, y: Key): Graph[Key, A] = | |
| 264 | if (is_edge(x, y)) this | |
| 265 | else | |
| 266 |       map_entry(y, { case (i, (preds, succs)) => (i, (preds + x, succs)) }).
 | |
| 267 |       map_entry(x, { case (i, (preds, succs)) => (i, (preds, succs + y)) })
 | |
| 268 | ||
| 269 | def del_edge(x: Key, y: Key): Graph[Key, A] = | |
| 270 | if (is_edge(x, y)) | |
| 271 |       map_entry(y, { case (i, (preds, succs)) => (i, (preds - x, succs)) }).
 | |
| 272 |       map_entry(x, { case (i, (preds, succs)) => (i, (preds, succs - y)) })
 | |
| 273 | else this | |
| 274 | ||
| 275 | ||
| 276 | /* irreducible paths -- Hasse diagram */ | |
| 277 | ||
| 75393 | 278 |   private def irreducible_preds(x_set: Keys, path: List[Key], z: Key): List[Key] = {
 | 
| 46611 | 279 | def red(x: Key)(x1: Key) = is_edge(x, x1) && x1 != z | 
| 280 | @tailrec def irreds(xs0: List[Key], xs1: List[Key]): List[Key] = | |
| 281 |       xs0 match {
 | |
| 282 | case Nil => xs1 | |
| 283 | case x :: xs => | |
| 60215 | 284 | if (!x_set(x) || x == z || path.contains(x) || | 
| 46611 | 285 | xs.exists(red(x)) || xs1.exists(red(x))) | 
| 286 | irreds(xs, xs1) | |
| 287 | else irreds(xs, x :: xs1) | |
| 288 | } | |
| 289 | irreds(imm_preds(z).toList, Nil) | |
| 290 | } | |
| 291 | ||
| 75393 | 292 |   def irreducible_paths(x: Key, y: Key): List[List[Key]] = {
 | 
| 46611 | 293 | val (_, x_set) = reachable(imm_succs, List(x)) | 
| 294 | def paths(path: List[Key])(ps: List[List[Key]], z: Key): List[List[Key]] = | |
| 295 | if (x == z) (z :: path) :: ps | |
| 73359 | 296 | else irreducible_preds(x_set, path, z).foldLeft(ps)(paths(z :: path)) | 
| 46611 | 297 | if ((x == y) && !is_edge(x, x)) List(Nil) else paths(Nil)(Nil, y) | 
| 298 | } | |
| 299 | ||
| 300 | ||
| 50445 
68c9a6538c0e
added graph operations for transitive closure and reduction in Scala -- unproven and thus better left out of the kernel-relevant ML module;
 wenzelm parents: 
49560diff
changeset | 301 | /* transitive closure and reduction */ | 
| 
68c9a6538c0e
added graph operations for transitive closure and reduction in Scala -- unproven and thus better left out of the kernel-relevant ML module;
 wenzelm parents: 
49560diff
changeset | 302 | |
| 75393 | 303 |   private def transitive_step(z: Key): Graph[Key, A] = {
 | 
| 50447 
2e22cdccdc38
clarified transitive_closure: proper cumulation of transitive steps, which is essential for Warshall-style algorithms;
 wenzelm parents: 
50445diff
changeset | 304 | val (preds, succs) = get_entry(z)._2 | 
| 50445 
68c9a6538c0e
added graph operations for transitive closure and reduction in Scala -- unproven and thus better left out of the kernel-relevant ML module;
 wenzelm parents: 
49560diff
changeset | 305 | var graph = this | 
| 50447 
2e22cdccdc38
clarified transitive_closure: proper cumulation of transitive steps, which is essential for Warshall-style algorithms;
 wenzelm parents: 
50445diff
changeset | 306 | for (x <- preds; y <- succs) graph = graph.add_edge(x, y) | 
| 50445 
68c9a6538c0e
added graph operations for transitive closure and reduction in Scala -- unproven and thus better left out of the kernel-relevant ML module;
 wenzelm parents: 
49560diff
changeset | 307 | graph | 
| 
68c9a6538c0e
added graph operations for transitive closure and reduction in Scala -- unproven and thus better left out of the kernel-relevant ML module;
 wenzelm parents: 
49560diff
changeset | 308 | } | 
| 
68c9a6538c0e
added graph operations for transitive closure and reduction in Scala -- unproven and thus better left out of the kernel-relevant ML module;
 wenzelm parents: 
49560diff
changeset | 309 | |
| 73359 | 310 | def transitive_closure: Graph[Key, A] = keys_iterator.foldLeft(this)(_.transitive_step(_)) | 
| 50447 
2e22cdccdc38
clarified transitive_closure: proper cumulation of transitive steps, which is essential for Warshall-style algorithms;
 wenzelm parents: 
50445diff
changeset | 311 | |
| 75393 | 312 |   def transitive_reduction_acyclic: Graph[Key, A] = {
 | 
| 50445 
68c9a6538c0e
added graph operations for transitive closure and reduction in Scala -- unproven and thus better left out of the kernel-relevant ML module;
 wenzelm parents: 
49560diff
changeset | 313 | val trans = this.transitive_closure | 
| 56372 
fadb0fef09d7
more explicit iterator terminology, in accordance to Scala 2.8 library;
 wenzelm parents: 
50452diff
changeset | 314 |     if (trans.iterator.exists({ case (x, (_, (_, succs))) => succs.contains(x) }))
 | 
| 50452 
bfb5964e3041
stateless dockable window for graphview, which is triggered by the active area of the corresponding diagnostic command;
 wenzelm parents: 
50447diff
changeset | 315 |       error("Cyclic graph")
 | 
| 50445 
68c9a6538c0e
added graph operations for transitive closure and reduction in Scala -- unproven and thus better left out of the kernel-relevant ML module;
 wenzelm parents: 
49560diff
changeset | 316 | |
| 
68c9a6538c0e
added graph operations for transitive closure and reduction in Scala -- unproven and thus better left out of the kernel-relevant ML module;
 wenzelm parents: 
49560diff
changeset | 317 | var graph = this | 
| 
68c9a6538c0e
added graph operations for transitive closure and reduction in Scala -- unproven and thus better left out of the kernel-relevant ML module;
 wenzelm parents: 
49560diff
changeset | 318 |     for {
 | 
| 56372 
fadb0fef09d7
more explicit iterator terminology, in accordance to Scala 2.8 library;
 wenzelm parents: 
50452diff
changeset | 319 | (x, (_, (_, succs))) <- iterator | 
| 50445 
68c9a6538c0e
added graph operations for transitive closure and reduction in Scala -- unproven and thus better left out of the kernel-relevant ML module;
 wenzelm parents: 
49560diff
changeset | 320 | y <- succs | 
| 
68c9a6538c0e
added graph operations for transitive closure and reduction in Scala -- unproven and thus better left out of the kernel-relevant ML module;
 wenzelm parents: 
49560diff
changeset | 321 | if trans.imm_preds(y).exists(z => trans.is_edge(x, z)) | 
| 
68c9a6538c0e
added graph operations for transitive closure and reduction in Scala -- unproven and thus better left out of the kernel-relevant ML module;
 wenzelm parents: 
49560diff
changeset | 322 | } graph = graph.del_edge(x, y) | 
| 
68c9a6538c0e
added graph operations for transitive closure and reduction in Scala -- unproven and thus better left out of the kernel-relevant ML module;
 wenzelm parents: 
49560diff
changeset | 323 | graph | 
| 
68c9a6538c0e
added graph operations for transitive closure and reduction in Scala -- unproven and thus better left out of the kernel-relevant ML module;
 wenzelm parents: 
49560diff
changeset | 324 | } | 
| 
68c9a6538c0e
added graph operations for transitive closure and reduction in Scala -- unproven and thus better left out of the kernel-relevant ML module;
 wenzelm parents: 
49560diff
changeset | 325 | |
| 
68c9a6538c0e
added graph operations for transitive closure and reduction in Scala -- unproven and thus better left out of the kernel-relevant ML module;
 wenzelm parents: 
49560diff
changeset | 326 | |
| 46611 | 327 | /* maintain acyclic graphs */ | 
| 328 | ||
| 329 | def add_edge_acyclic(x: Key, y: Key): Graph[Key, A] = | |
| 330 | if (is_edge(x, y)) this | |
| 331 |     else {
 | |
| 332 |       irreducible_paths(y, x) match {
 | |
| 333 | case Nil => add_edge(x, y) | |
| 334 | case cycles => throw new Graph.Cycles(cycles.map(x :: _)) | |
| 335 | } | |
| 336 | } | |
| 337 | ||
| 48348 | 338 | def add_deps_acyclic(y: Key, xs: List[Key]): Graph[Key, A] = | 
| 73359 | 339 | xs.foldLeft(this)(_.add_edge_acyclic(_, y)) | 
| 46611 | 340 | |
| 341 | def topological_order: List[Key] = all_succs(minimals) | |
| 342 | } |