| author | wenzelm | 
| Fri, 20 Apr 2012 20:21:22 +0200 | |
| changeset 47628 | 3275758d274e | 
| parent 47542 | 26d0a76fef0a | 
| child 48705 | dd32321d6eef | 
| permissions | -rw-r--r-- | 
| 
45709
 
87017fcbad83
clarified modules (again) -- NB: both Document and Protocol are specific to this particular prover;
 
wenzelm 
parents: 
45672 
diff
changeset
 | 
1  | 
/* Title: Pure/PIDE/protocol.scala  | 
| 
38412
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
Author: Makarius  | 
| 
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
|
| 
45709
 
87017fcbad83
clarified modules (again) -- NB: both Document and Protocol are specific to this particular prover;
 
wenzelm 
parents: 
45672 
diff
changeset
 | 
4  | 
Protocol message formats for interactive proof documents.  | 
| 
38412
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
5  | 
*/  | 
| 
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
7  | 
package isabelle  | 
| 
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
|
| 
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
9  | 
|
| 
45709
 
87017fcbad83
clarified modules (again) -- NB: both Document and Protocol are specific to this particular prover;
 
wenzelm 
parents: 
45672 
diff
changeset
 | 
10  | 
object Protocol  | 
| 
38412
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
11  | 
{
 | 
| 
38567
 
b670faa807c9
concentrate protocol message formats in Isar_Document;
 
wenzelm 
parents: 
38483 
diff
changeset
 | 
12  | 
/* document editing */  | 
| 
38412
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
13  | 
|
| 
44476
 
e8a87398f35d
propagate information about last command with exec state assignment through document model;
 
wenzelm 
parents: 
44474 
diff
changeset
 | 
14  | 
object Assign  | 
| 
 
e8a87398f35d
propagate information about last command with exec state assignment through document model;
 
wenzelm 
parents: 
44474 
diff
changeset
 | 
15  | 
  {
 | 
| 
44661
 
383c9d758a56
raw message function "assign_execs" avoids full overhead of decoding and caching message body;
 
wenzelm 
parents: 
44644 
diff
changeset
 | 
16  | 
def unapply(text: String): Option[(Document.Version_ID, Document.Assign)] =  | 
| 
 
383c9d758a56
raw message function "assign_execs" avoids full overhead of decoding and caching message body;
 
wenzelm 
parents: 
44644 
diff
changeset
 | 
17  | 
      try {
 | 
| 
 
383c9d758a56
raw message function "assign_execs" avoids full overhead of decoding and caching message body;
 
wenzelm 
parents: 
44644 
diff
changeset
 | 
18  | 
import XML.Decode._  | 
| 
 
383c9d758a56
raw message function "assign_execs" avoids full overhead of decoding and caching message body;
 
wenzelm 
parents: 
44644 
diff
changeset
 | 
19  | 
val body = YXML.parse_body(text)  | 
| 
47388
 
fe4b245af74c
discontinued obsolete last_execs (cf. cd3ab7625519);
 
wenzelm 
parents: 
47346 
diff
changeset
 | 
20  | 
Some(pair(long, list(pair(long, option(long))))(body))  | 
| 
44661
 
383c9d758a56
raw message function "assign_execs" avoids full overhead of decoding and caching message body;
 
wenzelm 
parents: 
44644 
diff
changeset
 | 
21  | 
}  | 
| 
 
383c9d758a56
raw message function "assign_execs" avoids full overhead of decoding and caching message body;
 
wenzelm 
parents: 
44644 
diff
changeset
 | 
22  | 
      catch {
 | 
| 
 
383c9d758a56
raw message function "assign_execs" avoids full overhead of decoding and caching message body;
 
wenzelm 
parents: 
44644 
diff
changeset
 | 
23  | 
case ERROR(_) => None  | 
| 
 
383c9d758a56
raw message function "assign_execs" avoids full overhead of decoding and caching message body;
 
wenzelm 
parents: 
44644 
diff
changeset
 | 
24  | 
case _: XML.XML_Atom => None  | 
| 
 
383c9d758a56
raw message function "assign_execs" avoids full overhead of decoding and caching message body;
 
wenzelm 
parents: 
44644 
diff
changeset
 | 
25  | 
case _: XML.XML_Body => None  | 
| 
38412
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
26  | 
}  | 
| 
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
27  | 
}  | 
| 
38567
 
b670faa807c9
concentrate protocol message formats in Isar_Document;
 
wenzelm 
parents: 
38483 
diff
changeset
 | 
28  | 
|
| 44676 | 29  | 
object Removed  | 
30  | 
  {
 | 
|
31  | 
def unapply(text: String): Option[List[Document.Version_ID]] =  | 
|
32  | 
      try {
 | 
|
33  | 
import XML.Decode._  | 
|
34  | 
Some(list(long)(YXML.parse_body(text)))  | 
|
35  | 
}  | 
|
36  | 
      catch {
 | 
|
37  | 
case ERROR(_) => None  | 
|
38  | 
case _: XML.XML_Atom => None  | 
|
39  | 
case _: XML.XML_Body => None  | 
|
40  | 
}  | 
|
41  | 
}  | 
|
42  | 
||
| 
38567
 
b670faa807c9
concentrate protocol message formats in Isar_Document;
 
wenzelm 
parents: 
38483 
diff
changeset
 | 
43  | 
|
| 46209 | 44  | 
/* command status */  | 
| 
38567
 
b670faa807c9
concentrate protocol message formats in Isar_Document;
 
wenzelm 
parents: 
38483 
diff
changeset
 | 
45  | 
|
| 
46227
 
4aa84f84d5e8
more precise rendering of overview_color/gutter_message/squiggly_underline based on cumulation of command status and warning/error messages;
 
wenzelm 
parents: 
46209 
diff
changeset
 | 
46  | 
object Status  | 
| 
 
4aa84f84d5e8
more precise rendering of overview_color/gutter_message/squiggly_underline based on cumulation of command status and warning/error messages;
 
wenzelm 
parents: 
46209 
diff
changeset
 | 
47  | 
  {
 | 
| 
 
4aa84f84d5e8
more precise rendering of overview_color/gutter_message/squiggly_underline based on cumulation of command status and warning/error messages;
 
wenzelm 
parents: 
46209 
diff
changeset
 | 
48  | 
val init = Status()  | 
| 
 
4aa84f84d5e8
more precise rendering of overview_color/gutter_message/squiggly_underline based on cumulation of command status and warning/error messages;
 
wenzelm 
parents: 
46209 
diff
changeset
 | 
49  | 
}  | 
| 
 
4aa84f84d5e8
more precise rendering of overview_color/gutter_message/squiggly_underline based on cumulation of command status and warning/error messages;
 
wenzelm 
parents: 
46209 
diff
changeset
 | 
50  | 
|
| 46166 | 51  | 
sealed case class Status(  | 
| 
47395
 
e6261a493f04
added static command status markup, to emphasize accepted but unassigned/unparsed commands (notably in overview panel);
 
wenzelm 
parents: 
47388 
diff
changeset
 | 
52  | 
private val accepted: Boolean = false,  | 
| 46166 | 53  | 
private val finished: Boolean = false,  | 
54  | 
private val failed: Boolean = false,  | 
|
55  | 
forks: Int = 0)  | 
|
| 
38567
 
b670faa807c9
concentrate protocol message formats in Isar_Document;
 
wenzelm 
parents: 
38483 
diff
changeset
 | 
56  | 
  {
 | 
| 
47395
 
e6261a493f04
added static command status markup, to emphasize accepted but unassigned/unparsed commands (notably in overview panel);
 
wenzelm 
parents: 
47388 
diff
changeset
 | 
57  | 
def is_unprocessed: Boolean = accepted && !finished && !failed && forks == 0  | 
| 46166 | 58  | 
def is_running: Boolean = forks != 0  | 
59  | 
def is_finished: Boolean = finished && forks == 0  | 
|
60  | 
def is_failed: Boolean = failed && forks == 0  | 
|
| 
46227
 
4aa84f84d5e8
more precise rendering of overview_color/gutter_message/squiggly_underline based on cumulation of command status and warning/error messages;
 
wenzelm 
parents: 
46209 
diff
changeset
 | 
61  | 
def + (that: Status): Status =  | 
| 
47395
 
e6261a493f04
added static command status markup, to emphasize accepted but unassigned/unparsed commands (notably in overview panel);
 
wenzelm 
parents: 
47388 
diff
changeset
 | 
62  | 
Status(accepted || that.accepted, finished || that.finished,  | 
| 
46910
 
3e068ef04b42
clarified command state -- markup within proper_range, excluding trailing whitespace;
 
wenzelm 
parents: 
46812 
diff
changeset
 | 
63  | 
failed || that.failed, forks + that.forks)  | 
| 46166 | 64  | 
}  | 
65  | 
||
66  | 
val command_status_markup: Set[String] =  | 
|
| 
47395
 
e6261a493f04
added static command status markup, to emphasize accepted but unassigned/unparsed commands (notably in overview panel);
 
wenzelm 
parents: 
47388 
diff
changeset
 | 
67  | 
Set(Isabelle_Markup.ACCEPTED, Isabelle_Markup.FINISHED, Isabelle_Markup.FAILED,  | 
| 46166 | 68  | 
Isabelle_Markup.FORKED, Isabelle_Markup.JOINED)  | 
69  | 
||
70  | 
def command_status(status: Status, markup: Markup): Status =  | 
|
71  | 
    markup match {
 | 
|
| 
47395
 
e6261a493f04
added static command status markup, to emphasize accepted but unassigned/unparsed commands (notably in overview panel);
 
wenzelm 
parents: 
47388 
diff
changeset
 | 
72  | 
case Markup(Isabelle_Markup.ACCEPTED, _) => status.copy(accepted = true)  | 
| 46166 | 73  | 
case Markup(Isabelle_Markup.FINISHED, _) => status.copy(finished = true)  | 
74  | 
case Markup(Isabelle_Markup.FAILED, _) => status.copy(failed = true)  | 
|
75  | 
case Markup(Isabelle_Markup.FORKED, _) => status.copy(forks = status.forks + 1)  | 
|
76  | 
case Markup(Isabelle_Markup.JOINED, _) => status.copy(forks = status.forks - 1)  | 
|
77  | 
case _ => status  | 
|
| 
38567
 
b670faa807c9
concentrate protocol message formats in Isar_Document;
 
wenzelm 
parents: 
38483 
diff
changeset
 | 
78  | 
}  | 
| 46166 | 79  | 
|
80  | 
def command_status(markups: List[Markup]): Status =  | 
|
| 
46227
 
4aa84f84d5e8
more precise rendering of overview_color/gutter_message/squiggly_underline based on cumulation of command status and warning/error messages;
 
wenzelm 
parents: 
46209 
diff
changeset
 | 
81  | 
(Status.init /: markups)(command_status(_, _))  | 
| 
38887
 
1261481ef5e5
Command.State: add reported positions to markup tree, according main message position or Markup.binding/entity/report occurrences in body;
 
wenzelm 
parents: 
38581 
diff
changeset
 | 
82  | 
|
| 46209 | 83  | 
|
84  | 
/* node status */  | 
|
85  | 
||
| 46688 | 86  | 
sealed case class Node_Status(  | 
87  | 
unprocessed: Int, running: Int, finished: Int, warned: Int, failed: Int)  | 
|
| 44866 | 88  | 
  {
 | 
| 46688 | 89  | 
def total: Int = unprocessed + running + finished + warned + failed  | 
| 44866 | 90  | 
}  | 
| 44613 | 91  | 
|
92  | 
def node_status(  | 
|
93  | 
state: Document.State, version: Document.Version, node: Document.Node): Node_Status =  | 
|
94  | 
  {
 | 
|
95  | 
var unprocessed = 0  | 
|
96  | 
var running = 0  | 
|
97  | 
var finished = 0  | 
|
| 46688 | 98  | 
var warned = 0  | 
| 44613 | 99  | 
var failed = 0  | 
100  | 
node.commands.foreach(command =>  | 
|
| 46166 | 101  | 
      {
 | 
| 46688 | 102  | 
val st = state.command_state(version, command)  | 
103  | 
val status = command_status(st.status)  | 
|
| 46166 | 104  | 
if (status.is_running) running += 1  | 
| 46688 | 105  | 
        else if (status.is_finished) {
 | 
106  | 
if (st.results.exists(p => is_warning(p._2))) warned += 1  | 
|
107  | 
else finished += 1  | 
|
108  | 
}  | 
|
| 46166 | 109  | 
else if (status.is_failed) failed += 1  | 
110  | 
else unprocessed += 1  | 
|
| 44613 | 111  | 
})  | 
| 46688 | 112  | 
Node_Status(unprocessed, running, finished, warned, failed)  | 
| 44613 | 113  | 
}  | 
114  | 
||
| 
38887
 
1261481ef5e5
Command.State: add reported positions to markup tree, according main message position or Markup.binding/entity/report occurrences in body;
 
wenzelm 
parents: 
38581 
diff
changeset
 | 
115  | 
|
| 
39441
 
4110cc1b8f9f
allow embedded reports in regular prover messages, to avoid side-effects for errors for example;
 
wenzelm 
parents: 
39439 
diff
changeset
 | 
116  | 
/* result messages */  | 
| 
39439
 
1c294d150ded
eliminated markup "location" in favour of more explicit "no_report", which is actually deleted from messages;
 
wenzelm 
parents: 
39181 
diff
changeset
 | 
117  | 
|
| 
 
1c294d150ded
eliminated markup "location" in favour of more explicit "no_report", which is actually deleted from messages;
 
wenzelm 
parents: 
39181 
diff
changeset
 | 
118  | 
def clean_message(body: XML.Body): XML.Body =  | 
| 45666 | 119  | 
    body filter { case XML.Elem(Markup(Isabelle_Markup.NO_REPORT, _), _) => false case _ => true } map
 | 
| 
39439
 
1c294d150ded
eliminated markup "location" in favour of more explicit "no_report", which is actually deleted from messages;
 
wenzelm 
parents: 
39181 
diff
changeset
 | 
120  | 
      { case XML.Elem(markup, ts) => XML.Elem(markup, clean_message(ts)) case t => t }
 | 
| 
 
1c294d150ded
eliminated markup "location" in favour of more explicit "no_report", which is actually deleted from messages;
 
wenzelm 
parents: 
39181 
diff
changeset
 | 
121  | 
|
| 
39441
 
4110cc1b8f9f
allow embedded reports in regular prover messages, to avoid side-effects for errors for example;
 
wenzelm 
parents: 
39439 
diff
changeset
 | 
122  | 
def message_reports(msg: XML.Tree): List[XML.Elem] =  | 
| 
 
4110cc1b8f9f
allow embedded reports in regular prover messages, to avoid side-effects for errors for example;
 
wenzelm 
parents: 
39439 
diff
changeset
 | 
123  | 
    msg match {
 | 
| 45666 | 124  | 
case elem @ XML.Elem(Markup(Isabelle_Markup.REPORT, _), _) => List(elem)  | 
| 
39441
 
4110cc1b8f9f
allow embedded reports in regular prover messages, to avoid side-effects for errors for example;
 
wenzelm 
parents: 
39439 
diff
changeset
 | 
125  | 
case XML.Elem(_, body) => body.flatMap(message_reports)  | 
| 
 
4110cc1b8f9f
allow embedded reports in regular prover messages, to avoid side-effects for errors for example;
 
wenzelm 
parents: 
39439 
diff
changeset
 | 
126  | 
case XML.Text(_) => Nil  | 
| 
 
4110cc1b8f9f
allow embedded reports in regular prover messages, to avoid side-effects for errors for example;
 
wenzelm 
parents: 
39439 
diff
changeset
 | 
127  | 
}  | 
| 
 
4110cc1b8f9f
allow embedded reports in regular prover messages, to avoid side-effects for errors for example;
 
wenzelm 
parents: 
39439 
diff
changeset
 | 
128  | 
|
| 
39439
 
1c294d150ded
eliminated markup "location" in favour of more explicit "no_report", which is actually deleted from messages;
 
wenzelm 
parents: 
39181 
diff
changeset
 | 
129  | 
|
| 39511 | 130  | 
/* specific messages */  | 
131  | 
||
| 39625 | 132  | 
def is_tracing(msg: XML.Tree): Boolean =  | 
| 
39622
 
53365ba766ac
Command.accumulate: refrain from adding tracing messages to markup tree -- potential scalability problem;
 
wenzelm 
parents: 
39511 
diff
changeset
 | 
133  | 
    msg match {
 | 
| 45666 | 134  | 
case XML.Elem(Markup(Isabelle_Markup.TRACING, _), _) => true  | 
| 
39622
 
53365ba766ac
Command.accumulate: refrain from adding tracing messages to markup tree -- potential scalability problem;
 
wenzelm 
parents: 
39511 
diff
changeset
 | 
135  | 
case _ => false  | 
| 
 
53365ba766ac
Command.accumulate: refrain from adding tracing messages to markup tree -- potential scalability problem;
 
wenzelm 
parents: 
39511 
diff
changeset
 | 
136  | 
}  | 
| 
 
53365ba766ac
Command.accumulate: refrain from adding tracing messages to markup tree -- potential scalability problem;
 
wenzelm 
parents: 
39511 
diff
changeset
 | 
137  | 
|
| 39511 | 138  | 
def is_warning(msg: XML.Tree): Boolean =  | 
139  | 
    msg match {
 | 
|
| 45666 | 140  | 
case XML.Elem(Markup(Isabelle_Markup.WARNING, _), _) => true  | 
| 39511 | 141  | 
case _ => false  | 
142  | 
}  | 
|
143  | 
||
144  | 
def is_error(msg: XML.Tree): Boolean =  | 
|
145  | 
    msg match {
 | 
|
| 45666 | 146  | 
case XML.Elem(Markup(Isabelle_Markup.ERROR, _), _) => true  | 
| 39511 | 147  | 
case _ => false  | 
148  | 
}  | 
|
| 
38887
 
1261481ef5e5
Command.State: add reported positions to markup tree, according main message position or Markup.binding/entity/report occurrences in body;
 
wenzelm 
parents: 
38581 
diff
changeset
 | 
149  | 
|
| 
39441
 
4110cc1b8f9f
allow embedded reports in regular prover messages, to avoid side-effects for errors for example;
 
wenzelm 
parents: 
39439 
diff
changeset
 | 
150  | 
def is_state(msg: XML.Tree): Boolean =  | 
| 
39170
 
04ad0fed81f5
Isar_Document.reported_positions: exclude proof state output;
 
wenzelm 
parents: 
39042 
diff
changeset
 | 
151  | 
    msg match {
 | 
| 45666 | 152  | 
case XML.Elem(Markup(Isabelle_Markup.WRITELN, _),  | 
153  | 
List(XML.Elem(Markup(Isabelle_Markup.STATE, _), _))) => true  | 
|
| 
39170
 
04ad0fed81f5
Isar_Document.reported_positions: exclude proof state output;
 
wenzelm 
parents: 
39042 
diff
changeset
 | 
154  | 
case _ => false  | 
| 
 
04ad0fed81f5
Isar_Document.reported_positions: exclude proof state output;
 
wenzelm 
parents: 
39042 
diff
changeset
 | 
155  | 
}  | 
| 
 
04ad0fed81f5
Isar_Document.reported_positions: exclude proof state output;
 
wenzelm 
parents: 
39042 
diff
changeset
 | 
156  | 
|
| 
47542
 
26d0a76fef0a
more robust Sendback handling: JVM/jEdit paranoia for case matching, treat Pretty body not just XML.Text, replace proper_range only (without trailing whitespace);
 
wenzelm 
parents: 
47395 
diff
changeset
 | 
157  | 
object Sendback  | 
| 
 
26d0a76fef0a
more robust Sendback handling: JVM/jEdit paranoia for case matching, treat Pretty body not just XML.Text, replace proper_range only (without trailing whitespace);
 
wenzelm 
parents: 
47395 
diff
changeset
 | 
158  | 
  {
 | 
| 
 
26d0a76fef0a
more robust Sendback handling: JVM/jEdit paranoia for case matching, treat Pretty body not just XML.Text, replace proper_range only (without trailing whitespace);
 
wenzelm 
parents: 
47395 
diff
changeset
 | 
159  | 
def unapply(msg: Any): Option[XML.Body] =  | 
| 
 
26d0a76fef0a
more robust Sendback handling: JVM/jEdit paranoia for case matching, treat Pretty body not just XML.Text, replace proper_range only (without trailing whitespace);
 
wenzelm 
parents: 
47395 
diff
changeset
 | 
160  | 
      msg match {
 | 
| 
 
26d0a76fef0a
more robust Sendback handling: JVM/jEdit paranoia for case matching, treat Pretty body not just XML.Text, replace proper_range only (without trailing whitespace);
 
wenzelm 
parents: 
47395 
diff
changeset
 | 
161  | 
case XML.Elem(Markup(Isabelle_Markup.SENDBACK, _), body) => Some(body)  | 
| 
 
26d0a76fef0a
more robust Sendback handling: JVM/jEdit paranoia for case matching, treat Pretty body not just XML.Text, replace proper_range only (without trailing whitespace);
 
wenzelm 
parents: 
47395 
diff
changeset
 | 
162  | 
case _ => None  | 
| 
 
26d0a76fef0a
more robust Sendback handling: JVM/jEdit paranoia for case matching, treat Pretty body not just XML.Text, replace proper_range only (without trailing whitespace);
 
wenzelm 
parents: 
47395 
diff
changeset
 | 
163  | 
}  | 
| 
 
26d0a76fef0a
more robust Sendback handling: JVM/jEdit paranoia for case matching, treat Pretty body not just XML.Text, replace proper_range only (without trailing whitespace);
 
wenzelm 
parents: 
47395 
diff
changeset
 | 
164  | 
}  | 
| 
 
26d0a76fef0a
more robust Sendback handling: JVM/jEdit paranoia for case matching, treat Pretty body not just XML.Text, replace proper_range only (without trailing whitespace);
 
wenzelm 
parents: 
47395 
diff
changeset
 | 
165  | 
|
| 39511 | 166  | 
|
167  | 
/* reported positions */  | 
|
168  | 
||
| 39627 | 169  | 
private val include_pos =  | 
| 45666 | 170  | 
Set(Isabelle_Markup.BINDING, Isabelle_Markup.ENTITY, Isabelle_Markup.REPORT,  | 
171  | 
Isabelle_Markup.POSITION)  | 
|
| 
39441
 
4110cc1b8f9f
allow embedded reports in regular prover messages, to avoid side-effects for errors for example;
 
wenzelm 
parents: 
39439 
diff
changeset
 | 
172  | 
|
| 
 
4110cc1b8f9f
allow embedded reports in regular prover messages, to avoid side-effects for errors for example;
 
wenzelm 
parents: 
39439 
diff
changeset
 | 
173  | 
def message_positions(command: Command, message: XML.Elem): Set[Text.Range] =  | 
| 
38887
 
1261481ef5e5
Command.State: add reported positions to markup tree, according main message position or Markup.binding/entity/report occurrences in body;
 
wenzelm 
parents: 
38581 
diff
changeset
 | 
174  | 
  {
 | 
| 
39441
 
4110cc1b8f9f
allow embedded reports in regular prover messages, to avoid side-effects for errors for example;
 
wenzelm 
parents: 
39439 
diff
changeset
 | 
175  | 
def positions(set: Set[Text.Range], tree: XML.Tree): Set[Text.Range] =  | 
| 
38887
 
1261481ef5e5
Command.State: add reported positions to markup tree, according main message position or Markup.binding/entity/report occurrences in body;
 
wenzelm 
parents: 
38581 
diff
changeset
 | 
176  | 
      tree match {
 | 
| 
39172
 
31b95e0da7b7
Isar_Document.reported_positions: slightly more robust treatment of positions outside the command range, notably parsing beyond EOF;
 
wenzelm 
parents: 
39170 
diff
changeset
 | 
177  | 
case XML.Elem(Markup(name, Position.Id_Range(id, raw_range)), body)  | 
| 
 
31b95e0da7b7
Isar_Document.reported_positions: slightly more robust treatment of positions outside the command range, notably parsing beyond EOF;
 
wenzelm 
parents: 
39170 
diff
changeset
 | 
178  | 
if include_pos(name) && id == command.id =>  | 
| 
 
31b95e0da7b7
Isar_Document.reported_positions: slightly more robust treatment of positions outside the command range, notably parsing beyond EOF;
 
wenzelm 
parents: 
39170 
diff
changeset
 | 
179  | 
val range = command.decode(raw_range).restrict(command.range)  | 
| 
39441
 
4110cc1b8f9f
allow embedded reports in regular prover messages, to avoid side-effects for errors for example;
 
wenzelm 
parents: 
39439 
diff
changeset
 | 
180  | 
body.foldLeft(if (range.is_singularity) set else set + range)(positions)  | 
| 
 
4110cc1b8f9f
allow embedded reports in regular prover messages, to avoid side-effects for errors for example;
 
wenzelm 
parents: 
39439 
diff
changeset
 | 
181  | 
case XML.Elem(Markup(name, _), body) => body.foldLeft(set)(positions)  | 
| 
39042
 
470fd769ae53
Isar_Document.reported_positions: more precise include/exclude, include root as last resort only;
 
wenzelm 
parents: 
38887 
diff
changeset
 | 
182  | 
case _ => set  | 
| 
38887
 
1261481ef5e5
Command.State: add reported positions to markup tree, according main message position or Markup.binding/entity/report occurrences in body;
 
wenzelm 
parents: 
38581 
diff
changeset
 | 
183  | 
}  | 
| 
39441
 
4110cc1b8f9f
allow embedded reports in regular prover messages, to avoid side-effects for errors for example;
 
wenzelm 
parents: 
39439 
diff
changeset
 | 
184  | 
val set = positions(Set.empty, message)  | 
| 
39170
 
04ad0fed81f5
Isar_Document.reported_positions: exclude proof state output;
 
wenzelm 
parents: 
39042 
diff
changeset
 | 
185  | 
if (set.isEmpty && !is_state(message))  | 
| 
39172
 
31b95e0da7b7
Isar_Document.reported_positions: slightly more robust treatment of positions outside the command range, notably parsing beyond EOF;
 
wenzelm 
parents: 
39170 
diff
changeset
 | 
186  | 
set ++ Position.Range.unapply(message.markup.properties).map(command.decode(_))  | 
| 
39042
 
470fd769ae53
Isar_Document.reported_positions: more precise include/exclude, include root as last resort only;
 
wenzelm 
parents: 
38887 
diff
changeset
 | 
187  | 
else set  | 
| 
38887
 
1261481ef5e5
Command.State: add reported positions to markup tree, according main message position or Markup.binding/entity/report occurrences in body;
 
wenzelm 
parents: 
38581 
diff
changeset
 | 
188  | 
}  | 
| 
38412
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
189  | 
}  | 
| 
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
190  | 
|
| 
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
191  | 
|
| 
45709
 
