wenzelm@38150: (* Title: Pure/PIDE/document.ML wenzelm@38150: Author: Makarius wenzelm@38150: wenzelm@38150: Document as collection of named nodes, each consisting of an editable wenzelm@52536: list of commands, associated with asynchronous execution process. wenzelm@38150: *) wenzelm@38150: wenzelm@38150: signature DOCUMENT = wenzelm@38150: sig wenzelm@48707: type node_header = string * Thy_Header.header * string list wenzelm@44157: datatype node_edit = wenzelm@48755: Clear | (* FIXME unused !? *) wenzelm@52530: Edits of (Document_ID.command option * Document_ID.command option) list | wenzelm@48707: Deps of node_header | wenzelm@52530: Perspective of Document_ID.command list wenzelm@44156: type edit = string * node_edit wenzelm@38418: type state wenzelm@38418: val init_state: state wenzelm@52530: val define_command: Document_ID.command -> string -> string -> state -> state wenzelm@52530: val remove_versions: Document_ID.version list -> state -> state wenzelm@47343: val discontinue_execution: state -> unit wenzelm@47404: val cancel_execution: state -> unit wenzelm@47420: val start_execution: state -> unit wenzelm@47628: val timing: bool Unsynchronized.ref wenzelm@52530: val update: Document_ID.version -> Document_ID.version -> edit list -> state -> wenzelm@52530: (Document_ID.command * Document_ID.exec list) list * state wenzelm@43713: val state: unit -> state wenzelm@43713: val change_state: (state -> state) -> unit wenzelm@38150: end; wenzelm@38150: wenzelm@38150: structure Document: DOCUMENT = wenzelm@38150: struct wenzelm@38150: wenzelm@38418: (** document structure **) wenzelm@38418: wenzelm@52530: fun err_dup kind id = error ("Duplicate " ^ kind ^ ": " ^ Document_ID.print id); wenzelm@52530: fun err_undef kind id = error ("Undefined " ^ kind ^ ": " ^ Document_ID.print id); wenzelm@52530: wenzelm@48707: type node_header = string * Thy_Header.header * string list; wenzelm@52530: type perspective = Inttab.set * Document_ID.command option; wenzelm@52530: structure Entries = Linear_Set(type key = Document_ID.command val ord = int_ord); wenzelm@38448: wenzelm@43668: abstype node = Node of wenzelm@48707: {header: node_header, (*master directory, theory header, errors*) wenzelm@52508: perspective: perspective, (*visible commands, last visible command*) wenzelm@52532: entries: Command.exec option Entries.T * bool, (*command entries with excecutions, stable*) wenzelm@52526: result: Command.eval option} (*result of last execution*) wenzelm@50862: and version = Version of node String_Graph.T (*development graph wrt. static imports*) wenzelm@38418: with wenzelm@38418: wenzelm@47406: fun make_node (header, perspective, entries, result) = wenzelm@47406: Node {header = header, perspective = perspective, entries = entries, result = result}; wenzelm@43668: wenzelm@47406: fun map_node f (Node {header, perspective, entries, result}) = wenzelm@47406: make_node (f (header, perspective, entries, result)); wenzelm@38418: wenzelm@44476: fun make_perspective command_ids : perspective = wenzelm@52526: (Inttab.make_set command_ids, try List.last command_ids); wenzelm@44441: wenzelm@51294: val no_header = ("", Thy_Header.make ("", Position.none) [] [], ["Bad theory header"]); wenzelm@44441: val no_perspective = make_perspective []; wenzelm@44386: wenzelm@49064: val empty_node = make_node (no_header, no_perspective, (Entries.empty, true), NONE); wenzelm@49064: val clear_node = wenzelm@49064: map_node (fn (header, _, _, _) => (header, no_perspective, (Entries.empty, true), NONE)); wenzelm@44386: wenzelm@44386: wenzelm@44386: (* basic components *) wenzelm@44386: wenzelm@48927: fun set_header header = wenzelm@48927: map_node (fn (_, perspective, entries, result) => (header, perspective, entries, result)); wenzelm@48927: wenzelm@48707: fun get_header (Node {header = (master, header, errors), ...}) = wenzelm@48707: if null errors then (master, header) wenzelm@48707: else error (cat_lines errors); wenzelm@48707: wenzelm@48927: fun read_header node span = wenzelm@48927: let wenzelm@51294: val (dir, {name = (name, _), imports, keywords}) = get_header node; wenzelm@48927: val {name = (_, pos), imports = imports', ...} = Thy_Header.read_tokens span; wenzelm@51294: in (dir, Thy_Header.make (name, pos) (map #1 imports ~~ map #2 imports') keywords) end; wenzelm@44180: wenzelm@44385: fun get_perspective (Node {perspective, ...}) = perspective; wenzelm@44441: fun set_perspective ids = wenzelm@47406: map_node (fn (header, _, entries, result) => (header, make_perspective ids, entries, result)); wenzelm@44385: wenzelm@52569: val visible_command = Inttab.defined o #1 o get_perspective; wenzelm@47345: val visible_last = #2 o get_perspective; wenzelm@47345: val visible_node = is_some o visible_last wenzelm@47345: wenzelm@44444: fun map_entries f = wenzelm@49064: map_node (fn (header, perspective, (entries, stable), result) => wenzelm@49064: (header, perspective, (f entries, stable), result)); wenzelm@49064: fun get_entries (Node {entries = (entries, _), ...}) = entries; wenzelm@49064: wenzelm@49064: fun entries_stable stable = wenzelm@49064: map_node (fn (header, perspective, (entries, _), result) => wenzelm@49064: (header, perspective, (entries, stable), result)); wenzelm@49064: fun stable_entries (Node {entries = (_, stable), ...}) = stable; wenzelm@44645: wenzelm@44645: fun iterate_entries f = Entries.iterate NONE f o get_entries; wenzelm@49064: fun iterate_entries_after start f (Node {entries = (entries, _), ...}) = wenzelm@44645: (case Entries.get_after entries start of wenzelm@44645: NONE => I wenzelm@44645: | SOME id => Entries.iterate (SOME id) f entries); wenzelm@43668: wenzelm@47339: fun get_result (Node {result, ...}) = result; wenzelm@44385: fun set_result result = wenzelm@47406: map_node (fn (header, perspective, entries, _) => (header, perspective, entries, result)); wenzelm@44180: wenzelm@52566: fun changed_result node node' = wenzelm@52566: (case (get_result node, get_result node') of wenzelm@52566: (SOME eval, SOME eval') => #exec_id eval <> #exec_id eval' wenzelm@52566: | (NONE, NONE) => false wenzelm@52566: | _ => true); wenzelm@52566: wenzelm@52533: fun finished_theory node = wenzelm@52534: (case Exn.capture (Command.eval_result_state o the) (get_result node) of wenzelm@52534: Exn.Res st => can (Toplevel.end_theory Position.none) st wenzelm@52533: | _ => false); wenzelm@52533: wenzelm@50862: fun get_node nodes name = String_Graph.get_node nodes name wenzelm@50862: handle String_Graph.UNDEF _ => empty_node; wenzelm@50862: fun default_node name = String_Graph.default_node (name, empty_node); wenzelm@50862: fun update_node name f = default_node name #> String_Graph.map_node name f; wenzelm@38418: wenzelm@38418: wenzelm@38448: (* node edits and associated executions *) wenzelm@38150: wenzelm@44157: datatype node_edit = wenzelm@44185: Clear | wenzelm@52530: Edits of (Document_ID.command option * Document_ID.command option) list | wenzelm@48707: Deps of node_header | wenzelm@52530: Perspective of Document_ID.command list; wenzelm@44157: wenzelm@44156: type edit = string * node_edit; wenzelm@38448: wenzelm@49064: val after_entry = Entries.get_after o get_entries; wenzelm@44479: wenzelm@49064: fun lookup_entry node id = wenzelm@49064: (case Entries.lookup (get_entries node) id of wenzelm@44480: NONE => NONE wenzelm@44480: | SOME (exec, _) => exec); wenzelm@44480: wenzelm@49064: fun the_entry node id = wenzelm@49064: (case Entries.lookup (get_entries node) id of wenzelm@38448: NONE => err_undef "command entry" id wenzelm@44476: | SOME (exec, _) => exec); wenzelm@38150: wenzelm@52532: fun the_default_entry node (SOME id) = (id, the_default Command.no_exec (the_entry node id)) wenzelm@52532: | the_default_entry _ NONE = (Document_ID.none, Command.no_exec); wenzelm@44476: wenzelm@52566: fun assign_entry (command_id, exec) node = wenzelm@52566: if is_none (Entries.lookup (get_entries node) command_id) then node wenzelm@52566: else map_entries (Entries.update (command_id, exec)) node; wenzelm@38418: wenzelm@38448: fun reset_after id entries = wenzelm@38448: (case Entries.get_after entries id of wenzelm@38448: NONE => entries wenzelm@38448: | SOME next => Entries.update (next, NONE) entries); wenzelm@38448: wenzelm@43668: val edit_node = map_entries o fold wenzelm@43668: (fn (id, SOME id2) => Entries.insert_after id (id2, NONE) wenzelm@43668: | (id, NONE) => Entries.delete_after id #> reset_after id); wenzelm@38418: wenzelm@38418: wenzelm@38418: (* version operations *) wenzelm@38418: wenzelm@50862: val empty_version = Version String_Graph.empty; wenzelm@44185: wenzelm@38418: fun nodes_of (Version nodes) = nodes; wenzelm@44185: val node_of = get_node o nodes_of; wenzelm@38418: wenzelm@44185: fun cycle_msg names = "Cyclic dependency of " ^ space_implode " via " (map quote names); wenzelm@44180: wenzelm@44157: fun edit_nodes (name, node_edit) (Version nodes) = wenzelm@44157: Version wenzelm@44436: (case node_edit of wenzelm@47406: Clear => update_node name clear_node nodes wenzelm@47406: | Edits edits => update_node name (edit_node edits) nodes wenzelm@48707: | Deps (master, header, errors) => wenzelm@44436: let wenzelm@48927: val imports = map fst (#imports header); wenzelm@48707: val errors1 = wenzelm@48707: (Thy_Header.define_keywords header; errors) wenzelm@48707: handle ERROR msg => errors @ [msg]; wenzelm@44436: val nodes1 = nodes wenzelm@44436: |> default_node name wenzelm@48927: |> fold default_node imports; wenzelm@44436: val nodes2 = nodes1 wenzelm@50862: |> String_Graph.Keys.fold wenzelm@50862: (fn dep => String_Graph.del_edge (dep, name)) (String_Graph.imm_preds nodes1 name); wenzelm@48707: val (nodes3, errors2) = wenzelm@50862: (String_Graph.add_deps_acyclic (name, imports) nodes2, errors1) wenzelm@50862: handle String_Graph.CYCLES cs => (nodes2, errors1 @ map cycle_msg cs); wenzelm@50862: in String_Graph.map_node name (set_header (master, header, errors2)) nodes3 end wenzelm@47406: | Perspective perspective => update_node name (set_perspective perspective) nodes); wenzelm@44444: wenzelm@44222: fun put_node (name, node) (Version nodes) = wenzelm@44222: Version (update_node name (K node) nodes); wenzelm@38418: wenzelm@38150: end; wenzelm@38150: wenzelm@38418: wenzelm@38418: wenzelm@47389: (** main state -- document structure and execution process **) wenzelm@38418: wenzelm@52536: type execution = wenzelm@52536: {version_id: Document_ID.version, wenzelm@52536: group: Future.group, wenzelm@52536: running: bool Unsynchronized.ref}; wenzelm@52536: wenzelm@52536: val no_execution = wenzelm@52536: {version_id = Document_ID.none, wenzelm@52536: group = Future.new_group NONE, wenzelm@52536: running = Unsynchronized.ref false}; wenzelm@52536: wenzelm@38418: abstype state = State of wenzelm@52536: {versions: version Inttab.table, (*version id -> document content*) wenzelm@52587: commands: (string * Token.T list lazy) Inttab.table, (*command id -> named command span*) wenzelm@52536: execution: execution} (*current execution process*) wenzelm@38418: with wenzelm@38418: wenzelm@44674: fun make_state (versions, commands, execution) = wenzelm@44674: State {versions = versions, commands = commands, execution = execution}; wenzelm@38418: wenzelm@44674: fun map_state f (State {versions, commands, execution}) = wenzelm@44674: make_state (f (versions, commands, execution)); wenzelm@38418: wenzelm@38418: val init_state = wenzelm@52536: make_state (Inttab.make [(Document_ID.none, empty_version)], Inttab.empty, no_execution); wenzelm@38418: wenzelm@47420: fun execution_of (State {execution, ...}) = execution; wenzelm@47420: wenzelm@38418: wenzelm@38418: (* document versions *) wenzelm@38418: wenzelm@52536: fun define_version version_id version = wenzelm@47420: map_state (fn (versions, commands, _) => wenzelm@47420: let wenzelm@52536: val versions' = Inttab.update_new (version_id, version) versions wenzelm@47420: handle Inttab.DUP dup => err_dup "document version" dup; wenzelm@52536: val execution' = wenzelm@52536: {version_id = version_id, wenzelm@52536: group = Future.new_group NONE, wenzelm@52536: running = Unsynchronized.ref true}; wenzelm@47420: in (versions', commands, execution') end); wenzelm@38418: wenzelm@52536: fun the_version (State {versions, ...}) version_id = wenzelm@52536: (case Inttab.lookup versions version_id of wenzelm@52536: NONE => err_undef "document version" version_id wenzelm@38418: | SOME version => version); wenzelm@38418: wenzelm@52536: fun delete_version version_id versions = wenzelm@52536: Inttab.delete version_id versions wenzelm@52536: handle Inttab.UNDEF _ => err_undef "document version" version_id; wenzelm@44673: wenzelm@38418: wenzelm@38418: (* commands *) wenzelm@38418: wenzelm@52536: fun define_command command_id name text = wenzelm@44674: map_state (fn (versions, commands, execution) => wenzelm@38418: let wenzelm@52536: val id = Document_ID.print command_id; wenzelm@52536: val span = wenzelm@52536: Lazy.lazy (fn () => wenzelm@52536: Position.setmp_thread_data (Position.id_only id) wenzelm@52536: (fn () => Thy_Syntax.parse_tokens (Keyword.get_lexicons ()) (Position.id id) text) ()); wenzelm@47395: val _ = wenzelm@52536: Position.setmp_thread_data (Position.id_only id) wenzelm@50201: (fn () => Output.status (Markup.markup_only Markup.accepted)) (); wenzelm@38418: val commands' = wenzelm@52536: Inttab.update_new (command_id, (name, span)) commands wenzelm@38418: handle Inttab.DUP dup => err_dup "command" dup; wenzelm@44674: in (versions, commands', execution) end); wenzelm@38418: wenzelm@52536: fun the_command (State {commands, ...}) command_id = wenzelm@52536: (case Inttab.lookup commands command_id of wenzelm@52536: NONE => err_undef "command" command_id wenzelm@44644: | SOME command => command); wenzelm@38418: wenzelm@52569: val the_command_name = #1 oo the_command; wenzelm@52569: wenzelm@47420: end; wenzelm@38418: wenzelm@52566: wenzelm@52566: (* remove_versions *) wenzelm@52566: wenzelm@52536: fun remove_versions version_ids state = state |> map_state (fn (versions, _, execution) => wenzelm@47420: let wenzelm@52536: val _ = wenzelm@52536: member (op =) version_ids (#version_id execution) andalso wenzelm@52536: error ("Attempt to remove execution version " ^ Document_ID.print (#version_id execution)); wenzelm@47343: wenzelm@52536: val versions' = fold delete_version version_ids versions; wenzelm@47420: val commands' = wenzelm@47420: (versions', Inttab.empty) |-> wenzelm@47420: Inttab.fold (fn (_, version) => nodes_of version |> wenzelm@50862: String_Graph.fold (fn (_, (node, _)) => node |> wenzelm@52536: iterate_entries (fn ((_, command_id), _) => wenzelm@52536: SOME o Inttab.insert (K true) (command_id, the_command state command_id)))); wenzelm@47420: in (versions', commands', execution) end); wenzelm@38418: wenzelm@38418: wenzelm@38888: wenzelm@47420: (** document execution **) wenzelm@47389: wenzelm@52536: val discontinue_execution = execution_of #> (fn {running, ...} => running := false); wenzelm@52536: val cancel_execution = execution_of #> (fn {group, ...} => Future.cancel_group group); wenzelm@52536: val terminate_execution = execution_of #> (fn {group, ...} => Future.terminate group); wenzelm@47345: wenzelm@47420: fun start_execution state = wenzelm@47420: let wenzelm@52536: val {version_id, group, running} = execution_of state; wenzelm@47420: val _ = wenzelm@52573: if not (! running) orelse Task_Queue.is_canceled group then [] wenzelm@52573: else wenzelm@52573: nodes_of (the_version state version_id) |> String_Graph.schedule wenzelm@52573: (fn deps => fn (name, node) => wenzelm@52573: if not (visible_node node) andalso finished_theory node then wenzelm@52573: Future.value () wenzelm@52573: else wenzelm@52573: (singleton o Future.forks) wenzelm@52573: {name = "theory:" ^ name, group = SOME (Future.new_group (SOME group)), wenzelm@52573: deps = map (Future.task_of o #2) deps, pri = ~2, interrupts = false} wenzelm@52573: (fn () => wenzelm@52573: iterate_entries (fn (_, opt_exec) => fn () => wenzelm@52573: (case opt_exec of wenzelm@52573: SOME exec => if ! running then SOME (Command.execute exec) else NONE wenzelm@52573: | NONE => NONE)) node ())); wenzelm@47420: in () end; wenzelm@47345: wenzelm@47345: wenzelm@47420: wenzelm@47420: (** document update **) wenzelm@38418: wenzelm@52566: (* exec state assignment *) wenzelm@52566: wenzelm@52566: type assign_update = Command.exec option Inttab.table; (*command id -> exec*) wenzelm@52566: wenzelm@52566: val assign_update_empty: assign_update = Inttab.empty; wenzelm@52566: fun assign_update_null (tab: assign_update) = Inttab.is_empty tab; wenzelm@52566: fun assign_update_defined (tab: assign_update) command_id = Inttab.defined tab command_id; wenzelm@52566: fun assign_update_apply (tab: assign_update) node = Inttab.fold assign_entry tab node; wenzelm@52566: wenzelm@52566: fun assign_update_new upd (tab: assign_update) = wenzelm@52566: Inttab.update_new upd tab wenzelm@52566: handle Inttab.DUP dup => err_dup "exec state assignment" dup; wenzelm@52566: wenzelm@52566: fun assign_update_result (tab: assign_update) = wenzelm@52566: Inttab.fold (fn (command_id, exec) => cons (command_id, Command.exec_ids exec)) tab []; wenzelm@52566: wenzelm@52566: wenzelm@52566: (* update *) wenzelm@52566: wenzelm@47628: val timing = Unsynchronized.ref false; wenzelm@47628: fun timeit msg e = cond_timeit (! timing) msg e; wenzelm@47628: wenzelm@38418: local wenzelm@38418: wenzelm@47335: fun make_required nodes = wenzelm@47335: let wenzelm@47335: val all_visible = wenzelm@50862: String_Graph.fold (fn (a, (node, _)) => visible_node node ? cons a) nodes [] wenzelm@50862: |> String_Graph.all_preds nodes wenzelm@47345: |> map (rpair ()) |> Symtab.make; wenzelm@47345: wenzelm@47335: val required = wenzelm@47345: Symtab.fold (fn (a, ()) => wenzelm@50862: exists (Symtab.defined all_visible) (String_Graph.immediate_succs nodes a) ? wenzelm@47345: Symtab.update (a, ())) all_visible Symtab.empty; wenzelm@47335: in Symtab.defined required end; wenzelm@47335: wenzelm@48927: fun init_theory deps node span = wenzelm@47335: let wenzelm@49173: (* FIXME provide files via Isabelle/Scala, not master_dir *) wenzelm@48927: val (dir, header) = read_header node span; wenzelm@47335: val master_dir = wenzelm@48735: (case try Url.explode dir of wenzelm@48735: SOME (Url.File path) => path wenzelm@47335: | _ => Path.current); wenzelm@48927: val imports = #imports header; wenzelm@47335: val parents = wenzelm@48927: imports |> map (fn (import, _) => wenzelm@47407: (case Thy_Info.lookup_theory import of wenzelm@47335: SOME thy => thy wenzelm@47335: | NONE => wenzelm@47339: Toplevel.end_theory (Position.file_only import) wenzelm@52536: (case get_result (snd (the (AList.lookup (op =) deps import))) of wenzelm@52536: NONE => Toplevel.toplevel wenzelm@52536: | SOME eval => Command.eval_result_state eval))); wenzelm@48927: val _ = Position.reports (map #2 imports ~~ map Theory.get_markup parents); wenzelm@47335: in Thy_Load.begin_theory master_dir header parents end; wenzelm@47335: wenzelm@47407: fun check_theory full name node = wenzelm@47407: is_some (Thy_Info.lookup_theory name) orelse wenzelm@48707: can get_header node andalso (not full orelse is_some (get_result node)); wenzelm@47335: wenzelm@52569: fun last_common state node0 node = wenzelm@44482: let wenzelm@44645: fun update_flags prev (visible, initial) = wenzelm@44645: let wenzelm@52569: val visible' = visible andalso prev <> visible_last node; wenzelm@44645: val initial' = initial andalso wenzelm@44645: (case prev of wenzelm@44645: NONE => true wenzelm@52569: | SOME command_id => not (Keyword.is_theory_begin (the_command_name state command_id))); wenzelm@44645: in (visible', initial') end; wenzelm@52569: wenzelm@52569: fun get_common ((prev, command_id), opt_exec) (_, same, flags, assign_update) = wenzelm@47630: if same then wenzelm@47630: let wenzelm@47630: val flags' = update_flags prev flags; wenzelm@47630: val same' = wenzelm@52586: (case (lookup_entry node0 command_id, opt_exec) of wenzelm@52586: (SOME (eval0, _), SOME (eval, _)) => Command.same_eval (eval0, eval) wenzelm@52586: | _ => false); wenzelm@52569: val assign_update' = assign_update |> wenzelm@52569: (case opt_exec of wenzelm@52569: SOME (eval, prints) => wenzelm@52569: let wenzelm@52569: val command_visible = visible_command node command_id; wenzelm@52569: val command_name = the_command_name state command_id; wenzelm@52569: in wenzelm@52570: (case Command.print command_visible command_name eval prints of wenzelm@52569: SOME prints' => assign_update_new (command_id, SOME (eval, prints')) wenzelm@52569: | NONE => I) wenzelm@52569: end wenzelm@52569: | NONE => I); wenzelm@52569: in SOME (prev, same', flags', assign_update') end wenzelm@47630: else NONE; wenzelm@52569: val (common, same, flags, assign_update') = wenzelm@52569: iterate_entries get_common node (NONE, true, (true, true), assign_update_empty); wenzelm@52569: val (common', flags') = wenzelm@52569: if same then wenzelm@52569: let val last = Entries.get_after (get_entries node) common wenzelm@52569: in (last, update_flags last flags) end wenzelm@52569: else (common, flags); wenzelm@52569: in (assign_update', common', flags') end; wenzelm@44645: wenzelm@52534: fun illegal_init _ = error "Illegal theory header after end of theory"; wenzelm@52534: wenzelm@52566: fun new_exec state proper_init command_visible command_id' (assign_update, command_exec, init) = wenzelm@47407: if not proper_init andalso is_none init then NONE wenzelm@44645: else wenzelm@44645: let wenzelm@52534: val (_, (eval, _)) = command_exec; wenzelm@52527: val (command_name, span) = the_command state command_id' ||> Lazy.force; wenzelm@52566: wenzelm@52534: val eval' = Command.eval (fn () => the_default illegal_init init span) span eval; wenzelm@52570: val prints' = perhaps (Command.print command_visible command_name eval') []; wenzelm@52566: val exec' = (eval', prints'); wenzelm@44482: wenzelm@52566: val assign_update' = assign_update_new (command_id', SOME exec') assign_update; wenzelm@52566: val init' = if Keyword.is_theory_begin command_name then NONE else init; wenzelm@52566: in SOME (assign_update', (command_id', (eval', prints')), init') end; wenzelm@38418: in wenzelm@38418: wenzelm@52536: fun update old_version_id new_version_id edits state = wenzelm@38418: let wenzelm@52536: val old_version = the_version state old_version_id; wenzelm@47628: val new_version = timeit "Document.edit_nodes" (fn () => fold edit_nodes edits old_version); wenzelm@38418: wenzelm@44544: val nodes = nodes_of new_version; wenzelm@44544: val is_required = make_required nodes; wenzelm@44544: wenzelm@47628: val _ = timeit "Document.terminate_execution" (fn () => terminate_execution state); wenzelm@47628: val updated = timeit "Document.update" (fn () => wenzelm@50862: nodes |> String_Graph.schedule wenzelm@44199: (fn deps => fn (name, node) => wenzelm@47406: (singleton o Future.forks) wenzelm@47406: {name = "Document.update", group = NONE, wenzelm@47406: deps = map (Future.task_of o #2) deps, pri = 0, interrupts = false} wenzelm@47406: (fn () => wenzelm@47406: let wenzelm@47406: val imports = map (apsnd Future.join) deps; wenzelm@52566: val imports_changed = exists (#3 o #1 o #2) imports; wenzelm@52514: val node_required = is_required name; wenzelm@47406: in wenzelm@52566: if imports_changed orelse AList.defined (op =) edits name orelse wenzelm@49064: not (stable_entries node) orelse not (finished_theory node) wenzelm@47406: then wenzelm@44482: let wenzelm@47406: val node0 = node_of old_version name; wenzelm@48927: val init = init_theory imports node; wenzelm@47407: val proper_init = wenzelm@47407: check_theory false name node andalso wenzelm@47407: forall (fn (name, (_, node)) => check_theory true name node) imports; wenzelm@47406: wenzelm@52569: val (print_execs, common, (still_visible, initial)) = wenzelm@52569: if imports_changed then (assign_update_empty, NONE, (true, true)) wenzelm@52569: else last_common state node0 node; wenzelm@44674: val common_command_exec = the_default_entry node common; wenzelm@44479: wenzelm@52569: val (updated_execs, (command_id', (eval', _)), _) = wenzelm@52569: (print_execs, common_command_exec, if initial then SOME init else NONE) wenzelm@52566: |> (still_visible orelse node_required) ? wenzelm@44645: iterate_entries_after common wenzelm@44482: (fn ((prev, id), _) => fn res => wenzelm@52569: if not node_required andalso prev = visible_last node then NONE wenzelm@52569: else new_exec state proper_init (visible_command node id) id res) node; wenzelm@44479: wenzelm@52566: val assigned_execs = wenzelm@52566: (node0, updated_execs) |-> iterate_entries_after common wenzelm@52566: (fn ((_, command_id0), exec0) => fn res => wenzelm@44482: if is_none exec0 then NONE wenzelm@52566: else if assign_update_defined updated_execs command_id0 then SOME res wenzelm@52566: else SOME (assign_update_new (command_id0, NONE) res)); wenzelm@44480: wenzelm@52530: val last_exec = wenzelm@52530: if command_id' = Document_ID.none then NONE else SOME command_id'; wenzelm@44482: val result = wenzelm@52566: if is_none last_exec orelse is_some (after_entry node last_exec) then NONE wenzelm@52526: else SOME eval'; wenzelm@44480: wenzelm@44482: val node' = node wenzelm@52566: |> assign_update_apply assigned_execs wenzelm@52566: |> entries_stable (assign_update_null updated_execs) wenzelm@47406: |> set_result result; wenzelm@52587: val assigned_node = SOME (name, node'); wenzelm@52566: val result_changed = changed_result node0 node'; wenzelm@52566: in wenzelm@52566: ((assign_update_result assigned_execs, assigned_node, result_changed), node') wenzelm@52566: end wenzelm@52566: else (([], NONE, false), node) wenzelm@47406: end)) wenzelm@47628: |> Future.joins |> map #1); wenzelm@44476: wenzelm@44197: val state' = state wenzelm@52566: |> define_version new_version_id (fold put_node (map_filter #2 updated) new_version); wenzelm@52566: in (maps #1 updated, state') end; wenzelm@38418: wenzelm@38418: end; wenzelm@38418: wenzelm@38418: wenzelm@43713: wenzelm@43713: (** global state **) wenzelm@43713: wenzelm@52508: val global_state = Synchronized.var "Document.global_state" init_state; wenzelm@43713: wenzelm@43713: fun state () = Synchronized.value global_state; wenzelm@43713: val change_state = Synchronized.change global_state; wenzelm@43713: wenzelm@38418: end; wenzelm@38418: