| author | blanchet | 
| Thu, 25 Jun 2015 12:41:43 +0200 | |
| changeset 60568 | a9b71c82647b | 
| parent 59447 | e7cbfe240078 | 
| child 67601 | b34be3010273 | 
| 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 | |
| 59384 | 58 | sealed case class Info( | 
| 59 | x: Double, y: Double, width2: Double, height2: Double, lines: List[String]) | |
| 60 |   {
 | |
| 61 | def left: Double = x - width2 | |
| 62 | def right: Double = x + width2 | |
| 63 | def width: Double = 2 * width2 | |
| 64 | def height: Double = 2 * height2 | |
| 65 | } | |
| 50469 | 66 | |
| 59384 | 67 | type Graph = isabelle.Graph[Vertex, Info] | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 68 | |
| 59384 | 69 | def make_graph(entries: List[((Vertex, Info), List[Vertex])]): Graph = | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 70 | isabelle.Graph.make(entries)(Vertex.Ordering) | 
| 59262 | 71 | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 72 | val empty_graph: Graph = make_graph(Nil) | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 73 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 74 | |
| 59307 | 75 | /* vertex x coordinate */ | 
| 76 | ||
| 59384 | 77 | private def vertex_left(graph: Graph, v: Vertex) = graph.get_node(v).left | 
| 78 | private def vertex_right(graph: Graph, v: Vertex) = graph.get_node(v).right | |
| 59307 | 79 | |
| 80 | private def move_vertex(graph: Graph, v: Vertex, dx: Double): Graph = | |
| 59384 | 81 | if (dx == 0.0) graph else graph.map_node(v, info => info.copy(x = info.x + dx)) | 
| 59307 | 82 | |
| 83 | ||
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 84 | /* layout */ | 
| 
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 | val empty: Layout = new Layout(Metrics.default, Graph_Display.empty_graph, empty_graph) | 
| 50470 | 87 | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 88 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 89 | private type Levels = Map[Vertex, Int] | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 90 | private val empty_levels: Levels = Map.empty | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 91 | |
| 59384 | 92 | def make(options: Options, metrics: Metrics, | 
| 93 | node_text: (Graph_Display.Node, XML.Body) => String, | |
| 94 | 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 | 95 |   {
 | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 96 | 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 | 97 |     else {
 | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 98 | /* initial graph */ | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 99 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 100 | val initial_graph = | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 101 | make_graph( | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 102 | input_graph.iterator.map( | 
| 59384 | 103 |             { case (a, (content, (_, bs))) =>
 | 
| 104 | val lines = split_lines(node_text(a, content)) | |
| 105 | val w2 = metrics.node_width2(lines) | |
| 106 | val h2 = metrics.node_height2(lines.length) | |
| 107 | ((Node(a), Info(0.0, 0.0, w2, h2, lines)), bs.toList.map(Node(_))) | |
| 108 | }).toList) | |
| 49557 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 109 | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 110 | val initial_levels: Levels = | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 111 |         (empty_levels /: initial_graph.topological_order) {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 112 | case (levels, vertex) => | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 113 | val level = | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 114 |               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 | 115 | levels + (vertex -> level) | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 116 | } | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 117 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 118 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 119 | /* graph with dummies */ | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 120 | |
| 59384 | 121 | val dummy_info = Info(0.0, 0.0, metrics.dummy_size2, metrics.dummy_size2, Nil) | 
| 122 | ||
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 123 | val (dummy_graph, dummy_levels) = | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 124 |         ((initial_graph, initial_levels) /: input_graph.edges_iterator) {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 125 | case ((graph, levels), (node1, node2)) => | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 126 | val v1 = Node(node1); val l1 = levels(v1) | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 127 | val v2 = Node(node2); val l2 = levels(v2) | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 128 | if (l2 - l1 <= 1) (graph, levels) | 
| 59263 | 129 |               else {
 | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 130 | val dummies_levels = | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 131 |                   (for { (l, i) <- ((l1 + 1) until l2).iterator.zipWithIndex }
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 132 | yield (Dummy(node1, node2, i), l)).toList | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 133 | val dummies = dummies_levels.map(_._1) | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 134 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 135 | val levels1 = (levels /: dummies_levels)(_ + _) | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 136 | val graph1 = | 
| 59384 | 137 | ((graph /: dummies)(_.new_node(_, dummy_info)).del_edge(v1, v2) /: | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 138 |                     (v1 :: dummies ::: List(v2)).sliding(2)) {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 139 | case (g, List(a, b)) => g.add_edge(a, b) } | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 140 | (graph1, levels1) | 
| 59263 | 141 | } | 
| 142 | } | |
| 49557 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 143 | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 144 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 145 | /* minimize edge crossings and initial coordinates */ | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 146 | |
| 59313 | 147 | val levels = minimize_crossings(options, dummy_graph, level_list(dummy_levels)) | 
| 50469 | 148 | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 149 | val levels_graph: Graph = | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 150 |         (((dummy_graph, 0.0) /: levels) {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 151 | case ((graph, y), level) => | 
| 59384 | 152 |             val num_lines = (0 /: level) { case (n, v) => n max graph.get_node(v).lines.length }
 | 
| 59312 
c4b9b04bfc6d
proper level distance according to number of edges, as in old browser;
 wenzelm parents: 
59311diff
changeset | 153 |             val num_edges = (0 /: level) { case (n, v) => n + graph.imm_succs(v).size }
 | 
| 59384 | 154 | |
| 155 | val y1 = y + metrics.node_height2(num_lines) | |
| 156 | ||
| 157 | val graph1 = | |
| 158 |               (((graph, 0.0) /: level) { case ((g, x), v) =>
 | |
| 159 | val info = g.get_node(v) | |
| 160 | val g1 = g.map_node(v, _ => info.copy(x = x + info.width2, y = y1)) | |
| 161 | val x1 = x + info.width + metrics.gap | |
| 162 | (g1, x1) | |
| 163 | })._1 | |
| 164 | ||
| 165 | val y2 = y1 + metrics.level_height2(num_lines, num_edges) | |
| 166 | ||
| 167 | (graph1, y2) | |
| 50474 
6ee044e2d1a7
initial layout coordinates more like old browser;
 wenzelm parents: 
50470diff
changeset | 168 | })._1 | 
| 
6ee044e2d1a7
initial layout coordinates more like old browser;
 wenzelm parents: 
50470diff
changeset | 169 | |
| 59265 
962ad3942ea7
more metrics, with integer coordinates for layout;
 wenzelm parents: 
59264diff
changeset | 170 | |
| 59307 | 171 | /* 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 | 172 | |
| 59307 | 173 | val output_graph = | 
| 59313 | 174 | rubberband(options, metrics, levels, | 
| 175 | 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 | 176 | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 177 | 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 | 178 | } | 
| 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 179 | } | 
| 49737 
dd6fc7c9504a
avoid somewhat heavy rebuilding of SortedMap via map where plain iteration is sufficient;
 wenzelm parents: 
49565diff
changeset | 180 | |
| 49744 | 181 | |
| 182 | ||
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 183 | /** edge crossings **/ | 
| 50469 | 184 | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 185 | private type Level = List[Vertex] | 
| 49737 
dd6fc7c9504a
avoid somewhat heavy rebuilding of SortedMap via map where plain iteration is sufficient;
 wenzelm parents: 
49565diff
changeset | 186 | |
| 59313 | 187 | private def minimize_crossings( | 
| 188 | options: Options, graph: Graph, levels: List[Level]): List[Level] = | |
| 49745 | 189 |   {
 | 
| 59306 | 190 | def resort(parent: Level, child: Level, top_down: Boolean): Level = | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 191 |       child.map(v => {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 192 | val ps = if (top_down) graph.imm_preds(v) else graph.imm_succs(v) | 
| 50469 | 193 | val weight = | 
| 50468 | 194 |             (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 | 195 | (v, weight) | 
| 49557 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 196 | }).sortBy(_._2).map(_._1) | 
| 50469 | 197 | |
| 59419 
2fb2194853cc
clarified iterations: enforce full top_down/bottom_up cycle for better stability of layout;
 wenzelm parents: 
59384diff
changeset | 198 | ((levels, count_crossings(graph, levels)) /: | 
| 
2fb2194853cc
clarified iterations: enforce full top_down/bottom_up cycle for better stability of layout;
 wenzelm parents: 
59384diff
changeset | 199 |       (1 to (2 * options.int("graphview_iterations_minimize_crossings")))) {
 | 
| 
2fb2194853cc
clarified iterations: enforce full top_down/bottom_up cycle for better stability of layout;
 wenzelm parents: 
59384diff
changeset | 200 | case ((old_levels, old_crossings), iteration) => | 
| 
2fb2194853cc
clarified iterations: enforce full top_down/bottom_up cycle for better stability of layout;
 wenzelm parents: 
59384diff
changeset | 201 | val top_down = (iteration % 2 == 1) | 
| 59306 | 202 | val new_levels = | 
| 203 | if (top_down) | |
| 204 | (List(old_levels.head) /: old_levels.tail)((tops, bot) => | |
| 205 | resort(tops.head, bot, top_down) :: tops).reverse | |
| 206 |           else {
 | |
| 207 | val rev_old_levels = old_levels.reverse | |
| 208 | (List(rev_old_levels.head) /: rev_old_levels.tail)((bots, top) => | |
| 209 | resort(bots.head, top, top_down) :: bots) | |
| 210 | } | |
| 211 | val new_crossings = count_crossings(graph, new_levels) | |
| 212 | if (new_crossings < old_crossings) | |
| 59419 
2fb2194853cc
clarified iterations: enforce full top_down/bottom_up cycle for better stability of layout;
 wenzelm parents: 
59384diff
changeset | 213 | (new_levels, new_crossings) | 
| 59306 | 214 | else | 
| 59419 
2fb2194853cc
clarified iterations: enforce full top_down/bottom_up cycle for better stability of layout;
 wenzelm parents: 
59384diff
changeset | 215 | (old_levels, old_crossings) | 
| 49557 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 216 | }._1 | 
| 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 217 | } | 
| 50469 | 218 | |
| 59307 | 219 | 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 | 220 |   {
 | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 221 |     val max_lev = (-1 /: levels) { case (m, (_, l)) => m max l }
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 222 | val buckets = new Array[Level](max_lev + 1) | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 223 |     for (l <- 0 to max_lev) { buckets(l) = Nil }
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 224 |     for ((v, l) <- levels) { buckets(l) = v :: buckets(l) }
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 225 | buckets.iterator.map(_.sorted(Vertex.Ordering)).toList | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 226 | } | 
| 50469 | 227 | |
| 59307 | 228 | private def count_crossings(graph: Graph, levels: List[Level]): Int = | 
| 59310 | 229 |     levels.iterator.sliding(2).map {
 | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 230 | case List(top, bot) => | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 231 |         top.iterator.zipWithIndex.map {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 232 | case (outer_parent, outer_parent_index) => | 
| 59310 | 233 | graph.imm_succs(outer_parent).iterator.map(bot.indexOf(_)).map(outer_child => | 
| 234 | (0 until outer_parent_index).iterator.map(inner_parent => | |
| 235 | graph.imm_succs(top(inner_parent)).iterator.map(bot.indexOf(_)). | |
| 236 | filter(inner_child => outer_child < inner_child).size).sum).sum | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 237 | }.sum | 
| 59447 
e7cbfe240078
complete pattern coverage, e.g. relevant for singleton graph;
 wenzelm parents: 
59419diff
changeset | 238 | case _ => 0 | 
| 59310 | 239 | }.sum | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 240 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 241 | |
| 
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 | /** pendulum method **/ | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 244 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 245 | /*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 | 246 | calculating coordinates with the "pendulum method". A "region" is a | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 247 | group of vertices which "stick together".*/ | 
| 59315 | 248 | private class Region(val content: List[Vertex]) | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 249 |   {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 250 | def distance(metrics: Metrics, graph: Graph, that: Region): Double = | 
| 59384 | 251 | vertex_left(graph, that.content.head) - | 
| 252 | vertex_right(graph, this.content.last) - | |
| 253 | metrics.gap | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 254 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 255 | def deflection(graph: Graph, top_down: Boolean): Double = | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 256 |       ((for (a <- content.iterator) yield {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 257 | val x = graph.get_node(a).x | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 258 | 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 | 259 | 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 | 260 | }).sum / content.length).round.toDouble | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 261 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 262 | def move(graph: Graph, dx: Double): Graph = | 
| 59307 | 263 | if (dx == 0.0) graph else (graph /: content)(move_vertex(_, _, dx)) | 
| 59315 | 264 | |
| 265 | def combine(that: Region): Region = new Region(content ::: that.content) | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 266 | } | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 267 | |
| 59313 | 268 | private def pendulum( | 
| 269 | 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 | 270 |   {
 | 
| 59315 | 271 | def combine_regions(graph: Graph, top_down: Boolean, level: List[Region]): List[Region] = | 
| 272 |       level match {
 | |
| 273 | case r1 :: rest => | |
| 274 | val rest1 = combine_regions(graph, top_down, rest) | |
| 275 |           rest1 match {
 | |
| 276 | case r2 :: rest2 => | |
| 277 | val d1 = r1.deflection(graph, top_down) | |
| 278 | val d2 = r2.deflection(graph, top_down) | |
| 279 | if (// Do regions touch? | |
| 280 | r1.distance(metrics, graph, r2) <= 0.0 && | |
| 281 | // Do they influence each other? | |
| 282 | (d1 <= 0.0 && d2 < d1 || d2 > 0.0 && d1 > d2 || d1 > 0.0 && d2 < 0.0)) | |
| 283 | r1.combine(r2) :: rest2 | |
| 284 | else r1 :: rest1 | |
| 285 | case _ => r1 :: rest1 | |
| 286 | } | |
| 287 | case _ => level | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 288 | } | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 289 | |
| 59315 | 290 | def deflect(level: List[Region], top_down: Boolean, graph: Graph): (Graph, Boolean) = | 
| 50469 | 291 |     {
 | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 292 |       ((graph, false) /: (0 until level.length)) {
 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 293 | case ((graph, moved), i) => | 
| 59263 | 294 | val r = level(i) | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 295 | val d = r.deflection(graph, top_down) | 
| 59316 | 296 | val dx = | 
| 297 | if (d < 0.0 && i > 0) | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 298 | - (level(i - 1).distance(metrics, graph, r) min (- d)) | 
| 59316 | 299 | else if (d >= 0.0 && i < level.length - 1) | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 300 | r.distance(metrics, graph, level(i + 1)) min d | 
| 59265 
962ad3942ea7
more metrics, with integer coordinates for layout;
 wenzelm parents: 
59264diff
changeset | 301 | else d | 
| 59316 | 302 | (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 | 303 | } | 
| 50469 | 304 | } | 
| 305 | ||
| 59315 | 306 | val initial_regions = levels.map(level => level.map(l => new Region(List(l)))) | 
| 50469 | 307 | |
| 59419 
2fb2194853cc
clarified iterations: enforce full top_down/bottom_up cycle for better stability of layout;
 wenzelm parents: 
59384diff
changeset | 308 | ((levels_graph, initial_regions, true) /: | 
| 
2fb2194853cc
clarified iterations: enforce full top_down/bottom_up cycle for better stability of layout;
 wenzelm parents: 
59384diff
changeset | 309 |       (1 to (2 * options.int("graphview_iterations_pendulum")))) {
 | 
| 
2fb2194853cc
clarified iterations: enforce full top_down/bottom_up cycle for better stability of layout;
 wenzelm parents: 
59384diff
changeset | 310 | case ((graph, regions, moved), iteration) => | 
| 
2fb2194853cc
clarified iterations: enforce full top_down/bottom_up cycle for better stability of layout;
 wenzelm parents: 
59384diff
changeset | 311 | val top_down = (iteration % 2 == 1) | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 312 |         if (moved) {
 | 
| 59315 | 313 | val (graph1, regions1, moved1) = | 
| 314 | ((graph, List.empty[List[Region]], false) /: | |
| 315 |               (if (top_down) regions else regions.reverse)) { case ((graph, tops, moved), bot) =>
 | |
| 316 | val bot1 = combine_regions(graph, top_down, bot) | |
| 317 | val (graph1, moved1) = deflect(bot1, top_down, graph) | |
| 318 | (graph1, bot1 :: tops, moved || moved1) | |
| 319 | } | |
| 59419 
2fb2194853cc
clarified iterations: enforce full top_down/bottom_up cycle for better stability of layout;
 wenzelm parents: 
59384diff
changeset | 320 | (graph1, regions1.reverse, moved1) | 
| 59265 
962ad3942ea7
more metrics, with integer coordinates for layout;
 wenzelm parents: 
59264diff
changeset | 321 | } | 
| 59419 
2fb2194853cc
clarified iterations: enforce full top_down/bottom_up cycle for better stability of layout;
 wenzelm parents: 
59384diff
changeset | 322 | else (graph, regions, moved) | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 323 | }._1 | 
| 49557 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 324 | } | 
| 59307 | 325 | |
| 326 | ||
| 327 | ||
| 328 | /** rubberband method **/ | |
| 329 | ||
| 330 | private def force_weight(graph: Graph, v: Vertex): Double = | |
| 331 |   {
 | |
| 332 | val preds = graph.imm_preds(v) | |
| 333 | val succs = graph.imm_succs(v) | |
| 334 | val n = preds.size + succs.size | |
| 335 | if (n == 0) 0.0 | |
| 336 |     else {
 | |
| 337 | val x = graph.get_node(v).x | |
| 338 | ((preds.iterator ++ succs.iterator).map(w => graph.get_node(w).x - x)).sum / n | |
| 339 | } | |
| 340 | } | |
| 341 | ||
| 59313 | 342 | private def rubberband( | 
| 343 | options: Options, metrics: Metrics, levels: List[Level], graph: Graph): Graph = | |
| 59307 | 344 |   {
 | 
| 59384 | 345 | val gap = metrics.gap | 
| 59307 | 346 | |
| 59419 
2fb2194853cc
clarified iterations: enforce full top_down/bottom_up cycle for better stability of layout;
 wenzelm parents: 
59384diff
changeset | 347 |     (graph /: (1 to (2 * options.int("graphview_iterations_rubberband")))) { case (graph, _) =>
 | 
| 59307 | 348 |       (graph /: levels) { case (graph, level) =>
 | 
| 349 | val m = level.length - 1 | |
| 350 |         (graph /: level.iterator.zipWithIndex) {
 | |
| 351 | case (g, (v, i)) => | |
| 352 | val d = force_weight(g, v) | |
| 59384 | 353 | if (d < 0.0 && (i == 0 || vertex_right(g, level(i - 1)) + gap < vertex_left(g, v) + d) || | 
| 354 | d > 0.0 && (i == m || vertex_left(g, level(i + 1)) - gap > vertex_right(g, v) + d)) | |
| 59307 | 355 | move_vertex(g, v, d.round.toDouble) | 
| 356 | else g | |
| 357 | } | |
| 358 | } | |
| 359 | } | |
| 360 | } | |
| 49557 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 Markus Kaiser <markus.kaiser@in.tum.de> parents: diff
changeset | 361 | } | 
| 59262 | 362 | |
| 363 | final class Layout private( | |
| 59290 | 364 | val metrics: Metrics, | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 365 | val input_graph: Graph_Display.Graph, | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 366 | val output_graph: Layout.Graph) | 
| 59262 | 367 | {
 | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 368 | /* vertex coordinates */ | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 369 | |
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 370 | def translate_vertex(v: Layout.Vertex, dx: Double, dy: Double): Layout = | 
| 
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 | 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 | 373 |     else {
 | 
| 59384 | 374 | val output_graph1 = | 
| 375 | output_graph.map_node(v, info => info.copy(x = info.x + dx, y = info.y + dy)) | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 376 | new Layout(metrics, input_graph, output_graph1) | 
| 59262 | 377 | } | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 378 | } | 
| 59262 | 379 | |
| 380 | ||
| 59384 | 381 | /* regular nodes */ | 
| 382 | ||
| 383 | def get_node(node: Graph_Display.Node): Layout.Info = | |
| 384 | output_graph.get_node(Layout.Node(node)) | |
| 385 | ||
| 386 | def nodes_iterator: Iterator[Layout.Info] = | |
| 387 | for ((_: Layout.Node, (info, _)) <- output_graph.iterator) yield info | |
| 388 | ||
| 389 | ||
| 59290 | 390 | /* dummies */ | 
| 391 | ||
| 59384 | 392 | def dummies_iterator: Iterator[Layout.Info] = | 
| 393 | for ((_: Layout.Dummy, (info, _)) <- output_graph.iterator) yield info | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 394 | |
| 59384 | 395 | def dummies_iterator(edge: Graph_Display.Edge): Iterator[Layout.Info] = | 
| 396 |     new Iterator[Layout.Info] {
 | |
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 397 | private var index = 0 | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 398 | def hasNext: Boolean = output_graph.defined(Layout.Dummy(edge._1, edge._2, index)) | 
| 59384 | 399 | def next: Layout.Info = | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 400 |       {
 | 
| 59384 | 401 | val info = output_graph.get_node(Layout.Dummy(edge._1, edge._2, index)) | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 402 | index += 1 | 
| 59384 | 403 | info | 
| 59302 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 404 | } | 
| 
4d985afc0565
explict layout graph structure, with dummies and coordinates;
 wenzelm parents: 
59292diff
changeset | 405 | } | 
| 59262 | 406 | } | 
| 407 |