| author | haftmann | 
| Mon, 13 Sep 2010 16:43:23 +0200 | |
| changeset 39379 | ab1b070aa412 | 
| parent 39243 | 307e3d07d19f | 
| child 40450 | 8eab60e1baeb | 
| permissions | -rw-r--r-- | 
| 28165 | 1 | (* Title: Pure/Concurrent/task_queue.ML | 
| 2 | Author: Makarius | |
| 3 | ||
| 4 | Ordered queue of grouped tasks. | |
| 5 | *) | |
| 6 | ||
| 7 | signature TASK_QUEUE = | |
| 8 | sig | |
| 29340 | 9 | type task | 
| 34277 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 wenzelm parents: 
32814diff
changeset | 10 | val dummy_task: task | 
| 29121 | 11 | val pri_of_task: task -> int | 
| 28196 | 12 | val str_of_task: task -> string | 
| 29340 | 13 | type group | 
| 32221 | 14 | val new_group: group option -> group | 
| 32052 | 15 | val group_id: group -> int | 
| 29340 | 16 | val eq_group: group * group -> bool | 
| 32221 | 17 | val cancel_group: group -> exn -> unit | 
| 18 | val is_canceled: group -> bool | |
| 32101 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 19 | val group_status: group -> exn list | 
| 28179 | 20 | val str_of_group: group -> string | 
| 28165 | 21 | type queue | 
| 22 | val empty: queue | |
| 34277 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 wenzelm parents: 
32814diff
changeset | 23 | val all_passive: queue -> bool | 
| 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 wenzelm parents: 
32814diff
changeset | 24 |   val status: queue -> {ready: int, pending: int, running: int, passive: int}
 | 
| 34280 
16bf3e9786a3
eliminated cache, which complicates the code without making a real difference (NB: deque_towards often disrupts caching, and daisy-chaining of workers already reduces queue overhead);
 wenzelm parents: 
34279diff
changeset | 25 | val cancel: queue -> group -> bool | 
| 
16bf3e9786a3
eliminated cache, which complicates the code without making a real difference (NB: deque_towards often disrupts caching, and daisy-chaining of workers already reduces queue overhead);
 wenzelm parents: 
34279diff
changeset | 26 | val cancel_all: queue -> group list | 
| 37854 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 27 | val enqueue_passive: group -> (unit -> bool) -> queue -> task * queue | 
| 32218 
222f26693757
enqueue/finish: return minimal/maximal state of this task;
 wenzelm parents: 
32192diff
changeset | 28 | val enqueue: group -> task list -> int -> (bool -> bool) -> queue -> (task * bool) * queue | 
| 29365 | 29 | val extend: task -> (bool -> bool) -> queue -> queue option | 
| 39243 
307e3d07d19f
Future.promise: more robust treatment of concurrent abort vs. fulfill (amending 047c96f41455);
 wenzelm parents: 
39232diff
changeset | 30 | val dequeue_passive: Thread.thread -> task -> queue -> bool * queue | 
| 32249 | 31 | val dequeue: Thread.thread -> queue -> (task * group * (bool -> bool) list) option * queue | 
| 32814 
81897d30b97f
added Task_Queue.depend (again) -- light-weight version for transitive graph;
 wenzelm parents: 
32784diff
changeset | 32 | val depend: task -> task list -> queue -> queue | 
| 32249 | 33 | val dequeue_towards: Thread.thread -> task list -> queue -> | 
| 32224 | 34 | (((task * group * (bool -> bool) list) option * task list) * queue) | 
| 32221 | 35 | val finish: task -> queue -> bool * queue | 
| 28165 | 36 | end; | 
| 37 | ||
| 35242 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 wenzelm parents: 
35012diff
changeset | 38 | structure Task_Queue: TASK_QUEUE = | 
| 28165 | 39 | struct | 
| 40 | ||
| 29121 | 41 | (* tasks *) | 
| 42 | ||
| 35242 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 wenzelm parents: 
35012diff
changeset | 43 | abstype task = Task of int option * serial | 
| 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 wenzelm parents: 
35012diff
changeset | 44 | with | 
| 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 wenzelm parents: 
35012diff
changeset | 45 | |
| 34277 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 wenzelm parents: 
32814diff
changeset | 46 | val dummy_task = Task (NONE, ~1); | 
| 29121 | 47 | fun new_task pri = Task (pri, serial ()); | 
| 28165 | 48 | |
| 34277 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 wenzelm parents: 
32814diff
changeset | 49 | fun pri_of_task (Task (pri, _)) = the_default 0 pri; | 
| 29121 | 50 | fun str_of_task (Task (_, i)) = string_of_int i; | 
| 28998 | 51 | |
| 34277 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 wenzelm parents: 
32814diff
changeset | 52 | fun task_ord (Task t1, Task t2) = prod_ord (rev_order o option_ord int_ord) int_ord (t1, t2); | 
| 35242 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 wenzelm parents: 
35012diff
changeset | 53 | val eq_task = is_equal o task_ord; | 
| 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 wenzelm parents: 
35012diff
changeset | 54 | |
| 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 wenzelm parents: 
35012diff
changeset | 55 | end; | 
| 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 wenzelm parents: 
35012diff
changeset | 56 | |
| 31971 
8c1b845ed105
renamed functor TableFun to Table, and GraphFun to Graph;
 wenzelm parents: 
