| author | desharna | 
| Mon, 06 Oct 2014 13:41:37 +0200 | |
| changeset 58585 | efc8b2c54a38 | 
| parent 57912 | dd9550f84106 | 
| child 59072 | 27c6936c6484 | 
| permissions | -rw-r--r-- | 
| 36676 | 1 | /* Title: Pure/PIDE/command.scala | 
| 2 | Author: Fabian Immler, TU Munich | |
| 3 | Author: Makarius | |
| 4 | ||
| 52536 | 5 | Prover commands with accumulated results from execution. | 
| 36676 | 6 | */ | 
| 34407 | 7 | |
| 34871 
e596a0b71f3c
incorporate "proofdocument" part into main Isabelle/Pure.jar -- except for html_panel.scala, which depends on external library (Lobo/Cobra browser);
 wenzelm parents: 
34865diff
changeset | 8 | package isabelle | 
| 34318 
c13e168a8ae6
original sources from Johannes Hölzl a48e0c6ab1aea77c52d596f7efc007a543d3d10c with minor modifications of directory layout;
 wenzelm parents: diff
changeset | 9 | |
| 34451 | 10 | |
| 45644 | 11 | import scala.collection.mutable | 
| 38872 | 12 | import scala.collection.immutable.SortedMap | 
| 13 | ||
| 14 | ||
| 34637 | 15 | object Command | 
| 16 | {
 | |
| 52849 | 17 | type Edit = (Option[Command], Option[Command]) | 
| 56746 | 18 | type Blob = Exn.Result[(Document.Node.Name, Option[(SHA1.Digest, Symbol.Text_Chunk)])] | 
| 55648 
38f264741609
tuned signature -- avoid obscure default arguments;
 wenzelm parents: 
55622diff
changeset | 19 | |
| 52849 | 20 | |
| 21 | ||
| 38361 | 22 | /** accumulated results from prover **/ | 
| 23 | ||
| 50507 | 24 | /* results */ | 
| 25 | ||
| 26 | object Results | |
| 27 |   {
 | |
| 51496 
cb677987b7e3
retain original tooltip range, to avoid repeated window popup when the mouse is moved over the same content;
 wenzelm parents: 
51494diff
changeset | 28 | type Entry = (Long, XML.Tree) | 
| 50507 | 29 | val empty = new Results(SortedMap.empty) | 
| 56299 
8201790fdeb9
more careful treatment of multiple command states (eval + prints): merge content that is actually required;
 wenzelm parents: 
56295diff
changeset | 30 | def make(es: List[Results.Entry]): Results = (empty /: es)(_ + _) | 
| 
8201790fdeb9
more careful treatment of multiple command states (eval + prints): merge content that is actually required;
 wenzelm parents: 
56295diff
changeset | 31 | def merge(rs: List[Results]): Results = (empty /: rs)(_ ++ _) | 
| 50507 | 32 | } | 
| 33 | ||
| 51494 | 34 | final class Results private(private val rep: SortedMap[Long, XML.Tree]) | 
| 50507 | 35 |   {
 | 
| 36 | def defined(serial: Long): Boolean = rep.isDefinedAt(serial) | |
| 37 | def get(serial: Long): Option[XML.Tree] = rep.get(serial) | |
| 56372 
fadb0fef09d7
more explicit iterator terminology, in accordance to Scala 2.8 library;
 wenzelm parents: 
56359diff
changeset | 38 | def iterator: Iterator[Results.Entry] = rep.iterator | 
| 50508 
5b7150395568
tuned implementation according to Library.insert/merge in ML;
 wenzelm parents: 
50507diff
changeset | 39 | |
| 51496 
cb677987b7e3
retain original tooltip range, to avoid repeated window popup when the mouse is moved over the same content;
 wenzelm parents: 
51494diff
changeset | 40 | def + (entry: Results.Entry): Results = | 
| 50508 
5b7150395568
tuned implementation according to Library.insert/merge in ML;
 wenzelm parents: 
50507diff
changeset | 41 | if (defined(entry._1)) this | 
| 
5b7150395568
tuned implementation according to Library.insert/merge in ML;
 wenzelm parents: 
50507diff
changeset | 42 | else new Results(rep + entry) | 
| 
5b7150395568
tuned implementation according to Library.insert/merge in ML;
 wenzelm parents: 
50507diff
changeset | 43 | |
| 
5b7150395568
tuned implementation according to Library.insert/merge in ML;
 wenzelm parents: 
50507diff
changeset | 44 | def ++ (other: Results): Results = | 
| 
5b7150395568
tuned implementation according to Library.insert/merge in ML;
 wenzelm parents: 
50507diff
changeset | 45 | if (this eq other) this | 
| 
5b7150395568
tuned implementation according to Library.insert/merge in ML;
 wenzelm parents: 
50507diff
changeset | 46 | else if (rep.isEmpty) other | 
| 56372 
fadb0fef09d7
more explicit iterator terminology, in accordance to Scala 2.8 library;
 wenzelm parents: 
56359diff
changeset | 47 | else (this /: other.iterator)(_ + _) | 
| 50540 | 48 | |
| 51494 | 49 | override def hashCode: Int = rep.hashCode | 
| 50 | override def equals(that: Any): Boolean = | |
| 51 |       that match {
 | |
| 52 | case other: Results => rep == other.rep | |
| 53 | case _ => false | |
| 54 | } | |
| 56372 
fadb0fef09d7
more explicit iterator terminology, in accordance to Scala 2.8 library;
 wenzelm parents: 
56359diff
changeset | 55 |     override def toString: String = iterator.mkString("Results(", ", ", ")")
 | 
| 50507 | 56 | } | 
| 57 | ||
| 58 | ||
| 55649 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 59 | /* markup */ | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 60 | |
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 61 | object Markup_Index | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 62 |   {
 | 
| 56746 | 63 | val markup: Markup_Index = Markup_Index(false, Symbol.Text_Chunk.Default) | 
| 55649 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 64 | } | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 65 | |
| 56746 | 66 | sealed case class Markup_Index(status: Boolean, chunk_name: Symbol.Text_Chunk.Name) | 
| 55649 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 67 | |
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 68 | object Markups | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 69 |   {
 | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 70 | val empty: Markups = new Markups(Map.empty) | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 71 | |
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 72 | def init(markup: Markup_Tree): Markups = | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 73 | new Markups(Map(Markup_Index.markup -> markup)) | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 74 | } | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 75 | |
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 76 | final class Markups private(private val rep: Map[Markup_Index, Markup_Tree]) | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 77 |   {
 | 
| 56489 | 78 | def is_empty: Boolean = rep.isEmpty | 
| 79 | ||
| 55649 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 80 | def apply(index: Markup_Index): Markup_Tree = | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 81 | rep.getOrElse(index, Markup_Tree.empty) | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 82 | |
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 83 | def add(index: Markup_Index, markup: Text.Markup): Markups = | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 84 | new Markups(rep + (index -> (this(index) + markup))) | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 85 | |
| 56475 | 86 | def redirection_iterator: Iterator[Document_ID.Generic] = | 
| 56746 | 87 | for (Markup_Index(_, Symbol.Text_Chunk.Id(id)) <- rep.keysIterator) | 
| 56475 | 88 | yield id | 
| 56470 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 wenzelm parents: 
56469diff
changeset | 89 | |
| 56475 | 90 | def redirect(other_id: Document_ID.Generic): Markups = | 
| 56489 | 91 |     {
 | 
| 92 | val rep1 = | |
| 56470 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 wenzelm parents: 
56469diff
changeset | 93 |         (for {
 | 
| 56746 | 94 | (Markup_Index(status, Symbol.Text_Chunk.Id(id)), markup) <- rep.iterator | 
| 56470 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 wenzelm parents: 
56469diff
changeset | 95 | if other_id == id | 
| 56746 | 96 | } yield (Markup_Index(status, Symbol.Text_Chunk.Default), markup)).toMap | 
| 56489 | 97 | if (rep1.isEmpty) Markups.empty else new Markups(rep1) | 
| 98 | } | |
| 56470 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 wenzelm parents: 
56469diff
changeset | 99 | |
| 55649 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 100 | override def hashCode: Int = rep.hashCode | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 101 | override def equals(that: Any): Boolean = | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 102 |       that match {
 | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 103 | case other: Markups => rep == other.rep | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 104 | case _ => false | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 105 | } | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 106 |     override def toString: String = rep.iterator.mkString("Markups(", ", ", ")")
 | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 107 | } | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 108 | |
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 109 | |
| 50507 | 110 | /* state */ | 
| 50501 
6f41f1646617
more careful handling of Dialog_Result, with active area and color feedback;
 wenzelm parents: 
50500diff
changeset | 111 | |
| 56299 
8201790fdeb9
more careful treatment of multiple command states (eval + prints): merge content that is actually required;
 wenzelm parents: 
56295diff
changeset | 112 | object State | 
| 
8201790fdeb9
more careful treatment of multiple command states (eval + prints): merge content that is actually required;
 wenzelm parents: 
56295diff
changeset | 113 |   {
 | 
| 
8201790fdeb9
more careful treatment of multiple command states (eval + prints): merge content that is actually required;
 wenzelm parents: 
56295diff
changeset | 114 | def merge_results(states: List[State]): Command.Results = | 
| 
8201790fdeb9
more careful treatment of multiple command states (eval + prints): merge content that is actually required;
 wenzelm parents: 
56295diff
changeset | 115 | Results.merge(states.map(_.results)) | 
| 
8201790fdeb9
more careful treatment of multiple command states (eval + prints): merge content that is actually required;
 wenzelm parents: 
56295diff
changeset | 116 | |
| 56301 
1da7b4c33db9
more frugal merge of markup trees: filter wrt. subsequent query;
 wenzelm parents: 
56299diff
changeset | 117 | def merge_markup(states: List[State], index: Markup_Index, | 
| 56743 | 118 | range: Text.Range, elements: Markup.Elements): Markup_Tree = | 
| 56301 
1da7b4c33db9
more frugal merge of markup trees: filter wrt. subsequent query;
 wenzelm parents: 
56299diff
changeset | 119 | Markup_Tree.merge(states.map(_.markup(index)), range, elements) | 
| 56299 
8201790fdeb9
more careful treatment of multiple command states (eval + prints): merge content that is actually required;
 wenzelm parents: 
56295diff
changeset | 120 | } | 
| 
8201790fdeb9
more careful treatment of multiple command states (eval + prints): merge content that is actually required;
 wenzelm parents: 
56295diff
changeset | 121 | |
| 43714 | 122 | sealed case class State( | 
| 50501 
6f41f1646617
more careful handling of Dialog_Result, with active area and color feedback;
 wenzelm parents: 
50500diff
changeset | 123 | command: Command, | 
| 
6f41f1646617
more careful handling of Dialog_Result, with active area and color feedback;
 wenzelm parents: 
50500diff
changeset | 124 | status: List[Markup] = Nil, | 
| 50507 | 125 | results: Results = Results.empty, | 
| 55649 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 126 | markups: Markups = Markups.empty) | 
| 38361 | 127 |   {
 | 
| 56395 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 wenzelm parents: 
56372diff
changeset | 128 | lazy val protocol_status: Protocol.Status = | 
| 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 wenzelm parents: 
56372diff
changeset | 129 |     {
 | 
| 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 wenzelm parents: 
56372diff
changeset | 130 | val warnings = | 
| 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 wenzelm parents: 
56372diff
changeset | 131 | if (results.iterator.exists(p => Protocol.is_warning(p._2))) | 
| 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 wenzelm parents: 
56372diff
changeset | 132 | List(Markup(Markup.WARNING, Nil)) | 
| 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 wenzelm parents: 
56372diff
changeset | 133 | else Nil | 
| 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 wenzelm parents: 
56372diff
changeset | 134 | val errors = | 
| 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 wenzelm parents: 
56372diff
changeset | 135 | if (results.iterator.exists(p => Protocol.is_error(p._2))) | 
| 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 wenzelm parents: 
56372diff
changeset | 136 | List(Markup(Markup.ERROR, Nil)) | 
| 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 wenzelm parents: 
56372diff
changeset | 137 | else Nil | 
| 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 wenzelm parents: 
56372diff
changeset | 138 | Protocol.Status.make((warnings ::: errors ::: status).iterator) | 
| 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 wenzelm parents: 
56372diff
changeset | 139 | } | 
| 55432 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 wenzelm parents: 
55431diff
changeset | 140 | |
| 55650 | 141 | def markup(index: Markup_Index): Markup_Tree = markups(index) | 
| 55432 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 wenzelm parents: 
55431diff
changeset | 142 | |
| 56489 | 143 | def redirect(other_command: Command): Option[State] = | 
| 144 |     {
 | |
| 145 | val markups1 = markups.redirect(other_command.id) | |
| 146 | if (markups1.is_empty) None | |
| 147 | else Some(new State(other_command, Nil, Results.empty, markups1)) | |
| 148 | } | |
| 49614 | 149 | |
| 51494 | 150 | def eq_content(other: State): Boolean = | 
| 151 | command.source == other.command.source && | |
| 152 | status == other.status && | |
| 153 | results == other.results && | |
| 55434 | 154 | markups == other.markups | 
| 38361 | 155 | |
| 55649 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 156 | private def add_status(st: Markup): State = | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 157 | copy(status = st :: status) | 
| 55432 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 wenzelm parents: 
55431diff
changeset | 158 | |
| 56746 | 159 | private def add_markup( | 
| 160 | status: Boolean, chunk_name: Symbol.Text_Chunk.Name, m: Text.Markup): State = | |
| 55649 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 161 |     {
 | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 162 | val markups1 = | 
| 56395 
0546e036d1c0
more direct warning within persistent Protocol.Status;
 wenzelm parents: 
56372diff
changeset | 163 | if (status || Protocol.liberal_status_elements(m.info.name)) | 
| 56462 
b64b0cb845fe
more explicit Command.Chunk types, less ooddities;
 wenzelm parents: 
56395diff
changeset | 164 | markups.add(Markup_Index(true, chunk_name), m) | 
| 55649 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 165 | else markups | 
| 56462 
b64b0cb845fe
more explicit Command.Chunk types, less ooddities;
 wenzelm parents: 
56395diff
changeset | 166 | copy(markups = markups1.add(Markup_Index(false, chunk_name), m)) | 
| 55649 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 167 | } | 
| 38361 | 168 | |
| 56470 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 wenzelm parents: 
56469diff
changeset | 169 | def accumulate( | 
| 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 wenzelm parents: 
56469diff
changeset | 170 | self_id: Document_ID.Generic => Boolean, | 
| 56746 | 171 | other_id: Document_ID.Generic => Option[(Symbol.Text_Chunk.Id, Symbol.Text_Chunk)], | 
| 56470 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 wenzelm parents: 
56469diff
changeset | 172 | message: XML.Elem): State = | 
| 38361 | 173 |       message match {
 | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
50163diff
changeset | 174 | case XML.Elem(Markup(Markup.STATUS, _), msgs) => | 
| 38714 | 175 | (this /: msgs)((state, msg) => | 
| 176 |             msg match {
 | |
| 46152 
793cecd4ffc0
accumulate status as regular markup for command range;
 wenzelm parents: 
45709diff
changeset | 177 | case elem @ XML.Elem(markup, Nil) => | 
| 55649 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 178 | state. | 
| 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 179 | add_status(markup). | 
| 56746 | 180 | add_markup(true, Symbol.Text_Chunk.Default, Text.Info(command.proper_range, elem)) | 
| 52536 | 181 | case _ => | 
| 56782 
433cf57550fa
more systematic Isabelle output, like in classic Isabelle/ML (without markup);
 wenzelm parents: 
56746diff
changeset | 182 |                 Output.warning("Ignored status message: " + msg)
 | 
| 52536 | 183 | state | 
| 38714 | 184 | }) | 
| 38581 
d503a0912e14
simplified Command.status again, reverting most of e5eed57913d0 (note that more complex information can be represented with full markup reports);
 wenzelm parents: 
38579diff
changeset | 185 | |
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
50163diff
changeset | 186 | case XML.Elem(Markup(Markup.REPORT, _), msgs) => | 
| 38572 
0fe2c01ef7da
Command.State: accumulate markup reports uniformly;
 wenzelm parents: 
38564diff
changeset | 187 | (this /: msgs)((state, msg) => | 
| 55432 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 wenzelm parents: 
55431diff
changeset | 188 |             {
 | 
| 56782 
433cf57550fa
more systematic Isabelle output, like in classic Isabelle/ML (without markup);
 wenzelm parents: 
56746diff
changeset | 189 |               def bad(): Unit = Output.warning("Ignored report message: " + msg)
 | 
| 55432 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 wenzelm parents: 
55431diff
changeset | 190 | |
| 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 wenzelm parents: 
55431diff
changeset | 191 |               msg match {
 | 
| 57911 
dcb758188aa6
clarified Position.Identified: do not require range from prover, default to command position;
 wenzelm parents: 
57910diff
changeset | 192 | case XML.Elem(Markup(name, atts @ Position.Identified(id, chunk_name)), args) => | 
| 56470 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 wenzelm parents: 
56469diff
changeset | 193 | |
| 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 wenzelm parents: 
56469diff
changeset | 194 | val target = | 
| 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 wenzelm parents: 
56469diff
changeset | 195 | if (self_id(id) && command.chunks.isDefinedAt(chunk_name)) | 
| 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 wenzelm parents: 
56469diff
changeset | 196 | Some((chunk_name, command.chunks(chunk_name))) | 
| 56746 | 197 | else if (chunk_name == Symbol.Text_Chunk.Default) other_id(id) | 
| 56470 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 wenzelm parents: 
56469diff
changeset | 198 | else None | 
| 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 wenzelm parents: 
56469diff
changeset | 199 | |
| 57911 
dcb758188aa6
clarified Position.Identified: do not require range from prover, default to command position;
 wenzelm parents: 
57910diff
changeset | 200 |                   (target, atts) match {
 | 
| 
dcb758188aa6
clarified Position.Identified: do not require range from prover, default to command position;
 wenzelm parents: 
57910diff
changeset | 201 | case (Some((target_name, target_chunk)), Position.Range(symbol_range)) => | 
| 56470 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 wenzelm parents: 
56469diff
changeset | 202 |                       target_chunk.incorporate(symbol_range) match {
 | 
| 55548 
a645277885cf
more uniform/robust restriction of reported positions, e.g. relevant for "bad" markup due to unclosed comment in ML file;
 wenzelm parents: 
55434diff
changeset | 203 | case Some(range) => | 
| 55822 
ccf2d784be97
incorporate chunk range that is 1 off end-of-input, for improved error positions (NB: command spans are tight, without trailing whitespace);
 wenzelm parents: 
55785diff
changeset | 204 | val props = Position.purge(atts) | 
| 
ccf2d784be97
incorporate chunk range that is 1 off end-of-input, for improved error positions (NB: command spans are tight, without trailing whitespace);
 wenzelm parents: 
55785diff
changeset | 205 | val info = Text.Info(range, XML.Elem(Markup(name, props), args)) | 
| 56470 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 wenzelm parents: 
56469diff
changeset | 206 | state.add_markup(false, target_name, info) | 
| 55548 
a645277885cf
more uniform/robust restriction of reported positions, e.g. relevant for "bad" markup due to unclosed comment in ML file;
 wenzelm parents: 
55434diff
changeset | 207 | case None => bad(); state | 
| 55432 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 wenzelm parents: 
55431diff
changeset | 208 | } | 
| 57911 
dcb758188aa6
clarified Position.Identified: do not require range from prover, default to command position;
 wenzelm parents: 
57910diff
changeset | 209 | case _ => | 
| 56514 
db929027e701
ignore other_id reports for now (see 8eda56033203): massive amounts of redirections to 'class' etc. makes it difficult to edit main HOL;
 wenzelm parents: 
56489diff
changeset | 210 | // silently ignore excessive reports | 
| 56476 | 211 | state | 
| 55432 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 wenzelm parents: 
55431diff
changeset | 212 | } | 
| 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 wenzelm parents: 
55431diff
changeset | 213 | |
| 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 wenzelm parents: 
55431diff
changeset | 214 | case XML.Elem(Markup(name, atts), args) | 
| 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 wenzelm parents: 
55431diff
changeset | 215 |                 if !atts.exists({ case (a, _) => Markup.POSITION_PROPERTIES(a) }) =>
 | 
| 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 wenzelm parents: 
55431diff
changeset | 216 | val range = command.proper_range | 
| 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 wenzelm parents: 
55431diff
changeset | 217 | val props = Position.purge(atts) | 
| 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 wenzelm parents: 
55431diff
changeset | 218 | val info: Text.Markup = Text.Info(range, XML.Elem(Markup(name, props), args)) | 
| 56746 | 219 | state.add_markup(false, Symbol.Text_Chunk.Default, info) | 
| 55432 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 wenzelm parents: 
55431diff
changeset | 220 | |
| 56470 
8eda56033203
accumulate markup reports for "other" command ids, which are later retargeted and merged for rendering (in erratic order);
 wenzelm parents: 
56469diff
changeset | 221 | case _ => bad(); state | 
| 55432 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 wenzelm parents: 
55431diff
changeset | 222 | } | 
| 38361 | 223 | }) | 
| 52930 
5fab62ae3532
retain original messages properties, e.g. for retrieval via Command.Results;
 wenzelm parents: 
52849diff
changeset | 224 | case XML.Elem(Markup(name, props), body) => | 
| 
5fab62ae3532
retain original messages properties, e.g. for retrieval via Command.Results;
 wenzelm parents: 
52849diff
changeset | 225 |           props match {
 | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
50163diff
changeset | 226 | case Markup.Serial(i) => | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
50163diff
changeset | 227 | val message1 = XML.Elem(Markup(Markup.message(name), props), body) | 
| 50163 
c62ce309dc26
more abstract Sendback operations, with explicit id/exec_id properties;
 wenzelm parents: 
50158diff
changeset | 228 | val message2 = XML.Elem(Markup(name, props), body) | 
| 
c62ce309dc26
more abstract Sendback operations, with explicit id/exec_id properties;
 wenzelm parents: 
50158diff
changeset | 229 | |
| 55433 
d2960d67f163
clarified message_positions: cover alt_id as well;
 wenzelm parents: 
55432diff
changeset | 230 | var st = copy(results = results + (i -> message1)) | 
| 
d2960d67f163
clarified message_positions: cover alt_id as well;
 wenzelm parents: 
55432diff
changeset | 231 |               if (Protocol.is_inlined(message)) {
 | 
| 
d2960d67f163
clarified message_positions: cover alt_id as well;
 wenzelm parents: 
55432diff
changeset | 232 |                 for {
 | 
| 56469 | 233 | (chunk_name, chunk) <- command.chunks.iterator | 
| 57911 
dcb758188aa6
clarified Position.Identified: do not require range from prover, default to command position;
 wenzelm parents: 
57910diff
changeset | 234 | range <- Protocol.message_positions( | 
| 
dcb758188aa6
clarified Position.Identified: do not require range from prover, default to command position;
 wenzelm parents: 
57910diff
changeset | 235 | self_id, command.position, chunk_name, chunk, message) | 
| 56462 
b64b0cb845fe
more explicit Command.Chunk types, less ooddities;
 wenzelm parents: 
56395diff
changeset | 236 | } st = st.add_markup(false, chunk_name, Text.Info(range, message2)) | 
| 55433 
d2960d67f163
clarified message_positions: cover alt_id as well;
 wenzelm parents: 
55432diff
changeset | 237 | } | 
| 
d2960d67f163
clarified message_positions: cover alt_id as well;
 wenzelm parents: 
55432diff
changeset | 238 | st | 
| 55432 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 wenzelm parents: 
55431diff
changeset | 239 | |
| 52536 | 240 | case _ => | 
| 56782 
433cf57550fa
more systematic Isabelle output, like in classic Isabelle/ML (without markup);
 wenzelm parents: 
56746diff
changeset | 241 |               Output.warning("Ignored message without serial number: " + message)
 | 
| 52536 | 242 | this | 
| 38872 | 243 | } | 
| 56295 
a40e67ce4f84
clarified valid_id: always standardize towards static command.id;
 wenzelm parents: 
55884diff
changeset | 244 | } | 
| 38361 | 245 | } | 
| 38367 | 246 | |
| 247 | ||
| 55431 
e0f20a44ff9d
common Command.Chunk for command source and auxiliary files (static Symbol.Index without actual String content);
 wenzelm parents: 
55430diff
changeset | 248 | |
| 
e0f20a44ff9d
common Command.Chunk for command source and auxiliary files (static Symbol.Index without actual String content);
 wenzelm parents: 
55430diff
changeset | 249 | /** static content **/ | 
| 
e0f20a44ff9d
common Command.Chunk for command source and auxiliary files (static Symbol.Index without actual String content);
 wenzelm parents: 
55430diff
changeset | 250 | |
| 45644 | 251 | /* make commands */ | 
| 252 | ||
| 55648 
38f264741609
tuned signature -- avoid obscure default arguments;
 wenzelm parents: 
55622diff
changeset | 253 | def apply( | 
| 
38f264741609
tuned signature -- avoid obscure default arguments;
 wenzelm parents: 
55622diff
changeset | 254 | id: Document_ID.Command, | 
| 
38f264741609
tuned signature -- avoid obscure default arguments;
 wenzelm parents: 
55622diff
changeset | 255 | node_name: Document.Node.Name, | 
| 
38f264741609
tuned signature -- avoid obscure default arguments;
 wenzelm parents: 
55622diff
changeset | 256 | blobs: List[Blob], | 
| 57905 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57904diff
changeset | 257 | span: Command_Span.Span): Command = | 
| 55648 
38f264741609
tuned signature -- avoid obscure default arguments;
 wenzelm parents: 
55622diff
changeset | 258 |   {
 | 
| 57901 
e1abca2527da
more explicit type Span in Scala, according to ML version;
 wenzelm parents: 
57842diff
changeset | 259 | val (source, span1) = span.compact_source | 
| 55648 
38f264741609
tuned signature -- avoid obscure default arguments;
 wenzelm parents: 
55622diff
changeset | 260 | new Command(id, node_name, blobs, span1, source, Results.empty, Markup_Tree.empty) | 
| 
38f264741609
tuned signature -- avoid obscure default arguments;
 wenzelm parents: 
55622diff
changeset | 261 | } | 
| 49414 | 262 | |
| 57901 
e1abca2527da
more explicit type Span in Scala, according to ML version;
 wenzelm parents: 
57842diff
changeset | 263 | val empty: Command = | 
| 57905 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57904diff
changeset | 264 | Command(Document_ID.none, Document.Node.Name.empty, Nil, Command_Span.empty) | 
| 55648 
38f264741609
tuned signature -- avoid obscure default arguments;
 wenzelm parents: 
55622diff
changeset | 265 | |
| 
38f264741609
tuned signature -- avoid obscure default arguments;
 wenzelm parents: 
55622diff
changeset | 266 | def unparsed( | 
| 
38f264741609
tuned signature -- avoid obscure default arguments;
 wenzelm parents: 
55622diff
changeset | 267 | id: Document_ID.Command, | 
| 
38f264741609
tuned signature -- avoid obscure default arguments;
 wenzelm parents: 
55622diff
changeset | 268 | source: String, | 
| 
38f264741609
tuned signature -- avoid obscure default arguments;
 wenzelm parents: 
55622diff
changeset | 269 | results: Results, | 
| 
38f264741609
tuned signature -- avoid obscure default arguments;
 wenzelm parents: 
55622diff
changeset | 270 | markup: Markup_Tree): Command = | 
| 
38f264741609
tuned signature -- avoid obscure default arguments;
 wenzelm parents: 
55622diff
changeset | 271 |   {
 | 
| 57905 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57904diff
changeset | 272 | val (source1, span1) = Command_Span.unparsed(source).compact_source | 
| 57901 
e1abca2527da
more explicit type Span in Scala, according to ML version;
 wenzelm parents: 
57842diff
changeset | 273 | new Command(id, Document.Node.Name.empty, Nil, span1, source1, results, markup) | 
| 55648 
38f264741609
tuned signature -- avoid obscure default arguments;
 wenzelm parents: 
55622diff
changeset | 274 | } | 
| 49414 | 275 | |
| 50501 
6f41f1646617
more careful handling of Dialog_Result, with active area and color feedback;
 wenzelm parents: 
50500diff
changeset | 276 | def unparsed(source: String): Command = | 
| 52530 
99dd8b4ef3fe
explicit module Document_ID as source of globally unique identifiers across ML/Scala;
 wenzelm parents: 
52527diff
changeset | 277 | unparsed(Document_ID.none, source, Results.empty, Markup_Tree.empty) | 
| 44384 | 278 | |
| 52530 
99dd8b4ef3fe
explicit module Document_ID as source of globally unique identifiers across ML/Scala;
 wenzelm parents: 
52527diff
changeset | 279 | def rich_text(id: Document_ID.Command, results: Results, body: XML.Body): Command = | 
| 49414 | 280 |   {
 | 
| 49466 | 281 | val text = XML.content(body) | 
| 282 | val markup = Markup_Tree.from_XML(body) | |
| 50501 
6f41f1646617
more careful handling of Dialog_Result, with active area and color feedback;
 wenzelm parents: 
50500diff
changeset | 283 | unparsed(id, text, results, markup) | 
| 49414 | 284 | } | 
| 49359 
c1262d7389fb
refined output panel: more value-oriented approach to update and caret focus;
 wenzelm parents: 
49037diff
changeset | 285 | |
| 44384 | 286 | |
| 287 | /* perspective */ | |
| 288 | ||
| 44474 | 289 | object Perspective | 
| 290 |   {
 | |
| 291 | val empty: Perspective = Perspective(Nil) | |
| 292 | } | |
| 44385 
e7fdb008aa7d
propagate editor perspective through document model;
 wenzelm parents: 
44384diff
changeset | 293 | |
| 44474 | 294 | sealed case class Perspective(commands: List[Command]) // visible commands in canonical order | 
| 44385 
e7fdb008aa7d
propagate editor perspective through document model;
 wenzelm parents: 
44384diff
changeset | 295 |   {
 | 
| 57615 
df1b3452d71c
more explicit discrimination of empty nodes -- suppress from Theories panel;
 wenzelm parents: 
56782diff
changeset | 296 | def is_empty: Boolean = commands.isEmpty | 
| 
df1b3452d71c
more explicit discrimination of empty nodes -- suppress from Theories panel;
 wenzelm parents: 
56782diff
changeset | 297 | |
| 44474 | 298 | def same(that: Perspective): Boolean = | 
| 299 |     {
 | |
| 300 | val cmds1 = this.commands | |
| 301 | val cmds2 = that.commands | |
| 48754 
c2c1e5944536
clarified undefined, unparsed, unfinished command spans;
 wenzelm parents: 
48745diff
changeset | 302 | require(!cmds1.exists(_.is_undefined)) | 
| 
c2c1e5944536
clarified undefined, unparsed, unfinished command spans;
 wenzelm parents: 
48745diff
changeset | 303 | require(!cmds2.exists(_.is_undefined)) | 
| 44474 | 304 | cmds1.length == cmds2.length && | 
| 305 |         (cmds1.iterator zip cmds2.iterator).forall({ case (c1, c2) => c1.id == c2.id })
 | |
| 306 | } | |
| 44385 
e7fdb008aa7d
propagate editor perspective through document model;
 wenzelm parents: 
44384diff
changeset | 307 | } | 
| 34318 
c13e168a8ae6
original sources from Johannes Hölzl a48e0c6ab1aea77c52d596f7efc007a543d3d10c with minor modifications of directory layout;
 wenzelm parents: diff
changeset | 308 | } | 
| 
c13e168a8ae6
original sources from Johannes Hölzl a48e0c6ab1aea77c52d596f7efc007a543d3d10c with minor modifications of directory layout;
 wenzelm parents: diff
