| author | wenzelm | 
| Mon, 13 Sep 2010 00:10:29 +0200 | |
| changeset 39293 | 651e5a3e8cfd | 
| parent 39173 | ed3946086358 | 
| child 39441 | 4110cc1b8f9f | 
| permissions | -rw-r--r-- | 
| 36676 | 1  | 
/* Title: Pure/PIDE/command.scala  | 
2  | 
Author: Fabian Immler, TU Munich  | 
|
3  | 
Author: Makarius  | 
|
4  | 
||
5  | 
Prover commands with semantic state.  | 
|
6  | 
*/  | 
|
| 34407 | 7  | 
|
| 
34871
 
e596a0b71f3c
incorporate "proofdocument" part into main Isabelle/Pure.jar -- except for html_panel.scala, which depends on external library (Lobo/Cobra browser);
 
wenzelm 
parents: 
34865 
diff
changeset
 | 
8  | 
package isabelle  | 
| 
34318
 
c13e168a8ae6
original sources from Johannes Hölzl a48e0c6ab1aea77c52d596f7efc007a543d3d10c with minor modifications of directory layout;
 
wenzelm 
parents:  
diff
changeset
 | 
9  | 
|
| 34451 | 10  | 
|
| 38872 | 11  | 
import scala.collection.immutable.SortedMap  | 
12  | 
||
13  | 
||
| 34637 | 14  | 
object Command  | 
15  | 
{
 | 
|
| 38361 | 16  | 
/** accumulated results from prover **/  | 
17  | 
||
| 38362 | 18  | 
case class State(  | 
| 38361 | 19  | 
val command: Command,  | 
| 38872 | 20  | 
val status: List[Markup] = Nil,  | 
21  | 
val results: SortedMap[Long, XML.Tree] = SortedMap.empty,  | 
|
22  | 
val markup: Markup_Tree = Markup_Tree.empty)  | 
|
| 38361 | 23  | 
  {
 | 
24  | 
/* content */  | 
|
25  | 
||
| 38714 | 26  | 
def add_status(st: Markup): State = copy(status = st :: status)  | 
27  | 
def add_markup(info: Text.Info[Any]): State = copy(markup = markup + info)  | 
|
| 38872 | 28  | 
def add_result(serial: Long, result: XML.Tree): State =  | 
29  | 
copy(results = results + (serial -> result))  | 
|
| 38361 | 30  | 
|
| 38658 | 31  | 
def root_info: Text.Info[Any] =  | 
| 
38581
 
d503a0912e14
simplified Command.status again, reverting most of e5eed57913d0 (note that more complex information can be represented with full markup reports);
 
wenzelm 
parents: 
38579 
diff
changeset
 | 
32  | 
new Text.Info(command.range,  | 
| 38658 | 33  | 
XML.Elem(Markup(Markup.STATUS, Nil), status.reverse.map(XML.Elem(_, Nil))))  | 
34  | 
def root_markup: Markup_Tree = markup + root_info  | 
|
| 38361 | 35  | 
|
36  | 
||
37  | 
/* message dispatch */  | 
|
38  | 
||
| 38872 | 39  | 
def accumulate(message: XML.Elem): Command.State =  | 
| 38361 | 40  | 
      message match {
 | 
| 38714 | 41  | 
case XML.Elem(Markup(Markup.STATUS, _), msgs) =>  | 
42  | 
(this /: msgs)((state, msg) =>  | 
|
43  | 
            msg match {
 | 
|
44  | 
case XML.Elem(markup, Nil) => state.add_status(markup)  | 
|
45  | 
              case _ => System.err.println("Ignored status message: " + msg); state
 | 
|
46  | 
})  | 
|
| 
38581
 
d503a0912e14
simplified Command.status again, reverting most of e5eed57913d0 (note that more complex information can be represented with full markup reports);
 
wenzelm 
parents: 
38579 
diff
changeset
 | 
47  | 
|
| 
38572
 
0fe2c01ef7da
Command.State: accumulate markup reports uniformly;
 
wenzelm 
parents: 
38564 
diff
changeset
 | 
48  | 
case XML.Elem(Markup(Markup.REPORT, _), msgs) =>  | 
| 
 
0fe2c01ef7da
Command.State: accumulate markup reports uniformly;
 
wenzelm 
parents: 
38564 
diff
changeset
 | 
49  | 
(this /: msgs)((state, msg) =>  | 
| 
 
0fe2c01ef7da
Command.State: accumulate markup reports uniformly;
 
wenzelm 
parents: 
38564 
diff
changeset
 | 
50  | 
            msg match {
 | 
| 
39173
 
ed3946086358
Command.State.accumulate: check actual source range;
 
wenzelm 
parents: 
39172 
diff
changeset
 | 
51  | 
case XML.Elem(Markup(name, atts @ Position.Id_Range(id, raw_range)), args)  | 
| 
 
ed3946086358
Command.State.accumulate: check actual source range;
 
wenzelm 
parents: 
39172 
diff
changeset
 | 
52  | 
if id == command.id && command.range.contains(command.decode(raw_range)) =>  | 
| 
 
ed3946086358
Command.State.accumulate: check actual source range;
 
wenzelm 
parents: 
39172 
diff
changeset
 | 
53  | 
val range = command.decode(raw_range)  | 
| 38872 | 54  | 
val props = Position.purge(atts)  | 
| 
39173
 
ed3946086358
Command.State.accumulate: check actual source range;
 
wenzelm 
parents: 
39172 
diff
changeset
 | 
55  | 
val info = Text.Info[Any](range, XML.Elem(Markup(name, props), args))  | 
| 38723 | 56  | 
state.add_markup(info)  | 
| 
38572
 
0fe2c01ef7da
Command.State: accumulate markup reports uniformly;
 
wenzelm 
parents: 
38564 
diff
changeset
 | 
57  | 
              case _ => System.err.println("Ignored report message: " + msg); state
 | 
| 38361 | 58  | 
})  | 
| 38872 | 59  | 
case XML.Elem(Markup(name, atts), body) =>  | 
60  | 
          atts match {
 | 
|
61  | 
case Markup.Serial(i) =>  | 
|
| 
38887
 
1261481ef5e5
Command.State: add reported positions to markup tree, according main message position or Markup.binding/entity/report occurrences in body;
 
wenzelm 
parents: 
38877 
diff
changeset
 | 
62  | 
val result = XML.Elem(Markup(name, Position.purge(atts)), body)  | 
| 
39172
 
31b95e0da7b7
Isar_Document.reported_positions: slightly more robust treatment of positions outside the command range, notably parsing beyond EOF;
 
wenzelm 
parents: 
38887 
diff
changeset
 | 
63  | 
(add_result(i, result) /: Isar_Document.reported_positions(command, message))(  | 
| 
 
31b95e0da7b7
Isar_Document.reported_positions: slightly more robust treatment of positions outside the command range, notably parsing beyond EOF;
 
wenzelm 
parents: 
38887 
diff
changeset
 | 
64  | 
(st, range) => st.add_markup(Text.Info(range, result)))  | 
| 38872 | 65  | 
            case _ => System.err.println("Ignored message without serial number: " + message); this
 | 
66  | 
}  | 
|
| 38361 | 67  | 
}  | 
68  | 
}  | 
|
| 38367 | 69  | 
|
70  | 
||
71  | 
/* unparsed dummy commands */  | 
|
72  | 
||
73  | 
def unparsed(source: String) =  | 
|
74  | 
new Command(Document.NO_ID, List(Token(Token.Kind.UNPARSED, source)))  | 
|
| 
34318
 
c13e168a8ae6
original sources from Johannes Hölzl a48e0c6ab1aea77c52d596f7efc007a543d3d10c with minor modifications of directory layout;
 
wenzelm 
parents:  
diff
changeset
 | 
75  | 
}  | 
| 
 