31632diff
changeset | 57 | structure Task_Graph = Graph(type key = task val ord = task_ord); | 
| 28165 | 58 | |
| 28998 | 59 | |
| 32101 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 60 | (* nested groups *) | 
| 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 61 | |
| 35242 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 wenzelm parents: 
35012diff
changeset | 62 | abstype group = Group of | 
| 32101 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 63 |  {parent: group option,
 | 
| 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 64 | id: serial, | 
| 35242 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 wenzelm parents: 
35012diff
changeset | 65 | status: exn list Synchronized.var} | 
| 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 wenzelm parents: 
35012diff
changeset | 66 | with | 
| 29121 | 67 | |
| 32101 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 68 | fun make_group (parent, id, status) = Group {parent = parent, id = id, status = status};
 | 
| 32052 | 69 | |
| 32251 | 70 | fun new_group parent = make_group (parent, serial (), Synchronized.var "group" []); | 
| 32101 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 71 | |
| 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 72 | fun group_id (Group {id, ...}) = id;
 | 
| 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 73 | fun eq_group (group1, group2) = group_id group1 = group_id group2; | 
| 28551 | 74 | |
| 32101 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 75 | fun group_ancestry (Group {parent, id, ...}) =
 | 
| 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 76 | id :: (case parent of NONE => [] | SOME group => group_ancestry group); | 
| 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 77 | |
| 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 78 | |
| 32221 | 79 | (* group status *) | 
| 80 | ||
| 32251 | 81 | fun cancel_group (Group {status, ...}) exn =
 | 
| 82 | Synchronized.change status | |
| 83 | (fn exns => | |
| 39232 
69c6d3e87660
more abstract treatment of interrupts in structure Exn -- hardly ever need to mention Interrupt literally;
 wenzelm parents: 
37854diff
changeset | 84 | if not (Exn.is_interrupt exn) orelse null exns then exn :: exns | 
| 
69c6d3e87660
more abstract treatment of interrupts in structure Exn -- hardly ever need to mention Interrupt literally;
 wenzelm parents: 
37854diff
changeset | 85 | else exns); | 
| 29121 | 86 | |
| 32251 | 87 | fun is_canceled (Group {parent, status, ...}) =
 | 
| 88 | not (null (Synchronized.value status)) orelse | |
| 89 | (case parent of NONE => false | SOME group => is_canceled group); | |
| 32221 | 90 | |
| 32251 | 91 | fun group_status (Group {parent, status, ...}) =
 | 
| 92 | Synchronized.value status @ | |
| 93 | (case parent of NONE => [] | SOME group => group_status group); | |
| 28165 | 94 | |
| 32101 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 95 | fun str_of_group group = | 
| 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 96 |   (is_canceled group ? enclose "(" ")") (string_of_int (group_id group));
 | 
| 28179 | 97 | |
| 35242 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 wenzelm parents: 
35012diff
changeset | 98 | end; | 
| 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 wenzelm parents: 
35012diff
changeset | 99 | |
| 28165 | 100 | |
| 28176 
01b21886e7f0
job: explicit 'ok' status -- false for canceled jobs;
 wenzelm parents: 
28171diff
changeset | 101 | (* jobs *) | 
| 28165 | 102 | |
| 103 | datatype job = | |
| 29365 | 104 | Job of (bool -> bool) list | | 
| 34277 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 wenzelm parents: 
32814diff
changeset | 105 | Running of Thread.thread | | 
| 37854 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 106 | Passive of unit -> bool; | 
| 28165 | 107 | |
| 29121 | 108 | type jobs = (group * job) Task_Graph.T; | 
| 28176 
01b21886e7f0
job: explicit 'ok' status -- false for canceled jobs;
 wenzelm parents: 
28171diff
changeset | 109 | |
| 29121 | 110 | fun get_group (jobs: jobs) task = #1 (Task_Graph.get_node jobs task); | 
| 111 | fun get_job (jobs: jobs) task = #2 (Task_Graph.get_node jobs task); | |
| 29365 | 112 | fun set_job task job (jobs: jobs) = Task_Graph.map_node task (fn (group, _) => (group, job)) jobs; | 
| 28202 
23cb9a974630
added focus, which indicates a particular collection of high-priority tasks;
 wenzelm parents: 
28196diff
changeset | 113 | |
| 32250 | 114 | fun add_job task dep (jobs: jobs) = | 
| 115 | Task_Graph.add_edge (dep, task) jobs handle Task_Graph.UNDEF _ => jobs; | |
| 116 | ||
| 32814 
81897d30b97f
added Task_Queue.depend (again) -- light-weight version for transitive graph;
 wenzelm parents: 
32784diff
changeset | 117 | fun add_dep task dep (jobs: jobs) = | 
| 
81897d30b97f
added Task_Queue.depend (again) -- light-weight version for transitive graph;
 wenzelm parents: 
32784diff
changeset | 118 | if Task_Graph.is_edge jobs (task, dep) then | 
| 
81897d30b97f
added Task_Queue.depend (again) -- light-weight version for transitive graph;
 wenzelm parents: 
32784diff
changeset | 119 | raise Fail "Cyclic dependency of future tasks" | 
| 
81897d30b97f
added Task_Queue.depend (again) -- light-weight version for transitive graph;
 wenzelm parents: 
32784diff
changeset | 120 | else add_job task dep jobs; | 
| 
81897d30b97f
added Task_Queue.depend (again) -- light-weight version for transitive graph;
 wenzelm parents: 
32784diff
changeset | 121 | |
| 32250 | 122 | fun get_deps (jobs: jobs) task = | 
| 123 | Task_Graph.imm_preds jobs task handle Task_Graph.UNDEF _ => []; | |
| 124 | ||
| 28176 
01b21886e7f0
job: explicit 'ok' status -- false for canceled jobs;
 wenzelm parents: 
28171diff
changeset | 125 | |
| 
01b21886e7f0
job: explicit 'ok' status -- false for canceled jobs;
 wenzelm parents: 
28171diff
changeset | 126 | (* queue of grouped jobs *) | 
| 
01b21886e7f0
job: explicit 'ok' status -- false for canceled jobs;
 wenzelm parents: 
28171diff
changeset | 127 | |
| 28165 | 128 | datatype queue = Queue of | 
| 28184 
5ed5cb73a2e9
eliminated cache, access queue efficiently via IntGraph.get_first;
 wenzelm parents: 
28179diff
changeset | 129 |  {groups: task list Inttab.table,   (*groups with presently active members*)
 | 
| 34280 
16bf3e9786a3
eliminated cache, which complicates the code without making a real difference (NB: deque_towards often disrupts caching, and daisy-chaining of workers already reduces queue overhead);
 wenzelm parents: 
34279diff
changeset | 130 | jobs: jobs}; (*job dependency graph*) | 
| 28165 | 131 | |
| 34280 
16bf3e9786a3
eliminated cache, which complicates the code without making a real difference (NB: deque_towards often disrupts caching, and daisy-chaining of workers already reduces queue overhead);
 wenzelm parents: 
34279diff
changeset | 132 | fun make_queue groups jobs = Queue {groups = groups, jobs = jobs};
 | 
| 28204 | 133 | |
| 34280 
16bf3e9786a3
eliminated cache, which complicates the code without making a real difference (NB: deque_towards often disrupts caching, and daisy-chaining of workers already reduces queue overhead);
 wenzelm parents: 
