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