| author | wenzelm | 
| Tue, 31 Jul 2018 21:06:09 +0200 | |
| changeset 68728 | c07f6fa02c59 | 
| parent 68381 | 2fd3a6d6ba2e | 
| child 68729 | 3a02b424d5fb | 
| permissions | -rw-r--r-- | 
| 36676 | 1  | 
/* Title: Pure/PIDE/command.scala  | 
2  | 
Author: Fabian Immler, TU Munich  | 
|
3  | 
Author: Makarius  | 
|
4  | 
||
| 52536 | 5  | 
Prover commands with accumulated results from execution.  | 
| 36676 | 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  | 
|
| 45644 | 11  | 
import scala.collection.mutable  | 
| 38872 | 12  | 
import scala.collection.immutable.SortedMap  | 
13  | 
||
14  | 
||
| 34637 | 15  | 
object Command  | 
16  | 
{
 | 
|
| 52849 | 17  | 
type Edit = (Option[Command], Option[Command])  | 
| 
59702
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
18  | 
|
| 56746 | 19  | 
type Blob = Exn.Result[(Document.Node.Name, Option[(SHA1.Digest, Symbol.Text_Chunk)])]  | 
| 
59702
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
20  | 
type Blobs_Info = (List[Blob], Int)  | 
| 
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
21  | 
val no_blobs: Blobs_Info = (Nil, -1)  | 
| 52849 | 22  | 
|
23  | 
||
| 38361 | 24  | 
/** accumulated results from prover **/  | 
25  | 
||
| 50507 | 26  | 
/* results */  | 
27  | 
||
28  | 
object Results  | 
|
29  | 
  {
 | 
|
| 
51496
 
cb677987b7e3
retain original tooltip range, to avoid repeated window popup when the mouse is moved over the same content;
 
wenzelm 
parents: 
51494 
diff
changeset
 | 
30  | 
type Entry = (Long, XML.Tree)  | 
| 68101 | 31  | 
val empty: Results = new Results(SortedMap.empty)  | 
| 59072 | 32  | 
def make(args: TraversableOnce[Results.Entry]): Results = (empty /: args)(_ + _)  | 
33  | 
def merge(args: TraversableOnce[Results]): Results = (empty /: args)(_ ++ _)  | 
|
| 65335 | 34  | 
|
35  | 
||
36  | 
/* XML data representation */  | 
|
37  | 
||
38  | 
val encode: XML.Encode.T[Results] = (results: Results) =>  | 
|
39  | 
    { import XML.Encode._; list(pair(long, tree))(results.rep.toList) }
 | 
|
40  | 
||
41  | 
val decode: XML.Decode.T[Results] = (body: XML.Body) =>  | 
|
42  | 
    { import XML.Decode._; make(list(pair(long, tree))(body)) }
 | 
|
| 50507 | 43  | 
}  | 
44  | 
||
| 51494 | 45  | 
final class Results private(private val rep: SortedMap[Long, XML.Tree])  | 
| 50507 | 46  | 
  {
 | 
| 64802 | 47  | 
def is_empty: Boolean = rep.isEmpty  | 
| 50507 | 48  | 
def defined(serial: Long): Boolean = rep.isDefinedAt(serial)  | 
49  | 
def get(serial: Long): Option[XML.Tree] = rep.get(serial)  | 
|
| 
56372
 
fadb0fef09d7
more explicit iterator terminology, in accordance to Scala 2.8 library;
 
wenzelm 
parents: 
56359 
diff
changeset
 | 
50  | 
def iterator: Iterator[Results.Entry] = rep.iterator  | 
| 
50508
 
5b7150395568
tuned implementation according to Library.insert/merge in ML;
 
wenzelm 
parents: 
50507 
diff
changeset
 | 
51  | 
|
| 
51496
 
cb677987b7e3
retain original tooltip range, to avoid repeated window popup when the mouse is moved over the same content;
 
wenzelm 
parents: 
51494 
diff
changeset
 | 
52  | 
def + (entry: Results.Entry): Results =  | 
| 
50508
 
5b7150395568
tuned implementation according to Library.insert/merge in ML;
 
wenzelm 
parents: 
50507 
diff
changeset
 | 
53  | 
if (defined(entry._1)) this  | 
| 
 
5b7150395568
tuned implementation according to Library.insert/merge in ML;
 
wenzelm 
parents: 
50507 
diff
changeset
 | 
54  | 
else new Results(rep + entry)  | 
| 
 
5b7150395568
tuned implementation according to Library.insert/merge in ML;
 
wenzelm 
parents: 
50507 
diff
changeset
 | 
55  | 
|
| 
 
5b7150395568
tuned implementation according to Library.insert/merge in ML;
 
wenzelm 
parents: 
50507 
diff
changeset
 | 
56  | 
def ++ (other: Results): Results =  | 
| 
 
5b7150395568
tuned implementation according to Library.insert/merge in ML;
 
wenzelm 
parents: 
50507 
diff
changeset
 | 
57  | 
if (this eq other) this  | 
| 
 
5b7150395568
tuned implementation according to Library.insert/merge in ML;
 
wenzelm 
parents: 
50507 
diff
changeset
 | 
58  | 
else if (rep.isEmpty) other  | 
| 
56372
 
fadb0fef09d7
more explicit iterator terminology, in accordance to Scala 2.8 library;
 
wenzelm 
parents: 
56359 
diff
changeset
 | 
59  | 
else (this /: other.iterator)(_ + _)  | 
| 50540 | 60  | 
|
| 51494 | 61  | 
override def hashCode: Int = rep.hashCode  | 
62  | 
override def equals(that: Any): Boolean =  | 
|
63  | 
      that match {
 | 
|
64  | 
case other: Results => rep == other.rep  | 
|
65  | 
case _ => false  | 
|
66  | 
}  | 
|
| 
56372
 
fadb0fef09d7
more explicit iterator terminology, in accordance to Scala 2.8 library;
 
wenzelm 
parents: 
56359 
diff
changeset
 | 
67  | 
    override def toString: String = iterator.mkString("Results(", ", ", ")")
 | 
| 50507 | 68  | 
}  | 
69  | 
||
70  | 
||
| 68101 | 71  | 
/* exports */  | 
72  | 
||
73  | 
object Exports  | 
|
74  | 
  {
 | 
|
75  | 
type Entry = (Long, Export.Entry)  | 
|
76  | 
val empty: Exports = new Exports(SortedMap.empty)  | 
|
77  | 
def merge(args: TraversableOnce[Exports]): Exports = (empty /: args)(_ ++ _)  | 
|
78  | 
}  | 
|
79  | 
||
80  | 
final class Exports private(private val rep: SortedMap[Long, Export.Entry])  | 
|
81  | 
  {
 | 
|
82  | 
def iterator: Iterator[Exports.Entry] = rep.iterator  | 
|
83  | 
||
84  | 
def + (entry: Exports.Entry): Exports =  | 
|
85  | 
if (rep.isDefinedAt(entry._1)) this  | 
|
86  | 
else new Exports(rep + entry)  | 
|
87  | 
||
88  | 
def ++ (other: Exports): Exports =  | 
|
89  | 
if (this eq other) this  | 
|
90  | 
else if (rep.isEmpty) other  | 
|
91  | 
else (this /: other.iterator)(_ + _)  | 
|
92  | 
||
93  | 
override def hashCode: Int = rep.hashCode  | 
|
94  | 
override def equals(that: Any): Boolean =  | 
|
95  | 
      that match {
 | 
|
96  | 
case other: Exports => rep == other.rep  | 
|
97  | 
case _ => false  | 
|
98  | 
}  | 
|
99  | 
    override def toString: String = iterator.mkString("Exports(", ", ", ")")
 | 
|
100  | 
}  | 
|
101  | 
||
102  | 
||
103  | 
/* markups */  | 
|
| 
55649
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
104  | 
|
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
105  | 
object Markup_Index  | 
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
106  | 
  {
 | 
| 56746 | 107  | 
val markup: Markup_Index = Markup_Index(false, Symbol.Text_Chunk.Default)  | 
| 
55649
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
108  | 
}  | 
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
109  | 
|
| 56746 | 110  | 
sealed case class Markup_Index(status: Boolean, chunk_name: Symbol.Text_Chunk.Name)  | 
| 
55649
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
111  | 
|
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
112  | 
object Markups  | 
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
113  | 
  {
 | 
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
114  | 
val empty: Markups = new Markups(Map.empty)  | 
| 65335 | 115  | 
def init(markup: Markup_Tree): Markups = new Markups(Map(Markup_Index.markup -> markup))  | 
116  | 
def merge(args: TraversableOnce[Markups]): Markups = (empty /: args)(_ ++ _)  | 
|
| 
55649
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
117  | 
|
| 65335 | 118  | 
|
119  | 
/* XML data representation */  | 
|
120  | 
||
121  | 
def encode(source_length: Int): XML.Encode.T[Markups] = (markups: Markups) =>  | 
|
122  | 
    {
 | 
|
123  | 
import XML.Encode._  | 
|
124  | 
||
125  | 
val markup_index: T[Markup_Index] = (index: Markup_Index) =>  | 
|
126  | 
pair(bool, Symbol.Text_Chunk.encode_name)(index.status, index.chunk_name)  | 
|
127  | 
||
128  | 
val markup_tree: T[Markup_Tree] =  | 
|
129  | 
_.to_XML(Text.Range(0, source_length), Symbol.spaces(source_length), Markup.Elements.full)  | 
|
130  | 
||
131  | 
list(pair(markup_index, markup_tree))(markups.rep.toList)  | 
|
132  | 
}  | 
|
133  | 
||
134  | 
val decode: XML.Decode.T[Markups] = (body: XML.Body) =>  | 
|
135  | 
    {
 | 
|
136  | 
import XML.Decode._  | 
|
137  | 
||
138  | 
val markup_index: T[Markup_Index] = (body: XML.Body) =>  | 
|
139  | 
      {
 | 
|
140  | 
val (status, chunk_name) = pair(bool, Symbol.Text_Chunk.decode_name)(body)  | 
|
141  | 
Markup_Index(status, chunk_name)  | 
|
142  | 
}  | 
|
143  | 
||
144  | 
(Markups.empty /: list(pair(markup_index, Markup_Tree.from_XML(_)))(body))(_ + _)  | 
|
145  | 
}  | 
|
| 
55649
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
146  | 
}  | 
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
147  | 
|
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
148  | 
final class Markups private(private val rep: Map[Markup_Index, Markup_Tree])  | 
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
149  | 
  {
 | 
| 56489 | 150  | 
def is_empty: Boolean = rep.isEmpty  | 
151  | 
||
| 
55649
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
152  | 
def apply(index: Markup_Index): Markup_Tree =  | 
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
153  | 
rep.getOrElse(index, Markup_Tree.empty)  | 
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
154  | 
|
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
155  | 
def add(index: Markup_Index, markup: Text.Markup): Markups =  | 
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
156  | 
new Markups(rep + (index -> (this(index) + markup)))  | 
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
157  | 
|
| 65335 | 158  | 
def + (entry: (Markup_Index, Markup_Tree)): Markups =  | 
159  | 
    {
 | 
|
160  | 
val (index, tree) = entry  | 
|
161  | 
new Markups(rep + (index -> (this(index).merge(tree, Text.Range.full, Markup.Elements.full))))  | 
|
162  | 
}  | 
|
163  | 
||
164  | 
def ++ (other: Markups): Markups =  | 
|
165  | 
if (this eq other) this  | 
|
166  | 
else if (rep.isEmpty) other  | 
|
167  | 
else (this /: other.rep.iterator)(_ + _)  | 
|
168  | 
||
| 56475 | 169  | 
def redirection_iterator: Iterator[Document_ID.Generic] =  | 
| 56746 | 170  | 
for (Markup_Index(_, Symbol.Text_Chunk.Id(id)) <- rep.keysIterator)  | 
| 56475 | 171  | 
yield id  | 
| 
56470
 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 
wenzelm 
parents: 
56469 
diff
changeset
 | 
172  | 
|
| 56475 | 173  | 
def redirect(other_id: Document_ID.Generic): Markups =  | 
| 56489 | 174  | 
    {
 | 
175  | 
val rep1 =  | 
|
| 
56470
 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 
wenzelm 
parents: 
56469 
diff
changeset
 | 
176  | 
        (for {
 | 
| 56746 | 177  | 
(Markup_Index(status, Symbol.Text_Chunk.Id(id)), markup) <- rep.iterator  | 
| 
56470
 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 
wenzelm 
parents: 
56469 
diff
changeset
 | 
178  | 
if other_id == id  | 
| 56746 | 179  | 
} yield (Markup_Index(status, Symbol.Text_Chunk.Default), markup)).toMap  | 
| 56489 | 180  | 
if (rep1.isEmpty) Markups.empty else new Markups(rep1)  | 
181  | 
}  | 
|
| 
56470
 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 
wenzelm 
parents: 
56469 
diff
changeset
 | 
182  | 
|
| 
55649
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
183  | 
override def hashCode: Int = rep.hashCode  | 
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
184  | 
override def equals(that: Any): Boolean =  | 
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
185  | 
      that match {
 | 
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
186  | 
case other: Markups => rep == other.rep  | 
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
187  | 
case _ => false  | 
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
188  | 
}  | 
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
189  | 
    override def toString: String = rep.iterator.mkString("Markups(", ", ", ")")
 | 
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
190  | 
}  | 
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
191  | 
|
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
192  | 
|
| 50507 | 193  | 
/* state */  | 
| 
50501
 
