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