| author | wenzelm | 
| Mon, 05 Jan 2015 23:29:38 +0100 | |
| changeset 59294 | 126293918a37 | 
| parent 59259 | 399506ee38a5 | 
| child 71601 | 97ccf48c2f0c | 
| permissions | -rw-r--r-- | 
| 59202 | 1  | 
/* Title: Tools/Graphview/model.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  | 
Internal graph representation.  | 
| 
 
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  | 
|
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
11  | 
import isabelle._  | 
| 55618 | 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  | 
|
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
15  | 
|
| 59218 | 16  | 
class Mutator_Container(val available: List[Mutator])  | 
17  | 
{
 | 
|
18  | 
val events = new Mutator_Event.Bus  | 
|
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
19  | 
|
| 59221 | 20  | 
private var _mutators : List[Mutator.Info] = Nil  | 
| 59218 | 21  | 
def apply() = _mutators  | 
| 59221 | 22  | 
def apply(mutators: List[Mutator.Info])  | 
| 59218 | 23  | 
  {
 | 
24  | 
_mutators = mutators  | 
|
| 59221 | 25  | 
events.event(Mutator_Event.New_List(mutators))  | 
| 59218 | 26  | 
}  | 
27  | 
||
| 59221 | 28  | 
def add(mutator: Mutator.Info)  | 
| 59218 | 29  | 
  {
 | 
30  | 
_mutators = _mutators ::: List(mutator)  | 
|
31  | 
events.event(Mutator_Event.Add(mutator))  | 
|
32  | 
}  | 
|
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
33  | 
}  | 
| 
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
34  | 
|
| 
49565
 
ea4308b7ef0f
ML support for generic graph display, with browser and graphview backends (via print modes);
 
wenzelm 
parents: 
49557 
diff
changeset
 | 
35  | 
|
| 
59259
 
399506ee38a5
clarified static full_graph vs. dynamic visible_graph;
 
wenzelm 
parents: 
59246 
diff
changeset
 | 
36  | 
class Model(val full_graph: Graph_Display.Graph)  | 
| 59218 | 37  | 
{
 | 
| 59221 | 38  | 
val Mutators =  | 
39  | 
new Mutator_Container(  | 
|
40  | 
List(  | 
|
41  | 
        Mutator.Node_Expression(".*", false, false, false),
 | 
|
42  | 
Mutator.Node_List(Nil, false, false, false),  | 
|
| 59246 | 43  | 
Mutator.Edge_Endpoints((Graph_Display.Node.dummy, Graph_Display.Node.dummy)),  | 
| 59221 | 44  | 
        Mutator.Add_Node_Expression(""),
 | 
45  | 
Mutator.Add_Transitive_Closure(true, true)))  | 
|
| 59218 | 46  | 
|
| 59221 | 47  | 
val Colors =  | 
48  | 
new Mutator_Container(  | 
|
49  | 
List(  | 
|
50  | 
        Mutator.Node_Expression(".*", false, false, false),
 | 
|
51  | 
Mutator.Node_List(Nil, false, false, false)))  | 
|
| 59218 | 52  | 
|
| 
59245
 
be4180f3c236
more formal Graph_Display.Node (with ordering) and Graph_Display.Edge;
 
wenzelm 
parents: 
59240 
diff
changeset
 | 
53  | 
def find_node(ident: String): Option[Graph_Display.Node] =  | 
| 
59259
 
399506ee38a5
clarified static full_graph vs. dynamic visible_graph;
 
wenzelm 
parents: 
59246 
diff
changeset
 | 
54  | 
full_graph.keys_iterator.find(node => node.ident == ident)  | 
| 
59245
 
be4180f3c236
more formal Graph_Display.Node (with ordering) and Graph_Display.Edge;
 
wenzelm 
parents: 
59240 
diff
changeset
 | 
55  | 
|
| 
59259
 
399506ee38a5
clarified static full_graph vs. dynamic visible_graph;
 
wenzelm 
parents: 
59246 
diff
changeset
 | 
56  | 
def make_visible_graph(): Graph_Display.Graph =  | 
| 
 
399506ee38a5
clarified static full_graph vs. dynamic visible_graph;
 
wenzelm 
parents: 
59246 
diff
changeset
 | 
57  | 
    (full_graph /: Mutators()) {
 | 
| 
 
399506ee38a5
clarified static full_graph vs. dynamic visible_graph;
 
wenzelm 
parents: 
59246 
diff
changeset
 | 
58  | 
case (g, m) => if (!m.enabled) g else m.mutator.mutate(full_graph, g)  | 
| 59218 | 59  | 
}  | 
60  | 
||
| 
59245
 
be4180f3c236
more formal Graph_Display.Node (with ordering) and Graph_Display.Edge;
 
wenzelm 
parents: 
59240 
diff
changeset
 | 
61  | 
private var _colors = Map.empty[Graph_Display.Node, Color]  | 
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
62  | 
def colors = _colors  | 
| 59218 | 63  | 
|
64  | 
private def build_colors()  | 
|
65  | 
  {
 | 
|
66  | 
_colors =  | 
|
| 
59245
 
be4180f3c236
more formal Graph_Display.Node (with ordering) and Graph_Display.Edge;
 
wenzelm 
parents: 
59240 
diff
changeset
 | 
67  | 
      (Map.empty[Graph_Display.Node, Color] /: Colors()) {
 | 
| 59221 | 68  | 
case (colors, m) =>  | 
| 
59259
 
399506ee38a5
clarified static full_graph vs. dynamic visible_graph;
 
wenzelm 
parents: 
59246 
diff
changeset
 | 
69  | 
          (colors /: m.mutator.mutate(full_graph, full_graph).keys_iterator) {
 | 
| 
59245
 
be4180f3c236
more formal Graph_Display.Node (with ordering) and Graph_Display.Edge;
 
wenzelm 
parents: 
59240 
diff
changeset
 | 
70  | 
case (colors, node) => colors + (node -> m.color)  | 
| 59218 | 71  | 
}  | 
72  | 
}  | 
|
| 
49557
 
61988f9df94d
added Graphview tool, based on Isabelle/Scala and Swing/Graphics2D;
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents:  
diff
changeset
 | 
73  | 
}  | 
| 
49733
 
38a68e6593be
prefer synchronous Mutator_Event.Bus on Swing_Thread;
 
wenzelm 
parents: 
49565 
diff
changeset
 | 
74  | 
  Colors.events += { case _ => build_colors() }
 | 
| 
49854
 
c541bbad7024
store colors after build
 
Markus Kaiser <markus.kaiser@in.tum.de> 
parents: 
49737 
diff
changeset
 | 
75  | 
}  |