changeset | 309 | |
| 38361 | 310 | |
| 46712 | 311 | final class Command private( | 
| 52530 
99dd8b4ef3fe
explicit module Document_ID as source of globally unique identifiers across ML/Scala;
 wenzelm parents: 
52527diff
changeset | 312 | val id: Document_ID.Command, | 
| 44615 | 313 | val node_name: Document.Node.Name, | 
| 54519 
5fed81762406
maintain blobs within document state: digest + text in ML, digest-only in Scala;
 wenzelm parents: 
54517diff
changeset | 314 | val blobs: List[Command.Blob], | 
| 57905 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57904diff
changeset | 315 | val span: Command_Span.Span, | 
| 49414 | 316 | val source: String, | 
| 50501 
6f41f1646617
more careful handling of Dialog_Result, with active area and color feedback;
 wenzelm parents: 
50500diff
changeset | 317 | val init_results: Command.Results, | 
| 49414 | 318 | val init_markup: Markup_Tree) | 
| 34451 | 319 | {
 | 
| 57910 | 320 | /* name */ | 
| 34859 | 321 | |
| 57901 
e1abca2527da
more explicit type Span in Scala, according to ML version;
 wenzelm parents: 
57842diff
changeset | 322 | def name: String = | 
| 57910 | 323 |     span.kind match { case Command_Span.Command_Span(name, _) => name case _ => "" }
 | 
| 324 | ||
| 325 | def position: Position.T = | |
| 326 |     span.kind match { case Command_Span.Command_Span(_, pos) => pos case _ => Position.none }
 | |
| 47012 
0e246130486b
clarified command span classification: strict Command.is_command, permissive Command.name;
 wenzelm parents: 
46910diff
changeset | 327 | |
| 57912 | 328 | override def toString: String = id + "/" + span.kind.toString | 
| 34495 | 329 | |
| 57910 | 330 | |
| 331 | /* classification */ | |
| 332 | ||
| 57905 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57904diff
changeset | 333 | def is_proper: Boolean = span.kind.isInstanceOf[Command_Span.Command_Span] | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57904diff
changeset | 334 | def is_ignored: Boolean = span.kind == Command_Span.Ignored_Span | 
| 57904 | 335 | |
| 336 | def is_undefined: Boolean = id == Document_ID.none | |
| 337 | val is_unparsed: Boolean = span.content.exists(_.is_unparsed) | |
| 338 | val is_unfinished: Boolean = span.content.exists(_.is_unfinished) | |
| 339 | ||
| 34859 | 340 | |
| 54519 
5fed81762406
maintain blobs within document state: digest + text in ML, digest-only in Scala;
 wenzelm parents: 
54517diff
changeset | 341 | /* blobs */ | 
| 
5fed81762406
maintain blobs within document state: digest + text in ML, digest-only in Scala;
 wenzelm parents: 
54517diff
changeset | 342 | |
| 54530 
2c1440f70028
ranges of thy_load commands count as visible within perspective;
 wenzelm parents: 
54524diff
changeset | 343 | def blobs_names: List[Document.Node.Name] = | 
| 
2c1440f70028
ranges of thy_load commands count as visible within perspective;
 wenzelm parents: 
54524diff
changeset | 344 | for (Exn.Res((name, _)) <- blobs) yield name | 
| 
2c1440f70028
ranges of thy_load commands count as visible within perspective;
 wenzelm parents: 
54524diff
changeset | 345 | |
| 57842 
8e4ae2db1849
more direct access to persistent blobs (see also 8953d4cc060a), avoiding fragile digest lookup from later version (which might have removed unused blobs already);
 wenzelm parents: 
57615diff
changeset | 346 | def blobs_defined: List[(Document.Node.Name, SHA1.Digest)] = | 
| 
8e4ae2db1849
more direct access to persistent blobs (see also 8953d4cc060a), avoiding fragile digest lookup from later version (which might have removed unused blobs already);
 wenzelm parents: 
57615diff
changeset | 347 | for (Exn.Res((name, Some((digest, _)))) <- blobs) yield (name, digest) | 
| 
8e4ae2db1849
more direct access to persistent blobs (see also 8953d4cc060a), avoiding fragile digest lookup from later version (which might have removed unused blobs already);
 wenzelm parents: 
57615diff
changeset | 348 | |
| 
8e4ae2db1849
more direct access to persistent blobs (see also 8953d4cc060a), avoiding fragile digest lookup from later version (which might have removed unused blobs already);
 wenzelm parents: 
57615diff
changeset | 349 | def blobs_changed(doc_blobs: Document.Blobs): Boolean = | 
| 
8e4ae2db1849
more direct access to persistent blobs (see also 8953d4cc060a), avoiding fragile digest lookup from later version (which might have removed unused blobs already);
 wenzelm parents: 
57615diff
changeset | 350 |     blobs.exists({ case Exn.Res((name, _)) => doc_blobs.changed(name) case _ => false })
 | 
| 54519 
5fed81762406
maintain blobs within document state: digest + text in ML, digest-only in Scala;
 wenzelm parents: 
54517diff
changeset | 351 | |
| 55432 
9c53198dbb1c
maintain multiple command chunks and markup trees: for main chunk and loaded files;
 wenzelm parents: 
55431diff
changeset | 352 | |
| 56462 
b64b0cb845fe
more explicit Command.Chunk types, less ooddities;
 wenzelm parents: 
56395diff
changeset | 353 | /* source chunks */ | 
| 55431 
e0f20a44ff9d
common Command.Chunk for command source and auxiliary files (static Symbol.Index without actual String content);
 wenzelm parents: 
55430diff
changeset | 354 | |
| 56746 | 355 | val chunk: Symbol.Text_Chunk = Symbol.Text_Chunk(source) | 
| 56473 | 356 | |
| 56746 | 357 | val chunks: Map[Symbol.Text_Chunk.Name, Symbol.Text_Chunk] = | 
| 358 | ((Symbol.Text_Chunk.Default -> chunk) :: | |
| 56473 | 359 | (for (Exn.Res((name, Some((_, file)))) <- blobs) | 
| 56746 | 360 | yield (Symbol.Text_Chunk.File(name.node) -> file))).toMap | 
| 56473 | 361 | |
| 46813 | 362 | def length: Int = source.length | 
| 56473 | 363 | def range: Text.Range = chunk.range | 
| 46813 | 364 | |
| 365 | val proper_range: Text.Range = | |
| 57901 
e1abca2527da
more explicit type Span in Scala, according to ML version;
 wenzelm parents: 
57842diff
changeset | 366 | Text.Range(0, | 
| 
e1abca2527da
more explicit type Span in Scala, according to ML version;
 wenzelm parents: 
57842diff
changeset | 367 | (length /: span.content.reverse.iterator.takeWhile(_.is_improper))(_ - _.source.length)) | 
| 46813 | 368 | |
| 38426 | 369 | def source(range: Text.Range): String = source.substring(range.start, range.stop) | 
| 38572 
0fe2c01ef7da
Command.State: accumulate markup reports uniformly;
 wenzelm parents: 
38564diff
changeset | 370 | |
| 38370 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 371 | |
| 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 372 | /* accumulated results */ | 
| 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 373 | |
| 50501 
6f41f1646617
more careful handling of Dialog_Result, with active area and color feedback;
 wenzelm parents: 
50500diff
changeset | 374 | val init_state: Command.State = | 
| 55649 
1532ab0dc67b
more general / abstract Command.Markups, with separate index for status elements;
 wenzelm parents: 
55648diff
changeset | 375 | Command.State(this, results = init_results, markups = Command.Markups.init(init_markup)) | 
| 52527 
dbac84eab3bc
separate exec_id assignment for Command.print states, without affecting result of eval;
 wenzelm parents: 
52524diff
changeset | 376 | |
| 
dbac84eab3bc
separate exec_id assignment for Command.print states, without affecting result of eval;
 wenzelm parents: 
52524diff
changeset | 377 | val empty_state: Command.State = Command.State(this) | 
| 34676 
9e725d34df7b
Command and Command_State handle results from prover as Accumulator
 immler@in.tum.de parents: 
34675diff
changeset | 378 | } |