| author | paulson | 
| Sun, 20 Sep 2020 15:45:25 +0100 | |
| changeset 72271 | 7e90e1d178b5 | 
| parent 71601 | 97ccf48c2f0c | 
| child 73344 | f5c147654661 | 
| permissions | -rw-r--r-- | 
| 59202 | 1  | 
/* Title: Tools/Graphview/mutator.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  | 
|
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
5  | 
Filters and add-operations on graphs.  | 
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
6  | 
*/  | 
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
7  | 
|
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
8  | 
package isabelle.graphview  | 
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
9  | 
|
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
10  | 
|
| 49558 | 11  | 
import isabelle._  | 
12  | 
||
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
13  | 
import java.awt.Color  | 
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
14  | 
import scala.collection.immutable.SortedSet  | 
| 
 
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  | 
|
| 59221 | 17  | 
object Mutator  | 
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
18  | 
{
 | 
| 59221 | 19  | 
sealed case class Info(enabled: Boolean, color: Color, mutator: Mutator)  | 
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
20  | 
|
| 59459 | 21  | 
def make(graphview: Graphview, m: Mutator): Info =  | 
22  | 
Info(true, graphview.Colors.next, m)  | 
|
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
23  | 
|
| 59218 | 24  | 
class Graph_Filter(  | 
25  | 
val name: String,  | 
|
26  | 
val description: String,  | 
|
| 
59245
 
be4180f3c236
more formal Graph_Display.Node (with ordering) and Graph_Display.Edge;
 
wenzelm 
parents: 
59240 
diff
changeset
 | 
27  | 
pred: Graph_Display.Graph => Graph_Display.Graph) extends Filter  | 
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
28  | 
  {
 | 
| 
59245
 
be4180f3c236
more formal Graph_Display.Node (with ordering) and Graph_Display.Edge;
 
wenzelm 
parents: 
59240 
diff
changeset
 | 
29  | 
def filter(graph: Graph_Display.Graph): Graph_Display.Graph = pred(graph)  | 
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
30  | 
}  | 
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
31  | 
|
| 59218 | 32  | 
class Graph_Mutator(  | 
33  | 
val name: String,  | 
|
34  | 
val description: String,  | 
|
| 
59245
 
be4180f3c236
more formal Graph_Display.Node (with ordering) and Graph_Display.Edge;
 
wenzelm 
parents: 
59240 
diff
changeset
 | 
35  | 
pred: (Graph_Display.Graph, Graph_Display.Graph) => Graph_Display.Graph) extends Mutator  | 
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
36  | 
  {
 | 
| 
59259
 
399506ee38a5
clarified static full_graph vs. dynamic visible_graph;
 
wenzelm 
parents: 
59246 
diff
changeset
 | 
37  | 
def mutate(full_graph: Graph_Display.Graph, graph: Graph_Display.Graph): Graph_Display.Graph =  | 
| 
 
399506ee38a5
clarified static full_graph vs. dynamic visible_graph;
 
wenzelm 
parents: 
59246 
diff
changeset
 | 
38  | 
pred(full_graph, graph)  | 
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
39  | 
}  | 
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
40  | 
|
| 59218 | 41  | 
class Node_Filter(  | 
42  | 
name: String,  | 
|
43  | 
description: String,  | 
|
| 
59245
 
be4180f3c236
more formal Graph_Display.Node (with ordering) and Graph_Display.Edge;
 
wenzelm 
parents: 
59240 
diff
changeset
 | 
44  | 
pred: (Graph_Display.Graph, Graph_Display.Node) => Boolean)  | 
| 59218 | 45  | 
extends Graph_Filter(name, description, g => g.restrict(pred(g, _)))  | 
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
46  | 
|
| 59218 | 47  | 
class Edge_Filter(  | 
48  | 
name: String,  | 
|
49  | 
description: String,  | 
|
| 59246 | 50  | 
pred: (Graph_Display.Graph, Graph_Display.Edge) => Boolean)  | 
| 59218 | 51  | 
extends Graph_Filter(  | 
52  | 
name,  | 
|
53  | 
description,  | 
|
54  | 
      g => (g /: g.dest) {
 | 
|
55  | 
case (graph, ((from, _), tos)) =>  | 
|
56  | 
(graph /: tos)((gr, to) =>  | 
|
| 59246 | 57  | 
if (pred(gr, (from, to))) gr else gr.del_edge(from, to))  | 
| 59218 | 58  | 
})  | 
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
59  | 
|
| 59218 | 60  | 
class Node_Family_Filter(  | 
61  | 
name: String,  | 
|
62  | 
description: String,  | 
|
63  | 
reverse: Boolean,  | 
|
64  | 
parents: Boolean,  | 
|
65  | 
children: Boolean,  | 
|
| 
59245
 
be4180f3c236
more formal Graph_Display.Node (with ordering) and Graph_Display.Edge;
 
wenzelm 
parents: 
59240 
diff
changeset
 | 
66  | 
pred: (Graph_Display.Graph, Graph_Display.Node) => Boolean)  | 
| 
49565
 
ea4308b7ef0f
ML support for generic graph display, with browser and graphview backends (via print modes);
 
wenzelm 
parents: 
49563 
diff
changeset
 | 
67  | 
extends Node_Filter(  | 
| 59218 | 68  | 
name,  | 
69  | 
description,  | 
|
70  | 
(g, k) =>  | 
|
71  | 
reverse !=  | 
|
72  | 
(pred(g, k) ||  | 
|
73  | 
(parents && g.all_preds(List(k)).exists(pred(g, _))) ||  | 
|
74  | 
(children && g.all_succs(List(k)).exists(pred(g, _)))))  | 
|
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
75  | 
|
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
76  | 
case class Identity()  | 
| 
49565
 
ea4308b7ef0f
ML support for generic graph display, with browser and graphview backends (via print modes);
 
wenzelm 
parents: 
49563 
diff
changeset
 | 
77  | 
extends Graph_Filter(  | 
| 59218 | 78  | 
"Identity",  | 
79  | 
"Does not change the graph.",  | 
|
80  | 
g => g)  | 
|
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
81  | 
|
| 59218 | 82  | 
case class Node_Expression(  | 
83  | 
regex: String,  | 
|
84  | 
reverse: Boolean,  | 
|
85  | 
parents: Boolean,  | 
|
86  | 
children: Boolean)  | 
|
87  | 
extends Node_Family_Filter(  | 
|
88  | 
"Filter by Name",  | 
|
89  | 
"Only shows or hides all nodes with any family member's name matching a regex.",  | 
|
90  | 
reverse,  | 
|
91  | 
parents,  | 
|
92  | 
children,  | 
|
| 
59245
 
be4180f3c236
more formal Graph_Display.Node (with ordering) and Graph_Display.Edge;
 
wenzelm 
parents: 
59240 
diff
changeset
 | 
93  | 
(g, node) => (regex.r findFirstIn node.toString).isDefined)  | 
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
94  | 
|
| 59218 | 95  | 
case class Node_List(  | 
| 
59245
 
be4180f3c236
more formal Graph_Display.Node (with ordering) and Graph_Display.Edge;
 
wenzelm 
parents: 
59240 
diff
changeset
 | 
96  | 
list: List[Graph_Display.Node],  | 
| 59218 | 97  | 
reverse: Boolean,  | 
98  | 
parents: Boolean,  | 
|
99  | 
children: Boolean)  | 
|
| 
49565
 
ea4308b7ef0f
ML support for generic graph display, with browser and graphview backends (via print modes);
 
wenzelm 
parents: 
49563 
diff
changeset
 | 
100  | 
extends Node_Family_Filter(  | 
| 59218 | 101  | 
"Filter by Name List",  | 
102  | 
"Only shows or hides all nodes with any family member's name matching any string in a comma separated list.",  | 
|
103  | 
reverse,  | 
|
104  | 
parents,  | 
|
105  | 
children,  | 
|
| 
59245
 
be4180f3c236
more formal Graph_Display.Node (with ordering) and Graph_Display.Edge;
 
wenzelm 
parents: 
59240 
diff
changeset
 | 
106  | 
(g, node) => list.contains(node))  | 
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
107  | 
|
| 59246 | 108  | 
case class Edge_Endpoints(edge: Graph_Display.Edge)  | 
| 
49565
 
ea4308b7ef0f
ML support for generic graph display, with browser and graphview backends (via print modes);
 
wenzelm 
parents: 
49563 
diff
changeset
 | 
109  | 
extends Edge_Filter(  | 
| 59218 | 110  | 
"Hide edge",  | 
| 59246 | 111  | 
"Hides specified edge.",  | 
112  | 
(g, e) => e != edge)  | 
|
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
113  | 
|
| 
59245
 
be4180f3c236
more formal Graph_Display.Node (with ordering) and Graph_Display.Edge;
 
wenzelm 
parents: 
59240 
diff
changeset
 | 
114  | 
private def add_node_group(from: Graph_Display.Graph, to: Graph_Display.Graph,  | 
| 
 
be4180f3c236
more formal Graph_Display.Node (with ordering) and Graph_Display.Edge;
 
wenzelm 
parents: 
59240 
diff
changeset
 | 
115  | 
keys: List[Graph_Display.Node]) =  | 
| 59218 | 116  | 
  {
 | 
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
117  | 
// Add Nodes  | 
| 59218 | 118  | 
val with_nodes =  | 
119  | 
(to /: keys)((graph, key) => graph.default_node(key, from.get_node(key)))  | 
|
120  | 
||
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
121  | 
// Add Edges  | 
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
122  | 
    (with_nodes /: keys) {
 | 
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
123  | 
      (gv, key) => {
 | 
| 
59245
 
be4180f3c236
more formal Graph_Display.Node (with ordering) and Graph_Display.Edge;
 
wenzelm 
parents: 
59240 
diff
changeset
 | 
124  | 
def add_edges(g: Graph_Display.Graph, keys: from.Keys, succs: Boolean) =  | 
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
125  | 
          (g /: keys) {
 | 
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
126  | 
            (graph, end) => {
 | 
| 
56372
 
fadb0fef09d7
more explicit iterator terminology, in accordance to Scala 2.8 library;
 
wenzelm 
parents: 
50475 
diff
changeset
 | 
127  | 
if (!graph.keys_iterator.contains(end)) graph  | 
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
128  | 
              else {
 | 
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
129  | 
if (succs) graph.add_edge(key, end)  | 
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
130  | 
else graph.add_edge(end, key)  | 
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
131  | 
}  | 
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
132  | 
}  | 
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
133  | 
}  | 
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
134  | 
|
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
135  | 
add_edges(  | 
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
136  | 
add_edges(gv, from.imm_preds(key), false),  | 
| 59218 | 137  | 
from.imm_succs(key), true)  | 
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
138  | 
}  | 
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
139  | 
}  | 
| 59218 | 140  | 
}  | 
141  | 
||
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
142  | 
case class Add_Node_Expression(regex: String)  | 
| 
49565
 
ea4308b7ef0f
ML support for generic graph display, with browser and graphview backends (via print modes);
 
wenzelm 
parents: 
49563 
diff
changeset
 | 
143  | 
extends Graph_Mutator(  | 
| 59218 | 144  | 
"Add by name",  | 
145  | 
"Adds every node whose name matches the regex. " +  | 
|
146  | 
"Adds all relevant edges.",  | 
|
| 
59259
 
399506ee38a5
clarified static full_graph vs. dynamic visible_graph;
 
wenzelm 
parents: 
59246 
diff
changeset
 | 
147  | 
(full_graph, graph) =>  | 
| 
 
399506ee38a5
clarified static full_graph vs. dynamic visible_graph;
 
wenzelm 
parents: 
59246 
diff
changeset
 | 
148  | 
add_node_group(full_graph, graph,  | 
| 
 
399506ee38a5
clarified static full_graph vs. dynamic visible_graph;
 
wenzelm 
parents: 
59246 
diff
changeset
 | 
149  | 
full_graph.keys.filter(node => (regex.r findFirstIn node.toString).isDefined).toList))  | 
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
150  | 
|
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
151  | 
case class Add_Transitive_Closure(parents: Boolean, children: Boolean)  | 
| 
49565
 
ea4308b7ef0f
ML support for generic graph display, with browser and graphview backends (via print modes);
 
wenzelm 
parents: 
49563 
diff
changeset
 | 
152  | 
extends Graph_Mutator(  | 
| 59218 | 153  | 
"Add transitive closure",  | 
154  | 
"Adds all family members of all current nodes.",  | 
|
| 
59259
 
399506ee38a5
clarified static full_graph vs. dynamic visible_graph;
 
wenzelm 
parents: 
59246 
diff
changeset
 | 
155  | 
      (full_graph, graph) => {
 | 
| 59218 | 156  | 
val withparents =  | 
| 
59259
 
399506ee38a5
clarified static full_graph vs. dynamic visible_graph;
 
wenzelm 
parents: 
59246 
diff
changeset
 | 
157  | 
if (parents) add_node_group(full_graph, graph, full_graph.all_preds(graph.keys))  | 
| 
59228
 
56b34fc7a015
more dynamic visualizer -- re-use Isabelle/jEdit options;
 
wenzelm 
parents: 
59221 
diff
changeset
 | 
158  | 
else graph  | 
| 
 
56b34fc7a015
more dynamic visualizer -- re-use Isabelle/jEdit options;
 
wenzelm 
parents: 
59221 
diff
changeset
 | 
159  | 
if (children)  | 
| 
59259
 
399506ee38a5
clarified static full_graph vs. dynamic visible_graph;
 
wenzelm 
parents: 
59246 
diff
changeset
 | 
160  | 
add_node_group(full_graph, withparents, full_graph.all_succs(graph.keys))  | 
| 59218 | 161  | 
else withparents  | 
162  | 
})  | 
|
| 59221 | 163  | 
}  | 
164  | 
||
165  | 
trait Mutator  | 
|
166  | 
{
 | 
|
167  | 
val name: String  | 
|
168  | 
val description: String  | 
|
| 
59259
 
399506ee38a5
clarified static full_graph vs. dynamic visible_graph;
 
wenzelm 
parents: 
59246 
diff
changeset
 | 
169  | 
def mutate(full_graph: Graph_Display.Graph, graph: Graph_Display.Graph): Graph_Display.Graph  | 
| 59221 | 170  | 
|
171  | 
override def toString: String = name  | 
|
172  | 
}  | 
|
173  | 
||
174  | 
trait Filter extends Mutator  | 
|
175  | 
{
 | 
|
| 71601 | 176  | 
def mutate(full_graph: Graph_Display.Graph, graph: Graph_Display.Graph): Graph_Display.Graph = filter(graph)  | 
| 
59245
 
be4180f3c236
more formal Graph_Display.Node (with ordering) and Graph_Display.Edge;
 
wenzelm 
parents: 
59240 
diff
changeset
 | 
177  | 
def filter(graph: Graph_Display.Graph): Graph_Display.Graph  | 
| 59221 | 178  | 
}  |