| author | wenzelm | 
| Wed, 21 Mar 2012 23:26:35 +0100 | |
| changeset 47069 | 451fc10a81f3 | 
| parent 46997 | 395b7277ed76 | 
| child 47388 | fe4b245af74c | 
| permissions | -rw-r--r-- | 
| 36676 | 1 | /* Title: Pure/PIDE/document.scala | 
| 2 | Author: Makarius | |
| 3 | ||
| 38150 
67fc24df3721
simplified/refined document model: collection of named nodes, without proper dependencies yet;
 wenzelm parents: 
37849diff
changeset | 4 | Document as collection of named nodes, each consisting of an editable | 
| 38418 
9a7af64d71bb
more explicit / functional ML version of document model;
 wenzelm parents: 
38417diff
changeset | 5 | list of commands, associated with asynchronous execution process. | 
| 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: 
34868diff
changeset | 8 | package isabelle | 
| 34318 
c13e168a8ae6
original sources from Johannes Hölzl a48e0c6ab1aea77c52d596f7efc007a543d3d10c with minor modifications of directory layout;
 wenzelm parents: diff
changeset | 9 | |
| 34760 | 10 | |
| 38150 
67fc24df3721
simplified/refined document model: collection of named nodes, without proper dependencies yet;
 wenzelm parents: 
37849diff
changeset | 11 | import scala.collection.mutable | 
| 37073 | 12 | |
| 13 | ||
| 34823 | 14 | object Document | 
| 34483 | 15 | {
 | 
| 38227 
6bbb42843b6e
concentrate structural document notions in document.scala;
 wenzelm parents: 
38220diff
changeset | 16 | /* formal identifiers */ | 
| 38150 
67fc24df3721
simplified/refined document model: collection of named nodes, without proper dependencies yet;
 wenzelm parents: 
37849diff
changeset | 17 | |
| 38363 
af7f41a8a0a8
clarified "state" (accumulated data) vs. "exec" (execution that produces data);
 wenzelm parents: 
38361diff
changeset | 18 | type ID = Long | 
| 43780 | 19 | val ID = Properties.Value.Long | 
| 38414 
49f1f657adc2
more basic Markup.parse_int/print_int (using signed_string_of_int) (ML);
 wenzelm parents: 
38374diff
changeset | 20 | |
| 38373 | 21 | type Version_ID = ID | 
| 38363 
af7f41a8a0a8
clarified "state" (accumulated data) vs. "exec" (execution that produces data);
 wenzelm parents: 
38361diff
changeset | 22 | type Command_ID = ID | 
| 38373 | 23 | type Exec_ID = ID | 
| 38150 
67fc24df3721
simplified/refined document model: collection of named nodes, without proper dependencies yet;
 wenzelm parents: 
