| author | wenzelm | 
| Mon, 10 Sep 2012 15:20:50 +0200 | |
| changeset 49245 | cb70157293c0 | 
| parent 48648 | f13eeeea1a69 | 
| child 49560 | 11430dd89e35 | 
| 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 | |
| 24 | def string[A]: Graph[String, A] = empty(Ordering.String) | |
| 25 | def int[A]: Graph[Int, A] = empty(Ordering.Int) | |
| 26 | def long[A]: Graph[Long, A] = empty(Ordering.Long) | |
| 46611 | 27 | } | 
| 28 | ||
| 29 | ||
| 46712 | 30 | final class Graph[Key, A] private(rep: SortedMap[Key, (A, (SortedSet[Key], SortedSet[Key]))]) | 
| 46611 | 31 | {
 | 
| 46661 
d2ac78ba805e
prefer sorted Map/Set for canonical order of results -- pass ordering via fresh copy of empty;
 wenzelm parents: 
46659diff
changeset | 32 | type Keys = SortedSet[Key] | 
| 46611 | 33 | type Entry = (A, (Keys, Keys)) | 
| 34 | ||
| 46661 
d2ac78ba805e
prefer sorted Map/Set for canonical order of results -- pass ordering via fresh copy of empty;
 wenzelm parents: 
46659diff
changeset | 35 | 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 | 36 | 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 | 37 | |
| 46666 
b01b6977a5e8
clarified signature -- avoid oddities of Iterable like Iterator.map;
 wenzelm parents: 
46661diff
changeset | 38 | |
| 
b01b6977a5e8
clarified signature -- avoid oddities of Iterable like Iterator.map;
 wenzelm parents: 
46661diff
changeset | 39 | /* graphs */ | 
| 46611 | 40 | |
| 41 | def is_empty: Boolean = rep.isEmpty | |
| 48507 | 42 | def defined(x: Key): Boolean = rep.isDefinedAt(x) | 
| 46611 | 43 | |
| 46666 
b01b6977a5e8
clarified signature -- avoid oddities of Iterable like Iterator.map;
 wenzelm parents: 
46661diff
changeset | 44 | def entries: Iterator[(Key, Entry)] = rep.iterator | 
| 
b01b6977a5e8
clarified signature -- avoid oddities of Iterable like Iterator.map;
 wenzelm parents: 
46661diff
changeset | 45 | def keys: Iterator[Key] = entries.map(_._1) | 
| 46611 | 46 | |
| 47 | def dest: List[(Key, List[Key])] = | |
| 46666 
b01b6977a5e8
clarified signature -- avoid oddities of Iterable like Iterator.map;
 wenzelm parents: 
46661diff
changeset | 48 | (for ((x, (_, (_, succs))) <- entries) yield (x, succs.toList)).toList | 
| 46611 | 49 | |
| 46666 
b01b6977a5e8
clarified signature -- avoid oddities of Iterable like Iterator.map;
 wenzelm parents: 
46661diff
changeset | 50 | override def toString: String = | 
| 
b01b6977a5e8
clarified signature -- avoid oddities of Iterable like Iterator.map;
 wenzelm parents: 
46661diff
changeset | 51 |     dest.map(p => p._1.toString + " -> " + p._2.map(_.toString).mkString("{", ", ", "}"))
 | 
| 
b01b6977a5e8
clarified signature -- avoid oddities of Iterable like Iterator.map;
 wenzelm parents: 
46661diff
changeset | 52 |       .mkString("Graph(", ", ", ")")
 | 
| 46611 | 53 | |
| 54 | private def get_entry(x: Key): Entry = | |
| 55 |     rep.get(x) match {
 | |
| 56 | case Some(entry) => entry | |
| 57 | case None => throw new Graph.Undefined(x) | |
| 58 | } | |
| 59 | ||
| 60 | private def map_entry(x: Key, f: Entry => Entry): Graph[Key, A] = | |
| 61 | new Graph[Key, A](rep + (x -> f(get_entry(x)))) | |
| 62 | ||
| 63 | ||
| 64 | /* nodes */ | |
| 65 | ||
| 66 | def get_node(x: Key): A = get_entry(x)._1 | |
| 67 | ||
| 68 | def map_node(x: Key, f: A => A): Graph[Key, A] = | |
| 69 |     map_entry(x, { case (i, ps) => (f(i), ps) })
 | |
| 70 | ||
| 71 | ||
| 72 | /* reachability */ | |
| 73 | ||
| 74 | /*nodes reachable from xs -- topologically sorted for acyclic graphs*/ | |
| 75 | def reachable(next: Key => Keys, xs: List[Key]): (List[List[Key]], Keys) = | |
| 76 |   {
 | |
| 48350 
09bf3b73e446
clarified topological ordering: preserve order of adjacency via reverse fold;
 wenzelm parents: 
48348diff
changeset | 77 | def reach(x: Key, reached: (List[Key], Keys)): (List[Key], Keys) = | 
| 46611 | 78 |     {
 | 
| 79 | val (rs, r_set) = reached | |
| 80 | if (r_set(x)) reached | |
| 81 |       else {
 | |
| 48350 
09bf3b73e446
clarified topological ordering: preserve order of adjacency via reverse fold;
 wenzelm parents: 
48348diff
changeset | 82 | val (rs1, r_set1) = (next(x) :\ (rs, r_set + x))(reach) | 
| 46611 | 83 | (x :: rs1, r_set1) | 
| 84 | } | |
| 85 | } | |
| 86 | def reachs(reached: (List[List[Key]], Keys), x: Key): (List[List[Key]], Keys) = | |
| 87 |     {
 | |
| 88 | val (rss, r_set) = reached | |
| 48350 
09bf3b73e446
clarified topological ordering: preserve order of adjacency via reverse fold;
 wenzelm parents: 
48348diff
changeset | 89 | val (rs, r_set1) = reach(x, (Nil, r_set)) | 
| 46611 | 90 | (rs :: rss, r_set1) | 
| 91 | } | |
| 46661 
d2ac78ba805e
prefer sorted Map/Set for canonical order of results -- pass ordering via fresh copy of empty;
 wenzelm parents: 
46659diff
changeset | 92 | ((List.empty[List[Key]], empty_keys) /: xs)(reachs) | 
| 46611 | 93 | } | 
| 94 | ||
| 95 | /*immediate*/ | |
| 96 | def imm_preds(x: Key): Keys = get_entry(x)._2._1 | |
| 97 | def imm_succs(x: Key): Keys = get_entry(x)._2._2 | |
| 98 | ||
| 99 | /*transitive*/ | |
| 100 | def all_preds(xs: List[Key]): List[Key] = reachable(imm_preds, xs)._1.flatten | |
| 101 | def all_succs(xs: List[Key]): List[Key] = reachable(imm_succs, xs)._1.flatten | |
| 102 | ||
| 46613 | 103 | /*strongly connected components; see: David King and John Launchbury, | 
| 104 | "Structuring Depth First Search Algorithms in Haskell"*/ | |
| 105 | def strong_conn: List[List[Key]] = | |
| 46666 
b01b6977a5e8
clarified signature -- avoid oddities of Iterable like Iterator.map;
 wenzelm parents: 
46661diff
changeset | 106 | reachable(imm_preds, all_succs(keys.toList))._1.filterNot(_.isEmpty).reverse | 
| 46613 | 107 | |
| 46611 | 108 | |
| 109 | /* minimal and maximal elements */ | |
| 110 | ||
| 111 | def minimals: List[Key] = | |
| 112 |     (List.empty[Key] /: rep) {
 | |
| 113 | case (ms, (m, (_, (preds, _)))) => if (preds.isEmpty) m :: ms else ms } | |
| 114 | ||
| 115 | def maximals: List[Key] = | |
| 116 |     (List.empty[Key] /: rep) {
 | |
| 117 | case (ms, (m, (_, (_, succs)))) => if (succs.isEmpty) m :: ms else ms } | |
| 118 | ||
| 119 | def is_minimal(x: Key): Boolean = imm_preds(x).isEmpty | |
| 120 | def is_maximal(x: Key): Boolean = imm_succs(x).isEmpty | |
| 121 | ||
| 122 | ||
| 46668 | 123 | /* node operations */ | 
| 46611 | 124 | |
| 125 | def new_node(x: Key, info: A): Graph[Key, A] = | |
| 126 |   {
 | |
| 48648 | 127 | 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 | 128 | else new Graph[Key, A](rep + (x -> (info, (empty_keys, empty_keys)))) | 
| 46611 | 129 | } | 
| 130 | ||
| 46613 | 131 | def default_node(x: Key, info: A): Graph[Key, A] = | 
| 48648 | 132 | if (defined(x)) this else new_node(x, info) | 
| 46613 | 133 | |
| 46661 
d2ac78ba805e
prefer sorted Map/Set for canonical order of results -- pass ordering via fresh copy of empty;
 wenzelm parents: 
46659diff
changeset | 134 | 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 | 135 | : SortedMap[Key, Entry] = | 
| 46611 | 136 |     map.get(y) match {
 | 
| 137 | case None => map | |
| 138 | case Some((i, (preds, succs))) => | |
| 139 | map + (y -> (i, if (fst) (preds - x, succs) else (preds, succs - x))) | |
| 140 | } | |
| 141 | ||
| 142 | def del_node(x: Key): Graph[Key, A] = | |
| 143 |   {
 | |
| 144 | val (preds, succs) = get_entry(x)._2 | |
| 145 | new Graph[Key, A]( | |
| 146 | (((rep - x) /: preds)(del_adjacent(false, x)) /: succs)(del_adjacent(true, x))) | |
| 147 | } | |
| 148 | ||
| 46614 
165886a4fe64
clarified Graph.restrict (formerly Graph.subgraph) based on public graph operations;
 wenzelm parents: 
46613diff
changeset | 149 | def restrict(pred: Key => Boolean): Graph[Key, A] = | 
| 46666 
b01b6977a5e8
clarified signature -- avoid oddities of Iterable like Iterator.map;
 wenzelm parents: 
46661diff
changeset | 150 |     (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 | 151 | |
| 46611 | 152 | |
| 46668 | 153 | /* edge operations */ | 
| 46611 | 154 | |
| 155 | def is_edge(x: Key, y: Key): Boolean = | |
| 48507 | 156 | defined(x) && defined(y) && imm_succs(x)(y) | 
| 46611 | 157 | |
| 158 | def add_edge(x: Key, y: Key): Graph[Key, A] = | |
| 159 | if (is_edge(x, y)) this | |
| 160 | else | |
| 161 |       map_entry(y, { case (i, (preds, succs)) => (i, (preds + x, succs)) }).
 | |
| 162 |       map_entry(x, { case (i, (preds, succs)) => (i, (preds, succs + y)) })
 | |
| 163 | ||
| 164 | def del_edge(x: Key, y: Key): Graph[Key, A] = | |
| 165 | if (is_edge(x, y)) | |
| 166 |       map_entry(y, { case (i, (preds, succs)) => (i, (preds - x, succs)) }).
 | |
| 167 |       map_entry(x, { case (i, (preds, succs)) => (i, (preds, succs - y)) })
 | |
| 168 | else this | |
| 169 | ||
| 170 | ||
| 171 | /* irreducible paths -- Hasse diagram */ | |
| 172 | ||
| 46661 
d2ac78ba805e
prefer sorted Map/Set for canonical order of results -- pass ordering via fresh copy of empty;
 wenzelm parents: 
46659diff
changeset | 173 | private def irreducible_preds(x_set: Keys, path: List[Key], z: Key): List[Key] = | 
| 46611 | 174 |   {
 | 
| 175 | def red(x: Key)(x1: Key) = is_edge(x, x1) && x1 != z | |
| 176 | @tailrec def irreds(xs0: List[Key], xs1: List[Key]): List[Key] = | |
| 177 |       xs0 match {
 | |
| 178 | case Nil => xs1 | |
| 179 | case x :: xs => | |
| 180 | if (!(x_set(x)) || x == z || path.contains(x) || | |
| 181 | xs.exists(red(x)) || xs1.exists(red(x))) | |
| 182 | irreds(xs, xs1) | |
| 183 | else irreds(xs, x :: xs1) | |
| 184 | } | |
| 185 | irreds(imm_preds(z).toList, Nil) | |
| 186 | } | |
| 187 | ||
| 188 | def irreducible_paths(x: Key, y: Key): List[List[Key]] = | |
| 189 |   {
 | |
| 190 | val (_, x_set) = reachable(imm_succs, List(x)) | |
| 191 | def paths(path: List[Key])(ps: List[List[Key]], z: Key): List[List[Key]] = | |
| 192 | if (x == z) (z :: path) :: ps | |
| 193 | else (ps /: irreducible_preds(x_set, path, z))(paths(z :: path)) | |
| 194 | if ((x == y) && !is_edge(x, x)) List(Nil) else paths(Nil)(Nil, y) | |
| 195 | } | |
| 196 | ||
| 197 | ||
| 198 | /* maintain acyclic graphs */ | |
| 199 | ||
| 200 | def add_edge_acyclic(x: Key, y: Key): Graph[Key, A] = | |
| 201 | if (is_edge(x, y)) this | |
| 202 |     else {
 | |
| 203 |       irreducible_paths(y, x) match {
 | |
| 204 | case Nil => add_edge(x, y) | |
| 205 | case cycles => throw new Graph.Cycles(cycles.map(x :: _)) | |
| 206 | } | |
| 207 | } | |
| 208 | ||
| 48348 | 209 | def add_deps_acyclic(y: Key, xs: List[Key]): Graph[Key, A] = | 
| 46611 | 210 | (this /: xs)(_.add_edge_acyclic(_, y)) | 
| 211 | ||
| 212 | def topological_order: List[Key] = all_succs(minimals) | |
| 213 | } |