6f41f1646617
more careful handling of Dialog_Result, with active area and color feedback;
 
wenzelm 
parents: 
50500 
diff
changeset
 | 
194  | 
|
| 
56299
 
8201790fdeb9
more careful treatment of multiple command states (eval + prints): merge content that is actually required;
 
wenzelm 
parents: 
56295 
diff
changeset
 | 
195  | 
object State  | 
| 
 
8201790fdeb9
more careful treatment of multiple command states (eval + prints): merge content that is actually required;
 
wenzelm 
parents: 
56295 
diff
changeset
 | 
196  | 
  {
 | 
| 
67824
 
661cd25304ae
more compact markup tree: output messages are already stored in command results (e.g. relevant for  XML data representation);
 
wenzelm 
parents: 
67446 
diff
changeset
 | 
197  | 
def get_result(states: List[State], serial: Long): Option[XML.Tree] =  | 
| 
 
661cd25304ae
more compact markup tree: output messages are already stored in command results (e.g. relevant for  XML data representation);
 
wenzelm 
parents: 
67446 
diff
changeset
 | 
198  | 
states.find(st => st.results.defined(serial)).map(st => st.results.get(serial).get)  | 
| 
 
661cd25304ae
more compact markup tree: output messages are already stored in command results (e.g. relevant for  XML data representation);
 
wenzelm 
parents: 
67446 
diff
changeset
 | 
199  | 
|
| 
 
661cd25304ae
more compact markup tree: output messages are already stored in command results (e.g. relevant for  XML data representation);
 
wenzelm 
parents: 
67446 
diff
changeset
 | 
200  | 
def get_result_proper(states: List[State], props: Properties.T): Option[Results.Entry] =  | 
| 
 
661cd25304ae
more compact markup tree: output messages are already stored in command results (e.g. relevant for  XML data representation);
 
wenzelm 
parents: 
67446 
diff
changeset
 | 
201  | 
      for {
 | 
| 
 
661cd25304ae
more compact markup tree: output messages are already stored in command results (e.g. relevant for  XML data representation);
 
wenzelm 
parents: 
67446 
diff
changeset
 | 
202  | 
serial <- Markup.Serial.unapply(props)  | 
| 
 
661cd25304ae
more compact markup tree: output messages are already stored in command results (e.g. relevant for  XML data representation);
 
wenzelm 
parents: 
67446 
diff
changeset
 | 
203  | 
tree @ XML.Elem(_, body) <- get_result(states, serial)  | 
| 
 
661cd25304ae
more compact markup tree: output messages are already stored in command results (e.g. relevant for  XML data representation);
 
wenzelm 
parents: 
67446 
diff
changeset
 | 
204  | 
if body.nonEmpty  | 
| 
 
661cd25304ae
more compact markup tree: output messages are already stored in command results (e.g. relevant for  XML data representation);
 
wenzelm 
parents: 
67446 
diff
changeset
 | 
205  | 
} yield (serial -> tree)  | 
| 
 
661cd25304ae
more compact markup tree: output messages are already stored in command results (e.g. relevant for  XML data representation);
 
wenzelm 
parents: 
67446 
diff
changeset
 | 
206  | 
|
| 65335 | 207  | 
def merge_results(states: List[State]): Results =  | 
| 
56299
 
8201790fdeb9
more careful treatment of multiple command states (eval + prints): merge content that is actually required;
 
wenzelm 
parents: 
56295 
diff
changeset
 | 
208  | 
Results.merge(states.map(_.results))  | 
| 
 
8201790fdeb9
more careful treatment of multiple command states (eval + prints): merge content that is actually required;
 
wenzelm 
parents: 
56295 
diff
changeset
 | 
209  | 
|
| 68101 | 210  | 
def merge_exports(states: List[State]): Exports =  | 
211  | 
Exports.merge(states.map(_.exports))  | 
|
212  | 
||
| 65335 | 213  | 
def merge_markups(states: List[State]): Markups =  | 
214  | 
Markups.merge(states.map(_.markups))  | 
|
215  | 
||
| 
56301
 
1da7b4c33db9
more frugal merge of markup trees: filter wrt. subsequent query;
 
wenzelm 
parents: 
56299 
diff
changeset
 | 
216  | 
def merge_markup(states: List[State], index: Markup_Index,  | 
| 56743 | 217  | 
range: Text.Range, elements: Markup.Elements): Markup_Tree =  | 
| 
56301
 
1da7b4c33db9
more frugal merge of markup trees: filter wrt. subsequent query;
 
wenzelm 
parents: 
56299 
diff
changeset
 | 
218  | 
Markup_Tree.merge(states.map(_.markup(index)), range, elements)  | 
| 65335 | 219  | 
|
220  | 
def merge(command: Command, states: List[State]): State =  | 
|
| 68101 | 221  | 
State(command, states.flatMap(_.status), merge_results(states),  | 
222  | 
merge_exports(states), merge_markups(states))  | 
|
| 65335 | 223  | 
|
224  | 
||
225  | 
/* XML data representation */  | 
|
226  | 
||
227  | 
val encode: XML.Encode.T[State] = (st: State) =>  | 
|
228  | 
    {
 | 
|
229  | 
import XML.Encode._  | 
|
230  | 
||
231  | 
val command = st.command  | 
|
232  | 
val blobs_names = command.blobs_names.map(_.node)  | 
|
233  | 
val blobs_index = command.blobs_index  | 
|
234  | 
require(command.blobs_ok)  | 
|
235  | 
||
236  | 
pair(long, pair(string, pair(pair(list(string), int), pair(Command_Span.encode,  | 
|
237  | 
pair(list(Markup.encode), pair(Results.encode, Markups.encode(command.source.length)))))))(  | 
|
238  | 
(command.id, (command.node_name.node, ((blobs_names, blobs_index), (command.span,  | 
|
239  | 
(st.status, (st.results, st.markups)))))))  | 
|
240  | 
}  | 
|
241  | 
||
242  | 
def decode(node_name: String => Document.Node.Name): XML.Decode.T[State] = (body: XML.Body) =>  | 
|
243  | 
    {
 | 
|
244  | 
import XML.Decode._  | 
|
245  | 
val (id, (node, ((blobs_names, blobs_index), (span, (status, (results, markups)))))) =  | 
|
246  | 
pair(long, pair(string, pair(pair(list(string), int), pair(Command_Span.decode,  | 
|
247  | 
pair(list(Markup.decode), pair(Results.decode, Markups.decode))))))(body)  | 
|
248  | 
||
249  | 
val blobs_info: Blobs_Info =  | 
|
250  | 
(blobs_names.map(name => Exn.Res((node_name(name), None)): Blob), blobs_index)  | 
|
251  | 
val command = Command(id, node_name(node), blobs_info, span)  | 
|
| 68101 | 252  | 
State(command, status, results, Exports.empty, markups)  | 
| 65335 | 253  | 
}  | 
| 
56299
 
8201790fdeb9
more careful treatment of multiple command states (eval + prints): merge content that is actually required;
 
wenzelm 
parents: 
56295 
diff
changeset
 | 
254  | 
}  | 
| 
 
