| author | wenzelm | 
| Sat, 01 Mar 2014 18:33:49 +0100 | |
| changeset 55824 | 22bc50a19afa | 
| parent 50452 | bfb5964e3041 | 
| child 56372 | fadb0fef09d7 | 
| permissions | -rw-r--r-- | 
| 46611 | 1 | /* Title: Pure/General/graph.scala | 
| 2 | Module: PIDE | |
| 3 | Author: Makarius | |
| 4 | ||
| 5 | Directed graphs. | |
| 6 | */ | |
| 7 | ||
| 8 | package isabelle | |
| 9 | ||
| 10 | ||
| 46661 
d2ac78ba805e
prefer sorted Map/Set for canonical order of results -- pass ordering via fresh copy of empty;
 wenzelm parents: 
46659diff
changeset | 11 | import scala.collection.immutable.{SortedMap, SortedSet}
 | 
| 46611 | 12 | import scala.annotation.tailrec | 
| 13 | ||
| 14 | ||
| 15 | object Graph | |
| 16 | {
 | |
| 48348 | 17 | class Duplicate[Key](val key: Key) extends Exception | 
| 18 | class Undefined[Key](val key: Key) extends Exception | |
| 19 | class Cycles[Key](val cycles: List[List[Key]]) extends Exception | |
| 46611 | 20 | |
| 46661 
d2ac78ba805e
prefer sorted Map/Set for canonical order of results -- pass ordering via fresh copy of empty;
 wenzelm parents: 
46659diff
changeset | 21 | 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 | 22 | new Graph[Key, A](SortedMap.empty(ord)) | 
| 46667 | 23 | |
| 49560 | 24 | def make[Key, A](entries: List[((Key, A), List[Key])])(implicit ord: Ordering[Key]) | 
| 25 | : Graph[Key, A] = | |
| 26 |   {
 | |
| 27 | val graph1 = | |
| 28 |       (empty[Key, A](ord) /: entries) { case (graph, ((x, info), _)) => graph.new_node(x, info) }
 | |
| 29 | val graph2 = | |
| 30 |       (graph1 /: entries) { case (graph, ((x, _), ys)) => (graph /: ys)(_.add_edge(x, _)) }
 | |
| 31 | graph2 | |
| 32 | } | |
| 33 | ||
| 46667 | 34 | def string[A]: Graph[String, A] = empty(Ordering.String) | 
| 35 | def int[A]: Graph[Int, A] = empty(Ordering.Int) | |
| 36 | def long[A]: Graph[Long, A] = empty(Ordering.Long) | |
| 49560 | 37 | |
| 38 | ||
| 39 | /* XML data representation */ | |
| 40 | ||
| 41 | def encode[Key, A](key: XML.Encode.T[Key], info: XML.Encode.T[A]): XML.Encode.T[Graph[Key, A]] = | |
| 42 |     ((graph: Graph[Key, A]) => {
 | |
| 43 | import XML.Encode._ | |
| 44 | list(pair(pair(key, info), list(key)))(graph.dest) | |
| 45 | }) | |
| 46 | ||
| 47 | def decode[Key, A](key: XML.Decode.T[Key], info: XML.Decode.T[A])(implicit ord: Ordering[Key]) | |
| 48 | : XML.Decode.T[Graph[Key, A]] = | |
| 49 |     ((body: XML.Body) => {
 | |
| 50 | import XML.Decode._ | |
| 51 | make(list(pair(pair(key, info), list(key)))(body))(ord) | |
| 52 | }) | |
| 46611 | 53 | } | 
| 54 | ||
| 55 | ||
| 46712 | 56 | final class Graph[Key, A] private(rep: SortedMap[Key, (A, (SortedSet[Key], SortedSet[Key]))]) | 
| 46611 | 57 | {
 | 
| 46661 
d2ac78ba805e
prefer sorted Map/Set for canonical order of results -- pass ordering via fresh copy of empty;
 wenzelm parents: 
46659diff
changeset | 58 | type Keys = SortedSet[Key] | 
| 46611 | 59 | type Entry = (A, (Keys, Keys)) | 
| 60 | ||
| 46661 
d2ac78ba805e
prefer sorted Map/Set for canonical order of results -- pass ordering via fresh copy of empty;
 wenzelm parents: 
46659diff
changeset | 61 | 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 | 62 | 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 | 63 | |
| 46666 
b01b6977a5e8
clarified signature -- avoid oddities of Iterable like Iterator.map;
 wenzelm parents: 
46661diff
changeset | 64 | |
| 
b01b6977a5e8
clarified signature -- avoid oddities of Iterable like Iterator.map;
 wenzelm parents: 
46661diff
changeset | 65 | /* graphs */ | 
| 46611 | 66 | |
| 67 | def is_empty: Boolean = rep.isEmpty | |
| 48507 | 68 | def defined(x: Key): Boolean = rep.isDefinedAt(x) | 
| 46611 | 69 | |
| 46666 
b01b6977a5e8
clarified signature -- avoid oddities of Iterable like Iterator.map;
 wenzelm parents: 
46661diff
changeset | 70 | def entries: Iterator[(Key, Entry)] = rep.iterator | 
| 
b01b6977a5e8
clarified signature -- avoid oddities of Iterable like Iterator.map;
 wenzelm parents: 
46661diff
changeset | 71 | def keys: Iterator[Key] = entries.map(_._1) | 
| 46611 | 72 | |
| 49560 | 73 | def dest: List[((Key, A), List[Key])] = | 
| 74 | (for ((x, (i, (_, succs))) <- entries) yield ((x, i), succs.toList)).toList | |
| 46611 | 75 | |
| 46666 
b01b6977a5e8
clarified signature -- avoid oddities of Iterable like Iterator.map;
 wenzelm parents: 
46661diff
changeset | 76 | override def toString: String = | 
| 49560 | 77 |     dest.map({ case ((x, _), ys) =>
 | 
| 78 |         x.toString + " -> " + ys.iterator.map(_.toString).mkString("{", ", ", "}") })
 | |
| 46666 
b01b6977a5e8
clarified signature -- avoid oddities of Iterable like Iterator.map;
 wenzelm parents: 
46661diff
changeset | 79 |       .mkString("Graph(", ", ", ")")
 | 
| 46611 | 80 | |
| 81 | private def get_entry(x: Key): Entry = | |
| 82 |     rep.get(x) match {
 | |
| 83 | case Some(entry) => entry | |
| 84 | case None => throw new Graph.Undefined(x) | |
| 85 | } | |
| 86 | ||
| 87 | private def map_entry(x: Key, f: Entry => Entry): Graph[Key, A] = | |
| 88 | new Graph[Key, A](rep + (x -> f(get_entry(x)))) | |
| 89 | ||
| 90 | ||
| 91 | /* nodes */ | |
| 92 | ||
| 93 | def get_node(x: Key): A = get_entry(x)._1 | |
| 94 | ||
| 95 | def map_node(x: Key, f: A => A): Graph[Key, A] = | |
| 96 |     map_entry(x, { case (i, ps) => (f(i), ps) })
 | |
| 97 | ||
| 98 | ||
| 99 | /* reachability */ | |
| 100 | ||
| 101 | /*nodes reachable from xs -- topologically sorted for acyclic graphs*/ | |
| 102 | def reachable(next: Key => Keys, xs: List[Key]): (List[List[Key]], Keys) = | |
| 103 |   {
 | |
| 48350 
09bf3b73e446
clarified topological ordering: preserve order of adjacency via reverse fold;
 wenzelm parents: 
48348diff
changeset | 104 | def reach(x: Key, reached: (List[Key], Keys)): (List[Key], Keys) = | 
| 46611 | 105 |     {
 | 
| 106 | val (rs, r_set) = reached | |
| 107 | if (r_set(x)) reached | |
| 108 |       else {
 | |
| 48350 
09bf3b73e446
clarified topological ordering: preserve order of adjacency via reverse fold;
 wenzelm parents: 
48348diff
changeset | 109 | val (rs1, r_set1) = (next(x) :\ (rs, r_set + x))(reach) | 
| 46611 | 110 | (x :: rs1, r_set1) | 
| 111 | } | |
| 112 | } | |
| 113 | def reachs(reached: (List[List[Key]], Keys), x: Key): (List[List[Key]], Keys) = | |
| 114 |     {
 | |
| 115 | val (rss, r_set) = reached | |
| 48350 
09bf3b73e446
clarified topological ordering: preserve order of adjacency via reverse fold;
 wenzelm parents: 
48348diff
changeset | 116 | val (rs, r_set1) = reach(x, (Nil, r_set)) | 
| 46611 | 117 | (rs :: rss, r_set1) | 
| 118 | } | |
| 46661 
d2ac78ba805e
prefer sorted Map/Set for canonical order of results -- pass ordering via fresh copy of empty;
 wenzelm parents: 
46659diff
changeset | 119 | ((List.empty[List[Key]], empty_keys) /: xs)(reachs) | 
| 46611 | 120 | } | 
| 121 | ||
| 122 | /*immediate*/ | |
| 123 | def imm_preds(x: Key): Keys = get_entry(x)._2._1 | |
| 124 | def imm_succs(x: Key): Keys = get_entry(x)._2._2 | |
| 125 | ||
| 126 | /*transitive*/ | |
| 127 | def all_preds(xs: List[Key]): List[Key] = reachable(imm_preds, xs)._1.flatten | |
| 128 | def all_succs(xs: List[Key]): List[Key] = reachable(imm_succs, xs)._1.flatten | |
| 129 | ||
| 46613 | 130 | /*strongly connected components; see: David King and John Launchbury, | 
| 131 | "Structuring Depth First Search Algorithms in Haskell"*/ | |
| 132 | def strong_conn: List[List[Key]] = | |
| 46666 
b01b6977a5e8
clarified signature -- avoid oddities of Iterable like Iterator.map;
 wenzelm parents: 
46661diff
changeset | 133 | reachable(imm_preds, all_succs(keys.toList))._1.filterNot(_.isEmpty).reverse | 
| 46613 | 134 | |
| 46611 | 135 | |
| 136 | /* minimal and maximal elements */ | |
| 137 | ||
| 138 | def minimals: List[Key] = | |
| 139 |     (List.empty[Key] /: rep) {
 | |
| 140 | case (ms, (m, (_, (preds, _)))) => if (preds.isEmpty) m :: ms else ms } | |
| 141 | ||
| 142 | def maximals: List[Key] = | |
| 143 |     (List.empty[Key] /: rep) {
 | |
| 144 | case (ms, (m, (_, (_, succs)))) => if (succs.isEmpty) m :: ms else ms } | |
| 145 | ||
| 146 | def is_minimal(x: Key): Boolean = imm_preds(x).isEmpty | |
| 147 | def is_maximal(x: Key): Boolean = imm_succs(x).isEmpty | |
| 148 | ||
| 149 | ||
| 46668 | 150 | /* node operations */ | 
| 46611 | 151 | |
| 152 | def new_node(x: Key, info: A): Graph[Key, A] = | |
| 153 |   {
 | |
| 48648 | 154 | 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 | 155 | else new Graph[Key, A](rep + (x -> (info, (empty_keys, empty_keys)))) | 
| 46611 | 156 | } | 
| 157 | ||
| 46613 | 158 | def default_node(x: Key, info: A): Graph[Key, A] = | 
| 48648 | 159 | if (defined(x)) this else new_node(x, info) | 
| 46613 | 160 | |
| 46661 
d2ac78ba805e
prefer sorted Map/Set for canonical order of results -- pass ordering via fresh copy of empty;
 wenzelm parents: 
46659diff
changeset | 161 | 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 | 162 | : SortedMap[Key, Entry] = | 
| 46611 | 163 |     map.get(y) match {
 | 
| 164 | case None => map | |
| 165 | case Some((i, (preds, succs))) => | |
| 166 | map + (y -> (i, if (fst) (preds - x, succs) else (preds, succs - x))) | |
| 167 | } | |
| 168 | ||
| 169 | def del_node(x: Key): Graph[Key, A] = | |
| 170 |   {
 | |
| 171 | val (preds, succs) = get_entry(x)._2 | |
| 172 | new Graph[Key, A]( | |
| 173 | (((rep - x) /: preds)(del_adjacent(false, x)) /: succs)(del_adjacent(true, x))) | |
| 174 | } | |
| 175 | ||
| 46614 
165886a4fe64
clarified Graph.restrict (formerly Graph.subgraph) based on public graph operations;
 wenzelm parents: 
46613diff
changeset | 176 | def restrict(pred: Key => Boolean): Graph[Key, A] = | 
| 46666 
b01b6977a5e8
clarified signature -- avoid oddities of Iterable like Iterator.map;
 wenzelm parents: 
46661diff
changeset | 177 |     (this /: entries){ 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 | 178 | |
| 46611 | 179 | |
| 46668 | 180 | /* edge operations */ | 
| 46611 | 181 | |
| 182 | def is_edge(x: Key, y: Key): Boolean = | |
| 48507 | 183 | defined(x) && defined(y) && imm_succs(x)(y) | 
| 46611 | 184 | |
| 185 | def add_edge(x: Key, y: Key): Graph[Key, A] = | |
| 186 | if (is_edge(x, y)) this | |
| 187 | else | |
| 188 |       map_entry(y, { case (i, (preds, succs)) => (i, (preds + x, succs)) }).
 | |
| 189 |       map_entry(x, { case (i, (preds, succs)) => (i, (preds, succs + y)) })
 | |
| 190 | ||
| 191 | def del_edge(x: Key, y: Key): Graph[Key, A] = | |
| 192 | if (is_edge(x, y)) | |
| 193 |       map_entry(y, { case (i, (preds, succs)) => (i, (preds - x, succs)) }).
 | |
| 194 |       map_entry(x, { case (i, (preds, succs)) => (i, (preds, succs - y)) })
 | |
| 195 | else this | |
| 196 | ||
| 197 | ||
| 198 | /* irreducible paths -- Hasse diagram */ | |
| 199 | ||
| 46661 
d2ac78ba805e
prefer sorted Map/Set for canonical order of results -- pass ordering via fresh copy of empty;
 wenzelm parents: 
46659diff
changeset | 200 | private def irreducible_preds(x_set: Keys, path: List[Key], z: Key): List[Key] = | 
| 46611 | 201 |   {
 | 
| 202 | def red(x: Key)(x1: Key) = is_edge(x, x1) && x1 != z | |
| 203 | @tailrec def irreds(xs0: List[Key], xs1: List[Key]): List[Key] = | |
| 204 |       xs0 match {
 | |
| 205 | case Nil => xs1 | |
| 206 | case x :: xs => | |
| 207 | if (!(x_set(x)) || x == z || path.contains(x) || | |
| 208 | xs.exists(red(x)) || xs1.exists(red(x))) | |
| 209 | irreds(xs, xs1) | |
| 210 | else irreds(xs, x :: xs1) | |
| 211 | } | |
| 212 | irreds(imm_preds(z).toList, Nil) | |
| 213 | } | |
| 214 | ||
| 215 | def irreducible_paths(x: Key, y: Key): List[List[Key]] = | |
| 216 |   {
 | |
| 217 | val (_, x_set) = reachable(imm_succs, List(x)) | |
| 218 | def paths(path: List[Key])(ps: List[List[Key]], z: Key): List[List[Key]] = | |
| 219 | if (x == z) (z :: path) :: ps | |
| 220 | else (ps /: irreducible_preds(x_set, path, z))(paths(z :: path)) | |
| 221 | if ((x == y) && !is_edge(x, x)) List(Nil) else paths(Nil)(Nil, y) | |
| 222 | } | |
| 223 | ||
| 224 | ||
| 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 | 225 | /* 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 | 226 | |
| 50447 
2e22cdccdc38
clarified transitive_closure: proper cumulation of transitive steps, which is essential for Warshall-style algorithms;
 wenzelm parents: 
50445diff
changeset | 227 | private def transitive_step(z: Key): 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 | 228 |   {
 | 
| 50447 
2e22cdccdc38
clarified transitive_closure: proper cumulation of transitive steps, which is essential for Warshall-style algorithms;
 wenzelm parents: 
50445diff
changeset | 229 | 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 | 230 | var graph = this | 
| 50447 
2e22cdccdc38
clarified transitive_closure: proper cumulation of transitive steps, which is essential for Warshall-style algorithms;
 wenzelm parents: 
50445diff
changeset | 231 | 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 | 232 | 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 | 233 | } | 
| 
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 | 234 | |
| 50447 
2e22cdccdc38
clarified transitive_closure: proper cumulation of transitive steps, which is essential for Warshall-style algorithms;
 wenzelm parents: 
50445diff
changeset | 235 | def transitive_closure: Graph[Key, A] = (this /: keys)(_.transitive_step(_)) | 
| 
2e22cdccdc38
clarified transitive_closure: proper cumulation of transitive steps, which is essential for Warshall-style algorithms;
 wenzelm parents: 
50445diff
changeset | 236 | |
| 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 | 237 | def transitive_reduction_acyclic: Graph[Key, A] = | 
| 
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 | 238 |   {
 | 
| 
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 | 239 | val trans = this.transitive_closure | 
| 50452 
bfb5964e3041
stateless dockable window for graphview, which is triggered by the active area of the corresponding diagnostic command;
 wenzelm parents: 
50447diff
changeset | 240 |     if (trans.entries.exists({ case (x, (_, (_, succs))) => succs.contains(x) }))
 | 
| 
bfb5964e3041
stateless dockable window for graphview, which is triggered by the active area of the corresponding diagnostic command;
 wenzelm parents: 
50447diff
changeset | 241 |       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 | 242 | |
| 
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 | 243 | 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 | 244 |     for {
 | 
| 
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 | 245 | (x, (_, (_, succs))) <- this.entries | 
| 
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 | 246 | 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 | 247 | 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 | 248 | } 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 | 249 | 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 | 250 | } | 
| 
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 | 251 | |
| 
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 | 252 | |
| 46611 | 253 | /* maintain acyclic graphs */ | 
| 254 | ||
| 255 | def add_edge_acyclic(x: Key, y: Key): Graph[Key, A] = | |
| 256 | if (is_edge(x, y)) this | |
| 257 |     else {
 | |
| 258 |       irreducible_paths(y, x) match {
 | |
| 259 | case Nil => add_edge(x, y) | |
| 260 | case cycles => throw new Graph.Cycles(cycles.map(x :: _)) | |
| 261 | } | |
| 262 | } | |
| 263 | ||
| 48348 | 264 | def add_deps_acyclic(y: Key, xs: List[Key]): Graph[Key, A] = | 
| 46611 | 265 | (this /: xs)(_.add_edge_acyclic(_, y)) | 
| 266 | ||
| 267 | def topological_order: List[Key] = all_succs(minimals) | |
| 268 | } |