37849diff
changeset | 24 | |
| 43662 
e3175ec00311
Document.no_id/new_id as in ML (new_id *could* be session-specific but it isn't right now);
 wenzelm parents: 
43427diff
changeset | 25 | val no_id: ID = 0 | 
| 45243 | 26 | val new_id = Counter() | 
| 38355 
8cb265fb12fe
represent document ids by (long) int, to benefit from the somewhat faster Inttab in ML (LinearSet in Scala is invariably indexed by native object ids);
 wenzelm parents: 
38227diff
changeset | 27 | |
| 38150 
67fc24df3721
simplified/refined document model: collection of named nodes, without proper dependencies yet;
 wenzelm parents: 
37849diff
changeset | 28 | |
| 34818 
7df68a8f0e3e
register Proof_Document instances as session entities -- handle Markup.EDIT messages locally;
 wenzelm parents: 
34815diff
changeset | 29 | |
| 38424 | 30 | /** document structure **/ | 
| 31 | ||
| 46723 | 32 | /* individual nodes */ | 
| 38424 | 33 | |
| 44615 | 34 | type Edit[A, B] = (Node.Name, Node.Edit[A, B]) | 
| 44384 | 35 | type Edit_Text = Edit[Text.Edit, Text.Perspective] | 
| 36 | type Edit_Command = Edit[(Option[Command], Option[Command]), Command.Perspective] | |
| 38151 | 37 | |
| 46737 | 38 | type Node_Header = Exn.Result[Node.Deps] | 
| 44182 | 39 | |
| 38151 | 40 | object Node | 
| 41 |   {
 | |
| 46938 
cda018294515
some support for outer syntax keyword declarations within theory header;
 wenzelm parents: 
46814diff
changeset | 42 | sealed case class Deps( | 
| 
cda018294515
some support for outer syntax keyword declarations within theory header;
 wenzelm parents: 
46814diff
changeset | 43 | imports: List[Name], | 
| 46940 | 44 | keywords: List[Outer_Syntax.Decl], | 
| 46938 
cda018294515
some support for outer syntax keyword declarations within theory header;
 wenzelm parents: 
46814diff
changeset | 45 | uses: List[(String, Boolean)]) | 
| 46737 | 46 | |
| 44957 
098dd95349e7
more elaborate Node_Renderer, which paints node_name.theory only;
 wenzelm parents: 
44676diff
changeset | 47 | object Name | 
| 
098dd95349e7
more elaborate Node_Renderer, which paints node_name.theory only;
 wenzelm parents: 
44676diff
changeset | 48 |     {
 | 
| 
098dd95349e7
more elaborate Node_Renderer, which paints node_name.theory only;
 wenzelm parents: 
44676diff
changeset | 49 |       val empty = Name("", "", "")
 | 
| 44959 | 50 | def apply(path: Path): Name = | 
| 51 |       {
 | |
| 52 | val node = path.implode | |
| 53 | val dir = path.dir.implode | |
| 54 |         val theory = Thy_Header.thy_name(node) getOrElse error("Bad theory file name: " + path)
 | |
| 55 | Name(node, dir, theory) | |
| 56 | } | |
| 46723 | 57 | |
| 58 | object Ordering extends scala.math.Ordering[Name] | |
| 59 |       {
 | |
| 60 | def compare(name1: Name, name2: Name): Int = name1.node compare name2.node | |
| 61 | } | |
| 44957 
098dd95349e7
more elaborate Node_Renderer, which paints node_name.theory only;
 wenzelm parents: 
44676diff
changeset | 62 | } | 
| 44616 | 63 | sealed case class Name(node: String, dir: String, theory: String) | 
| 44615 | 64 |     {
 | 
| 65 | override def hashCode: Int = node.hashCode | |
| 66 | override def equals(that: Any): Boolean = | |
| 67 |         that match {
 | |
| 68 | case other: Name => node == other.node | |
| 69 | case _ => false | |
| 70 | } | |
| 44642 | 71 | override def toString: String = theory | 
| 44615 | 72 | } | 
| 73 | ||
| 44384 | 74 | sealed abstract class Edit[A, B] | 
| 44156 | 75 |     {
 | 
| 44383 | 76 | def foreach(f: A => Unit) | 
| 77 |       {
 | |
| 44156 | 78 |         this match {
 | 
| 44383 | 79 | case Edits(es) => es.foreach(f) | 
| 80 | case _ => | |
| 44156 | 81 | } | 
| 44383 | 82 | } | 
| 44156 | 83 | } | 
| 44384 | 84 | case class Clear[A, B]() extends Edit[A, B] | 
| 85 | case class Edits[A, B](edits: List[A]) extends Edit[A, B] | |
| 86 | case class Header[A, B](header: Node_Header) extends Edit[A, B] | |
| 87 | case class Perspective[A, B](perspective: B) extends Edit[A, B] | |
| 44156 | 88 | |
| 38426 | 89 | def command_starts(commands: Iterator[Command], offset: Text.Offset = 0) | 
| 90 | : Iterator[(Command, Text.Offset)] = | |
| 38227 
6bbb42843b6e
concentrate structural document notions in document.scala;
 wenzelm parents: 
38220diff
changeset | 91 |     {
 | 
| 
6bbb42843b6e
concentrate structural document notions in document.scala;
 wenzelm parents: 
38220diff
changeset | 92 | var i = offset | 
| 
6bbb42843b6e
concentrate structural document notions in document.scala;
 wenzelm parents: 
38220diff
changeset | 93 |       for (command <- commands) yield {
 | 
| 
6bbb42843b6e
concentrate structural document notions in document.scala;
 wenzelm parents: 
38220diff
changeset | 94 | val start = i | 
| 
6bbb42843b6e
concentrate structural document notions in document.scala;
 wenzelm parents: 
38220diff
changeset | 95 | i += command.length | 
| 
6bbb42843b6e
concentrate structural document notions in document.scala;
 wenzelm parents: 
38220diff
changeset | 96 | (command, start) | 
| 
6bbb42843b6e
concentrate structural document notions in document.scala;
 wenzelm parents: 
38220diff
changeset | 97 | } | 
| 
6bbb42843b6e
concentrate structural document notions in document.scala;
 wenzelm parents: 
38220diff
changeset | 98 | } | 
| 46680 | 99 | |
| 100 | private val block_size = 1024 | |
| 101 | ||
| 102 | val empty: Node = new Node() | |
| 38151 | 103 | } | 
| 104 | ||
| 46712 | 105 | final class Node private( | 
| 46677 | 106 |     val header: Node_Header = Exn.Exn(ERROR("Bad theory header")),
 | 
| 107 | val perspective: Command.Perspective = Command.Perspective.empty, | |
| 108 | val blobs: Map[String, Blob] = Map.empty, | |
| 109 | val commands: Linear_Set[Command] = Linear_Set.empty) | |
| 38151 | 110 |   {
 | 
| 46680 | 111 | def clear: Node = new Node(header = header) | 
| 112 | ||
| 113 | def update_header(new_header: Node_Header): Node = | |
| 114 | new Node(new_header, perspective, blobs, commands) | |
| 115 | ||
| 116 | def update_perspective(new_perspective: Command.Perspective): Node = | |
| 117 | new Node(header, new_perspective, blobs, commands) | |
| 118 | ||
| 119 | def update_blobs(new_blobs: Map[String, Blob]): Node = | |
| 120 | new Node(header, perspective, new_blobs, commands) | |
| 121 | ||
| 122 | def update_commands(new_commands: Linear_Set[Command]): Node = | |
| 123 | new Node(header, perspective, blobs, new_commands) | |
| 44443 
35d67b2056cc
clarified Document.Node.clear -- retain header (cf. ML version);
 wenzelm parents: 
44442diff
changeset | 124 | |
| 46942 
f5c2d66faa04
basic support for outer syntax keywords in theory header;
 wenzelm parents: 
46941diff
changeset | 125 | def imports: List[Node.Name] = | 
| 
f5c2d66faa04
basic support for outer syntax keywords in theory header;
 wenzelm parents: 
46941diff
changeset | 126 |       header match { case Exn.Res(deps) => deps.imports case _ => Nil }
 | 
| 
f5c2d66faa04
basic support for outer syntax keywords in theory header;
 wenzelm parents: 
46941diff
changeset | 127 | |
| 
f5c2d66faa04
basic support for outer syntax keywords in theory header;
 wenzelm parents: 
46941diff
changeset | 128 | def keywords: List[Outer_Syntax.Decl] = | 
| 
f5c2d66faa04
basic support for outer syntax keywords in theory header;
 wenzelm parents: 
46941diff
changeset | 129 |       header match { case Exn.Res(deps) => deps.keywords case _ => Nil }
 | 
| 
f5c2d66faa04
basic support for outer syntax keywords in theory header;
 wenzelm parents: 
46941diff
changeset | 130 | |
| 44443 
35d67b2056cc
clarified Document.Node.clear -- retain header (cf. ML version);
 wenzelm parents: 
44442diff
changeset | 131 | |
| 43697 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43662diff
changeset | 132 | /* commands */ | 
| 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43662diff
changeset | 133 | |
| 38879 
dde403450419
Document.Node: significant speedup of command_range etc. via lazy full_index;
 wenzelm parents: 
38872diff
changeset | 134 | private lazy val full_index: (Array[(Command, Text.Offset)], Text.Range) = | 
| 
dde403450419
Document.Node: significant speedup of command_range etc. via lazy full_index;
 wenzelm parents: 
38872diff
changeset | 135 |     {
 | 
| 
dde403450419
Document.Node: significant speedup of command_range etc. via lazy full_index;
 wenzelm parents: 
38872diff
changeset | 136 | val blocks = new mutable.ListBuffer[(Command, Text.Offset)] | 
| 
dde403450419
Document.Node: significant speedup of command_range etc. via lazy full_index;
 wenzelm parents: 
38872diff
changeset | 137 | var next_block = 0 | 
| 
dde403450419
Document.Node: significant speedup of command_range etc. via lazy full_index;
 wenzelm parents: 
38872diff
changeset | 138 | var last_stop = 0 | 
| 
dde403450419
Document.Node: significant speedup of command_range etc. via lazy full_index;
 wenzelm parents: 
38872diff
changeset | 139 |       for ((command, start) <- Node.command_starts(commands.iterator)) {
 | 
| 
dde403450419
Document.Node: significant speedup of command_range etc. via lazy full_index;
 wenzelm parents: 
38872diff
changeset | 140 | last_stop = start + command.length | 
| 38885 
06582837872b
Node.full_index: allow command spans larger than block_size;
 wenzelm parents: 
38882diff
changeset | 141 |         while (last_stop + 1 > next_block) {
 | 
| 38879 
dde403450419
Document.Node: significant speedup of command_range etc. via lazy full_index;
 wenzelm parents: 
38872diff
changeset | 142 | blocks += (command -> start) | 
| 46680 | 143 | next_block += Node.block_size | 
| 38879 
dde403450419
Document.Node: significant speedup of command_range etc. via lazy full_index;
 wenzelm parents: 
38872diff
changeset | 144 | } | 
| 
dde403450419
Document.Node: significant speedup of command_range etc. via lazy full_index;
 wenzelm parents: 
38872diff
changeset | 145 | } | 
| 
dde403450419
Document.Node: significant speedup of command_range etc. via lazy full_index;
 wenzelm parents: 
38872diff
changeset | 146 | (blocks.toArray, Text.Range(0, last_stop)) | 
| 
dde403450419
Document.Node: significant speedup of command_range etc. via lazy full_index;
 wenzelm parents: 
38872diff
changeset | 147 | } | 
| 38151 | 148 | |
| 38879 
dde403450419
Document.Node: significant speedup of command_range etc. via lazy full_index;
 wenzelm parents: 
38872diff
changeset | 149 | def full_range: Text.Range = full_index._2 | 
| 38151 | 150 | |
| 38569 | 151 | def command_range(i: Text.Offset = 0): Iterator[(Command, Text.Offset)] = | 
| 38879 
dde403450419
Document.Node: significant speedup of command_range etc. via lazy full_index;
 wenzelm parents: 
38872diff
changeset | 152 |     {
 | 
| 
dde403450419
Document.Node: significant speedup of command_range etc. via lazy full_index;
 wenzelm parents: 
38872diff
changeset | 153 |       if (!commands.isEmpty && full_range.contains(i)) {
 | 
| 46680 | 154 | val (cmd0, start0) = full_index._1(i / Node.block_size) | 
| 38879 
dde403450419
Document.Node: significant speedup of command_range etc. via lazy full_index;
 wenzelm parents: 
38872diff
changeset | 155 |         Node.command_starts(commands.iterator(cmd0), start0) dropWhile {
 | 
| 
dde403450419
Document.Node: significant speedup of command_range etc. via lazy full_index;
 wenzelm parents: 
38872diff
changeset | 156 | case (cmd, start) => start + cmd.length <= i } | 
| 
dde403450419
Document.Node: significant speedup of command_range etc. via lazy full_index;
 wenzelm parents: 
38872diff
changeset | 157 | } | 
| 
dde403450419
Document.Node: significant speedup of command_range etc. via lazy full_index;
 wenzelm parents: 
38872diff
changeset | 158 | else Iterator.empty | 
| 
dde403450419
Document.Node: significant speedup of command_range etc. via lazy full_index;
 wenzelm parents: 
38872diff
changeset | 159 | } | 
| 38151 | 160 | |
| 38569 | 161 | def command_range(range: Text.Range): Iterator[(Command, Text.Offset)] = | 
| 162 |       command_range(range.start) takeWhile { case (_, start) => start < range.stop }
 | |
| 38151 | 163 | |
| 38426 | 164 | def command_at(i: Text.Offset): Option[(Command, Text.Offset)] = | 
| 38151 | 165 |     {
 | 
| 166 | val range = command_range(i) | |
| 167 | if (range.hasNext) Some(range.next) else None | |
| 168 | } | |
| 169 | ||
| 38879 
dde403450419
Document.Node: significant speedup of command_range etc. via lazy full_index;
 wenzelm parents: 
38872diff
changeset | 170 | def command_start(cmd: Command): Option[Text.Offset] = | 
| 46677 | 171 | Node.command_starts(commands.iterator).find(_._1 == cmd).map(_._2) | 
| 38151 | 172 | } | 
| 173 | ||
| 174 | ||
| 46723 | 175 | /* development graph */ | 
| 176 | ||
| 177 | object Nodes | |
| 178 |   {
 | |
| 179 | val empty: Nodes = new Nodes(Graph.empty(Node.Name.Ordering)) | |
| 180 | } | |
| 181 | ||
| 182 | final class Nodes private(graph: Graph[Node.Name, Node]) | |
| 183 |   {
 | |
| 184 | def get_name(s: String): Option[Node.Name] = | |
| 185 | graph.keys.find(name => name.node == s) | |
| 186 | ||
| 187 | def apply(name: Node.Name): Node = | |
| 188 | graph.default_node(name, Node.empty).get_node(name) | |
| 189 | ||
| 190 | def + (entry: (Node.Name, Node)): Nodes = | |
| 191 |     {
 | |
| 192 | val (name, node) = entry | |
| 46942 
f5c2d66faa04
basic support for outer syntax keywords in theory header;
 wenzelm parents: 
46941diff
changeset | 193 | val imports = node.imports | 
| 46723 | 194 | val graph1 = | 
| 46739 
6024353549ca
clarified document nodes (full import graph) vs. node_status (non-preloaded theories);
 wenzelm parents: 
46737diff
changeset | 195 | (graph.default_node(name, Node.empty) /: imports)((g, p) => g.default_node(p, Node.empty)) | 
| 46723 | 196 | val graph2 = (graph1 /: graph1.imm_preds(name))((g, dep) => g.del_edge(dep, name)) | 
| 46739 
6024353549ca
clarified document nodes (full import graph) vs. node_status (non-preloaded theories);
 wenzelm parents: 
46737diff
changeset | 197 | val graph3 = (graph2 /: imports)((g, dep) => g.add_edge(dep, name)) | 
| 46723 | 198 | new Nodes(graph3.map_node(name, _ => node)) | 
| 199 | } | |
| 200 | ||
| 201 | def entries: Iterator[(Node.Name, Node)] = | |
| 202 |       graph.entries.map({ case (name, (node, _)) => (name, node) })
 | |
| 203 | ||
| 46942 
f5c2d66faa04
basic support for outer syntax keywords in theory header;
 wenzelm parents: 
46941diff
changeset | 204 | def descendants(names: List[Node.Name]): List[Node.Name] = graph.all_succs(names) | 
| 46723 | 205 | def topological_order: List[Node.Name] = graph.topological_order | 
| 206 | } | |
| 207 | ||
| 208 | ||
| 38424 | 209 | |
| 210 | /** versioning **/ | |
| 211 | ||
| 212 | /* particular document versions */ | |
| 34485 | 213 | |
| 38417 | 214 | object Version | 
| 215 |   {
 | |
| 46681 | 216 | val init: Version = new Version() | 
| 217 | ||
| 46941 | 218 | def make(syntax: Outer_Syntax, nodes: Nodes): Version = | 
| 219 | new Version(new_id(), syntax, nodes) | |
| 38417 | 220 | } | 
| 221 | ||
| 46712 | 222 | final class Version private( | 
| 46677 | 223 | val id: Version_ID = no_id, | 
| 46941 | 224 | val syntax: Outer_Syntax = Outer_Syntax.empty, | 
| 46723 | 225 | val nodes: Nodes = Nodes.empty) | 
| 46941 | 226 |   {
 | 
| 227 | def is_init: Boolean = id == no_id | |
| 228 | } | |
| 34660 | 229 | |
| 34859 | 230 | |
| 38424 | 231 | /* changes of plain text, eventually resulting in document edits */ | 
| 38227 
6bbb42843b6e
concentrate structural document notions in document.scala;
 wenzelm parents: 
38220diff
changeset | 232 | |
| 
6bbb42843b6e
concentrate structural document notions in document.scala;
 wenzelm parents: 
38220diff
changeset | 233 | object Change | 
| 
6bbb42843b6e
concentrate structural document notions in document.scala;
 wenzelm parents: 
38220diff
changeset | 234 |   {
 | 
| 46678 | 235 | val init: Change = new Change() | 
| 236 | ||
| 237 | def make(previous: Future[Version], edits: List[Edit_Text], version: Future[Version]): Change = | |
| 238 | new Change(Some(previous), edits, version) | |
| 38227 
6bbb42843b6e
concentrate structural document notions in document.scala;
 wenzelm parents: 
38220diff
changeset | 239 | } | 
| 
6bbb42843b6e
concentrate structural document notions in document.scala;
 wenzelm parents: 
38220diff
changeset | 240 | |
| 46712 | 241 | final class Change private( | 
| 46677 | 242 | val previous: Option[Future[Version]] = Some(Future.value(Version.init)), | 
| 243 | val edits: List[Edit_Text] = Nil, | |
| 244 | val version: Future[Version] = Future.value(Version.init)) | |
| 38227 
6bbb42843b6e
concentrate structural document notions in document.scala;
 wenzelm parents: 
38220diff
changeset | 245 |   {
 | 
| 44672 | 246 | def is_finished: Boolean = | 
| 247 |       (previous match { case None => true case Some(future) => future.is_finished }) &&
 | |
| 248 | version.is_finished | |
| 249 | ||
| 46678 | 250 | def truncate: Change = new Change(None, Nil, version) | 
| 38227 
6bbb42843b6e
concentrate structural document notions in document.scala;
 wenzelm parents: 
38220diff
changeset | 251 | } | 
| 
6bbb42843b6e
concentrate structural document notions in document.scala;
 wenzelm parents: 
38220diff
changeset | 252 | |
| 
6bbb42843b6e
concentrate structural document notions in document.scala;
 wenzelm parents: 
38220diff
changeset | 253 | |
| 38841 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 254 | /* history navigation */ | 
| 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 255 | |
| 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 256 | object History | 
| 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 257 |   {
 | 
| 46679 | 258 | val init: History = new History() | 
| 38841 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 259 | } | 
| 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 260 | |
| 46712 | 261 | final class History private( | 
| 46679 | 262 | val undo_list: List[Change] = List(Change.init)) // non-empty list | 
| 38841 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 263 |   {
 | 
| 46679 | 264 | def tip: Change = undo_list.head | 
| 265 | def + (change: Change): History = new History(change :: undo_list) | |
| 38841 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 266 | |
| 46679 | 267 | def prune(check: Change => Boolean, retain: Int): Option[(List[Change], History)] = | 
| 268 |     {
 | |
| 269 | val n = undo_list.iterator.zipWithIndex.find(p => check(p._1)).get._2 + 1 | |
| 270 | val (retained, dropped) = undo_list.splitAt(n max retain) | |
| 271 | ||
| 272 |       retained.splitAt(retained.length - 1) match {
 | |
| 273 | case (prefix, List(last)) => Some(dropped, new History(prefix ::: List(last.truncate))) | |
| 274 | case _ => None | |
| 275 | } | |
| 276 | } | |
| 38841 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 277 | } | 
| 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 278 | |
| 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 279 | |
| 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 280 | |
| 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 281 | /** global state -- document structure, execution process, editing history **/ | 
| 38424 | 282 | |
| 283 | abstract class Snapshot | |
| 284 |   {
 | |
| 44582 | 285 | val state: State | 
| 38424 | 286 | val version: Version | 
| 287 | val node: Node | |
| 288 | val is_outdated: Boolean | |
| 38427 | 289 | def convert(i: Text.Offset): Text.Offset | 
| 39177 
0468964aec11
simplified Markup_Tree.select: Stream instead of Iterator (again), explicit Option instead of default;
 wenzelm parents: 
38976diff
changeset | 290 | def convert(range: Text.Range): Text.Range | 
| 38427 | 291 | def revert(i: Text.Offset): Text.Offset | 
| 39177 
0468964aec11
simplified Markup_Tree.select: Stream instead of Iterator (again), explicit Option instead of default;
 wenzelm parents: 
38976diff
changeset | 292 | def revert(range: Text.Range): Text.Range | 
| 46178 
1c5c88f6feb5
clarified Isabelle_Rendering vs. physical painting;
 wenzelm parents: 
46152diff
changeset | 293 | def cumulate_markup[A](range: Text.Range, info: A, elements: Option[Set[String]], | 
| 
1c5c88f6feb5
clarified Isabelle_Rendering vs. physical painting;
 wenzelm parents: 
46152diff
changeset | 294 | result: PartialFunction[(A, Text.Markup), A]): Stream[Text.Info[A]] | 
| 
1c5c88f6feb5
clarified Isabelle_Rendering vs. physical painting;
 wenzelm parents: 
46152diff
changeset | 295 | def select_markup[A](range: Text.Range, elements: Option[Set[String]], | 
| 46198 
cd040c5772de
improved select_markup: include filtering of defined results;
 wenzelm parents: 
46178diff
changeset | 296 | result: PartialFunction[Text.Markup, A]): Stream[Text.Info[A]] | 
| 38424 | 297 | } | 
| 298 | ||
| 44476 
e8a87398f35d
propagate information about last command with exec state assignment through document model;
 wenzelm parents: 
44475diff
changeset | 299 | type Assign = | 
| 44479 
9a04e7502e22
refined document state assignment: observe perspective, more explicit assignment message;
 wenzelm parents: 
44477diff
changeset | 300 | (List[(Document.Command_ID, Option[Document.Exec_ID])], // exec state assignment | 
| 44476 
e8a87398f35d
propagate information about last command with exec state assignment through document model;
 wenzelm parents: 
44475diff
changeset | 301 | List[(String, Option[Document.Command_ID])]) // last exec | 
| 
e8a87398f35d
propagate information about last command with exec state assignment through document model;
 wenzelm parents: 
44475diff
changeset | 302 | |
| 38370 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 303 | object State | 
| 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 304 |   {
 | 
| 38373 | 305 | class Fail(state: State) extends Exception | 
| 306 | ||
| 44479 
9a04e7502e22
refined document state assignment: observe perspective, more explicit assignment message;
 wenzelm parents: 
44477diff
changeset | 307 | object Assignment | 
| 
9a04e7502e22
refined document state assignment: observe perspective, more explicit assignment message;
 wenzelm parents: 
44477diff
changeset | 308 |     {
 | 
| 46683 | 309 | val init: Assignment = new Assignment() | 
| 44479 
9a04e7502e22
refined document state assignment: observe perspective, more explicit assignment message;
 wenzelm parents: 
44477diff
changeset | 310 | } | 
| 38370 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 311 | |
| 46712 | 312 | final class Assignment private( | 
| 46677 | 313 | val command_execs: Map[Command_ID, Exec_ID] = Map.empty, | 
| 314 | val last_execs: Map[String, Option[Command_ID]] = Map.empty, | |
| 315 | val is_finished: Boolean = false) | |
| 38370 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 316 |     {
 | 
| 44477 | 317 |       def check_finished: Assignment = { require(is_finished); this }
 | 
| 46683 | 318 | def unfinished: Assignment = new Assignment(command_execs, last_execs, false) | 
| 44479 
9a04e7502e22
refined document state assignment: observe perspective, more explicit assignment message;
 wenzelm parents: 
44477diff
changeset | 319 | |
| 
9a04e7502e22
refined document state assignment: observe perspective, more explicit assignment message;
 wenzelm parents: 
44477diff
changeset | 320 | def assign(add_command_execs: List[(Command_ID, Option[Exec_ID])], | 
| 44477 | 321 | add_last_execs: List[(String, Option[Command_ID])]): Assignment = | 
| 38370 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 322 |       {
 | 
| 38976 
a4a465dc89d9
Document.State.Assignment: eliminated promise in favour of plain values -- signalling is done via event bus in Session;
 wenzelm parents: 
38885diff
changeset | 323 | require(!is_finished) | 
| 44479 
9a04e7502e22
refined document state assignment: observe perspective, more explicit assignment message;
 wenzelm parents: 
44477diff
changeset | 324 | val command_execs1 = | 
| 
9a04e7502e22
refined document state assignment: observe perspective, more explicit assignment message;
 wenzelm parents: 
44477diff
changeset | 325 |           (command_execs /: add_command_execs) {
 | 
| 
9a04e7502e22
refined document state assignment: observe perspective, more explicit assignment message;
 wenzelm parents: 
44477diff
changeset | 326 | case (res, (command_id, None)) => res - command_id | 
| 
9a04e7502e22
refined document state assignment: observe perspective, more explicit assignment message;
 wenzelm parents: 
44477diff
changeset | 327 | case (res, (command_id, Some(exec_id))) => res + (command_id -> exec_id) | 
| 
9a04e7502e22
refined document state assignment: observe perspective, more explicit assignment message;
 wenzelm parents: 
44477diff
changeset | 328 | } | 
| 46683 | 329 | new Assignment(command_execs1, last_execs ++ add_last_execs, true) | 
| 38150 
67fc24df3721
simplified/refined document model: collection of named nodes, without proper dependencies yet;
 wenzelm parents: 
37849diff
changeset | 330 | } | 
| 38370 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 331 | } | 
| 46682 | 332 | |
| 333 | val init: State = | |
| 334 | State().define_version(Version.init, Assignment.init).assign(Version.init.id)._2 | |
| 38370 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 335 | } | 
| 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 336 | |
| 46712 | 337 | final case class State private( | 
| 46677 | 338 | val versions: Map[Version_ID, Version] = Map.empty, | 
| 339 | val commands: Map[Command_ID, Command.State] = Map.empty, // static markup from define_command | |
| 340 | val execs: Map[Exec_ID, Command.State] = Map.empty, // dynamic markup from execution | |
| 341 | val assignments: Map[Version_ID, State.Assignment] = Map.empty, | |
| 38841 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 342 | val history: History = History.init) | 
| 38370 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 343 |   {
 | 
| 38373 | 344 | private def fail[A]: A = throw new State.Fail(this) | 
| 38370 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 345 | |
| 44479 
9a04e7502e22
refined document state assignment: observe perspective, more explicit assignment message;
 wenzelm parents: 
44477diff
changeset | 346 | def define_version(version: Version, assignment: State.Assignment): State = | 
| 38370 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 347 |     {
 | 
| 38417 | 348 | val id = version.id | 
| 349 | copy(versions = versions + (id -> version), | |
| 44479 
9a04e7502e22
refined document state assignment: observe perspective, more explicit assignment message;
 wenzelm parents: 
44477diff
changeset | 350 | assignments = assignments + (id -> assignment.unfinished)) | 
| 38370 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 351 | } | 
| 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 352 | |
| 38373 | 353 | def define_command(command: Command): State = | 
| 354 |     {
 | |
| 355 | val id = command.id | |
| 356 | copy(commands = commands + (id -> command.empty_state)) | |
| 357 | } | |
| 358 | ||
| 44582 | 359 | def defined_command(id: Command_ID): Boolean = commands.isDefinedAt(id) | 
| 360 | ||
| 44607 
274eff0ea12e
maintain name of *the* enclosing node as part of command -- avoid full document traversal;
 wenzelm parents: 
44582diff
changeset | 361 | def find_command(version: Version, id: ID): Option[(Node, Command)] = | 
| 44582 | 362 |       commands.get(id) orElse execs.get(id) match {
 | 
| 363 | case None => None | |
| 364 | case Some(st) => | |
| 365 | val command = st.command | |
| 46681 | 366 | val node = version.nodes(command.node_name) | 
| 367 | Some((node, command)) | |
| 44582 | 368 | } | 
| 38373 | 369 | |
| 38417 | 370 | def the_version(id: Version_ID): Version = versions.getOrElse(id, fail) | 
| 44479 
9a04e7502e22
refined document state assignment: observe perspective, more explicit assignment message;
 wenzelm parents: 
44477diff
changeset | 371 | def the_command(id: Command_ID): Command.State = commands.getOrElse(id, fail) // FIXME rename | 
| 44446 | 372 | def the_exec(id: Exec_ID): Command.State = execs.getOrElse(id, fail) | 
| 43722 | 373 | def the_assignment(version: Version): State.Assignment = | 
| 374 | assignments.getOrElse(version.id, fail) | |
| 38370 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 375 | |
| 38872 | 376 | def accumulate(id: ID, message: XML.Elem): (Command.State, State) = | 
| 38370 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 377 |       execs.get(id) match {
 | 
| 44446 | 378 | case Some(st) => | 
| 46152 
793cecd4ffc0
accumulate status as regular markup for command range;
 wenzelm parents: 
45900diff
changeset | 379 | val new_st = st + message | 
| 44446 | 380 | (new_st, copy(execs = execs + (id -> new_st))) | 
| 38370 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 381 | case None => | 
| 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 382 |           commands.get(id) match {
 | 
| 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 383 | case Some(st) => | 
| 46152 
793cecd4ffc0
accumulate status as regular markup for command range;
 wenzelm parents: 
45900diff
changeset | 384 | val new_st = st + message | 
| 38370 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 385 | (new_st, copy(commands = commands + (id -> new_st))) | 
| 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 386 | case None => fail | 
| 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 387 | } | 
| 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 388 | } | 
| 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 389 | |
| 46682 | 390 | def assign(id: Version_ID, arg: Assign = (Nil, Nil)): (List[Command], State) = | 
| 38370 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 391 |     {
 | 
| 38417 | 392 | val version = the_version(id) | 
| 44477 | 393 | val (command_execs, last_execs) = arg | 
| 38370 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 394 | |
| 44543 | 395 | val (changed_commands, new_execs) = | 
| 44479 
9a04e7502e22
refined document state assignment: observe perspective, more explicit assignment message;
 wenzelm parents: 
44477diff
changeset | 396 |         ((Nil: List[Command], execs) /: command_execs) {
 | 
| 44543 | 397 | case ((commands1, execs1), (cmd_id, exec)) => | 
| 44479 
9a04e7502e22
refined document state assignment: observe perspective, more explicit assignment message;
 wenzelm parents: 
44477diff
changeset | 398 | val st = the_command(cmd_id) | 
| 44543 | 399 | val commands2 = st.command :: commands1 | 
| 400 | val execs2 = | |
| 401 |               exec match {
 | |
| 402 | case None => execs1 | |
| 403 | case Some(exec_id) => execs1 + (exec_id -> st) | |
| 404 | } | |
| 405 | (commands2, execs2) | |
| 38370 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 406 | } | 
| 44477 | 407 | val new_assignment = the_assignment(version).assign(command_execs, last_execs) | 
| 43722 | 408 | val new_state = copy(assignments = assignments + (id -> new_assignment), execs = new_execs) | 
| 44476 
e8a87398f35d
propagate information about last command with exec state assignment through document model;
 wenzelm parents: 
44475diff
changeset | 409 | |
| 44543 | 410 | (changed_commands, new_state) | 
| 38370 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 411 | } | 
| 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 412 | |
| 38417 | 413 | def is_assigned(version: Version): Boolean = | 
| 43722 | 414 |       assignments.get(version.id) match {
 | 
| 38370 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 415 | case Some(assgn) => assgn.is_finished | 
| 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 416 | case None => false | 
| 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 417 | } | 
| 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 418 | |
| 44436 
546adfa8a6fc
update_perspective without actual edits, bypassing the full state assignment protocol;
 wenzelm parents: 
44385diff
changeset | 419 | def is_stable(change: Change): Boolean = | 
| 
546adfa8a6fc
update_perspective without actual edits, bypassing the full state assignment protocol;
 wenzelm parents: 
44385diff
changeset | 420 | change.is_finished && is_assigned(change.version.get_finished) | 
| 
546adfa8a6fc
update_perspective without actual edits, bypassing the full state assignment protocol;
 wenzelm parents: 
44385diff
changeset | 421 | |
| 46944 
9fc22eb6408c
more recent recent_syntax, e.g. relevant for document rendering during startup;
 wenzelm parents: 
46942diff
changeset | 422 | def recent_finished: Change = history.undo_list.find(_.is_finished) getOrElse fail | 
| 44672 | 423 | def recent_stable: Change = history.undo_list.find(is_stable) getOrElse fail | 
| 44436 
546adfa8a6fc
update_perspective without actual edits, bypassing the full state assignment protocol;
 wenzelm parents: 
44385diff
changeset | 424 | def tip_stable: Boolean = is_stable(history.tip) | 
| 44475 | 425 | def tip_version: Version = history.tip.version.get_finished | 
| 44436 
546adfa8a6fc
update_perspective without actual edits, bypassing the full state assignment protocol;
 wenzelm parents: 
44385diff
changeset | 426 | |
| 44615 | 427 | def last_exec_offset(name: Node.Name): Text.Offset = | 
| 44484 
470f2ee5950e
tuned Session.edit_node: update_perspective based on last_exec_offset;
 wenzelm parents: 
44479diff
changeset | 428 |     {
 | 
| 
470f2ee5950e
tuned Session.edit_node: update_perspective based on last_exec_offset;
 wenzelm parents: 
44479diff
changeset | 429 | val version = tip_version | 
| 44615 | 430 |       the_assignment(version).last_execs.get(name.node) match {
 | 
| 44484 
470f2ee5950e
tuned Session.edit_node: update_perspective based on last_exec_offset;
 wenzelm parents: 
44479diff
changeset | 431 | case Some(Some(id)) => | 
| 
470f2ee5950e
tuned Session.edit_node: update_perspective based on last_exec_offset;
 wenzelm parents: 
44479diff
changeset | 432 | val node = version.nodes(name) | 
| 
470f2ee5950e
tuned Session.edit_node: update_perspective based on last_exec_offset;
 wenzelm parents: 
44479diff
changeset | 433 | val cmd = the_command(id).command | 
| 
470f2ee5950e
tuned Session.edit_node: update_perspective based on last_exec_offset;
 wenzelm parents: 
44479diff
changeset | 434 |           node.command_start(cmd) match {
 | 
| 
470f2ee5950e
tuned Session.edit_node: update_perspective based on last_exec_offset;
 wenzelm parents: 
44479diff
changeset | 435 | case None => 0 | 
| 
470f2ee5950e
tuned Session.edit_node: update_perspective based on last_exec_offset;
 wenzelm parents: 
44479diff
changeset | 436 | case Some(start) => start + cmd.length | 
| 
470f2ee5950e
tuned Session.edit_node: update_perspective based on last_exec_offset;
 wenzelm parents: 
44479diff
changeset | 437 | } | 
| 
470f2ee5950e
tuned Session.edit_node: update_perspective based on last_exec_offset;
 wenzelm parents: 
44479diff
changeset | 438 | case _ => 0 | 
| 
470f2ee5950e
tuned Session.edit_node: update_perspective based on last_exec_offset;
 wenzelm parents: 
44479diff
changeset | 439 | } | 
| 
470f2ee5950e
tuned Session.edit_node: update_perspective based on last_exec_offset;
 wenzelm parents: 
44479diff
changeset | 440 | } | 
| 
470f2ee5950e
tuned Session.edit_node: update_perspective based on last_exec_offset;
 wenzelm parents: 
44479diff
changeset | 441 | |
| 44446 | 442 | def continue_history( | 
| 443 | previous: Future[Version], | |
| 40479 | 444 | edits: List[Edit_Text], | 
| 43722 | 445 | version: Future[Version]): (Change, State) = | 
| 38841 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 446 |     {
 | 
| 46678 | 447 | val change = Change.make(previous, edits, version) | 
| 38841 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 448 | (change, copy(history = history + change)) | 
| 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 449 | } | 
| 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 450 | |
| 44672 | 451 | def prune_history(retain: Int = 0): (List[Version], State) = | 
| 452 |     {
 | |
| 46679 | 453 |       history.prune(is_stable, retain) match {
 | 
| 454 | case Some((dropped, history1)) => | |
| 44672 | 455 | val dropped_versions = dropped.map(change => change.version.get_finished) | 
| 46679 | 456 | val state1 = copy(history = history1) | 
| 44672 | 457 | (dropped_versions, state1) | 
| 46679 | 458 | case None => fail | 
| 44672 | 459 | } | 
| 460 | } | |
| 461 | ||
| 44676 | 462 | def removed_versions(removed: List[Version_ID]): State = | 
| 463 |     {
 | |
| 464 | val versions1 = versions -- removed | |
| 465 | val assignments1 = assignments -- removed | |
| 466 | var commands1 = Map.empty[Command_ID, Command.State] | |
| 467 | var execs1 = Map.empty[Exec_ID, Command.State] | |
| 468 |       for {
 | |
| 469 | (version_id, version) <- versions1.iterator | |
| 46997 | 470 | command_execs = assignments1(version_id).command_execs | 
| 46723 | 471 | (_, node) <- version.nodes.entries | 
| 44676 | 472 | command <- node.commands.iterator | 
| 473 |       } {
 | |
| 474 | val id = command.id | |
| 475 | if (!commands1.isDefinedAt(id) && commands.isDefinedAt(id)) | |
| 476 | commands1 += (id -> commands(id)) | |
| 477 |         if (command_execs.isDefinedAt(id)) {
 | |
| 478 | val exec_id = command_execs(id) | |
| 479 | if (!execs1.isDefinedAt(exec_id) && execs.isDefinedAt(exec_id)) | |
| 480 | execs1 += (exec_id -> execs(exec_id)) | |
| 481 | } | |
| 482 | } | |
| 483 | copy(versions = versions1, commands = commands1, execs = execs1, assignments = assignments1) | |
| 484 | } | |
| 485 | ||
| 44613 | 486 | def command_state(version: Version, command: Command): Command.State = | 
| 487 |     {
 | |
| 488 | require(is_assigned(version)) | |
| 489 |       try {
 | |
| 490 | the_exec(the_assignment(version).check_finished.command_execs | |
| 491 | .getOrElse(command.id, fail)) | |
| 492 | } | |
| 493 |       catch { case _: State.Fail => command.empty_state }
 | |
| 494 | } | |
| 495 | ||
| 38841 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 496 | |
| 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 497 | // persistent user-view | 
| 44615 | 498 | def snapshot(name: Node.Name, pending_edits: List[Text.Edit]): Snapshot = | 
| 38370 
8b15d0f98962
explicit Document.State value, instead of individual state variables in Session, Command, Document;
 wenzelm parents: 
38367diff
changeset | 499 |     {
 | 
| 44672 | 500 | val stable = recent_stable | 
| 44436 
546adfa8a6fc
update_perspective without actual edits, bypassing the full state assignment protocol;
 wenzelm parents: 
44385diff
changeset | 501 | val latest = history.tip | 
| 38841 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 502 | |
| 44156 | 503 | // FIXME proper treatment of removed nodes | 
| 38841 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 504 | val edits = | 
| 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 505 | (pending_edits /: history.undo_list.takeWhile(_ != stable))((edits, change) => | 
| 44156 | 506 | (for ((a, Node.Edits(es)) <- change.edits if a == name) yield es).flatten ::: edits) | 
| 38841 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 507 | lazy val reverse_edits = edits.reverse | 
| 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 508 | |
| 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 509 | new Snapshot | 
| 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 510 |       {
 | 
| 44582 | 511 | val state = State.this | 
| 39698 | 512 | val version = stable.version.get_finished | 
| 38841 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 513 | val node = version.nodes(name) | 
| 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 514 | val is_outdated = !(pending_edits.isEmpty && latest == stable) | 
| 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 515 | |
| 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 516 | def convert(offset: Text.Offset) = (offset /: edits)((i, edit) => edit.convert(i)) | 
| 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 517 | def revert(offset: Text.Offset) = (offset /: reverse_edits)((i, edit) => edit.revert(i)) | 
| 43425 | 518 | def convert(range: Text.Range) = (range /: edits)((r, edit) => edit.convert(r)) | 
| 519 | def revert(range: Text.Range) = (range /: reverse_edits)((r, edit) => edit.revert(r)) | |
| 39177 
0468964aec11
simplified Markup_Tree.select: Stream instead of Iterator (again), explicit Option instead of default;
 wenzelm parents: 
38976diff
changeset | 520 | |
| 46178 
1c5c88f6feb5
clarified Isabelle_Rendering vs. physical painting;
 wenzelm parents: 
46152diff
changeset | 521 | def cumulate_markup[A](range: Text.Range, info: A, elements: Option[Set[String]], | 
| 
1c5c88f6feb5
clarified Isabelle_Rendering vs. physical painting;
 wenzelm parents: 
46152diff
changeset | 522 | result: PartialFunction[(A, Text.Markup), A]): Stream[Text.Info[A]] = | 
| 45459 | 523 |         {
 | 
| 45467 | 524 | val former_range = revert(range) | 
| 45459 | 525 |           for {
 | 
| 526 | (command, command_start) <- node.command_range(former_range).toStream | |
| 46208 | 527 | Text.Info(r0, a) <- state.command_state(version, command).markup. | 
| 46178 
1c5c88f6feb5
clarified Isabelle_Rendering vs. physical painting;
 wenzelm parents: 
46152diff
changeset | 528 | cumulate[A]((former_range - command_start).restrict(command.range), info, elements, | 
| 
1c5c88f6feb5
clarified Isabelle_Rendering vs. physical painting;
 wenzelm parents: 
46152diff
changeset | 529 |                 {
 | 
| 
1c5c88f6feb5
clarified Isabelle_Rendering vs. physical painting;
 wenzelm parents: 
46152diff
changeset | 530 | case (a, Text.Info(r0, b)) | 
| 
1c5c88f6feb5
clarified Isabelle_Rendering vs. physical painting;
 wenzelm parents: 
46152diff
changeset | 531 | if result.isDefinedAt((a, Text.Info(convert(r0 + command_start), b))) => | 
| 
1c5c88f6feb5
clarified Isabelle_Rendering vs. physical painting;
 wenzelm parents: 
46152diff
changeset | 532 | result((a, Text.Info(convert(r0 + command_start), b))) | 
| 
1c5c88f6feb5
clarified Isabelle_Rendering vs. physical painting;
 wenzelm parents: 
46152diff
changeset | 533 | }) | 
| 45459 | 534 | } yield Text.Info(convert(r0 + command_start), a) | 
| 535 | } | |
| 536 | ||
| 46178 
1c5c88f6feb5
clarified Isabelle_Rendering vs. physical painting;
 wenzelm parents: 
46152diff
changeset | 537 | def select_markup[A](range: Text.Range, elements: Option[Set[String]], | 
| 46198 
cd040c5772de
improved select_markup: include filtering of defined results;
 wenzelm parents: 
46178diff
changeset | 538 | result: PartialFunction[Text.Markup, A]): Stream[Text.Info[A]] = | 
| 38845 
a9e37daf5bd0
added Document.Snapshot.select_markup, which includes command iteration, range conversion etc.;
 wenzelm parents: 
38841diff
changeset | 539 |         {
 | 
| 45468 | 540 | val result1 = | 
| 541 |             new PartialFunction[(Option[A], Text.Markup), Option[A]] {
 | |
| 542 | def isDefinedAt(arg: (Option[A], Text.Markup)): Boolean = result.isDefinedAt(arg._2) | |
| 543 | def apply(arg: (Option[A], Text.Markup)): Option[A] = Some(result(arg._2)) | |
| 544 | } | |
| 46198 
cd040c5772de
improved select_markup: include filtering of defined results;
 wenzelm parents: 
46178diff
changeset | 545 | for (Text.Info(r, Some(x)) <- cumulate_markup(range, None, elements, result1)) | 
| 
cd040c5772de
improved select_markup: include filtering of defined results;
 wenzelm parents: 
46178diff
changeset | 546 | yield Text.Info(r, x) | 
| 38845 
a9e37daf5bd0
added Document.Snapshot.select_markup, which includes command iteration, range conversion etc.;
 wenzelm parents: 
38841diff
changeset | 547 | } | 
| 38841 
4df7b76249a0
include Document.History in Document.State -- just one universal session state maintained by main actor;
 wenzelm parents: 
38569diff
changeset | 548 | } | 
| 38150 
67fc24df3721
simplified/refined document model: collection of named nodes, without proper dependencies yet;
 wenzelm parents: 
37849diff
changeset | 549 | } | 
| 34485 | 550 | } | 
| 34840 
6c5560d48561
more precise treatment of document/state assigment;
 wenzelm parents: 
34838diff
changeset | 551 | } | 
| 
6c5560d48561
more precise treatment of document/state assigment;
 wenzelm parents: 
34838diff
changeset | 552 |