| author | wenzelm | 
| Mon, 02 Dec 2019 13:03:09 +0100 | |
| changeset 71211 | 7d522732b7f2 | 
| parent 70662 | 0f9a4e8ee1ab | 
| child 73101 | 3d5d949cd865 | 
| permissions | -rw-r--r-- | 
| 52605 | 1 | (* Title: Pure/PIDE/execution.ML | 
| 52596 | 2 | Author: Makarius | 
| 3 | ||
| 52606 | 4 | Global management of execution. Unique running execution serves as | 
| 53192 | 5 | barrier for further exploration of forked command execs. | 
| 52596 | 6 | *) | 
| 7 | ||
| 52605 | 8 | signature EXECUTION = | 
| 52596 | 9 | sig | 
| 52606 | 10 | val start: unit -> Document_ID.execution | 
| 11 | val discontinue: unit -> unit | |
| 12 | val is_running: Document_ID.execution -> bool | |
| 68354 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 13 | val active_tasks: string -> Future.task list | 
| 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 14 | val worker_task_active: bool -> string -> unit | 
| 52784 
4ba2e8b9972f
de-assign execs that were not registered as running yet -- observe change of perspective more thoroughly;
 wenzelm parents: 
52775diff
changeset | 15 | val is_running_exec: Document_ID.exec -> bool | 
| 53375 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53193diff
changeset | 16 | val running: Document_ID.execution -> Document_ID.exec -> Future.group list -> bool | 
| 66370 | 17 | val snapshot: Document_ID.exec list -> Future.task list | 
| 66371 
6ce1afc01040
more thorough Execution.join, under the assumption that nested Execution.fork only happens from given exed_ids;
 wenzelm parents: 
66370diff
changeset | 18 | val join: Document_ID.exec list -> unit | 
| 53192 | 19 | val peek: Document_ID.exec -> Future.group list | 
| 52655 
3b2b1ef13979
more careful termination of removed execs, leaving running execs undisturbed;
 wenzelm parents: 
52608diff
changeset | 20 | val cancel: Document_ID.exec -> unit | 
| 56292 | 21 |   type params = {name: string, pos: Position.T, pri: int}
 | 