8201790fdeb9
more careful treatment of multiple command states (eval + prints): merge content that is actually required;
 
wenzelm 
parents: 
56295 
diff
changeset
 | 
255  | 
|
| 43714 | 256  | 
sealed case class State(  | 
| 
50501
 
6f41f1646617
more careful handling of Dialog_Result, with active area and color feedback;
 
wenzelm 
parents: 
50500 
diff
changeset
 | 
257  | 
command: Command,  | 
| 
 
6f41f1646617
more careful handling of Dialog_Result, with active area and color feedback;
 
wenzelm 
parents: 
50500 
diff
changeset
 | 
258  | 
status: List[Markup] = Nil,  | 
| 50507 | 259  | 
results: Results = Results.empty,  | 
| 68101 | 260  | 
exports: Exports = Exports.empty,  | 
| 
55649
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
261  | 
markups: Markups = Markups.empty)  | 
| 38361 | 262  | 
  {
 | 
| 68323 | 263  | 
def initialized: Boolean = status.exists(markup => markup.name == Markup.INITIALIZED)  | 
| 68335 | 264  | 
def consolidated: Boolean = status.exists(markup => markup.name == Markup.CONSOLIDATED)  | 
| 
66379
 
6392766f3c25
maintain "consolidated" status of theory nodes, which means all evals are finished (but not necessarily prints nor imports);
 
wenzelm 
parents: 
65522 
diff
changeset
 | 
265  | 
|
| 
68381
 
2fd3a6d6ba2e
less wasteful consolidation, based on PIDE front-end state and recent changes;
 
wenzelm 
parents: 
68335 
diff
changeset
 | 
266  | 
lazy val maybe_consolidated: Boolean = Protocol.maybe_consolidated(status)  | 
| 
 
2fd3a6d6ba2e
less wasteful consolidation, based on PIDE front-end state and recent changes;
 
wenzelm 
parents: 
68335 
diff
changeset
 | 
267  | 
|
| 
56395
 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 
wenzelm 
parents: 
56372 
diff
changeset
 | 
268  | 
lazy val protocol_status: Protocol.Status =  | 
| 
 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 
wenzelm 
parents: 
56372 
diff
changeset
 | 
269  | 
    {
 | 
| 
 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 
wenzelm 
parents: 
56372 
diff
changeset
 | 
270  | 
val warnings =  | 
| 
59203
 
5f0bd5afc16d
explicit message channel for "legacy", which is nonetheless a variant of "warning";
 
wenzelm 
parents: 
59072 
diff
changeset
 | 
271  | 
if (results.iterator.exists(p => Protocol.is_warning(p._2) || Protocol.is_legacy(p._2)))  | 
| 
56395
 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 
wenzelm 
parents: 
56372 
diff
changeset
 | 
272  | 
List(Markup(Markup.WARNING, Nil))  | 
| 
 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 
wenzelm 
parents: 
56372 
diff
changeset
 | 
273  | 
else Nil  | 
| 
 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 
wenzelm 
parents: 
56372 
diff
changeset
 | 
274  | 
val errors =  | 
| 
 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 
wenzelm 
parents: 
56372 
diff
changeset
 | 
275  | 
if (results.iterator.exists(p => Protocol.is_error(p._2)))  | 
| 
 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 
wenzelm 
parents: 
56372 
diff
changeset
 | 
276  | 
List(Markup(Markup.ERROR, Nil))  | 
| 
 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 
wenzelm 
parents: 
56372 
diff
changeset
 | 
277  | 
else Nil  | 
| 
 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 
wenzelm 
parents: 
56372 
diff
changeset
 | 
278  | 
Protocol.Status.make((warnings ::: errors ::: status).iterator)  | 
| 
 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 
wenzelm 
parents: 
56372 
diff
changeset
 | 
279  | 
}  | 
| 
55432
 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 
wenzelm 
parents: 
55431 
diff
changeset
 | 
280  | 
|
| 55650 | 281  | 
def markup(index: Markup_Index): Markup_Tree = markups(index)  | 
| 
55432
 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 
wenzelm 
parents: 
55431 
diff
changeset
 | 
282  | 
|
| 56489 | 283  | 
def redirect(other_command: Command): Option[State] =  | 
284  | 
    {
 | 
|
285  | 
val markups1 = markups.redirect(other_command.id)  | 
|
286  | 
if (markups1.is_empty) None  | 
|
| 68101 | 287  | 
else Some(new State(other_command, markups = markups1))  | 
| 56489 | 288  | 
}  | 
| 49614 | 289  | 
|
| 
55649
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
290  | 
private def add_status(st: Markup): State =  | 
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
291  | 
copy(status = st :: status)  | 
| 
55432
 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 
wenzelm 
parents: 
55431 
diff
changeset
 | 
292  | 
|
| 67826 | 293  | 
private def add_result(entry: Results.Entry): State =  | 
294  | 
copy(results = results + entry)  | 
|
295  | 
||
| 68114 | 296  | 
def add_export(entry: Exports.Entry): Option[State] =  | 
297  | 
if (command.node_name.theory == entry._2.theory_name) Some(copy(exports = exports + entry))  | 
|
298  | 
else None  | 
|
| 68101 | 299  | 
|
| 56746 | 300  | 
private def add_markup(  | 
301  | 
status: Boolean, chunk_name: Symbol.Text_Chunk.Name, m: Text.Markup): State =  | 
|
| 
55649
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
302  | 
    {
 | 
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
303  | 
val markups1 =  | 
| 
56395
 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 
wenzelm 
parents: 
56372 
diff
changeset
 | 
304  | 
if (status || Protocol.liberal_status_elements(m.info.name))  | 
| 
56462
 
b64b0cb845fe
more explicit Command.Chunk types, less ooddities;
 
wenzelm 
parents: 
56395 
diff
changeset
 | 
305  | 
markups.add(Markup_Index(true, chunk_name), m)  | 
| 
55649
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
306  | 
else markups  | 
| 
56462
 
b64b0cb845fe
more explicit Command.Chunk types, less ooddities;
 
wenzelm 
parents: 
56395 
diff
changeset
 | 
307  | 
copy(markups = markups1.add(Markup_Index(false, chunk_name), m))  | 
| 
55649
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
308  | 
}  | 
| 38361 | 309  | 
|
| 
56470
 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 
wenzelm 
parents: 
56469 
diff
changeset
 | 
310  | 
def accumulate(  | 
| 
 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 
wenzelm 
parents: 
56469 
diff
changeset
 | 
311  | 
self_id: Document_ID.Generic => Boolean,  | 
| 56746 | 312  | 
other_id: Document_ID.Generic => Option[(Symbol.Text_Chunk.Id, Symbol.Text_Chunk)],  | 
| 67825 | 313  | 
message: XML.Elem,  | 
314  | 
xml_cache: XML.Cache): State =  | 
|
| 38361 | 315  | 
      message match {
 | 
| 
50201
 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 
wenzelm 
parents: 
50163 
diff
changeset
 | 
316  | 
case XML.Elem(Markup(Markup.STATUS, _), msgs) =>  | 
| 38714 | 317  | 
(this /: msgs)((state, msg) =>  | 
318  | 
            msg match {
 | 
|
| 
46152
 
793cecd4ffc0
accumulate status as regular markup for command range;
 
wenzelm 
parents: 
45709 
diff
changeset
 | 
319  | 
case elem @ XML.Elem(markup, Nil) =>  | 
| 
55649
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
320  | 
state.  | 
| 
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
321  | 
add_status(markup).  | 
| 68728 | 322  | 
add_markup(true, Symbol.Text_Chunk.Default, Text.Info(command.core_range, elem))  | 
| 52536 | 323  | 
case _ =>  | 
| 
56782
 
433cf57550fa
more systematic Isabelle output, like in classic Isabelle/ML (without markup);
 
wenzelm 
parents: 
56746 
diff
changeset
 | 
324  | 
                Output.warning("Ignored status message: " + msg)
 | 
| 52536 | 325  | 
state  | 
| 38714 | 326  | 
})  | 
| 
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
 | 
327  | 
|
| 
50201
 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 
wenzelm 
parents: 
50163 
diff
changeset
 | 
328  | 
case XML.Elem(Markup(Markup.REPORT, _), msgs) =>  | 
| 
38572
 
0fe2c01ef7da
Command.State: accumulate markup reports uniformly;
 
wenzelm 
parents: 
38564 
diff
changeset
 | 
329  | 
(this /: msgs)((state, msg) =>  | 
| 
55432
 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 
wenzelm 
parents: 
55431 
diff
changeset
 | 
330  | 
            {
 | 
| 
56782
 
433cf57550fa
more systematic Isabelle output, like in classic Isabelle/ML (without markup);
 
wenzelm 
parents: 
56746 
diff
changeset
 | 
331  | 
              def bad(): Unit = Output.warning("Ignored report message: " + msg)
 | 
| 
55432
 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 
wenzelm 
parents: 
55431 
diff
changeset
 | 
332  | 
|
| 
 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 
wenzelm 
parents: 
55431 
diff
changeset
 | 
333  | 
              msg match {
 | 
| 
57911
 
dcb758188aa6
clarified Position.Identified: do not require range from prover, default to command position;
 
wenzelm 
parents: 
57910 
diff
changeset
 | 
334  | 
case XML.Elem(Markup(name, atts @ Position.Identified(id, chunk_name)), args) =>  | 
| 
56470
 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 
wenzelm 
parents: 
56469 
diff
changeset
 | 
335  | 
|
| 
 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 
wenzelm 
parents: 
56469 
diff
changeset
 | 
336  | 
val target =  | 
| 
 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 
wenzelm 
parents: 
56469 
diff
changeset
 | 
337  | 
if (self_id(id) && command.chunks.isDefinedAt(chunk_name))  | 
| 
 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 
wenzelm 
parents: 
56469 
diff
changeset
 | 
338  | 
Some((chunk_name, command.chunks(chunk_name)))  | 
| 56746 | 339  | 
else if (chunk_name == Symbol.Text_Chunk.Default) other_id(id)  | 
| 
56470
 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 
wenzelm 
parents: 
56469 
diff
changeset
 | 
340  | 
else None  | 
| 
 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 
wenzelm 
parents: 
56469 
diff
changeset
 | 
341  | 
|
| 
57911
 
dcb758188aa6
clarified Position.Identified: do not require range from prover, default to command position;
 
wenzelm 
parents: 
57910 
diff
changeset
 | 
342  | 
                  (target, atts) match {
 | 
| 
 
dcb758188aa6
clarified Position.Identified: do not require range from prover, default to command position;
 
wenzelm 
parents: 
57910 
diff
changeset
 | 
343  | 
case (Some((target_name, target_chunk)), Position.Range(symbol_range)) =>  | 
| 
56470
 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 
wenzelm 
parents: 
56469 
diff
changeset
 | 
344  | 
                      target_chunk.incorporate(symbol_range) match {
 | 
| 
55548
 
a645277885cf
more uniform/robust restriction of reported positions, e.g. relevant for "bad" markup due to unclosed comment in ML file;
 
wenzelm 
parents: 
55434 
diff
changeset
 | 
345  | 
case Some(range) =>  | 
| 
55822
 
ccf2d784be97
incorporate chunk range that is 1 off end-of-input, for improved error positions (NB: command spans are tight, without trailing whitespace);
 
wenzelm 
parents: 
55785 
diff
changeset
 | 
346  | 
val props = Position.purge(atts)  | 
| 67825 | 347  | 
val elem = xml_cache.elem(XML.Elem(Markup(name, props), args))  | 
348  | 
state.add_markup(false, target_name, Text.Info(range, elem))  | 
|
| 
55548
 
a645277885cf
more uniform/robust restriction of reported positions, e.g. relevant for "bad" markup due to unclosed comment in ML file;
 
wenzelm 
parents: 
55434 
diff
changeset
 | 
349  | 
case None => bad(); state  | 
| 
55432
 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 
wenzelm 
parents: 
55431 
diff
changeset
 | 
350  | 
}  | 
| 
57911
 
dcb758188aa6
clarified Position.Identified: do not require range from prover, default to command position;
 
wenzelm 
parents: 
57910 
diff
changeset
 | 
351  | 
case _ =>  | 
| 
56514
 
db929027e701
ignore other_id reports for now (see 8eda56033203): massive amounts of redirections to 'class' etc. makes it difficult to edit main HOL;
 
wenzelm 
parents: 
56489 
diff
changeset
 | 
352  | 
// silently ignore excessive reports  | 
| 56476 | 353  | 
state  | 
| 
55432
 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 
wenzelm 
parents: 
55431 
diff
changeset
 | 
354  | 
}  | 
| 
 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 
wenzelm 
parents: 
55431 
diff
changeset
 | 
355  | 
|
| 
 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 
wenzelm 
parents: 
55431 
diff
changeset
 | 
356  | 
case XML.Elem(Markup(name, atts), args)  | 
| 
 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 
wenzelm 
parents: 
55431 
diff
changeset
 | 
357  | 
                if !atts.exists({ case (a, _) => Markup.POSITION_PROPERTIES(a) }) =>
 | 
| 68728 | 358  | 
val range = command.core_range  | 
| 
55432
 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 
wenzelm 
parents: 
55431 
diff
changeset
 | 
359  | 
val props = Position.purge(atts)  | 
| 67825 | 360  | 
val elem = xml_cache.elem(XML.Elem(Markup(name, props), args))  | 
361  | 
state.add_markup(false, Symbol.Text_Chunk.Default, Text.Info(range, elem))  | 
|
| 
55432
 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 
wenzelm 
parents: 
55431 
diff
changeset
 | 
362  | 
|
| 
56470
 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 
wenzelm 
parents: 
56469 
diff
changeset
 | 
363  | 
case _ => bad(); state  | 
| 
55432
 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 
wenzelm 
parents: 
55431 
diff
changeset
 | 
364  | 
}  | 
| 38361 | 365  | 
})  | 
| 
52930
 
5fab62ae3532
retain original messages properties, e.g. for retrieval via Command.Results;
 
wenzelm 
parents: 
52849 
diff
changeset
 | 
366  | 
case XML.Elem(Markup(name, props), body) =>  | 
| 
 
5fab62ae3532
retain original messages properties, e.g. for retrieval via Command.Results;
 
wenzelm 
parents: 
52849 
diff
changeset
 | 
367  | 
          props match {
 | 
| 
50201
 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 
wenzelm 
parents: 
50163 
diff
changeset
 | 
368  | 
case Markup.Serial(i) =>  | 
| 67825 | 369  | 
val markup_message =  | 
370  | 
xml_cache.elem(XML.Elem(Markup(Markup.message(name), props), body))  | 
|
371  | 
val message_markup =  | 
|
372  | 
xml_cache.elem(XML.elem(Markup(name, props.filter(p => p._1 == Markup.SERIAL))))  | 
|
| 
50163
 
c62ce309dc26
more abstract Sendback operations, with explicit id/exec_id properties;
 
wenzelm 
parents: 
50158 
diff
changeset
 | 
373  | 
|
| 67826 | 374  | 
var st = add_result(i -> markup_message)  | 
| 
55433
 
d2960d67f163
clarified message_positions: cover alt_id as well;
 
wenzelm 
parents: 
55432 
diff
changeset
 | 
375  | 
              if (Protocol.is_inlined(message)) {
 | 
| 
 
d2960d67f163
clarified message_positions: cover alt_id as well;
 
wenzelm 
parents: 
55432 
diff
changeset
 | 
376  | 
                for {
 | 
| 56469 | 377  | 
(chunk_name, chunk) <- command.chunks.iterator  | 
| 59713 | 378  | 
range <- Protocol_Message.positions(  | 
| 59735 | 379  | 
self_id, command.span.position, chunk_name, chunk, message)  | 
| 
67824
 
661cd25304ae
more compact markup tree: output messages are already stored in command results (e.g. relevant for  XML data representation);
 
wenzelm 
parents: 
67446 
diff
changeset
 | 
380  | 
} st = st.add_markup(false, chunk_name, Text.Info(range, message_markup))  | 
| 
55433
 
d2960d67f163
clarified message_positions: cover alt_id as well;
 
wenzelm 
parents: 
55432 
diff
changeset
 | 
381  | 
}  | 
| 
 
d2960d67f163
clarified message_positions: cover alt_id as well;
 
wenzelm 
parents: 
55432 
diff
changeset
 | 
382  | 
st  | 
| 
55432
 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 
wenzelm 
parents: 
55431 
diff
changeset
 | 
383  | 
|
| 52536 | 384  | 
case _ =>  | 
| 
56782
 
433cf57550fa
more systematic Isabelle output, like in classic Isabelle/ML (without markup);
 
wenzelm 
parents: 
56746 
diff
changeset
 | 
385  | 
              Output.warning("Ignored message without serial number: " + message)
 | 
| 52536 | 386  | 
this  | 
| 38872 | 387  | 
}  | 
| 68101 | 388  | 
}  | 
| 38361 | 389  | 
}  | 
| 38367 | 390  | 
|
391  | 
||
| 
55431
 
