separate file_node vs. file_path, e.g. relevant on Windows for hyperlink to the latter;
1 (* Title: Pure/PIDE/document.ML
4 Document as collection of named nodes, each consisting of an editable
5 list of commands, associated with asynchronous execution process.
10 val timing: bool Unsynchronized.ref
11 type node_header = string * Thy_Header.header * string list
12 type overlay = Document_ID.command * (string * string list)
14 Edits of (Document_ID.command option * Document_ID.command option) list |
16 Perspective of bool * Document_ID.command list * overlay list
17 type edit = string * node_edit
20 val define_blob: string -> string -> state -> state
21 type blob_digest = (string * string * string option) Exn.result
22 val define_command: Document_ID.command -> string -> blob_digest list -> string ->
24 val remove_versions: Document_ID.version list -> state -> state
25 val start_execution: state -> state
26 val update: Document_ID.version -> Document_ID.version -> edit list -> state ->
27 Document_ID.exec list * (Document_ID.command * Document_ID.exec list) list * state
28 val state: unit -> state
29 val change_state: (state -> state) -> unit
32 structure Document: DOCUMENT =
35 val timing = Unsynchronized.ref false;
36 fun timeit msg e = cond_timeit (! timing) msg e;
40 (** document structure **)
42 fun err_dup kind id = error ("Duplicate " ^ kind ^ ": " ^ Document_ID.print id);
43 fun err_undef kind id = error ("Undefined " ^ kind ^ ": " ^ Document_ID.print id);
45 type node_header = string * Thy_Header.header * string list;
48 {required: bool, (*required node*)
49 visible: Inttab.set, (*visible commands*)
50 visible_last: Document_ID.command option, (*last visible command*)
51 overlays: (string * string list) list Inttab.table}; (*command id -> print function with args*)
53 structure Entries = Linear_Set(type key = Document_ID.command val ord = int_ord);
55 abstype node = Node of
56 {header: node_header, (*master directory, theory header, errors*)
57 perspective: perspective, (*command perspective*)
58 entries: Command.exec option Entries.T, (*command entries with excecutions*)
59 result: Command.eval option} (*result of last execution*)
60 and version = Version of node String_Graph.T (*development graph wrt. static imports*)
63 fun make_node (header, perspective, entries, result) =
64 Node {header = header, perspective = perspective, entries = entries, result = result};
66 fun map_node f (Node {header, perspective, entries, result}) =
67 make_node (f (header, perspective, entries, result));
69 fun make_perspective (required, command_ids, overlays) : perspective =
71 visible = Inttab.make_set command_ids,
72 visible_last = try List.last command_ids,
73 overlays = Inttab.make_list overlays};
75 val no_header = ("", Thy_Header.make ("", Position.none) [] [], ["No theory header"]);
76 val no_perspective = make_perspective (false, [], []);
78 val empty_node = make_node (no_header, no_perspective, Entries.empty, NONE);
81 (* basic components *)
83 fun master_directory (Node {header = (master, _, _), ...}) =
84 (case try Url.explode master of
85 SOME (Url.File path) => path
88 fun set_header header =
89 map_node (fn (_, perspective, entries, result) => (header, perspective, entries, result));
91 fun get_header (Node {header = (master, header, errors), ...}) =
92 if null errors then (master, header)
93 else error (cat_lines errors);
95 fun read_header node span =
97 val {name = (name, _), imports, keywords} = #2 (get_header node);
98 val {name = (_, pos), imports = imports', ...} = Thy_Header.read_tokens span;
99 in Thy_Header.make (name, pos) (map #1 imports ~~ map #2 imports') keywords end;
101 fun get_perspective (Node {perspective, ...}) = perspective;
102 fun set_perspective args =
103 map_node (fn (header, _, entries, result) => (header, make_perspective args, entries, result));
105 val required_node = #required o get_perspective;
106 val visible_command = Inttab.defined o #visible o get_perspective;
107 val visible_last = #visible_last o get_perspective;
108 val visible_node = is_some o visible_last
109 val overlays = Inttab.lookup_list o #overlays o get_perspective;
112 map_node (fn (header, perspective, entries, result) => (header, perspective, f entries, result));
113 fun get_entries (Node {entries, ...}) = entries;
115 fun iterate_entries f = Entries.iterate NONE f o get_entries;
116 fun iterate_entries_after start f (Node {entries, ...}) =
117 (case Entries.get_after entries start of
119 | SOME id => Entries.iterate (SOME id) f entries);
121 fun get_result (Node {result, ...}) = result;
122 fun set_result result =
123 map_node (fn (header, perspective, entries, _) => (header, perspective, entries, result));
125 fun changed_result node node' =
126 (case (get_result node, get_result node') of
127 (SOME eval, SOME eval') => not (Command.eval_eq (eval, eval'))
128 | (NONE, NONE) => false
131 fun pending_result node =
132 (case get_result node of
133 SOME eval => not (Command.eval_finished eval)
136 fun get_node nodes name = String_Graph.get_node nodes name
137 handle String_Graph.UNDEF _ => empty_node;
138 fun default_node name = String_Graph.default_node (name, empty_node);
139 fun update_node name f = default_node name #> String_Graph.map_node name f;
142 (* node edits and associated executions *)
144 type overlay = Document_ID.command * (string * string list);
147 Edits of (Document_ID.command option * Document_ID.command option) list |
148 Deps of node_header |
149 Perspective of bool * Document_ID.command list * overlay list;
151 type edit = string * node_edit;
153 val after_entry = Entries.get_after o get_entries;
155 fun lookup_entry node id =
156 (case Entries.lookup (get_entries node) id of
158 | SOME (exec, _) => exec);
160 fun the_entry node id =
161 (case Entries.lookup (get_entries node) id of
162 NONE => err_undef "command entry" id
163 | SOME (exec, _) => exec);
165 fun the_default_entry node (SOME id) = (id, the_default Command.no_exec (the_entry node id))
166 | the_default_entry _ NONE = (Document_ID.none, Command.no_exec);
168 fun assign_entry (command_id, exec) node =
169 if is_none (Entries.lookup (get_entries node) command_id) then node
170 else map_entries (Entries.update (command_id, exec)) node;
172 fun reset_after id entries =
173 (case Entries.get_after entries id of
175 | SOME next => Entries.update (next, NONE) entries);
177 val edit_node = map_entries o fold
178 (fn (id, SOME id2) => Entries.insert_after id (id2, NONE)
179 | (id, NONE) => Entries.delete_after id #> reset_after id);
182 (* version operations *)
184 val empty_version = Version String_Graph.empty;
186 fun nodes_of (Version nodes) = nodes;
187 val node_of = get_node o nodes_of;
189 fun cycle_msg names = "Cyclic dependency of " ^ space_implode " via " (map quote names);
191 fun edit_nodes (name, node_edit) (Version nodes) =
194 Edits edits => update_node name (edit_node edits) nodes
195 | Deps (master, header, errors) =>
197 val imports = map fst (#imports header);
199 (Thy_Header.define_keywords header; errors)
200 handle ERROR msg => errors @ [msg];
203 |> fold default_node imports;
205 |> String_Graph.Keys.fold
206 (fn dep => String_Graph.del_edge (dep, name)) (String_Graph.imm_preds nodes1 name);
207 val (nodes3, errors2) =
208 (String_Graph.add_deps_acyclic (name, imports) nodes2, errors1)
209 handle String_Graph.CYCLES cs => (nodes2, errors1 @ map cycle_msg cs);
210 in String_Graph.map_node name (set_header (master, header, errors2)) nodes3 end
211 | Perspective perspective => update_node name (set_perspective perspective) nodes);
213 fun put_node (name, node) (Version nodes) =
214 Version (update_node name (K node) nodes);
220 (** main state -- document structure and execution process **)
223 (string * string * string option) Exn.result; (* file node name, path, raw digest*)
226 {version_id: Document_ID.version, (*static version id*)
227 execution_id: Document_ID.execution, (*dynamic execution id*)
228 delay_request: unit future, (*pending event timer request*)
229 frontier: Future.task Symtab.table}; (*node name -> running execution task*)
231 val no_execution: execution =
232 {version_id = Document_ID.none, execution_id = Document_ID.none,
233 delay_request = Future.value (), frontier = Symtab.empty};
235 fun new_execution version_id delay_request frontier : execution =
236 {version_id = version_id, execution_id = Execution.start (),
237 delay_request = delay_request, frontier = frontier};
239 abstype state = State of
240 {versions: version Inttab.table, (*version id -> document content*)
241 blobs: (SHA1.digest * string list) Symtab.table, (*raw digest -> digest, lines*)
242 commands: (string * blob_digest list * Token.T list lazy) Inttab.table,
243 (*command id -> name, inlined files, command span*)
244 execution: execution} (*current execution process*)
247 fun make_state (versions, blobs, commands, execution) =
248 State {versions = versions, blobs = blobs, commands = commands, execution = execution};
250 fun map_state f (State {versions, blobs, commands, execution}) =
251 make_state (f (versions, blobs, commands, execution));
254 make_state (Inttab.make [(Document_ID.none, empty_version)],
255 Symtab.empty, Inttab.empty, no_execution);
258 (* document versions *)
260 fun define_version version_id version =
261 map_state (fn (versions, blobs, commands, {delay_request, frontier, ...}) =>
263 val versions' = Inttab.update_new (version_id, version) versions
264 handle Inttab.DUP dup => err_dup "document version" dup;
265 val execution' = new_execution version_id delay_request frontier;
266 in (versions', blobs, commands, execution') end);
268 fun the_version (State {versions, ...}) version_id =
269 (case Inttab.lookup versions version_id of
270 NONE => err_undef "document version" version_id
271 | SOME version => version);
273 fun delete_version version_id versions =
274 Inttab.delete version_id versions
275 handle Inttab.UNDEF _ => err_undef "document version" version_id;
280 fun define_blob digest text =
281 map_state (fn (versions, blobs, commands, execution) =>
283 val sha1_digest = SHA1.digest text;
285 digest = SHA1.rep sha1_digest orelse
286 error ("Ill-defined blob: bad SHA1 digest " ^ digest ^ " vs. " ^ SHA1.rep sha1_digest);
287 val blobs' = Symtab.update (digest, (sha1_digest, split_lines text)) blobs;
288 in (versions, blobs', commands, execution) end);
290 fun the_blob (State {blobs, ...}) digest =
291 (case Symtab.lookup blobs digest of
292 NONE => error ("Undefined blob: " ^ digest)
293 | SOME content => content);
295 fun resolve_blob state (blob_digest: blob_digest) =
296 blob_digest |> Exn.map_result (fn (file, path, raw_digest) =>
297 (file, path, Option.map (the_blob state) raw_digest));
302 fun define_command command_id name command_blobs text =
303 map_state (fn (versions, blobs, commands, execution) =>
305 val id = Document_ID.print command_id;
308 Position.setmp_thread_data (Position.id_only id)
309 (fn () => Thy_Syntax.parse_tokens (Keyword.get_lexicons ()) (Position.id id) text) ());
311 Position.setmp_thread_data (Position.id_only id)
312 (fn () => Output.status (Markup.markup_only Markup.accepted)) ();
314 Inttab.update_new (command_id, (name, command_blobs, span)) commands
315 handle Inttab.DUP dup => err_dup "command" dup;
316 in (versions, blobs, commands', execution) end);
318 fun the_command (State {commands, ...}) command_id =
319 (case Inttab.lookup commands command_id of
320 NONE => err_undef "command" command_id
321 | SOME command => command);
323 val the_command_name = #1 oo the_command;
328 (* remove_versions *)
330 fun remove_versions version_ids state = state |> map_state (fn (versions, _, _, execution) =>
333 member (op =) version_ids (#version_id execution) andalso
334 error ("Attempt to remove execution version " ^ Document_ID.print (#version_id execution));
336 val versions' = fold delete_version version_ids versions;
338 (versions', Inttab.empty) |->
339 Inttab.fold (fn (_, version) => nodes_of version |>
340 String_Graph.fold (fn (_, (node, _)) => node |>
341 iterate_entries (fn ((_, command_id), _) =>
342 SOME o Inttab.insert (K true) (command_id, the_command state command_id))));
344 (commands', Symtab.empty) |->
345 Inttab.fold (fn (_, (_, blobs, _)) => blobs |>
346 fold (fn Exn.Res (_, _, SOME b) => Symtab.update (b, the_blob state b) | _ => I));
348 in (versions', blobs', commands', execution) end);
351 (* document execution *)
353 fun start_execution state = state |> map_state (fn (versions, blobs, commands, execution) =>
354 timeit "Document.start_execution" (fn () =>
356 val {version_id, execution_id, delay_request, frontier} = execution;
358 val delay = seconds (Options.default_real "editor_execution_delay");
360 val _ = Future.cancel delay_request;
361 val delay_request' = Event_Timer.future (Time.+ (Time.now (), delay));
364 nodes_of (the_version state version_id) |> String_Graph.schedule
365 (fn deps => fn (name, node) =>
366 if visible_node node orelse pending_result node then
369 Future.task_of delay_request' :: the_list (Symtab.lookup frontier name);
371 iterate_entries (fn (_, opt_exec) => fn () =>
374 if Execution.is_running execution_id
375 then SOME (Command.exec execution_id exec)
377 | NONE => NONE)) node ()
378 handle exn => if Exn.is_interrupt exn then () (*sic!*) else reraise exn;
380 (singleton o Future.forks)
381 {name = "theory:" ^ name, group = SOME (Future.new_group NONE),
382 deps = more_deps @ map #2 (maps #2 deps),
383 pri = 0, interrupts = false} body;
384 in [(name, Future.task_of future)] end
386 val frontier' = (fold o fold) Symtab.update new_tasks frontier;
388 {version_id = version_id, execution_id = execution_id,
389 delay_request = delay_request', frontier = frontier'};
390 in (versions, blobs, commands, execution') end));
394 (** document update **)
396 (* exec state assignment *)
398 type assign_update = Command.exec option Inttab.table; (*command id -> exec*)
400 val assign_update_empty: assign_update = Inttab.empty;
401 fun assign_update_defined (tab: assign_update) command_id = Inttab.defined tab command_id;
402 fun assign_update_apply (tab: assign_update) node = Inttab.fold assign_entry tab node;
404 fun assign_update_new upd (tab: assign_update) =
405 Inttab.update_new upd tab
406 handle Inttab.DUP dup => err_dup "exec state assignment" dup;
408 fun assign_update_result (tab: assign_update) =
409 Inttab.fold (fn (command_id, exec) => cons (command_id, Command.exec_ids exec)) tab [];
416 fun make_required nodes =
419 String_Graph.fold (fn (a, (node, _)) => P node ? cons a) nodes []
420 |> String_Graph.all_preds nodes
423 val all_visible = all_preds visible_node;
424 val all_required = all_preds required_node;
426 Symtab.fold (fn (a, ()) =>
427 exists (Symtab.defined all_visible) (String_Graph.immediate_succs nodes a) ?
428 Symtab.update (a, ())) all_visible all_required
431 fun loaded_theory name =
432 (case try (unsuffix ".thy") name of
433 SOME a => Thy_Info.lookup_theory a
436 fun init_theory deps node span =
438 val master_dir = master_directory node;
439 val header = read_header node span;
440 val imports = #imports header;
442 imports |> map (fn (import, _) =>
443 (case loaded_theory import of
446 Toplevel.end_theory (Position.file_only import)
447 (case get_result (snd (the (AList.lookup (op =) deps import))) of
448 NONE => Toplevel.toplevel
449 | SOME eval => Command.eval_result_state eval)));
450 val _ = Position.reports (map #2 imports ~~ map Theory.get_markup parents);
451 in Resources.begin_theory master_dir header parents end;
453 fun check_theory full name node =
454 is_some (loaded_theory name) orelse
455 can get_header node andalso (not full orelse is_some (get_result node));
457 fun last_common state node_required node0 node =
459 fun update_flags prev (visible, initial) =
461 val visible' = visible andalso prev <> visible_last node;
462 val initial' = initial andalso
465 | SOME command_id => not (Keyword.is_theory_begin (the_command_name state command_id)));
466 in (visible', initial') end;
468 fun get_common ((prev, command_id), opt_exec) (_, ok, flags, assign_update) =
471 val flags' as (visible', _) = update_flags prev flags;
473 (case (lookup_entry node0 command_id, opt_exec) of
474 (SOME (eval0, _), SOME (eval, _)) =>
475 Command.eval_eq (eval0, eval) andalso
476 (visible' orelse node_required orelse Command.eval_running eval)
478 val assign_update' = assign_update |> ok' ?
480 SOME (eval, prints) =>
482 val command_visible = visible_command node command_id;
483 val command_overlays = overlays node command_id;
484 val command_name = the_command_name state command_id;
486 (case Command.print command_visible command_overlays command_name eval prints of
487 SOME prints' => assign_update_new (command_id, SOME (eval, prints'))
491 in SOME (prev, ok', flags', assign_update') end
493 val (common, ok, flags, assign_update') =
494 iterate_entries get_common node (NONE, true, (true, true), assign_update_empty);
495 val (common', flags') =
497 let val last = Entries.get_after (get_entries node) common
498 in (last, update_flags last flags) end
499 else (common, flags);
500 in (assign_update', common', flags') end;
502 fun illegal_init _ = error "Illegal theory header after end of theory";
504 fun new_exec state node proper_init command_id' (assign_update, command_exec, init) =
505 if not proper_init andalso is_none init then NONE
508 val (_, (eval, _)) = command_exec;
510 val command_visible = visible_command node command_id';
511 val command_overlays = overlays node command_id';
512 val (command_name, blob_digests, span0) = the_command state command_id';
513 val blobs = map (resolve_blob state) blob_digests;
514 val span = Lazy.force span0;
517 Command.eval (fn () => the_default illegal_init init span)
518 (master_directory node) blobs span eval;
519 val prints' = perhaps (Command.print command_visible command_overlays command_name eval') [];
520 val exec' = (eval', prints');
522 val assign_update' = assign_update_new (command_id', SOME exec') assign_update;
523 val init' = if Keyword.is_theory_begin command_name then NONE else init;
524 in SOME (assign_update', (command_id', (eval', prints')), init') end;
526 fun removed_execs node0 (command_id, exec_ids) =
527 subtract (op =) exec_ids (Command.exec_ids (lookup_entry node0 command_id));
531 fun update old_version_id new_version_id edits state =
533 val old_version = the_version state old_version_id;
534 val new_version = timeit "Document.edit_nodes" (fn () => fold edit_nodes edits old_version);
536 val nodes = nodes_of new_version;
537 val required = make_required nodes;
538 val required0 = make_required (nodes_of old_version);
539 val edited = fold (fn (name, _) => Symtab.update (name, ())) edits Symtab.empty;
541 val updated = timeit "Document.update" (fn () =>
542 nodes |> String_Graph.schedule
543 (fn deps => fn (name, node) =>
544 (singleton o Future.forks)
545 {name = "Document.update", group = NONE,
546 deps = map (Future.task_of o #2) deps, pri = 1, interrupts = false}
547 (fn () => timeit ("Document.update " ^ name) (fn () =>
549 val imports = map (apsnd Future.join) deps;
550 val imports_result_changed = exists (#4 o #1 o #2) imports;
551 val node_required = Symtab.defined required name;
553 if Symtab.defined edited name orelse visible_node node orelse
554 imports_result_changed orelse Symtab.defined required0 name <> node_required
557 val node0 = node_of old_version name;
558 val init = init_theory imports node;
560 check_theory false name node andalso
561 forall (fn (name, (_, node)) => check_theory true name node) imports;
563 val (print_execs, common, (still_visible, initial)) =
564 if imports_result_changed then (assign_update_empty, NONE, (true, true))
565 else last_common state node_required node0 node;
566 val common_command_exec = the_default_entry node common;
568 val (updated_execs, (command_id', (eval', _)), _) =
569 (print_execs, common_command_exec, if initial then SOME init else NONE)
570 |> (still_visible orelse node_required) ?
571 iterate_entries_after common
572 (fn ((prev, id), _) => fn res =>
573 if not node_required andalso prev = visible_last node then NONE
574 else new_exec state node proper_init id res) node;
577 (node0, updated_execs) |-> iterate_entries_after common
578 (fn ((_, command_id0), exec0) => fn res =>
579 if is_none exec0 then NONE
580 else if assign_update_defined updated_execs command_id0 then SOME res
581 else SOME (assign_update_new (command_id0, NONE) res));
584 if command_id' = Document_ID.none then NONE else SOME command_id';
586 if is_none last_exec orelse is_some (after_entry node last_exec) then NONE
589 val assign_update = assign_update_result assigned_execs;
590 val removed = maps (removed_execs node0) assign_update;
591 val _ = List.app Execution.cancel removed;
594 |> assign_update_apply assigned_execs
595 |> set_result result;
596 val assigned_node = SOME (name, node');
597 val result_changed = changed_result node0 node';
598 in ((removed, assign_update, assigned_node, result_changed), node') end
599 else (([], [], NONE, false), node)
601 |> Future.joins |> map #1);
603 val removed = maps #1 updated;
604 val assign_update = maps #2 updated;
605 val assigned_nodes = map_filter #3 updated;
608 |> define_version new_version_id (fold put_node assigned_nodes new_version);
610 in (removed, assign_update, state') end;
618 val global_state = Synchronized.var "Document.global_state" init_state;
620 fun state () = Synchronized.value global_state;
621 val change_state = Synchronized.change global_state;