| 22 | val fork: params -> (unit -> 'a) -> 'a future | |
| 56305 | 23 | val print: params -> (unit -> unit) -> unit | 
| 56292 | 24 | val fork_prints: Document_ID.exec -> unit | 
| 53192 | 25 | val purge: Document_ID.exec list -> unit | 
| 26 | val reset: unit -> Future.group list | |
| 27 | val shutdown: unit -> unit | |
| 52596 | 28 | end; | 
| 29 | ||
| 52605 | 30 | structure Execution: EXECUTION = | 
| 52596 | 31 | struct | 
| 32 | ||
| 52787 
c143ad7811fc
pro-forma Execution.reset, despite lack of final join/commit;
 wenzelm parents: 
52784diff
changeset | 33 | (* global state *) | 
| 
c143ad7811fc
pro-forma Execution.reset, despite lack of final join/commit;
 wenzelm parents: 
52784diff
changeset | 34 | |
| 56292 | 35 | type print = {name: string, pri: int, body: unit -> unit};
 | 
| 68353 | 36 | type execs = (Future.group list * print list) (*active forks, prints*) Inttab.table; | 
| 37 | val init_execs: execs = Inttab.make [(Document_ID.none, ([], []))]; | |
| 38 | ||
| 39 | datatype state = | |
| 40 | State of | |
| 41 |    {execution_id: Document_ID.execution,  (*overall document execution*)
 | |
| 68354 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 42 | nodes: Future.task list Symtab.table, (*active nodes*) | 
| 68353 | 43 | execs: execs}; (*running command execs*) | 
| 44 | ||
| 68354 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 45 | fun make_state (execution_id, nodes, execs) = | 
| 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 46 |   State {execution_id = execution_id, nodes = nodes, execs = execs};
 | 
| 52787 
c143ad7811fc
pro-forma Execution.reset, despite lack of final join/commit;
 wenzelm parents: 
52784diff
changeset | 47 | |
| 68353 | 48 | local | 
| 49 | val state = | |
| 68354 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 50 | Synchronized.var "Execution.state" (make_state (Document_ID.none, Symtab.empty, init_execs)); | 
| 68353 | 51 | in | 
| 52 | ||
| 53 | fun get_state () = let val State args = Synchronized.value state in args end; | |
| 52602 
00170ef1dc39
strictly monotonic Document.update: avoid disruptive cancel_execution, merely discontinue_execution and cancel/terminate old execs individually;
 wenzelm parents: 
52596diff
changeset | 54 | |
| 68353 | 55 | fun change_state_result f = | 
| 68354 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 56 |   Synchronized.change_result state (fn (State {execution_id, nodes, execs}) =>
 | 
| 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 57 | let val (result, args') = f (execution_id, nodes, execs) | 
| 68353 | 58 | in (result, make_state args') end); | 
| 59 | ||
| 60 | fun change_state f = | |
| 68354 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 61 |   Synchronized.change state (fn (State {execution_id, nodes, execs}) =>
 | 
| 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 62 | make_state (f (execution_id, nodes, execs))); | 
| 68353 | 63 | |
| 64 | end; | |
| 56291 
e79f76a48449
added Execution.print: accumulate print operations for some command execution, which are applied later and print time;
 wenzelm parents: 
54678diff
changeset | 65 | |
| 
e79f76a48449
added Execution.print: accumulate print operations for some command execution, which are applied later and print time;
 wenzelm parents: 
54678diff
changeset | 66 | fun unregistered exec_id = "Unregistered execution: " ^ Document_ID.print exec_id; | 
| 53192 | 67 | |
| 52604 | 68 | |
| 52606 | 69 | (* unique running execution *) | 
| 52604 | 70 | |
| 52606 | 71 | fun start () = | 
| 52602 
00170ef1dc39
strictly monotonic Document.update: avoid disruptive cancel_execution, merely discontinue_execution and cancel/terminate old execs individually;
 wenzelm parents: 
52596diff
changeset | 72 | let | 
| 52606 | 73 | val execution_id = Document_ID.make (); | 
| 68354 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 74 | val _ = change_state (fn (_, nodes, execs) => (execution_id, nodes, execs)); | 
| 52606 | 75 | in execution_id end; | 
| 52604 | 76 | |
| 68354 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 77 | fun discontinue () = | 
| 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 78 | change_state (fn (_, nodes, execs) => (Document_ID.none, nodes, execs)); | 
| 52606 | 79 | |
| 68353 | 80 | fun is_running execution_id = | 
| 81 | execution_id = #execution_id (get_state ()); | |
| 52604 | 82 | |
| 83 | ||
| 68354 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 84 | (* active nodes *) | 
| 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 85 | |
| 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 86 | fun active_tasks node_name = | 
| 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 87 | Symtab.lookup_list (#nodes (get_state ())) node_name; | 
| 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 88 | |
| 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 89 | fun worker_task_active insert node_name = | 
| 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 90 | change_state (fn (execution_id, nodes, execs) => | 
| 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 91 | let | 
| 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 92 | val nodes' = nodes | 
| 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 93 | |> (if insert then Symtab.insert_list else Symtab.remove_list) | 
| 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 94 | Task_Queue.eq_task (node_name, the (Future.worker_task ())); | 
| 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 95 | in (execution_id, nodes', execs) end); | 
| 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 96 | |
| 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 97 | |
| 66371 
6ce1afc01040
more thorough Execution.join, under the assumption that nested Execution.fork only happens from given exed_ids;
 wenzelm parents: 
66370diff
changeset | 98 | (* running execs *) | 
| 52604 | 99 | |
| 52784 
4ba2e8b9972f
de-assign execs that were not registered as running yet -- observe change of perspective more thoroughly;
 wenzelm parents: 
52775diff
changeset | 100 | fun is_running_exec exec_id = | 
| 68353 | 101 | Inttab.defined (#execs (get_state ())) exec_id; | 
| 52784 
4ba2e8b9972f
de-assign execs that were not registered as running yet -- observe change of perspective more thoroughly;
 wenzelm parents: 
52775diff
changeset | 102 | |
| 53375 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53193diff
changeset | 103 | fun running execution_id exec_id groups = | 
| 68354 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 104 | change_state_result (fn (execution_id', nodes, execs) => | 
| 56291 
e79f76a48449
added Execution.print: accumulate print operations for some command execution, which are applied later and print time;
 wenzelm parents: 
54678diff
changeset | 105 | let | 
| 59348 
8a6788917b32
do not crash into already running exec, instead join its lazy result in the subsequent step (amending 59f1591a11cb);
 wenzelm parents: 
56333diff
changeset | 106 | val ok = execution_id = execution_id' andalso not (Inttab.defined execs exec_id); | 
| 
8a6788917b32
do not crash into already running exec, instead join its lazy result in the subsequent step (amending 59f1591a11cb);
 wenzelm parents: 
56333diff
changeset | 107 | val execs' = execs |> ok ? Inttab.update (exec_id, (groups, [])); | 
| 68354 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 108 | in (ok, (execution_id', nodes, execs')) end); | 
| 53192 | 109 | |
| 66371 
6ce1afc01040
more thorough Execution.join, under the assumption that nested Execution.fork only happens from given exed_ids;
 wenzelm parents: 
66370diff
changeset | 110 | |
| 
6ce1afc01040
more thorough Execution.join, under the assumption that nested Execution.fork only happens from given exed_ids;
 wenzelm parents: 
66370diff
changeset | 111 | (* exec groups and tasks *) | 
| 
6ce1afc01040
more thorough Execution.join, under the assumption that nested Execution.fork only happens from given exed_ids;
 wenzelm parents: 
66370diff
changeset | 112 | |
| 68353 | 113 | fun exec_groups (execs: execs) exec_id = | 
| 66370 | 114 | (case Inttab.lookup execs exec_id of | 
| 56291 
e79f76a48449
added Execution.print: accumulate print operations for some command execution, which are applied later and print time;
 wenzelm parents: 
54678diff
changeset | 115 | SOME (groups, _) => groups | 
| 
e79f76a48449
added Execution.print: accumulate print operations for some command execution, which are applied later and print time;
 wenzelm parents: 
54678diff
changeset | 116 | | NONE => []); | 
| 53192 | 117 | |
| 68379 | 118 | fun snapshot [] = [] | 
| 119 | | snapshot exec_ids = | |
| 120 | change_state_result | |
| 121 | (`(fn (_, _, execs) => Future.snapshot (maps (exec_groups execs) exec_ids))); | |
| 66370 | 122 | |
| 66371 
6ce1afc01040
more thorough Execution.join, under the assumption that nested Execution.fork only happens from given exed_ids;
 wenzelm parents: 
66370diff
changeset | 123 | fun join exec_ids = | 
| 
6ce1afc01040
more thorough Execution.join, under the assumption that nested Execution.fork only happens from given exed_ids;
 wenzelm parents: 
66370diff
changeset | 124 | (case snapshot exec_ids of | 
| 
6ce1afc01040
more thorough Execution.join, under the assumption that nested Execution.fork only happens from given exed_ids;
 wenzelm parents: 
66370diff
changeset | 125 | [] => () | 
| 
6ce1afc01040
more thorough Execution.join, under the assumption that nested Execution.fork only happens from given exed_ids;
 wenzelm parents: 
66370diff
changeset | 126 | | tasks => | 
| 
6ce1afc01040
more thorough Execution.join, under the assumption that nested Execution.fork only happens from given exed_ids;
 wenzelm parents: 
66370diff
changeset | 127 | ((singleton o Future.forks) | 
| 
6ce1afc01040
more thorough Execution.join, under the assumption that nested Execution.fork only happens from given exed_ids;
 wenzelm parents: 
66370diff
changeset | 128 |         {name = "Execution.join", group = SOME (Future.new_group NONE),
 | 
| 
6ce1afc01040
more thorough Execution.join, under the assumption that nested Execution.fork only happens from given exed_ids;
 wenzelm parents: 
66370diff
changeset | 129 | deps = tasks, pri = 0, interrupts = false} I | 
| 
6ce1afc01040
more thorough Execution.join, under the assumption that nested Execution.fork only happens from given exed_ids;
 wenzelm parents: 
66370diff
changeset | 130 | |> Future.join; join exec_ids)); | 
| 
6ce1afc01040
more thorough Execution.join, under the assumption that nested Execution.fork only happens from given exed_ids;
 wenzelm parents: 
66370diff
changeset | 131 | |
| 68353 | 132 | fun peek exec_id = exec_groups (#execs (get_state ())) exec_id; | 
| 66370 | 133 | |
| 53192 | 134 | fun cancel exec_id = List.app Future.cancel_group (peek exec_id); | 
| 135 | ||
| 136 | ||
| 137 | (* fork *) | |
| 138 | ||
| 139 | fun status task markups = | |
| 56296 
5413b6379c0e
less markup by default -- this is stored persistently in Isabelle/Scala;
 wenzelm parents: 
56292diff
changeset | 140 | let | 
| 
5413b6379c0e
less markup by default -- this is stored persistently in Isabelle/Scala;
 wenzelm parents: 
56292diff
changeset | 141 | val props = | 
| 
5413b6379c0e
less markup by default -- this is stored persistently in Isabelle/Scala;
 wenzelm parents: 
56292diff
changeset | 142 | if ! Multithreading.trace >= 2 | 
| 
5413b6379c0e
less markup by default -- this is stored persistently in Isabelle/Scala;
 wenzelm parents: 
56292diff
changeset | 143 | then [(Markup.taskN, Task_Queue.str_of_task task)] else []; | 
| 70662 | 144 | in Output.status (map (Markup.markup_only o Markup.properties props) markups) end; | 
| 53192 | 145 | |
| 56292 | 146 | type params = {name: string, pos: Position.T, pri: int};
 | 
| 147 | ||
| 148 | fun fork ({name, pos, pri}: params) e =
 | |
| 62923 | 149 | Thread_Attributes.uninterruptible (fn _ => Position.setmp_thread_data pos (fn () => | 
| 53192 | 150 | let | 
| 151 | val exec_id = the_default 0 (Position.parse_id pos); | |
| 152 | val group = Future.worker_subgroup (); | |
| 68354 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 153 | val _ = change_state (fn (execution_id, nodes, execs) => | 
| 56291 
e79f76a48449
added Execution.print: accumulate print operations for some command execution, which are applied later and print time;
 wenzelm parents: 
54678diff
changeset | 154 | (case Inttab.lookup execs exec_id of | 
| 
e79f76a48449
added Execution.print: accumulate print operations for some command execution, which are applied later and print time;
 wenzelm parents: 
54678diff
changeset | 155 | SOME (groups, prints) => | 
| 68354 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 156 | let val execs' = Inttab.update (exec_id, (group :: groups, prints)) execs | 
| 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 157 | in (execution_id, nodes, execs') end | 
| 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 158 | | NONE => raise Fail (unregistered exec_id))); | 
| 52655 
3b2b1ef13979
more careful termination of removed execs, leaving running execs undisturbed;
 wenzelm parents: 
52608diff
changeset | 159 | |
| 53192 | 160 | val future = | 
| 161 | (singleton o Future.forks) | |
| 162 |           {name = name, group = SOME group, deps = [], pri = pri, interrupts = false}
 | |
| 163 | (fn () => | |
| 164 | let | |
| 165 | val task = the (Future.worker_task ()); | |
| 166 | val _ = status task [Markup.running]; | |
| 167 | val result = | |
| 168 | Exn.capture (Future.interruptible_task e) () | |
| 61077 
06cca32aa519
thread context for exceptions from forks, e.g. relevant when printing errors;
 wenzelm parents: 
59467diff
changeset | 169 | |> Future.identify_result pos | 
| 
06cca32aa519
thread context for exceptions from forks, e.g. relevant when printing errors;
 wenzelm parents: 
59467diff
changeset | 170 | |> Exn.map_exn Runtime.thread_context; | 
| 68880 
8b98db8fd183
clarified bracketing of messages: [forked [running finished] joined];
 wenzelm parents: 
68379diff
changeset | 171 | val errors = | 
| 
8b98db8fd183
clarified bracketing of messages: [forked [running finished] joined];
 wenzelm parents: 
68379diff
changeset | 172 | Exn.capture (fn () => | 
| 
8b98db8fd183
clarified bracketing of messages: [forked [running finished] joined];
 wenzelm parents: 
68379diff
changeset | 173 | (case result of | 
| 
8b98db8fd183
clarified bracketing of messages: [forked [running finished] joined];
 wenzelm parents: 
68379diff
changeset | 174 | Exn.Exn exn => | 
| 
8b98db8fd183
clarified bracketing of messages: [forked [running finished] joined];
 wenzelm parents: 
68379diff
changeset | 175 | (status task [Markup.failed]; | 
| 
8b98db8fd183
clarified bracketing of messages: [forked [running finished] joined];
 wenzelm parents: 
68379diff
changeset | 176 | status task [Markup.finished]; | 
| 
8b98db8fd183
clarified bracketing of messages: [forked [running finished] joined];
 wenzelm parents: 
68379diff
changeset | 177 | Output.report [Markup.markup_only (Markup.bad ())]; | 
| 
8b98db8fd183
clarified bracketing of messages: [forked [running finished] joined];
 wenzelm parents: 
68379diff
changeset | 178 | if exec_id = 0 then () | 
| 
8b98db8fd183
clarified bracketing of messages: [forked [running finished] joined];
 wenzelm parents: 
68379diff
changeset | 179 | else List.app (Future.error_message pos) (Runtime.exn_messages exn)) | 
| 
8b98db8fd183
clarified bracketing of messages: [forked [running finished] joined];
 wenzelm parents: 
68379diff
changeset | 180 | | Exn.Res _ => | 
| 
8b98db8fd183
clarified bracketing of messages: [forked [running finished] joined];
 wenzelm parents: 
68379diff
changeset | 181 | status task [Markup.finished])) (); | 
| 54646 
2b38549374ba
more robust status reports: avoid loosing information about physical events (see also 28d207ba9340, 2bc5924b117f, 9edfd36a0355) -- NB: TTY and Proof General ignore Output.status and Output.report;
 wenzelm parents: 
53375diff
changeset | 182 | val _ = status task [Markup.joined]; | 
| 68880 
8b98db8fd183
clarified bracketing of messages: [forked [running finished] joined];
 wenzelm parents: 
68379diff
changeset | 183 | in Exn.release errors; Exn.release result end); | 
| 54646 
2b38549374ba
more robust status reports: avoid loosing information about physical events (see also 28d207ba9340, 2bc5924b117f, 9edfd36a0355) -- NB: TTY and Proof General ignore Output.status and Output.report;
 wenzelm parents: 
53375diff
changeset | 184 | |
| 53192 | 185 | val _ = status (Future.task_of future) [Markup.forked]; | 
| 186 | in future end)) (); | |
| 52604 | 187 | |
| 53192 | 188 | |
| 56291 
e79f76a48449
added Execution.print: accumulate print operations for some command execution, which are applied later and print time;
 wenzelm parents: 
54678diff
changeset | 189 | (* print *) | 
| 
e79f76a48449
added Execution.print: accumulate print operations for some command execution, which are applied later and print time;
 wenzelm parents: 
54678diff
changeset | 190 | |
| 56305 | 191 | fun print ({name, pos, pri}: params) e =
 | 
| 68354 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 192 | change_state (fn (execution_id, nodes, execs) => | 
| 56291 
e79f76a48449
added Execution.print: accumulate print operations for some command execution, which are applied later and print time;
 wenzelm parents: 
54678diff
changeset | 193 | let | 
| 56292 | 194 | val exec_id = the_default 0 (Position.parse_id pos); | 
| 56305 | 195 |       val print = {name = name, pri = pri, body = e};
 | 
| 56291 
e79f76a48449
added Execution.print: accumulate print operations for some command execution, which are applied later and print time;
 wenzelm parents: 
54678diff
changeset | 196 | in | 
| 
e79f76a48449
added Execution.print: accumulate print operations for some command execution, which are applied later and print time;
 wenzelm parents: 
54678diff
changeset | 197 | (case Inttab.lookup execs exec_id of | 
| 68354 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 198 | SOME (groups, prints) => | 
| 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 199 | let val execs' = Inttab.update (exec_id, (groups, print :: prints)) execs | 
| 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 200 | in (execution_id, nodes, execs') end | 
| 56291 
e79f76a48449
added Execution.print: accumulate print operations for some command execution, which are applied later and print time;
 wenzelm parents: 
54678diff
changeset | 201 | | NONE => raise Fail (unregistered exec_id)) | 
| 68354 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 202 | end); | 
| 56291 
e79f76a48449
added Execution.print: accumulate print operations for some command execution, which are applied later and print time;
 wenzelm parents: 
54678diff
changeset | 203 | |
| 56292 | 204 | fun fork_prints exec_id = | 
| 68353 | 205 | (case Inttab.lookup (#execs (get_state ())) exec_id of | 
| 56291 
e79f76a48449
added Execution.print: accumulate print operations for some command execution, which are applied later and print time;
 wenzelm parents: 
54678diff
changeset | 206 | SOME (_, prints) => | 
| 68130 
6fb85346cb79
clarified future scheduling parameters, with support for parallel_limit;
 wenzelm parents: 
67659diff
changeset | 207 | if Future.relevant prints then | 
| 56296 
5413b6379c0e
less markup by default -- this is stored persistently in Isabelle/Scala;
 wenzelm parents: 
56292diff
changeset | 208 | let val pos = Position.thread_data () in | 
| 56292 | 209 |           List.app (fn {name, pri, body} =>
 | 
| 210 |             ignore (fork {name = name, pos = pos, pri = pri} body)) (rev prints)
 | |
| 211 | end | |
| 67659 | 212 |       else List.app (fn {body, ...} => body ()) (rev prints)
 | 
| 56291 
e79f76a48449
added Execution.print: accumulate print operations for some command execution, which are applied later and print time;
 wenzelm parents: 
54678diff
changeset | 213 | | NONE => raise Fail (unregistered exec_id)); | 
| 
e79f76a48449
added Execution.print: accumulate print operations for some command execution, which are applied later and print time;
 wenzelm parents: 
54678diff
changeset | 214 | |
| 
e79f76a48449
added Execution.print: accumulate print operations for some command execution, which are applied later and print time;
 wenzelm parents: 
54678diff
changeset | 215 | |
| 53192 | 216 | (* cleanup *) | 
| 217 | ||
| 218 | fun purge exec_ids = | |
| 68354 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 219 | change_state (fn (execution_id, nodes, execs) => | 
| 53192 | 220 | let | 
| 221 | val execs' = fold Inttab.delete_safe exec_ids execs; | |
| 222 | val () = | |
| 56291 
e79f76a48449
added Execution.print: accumulate print operations for some command execution, which are applied later and print time;
 wenzelm parents: 
54678diff
changeset | 223 | (execs', ()) |-> Inttab.fold (fn (exec_id, (groups, _)) => fn () => | 
| 53192 | 224 | if Inttab.defined execs' exec_id then () | 
| 225 | else groups |> List.app (fn group => | |
| 226 | if Task_Queue.is_canceled group then () | |
| 53375 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53193diff
changeset | 227 |             else raise Fail ("Attempt to purge valid execution: " ^ Document_ID.print exec_id)));
 | 
| 68354 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 228 | in (execution_id, nodes, execs') end); | 
| 53192 | 229 | |
| 230 | fun reset () = | |
| 68354 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 231 | change_state_result (fn (_, _, execs) => | 
| 56291 
e79f76a48449
added Execution.print: accumulate print operations for some command execution, which are applied later and print time;
 wenzelm parents: 
54678diff
changeset | 232 | let val groups = Inttab.fold (append o #1 o #2) execs [] | 
| 68354 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 wenzelm parents: 
68353diff
changeset | 233 | in (groups, (Document_ID.none, Symtab.empty, init_execs)) end); | 
| 53192 | 234 | |
| 235 | fun shutdown () = | |
| 236 | (Future.shutdown (); | |
| 237 | (case maps Task_Queue.group_status (reset ()) of | |
| 238 | [] => () | |
| 239 | | exns => raise Par_Exn.make exns)); | |
| 240 | ||
| 241 | end; |