e0f20a44ff9d
common Command.Chunk for command source and auxiliary files (static Symbol.Index without actual String content);
 
wenzelm 
parents: 
55430 
diff
changeset
 | 
392  | 
|
| 
 
e0f20a44ff9d
common Command.Chunk for command source and auxiliary files (static Symbol.Index without actual String content);
 
wenzelm 
parents: 
55430 
diff
changeset
 | 
393  | 
/** static content **/  | 
| 
 
e0f20a44ff9d
common Command.Chunk for command source and auxiliary files (static Symbol.Index without actual String content);
 
wenzelm 
parents: 
55430 
diff
changeset
 | 
394  | 
|
| 45644 | 395  | 
/* make commands */  | 
396  | 
||
| 
55648
 
38f264741609
tuned signature -- avoid obscure default arguments;
 
wenzelm 
parents: 
55622 
diff
changeset
 | 
397  | 
def apply(  | 
| 
 
38f264741609
tuned signature -- avoid obscure default arguments;
 
wenzelm 
parents: 
55622 
diff
changeset
 | 
398  | 
id: Document_ID.Command,  | 
| 
 
38f264741609
tuned signature -- avoid obscure default arguments;
 
wenzelm 
parents: 
55622 
diff
changeset
 | 
399  | 
node_name: Document.Node.Name,  | 
| 
59702
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
400  | 
blobs_info: Blobs_Info,  | 
| 
57905
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57904 
diff
changeset
 | 
401  | 
span: Command_Span.Span): Command =  | 
| 
55648
 
