| author | wenzelm | 
| Thu, 30 Jun 2011 14:51:32 +0200 | |
| changeset 43605 | 4f119a9ed37c | 
| parent 43593 | 11140987d415 | 
| child 43673 | 29eb1cd29961 | 
| permissions | -rw-r--r-- | 
| 27958 | 1 | /* Title: Pure/General/markup.scala | 
| 2 | Author: Makarius | |
| 3 | ||
| 4 | Common markup elements. | |
| 5 | */ | |
| 6 | ||
| 7 | package isabelle | |
| 8 | ||
| 27970 | 9 | |
| 32450 | 10 | object Markup | 
| 11 | {
 | |
| 38722 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 12 | /* plain values */ | 
| 38414 
49f1f657adc2
more basic Markup.parse_int/print_int (using signed_string_of_int) (ML);
 wenzelm parents: 
38355diff
changeset | 13 | |
| 40392 | 14 | object Int | 
| 15 |   {
 | |
| 16 | def apply(x: scala.Int): String = x.toString | |
| 38414 
49f1f657adc2
more basic Markup.parse_int/print_int (using signed_string_of_int) (ML);
 wenzelm parents: 
38355diff
changeset | 17 | def unapply(s: String): Option[scala.Int] = | 
| 
49f1f657adc2
more basic Markup.parse_int/print_int (using signed_string_of_int) (ML);
 wenzelm parents: 
38355diff
changeset | 18 |       try { Some(Integer.parseInt(s)) }
 | 
| 
49f1f657adc2
more basic Markup.parse_int/print_int (using signed_string_of_int) (ML);
 wenzelm parents: 
38355diff
changeset | 19 |       catch { case _: NumberFormatException => None }
 | 
| 
49f1f657adc2
more basic Markup.parse_int/print_int (using signed_string_of_int) (ML);
 wenzelm parents: 
38355diff
changeset | 20 | } | 
| 
49f1f657adc2
more basic Markup.parse_int/print_int (using signed_string_of_int) (ML);
 wenzelm parents: 
38355diff
changeset | 21 | |
| 40392 | 22 | object Long | 
| 23 |   {
 | |
| 24 | def apply(x: scala.Long): String = x.toString | |
| 38414 
49f1f657adc2
more basic Markup.parse_int/print_int (using signed_string_of_int) (ML);
 wenzelm parents: 
38355diff
changeset | 25 | def unapply(s: String): Option[scala.Long] = | 
| 
49f1f657adc2
more basic Markup.parse_int/print_int (using signed_string_of_int) (ML);
 wenzelm parents: 
38355diff
changeset | 26 |       try { Some(java.lang.Long.parseLong(s)) }
 | 
| 
49f1f657adc2
more basic Markup.parse_int/print_int (using signed_string_of_int) (ML);
 wenzelm parents: 
38355diff
changeset | 27 |       catch { case _: NumberFormatException => None }
 | 
| 
49f1f657adc2
more basic Markup.parse_int/print_int (using signed_string_of_int) (ML);
 wenzelm parents: 
38355diff
changeset | 28 | } | 
| 
49f1f657adc2
more basic Markup.parse_int/print_int (using signed_string_of_int) (ML);
 wenzelm parents: 
38355diff
changeset | 29 | |
| 40392 | 30 | object Double | 
| 31 |   {
 | |
| 32 | def apply(x: scala.Double): String = x.toString | |
| 33 | def unapply(s: String): Option[scala.Double] = | |
| 34 |       try { Some(java.lang.Double.parseDouble(s)) }
 | |
| 35 |       catch { case _: NumberFormatException => None }
 | |
| 36 | } | |
| 37 | ||
| 38414 
49f1f657adc2
more basic Markup.parse_int/print_int (using signed_string_of_int) (ML);
 wenzelm parents: 
38355diff
changeset | 38 | |
| 38722 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 39 | /* named properties */ | 
| 36683 | 40 | |
| 38722 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 41 | class Property(val name: String) | 
| 38355 
8cb265fb12fe
represent document ids by (long) int, to benefit from the somewhat faster Inttab in ML (LinearSet in Scala is invariably indexed by native object ids);
 wenzelm parents: 
38259diff
changeset | 42 |   {
 | 
| 38722 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 43 | def apply(value: String): List[(String, String)] = List((name, value)) | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 44 | def unapply(props: List[(String, String)]): Option[String] = | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 45 | props.find(_._1 == name).map(_._2) | 
| 38355 
8cb265fb12fe
represent document ids by (long) int, to benefit from the somewhat faster Inttab in ML (LinearSet in Scala is invariably indexed by native object ids);
 wenzelm parents: 
38259diff
changeset | 46 | } | 
| 
8cb265fb12fe
represent document ids by (long) int, to benefit from the somewhat faster Inttab in ML (LinearSet in Scala is invariably indexed by native object ids);
 wenzelm parents: 
38259diff
changeset | 47 | |
| 38722 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 48 | class Int_Property(name: String) | 
| 36683 | 49 |   {
 | 
| 38722 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 50 | def apply(value: scala.Int): List[(String, String)] = List((name, Int(value))) | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 51 | def unapply(props: List[(String, String)]): Option[Int] = | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 52 |       props.find(_._1 == name) match {
 | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 53 | case None => None | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 54 | case Some((_, value)) => Int.unapply(value) | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 55 | } | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 56 | } | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 57 | |
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 58 | class Long_Property(name: String) | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 59 |   {
 | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 60 | def apply(value: scala.Long): List[(String, String)] = List((name, Long(value))) | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 61 | def unapply(props: List[(String, String)]): Option[Long] = | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 62 |       props.find(_._1 == name) match {
 | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 63 | case None => None | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 64 | case Some((_, value)) => Long.unapply(value) | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38721diff
changeset | 65 | } | 
| 36683 | 66 | } | 
| 67 | ||
| 40392 | 68 | class Double_Property(name: String) | 
| 69 |   {
 | |
| 70 | def apply(value: scala.Double): List[(String, String)] = List((name, Double(value))) | |
| 71 | def unapply(props: List[(String, String)]): Option[Double] = | |
| 72 |       props.find(_._1 == name) match {
 | |
| 73 | case None => None | |
| 74 | case Some((_, value)) => Double.unapply(value) | |
| 75 | } | |
| 76 | } | |
| 77 | ||
| 36683 | 78 | |
| 38474 
e498dc2eb576
uniform Markup.empty/Markup.Empty in ML and Scala;
 wenzelm parents: 
38429diff
changeset | 79 | /* empty */ | 
| 
e498dc2eb576
uniform Markup.empty/Markup.Empty in ML and Scala;
 wenzelm parents: 
38429diff
changeset | 80 | |
| 
e498dc2eb576
uniform Markup.empty/Markup.Empty in ML and Scala;
 wenzelm parents: 
38429diff
changeset | 81 |   val Empty = Markup("", Nil)
 | 
| 
e498dc2eb576
uniform Markup.empty/Markup.Empty in ML and Scala;
 wenzelm parents: 
38429diff
changeset | 82 | |
| 
e498dc2eb576
uniform Markup.empty/Markup.Empty in ML and Scala;
 wenzelm parents: 
38429diff
changeset | 83 | |
| 38721 
ca8b14fa0d0d
added some proof state markup, notably number of subgoals (e.g. for indentation);
 wenzelm parents: 
38474diff
changeset | 84 | /* misc properties */ | 
| 29184 | 85 | |
| 86 | val NAME = "name" | |
| 42136 | 87 | val Name = new Property(NAME) | 
| 88 | ||
| 29184 | 89 | val KIND = "kind" | 
| 42136 | 90 | val Kind = new Property(KIND) | 
| 29184 | 91 | |
| 92 | ||
| 33088 
757d7787b10c
markup for formal entities, with "def" or "ref" occurrences;
 wenzelm parents: 
32450diff
changeset | 93 | /* formal entities */ | 
| 
757d7787b10c
markup for formal entities, with "def" or "ref" occurrences;
 wenzelm parents: 
32450diff
changeset | 94 | |
| 38887 
1261481ef5e5
Command.State: add reported positions to markup tree, according main message position or Markup.binding/entity/report occurrences in body;
 wenzelm parents: 
38872diff
changeset | 95 | val BINDING = "binding" | 
| 33088 
757d7787b10c
markup for formal entities, with "def" or "ref" occurrences;
 wenzelm parents: 
32450diff
changeset | 96 | val ENTITY = "entity" | 
| 
757d7787b10c
markup for formal entities, with "def" or "ref" occurrences;
 wenzelm parents: 
32450diff
changeset | 97 | val DEF = "def" | 
| 
757d7787b10c
markup for formal entities, with "def" or "ref" occurrences;
 wenzelm parents: 
32450diff
changeset | 98 | val REF = "ref" | 
| 
757d7787b10c
markup for formal entities, with "def" or "ref" occurrences;
 wenzelm parents: 
32450diff
changeset | 99 | |
| 42202 | 100 | object Entity | 
| 101 |   {
 | |
| 102 | def unapply(markup: Markup): Option[(String, String)] = | |
| 103 |       markup match {
 | |
| 104 | case Markup(ENTITY, props @ Kind(kind)) => | |
| 105 |           props match {
 | |
| 106 | case Name(name) => Some(kind, name) | |
| 107 | case _ => None | |
| 108 | } | |
| 109 | case _ => None | |
| 110 | } | |
| 111 | } | |
| 112 | ||
| 33088 
757d7787b10c
markup for formal entities, with "def" or "ref" occurrences;
 wenzelm parents: 
32450diff
changeset | 113 | |
| 27970 | 114 | /* position */ | 
| 115 | ||
| 116 | val LINE = "line" | |
| 117 | val COLUMN = "column" | |
| 118 | val OFFSET = "offset" | |
| 119 | val END_OFFSET = "end_offset" | |
| 120 | val FILE = "file" | |
| 121 | val ID = "id" | |
| 122 | ||
| 42327 
7c7cc7590eb3
Name_Space.entry_markup: keep def position as separate properties;
 wenzelm parents: 
42202diff
changeset | 123 | val DEF_LINE = "def_line" | 
| 
7c7cc7590eb3
Name_Space.entry_markup: keep def position as separate properties;
 wenzelm parents: 
42202diff
changeset | 124 | val DEF_COLUMN = "def_column" | 
| 
7c7cc7590eb3
Name_Space.entry_markup: keep def position as separate properties;
 wenzelm parents: 
42202diff
changeset | 125 | val DEF_OFFSET = "def_offset" | 
| 
7c7cc7590eb3
Name_Space.entry_markup: keep def position as separate properties;
 wenzelm parents: 
42202diff
changeset | 126 | val DEF_END_OFFSET = "def_end_offset" | 
| 
7c7cc7590eb3
Name_Space.entry_markup: keep def position as separate properties;
 wenzelm parents: 
42202diff
changeset | 127 | val DEF_FILE = "def_file" | 
| 
7c7cc7590eb3
Name_Space.entry_markup: keep def position as separate properties;
 wenzelm parents: 
42202diff
changeset | 128 | val DEF_ID = "def_id" | 
| 
7c7cc7590eb3
Name_Space.entry_markup: keep def position as separate properties;
 wenzelm parents: 
42202diff
changeset | 129 | |
| 41483 | 130 | val POSITION_PROPERTIES = Set(LINE, COLUMN, OFFSET, END_OFFSET, FILE, ID) | 
| 43593 | 131 | val POSITION = "position" | 
| 29205 | 132 | |
| 43593 | 133 | |
| 134 | /* path */ | |
| 135 | ||
| 136 | val PATH = "path" | |
| 137 | ||
| 138 | object Path | |
| 139 |   {
 | |
| 140 | def unapply(markup: Markup): Option[String] = | |
| 141 |       markup match {
 | |
| 142 | case Markup(PATH, Name(name)) => Some(name) | |
| 143 | case _ => None | |
| 144 | } | |
| 145 | } | |
| 29184 | 146 | |
| 147 | ||
| 36683 | 148 | /* pretty printing */ | 
| 149 | ||
| 38724 | 150 |   val Indent = new Int_Property("indent")
 | 
| 36683 | 151 | val BLOCK = "block" | 
| 38724 | 152 |   val Width = new Int_Property("width")
 | 
| 36683 | 153 | val BREAK = "break" | 
| 154 | ||
| 155 | ||
| 33985 | 156 | /* hidden text */ | 
| 157 | ||
| 158 | val HIDDEN = "hidden" | |
| 159 | ||
| 160 | ||
| 29184 | 161 | /* logical entities */ | 
| 162 | ||
| 43551 | 163 | val CLASS = "class" | 
| 43552 | 164 | val TYPE = "type" | 
| 29184 | 165 | val FIXED = "fixed" | 
| 43552 | 166 | val CONSTANT = "constant" | 
| 167 | ||
| 29184 | 168 | val DYNAMIC_FACT = "dynamic_fact" | 
| 169 | ||
| 170 | ||
| 171 | /* inner syntax */ | |
| 172 | ||
| 173 | val TFREE = "tfree" | |
| 174 | val TVAR = "tvar" | |
| 175 | val FREE = "free" | |
| 176 | val SKOLEM = "skolem" | |
| 177 | val BOUND = "bound" | |
| 178 | val VAR = "var" | |
| 179 | val NUM = "num" | |
| 180 | val FLOAT = "float" | |
| 181 | val XNUM = "xnum" | |
| 182 | val XSTR = "xstr" | |
| 183 | val LITERAL = "literal" | |
| 43432 
224006e5ac46
inner literal/delimiter corresponds to outer keyword/operator;
 wenzelm parents: 
43386diff
changeset | 184 | val DELIMITER = "delimiter" | 
| 43386 
4e78dd88c64f
more foreground markup, using actual CSS color names;
 wenzelm parents: 
42492diff
changeset | 185 | val INNER_STRING = "inner_string" | 
| 29184 | 186 | val INNER_COMMENT = "inner_comment" | 
| 187 | ||
| 39168 
e3ac771235f7
report token range after inner parse error -- often provides important clues about misunderstanding concerning lexical phase;
 wenzelm parents: 
38887diff
changeset | 188 | val TOKEN_RANGE = "token_range" | 
| 
e3ac771235f7
report token range after inner parse error -- often provides important clues about misunderstanding concerning lexical phase;
 wenzelm parents: 
38887diff
changeset | 189 | |
| 29184 | 190 | val SORT = "sort" | 
| 191 | val TYP = "typ" | |
| 192 | val TERM = "term" | |
| 193 | val PROP = "prop" | |
| 194 | ||
| 195 | val ATTRIBUTE = "attribute" | |
| 196 | val METHOD = "method" | |
| 197 | ||
| 198 | ||
| 199 | /* embedded source text */ | |
| 200 | ||
| 201 | val ML_SOURCE = "ML_source" | |
| 202 | val DOC_SOURCE = "doc_source" | |
| 203 | ||
| 204 | val ANTIQ = "antiq" | |
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43552diff
changeset | 205 | val ML_ANTIQUOTATION = "ML antiquotation" | 
| 43564 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43560diff
changeset | 206 | val DOCUMENT_ANTIQUOTATION = "document antiquotation" | 
| 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43560diff
changeset | 207 | val DOCUMENT_ANTIQUOTATION_OPTION = "document antiquotation option" | 
| 29184 | 208 | |
| 209 | ||
| 30615 | 210 | /* ML syntax */ | 
| 211 | ||
| 212 | val ML_KEYWORD = "ML_keyword" | |
| 37195 | 213 | val ML_DELIMITER = "ML_delimiter" | 
| 30615 | 214 | val ML_IDENT = "ML_ident" | 
| 215 | val ML_TVAR = "ML_tvar" | |
| 216 | val ML_NUMERAL = "ML_numeral" | |
| 217 | val ML_CHAR = "ML_char" | |
| 218 | val ML_STRING = "ML_string" | |
| 219 | val ML_COMMENT = "ML_comment" | |
| 220 | val ML_MALFORMED = "ML_malformed" | |
| 221 | ||
| 30702 | 222 | val ML_DEF = "ML_def" | 
| 31472 | 223 | val ML_OPEN = "ML_open" | 
| 224 | val ML_STRUCT = "ML_struct" | |
| 30702 | 225 | val ML_TYPING = "ML_typing" | 
| 226 | ||
| 30615 | 227 | |
| 29184 | 228 | /* outer syntax */ | 
| 229 | ||
| 230 | val KEYWORD_DECL = "keyword_decl" | |
| 231 | val COMMAND_DECL = "command_decl" | |
| 232 | ||
| 233 | val KEYWORD = "keyword" | |
| 37194 | 234 | val OPERATOR = "operator" | 
| 29184 | 235 | val COMMAND = "command" | 
| 236 | val IDENT = "ident" | |
| 237 | val STRING = "string" | |
| 238 | val ALTSTRING = "altstring" | |
| 239 | val VERBATIM = "verbatim" | |
| 240 | val COMMENT = "comment" | |
| 241 | val CONTROL = "control" | |
| 242 | val MALFORMED = "malformed" | |
| 243 | ||
| 29185 | 244 | val COMMAND_SPAN = "command_span" | 
| 245 | val IGNORED_SPAN = "ignored_span" | |
| 246 | val MALFORMED_SPAN = "malformed_span" | |
| 247 | ||
| 29184 | 248 | |
| 40394 
6dcb6cbf0719
somewhat more uniform timing markup in ML vs. Scala;
 wenzelm parents: 
40392diff
changeset | 249 | /* timing */ | 
| 
6dcb6cbf0719
somewhat more uniform timing markup in ML vs. Scala;
 wenzelm parents: 
40392diff
changeset | 250 | |
| 
6dcb6cbf0719
somewhat more uniform timing markup in ML vs. Scala;
 wenzelm parents: 
40392diff
changeset | 251 | val TIMING = "timing" | 
| 
6dcb6cbf0719
somewhat more uniform timing markup in ML vs. Scala;
 wenzelm parents: 
40392diff
changeset | 252 | val ELAPSED = "elapsed" | 
| 
6dcb6cbf0719
somewhat more uniform timing markup in ML vs. Scala;
 wenzelm parents: 
40392diff
changeset | 253 | val CPU = "cpu" | 
| 
6dcb6cbf0719
somewhat more uniform timing markup in ML vs. Scala;
 wenzelm parents: 
40392diff
changeset | 254 | val GC = "gc" | 
| 
6dcb6cbf0719
somewhat more uniform timing markup in ML vs. Scala;
 wenzelm parents: 
40392diff
changeset | 255 | |
| 
6dcb6cbf0719
somewhat more uniform timing markup in ML vs. Scala;
 wenzelm parents: 
40392diff
changeset | 256 | object Timing | 
| 
6dcb6cbf0719
somewhat more uniform timing markup in ML vs. Scala;
 wenzelm parents: 
40392diff
changeset | 257 |   {
 | 
| 
6dcb6cbf0719
somewhat more uniform timing markup in ML vs. Scala;
 wenzelm parents: 
40392diff
changeset | 258 | def apply(timing: isabelle.Timing): Markup = | 
| 
6dcb6cbf0719
somewhat more uniform timing markup in ML vs. Scala;
 wenzelm parents: 
40392diff
changeset | 259 | Markup(TIMING, List( | 
| 40848 
8662b9b1f123
more abstract/uniform handling of time, preferring seconds as Double;
 wenzelm parents: 
40394diff
changeset | 260 | (ELAPSED, Double(timing.elapsed.seconds)), | 
| 
8662b9b1f123
more abstract/uniform handling of time, preferring seconds as Double;
 wenzelm parents: 
40394diff
changeset | 261 | (CPU, Double(timing.cpu.seconds)), | 
| 
8662b9b1f123
more abstract/uniform handling of time, preferring seconds as Double;
 wenzelm parents: 
40394diff
changeset | 262 | (GC, Double(timing.gc.seconds)))) | 
| 40394 
6dcb6cbf0719
somewhat more uniform timing markup in ML vs. Scala;
 wenzelm parents: 
40392diff
changeset | 263 | def unapply(markup: Markup): Option[isabelle.Timing] = | 
| 
6dcb6cbf0719
somewhat more uniform timing markup in ML vs. Scala;
 wenzelm parents: 
40392diff
changeset | 264 |       markup match {
 | 
| 
6dcb6cbf0719
somewhat more uniform timing markup in ML vs. Scala;
 wenzelm parents: 
40392diff
changeset | 265 | case Markup(TIMING, List( | 
| 40848 
8662b9b1f123
more abstract/uniform handling of time, preferring seconds as Double;
 wenzelm parents: 
40394diff
changeset | 266 | (ELAPSED, Double(elapsed)), (CPU, Double(cpu)), (GC, Double(gc)))) => | 
| 
8662b9b1f123
more abstract/uniform handling of time, preferring seconds as Double;
 wenzelm parents: 
40394diff
changeset | 267 | Some(new isabelle.Timing(Time.seconds(elapsed), Time.seconds(cpu), Time.seconds(gc))) | 
| 40394 
6dcb6cbf0719
somewhat more uniform timing markup in ML vs. Scala;
 wenzelm parents: 
40392diff
changeset | 268 | case _ => None | 
| 
6dcb6cbf0719
somewhat more uniform timing markup in ML vs. Scala;
 wenzelm parents: 
40392diff
changeset | 269 | } | 
| 
6dcb6cbf0719
somewhat more uniform timing markup in ML vs. Scala;
 wenzelm parents: 
40392diff
changeset | 270 | } | 
| 
6dcb6cbf0719
somewhat more uniform timing markup in ML vs. Scala;
 wenzelm parents: 
40392diff
changeset | 271 | |
| 
6dcb6cbf0719
somewhat more uniform timing markup in ML vs. Scala;
 wenzelm parents: 
40392diff
changeset | 272 | |
| 29184 | 273 | /* toplevel */ | 
| 274 | ||
| 38721 
ca8b14fa0d0d
added some proof state markup, notably number of subgoals (e.g. for indentation);
 wenzelm parents: 
38474diff
changeset | 275 | val SUBGOALS = "subgoals" | 
| 
ca8b14fa0d0d
added some proof state markup, notably number of subgoals (e.g. for indentation);
 wenzelm parents: 
38474diff
changeset | 276 | val PROOF_STATE = "proof_state" | 
| 
ca8b14fa0d0d
added some proof state markup, notably number of subgoals (e.g. for indentation);
 wenzelm parents: 
38474diff
changeset | 277 | |
| 29184 | 278 | val STATE = "state" | 
| 279 | val SUBGOAL = "subgoal" | |
| 280 | val SENDBACK = "sendback" | |
| 281 | val HILITE = "hilite" | |
| 282 | ||
| 283 | ||
| 284 | /* command status */ | |
| 285 | ||
| 29417 | 286 | val TASK = "task" | 
| 287 | ||
| 37186 
349e9223c685
explicit markup for forked goals, as indicated by Goal.fork;
 wenzelm parents: 
37121diff
changeset | 288 | val FORKED = "forked" | 
| 
349e9223c685
explicit markup for forked goals, as indicated by Goal.fork;
 wenzelm parents: 
37121diff
changeset | 289 | val JOINED = "joined" | 
| 29184 | 290 | val FAILED = "failed" | 
| 291 | val FINISHED = "finished" | |
| 29488 | 292 | |
| 293 | ||
| 294 | /* interactive documents */ | |
| 295 | ||
| 38414 
49f1f657adc2
more basic Markup.parse_int/print_int (using signed_string_of_int) (ML);
 wenzelm parents: 
38355diff
changeset | 296 | val VERSION = "version" | 
| 
49f1f657adc2
more basic Markup.parse_int/print_int (using signed_string_of_int) (ML);
 wenzelm parents: 
38355diff
changeset | 297 | val EXEC = "exec" | 
| 
49f1f657adc2
more basic Markup.parse_int/print_int (using signed_string_of_int) (ML);
 wenzelm parents: 
38355diff
changeset | 298 | val ASSIGN = "assign" | 
| 29488 | 299 | val EDIT = "edit" | 
| 29184 | 300 | |
| 27970 | 301 | |
| 302 | /* messages */ | |
| 303 | ||
| 38872 | 304 |   val Serial = new Long_Property("serial")
 | 
| 27970 | 305 | |
| 29195 | 306 | val MESSAGE = "message" | 
| 29522 
793766d4c1b5
moved message markup into Scala layer -- reduced redundancy;
 wenzelm parents: 
29488diff
changeset | 307 | |
| 
793766d4c1b5
moved message markup into Scala layer -- reduced redundancy;
 wenzelm parents: 
29488diff
changeset | 308 | val INIT = "init" | 
| 
793766d4c1b5
moved message markup into Scala layer -- reduced redundancy;
 wenzelm parents: 
29488diff
changeset | 309 | val STATUS = "status" | 
| 39525 | 310 | val REPORT = "report" | 
| 29522 
793766d4c1b5
moved message markup into Scala layer -- reduced redundancy;
 wenzelm parents: 
29488diff
changeset | 311 | val WRITELN = "writeln" | 
| 
793766d4c1b5
moved message markup into Scala layer -- reduced redundancy;
 wenzelm parents: 
29488diff
changeset | 312 | val TRACING = "tracing" | 
| 
793766d4c1b5
moved message markup into Scala layer -- reduced redundancy;
 wenzelm parents: 
29488diff
changeset | 313 | val WARNING = "warning" | 
| 
793766d4c1b5
moved message markup into Scala layer -- reduced redundancy;
 wenzelm parents: 
29488diff
changeset | 314 | val ERROR = "error" | 
| 
793766d4c1b5
moved message markup into Scala layer -- reduced redundancy;
 wenzelm parents: 
29488diff
changeset | 315 | val SYSTEM = "system" | 
| 
793766d4c1b5
moved message markup into Scala layer -- reduced redundancy;
 wenzelm parents: 
29488diff
changeset | 316 | val STDOUT = "stdout" | 
| 
793766d4c1b5
moved message markup into Scala layer -- reduced redundancy;
 wenzelm parents: 
29488diff
changeset | 317 | val EXIT = "exit" | 
| 29195 | 318 | |
| 39439 
1c294d150ded
eliminated markup "location" in favour of more explicit "no_report", which is actually deleted from messages;
 wenzelm parents: 
39171diff
changeset | 319 | val NO_REPORT = "no_report" | 
| 
1c294d150ded
eliminated markup "location" in favour of more explicit "no_report", which is actually deleted from messages;
 wenzelm parents: 
39171diff
changeset | 320 | |
| 39171 
525a13b9ac74
highlight bad range of nested error (here from inner parsing);
 wenzelm parents: 
39168diff
changeset | 321 | val BAD = "bad" | 
| 
525a13b9ac74
highlight bad range of nested error (here from inner parsing);
 wenzelm parents: 
39168diff
changeset | 322 | |
| 39591 | 323 | val READY = "ready" | 
| 31384 
ce169bd37fc0
IsabelleProcess: emit status "ready" after initialization and reports;
 wenzelm parents: 
30702diff
changeset | 324 | |
| 27970 | 325 | |
| 34119 | 326 | /* system data */ | 
| 27970 | 327 | |
| 38231 | 328 |   val Data = Markup("data", Nil)
 | 
| 27958 | 329 | } | 
| 38230 
ed147003de4b
simplified type XML.Tree: embed Markup directly, avoid slightly odd triple;
 wenzelm parents: 
37195diff
changeset | 330 | |
| 
ed147003de4b
simplified type XML.Tree: embed Markup directly, avoid slightly odd triple;
 wenzelm parents: 
37195diff
changeset | 331 | sealed case class Markup(name: String, properties: List[(String, String)]) |