| author | wenzelm | 
| Sat, 10 Jan 2015 16:35:21 +0100 | |
| changeset 59342 | fd9102b419f5 | 
| parent 59316 | a1238edd8b36 | 
| child 59384 | c75327a34960 | 
| permissions | -rw-r--r-- | 
| 59232 | 1 | /* Title: Tools/Graphview/layout.scala | 
| 49557 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 2 | Author: Markus Kaiser, TU Muenchen | 
| 59240 | 3 | Author: Makarius | 
| 49557 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 4 | |
| 59307 | 5 | DAG layout according to: | 
| 59261 | 6 | |
| 7 | Georg Sander, "Graph Layout through the VCG Tool", in: Graph Drawing, | |
| 8 | DIMACS International Workshop (GD'94), Springer LNCS 894, 1995. | |
| 9 | ||
| 10 | http://dx.doi.org/10.1007/3-540-58950-3_371 | |
| 11 | ftp://ftp.cs.uni-sb.de/pub/graphics/vcg/doc/tr-A03-94.ps.gz | |
| 49557 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 12 | */ | 
| 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 13 | |
| 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 14 | package isabelle.graphview | 
| 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 15 | |
| 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 16 | |
| 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 17 | import isabelle._ | 
| 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 18 | |
| 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 19 | |
| 59232 | 20 | object Layout | 
| 49744 | 21 | {
 | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 22 | /* graph structure */ | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 23 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 24 | object Vertex | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 25 |   {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 26 | object Ordering extends scala.math.Ordering[Vertex] | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 27 |     {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 28 | def compare(v1: Vertex, v2: Vertex): Int = | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 29 |         (v1, v2) match {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 30 | case (Node(a), Node(b)) => Graph_Display.Node.Ordering.compare(a, b) | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 31 | case (Dummy(a1, a2, i), Dummy(b1, b2, j)) => | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 32 |             Graph_Display.Node.Ordering.compare(a1, b1) match {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 33 | case 0 => | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 34 |                 Graph_Display.Node.Ordering.compare(a2, b2) match {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 35 | case 0 => i compare j | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 36 | case ord => ord | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 37 | } | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 38 | case ord => ord | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 39 | } | 
| 59311 
a269cc01e8eb
clarified Vertex.Ordering, to approximate situation before 4d985afc0565, which is relevant for level arrangement;
 wenzelm parents: 
59310diff
changeset | 40 | case (Node(a), Dummy(b, _, _)) => | 
| 
a269cc01e8eb
clarified Vertex.Ordering, to approximate situation before 4d985afc0565, which is relevant for level arrangement;
 wenzelm parents: 
59310diff
changeset | 41 |             Graph_Display.Node.Ordering.compare(a, b) match {
 | 
| 
a269cc01e8eb
clarified Vertex.Ordering, to approximate situation before 4d985afc0565, which is relevant for level arrangement;
 wenzelm parents: 
59310diff
changeset | 42 | case 0 => -1 | 
| 
a269cc01e8eb
clarified Vertex.Ordering, to approximate situation before 4d985afc0565, which is relevant for level arrangement;
 wenzelm parents: 
59310diff
changeset | 43 | case ord => ord | 
| 
a269cc01e8eb
clarified Vertex.Ordering, to approximate situation before 4d985afc0565, which is relevant for level arrangement;
 wenzelm parents: 
59310diff
changeset | 44 | } | 
| 
a269cc01e8eb
clarified Vertex.Ordering, to approximate situation before 4d985afc0565, which is relevant for level arrangement;
 wenzelm parents: 
59310diff
changeset | 45 | case (Dummy(a, _, _), Node(b)) => | 
| 
a269cc01e8eb
clarified Vertex.Ordering, to approximate situation before 4d985afc0565, which is relevant for level arrangement;
 wenzelm parents: 
59310diff
changeset | 46 |             Graph_Display.Node.Ordering.compare(a, b) match {
 | 
| 
a269cc01e8eb
clarified Vertex.Ordering, to approximate situation before 4d985afc0565, which is relevant for level arrangement;
 wenzelm parents: 
59310diff
changeset | 47 | case 0 => 1 | 
| 
a269cc01e8eb
clarified Vertex.Ordering, to approximate situation before 4d985afc0565, which is relevant for level arrangement;
 wenzelm parents: 
59310diff
changeset | 48 | case ord => ord | 
| 
a269cc01e8eb
clarified Vertex.Ordering, to approximate situation before 4d985afc0565, which is relevant for level arrangement;
 wenzelm parents: 
59310diff
changeset | 49 | } | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 50 | } | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 51 | } | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 52 | } | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 53 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 54 | sealed abstract class Vertex | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 55 | case class Node(node: Graph_Display.Node) extends Vertex | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 56 | case class Dummy(node1: Graph_Display.Node, node2: Graph_Display.Node, index: Int) extends Vertex | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 57 | |
| 59262 | 58 |   object Point { val zero: Point = Point(0.0, 0.0) }
 | 
| 59 | case class Point(x: Double, y: Double) | |
| 50469 | 60 | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 61 | type Graph = isabelle.Graph[Vertex, Point] | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 62 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 63 | def make_graph(entries: List[((Vertex, Point), List[Vertex])]): Graph = | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 64 | isabelle.Graph.make(entries)(Vertex.Ordering) | 
| 59262 | 65 | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 66 | val empty_graph: Graph = make_graph(Nil) | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 67 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 68 | |
| 59307 | 69 | /* vertex x coordinate */ | 
| 70 | ||
| 71 | private def vertex_left(metrics: Metrics, graph: Graph, v: Vertex): Double = | |
| 72 | graph.get_node(v).x - metrics.box_width2(v) | |
| 73 | ||
| 74 | private def vertex_right(metrics: Metrics, graph: Graph, v: Vertex): Double = | |
| 75 | graph.get_node(v).x + metrics.box_width2(v) | |
| 76 | ||
| 77 | private def move_vertex(graph: Graph, v: Vertex, dx: Double): Graph = | |
| 78 | if (dx == 0.0) graph else graph.map_node(v, p => Point(p.x + dx, p.y)) | |
| 79 | ||
| 80 | ||
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 81 | /* layout */ | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 82 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 83 | val empty: Layout = new Layout(Metrics.default, Graph_Display.empty_graph, empty_graph) | 
| 50470 | 84 | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 85 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 86 | private type Levels = Map[Vertex, Int] | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 87 | private val empty_levels: Levels = Map.empty | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 88 | |
| 59313 | 89 | def make(options: Options, metrics: Metrics, input_graph: Graph_Display.Graph): Layout = | 
| 49737 
dd6fc7c9504a
avoid somewhat heavy rebuilding of SortedMap via map where plain iteration is sufficient;
 wenzelm parents: 
49565diff
changeset | 90 |   {
 | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 91 | if (input_graph.is_empty) empty | 
| 49557 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 92 |     else {
 | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 93 | /* initial graph */ | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 94 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 95 | val initial_graph = | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 96 | make_graph( | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 97 | input_graph.iterator.map( | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 98 |             { case (a, (_, (_, bs))) => ((Node(a), Point.zero), bs.toList.map(Node(_))) }).toList)
 | 
| 49557 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 99 | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 100 | val initial_levels: Levels = | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 101 |         (empty_levels /: initial_graph.topological_order) {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 102 | case (levels, vertex) => | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 103 | val level = | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 104 |               1 + (-1 /: initial_graph.imm_preds(vertex)) { case (m, v) => m max levels(v) }
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 105 | levels + (vertex -> level) | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 106 | } | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 107 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 108 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 109 | /* graph with dummies */ | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 110 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 111 | val (dummy_graph, dummy_levels) = | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 112 |         ((initial_graph, initial_levels) /: input_graph.edges_iterator) {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 113 | case ((graph, levels), (node1, node2)) => | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 114 | val v1 = Node(node1); val l1 = levels(v1) | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 115 | val v2 = Node(node2); val l2 = levels(v2) | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 116 | if (l2 - l1 <= 1) (graph, levels) | 
| 59263 | 117 |               else {
 | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 118 | val dummies_levels = | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 119 |                   (for { (l, i) <- ((l1 + 1) until l2).iterator.zipWithIndex }
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 120 | yield (Dummy(node1, node2, i), l)).toList | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 121 | val dummies = dummies_levels.map(_._1) | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 122 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 123 | val levels1 = (levels /: dummies_levels)(_ + _) | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 124 | val graph1 = | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 125 | ((graph /: dummies)(_.new_node(_, Point.zero)).del_edge(v1, v2) /: | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 126 |                     (v1 :: dummies ::: List(v2)).sliding(2)) {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 127 | case (g, List(a, b)) => g.add_edge(a, b) } | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 128 | (graph1, levels1) | 
| 59263 | 129 | } | 
| 130 | } | |
| 49557 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 131 | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 132 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 133 | /* minimize edge crossings and initial coordinates */ | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 134 | |
| 59313 | 135 | val levels = minimize_crossings(options, dummy_graph, level_list(dummy_levels)) | 
| 50469 | 136 | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 137 | val levels_graph: Graph = | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 138 |         (((dummy_graph, 0.0) /: levels) {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 139 | case ((graph, y), level) => | 
| 59312 
c4b9b04bfc6d
proper level distance according to number of edges, as in old browser;
 wenzelm parents: 
59311diff
changeset | 140 |             val num_edges = (0 /: level) { case (n, v) => n + graph.imm_succs(v).size }
 | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 141 |             ((((graph, 0.0) /: level) {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 142 | case ((g, x), v) => | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 143 | val w2 = metrics.box_width2(v) | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 144 | (g.map_node(v, _ => Point(x + w2, y)), x + 2 * w2 + metrics.box_gap) | 
| 59312 
c4b9b04bfc6d
proper level distance according to number of edges, as in old browser;
 wenzelm parents: 
59311diff
changeset | 145 | })._1, y + metrics.box_height(num_edges)) | 
| 50474 
6ee044e2d1a7
initial layout coordinates more like old browser;
 wenzelm parents: 
50470diff
changeset | 146 | })._1 | 
| 
6ee044e2d1a7
initial layout coordinates more like old browser;
 wenzelm parents: 
50470diff
changeset | 147 | |
| 59265 
962ad3942ea7
more metrics, with integer coordinates for layout;
 wenzelm parents: 
59264diff
changeset | 148 | |
| 59307 | 149 | /* pendulum/rubberband layout */ | 
| 49557 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 150 | |
| 59307 | 151 | val output_graph = | 
| 59313 | 152 | rubberband(options, metrics, levels, | 
| 153 | pendulum(options, metrics, levels, levels_graph)) | |
| 49557 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 154 | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 155 | new Layout(metrics, input_graph, output_graph) | 
| 49557 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 156 | } | 
| 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 157 | } | 
| 49737 
dd6fc7c9504a
avoid somewhat heavy rebuilding of SortedMap via map where plain iteration is sufficient;
 wenzelm parents: 
49565diff
changeset | 158 | |
| 49744 | 159 | |
| 160 | ||
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 161 | /** edge crossings **/ | 
| 50469 | 162 | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 163 | private type Level = List[Vertex] | 
| 49737 
dd6fc7c9504a
avoid somewhat heavy rebuilding of SortedMap via map where plain iteration is sufficient;
 wenzelm parents: 
49565diff
changeset | 164 | |
| 59313 | 165 | private def minimize_crossings( | 
| 166 | options: Options, graph: Graph, levels: List[Level]): List[Level] = | |
| 49745 | 167 |   {
 | 
| 59306 | 168 | def resort(parent: Level, child: Level, top_down: Boolean): Level = | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 169 |       child.map(v => {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 170 | val ps = if (top_down) graph.imm_preds(v) else graph.imm_succs(v) | 
| 50469 | 171 | val weight = | 
| 50468 | 172 |             (0.0 /: ps) { (w, p) => w + (0 max parent.indexOf(p)) } / (ps.size max 1)
 | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 173 | (v, weight) | 
| 49557 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 174 | }).sortBy(_._2).map(_._1) | 
| 50469 | 175 | |
| 59313 | 176 | ((levels, count_crossings(graph, levels), true) /: | 
| 177 |       (1 to options.int("graphview_iterations_minimize_crossings"))) {
 | |
| 59306 | 178 | case ((old_levels, old_crossings, top_down), _) => | 
| 179 | val new_levels = | |
| 180 | if (top_down) | |
| 181 | (List(old_levels.head) /: old_levels.tail)((tops, bot) => | |
| 182 | resort(tops.head, bot, top_down) :: tops).reverse | |
| 183 |           else {
 | |
| 184 | val rev_old_levels = old_levels.reverse | |
| 185 | (List(rev_old_levels.head) /: rev_old_levels.tail)((bots, top) => | |
| 186 | resort(bots.head, top, top_down) :: bots) | |
| 187 | } | |
| 188 | val new_crossings = count_crossings(graph, new_levels) | |
| 189 | if (new_crossings < old_crossings) | |
| 190 | (new_levels, new_crossings, !top_down) | |
| 191 | else | |
| 192 | (old_levels, old_crossings, !top_down) | |
| 49557 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 193 | }._1 | 
| 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 194 | } | 
| 50469 | 195 | |
| 59307 | 196 | private def level_list(levels: Levels): List[Level] = | 
| 49557 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 197 |   {
 | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 198 |     val max_lev = (-1 /: levels) { case (m, (_, l)) => m max l }
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 199 | val buckets = new Array[Level](max_lev + 1) | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 200 |     for (l <- 0 to max_lev) { buckets(l) = Nil }
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 201 |     for ((v, l) <- levels) { buckets(l) = v :: buckets(l) }
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 202 | buckets.iterator.map(_.sorted(Vertex.Ordering)).toList | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 203 | } | 
| 50469 | 204 | |
| 59307 | 205 | private def count_crossings(graph: Graph, levels: List[Level]): Int = | 
| 59310 | 206 |     levels.iterator.sliding(2).map {
 | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 207 | case List(top, bot) => | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 208 |         top.iterator.zipWithIndex.map {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 209 | case (outer_parent, outer_parent_index) => | 
| 59310 | 210 | graph.imm_succs(outer_parent).iterator.map(bot.indexOf(_)).map(outer_child => | 
| 211 | (0 until outer_parent_index).iterator.map(inner_parent => | |
| 212 | graph.imm_succs(top(inner_parent)).iterator.map(bot.indexOf(_)). | |
| 213 | filter(inner_child => outer_child < inner_child).size).sum).sum | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 214 | }.sum | 
| 59310 | 215 | }.sum | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 216 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 217 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 218 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 219 | /** pendulum method **/ | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 220 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 221 | /*This is an auxiliary class which is used by the layout algorithm when | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 222 | calculating coordinates with the "pendulum method". A "region" is a | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 223 | group of vertices which "stick together".*/ | 
| 59315 | 224 | private class Region(val content: List[Vertex]) | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 225 |   {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 226 | def distance(metrics: Metrics, graph: Graph, that: Region): Double = | 
| 59307 | 227 | vertex_left(metrics, graph, that.content.head) - | 
| 228 | vertex_right(metrics, graph, this.content.last) - | |
| 229 | metrics.box_gap | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 230 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 231 | def deflection(graph: Graph, top_down: Boolean): Double = | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 232 |       ((for (a <- content.iterator) yield {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 233 | val x = graph.get_node(a).x | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 234 | val bs = if (top_down) graph.imm_preds(a) else graph.imm_succs(a) | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 235 | bs.iterator.map(graph.get_node(_).x - x).sum / (bs.size max 1) | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 236 | }).sum / content.length).round.toDouble | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 237 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 238 | def move(graph: Graph, dx: Double): Graph = | 
| 59307 | 239 | if (dx == 0.0) graph else (graph /: content)(move_vertex(_, _, dx)) | 
| 59315 | 240 | |
| 241 | def combine(that: Region): Region = new Region(content ::: that.content) | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 242 | } | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 243 | |
| 59313 | 244 | private def pendulum( | 
| 245 | options: Options, metrics: Metrics, levels: List[Level], levels_graph: Graph): Graph = | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 246 |   {
 | 
| 59315 | 247 | def combine_regions(graph: Graph, top_down: Boolean, level: List[Region]): List[Region] = | 
| 248 |       level match {
 | |
| 249 | case r1 :: rest => | |
| 250 | val rest1 = combine_regions(graph, top_down, rest) | |
| 251 |           rest1 match {
 | |
| 252 | case r2 :: rest2 => | |
| 253 | val d1 = r1.deflection(graph, top_down) | |
| 254 | val d2 = r2.deflection(graph, top_down) | |
| 255 | if (// Do regions touch? | |
| 256 | r1.distance(metrics, graph, r2) <= 0.0 && | |
| 257 | // Do they influence each other? | |
| 258 | (d1 <= 0.0 && d2 < d1 || d2 > 0.0 && d1 > d2 || d1 > 0.0 && d2 < 0.0)) | |
| 259 | r1.combine(r2) :: rest2 | |
| 260 | else r1 :: rest1 | |
| 261 | case _ => r1 :: rest1 | |
| 262 | } | |
| 263 | case _ => level | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 264 | } | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 265 | |
| 59315 | 266 | def deflect(level: List[Region], top_down: Boolean, graph: Graph): (Graph, Boolean) = | 
| 50469 | 267 |     {
 | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 268 |       ((graph, false) /: (0 until level.length)) {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 269 | case ((graph, moved), i) => | 
| 59263 | 270 | val r = level(i) | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 271 | val d = r.deflection(graph, top_down) | 
| 59316 | 272 | val dx = | 
| 273 | if (d < 0.0 && i > 0) | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 274 | - (level(i - 1).distance(metrics, graph, r) min (- d)) | 
| 59316 | 275 | else if (d >= 0.0 && i < level.length - 1) | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 276 | r.distance(metrics, graph, level(i + 1)) min d | 
| 59265 
962ad3942ea7
more metrics, with integer coordinates for layout;
 wenzelm parents: 
59264diff
changeset | 277 | else d | 
| 59316 | 278 | (r.move(graph, dx), moved || d != 0.0) | 
| 49557 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 279 | } | 
| 50469 | 280 | } | 
| 281 | ||
| 59315 | 282 | val initial_regions = levels.map(level => level.map(l => new Region(List(l)))) | 
| 50469 | 283 | |
| 59313 | 284 | ((levels_graph, initial_regions, true, true) /: | 
| 285 |       (1 to options.int("graphview_iterations_pendulum"))) {
 | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 286 | case ((graph, regions, top_down, moved), _) => | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 287 |         if (moved) {
 | 
| 59315 | 288 | val (graph1, regions1, moved1) = | 
| 289 | ((graph, List.empty[List[Region]], false) /: | |
| 290 |               (if (top_down) regions else regions.reverse)) { case ((graph, tops, moved), bot) =>
 | |
| 291 | val bot1 = combine_regions(graph, top_down, bot) | |
| 292 | val (graph1, moved1) = deflect(bot1, top_down, graph) | |
| 293 | (graph1, bot1 :: tops, moved || moved1) | |
| 294 | } | |
| 295 | (graph1, regions1.reverse, !top_down, moved1) | |
| 59265 
962ad3942ea7
more metrics, with integer coordinates for layout;
 wenzelm parents: 
59264diff
changeset | 296 | } | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 297 | else (graph, regions, !top_down, moved) | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 298 | }._1 | 
| 49557 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 299 | } | 
| 59307 | 300 | |
| 301 | ||
| 302 | ||
| 303 | /** rubberband method **/ | |
| 304 | ||
| 305 | private def force_weight(graph: Graph, v: Vertex): Double = | |
| 306 |   {
 | |
| 307 | val preds = graph.imm_preds(v) | |
| 308 | val succs = graph.imm_succs(v) | |
| 309 | val n = preds.size + succs.size | |
| 310 | if (n == 0) 0.0 | |
| 311 |     else {
 | |
| 312 | val x = graph.get_node(v).x | |
| 313 | ((preds.iterator ++ succs.iterator).map(w => graph.get_node(w).x - x)).sum / n | |
| 314 | } | |
| 315 | } | |
| 316 | ||
| 59313 | 317 | private def rubberband( | 
| 318 | options: Options, metrics: Metrics, levels: List[Level], graph: Graph): Graph = | |
| 59307 | 319 |   {
 | 
| 320 | val gap = metrics.box_gap | |
| 321 | def left(g: Graph, v: Vertex) = vertex_left(metrics, g, v) | |
| 322 | def right(g: Graph, v: Vertex) = vertex_right(metrics, g, v) | |
| 323 | ||
| 59313 | 324 |     (graph /: (1 to options.int("graphview_iterations_rubberband"))) { case (graph, _) =>
 | 
| 59307 | 325 |       (graph /: levels) { case (graph, level) =>
 | 
| 326 | val m = level.length - 1 | |
| 327 |         (graph /: level.iterator.zipWithIndex) {
 | |
| 328 | case (g, (v, i)) => | |
| 329 | val d = force_weight(g, v) | |
| 330 | if (d < 0.0 && (i == 0 || right(g, level(i - 1)) + gap < left(g, v) + d) || | |
| 331 | d > 0.0 && (i == m || left(g, level(i + 1)) - gap > right(g, v) + d)) | |
| 332 | move_vertex(g, v, d.round.toDouble) | |
| 333 | else g | |
| 334 | } | |
| 335 | } | |
| 336 | } | |
| 337 | } | |
| 49557 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 338 | } | 
| 59262 | 339 | |
| 340 | final class Layout private( | |
| 59290 | 341 | val metrics: Metrics, | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 342 | val input_graph: Graph_Display.Graph, | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 343 | val output_graph: Layout.Graph) | 
| 59262 | 344 | {
 | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 345 | /* vertex coordinates */ | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 346 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 347 | def get_vertex(v: Layout.Vertex): Layout.Point = | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 348 | if (output_graph.defined(v)) output_graph.get_node(v) | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 349 | else Layout.Point.zero | 
| 59290 | 350 | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 351 | def translate_vertex(v: Layout.Vertex, dx: Double, dy: Double): Layout = | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 352 |   {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 353 | if ((dx == 0.0 && dy == 0.0) || !output_graph.defined(v)) this | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 354 |     else {
 | 
| 59304 | 355 | val output_graph1 = output_graph.map_node(v, p => Layout.Point(p.x + dx, p.y + dy)) | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 356 | new Layout(metrics, input_graph, output_graph1) | 
| 59262 | 357 | } | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 358 | } | 
| 59262 | 359 | |
| 360 | ||
| 59290 | 361 | /* dummies */ | 
| 362 | ||
| 59292 | 363 | def dummies_iterator: Iterator[Layout.Point] = | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 364 | for ((_: Layout.Dummy, (p, _)) <- output_graph.iterator) yield p | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 365 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 366 | def dummies_iterator(edge: Graph_Display.Edge): Iterator[Layout.Point] = | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 367 |     new Iterator[Layout.Point] {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 368 | private var index = 0 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 369 | def hasNext: Boolean = output_graph.defined(Layout.Dummy(edge._1, edge._2, index)) | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 370 | def next: Layout.Point = | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 371 |       {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 372 | val p = output_graph.get_node(Layout.Dummy(edge._1, edge._2, index)) | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 373 | index += 1 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 374 | p | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 375 | } | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 376 | } | 
| 59262 | 377 | } | 
| 378 |