38f264741609
tuned signature -- avoid obscure default arguments;
 
wenzelm 
parents: 
55622 
diff
changeset
 | 
402  | 
  {
 | 
| 
57901
 
e1abca2527da
more explicit type Span in Scala, according to ML version;
 
wenzelm 
parents: 
57842 
diff
changeset
 | 
403  | 
val (source, span1) = span.compact_source  | 
| 
59702
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
404  | 
new Command(id, node_name, blobs_info, span1, source, Results.empty, Markup_Tree.empty)  | 
| 
55648
 
38f264741609
tuned signature -- avoid obscure default arguments;
 
wenzelm 
parents: 
55622 
diff
changeset
 | 
405  | 
}  | 
| 49414 | 406  | 
|
| 
57901
 
e1abca2527da
more explicit type Span in Scala, according to ML version;
 
wenzelm 
parents: 
57842 
diff
changeset
 | 
407  | 
val empty: Command =  | 
| 
59702
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
408  | 
Command(Document_ID.none, Document.Node.Name.empty, no_blobs, Command_Span.empty)  | 
| 
55648
 
38f264741609
tuned signature -- avoid obscure default arguments;
 
wenzelm 
parents: 
55622 
diff
changeset
 | 
409  | 
|
| 
 
38f264741609
tuned signature -- avoid obscure default arguments;
 