87017fcbad83
clarified modules (again) -- NB: both Document and Protocol are specific to this particular prover;
 
wenzelm 
parents: 
45672 
diff
changeset
 | 
192  | 
trait Protocol extends Isabelle_Process  | 
| 
38412
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
193  | 
{
 | 
| 
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
194  | 
/* commands */  | 
| 
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
195  | 
|
| 
44644
 
317e4962dd0f
clarified define_command: store name as structural information;
 
wenzelm 
parents: 
44615 
diff
changeset
 | 
196  | 
def define_command(command: Command): Unit =  | 
| 
45709
 
87017fcbad83
clarified modules (again) -- NB: both Document and Protocol are specific to this particular prover;
 
wenzelm 
parents: 
45672 
diff
changeset
 | 
197  | 
    input("Document.define_command",
 | 
| 
44644
 
317e4962dd0f
clarified define_command: store name as structural information;
 
wenzelm 
parents: 
44615 
diff
changeset
 | 
198  | 
Document.ID(command.id), Symbol.encode(command.name), Symbol.encode(command.source))  | 
| 
38412
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
199  | 
|
| 
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
200  | 
|
| 38417 | 201  | 
/* document versions */  | 
| 
38412
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
202  | 
|
| 
47343
 
b8aeab386414
less aggressive discontinue_execution before document update, to avoid unstable execs that need to be re-assigned;
 
wenzelm 
parents: 
46938 
diff
changeset
 | 
203  | 
  def discontinue_execution() { input("Document.discontinue_execution") }
 | 
| 
 
b8aeab386414
less aggressive discontinue_execution before document update, to avoid unstable execs that need to be re-assigned;
 
wenzelm 
parents: 
46938 
diff
changeset
 | 
204  | 
|
| 
 
b8aeab386414
less aggressive discontinue_execution before document update, to avoid unstable execs that need to be re-assigned;
 
wenzelm 
parents: 
46938 
diff
changeset
 | 
205  | 
  def cancel_execution() { input("Document.cancel_execution") }
 | 
| 
44612
 
990ac978854c
explicit cancel_execution before queueing new edits -- potential performance improvement for machines with few cores;
 
wenzelm 
parents: 
44481 
diff
changeset
 | 
206  | 
|
| 44481 | 207  | 
def update(old_id: Document.Version_ID, new_id: Document.Version_ID,  | 
| 44383 | 208  | 
edits: List[Document.Edit_Command])  | 
| 
38412
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
209  | 
  {
 | 
| 
44157
 
a21d3e1e64fd
uniform treatment of header edits as document edits;
 
wenzelm 
parents: 
44156 
diff
changeset
 | 
210  | 
val edits_yxml =  | 
| 43767 | 211  | 
    { import XML.Encode._
 | 
| 44383 | 212  | 
def id: T[Command] = (cmd => long(cmd.id))  | 
| 
46770
 
44c28a33c461
retain original "uses" (again) -- still required for Thy_Load.use_file etc. in ML (notably for maintaining required/provided);
 
wenzelm 
parents: 
46755 
diff
changeset
 | 
213  | 
def symbol_string: T[String] = (s => string(Symbol.encode(s)))  | 
| 46737 | 214  | 
def encode_edit(name: Document.Node.Name)  | 
| 
44979
 
68b990e950b1
explicit master_dir as part of header -- still required (for Cygwin) since Scala layer does not pass file content yet;
 
wenzelm 
parents: 
44866 
diff
changeset
 | 
215  | 
: T[Document.Node.Edit[(Option[Command], Option[Command]), Command.Perspective]] =  | 
| 
 
68b990e950b1
explicit master_dir as part of header -- still required (for Cygwin) since Scala layer does not pass file content yet;
 
wenzelm 
parents: 
44866 
diff
changeset
 | 
216  | 
variant(List(  | 
| 
 
68b990e950b1
explicit master_dir as part of header -- still required (for Cygwin) since Scala layer does not pass file content yet;
 
wenzelm 
parents: 
44866 
diff
changeset
 | 
217  | 
          { case Document.Node.Clear() => (Nil, Nil) },
 | 
| 
 
68b990e950b1
explicit master_dir as part of header -- still required (for Cygwin) since Scala layer does not pass file content yet;
 
wenzelm 
parents: 
44866 
diff
changeset
 | 
218  | 
          { case Document.Node.Edits(a) => (Nil, list(pair(option(id), option(id)))(a)) },
 | 
| 46737 | 219  | 
          { case Document.Node.Header(Exn.Res(deps)) =>
 | 
| 
46770
 
44c28a33c461
retain original "uses" (again) -- still required for Thy_Load.use_file etc. in ML (notably for maintaining required/provided);
 
wenzelm 
parents: 
46755 
diff
changeset
 | 
220  | 
val dir = Isabelle_System.posix_path(name.dir)  | 
| 46737 | 221  | 
val imports = deps.imports.map(_.node)  | 
| 
46770
 
44c28a33c461
retain original "uses" (again) -- still required for Thy_Load.use_file etc. in ML (notably for maintaining required/provided);
 
wenzelm 
parents: 
46755 
diff
changeset
 | 
222  | 
// FIXME val uses = deps.uses.map(p => (Isabelle_System.posix_path(p._1), p._2))  | 
| 
 
44c28a33c461
retain original "uses" (again) -- still required for Thy_Load.use_file etc. in ML (notably for maintaining required/provided);
 
wenzelm 
parents: 
46755 
diff
changeset
 | 
223  | 
val uses = deps.uses  | 
| 
44979
 
68b990e950b1
explicit master_dir as part of header -- still required (for Cygwin) since Scala layer does not pass file content yet;
 
wenzelm 
parents: 
44866 
diff
changeset
 | 
224  | 
(Nil,  | 
| 
46938
 
cda018294515
some support for outer syntax keyword declarations within theory header;
 
wenzelm 
parents: 
46910 
diff
changeset
 | 
225  | 
pair(pair(pair(pair(symbol_string, symbol_string), list(symbol_string)),  | 
| 
 
cda018294515
some support for outer syntax keyword declarations within theory header;
 
wenzelm 
parents: 
46910 
diff
changeset
 | 
226  | 
list(pair(symbol_string, option(pair(symbol_string, list(symbol_string)))))),  | 
| 
 
cda018294515
some support for outer syntax keyword declarations within theory header;
 
wenzelm 
parents: 
46910 
diff
changeset
 | 
227  | 
list(pair(symbol_string, bool)))(  | 
| 
 
cda018294515
some support for outer syntax keyword declarations within theory header;
 
wenzelm 
parents: 
46910 
diff
changeset
 | 
228  | 
(((dir, name.theory), imports), deps.keywords), uses)) },  | 
| 46755 | 229  | 
          { case Document.Node.Header(Exn.Exn(e)) => (List(Symbol.encode(Exn.message(e))), Nil) },
 | 
| 
44979
 
68b990e950b1
explicit master_dir as part of header -- still required (for Cygwin) since Scala layer does not pass file content yet;
 
wenzelm 
parents: 
44866 
diff
changeset
 | 
230  | 
          { case Document.Node.Perspective(a) => (a.commands.map(c => long_atom(c.id)), Nil) }))
 | 
| 
 
68b990e950b1
explicit master_dir as part of header -- still required (for Cygwin) since Scala layer does not pass file content yet;
 
wenzelm 
parents: 
44866 
diff
changeset
 | 
231  | 
def encode: T[List[Document.Edit_Command]] = list((node_edit: Document.Edit_Command) =>  | 
| 
 
68b990e950b1
explicit master_dir as part of header -- still required (for Cygwin) since Scala layer does not pass file content yet;
 
wenzelm 
parents: 
44866 
diff
changeset
 | 
232  | 
      {
 | 
| 
 
68b990e950b1
explicit master_dir as part of header -- still required (for Cygwin) since Scala layer does not pass file content yet;
 
wenzelm 
parents: 
44866 
diff
changeset
 | 
233  | 
val (name, edit) = node_edit  | 
| 46737 | 234  | 
pair(string, encode_edit(name))(name.node, edit)  | 
| 
44979
 
68b990e950b1
explicit master_dir as part of header -- still required (for Cygwin) since Scala layer does not pass file content yet;
 
wenzelm 
parents: 
44866 
diff
changeset
 | 
235  | 
})  | 
| 
44157
 
a21d3e1e64fd
uniform treatment of header edits as document edits;
 
wenzelm 
parents: 
44156 
diff
changeset
 | 
236  | 
YXML.string_of_body(encode(edits)) }  | 
| 
45709
 
87017fcbad83
clarified modules (again) -- NB: both Document and Protocol are specific to this particular prover;
 
wenzelm 
parents: 
45672 
diff
changeset
 | 
237  | 
    input("Document.update", Document.ID(old_id), Document.ID(new_id), edits_yxml)
 | 
| 
38412
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
238  | 
}  | 
| 43748 | 239  | 
|
| 44673 | 240  | 
def remove_versions(versions: List[Document.Version])  | 
241  | 
  {
 | 
|
242  | 
val versions_yxml =  | 
|
243  | 
      { import XML.Encode._
 | 
|
244  | 
YXML.string_of_body(list(long)(versions.map(_.id))) }  | 
|
| 
45709
 
87017fcbad83
clarified modules (again) -- NB: both Document and Protocol are specific to this particular prover;
 
wenzelm 
parents: 
45672 
diff
changeset
 | 
245  | 
    input("Document.remove_versions", versions_yxml)
 | 
| 44673 | 246  | 
}  | 
247  | 
||
| 43748 | 248  | 
|
249  | 
/* method invocation service */  | 
|
250  | 
||
251  | 
def invoke_scala(id: String, tag: Invoke_Scala.Tag.Value, res: String)  | 
|
252  | 
  {
 | 
|
| 
45709
 
87017fcbad83
clarified modules (again) -- NB: both Document and Protocol are specific to this particular prover;
 
wenzelm 
parents: 
45672 
diff
changeset
 | 
253  | 
    input("Document.invoke_scala", id, tag.toString, res)
 | 
| 43748 | 254  | 
}  | 
| 
38412
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
255  | 
}  |