c13e168a8ae6
original sources from Johannes Hölzl a48e0c6ab1aea77c52d596f7efc007a543d3d10c with minor modifications of directory layout;
 
wenzelm 
parents:  
diff
changeset
 | 
76  | 
|
| 38361 | 77  | 
|
| 34697 | 78  | 
class Command(  | 
| 
38150
 
67fc24df3721
simplified/refined document model: collection of named nodes, without proper dependencies yet;
 
wenzelm 
parents: 
37373 
diff
changeset
 | 
79  | 
val id: Document.Command_ID,  | 
| 38373 | 80  | 
val span: List[Token])  | 
| 34451 | 81  | 
{
 | 
| 34859 | 82  | 
/* classification */  | 
| 
34500
 
384427c750c8
state_results: separate buffer for messages from running command;
 
wenzelm 
parents: 
34497 
diff
changeset
 | 
83  | 
|
| 36012 | 84  | 
def is_command: Boolean = !span.isEmpty && span.head.is_command  | 
| 34865 | 85  | 
def is_ignored: Boolean = span.forall(_.is_ignored)  | 
| 34859 | 86  | 
def is_malformed: Boolean = !is_command && !is_ignored  | 
87  | 
||
| 38367 | 88  | 
def is_unparsed = id == Document.NO_ID  | 
89  | 
||
| 36012 | 90  | 
def name: String = if (is_command) span.head.content else ""  | 
| 37129 | 91  | 
override def toString =  | 
| 
37373
 
25078ba44436
tuned Command.toString -- preserving uniqueness allows the Scala toplevel to print Linear_Set[Command] results without crashing;
 
wenzelm 
parents: 
37197 
diff
changeset
 | 
92  | 
id + "/" + (if (is_command) name else if (is_ignored) "IGNORED" else "MALFORMED")  | 
| 34495 | 93  | 
|
| 34859 | 94  | 
|
95  | 
/* source text */  | 
|
| 34451 | 96  | 
|
| 34859 | 97  | 
val source: String = span.map(_.source).mkString  | 
| 38426 | 98  | 
def source(range: Text.Range): String = source.substring(range.start, range.stop)  | 
| 34859 | 99  | 
def length: Int = source.length  | 
| 
38572
 
0fe2c01ef7da
Command.State: accumulate markup reports uniformly;
 
wenzelm 
parents: 
38564 
diff
changeset
 | 
100  | 
|
| 38877 | 101  | 
val newlines =  | 
102  | 
    (0 /: Symbol.iterator(source)) {
 | 
|
103  | 
case (n, s) => if (Symbol.is_physical_newline(s)) n + 1 else n }  | 
|
104  | 
||
| 
38479
 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 
wenzelm 
parents: 
38476 
diff
changeset
 | 
105  | 
val range: Text.Range = Text.Range(0, length)  | 
| 
34855
 
81d0410dc3ac
iterators for ranges of commands/starts -- avoid extra array per document;
 
wenzelm 
parents: 
34835 
diff
changeset
 | 
106  | 
|
| 34859 | 107  | 
lazy val symbol_index = new Symbol.Index(source)  | 
| 38579 | 108  | 
def decode(i: Text.Offset): Text.Offset = symbol_index.decode(i)  | 
109  | 
def decode(r: Text.Range): Text.Range = symbol_index.decode(r)  | 
|
| 
38370
 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 
wenzelm 
parents: 
38367 
diff
changeset
 | 
110  | 
|
| 
 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 
wenzelm 
parents: 
38367 
diff
changeset
 | 
111  | 
|
| 
 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 
wenzelm 
parents: 
38367 
diff
changeset
 | 
112  | 
/* accumulated results */  | 
| 
 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 
wenzelm 
parents: 
38367 
diff
changeset
 | 
113  | 
|
| 38872 | 114  | 
val empty_state: Command.State = Command.State(this)  | 
| 
34676
 
9e725d34df7b
Command and Command_State handle results from prover as Accumulator
 
immler@in.tum.de 
parents: 
34675 
diff
changeset
 | 
115  | 
}  |