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