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