34279diff
changeset | 134 | val empty = make_queue Inttab.empty Task_Graph.empty; | 
| 34277 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 wenzelm parents: 
32814diff
changeset | 135 | |
| 37854 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 136 | |
| 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 137 | (* job status *) | 
| 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 138 | |
| 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 139 | fun ready_job task ((group, Job list), ([], _)) = SOME (task, group, rev list) | 
| 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 140 | | ready_job task ((group, Passive abort), ([], _)) = | 
| 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 141 | if is_canceled group then SOME (task, group, [fn _ => abort ()]) | 
| 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 142 | else NONE | 
| 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 143 | | ready_job _ _ = NONE; | 
| 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 144 | |
| 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 145 | fun active_job (_, Job _) = SOME () | 
| 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 146 | | active_job (_, Running _) = SOME () | 
| 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 147 | | active_job (group, Passive _) = if is_canceled group then SOME () else NONE; | 
| 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 148 | |
| 34277 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 wenzelm parents: 
32814diff
changeset | 149 | fun all_passive (Queue {jobs, ...}) =
 | 
| 37854 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 150 | is_none (Task_Graph.get_first (active_job o #1 o #2) jobs); | 
| 28165 | 151 | |
| 152 | ||
| 32101 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 153 | (* queue status *) | 
| 32052 | 154 | |
| 155 | fun status (Queue {jobs, ...}) =
 | |
| 156 | let | |
| 34277 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 wenzelm parents: 
32814diff
changeset | 157 | val (x, y, z, w) = | 
| 37854 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 158 | Task_Graph.fold (fn (_, ((_, job), (deps, _))) => fn (x, y, z, w) => | 
| 32052 | 159 | (case job of | 
| 37854 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 160 | Job _ => if null deps then (x + 1, y, z, w) else (x, y + 1, z, w) | 
| 34277 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 wenzelm parents: 
32814diff
changeset | 161 | | Running _ => (x, y, z + 1, w) | 
| 37854 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 162 | | Passive _ => (x, y, z, w + 1))) | 
| 34277 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 wenzelm parents: 
32814diff
changeset | 163 | jobs (0, 0, 0, 0); | 
| 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 wenzelm parents: 
32814diff
changeset | 164 |   in {ready = x, pending = y, running = z, passive = w} end;
 | 
| 32052 | 165 | |
| 166 | ||
| 32101 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 167 | (* cancel -- peers and sub-groups *) | 
| 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 168 | |
| 34280 
16bf3e9786a3
eliminated cache, which complicates the code without making a real difference (NB: deque_towards often disrupts caching, and daisy-chaining of workers already reduces queue overhead);
 wenzelm parents: 
34279diff
changeset | 169 | fun cancel (Queue {groups, jobs}) group =
 | 
| 32101 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 170 | let | 
| 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 171 | val _ = cancel_group group Exn.Interrupt; | 
| 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 172 | val tasks = Inttab.lookup_list groups (group_id group); | 
| 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 173 | val running = | 
| 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 174 | fold (get_job jobs #> (fn Running t => insert Thread.equal t | _ => I)) tasks []; | 
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
35242diff
changeset | 175 | val _ = List.app Simple_Thread.interrupt running; | 
| 34280 
16bf3e9786a3
eliminated cache, which complicates the code without making a real difference (NB: deque_towards often disrupts caching, and daisy-chaining of workers already reduces queue overhead);
 wenzelm parents: 
34279diff
changeset | 176 | in null running end; | 
| 32101 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 177 | |
| 37854 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 178 | fun cancel_all (Queue {jobs, ...}) =
 | 
| 32101 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 179 | let | 
| 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 180 | fun cancel_job (group, job) (groups, running) = | 
| 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 181 | (cancel_group group Exn.Interrupt; | 
| 34277 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 wenzelm parents: 
32814diff
changeset | 182 | (case job of | 
| 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 wenzelm parents: 
32814diff
changeset | 183 | Running t => (insert eq_group group groups, insert Thread.equal t running) | 
| 32101 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 184 | | _ => (groups, running))); | 
| 34277 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 wenzelm parents: 
32814diff
changeset | 185 | val (running_groups, running) = Task_Graph.fold (cancel_job o #1 o #2) jobs ([], []); | 
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
35242diff
changeset | 186 | val _ = List.app Simple_Thread.interrupt running; | 
| 34280 
16bf3e9786a3
eliminated cache, which complicates the code without making a real difference (NB: deque_towards often disrupts caching, and daisy-chaining of workers already reduces queue overhead);
 wenzelm parents: 
34279diff
changeset | 187 | in running_groups end; | 
| 32101 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 188 | |
| 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 189 | |
| 28185 
0f20cbce4935
simplified dequeue: provide Thread.self internally;
 wenzelm parents: 
28184diff
changeset | 190 | (* enqueue *) | 
| 28165 | 191 | |
| 37854 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 192 | fun enqueue_passive group abort (Queue {groups, jobs}) =
 | 
| 34277 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 wenzelm parents: 
32814diff
changeset | 193 | let | 
| 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 wenzelm parents: 
32814diff
changeset | 194 | val task = new_task NONE; | 
| 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 wenzelm parents: 
32814diff
changeset | 195 | val groups' = groups | 
| 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 wenzelm parents: 
32814diff
changeset | 196 | |> fold (fn gid => Inttab.cons_list (gid, task)) (group_ancestry group); | 
| 37854 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 197 | val jobs' = jobs |> Task_Graph.new_node (task, (group, Passive abort)); | 
| 34280 
16bf3e9786a3
eliminated cache, which complicates the code without making a real difference (NB: deque_towards often disrupts caching, and daisy-chaining of workers already reduces queue overhead);
 wenzelm parents: 
34279diff
changeset | 198 | in (task, make_queue groups' jobs') end; | 
| 34277 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 wenzelm parents: 
32814diff
changeset | 199 | |
| 34280 
16bf3e9786a3
eliminated cache, which complicates the code without making a real difference (NB: deque_towards often disrupts caching, and daisy-chaining of workers already reduces queue overhead);
 wenzelm parents: 
34279diff
changeset | 200 | fun enqueue group deps pri job (Queue {groups, jobs}) =
 | 
| 28165 | 201 | let | 
| 34277 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 wenzelm parents: 
32814diff
changeset | 202 | val task = new_task (SOME pri); | 
| 32101 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 203 | val groups' = groups | 
| 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 204 | |> fold (fn gid => Inttab.cons_list (gid, task)) (group_ancestry group); | 
| 28185 
0f20cbce4935
simplified dequeue: provide Thread.self internally;
 wenzelm parents: 
28184diff
changeset | 205 | val jobs' = jobs | 
| 32101 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 wenzelm parents: 
32099diff
changeset | 206 | |> Task_Graph.new_node (task, (group, Job [job])) | 
| 32190 
4fc7a882b41e
enqueue: maintain transitive closure, which simplifies dequeue_towards;
 wenzelm parents: 
32101diff
changeset | 207 | |> fold (add_job task) deps | 
| 
4fc7a882b41e
enqueue: maintain transitive closure, which simplifies dequeue_towards;
 wenzelm parents: 
32101diff
changeset | 208 | |> fold (fold (add_job task) o get_deps jobs) deps; | 
| 32218 
222f26693757
enqueue/finish: return minimal/maximal state of this task;
 wenzelm parents: 
32192diff
changeset | 209 | val minimal = null (get_deps jobs' task); | 
| 34280 
16bf3e9786a3
eliminated cache, which complicates the code without making a real difference (NB: deque_towards often disrupts caching, and daisy-chaining of workers already reduces queue overhead);
 wenzelm parents: 
34279diff
changeset | 210 | in ((task, minimal), make_queue groups' jobs') end; | 
| 28165 | 211 | |
| 34280 
16bf3e9786a3
eliminated cache, which complicates the code without making a real difference (NB: deque_towards often disrupts caching, and daisy-chaining of workers already reduces queue overhead);
 wenzelm parents: 
34279diff
changeset | 212 | fun extend task job (Queue {groups, jobs}) =
 | 
| 29365 | 213 | (case try (get_job jobs) task of | 
| 34280 
16bf3e9786a3
eliminated cache, which complicates the code without making a real difference (NB: deque_towards often disrupts caching, and daisy-chaining of workers already reduces queue overhead);
 wenzelm parents: 
34279diff
changeset | 214 | SOME (Job list) => SOME (make_queue groups (set_job task (Job (job :: list)) jobs)) | 
| 29365 | 215 | | _ => NONE); | 
| 216 | ||
| 28185 
0f20cbce4935
simplified dequeue: provide Thread.self internally;
 wenzelm parents: 
28184diff
changeset | 217 | |
| 
0f20cbce4935
simplified dequeue: provide Thread.self internally;
 wenzelm parents: 
28184diff
changeset | 218 | (* dequeue *) | 
| 
0f20cbce4935
simplified dequeue: provide Thread.self internally;
 wenzelm parents: 
28184diff
changeset | 219 | |
| 39243 
307e3d07d19f
Future.promise: more robust treatment of concurrent abort vs. fulfill (amending 047c96f41455);
 wenzelm parents: 
39232diff
changeset | 220 | fun dequeue_passive thread task (queue as Queue {groups, jobs}) =
 | 
| 
307e3d07d19f
Future.promise: more robust treatment of concurrent abort vs. fulfill (amending 047c96f41455);
 wenzelm parents: 
39232diff
changeset | 221 | (case try (get_job jobs) task of | 
| 
307e3d07d19f
Future.promise: more robust treatment of concurrent abort vs. fulfill (amending 047c96f41455);
 wenzelm parents: 
39232diff
changeset | 222 | SOME (Passive _) => | 
| 
307e3d07d19f
Future.promise: more robust treatment of concurrent abort vs. fulfill (amending 047c96f41455);
 wenzelm parents: 
39232diff
changeset | 223 | let val jobs' = set_job task (Running thread) jobs | 
| 
307e3d07d19f
Future.promise: more robust treatment of concurrent abort vs. fulfill (amending 047c96f41455);
 wenzelm parents: 
39232diff
changeset | 224 | in (true, make_queue groups jobs') end | 
| 
307e3d07d19f
Future.promise: more robust treatment of concurrent abort vs. fulfill (amending 047c96f41455);
 wenzelm parents: 
39232diff
changeset | 225 | | _ => (false, queue)); | 
| 
307e3d07d19f
Future.promise: more robust treatment of concurrent abort vs. fulfill (amending 047c96f41455);
 wenzelm parents: 
39232diff
changeset | 226 | |
| 34280 
16bf3e9786a3
eliminated cache, which complicates the code without making a real difference (NB: deque_towards often disrupts caching, and daisy-chaining of workers already reduces queue overhead);
 wenzelm parents: 
34279diff
changeset | 227 | fun dequeue thread (queue as Queue {groups, jobs}) =
 | 
| 37854 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 228 | (case Task_Graph.get_first (uncurry ready_job) jobs of | 
| 39243 
307e3d07d19f
Future.promise: more robust treatment of concurrent abort vs. fulfill (amending 047c96f41455);
 wenzelm parents: 
39232diff
changeset | 229 | SOME (result as (task, _, _)) => | 
| 37854 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 230 | let val jobs' = set_job task (Running thread) jobs | 
| 39243 
307e3d07d19f
Future.promise: more robust treatment of concurrent abort vs. fulfill (amending 047c96f41455);
 wenzelm parents: 
39232diff
changeset | 231 | in (SOME result, make_queue groups jobs') end | 
| 
307e3d07d19f
Future.promise: more robust treatment of concurrent abort vs. fulfill (amending 047c96f41455);
 wenzelm parents: 
39232diff
changeset | 232 | | NONE => (NONE, queue)); | 
| 28202 
23cb9a974630
added focus, which indicates a particular collection of high-priority tasks;
 wenzelm parents: 
28196diff
changeset | 233 | |
| 28165 | 234 | |
| 32055 
6a46898aa805
recovered a version of dequeue_towards (cf. bb7b5a5942c7);
 wenzelm parents: 
32052diff
changeset | 235 | (* dequeue_towards -- adhoc dependencies *) | 
| 
6a46898aa805
recovered a version of dequeue_towards (cf. bb7b5a5942c7);
 wenzelm parents: 
32052diff
changeset | 236 | |
| 34280 
16bf3e9786a3
eliminated cache, which complicates the code without making a real difference (NB: deque_towards often disrupts caching, and daisy-chaining of workers already reduces queue overhead);
 wenzelm parents: 
34279diff
changeset | 237 | fun depend task deps (Queue {groups, jobs}) =
 | 
| 
16bf3e9786a3
eliminated cache, which complicates the code without making a real difference (NB: deque_towards often disrupts caching, and daisy-chaining of workers already reduces queue overhead);
 wenzelm parents: 
34279diff
changeset | 238 | make_queue groups (fold (add_dep task) deps jobs); | 
| 32814 
81897d30b97f
added Task_Queue.depend (again) -- light-weight version for transitive graph;
 wenzelm parents: 
32784diff
changeset | 239 | |
| 34280 
16bf3e9786a3
eliminated cache, which complicates the code without making a real difference (NB: deque_towards often disrupts caching, and daisy-chaining of workers already reduces queue overhead);
 wenzelm parents: 
34279diff
changeset | 240 | fun dequeue_towards thread deps (queue as Queue {groups, jobs}) =
 | 
| 32055 
6a46898aa805
recovered a version of dequeue_towards (cf. bb7b5a5942c7);
 wenzelm parents: 
32052diff
changeset | 241 | let | 
| 37854 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 wenzelm parents: 
37216diff
changeset | 242 | fun ready task = ready_job task (Task_Graph.get_entry jobs task); | 
| 32055 
6a46898aa805
recovered a version of dequeue_towards (cf. bb7b5a5942c7);
 wenzelm parents: 
32052diff
changeset | 243 | val tasks = filter (can (Task_Graph.get_node jobs)) deps; | 
| 32192 | 244 | fun result (res as (task, _, _)) = | 
| 34280 
16bf3e9786a3
eliminated cache, which complicates the code without making a real difference (NB: deque_towards often disrupts caching, and daisy-chaining of workers already reduces queue overhead);
 wenzelm parents: 
34279diff
changeset | 245 | let val jobs' = set_job task (Running thread) jobs | 
| 
16bf3e9786a3
eliminated cache, which complicates the code without making a real difference (NB: deque_towards often disrupts caching, and daisy-chaining of workers already reduces queue overhead);
 wenzelm parents: 
34279diff
changeset | 246 | in ((SOME res, tasks), make_queue groups jobs') end; | 
| 32055 
6a46898aa805
recovered a version of dequeue_towards (cf. bb7b5a5942c7);
 wenzelm parents: 
32052diff
changeset | 247 | in | 
| 32093 
30996b775a7f
tuned dequeu_towards: try immediate tasks before expensive all_preds;
 wenzelm parents: 
32055diff
changeset | 248 | (case get_first ready tasks of | 
| 32192 | 249 | SOME res => result res | 
| 250 | | NONE => | |
| 32250 | 251 | (case get_first (get_first ready o get_deps jobs) tasks of | 
| 32192 | 252 | SOME res => result res | 
| 32224 | 253 | | NONE => ((NONE, tasks), queue))) | 
| 32055 
6a46898aa805
recovered a version of dequeue_towards (cf. bb7b5a5942c7);
 wenzelm parents: 
32052diff
changeset | 254 | end; | 
| 
6a46898aa805
recovered a version of dequeue_towards (cf. bb7b5a5942c7);
 wenzelm parents: 
32052diff
changeset | 255 | |
| 
6a46898aa805
recovered a version of dequeue_towards (cf. bb7b5a5942c7);
 wenzelm parents: 
32052diff
changeset | 256 | |
| 32221 | 257 | (* finish *) | 
| 258 | ||
| 34280 
16bf3e9786a3
eliminated cache, which complicates the code without making a real difference (NB: deque_towards often disrupts caching, and daisy-chaining of workers already reduces queue overhead);
 wenzelm parents: 
34279diff
changeset | 259 | fun finish task (Queue {groups, jobs}) =
 | 
| 32221 | 260 | let | 
| 261 | val group = get_group jobs task; | |
| 262 | val groups' = groups | |
| 35242 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 wenzelm parents: 
35012diff
changeset | 263 | |> fold (fn gid => Inttab.remove_list eq_task (gid, task)) (group_ancestry group); | 
| 32221 | 264 | val jobs' = Task_Graph.del_node task jobs; | 
| 265 | val maximal = null (Task_Graph.imm_succs jobs task); | |
| 34280 
16bf3e9786a3
eliminated cache, which complicates the code without making a real difference (NB: deque_towards often disrupts caching, and daisy-chaining of workers already reduces queue overhead);
 wenzelm parents: 
34279diff
changeset | 266 | in (maximal, make_queue groups' jobs') end; | 
| 32221 | 267 | |
| 28165 | 268 | end; |