wenzelm 
parents: 
55622 
diff
changeset
 | 
410  | 
def unparsed(  | 
| 
 
38f264741609
tuned signature -- avoid obscure default arguments;
 
wenzelm 
parents: 
55622 
diff
changeset
 | 
411  | 
id: Document_ID.Command,  | 
| 
 
38f264741609
tuned signature -- avoid obscure default arguments;
 
wenzelm 
parents: 
55622 
diff
changeset
 | 
412  | 
source: String,  | 
| 
 
38f264741609
tuned signature -- avoid obscure default arguments;
 
wenzelm 
parents: 
55622 
diff
changeset
 | 
413  | 
results: Results,  | 
| 
 
38f264741609
tuned signature -- avoid obscure default arguments;
 
wenzelm 
parents: 
55622 
diff
changeset
 | 
414  | 
markup: Markup_Tree): Command =  | 
| 
 
38f264741609
tuned signature -- avoid obscure default arguments;
 
wenzelm 
parents: 
55622 
diff
changeset
 | 
415  | 
  {
 | 
| 
57905
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57904 
diff
changeset
 | 
416  | 
val (source1, span1) = Command_Span.unparsed(source).compact_source  | 
| 
59702
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
417  | 
new Command(id, Document.Node.Name.empty, no_blobs, span1, source1, results, markup)  | 
| 
55648
 
38f264741609
tuned signature -- avoid obscure default arguments;
 
wenzelm 
parents: 
55622 
diff
changeset
 | 
418  | 
}  | 
| 49414 | 419  | 
|
| 65341 | 420  | 
def text(source: String): Command =  | 
| 
52530
 
99dd8b4ef3fe
explicit module Document_ID as source of globally unique identifiers across ML/Scala;
 
wenzelm 
parents: 
52527 
diff
changeset
 | 
421  | 
unparsed(Document_ID.none, source, Results.empty, Markup_Tree.empty)  | 
| 44384 | 422  | 
|
| 
52530
 
99dd8b4ef3fe
explicit module Document_ID as source of globally unique identifiers across ML/Scala;
 
wenzelm 
parents: 
52527 
diff
changeset
 | 
423  | 
def rich_text(id: Document_ID.Command, results: Results, body: XML.Body): Command =  | 
| 49414 | 424  | 
  {
 | 
| 49466 | 425  | 
val text = XML.content(body)  | 
426  | 
val markup = Markup_Tree.from_XML(body)  | 
|
| 
50501
 
6f41f1646617
more careful handling of Dialog_Result, with active area and color feedback;
 
wenzelm 
parents: 
50500 
diff
changeset
 | 
427  | 
unparsed(id, text, results, markup)  | 
| 49414 | 428  | 
}  | 
| 
49359
 
c1262d7389fb
refined output panel: more value-oriented approach to update and caret focus;
 
wenzelm 
parents: 
49037 
diff
changeset
 | 
429  | 
|
| 44384 | 430  | 
|
431  | 
/* perspective */  | 
|
432  | 
||
| 44474 | 433  | 
object Perspective  | 
434  | 
  {
 | 
|
435  | 
val empty: Perspective = Perspective(Nil)  | 
|
436  | 
}  | 
|
| 
44385
 
e7fdb008aa7d
propagate editor perspective through document model;
 
wenzelm 
parents: 
44384 
diff
changeset
 | 
437  | 
|
| 44474 | 438  | 
sealed case class Perspective(commands: List[Command]) // visible commands in canonical order  | 
| 
44385
 
e7fdb008aa7d
propagate editor perspective through document model;
 
wenzelm 
parents: 
44384 
diff
changeset
 | 
439  | 
  {
 | 
| 
57615
 
df1b3452d71c
more explicit discrimination of empty nodes -- suppress from Theories panel;
 
wenzelm 
parents: 
56782 
diff
changeset
 | 
440  | 
def is_empty: Boolean = commands.isEmpty  | 
| 
 
df1b3452d71c
more explicit discrimination of empty nodes -- suppress from Theories panel;
 
wenzelm 
parents: 
56782 
diff
changeset
 | 
441  | 
|
| 44474 | 442  | 
def same(that: Perspective): Boolean =  | 
443  | 
    {
 | 
|
444  | 
val cmds1 = this.commands  | 
|
445  | 
val cmds2 = that.commands  | 
|
| 
48754
 
c2c1e5944536
clarified undefined, unparsed, unfinished command spans;
 
wenzelm 
parents: 
48745 
diff
changeset
 | 
446  | 
require(!cmds1.exists(_.is_undefined))  | 
| 
 
c2c1e5944536
clarified undefined, unparsed, unfinished command spans;
 
wenzelm 
parents: 
48745 
diff
changeset
 | 
447  | 
require(!cmds2.exists(_.is_undefined))  | 
| 44474 | 448  | 
cmds1.length == cmds2.length &&  | 
449  | 
        (cmds1.iterator zip cmds2.iterator).forall({ case (c1, c2) => c1.id == c2.id })
 | 
|
450  | 
}  | 
|
| 
44385
 
e7fdb008aa7d
propagate editor perspective through document model;
 
wenzelm 
parents: 
44384 
diff
changeset
 | 
451  | 
}  | 
| 
59689
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
452  | 
|
| 
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
453  | 
|
| 
59702
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
454  | 
/* blobs: inlined errors and auxiliary files */  | 
| 
59689
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
455  | 
|
| 
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
456  | 
private def clean_tokens(tokens: List[Token]): List[(Token, Int)] =  | 
| 
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
457  | 
  {
 | 
| 
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
458  | 
def clean(toks: List[(Token, Int)]): List[(Token, Int)] =  | 
| 
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
459  | 
      toks match {
 | 
| 
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
460  | 
case (t1, i1) :: (t2, i2) :: rest =>  | 
| 67446 | 461  | 
if (t1.is_keyword && t1.source == "%" && t2.is_name) clean(rest)  | 
| 
59689
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
462  | 
else (t1, i1) :: clean((t2, i2) :: rest)  | 
| 
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
463  | 
case _ => toks  | 
| 
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
464  | 
}  | 
| 
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
465  | 
    clean(tokens.zipWithIndex.filter({ case (t, _) => t.is_proper }))
 | 
| 
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
466  | 
}  | 
| 
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
467  | 
|
| 
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
468  | 
private def find_file(tokens: List[(Token, Int)]): Option[(String, Int)] =  | 
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
469  | 
    if (tokens.exists({ case (t, _) => t.is_command })) {
 | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
470  | 
      tokens.dropWhile({ case (t, _) => !t.is_command }).
 | 
| 
64471
 
c40c2975fb02
more uniform path syntax, as in ML (see 5a7c919a4ada);
 
wenzelm 
parents: 
63584 
diff
changeset
 | 
471  | 
        collectFirst({ case (t, i) if t.is_embedded => (t.content, i) })
 | 
| 
59689
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
472  | 
}  | 
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
473  | 
else None  | 
| 
59689
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
474  | 
|
| 
63584
 
68751fe1c036
tuned signature -- prover-independence is presently theoretical;
 
