| author | wenzelm | 
| Wed, 30 Dec 2015 20:12:26 +0100 | |
| changeset 61992 | 6d02bb8b5fe1 | 
| parent 61864 | 3a5992c3410c | 
| child 62806 | de9bf8171626 | 
| permissions | -rw-r--r-- | 
| 45670 | 1 | /* Title: Pure/PIDE/markup.scala | 
| 45673 
cd41e3903fbf
separate compilation of PIDE vs. Pure sources, which enables independent Scala library;
 wenzelm parents: 
45670diff
changeset | 2 | Module: PIDE | 
| 27958 | 3 | Author: Makarius | 
| 4 | ||
| 56743 | 5 | Quasi-abstract markup elements. | 
| 27958 | 6 | */ | 
| 7 | ||
| 8 | package isabelle | |
| 9 | ||
| 27970 | 10 | |
| 32450 | 11 | object Markup | 
| 12 | {
 | |
| 56743 | 13 | /* elements */ | 
| 14 | ||
| 15 | object Elements | |
| 16 |   {
 | |
| 17 | def apply(elems: Set[String]): Elements = new Elements(elems) | |
| 18 | def apply(elems: String*): Elements = apply(Set(elems: _*)) | |
| 19 | val empty: Elements = apply() | |
| 20 | val full: Elements = | |
| 21 | new Elements(Set.empty) | |
| 22 |       {
 | |
| 23 | override def apply(elem: String): Boolean = true | |
| 24 | override def toString: String = "Elements.full" | |
| 25 | } | |
| 26 | } | |
| 27 | ||
| 28 | sealed class Elements private[Markup](private val rep: Set[String]) | |
| 29 |   {
 | |
| 30 | def apply(elem: String): Boolean = rep.contains(elem) | |
| 31 | def + (elem: String): Elements = new Elements(rep + elem) | |
| 32 | def ++ (elems: Elements): Elements = new Elements(rep ++ elems.rep) | |
| 33 |     override def toString: String = rep.mkString("Elements(", ",", ")")
 | |
| 34 | } | |
| 35 | ||
| 36 | ||
| 45666 | 37 | /* properties */ | 
| 29184 | 38 | |
| 39 | val NAME = "name" | |
| 43780 | 40 | val Name = new Properties.String(NAME) | 
| 42136 | 41 | |
| 29184 | 42 | val KIND = "kind" | 
| 43780 | 43 | val Kind = new Properties.String(KIND) | 
| 29184 | 44 | |
| 60744 | 45 | val SERIAL = "serial" | 
| 46 | val Serial = new Properties.Long(SERIAL) | |
| 47 | ||
| 52854 
92932931bd82
more general Output.result: allow to update arbitrary properties;
 wenzelm parents: 
52800diff
changeset | 48 | val INSTANCE = "instance" | 
| 
92932931bd82
more general Output.result: allow to update arbitrary properties;
 wenzelm parents: 
52800diff
changeset | 49 | val Instance = new Properties.String(INSTANCE) | 
| 
92932931bd82
more general Output.result: allow to update arbitrary properties;
 wenzelm parents: 
52800diff
changeset | 50 | |
| 29184 | 51 | |
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 52 | /* basic markup */ | 
| 29184 | 53 | |
| 45666 | 54 |   val Empty = Markup("", Nil)
 | 
| 55 |   val Broken = Markup("broken", Nil)
 | |
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 56 | |
| 55551 | 57 | class Markup_String(val name: String, prop: String) | 
| 58 |   {
 | |
| 59 | private val Prop = new Properties.String(prop) | |
| 60 | ||
| 61 | def apply(s: String): Markup = Markup(name, Prop(s)) | |
| 62 | def unapply(markup: Markup): Option[String] = | |
| 63 | if (markup.name == name) Prop.unapply(markup.properties) else None | |
| 64 | } | |
| 65 | ||
| 66 | class Markup_Int(val name: String, prop: String) | |
| 67 |   {
 | |
| 68 | private val Prop = new Properties.Int(prop) | |
| 69 | ||
| 70 | def apply(i: Int): Markup = Markup(name, Prop(i)) | |
| 71 | def unapply(markup: Markup): Option[Int] = | |
| 72 | if (markup.name == name) Prop.unapply(markup.properties) else None | |
| 73 | } | |
| 74 | ||
| 60744 | 75 | class Markup_Long(val name: String, prop: String) | 
| 76 |   {
 | |
| 77 | private val Prop = new Properties.Long(prop) | |
| 78 | ||
| 79 | def apply(i: Long): Markup = Markup(name, Prop(i)) | |
| 80 | def unapply(markup: Markup): Option[Long] = | |
| 81 | if (markup.name == name) Prop.unapply(markup.properties) else None | |
| 82 | } | |
| 83 | ||
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 84 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 85 | /* formal entities */ | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 86 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 87 | val BINDING = "binding" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 88 | val ENTITY = "entity" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 89 | val DEF = "def" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 90 | val REF = "ref" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 91 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 92 | object Entity | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 93 |   {
 | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 94 | def unapply(markup: Markup): Option[(String, String)] = | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 95 |       markup match {
 | 
| 55551 | 96 | case Markup(ENTITY, props) => | 
| 97 |           (props, props) match {
 | |
| 55615 
bf4bbe72f740
completion of keywords and symbols based on language context;
 wenzelm parents: 
55553diff
changeset | 98 | case (Kind(kind), Name(name)) => Some((kind, name)) | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 99 | case _ => None | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 100 | } | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 101 | case _ => None | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 102 | } | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 103 | } | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 104 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 105 | |
| 55672 
5e25cc741ab9
support for completion within the formal context;
 wenzelm parents: 
55666diff
changeset | 106 | /* completion */ | 
| 
5e25cc741ab9
support for completion within the formal context;
 wenzelm parents: 
55666diff
changeset | 107 | |
| 
5e25cc741ab9
support for completion within the formal context;
 wenzelm parents: 
55666diff
changeset | 108 | val COMPLETION = "completion" | 
| 55914 
c5b752d549e3
clarified init_assignable: make double-sure that initial values are reset;
 wenzelm parents: 
55837diff
changeset | 109 | val NO_COMPLETION = "no_completion" | 
| 55672 
5e25cc741ab9
support for completion within the formal context;
 wenzelm parents: 
55666diff
changeset | 110 | |
| 
5e25cc741ab9
support for completion within the formal context;
 wenzelm parents: 
55666diff
changeset | 111 | |
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 112 | /* position */ | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 113 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 114 | val LINE = "line" | 
| 58978 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 wenzelm parents: 
58853diff
changeset | 115 | val END_LINE = "line" | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 116 | val OFFSET = "offset" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 117 | val END_OFFSET = "end_offset" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 118 | val FILE = "file" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 119 | val ID = "id" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 120 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 121 | val DEF_LINE = "def_line" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 122 | val DEF_OFFSET = "def_offset" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 123 | val DEF_END_OFFSET = "def_end_offset" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 124 | val DEF_FILE = "def_file" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 125 | val DEF_ID = "def_id" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 126 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 127 | val POSITION_PROPERTIES = Set(LINE, OFFSET, END_OFFSET, FILE, ID) | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 128 | val POSITION = "position" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 129 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 130 | |
| 58464 | 131 | /* expression */ | 
| 132 | ||
| 133 | val EXPRESSION = "expression" | |
| 134 | ||
| 135 | ||
| 58544 
340f130b3d38
bibtex support in ML: document antiquotation @{cite} with markup;
 wenzelm parents: 
58464diff
changeset | 136 | /* citation */ | 
| 
340f130b3d38
bibtex support in ML: document antiquotation @{cite} with markup;
 wenzelm parents: 
58464diff
changeset | 137 | |
| 
340f130b3d38
bibtex support in ML: document antiquotation @{cite} with markup;
 wenzelm parents: 
58464diff
changeset | 138 | val CITATION = "citation" | 
| 58545 
30b75b7958d6
citation tooltip/hyperlink based on open buffers with .bib files;
 wenzelm parents: 
58544diff
changeset | 139 | val Citation = new Markup_String(CITATION, NAME) | 
| 58544 
340f130b3d38
bibtex support in ML: document antiquotation @{cite} with markup;
 wenzelm parents: 
58464diff
changeset | 140 | |
| 
340f130b3d38
bibtex support in ML: document antiquotation @{cite} with markup;
 wenzelm parents: 
58464diff
changeset | 141 | |
| 55550 | 142 | /* embedded languages */ | 
| 143 | ||
| 55666 | 144 |   val Symbols = new Properties.Boolean("symbols")
 | 
| 145 |   val Antiquotes = new Properties.Boolean("antiquotes")
 | |
| 55828 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
55765diff
changeset | 146 |   val Delimited = new Properties.Boolean("delimited")
 | 
| 55615 
bf4bbe72f740
completion of keywords and symbols based on language context;
 wenzelm parents: 
55553diff
changeset | 147 | |
| 55550 | 148 | val LANGUAGE = "language" | 
| 55615 
bf4bbe72f740
completion of keywords and symbols based on language context;
 wenzelm parents: 
55553diff
changeset | 149 | object Language | 
| 
bf4bbe72f740
completion of keywords and symbols based on language context;
 wenzelm parents: 
55553diff
changeset | 150 |   {
 | 
| 55616 | 151 | val ML = "ML" | 
| 56278 
2576d3a40ed6
separate tokenization and language context for SML: no symbols, no antiquotes;
 wenzelm parents: 
56202diff
changeset | 152 | val SML = "SML" | 
| 56843 
b2bfcd8cda80
support for path completion based on file-system content;
 wenzelm parents: 
56743diff
changeset | 153 | val PATH = "path" | 
| 55616 | 154 | val UNKNOWN = "unknown" | 
| 155 | ||
| 55828 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
55765diff
changeset | 156 | def unapply(markup: Markup): Option[(String, Boolean, Boolean, Boolean)] = | 
| 55615 
bf4bbe72f740
completion of keywords and symbols based on language context;
 wenzelm parents: 
55553diff
changeset | 157 |       markup match {
 | 
| 
bf4bbe72f740
completion of keywords and symbols based on language context;
 wenzelm parents: 
55553diff
changeset | 158 | case Markup(LANGUAGE, props) => | 
| 55828 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
55765diff
changeset | 159 |           (props, props, props, props) match {
 | 
| 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
55765diff
changeset | 160 | case (Name(name), Symbols(symbols), Antiquotes(antiquotes), Delimited(delimited)) => | 
| 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
55765diff
changeset | 161 | Some((name, symbols, antiquotes, delimited)) | 
| 55615 
bf4bbe72f740
completion of keywords and symbols based on language context;
 wenzelm parents: 
55553diff
changeset | 162 | case _ => None | 
| 
bf4bbe72f740
completion of keywords and symbols based on language context;
 wenzelm parents: 
55553diff
changeset | 163 | } | 
| 
bf4bbe72f740
completion of keywords and symbols based on language context;
 wenzelm parents: 
55553diff
changeset | 164 | case _ => None | 
| 
bf4bbe72f740
completion of keywords and symbols based on language context;
 wenzelm parents: 
55553diff
changeset | 165 | } | 
| 
bf4bbe72f740
completion of keywords and symbols based on language context;
 wenzelm parents: 
55553diff
changeset | 166 | } | 
| 55550 | 167 | |
| 168 | ||
| 54702 
3daeba5130f0
added document antiquotation @{url}, which produces formal markup for LaTeX and PIDE;
 wenzelm parents: 
53378diff
changeset | 169 | /* external resources */ | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 170 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 171 | val PATH = "path" | 
| 55551 | 172 | val Path = new Markup_String(PATH, NAME) | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 173 | |
| 54702 
3daeba5130f0
added document antiquotation @{url}, which produces formal markup for LaTeX and PIDE;
 wenzelm parents: 
53378diff
changeset | 174 | val URL = "url" | 
| 55551 | 175 | val Url = new Markup_String(URL, NAME) | 
| 54702 
3daeba5130f0
added document antiquotation @{url}, which produces formal markup for LaTeX and PIDE;
 wenzelm parents: 
53378diff
changeset | 176 | |
| 61660 
78b371644654
added antiquotation @{doc}, e.g. useful for demonstration purposes;
 wenzelm parents: 
61598diff
changeset | 177 | val DOC = "doc" | 
| 
78b371644654
added antiquotation @{doc}, e.g. useful for demonstration purposes;
 wenzelm parents: 
61598diff
changeset | 178 | val Doc = new Markup_String(DOC, NAME) | 
| 
78b371644654
added antiquotation @{doc}, e.g. useful for demonstration purposes;
 wenzelm parents: 
61598diff
changeset | 179 | |
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 180 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 181 | /* pretty printing */ | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 182 | |
| 61864 | 183 |   val Consistent = new Properties.Boolean("consistent")
 | 
| 61862 
e2a9e46ac0fb
support pretty break indent, like underlying ML systems;
 wenzelm parents: 
61660diff
changeset | 184 |   val Indent = new Properties.Int("indent")
 | 
| 61864 | 185 |   val Width = new Properties.Int("width")
 | 
| 61862 
e2a9e46ac0fb
support pretty break indent, like underlying ML systems;
 wenzelm parents: 
61660diff
changeset | 186 | |
| 61864 | 187 | object Block | 
| 188 |   {
 | |
| 189 | val name = "block" | |
| 190 | def apply(c: Boolean, i: Int): Markup = | |
| 191 | Markup(name, | |
| 192 | (if (c) Consistent(c) else Nil) ::: | |
| 193 | (if (i != 0) Indent(i) else Nil)) | |
| 194 | def unapply(markup: Markup): Option[(Boolean, Int)] = | |
| 195 |       if (markup.name == name) {
 | |
| 196 | val c = Consistent.unapply(markup.properties).getOrElse(false) | |
| 197 | val i = Indent.unapply(markup.properties).getOrElse(0) | |
| 198 | Some((c, i)) | |
| 199 | } | |
| 200 | else None | |
| 201 | } | |
| 202 | ||
| 61862 
e2a9e46ac0fb
support pretty break indent, like underlying ML systems;
 wenzelm parents: 
61660diff
changeset | 203 | object Break | 
| 
e2a9e46ac0fb
support pretty break indent, like underlying ML systems;
 wenzelm parents: 
61660diff
changeset | 204 |   {
 | 
| 
e2a9e46ac0fb
support pretty break indent, like underlying ML systems;
 wenzelm parents: 
61660diff
changeset | 205 | val name = "break" | 
| 61864 | 206 | def apply(w: Int, i: Int): Markup = | 
| 207 | Markup(name, | |
| 208 | (if (w != 0) Width(w) else Nil) ::: | |
| 209 | (if (i != 0) Indent(i) else Nil)) | |
| 61862 
e2a9e46ac0fb
support pretty break indent, like underlying ML systems;
 wenzelm parents: 
61660diff
changeset | 210 | def unapply(markup: Markup): Option[(Int, Int)] = | 
| 
e2a9e46ac0fb
support pretty break indent, like underlying ML systems;
 wenzelm parents: 
61660diff
changeset | 211 |       if (markup.name == name) {
 | 
| 61864 | 212 | val w = Width.unapply(markup.properties).getOrElse(0) | 
| 213 | val i = Indent.unapply(markup.properties).getOrElse(0) | |
| 214 | Some((w, i)) | |
| 61862 
e2a9e46ac0fb
support pretty break indent, like underlying ML systems;
 wenzelm parents: 
61660diff
changeset | 215 | } | 
| 
e2a9e46ac0fb
support pretty break indent, like underlying ML systems;
 wenzelm parents: 
61660diff
changeset | 216 | else None | 
| 
e2a9e46ac0fb
support pretty break indent, like underlying ML systems;
 wenzelm parents: 
61660diff
changeset | 217 | } | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 218 | |
| 51570 
3633828d80fc
basic support for Pretty.item, which is considered as logical markup and interpreted in Isabelle/Scala, but ignored elsewhere (TTY, latex etc.);
 wenzelm parents: 
50975diff
changeset | 219 | val ITEM = "item" | 
| 51574 
2b58d7b139d6
ghost bullet via markup, which is painted as bar under text (normally space);
 wenzelm parents: 
51570diff
changeset | 220 | val BULLET = "bullet" | 
| 51570 
3633828d80fc
basic support for Pretty.item, which is considered as logical markup and interpreted in Isabelle/Scala, but ignored elsewhere (TTY, latex etc.);
 wenzelm parents: 
50975diff
changeset | 221 | |
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 222 | val SEPARATOR = "separator" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 223 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 224 | |
| 56548 | 225 | /* text properties */ | 
| 226 | ||
| 227 | val WORDS = "words" | |
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 228 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 229 | val HIDDEN = "hidden" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 230 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 231 | |
| 58048 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
57595diff
changeset | 232 | /* misc entities */ | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 233 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 234 | val CLASS = "class" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 235 | val TYPE_NAME = "type_name" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 236 | val FIXED = "fixed" | 
| 53378 
07990ba8c0ea
cases: more position information and PIDE markup;
 wenzelm parents: 
53055diff
changeset | 237 | val CASE = "case" | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 238 | val CONSTANT = "constant" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 239 | val DYNAMIC_FACT = "dynamic_fact" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 240 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 241 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 242 | /* inner syntax */ | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 243 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 244 | val TFREE = "tfree" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 245 | val TVAR = "tvar" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 246 | val FREE = "free" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 247 | val SKOLEM = "skolem" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 248 | val BOUND = "bound" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 249 | val VAR = "var" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 250 | val NUMERAL = "numeral" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 251 | val LITERAL = "literal" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 252 | val DELIMITER = "delimiter" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 253 | val INNER_STRING = "inner_string" | 
| 55033 | 254 | val INNER_CARTOUCHE = "inner_cartouche" | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 255 | val INNER_COMMENT = "inner_comment" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 256 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 257 | val TOKEN_RANGE = "token_range" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 258 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 259 | val SORTING = "sorting" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 260 | val TYPING = "typing" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 261 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 262 | val ATTRIBUTE = "attribute" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 263 | val METHOD = "method" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 264 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 265 | |
| 55550 | 266 | /* antiquotations */ | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 267 | |
| 55526 | 268 | val ANTIQUOTED = "antiquoted" | 
| 269 | val ANTIQUOTE = "antiquote" | |
| 270 | ||
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 271 | val ML_ANTIQUOTATION = "ML_antiquotation" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 272 | val DOCUMENT_ANTIQUOTATION = "document_antiquotation" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 273 | val DOCUMENT_ANTIQUOTATION_OPTION = "document_antiquotation_option" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 274 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 275 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 276 | /* text structure */ | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 277 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 278 | val PARAGRAPH = "paragraph" | 
| 50545 
00bdc48c5f71
explicit text_fold markup, which is used by default in Pretty.chunks/chunks2;
 wenzelm parents: 
50543diff
changeset | 279 | val TEXT_FOLD = "text_fold" | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 280 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 281 | |
| 61449 | 282 | /* Markdown document structure */ | 
| 283 | ||
| 284 | val MARKDOWN_PARAGRAPH = "markdown_paragraph" | |
| 285 |   val Markdown_List = new Markup_String("markdown_list", "kind")
 | |
| 286 |   val Markdown_Item = new Markup_Int("markdown_item", "depth")
 | |
| 287 | ||
| 288 | ||
| 60744 | 289 | /* ML */ | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 290 | |
| 55505 | 291 | val ML_KEYWORD1 = "ML_keyword1" | 
| 292 | val ML_KEYWORD2 = "ML_keyword2" | |
| 293 | val ML_KEYWORD3 = "ML_keyword3" | |
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 294 | val ML_DELIMITER = "ML_delimiter" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 295 | val ML_TVAR = "ML_tvar" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 296 | val ML_NUMERAL = "ML_numeral" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 297 | val ML_CHAR = "ML_char" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 298 | val ML_STRING = "ML_string" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 299 | val ML_COMMENT = "ML_comment" | 
| 56278 
2576d3a40ed6
separate tokenization and language context for SML: no symbols, no antiquotes;
 wenzelm parents: 
56202diff
changeset | 300 | val SML_STRING = "SML_string" | 
| 
2576d3a40ed6
separate tokenization and language context for SML: no symbols, no antiquotes;
 wenzelm parents: 
56202diff
changeset | 301 | val SML_COMMENT = "SML_comment" | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 302 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 303 | val ML_DEF = "ML_def" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 304 | val ML_OPEN = "ML_open" | 
| 55837 | 305 | val ML_STRUCTURE = "ML_structure" | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 306 | val ML_TYPING = "ML_typing" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 307 | |
| 60744 | 308 | val ML_BREAKPOINT = "ML_breakpoint" | 
| 309 | ||
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 310 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 311 | /* outer syntax */ | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 312 | |
| 55744 
4a4e5686e091
clarified token markup: keyword1/keyword2 is for syntax, and "command" the entity kind;
 wenzelm parents: 
55694diff
changeset | 313 | val COMMAND = "command" | 
| 
4a4e5686e091
clarified token markup: keyword1/keyword2 is for syntax, and "command" the entity kind;
 wenzelm parents: 
55694diff
changeset | 314 | val KEYWORD1 = "keyword1" | 
| 
4a4e5686e091
clarified token markup: keyword1/keyword2 is for syntax, and "command" the entity kind;
 wenzelm parents: 
55694diff
changeset | 315 | val KEYWORD2 = "keyword2" | 
| 55765 | 316 | val KEYWORD3 = "keyword3" | 
| 55919 
2eb8c13339a5
more explicit quasi_keyword markup, for Args.$$$ material, which is somewhere in between of outer and inner syntax;
 wenzelm parents: 
55914diff
changeset | 317 | val QUASI_KEYWORD = "quasi_keyword" | 
| 56202 | 318 | val IMPROPER = "improper" | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 319 | val OPERATOR = "operator" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 320 | val STRING = "string" | 
| 59081 | 321 | val ALT_STRING = "alt_string" | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 322 | val VERBATIM = "verbatim" | 
| 55033 | 323 | val CARTOUCHE = "cartouche" | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 324 | val COMMENT = "comment" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 325 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 326 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 327 | /* timing */ | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 328 | |
| 50781 | 329 |   val Elapsed = new Properties.Double("elapsed")
 | 
| 330 |   val CPU = new Properties.Double("cpu")
 | |
| 331 |   val GC = new Properties.Double("gc")
 | |
| 332 | ||
| 333 | object Timing_Properties | |
| 334 |   {
 | |
| 335 | def apply(timing: isabelle.Timing): Properties.T = | |
| 336 | Elapsed(timing.elapsed.seconds) ::: CPU(timing.cpu.seconds) ::: GC(timing.gc.seconds) | |
| 51662 
3391a493f39a
just one timing protocol function, with 3 implementations: TTY/PG, PIDE/document, build;
 wenzelm parents: 
51574diff
changeset | 337 | |
| 50781 | 338 | def unapply(props: Properties.T): Option[isabelle.Timing] = | 
| 339 |       (props, props, props) match {
 | |
| 340 | case (Elapsed(elapsed), CPU(cpu), GC(gc)) => | |
| 341 | Some(new isabelle.Timing(Time.seconds(elapsed), Time.seconds(cpu), Time.seconds(gc))) | |
| 342 | case _ => None | |
| 343 | } | |
| 344 | } | |
| 345 | ||
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 346 | val TIMING = "timing" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 347 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 348 | object Timing | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 349 |   {
 | 
| 50781 | 350 | def apply(timing: isabelle.Timing): Markup = Markup(TIMING, Timing_Properties(timing)) | 
| 51662 
3391a493f39a
just one timing protocol function, with 3 implementations: TTY/PG, PIDE/document, build;
 wenzelm parents: 
51574diff
changeset | 351 | |
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 352 | def unapply(markup: Markup): Option[isabelle.Timing] = | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 353 |       markup match {
 | 
| 50781 | 354 | case Markup(TIMING, Timing_Properties(timing)) => Some(timing) | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 355 | case _ => None | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 356 | } | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 357 | } | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 358 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 359 | |
| 51662 
3391a493f39a
just one timing protocol function, with 3 implementations: TTY/PG, PIDE/document, build;
 wenzelm parents: 
51574diff
changeset | 360 | /* command timing */ | 
| 
3391a493f39a
just one timing protocol function, with 3 implementations: TTY/PG, PIDE/document, build;
 wenzelm parents: 
51574diff
changeset | 361 | |
| 51818 
517f232e867d
clarified module dependencies: avoid Properties and Document introding minimal "PIDE";
 wenzelm parents: 
51662diff
changeset | 362 | val COMMAND_TIMING = "command_timing" | 
| 51662 
3391a493f39a
just one timing protocol function, with 3 implementations: TTY/PG, PIDE/document, build;
 wenzelm parents: 
51574diff
changeset | 363 | |
| 
3391a493f39a
just one timing protocol function, with 3 implementations: TTY/PG, PIDE/document, build;
 wenzelm parents: 
51574diff
changeset | 364 | |
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 365 | /* toplevel */ | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 366 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 367 | val SUBGOALS = "subgoals" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 368 | val PROOF_STATE = "proof_state" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 369 | |
| 50543 | 370 | val GOAL = "goal" | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 371 | val SUBGOAL = "subgoal" | 
| 50215 | 372 | |
| 50450 
358b6020f8b6
generalized notion of active area, where sendback is just one application;
 wenzelm parents: 
50255diff
changeset | 373 | |
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 374 | /* command status */ | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 375 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 376 | val TASK = "task" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 377 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 378 | val ACCEPTED = "accepted" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 379 | val FORKED = "forked" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 380 | val JOINED = "joined" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 381 | val RUNNING = "running" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 382 | val FINISHED = "finished" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 383 | val FAILED = "failed" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 384 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 385 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 386 | /* interactive documents */ | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 387 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 388 | val VERSION = "version" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 389 | val ASSIGN = "assign" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 390 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 391 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 392 | /* prover process */ | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 393 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 394 | val PROVER_COMMAND = "prover_command" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 395 | val PROVER_ARG = "prover_arg" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 396 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 397 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 398 | /* messages */ | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 399 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 400 | val INIT = "init" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 401 | val STATUS = "status" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 402 | val REPORT = "report" | 
| 50500 
c94bba7906d2
identify dialogs via official serial and maintain as result message;
 wenzelm parents: 
50499diff
changeset | 403 | val RESULT = "result" | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 404 | val WRITELN = "writeln" | 
| 59184 
830bb7ddb3ab
explicit message channels for "state", "information";
 wenzelm parents: 
59112diff
changeset | 405 | val STATE = "state" | 
| 
830bb7ddb3ab
explicit message channels for "state", "information";
 wenzelm parents: 
59112diff
changeset | 406 | val INFORMATION = "information" | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 407 | val TRACING = "tracing" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 408 | val WARNING = "warning" | 
| 59203 
5f0bd5afc16d
explicit message channel for "legacy", which is nonetheless a variant of "warning";
 wenzelm parents: 
59186diff
changeset | 409 | val LEGACY = "legacy" | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 410 | val ERROR = "error" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 411 | val PROTOCOL = "protocol" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 412 | val SYSTEM = "system" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 413 | val STDOUT = "stdout" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 414 | val STDERR = "stderr" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 415 | val EXIT = "exit" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 416 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 417 | val WRITELN_MESSAGE = "writeln_message" | 
| 59184 
830bb7ddb3ab
explicit message channels for "state", "information";
 wenzelm parents: 
59112diff
changeset | 418 | val STATE_MESSAGE = "state_message" | 
| 
830bb7ddb3ab
explicit message channels for "state", "information";
 wenzelm parents: 
59112diff
changeset | 419 | val INFORMATION_MESSAGE = "information_message" | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 420 | val TRACING_MESSAGE = "tracing_message" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 421 | val WARNING_MESSAGE = "warning_message" | 
| 59203 
5f0bd5afc16d
explicit message channel for "legacy", which is nonetheless a variant of "warning";
 wenzelm parents: 
59186diff
changeset | 422 | val LEGACY_MESSAGE = "legacy_message" | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 423 | val ERROR_MESSAGE = "error_message" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 424 | |
| 59184 
830bb7ddb3ab
explicit message channels for "state", "information";
 wenzelm parents: 
59112diff
changeset | 425 | val messages = Map( | 
| 
830bb7ddb3ab
explicit message channels for "state", "information";
 wenzelm parents: 
59112diff
changeset | 426 | WRITELN -> WRITELN_MESSAGE, | 
| 
830bb7ddb3ab
explicit message channels for "state", "information";
 wenzelm parents: 
59112diff
changeset | 427 | STATE -> STATE_MESSAGE, | 
| 
830bb7ddb3ab
explicit message channels for "state", "information";
 wenzelm parents: 
59112diff
changeset | 428 | INFORMATION -> INFORMATION_MESSAGE, | 
| 
830bb7ddb3ab
explicit message channels for "state", "information";
 wenzelm parents: 
59112diff
changeset | 429 | TRACING -> TRACING_MESSAGE, | 
| 
830bb7ddb3ab
explicit message channels for "state", "information";
 wenzelm parents: 
59112diff
changeset | 430 | WARNING -> WARNING_MESSAGE, | 
| 59203 
5f0bd5afc16d
explicit message channel for "legacy", which is nonetheless a variant of "warning";
 wenzelm parents: 
59186diff
changeset | 431 | LEGACY -> LEGACY_MESSAGE, | 
| 59184 
830bb7ddb3ab
explicit message channels for "state", "information";
 wenzelm parents: 
59112diff
changeset | 432 | ERROR -> ERROR_MESSAGE) | 
| 
830bb7ddb3ab
explicit message channels for "state", "information";
 wenzelm parents: 
59112diff
changeset | 433 | |
| 52876 | 434 | val message: String => String = messages.withDefault((s: String) => s) | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 435 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 436 |   val Return_Code = new Properties.Int("return_code")
 | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 437 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 438 | val NO_REPORT = "no_report" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 439 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 440 | val BAD = "bad" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 441 | |
| 50500 
c94bba7906d2
identify dialogs via official serial and maintain as result message;
 wenzelm parents: 
50499diff
changeset | 442 | val INTENSIFY = "intensify" | 
| 
c94bba7906d2
identify dialogs via official serial and maintain as result message;
 wenzelm parents: 
50499diff
changeset | 443 | |
| 
c94bba7906d2
identify dialogs via official serial and maintain as result message;
 wenzelm parents: 
50499diff
changeset | 444 | |
| 
c94bba7906d2
identify dialogs via official serial and maintain as result message;
 wenzelm parents: 
50499diff
changeset | 445 | /* active areas */ | 
| 
c94bba7906d2
identify dialogs via official serial and maintain as result message;
 wenzelm parents: 
50499diff
changeset | 446 | |
| 50715 
8cfd585b9162
prefer old graph browser in Isabelle/jEdit, which still produces better layout;
 wenzelm parents: 
50545diff
changeset | 447 | val BROWSER = "browser" | 
| 50500 
c94bba7906d2
identify dialogs via official serial and maintain as result message;
 wenzelm parents: 
50499diff
changeset | 448 | val GRAPHVIEW = "graphview" | 
| 
c94bba7906d2
identify dialogs via official serial and maintain as result message;
 wenzelm parents: 
50499diff
changeset | 449 | |
| 
c94bba7906d2
identify dialogs via official serial and maintain as result message;
 wenzelm parents: 
50499diff
changeset | 450 | val SENDBACK = "sendback" | 
| 
c94bba7906d2
identify dialogs via official serial and maintain as result message;
 wenzelm parents: 
50499diff
changeset | 451 | val PADDING = "padding" | 
| 52697 
6fb98a20c349
explicit padding on command boundary for "auto" generated sendback -- do not replace the corresponding goal command, but append to it;
 wenzelm parents: 
52643diff
changeset | 452 | val PADDING_LINE = (PADDING, "line") | 
| 
6fb98a20c349
explicit padding on command boundary for "auto" generated sendback -- do not replace the corresponding goal command, but append to it;
 wenzelm parents: 
52643diff
changeset | 453 | val PADDING_COMMAND = (PADDING, "command") | 
| 50500 
c94bba7906d2
identify dialogs via official serial and maintain as result message;
 wenzelm parents: 
50499diff
changeset | 454 | |
| 
c94bba7906d2
identify dialogs via official serial and maintain as result message;
 wenzelm parents: 
50499diff
changeset | 455 | val DIALOG = "dialog" | 
| 50503 
50f141b34bb7
enable Isabelle/ML to produce uninterpreted result messages as well;
 wenzelm parents: 
50501diff
changeset | 456 | val Result = new Properties.String(RESULT) | 
| 50500 
c94bba7906d2
identify dialogs via official serial and maintain as result message;
 wenzelm parents: 
50499diff
changeset | 457 | |
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 458 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 459 | /* protocol message functions */ | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 460 | |
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 461 | val FUNCTION = "function" | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 462 | val Function = new Properties.String(FUNCTION) | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 463 | |
| 52563 | 464 | val Assign_Update: Properties.T = List((FUNCTION, "assign_update")) | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 465 | val Removed_Versions: Properties.T = List((FUNCTION, "removed_versions")) | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 466 | |
| 52111 
1fd184eaa310
explicit management of Session.Protocol_Handlers, with protocol state and functions;
 wenzelm parents: 
51818diff
changeset | 467 | object Protocol_Handler | 
| 
1fd184eaa310
explicit management of Session.Protocol_Handlers, with protocol state and functions;
 wenzelm parents: 
51818diff
changeset | 468 |   {
 | 
| 
1fd184eaa310
explicit management of Session.Protocol_Handlers, with protocol state and functions;
 wenzelm parents: 
51818diff
changeset | 469 | def unapply(props: Properties.T): Option[(String)] = | 
| 
1fd184eaa310
explicit management of Session.Protocol_Handlers, with protocol state and functions;
 wenzelm parents: 
51818diff
changeset | 470 |       props match {
 | 
| 
1fd184eaa310
explicit management of Session.Protocol_Handlers, with protocol state and functions;
 wenzelm parents: 
51818diff
changeset | 471 | case List((FUNCTION, "protocol_handler"), (NAME, name)) => Some(name) | 
| 
1fd184eaa310
explicit management of Session.Protocol_Handlers, with protocol state and functions;
 wenzelm parents: 
51818diff
changeset | 472 | case _ => None | 
| 
1fd184eaa310
explicit management of Session.Protocol_Handlers, with protocol state and functions;
 wenzelm parents: 
51818diff
changeset | 473 | } | 
| 
1fd184eaa310
explicit management of Session.Protocol_Handlers, with protocol state and functions;
 wenzelm parents: 
51818diff
changeset | 474 | } | 
| 
1fd184eaa310
explicit management of Session.Protocol_Handlers, with protocol state and functions;
 wenzelm parents: 
51818diff
changeset | 475 | |
| 
1fd184eaa310
explicit management of Session.Protocol_Handlers, with protocol state and functions;
 wenzelm parents: 
51818diff
changeset | 476 | val INVOKE_SCALA = "invoke_scala" | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 477 | object Invoke_Scala | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 478 |   {
 | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 479 | def unapply(props: Properties.T): Option[(String, String)] = | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 480 |       props match {
 | 
| 52111 
1fd184eaa310
explicit management of Session.Protocol_Handlers, with protocol state and functions;
 wenzelm parents: 
51818diff
changeset | 481 | case List((FUNCTION, INVOKE_SCALA), (NAME, name), (ID, id)) => Some((name, id)) | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 482 | case _ => None | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 483 | } | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 484 | } | 
| 52111 
1fd184eaa310
explicit management of Session.Protocol_Handlers, with protocol state and functions;
 wenzelm parents: 
51818diff
changeset | 485 | |
| 
1fd184eaa310
explicit management of Session.Protocol_Handlers, with protocol state and functions;
 wenzelm parents: 
51818diff
changeset | 486 | val CANCEL_SCALA = "cancel_scala" | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 487 | object Cancel_Scala | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 488 |   {
 | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 489 | def unapply(props: Properties.T): Option[String] = | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 490 |       props match {
 | 
| 52111 
1fd184eaa310
explicit management of Session.Protocol_Handlers, with protocol state and functions;
 wenzelm parents: 
51818diff
changeset | 491 | case List((FUNCTION, CANCEL_SCALA), (ID, id)) => Some(id) | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 492 | case _ => None | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 493 | } | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49613diff
changeset | 494 | } | 
| 50255 | 495 | |
| 496 | object ML_Statistics | |
| 497 |   {
 | |
| 498 | def unapply(props: Properties.T): Option[Properties.T] = | |
| 499 |       props match {
 | |
| 500 | case (FUNCTION, "ML_statistics") :: stats => Some(stats) | |
| 501 | case _ => None | |
| 502 | } | |
| 503 | } | |
| 50975 | 504 | |
| 505 | object Task_Statistics | |
| 506 |   {
 | |
| 507 | def unapply(props: Properties.T): Option[Properties.T] = | |
| 508 |       props match {
 | |
| 509 | case (FUNCTION, "task_statistics") :: stats => Some(stats) | |
| 510 | case _ => None | |
| 511 | } | |
| 512 | } | |
| 55316 
885500f4aa6a
interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
 Lars Hupel <lars.hupel@mytum.de> parents: 
55033diff
changeset | 513 | |
| 59367 | 514 | val LOADING_THEORY = "loading_theory" | 
| 56616 
abc2da18d08d
added protocol command "use_theories", with core functionality of batch build;
 wenzelm parents: 
56548diff
changeset | 515 | object Loading_Theory | 
| 
abc2da18d08d
added protocol command "use_theories", with core functionality of batch build;
 wenzelm parents: 
56548diff
changeset | 516 |   {
 | 
| 
abc2da18d08d
added protocol command "use_theories", with core functionality of batch build;
 wenzelm parents: 
56548diff
changeset | 517 | def unapply(props: Properties.T): Option[String] = | 
| 
abc2da18d08d
added protocol command "use_theories", with core functionality of batch build;
 wenzelm parents: 
56548diff
changeset | 518 |       props match {
 | 
| 59367 | 519 | case List((FUNCTION, LOADING_THEORY), (NAME, name)) => Some(name) | 
| 56616 
abc2da18d08d
added protocol command "use_theories", with core functionality of batch build;
 wenzelm parents: 
56548diff
changeset | 520 | case _ => None | 
| 
abc2da18d08d
added protocol command "use_theories", with core functionality of batch build;
 wenzelm parents: 
56548diff
changeset | 521 | } | 
| 
abc2da18d08d
added protocol command "use_theories", with core functionality of batch build;
 wenzelm parents: 
56548diff
changeset | 522 | } | 
| 
abc2da18d08d
added protocol command "use_theories", with core functionality of batch build;
 wenzelm parents: 
56548diff
changeset | 523 | |
| 59366 
e94df7f6b608
clarified build_theories: proper protocol handler;
 wenzelm parents: 
59364diff
changeset | 524 | val BUILD_THEORIES_RESULT = "build_theories_result" | 
| 59364 | 525 | object Build_Theories_Result | 
| 56616 
abc2da18d08d
added protocol command "use_theories", with core functionality of batch build;
 wenzelm parents: 
56548diff
changeset | 526 |   {
 | 
| 59369 
7090199d3f78
more informative build_theories_result: cumulative Runtime.exn_message;
 wenzelm parents: 
59367diff
changeset | 527 | def unapply(props: Properties.T): Option[String] = | 
| 56616 
abc2da18d08d
added protocol command "use_theories", with core functionality of batch build;
 wenzelm parents: 
56548diff
changeset | 528 |       props match {
 | 
| 60831 | 529 | case List((FUNCTION, BUILD_THEORIES_RESULT), (ID, id)) => Some(id) | 
| 56616 
abc2da18d08d
added protocol command "use_theories", with core functionality of batch build;
 wenzelm parents: 
56548diff
changeset | 530 | case _ => None | 
| 
abc2da18d08d
added protocol command "use_theories", with core functionality of batch build;
 wenzelm parents: 
56548diff
changeset | 531 | } | 
| 
abc2da18d08d
added protocol command "use_theories", with core functionality of batch build;
 wenzelm parents: 
56548diff
changeset | 532 | } | 
| 
abc2da18d08d
added protocol command "use_theories", with core functionality of batch build;
 wenzelm parents: 
56548diff
changeset | 533 | |
| 56864 | 534 | val PRINT_OPERATIONS = "print_operations" | 
| 535 | ||
| 55553 
99409ccbe04a
more standard names for protocol and markup elements;
 wenzelm parents: 
55551diff
changeset | 536 | |
| 60830 | 537 | /* debugger output */ | 
| 538 | ||
| 60842 | 539 | val DEBUGGER_STATE = "debugger_state" | 
| 540 | object Debugger_State | |
| 541 |   {
 | |
| 542 | def unapply(props: Properties.T): Option[String] = | |
| 543 |       props match {
 | |
| 544 | case List((FUNCTION, DEBUGGER_STATE), (NAME, name)) => Some(name) | |
| 545 | case _ => None | |
| 546 | } | |
| 547 | } | |
| 548 | ||
| 60830 | 549 | val DEBUGGER_OUTPUT = "debugger_output" | 
| 550 | object Debugger_Output | |
| 551 |   {
 | |
| 60834 | 552 | def unapply(props: Properties.T): Option[String] = | 
| 60830 | 553 |       props match {
 | 
| 60834 | 554 | case List((FUNCTION, DEBUGGER_OUTPUT), (NAME, name)) => Some(name) | 
| 60830 | 555 | case _ => None | 
| 556 | } | |
| 557 | } | |
| 558 | ||
| 559 | ||
| 55316 
885500f4aa6a
interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
 Lars Hupel <lars.hupel@mytum.de> parents: 
55033diff
changeset | 560 | /* simplifier trace */ | 
| 
885500f4aa6a
interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
 Lars Hupel <lars.hupel@mytum.de> parents: 
55033diff
changeset | 561 | |
| 57594 
037f3b251df5
regular message to refer to Simplifier Trace panel (unused);
 wenzelm parents: 
56864diff
changeset | 562 | val SIMP_TRACE_PANEL = "simp_trace_panel" | 
| 55316 
885500f4aa6a
interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
 Lars Hupel <lars.hupel@mytum.de> parents: 
55033diff
changeset | 563 | |
| 55553 
99409ccbe04a
more standard names for protocol and markup elements;
 wenzelm parents: 
55551diff
changeset | 564 | val SIMP_TRACE_LOG = "simp_trace_log" | 
| 
99409ccbe04a
more standard names for protocol and markup elements;
 wenzelm parents: 
55551diff
changeset | 565 | val SIMP_TRACE_STEP = "simp_trace_step" | 
| 
99409ccbe04a
more standard names for protocol and markup elements;
 wenzelm parents: 
55551diff
changeset | 566 | val SIMP_TRACE_RECURSE = "simp_trace_recurse" | 
| 
99409ccbe04a
more standard names for protocol and markup elements;
 wenzelm parents: 
55551diff
changeset | 567 | val SIMP_TRACE_HINT = "simp_trace_hint" | 
| 
99409ccbe04a
more standard names for protocol and markup elements;
 wenzelm parents: 
55551diff
changeset | 568 | val SIMP_TRACE_IGNORE = "simp_trace_ignore" | 
| 55316 
885500f4aa6a
interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
 Lars Hupel <lars.hupel@mytum.de> parents: 
55033diff
changeset | 569 | |
| 55553 
99409ccbe04a
more standard names for protocol and markup elements;
 wenzelm parents: 
55551diff
changeset | 570 | val SIMP_TRACE_CANCEL = "simp_trace_cancel" | 
| 
99409ccbe04a
more standard names for protocol and markup elements;
 wenzelm parents: 
55551diff
changeset | 571 | object Simp_Trace_Cancel | 
| 55316 
885500f4aa6a
interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
 Lars Hupel <lars.hupel@mytum.de> parents: 
55033diff
changeset | 572 |   {
 | 
| 
885500f4aa6a
interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
 Lars Hupel <lars.hupel@mytum.de> parents: 
55033diff
changeset | 573 | def unapply(props: Properties.T): Option[Long] = | 
| 
885500f4aa6a
interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
 Lars Hupel <lars.hupel@mytum.de> parents: 
55033diff
changeset | 574 |       props match {
 | 
| 55553 
99409ccbe04a
more standard names for protocol and markup elements;
 wenzelm parents: 
55551diff
changeset | 575 | case (FUNCTION, SIMP_TRACE_CANCEL) :: Serial(i) => Some(i) | 
| 55316 
885500f4aa6a
interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
 Lars Hupel <lars.hupel@mytum.de> parents: 
55033diff
changeset | 576 | case _ => None | 
| 
885500f4aa6a
interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
 Lars Hupel <lars.hupel@mytum.de> parents: 
55033diff
changeset | 577 | } | 
| 
885500f4aa6a
interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
 Lars Hupel <lars.hupel@mytum.de> parents: 
55033diff
changeset | 578 | } | 
| 45666 | 579 | } | 
| 43721 
fad8634cee62
echo prover input via raw_messages, for improved protocol tracing;
 wenzelm parents: 
43710diff
changeset | 580 | |
| 
fad8634cee62
echo prover input via raw_messages, for improved protocol tracing;
 wenzelm parents: 
43710diff
changeset | 581 | |
| 45666 | 582 | sealed case class Markup(name: String, properties: Properties.T) | 
| 59707 | 583 | {
 | 
| 584 | def markup(s: String): String = | |
| 585 | YXML.string_of_tree(XML.Elem(this, List(XML.Text(s)))) | |
| 586 | } |