wenzelm 
parents: 
62969 
diff
changeset
 | 
475  | 
def span_files(syntax: Outer_Syntax, span: Command_Span.Span): (List[String], Int) =  | 
| 59735 | 476  | 
    syntax.load_command(span.name) match {
 | 
477  | 
case Some(exts) =>  | 
|
478  | 
        find_file(clean_tokens(span.content)) match {
 | 
|
479  | 
case Some((file, i)) =>  | 
|
480  | 
if (exts.isEmpty) (List(file), i)  | 
|
481  | 
else (exts.map(ext => file + "." + ext), i)  | 
|
| 
59689
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
482  | 
case None => (Nil, -1)  | 
| 
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
483  | 
}  | 
| 59735 | 484  | 
case None => (Nil, -1)  | 
| 
59689
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
485  | 
}  | 
| 
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
486  | 
|
| 
59702
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
487  | 
def blobs_info(  | 
| 59699 | 488  | 
resources: Resources,  | 
| 
63584
 
68751fe1c036
tuned signature -- prover-independence is presently theoretical;
 
wenzelm 
parents: 
62969 
diff
changeset
 | 
489  | 
syntax: Outer_Syntax,  | 
| 59699 | 490  | 
get_blob: Document.Node.Name => Option[Document.Blob],  | 
| 
59702
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
491  | 
can_import: Document.Node.Name => Boolean,  | 
| 59699 | 492  | 
node_name: Document.Node.Name,  | 
| 
59702
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
493  | 
span: Command_Span.Span): Blobs_Info =  | 
| 
59689
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
494  | 
  {
 | 
| 59735 | 495  | 
    span.name match {
 | 
| 
59702
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
496  | 
// inlined errors  | 
| 59735 | 497  | 
case Thy_Header.THEORY =>  | 
| 64825 | 498  | 
val reader = Scan.char_reader(Token.implode(span.content))  | 
| 66768 | 499  | 
val imports = resources.check_thy_reader(node_name, reader).imports  | 
500  | 
val raw_imports =  | 
|
501  | 
          try {
 | 
|
502  | 
val imports1 = Thy_Header.read(reader, Token.Pos.none).imports  | 
|
503  | 
            if (imports.length == imports1.length) imports1.map(_._1) else error("")
 | 
|
504  | 
}  | 
|
505  | 
          catch { case exn: Throwable => List.fill(imports.length)("") }
 | 
|
506  | 
||
| 
59715
 
4f0d0e4ad68d
avoid duplicate header errors, more precise positions;
 
wenzelm 
parents: 
59713 
diff
changeset
 | 
507  | 
val errors =  | 
| 66768 | 508  | 
          for { ((import_name, pos), s) <- imports zip raw_imports if !can_import(import_name) }
 | 
509  | 
          yield {
 | 
|
510  | 
val completion =  | 
|
511  | 
              if (Thy_Header.is_base_name(s)) {
 | 
|
512  | 
val completed = Completion.completed(import_name.theory_base_name)  | 
|
| 66966 | 513  | 
val qualifier = resources.session_base.theory_qualifier(node_name)  | 
| 66768 | 514  | 
val dir = node_name.master_dir  | 
515  | 
                for {
 | 
|
| 
68306
 
d575281e18d0
clarified signature: Known.theories retains Document.Node.Entry (with header);
 
wenzelm 
parents: 
68114 
diff
changeset
 | 
516  | 
known_name <- resources.session_base.known.theory_names  | 
| 66768 | 517  | 
if completed(known_name.theory_base_name)  | 
| 66966 | 518  | 
}  | 
519  | 
                yield {
 | 
|
520  | 
resources.standard_import(  | 
|
521  | 
resources.session_base, qualifier, dir, known_name.theory)  | 
|
522  | 
}  | 
|
| 66768 | 523  | 
}.sorted  | 
524  | 
else Nil  | 
|
| 
59715
 
4f0d0e4ad68d
avoid duplicate header errors, more precise positions;
 
wenzelm 
parents: 
59713 
diff
changeset
 | 
525  | 
val msg =  | 
| 
 
4f0d0e4ad68d
avoid duplicate header errors, more precise positions;
 
wenzelm 
parents: 
59713 
diff
changeset
 | 
526  | 
"Bad theory import " +  | 
| 66768 | 527  | 
Markup.Path(import_name.node).markup(quote(import_name.toString)) +  | 
528  | 
Position.here(pos) + Completion.report_theories(pos, completion)  | 
|
| 
59715
 
4f0d0e4ad68d
avoid duplicate header errors, more precise positions;
 
wenzelm 
parents: 
59713 
diff
changeset
 | 
529  | 
Exn.Exn(ERROR(msg)): Command.Blob  | 
| 
59708
 
aed304412e43
more markup, which helps to create missing imports;
 
wenzelm 
parents: 
59706 
diff
changeset
 | 
530  | 
}  | 
| 
59715
 
4f0d0e4ad68d
avoid duplicate header errors, more precise positions;
 
wenzelm 
parents: 
59713 
diff
changeset
 | 
531  | 
(errors, -1)  | 
| 
59702
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
532  | 
|
| 
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
533  | 
// auxiliary files  | 
| 
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
534  | 
case _ =>  | 
| 
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
535  | 
val (files, index) = span_files(syntax, span)  | 
| 
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
536  | 
val blobs =  | 
| 
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
537  | 
files.map(file =>  | 
| 
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
538  | 
            (Exn.capture {
 | 
| 65488 | 539  | 
val name = Document.Node.Name(resources.append(node_name, Path.explode(file)))  | 
| 
59702
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
540  | 
val blob = get_blob(name).map(blob => ((blob.bytes.sha1_digest, blob.chunk)))  | 
| 
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
541  | 
(name, blob)  | 
| 
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
542  | 
}).user_error)  | 
| 
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
543  | 
(blobs, index)  | 
| 
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
544  | 
}  | 
| 
59689
 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 
wenzelm 
parents: 
59684 
diff
changeset
 | 
545  | 
}  | 
| 
34318
 
c13e168a8ae6
original sources from Johannes Hölzl a48e0c6ab1aea77c52d596f7efc007a543d3d10c with minor modifications of directory layout;
 
wenzelm 
parents:  
diff
changeset
 | 
546  | 
}  | 
| 
 
c13e168a8ae6
original sources from Johannes Hölzl a48e0c6ab1aea77c52d596f7efc007a543d3d10c with minor modifications of directory layout;
 
wenzelm 
parents:  
diff
changeset
 | 
547  | 
|
| 38361 | 548  | 
|
| 46712 | 549  | 
final class Command private(  | 
| 
52530
 
99dd8b4ef3fe
explicit module Document_ID as source of globally unique identifiers across ML/Scala;
 
wenzelm 
parents: 
52527 
diff
changeset
 | 
550  | 
val id: Document_ID.Command,  | 
| 44615 | 551  | 
val node_name: Document.Node.Name,  | 
| 
59702
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
552  | 
val blobs_info: Command.Blobs_Info,  | 
| 
57905
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57904 
diff
changeset
 | 
553  | 
val span: Command_Span.Span,  | 
| 49414 | 554  | 
val source: String,  | 
| 
50501
 
6f41f1646617
more careful handling of Dialog_Result, with active area and color feedback;
 
wenzelm 
parents: 
50500 
diff
changeset
 | 
555  | 
val init_results: Command.Results,  | 
| 49414 | 556  | 
val init_markup: Markup_Tree)  | 
| 34451 | 557  | 
{
 | 
| 57912 | 558  | 
override def toString: String = id + "/" + span.kind.toString  | 
| 34495 | 559  | 
|
| 57910 | 560  | 
|
561  | 
/* classification */  | 
|
562  | 
||
| 
57905
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57904 
diff
changeset
 | 
563  | 
def is_proper: Boolean = span.kind.isInstanceOf[Command_Span.Command_Span]  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57904 
diff
changeset
 | 
564  | 
def is_ignored: Boolean = span.kind == Command_Span.Ignored_Span  | 
| 57904 | 565  | 
|
566  | 
def is_undefined: Boolean = id == Document_ID.none  | 
|
567  | 
val is_unparsed: Boolean = span.content.exists(_.is_unparsed)  | 
|
568  | 
val is_unfinished: Boolean = span.content.exists(_.is_unfinished)  | 
|
569  | 
||
| 68323 | 570  | 
def potentially_initialized: Boolean = span.name == Thy_Header.THEORY  | 
571  | 
||
| 34859 | 572  | 
|
| 
54519
 
5fed81762406
maintain blobs within document state: digest + text in ML, digest-only in Scala;
 
wenzelm 
parents: 
54517 
diff
changeset
 | 
573  | 
/* blobs */  | 
| 
 
5fed81762406
maintain blobs within document state: digest + text in ML, digest-only in Scala;
 
wenzelm 
parents: 
54517 
diff
changeset
 | 
574  | 
|
| 
59702
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
575  | 
def blobs: List[Command.Blob] = blobs_info._1  | 
| 
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
576  | 
def blobs_index: Int = blobs_info._2  | 
| 
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59699 
diff
changeset
 | 
577  | 
|
| 65335 | 578  | 
def blobs_ok: Boolean =  | 
579  | 
    blobs.forall({ case Exn.Res(_) => true case _ => false })
 | 
|
580  | 
||
| 
54530
 
2c1440f70028
ranges of thy_load commands count as visible within perspective;
 
wenzelm 
parents: 
54524 
diff
changeset
 | 
581  | 
def blobs_names: List[Document.Node.Name] =  | 
| 
 
2c1440f70028
ranges of thy_load commands count as visible within perspective;
 
wenzelm 
parents: 
54524 
diff
changeset
 | 
582  | 
for (Exn.Res((name, _)) <- blobs) yield name  | 
| 
 
2c1440f70028
ranges of thy_load commands count as visible within perspective;
 
wenzelm 
parents: 
54524 
diff
changeset
 | 
583  | 
|
| 
60916
 
a6e2a667b0a8
resolve undefined blobs by default, e.g. relevant for ML debugger to avoid reset of breakpoints after reload;
 
wenzelm 
parents: 
60215 
diff
changeset
 | 
584  | 
def blobs_undefined: List[Document.Node.Name] =  | 
| 
 
a6e2a667b0a8
resolve undefined blobs by default, e.g. relevant for ML debugger to avoid reset of breakpoints after reload;
 
wenzelm 
parents: 
60215 
diff
changeset
 | 
585  | 
for (Exn.Res((name, None)) <- blobs) yield name  | 
| 
 
a6e2a667b0a8
resolve undefined blobs by default, e.g. relevant for ML debugger to avoid reset of breakpoints after reload;
 
wenzelm 
parents: 
60215 
diff
changeset
 | 
586  | 
|
| 
57842
 
8e4ae2db1849
more direct access to persistent blobs (see also 8953d4cc060a), avoiding fragile digest lookup from later version (which might have removed unused blobs already);
 
wenzelm 
parents: 
57615 
diff
changeset
 | 
587  | 
def blobs_defined: List[(Document.Node.Name, SHA1.Digest)] =  | 
| 
 
8e4ae2db1849
more direct access to persistent blobs (see also 8953d4cc060a), avoiding fragile digest lookup from later version (which might have removed unused blobs already);
 
wenzelm 
parents: 
57615 
diff
changeset
 | 
588  | 
for (Exn.Res((name, Some((digest, _)))) <- blobs) yield (name, digest)  | 
| 
 
8e4ae2db1849
more direct access to persistent blobs (see also 8953d4cc060a), avoiding fragile digest lookup from later version (which might have removed unused blobs already);
 
wenzelm 
parents: 
57615 
diff
changeset
 | 
589  | 
|
| 
 
8e4ae2db1849
more direct access to persistent blobs (see also 8953d4cc060a), avoiding fragile digest lookup from later version (which might have removed unused blobs already);
 
wenzelm 
parents: 
57615 
diff
changeset
 | 
590  | 
def blobs_changed(doc_blobs: Document.Blobs): Boolean =  | 
| 
 
8e4ae2db1849
more direct access to persistent blobs (see also 8953d4cc060a), avoiding fragile digest lookup from later version (which might have removed unused blobs already);
 
wenzelm 
parents: 
57615 
diff
changeset
 | 
591  | 
    blobs.exists({ case Exn.Res((name, _)) => doc_blobs.changed(name) case _ => false })
 | 
| 
54519
 
5fed81762406
maintain blobs within document state: digest + text in ML, digest-only in Scala;
 
wenzelm 
parents: 
54517 
diff
changeset
 | 
592  | 
|
| 
55432
 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 
wenzelm 
parents: 
55431 
diff
changeset
 | 
593  | 
|
| 
56462
 
b64b0cb845fe
more explicit Command.Chunk types, less ooddities;
 
wenzelm 
parents: 
56395 
diff
changeset
 | 
594  | 
/* source chunks */  | 
| 
55431
 
e0f20a44ff9d
common Command.Chunk for command source and auxiliary files (static Symbol.Index without actual String content);
 
wenzelm 
parents: 
55430 
diff
changeset
 | 
595  | 
|
| 56746 | 596  | 
val chunk: Symbol.Text_Chunk = Symbol.Text_Chunk(source)  | 
| 56473 | 597  | 
|
| 56746 | 598  | 
val chunks: Map[Symbol.Text_Chunk.Name, Symbol.Text_Chunk] =  | 
599  | 
((Symbol.Text_Chunk.Default -> chunk) ::  | 
|
| 56473 | 600  | 
(for (Exn.Res((name, Some((_, file)))) <- blobs)  | 
| 60215 | 601  | 
yield Symbol.Text_Chunk.File(name.node) -> file)).toMap  | 
| 56473 | 602  | 
|
| 46813 | 603  | 
def length: Int = source.length  | 
| 56473 | 604  | 
def range: Text.Range = chunk.range  | 
| 46813 | 605  | 
|
| 68728 | 606  | 
val core_range: Text.Range =  | 
| 
57901
 
e1abca2527da
more explicit type Span in Scala, according to ML version;
 
wenzelm 
parents: 
57842 
diff
changeset
 | 
607  | 
Text.Range(0,  | 
| 
 
e1abca2527da
more explicit type Span in Scala, according to ML version;
 
wenzelm 
parents: 
57842 
diff
changeset
 | 
608  | 
(length /: span.content.reverse.iterator.takeWhile(_.is_improper))(_ - _.source.length))  | 
| 46813 | 609  | 
|
| 65522 | 610  | 
def source(range: Text.Range): String = range.substring(source)  | 
| 
38572
 
0fe2c01ef7da
Command.State: accumulate markup reports uniformly;
 
wenzelm 
parents: 
38564 
diff
changeset
 | 
611  | 
|
| 
38370
 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 
wenzelm 
parents: 
38367 
diff
changeset
 | 
612  | 
|
| 
 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 
wenzelm 
parents: 
38367 
diff
changeset
 | 
613  | 
/* accumulated results */  | 
| 
 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 
wenzelm 
parents: 
38367 
diff
changeset
 | 
614  | 
|
| 
50501
 
6f41f1646617
more careful handling of Dialog_Result, with active area and color feedback;
 
wenzelm 
parents: 
50500 
diff
changeset
 | 
615  | 
val init_state: Command.State =  | 
| 
55649
 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 
wenzelm 
parents: 
55648 
diff
changeset
 | 
616  | 
Command.State(this, results = init_results, markups = Command.Markups.init(init_markup))  | 
| 
52527
 
dbac84eab3bc
separate exec_id assignment for Command.print states, without affecting result of eval;
 
wenzelm 
parents: 
52524 
diff
changeset
 | 
617  | 
|
| 
 
dbac84eab3bc
separate exec_id assignment for Command.print states, without affecting result of eval;
 
wenzelm 
parents: 
52524 
diff
changeset
 | 
618  | 
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
 | 
619  | 
}  |