| author | wenzelm | 
| Thu, 24 Apr 2014 15:19:11 +0200 | |
| changeset 56701 | ac5b66fa2a56 | 
| parent 54649 | 99b9249b3e05 | 
| child 59332 | 56fec2b5b826 | 
| 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 group  | 
| 32221 | 10  | 
val new_group: group option -> group  | 
| 32052 | 11  | 
val group_id: group -> int  | 
| 29340 | 12  | 
val eq_group: group * group -> bool  | 
| 32221 | 13  | 
val cancel_group: group -> exn -> unit  | 
14  | 
val is_canceled: group -> bool  | 
|
| 
53190
 
5d92649a310e
simplified Goal.forked_proofs: status is determined via group instead of dummy future (see also Pure/PIDE/execution.ML);
 
wenzelm 
parents: 
52558 
diff
changeset
 | 
15  | 
val group_status: group -> exn list  | 
| 28179 | 16  | 
val str_of_group: group -> string  | 
| 43951 | 17  | 
val str_of_groups: group -> string  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
18  | 
type task  | 
| 
45136
 
2afb928c71ca
static dummy_task (again) to avoid a few extra allocations;
 
wenzelm 
parents: 
44341 
diff
changeset
 | 
19  | 
val dummy_task: task  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
20  | 
val group_of_task: task -> group  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
21  | 
val name_of_task: task -> string  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
22  | 
val pri_of_task: task -> int  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
23  | 
val str_of_task: task -> string  | 
| 43951 | 24  | 
val str_of_task_groups: task -> string  | 
| 50975 | 25  | 
val task_statistics: task -> Properties.T  | 
| 
41695
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
26  | 
val running: task -> (unit -> 'a) -> 'a  | 
| 
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
27  | 
val joining: task -> (unit -> 'a) -> 'a  | 
| 
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
28  | 
val waiting: task -> task list -> (unit -> 'a) -> 'a  | 
| 28165 | 29  | 
type queue  | 
30  | 
val empty: queue  | 
|
| 
47404
 
e6e5750f1311
simplified Future.cancel/cancel_group (again) -- running threads only;
 
wenzelm 
parents: 
45642 
diff
changeset
 | 
31  | 
val group_tasks: queue -> group -> task list  | 
| 41708 | 32  | 
val known_task: queue -> task -> bool  | 
| 
34277
 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 
wenzelm 
parents: 
32814 
diff
changeset
 | 
33  | 
val all_passive: queue -> bool  | 
| 
 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 
wenzelm 
parents: 
32814 
diff
changeset
 | 
34  | 
  val status: queue -> {ready: int, pending: int, running: int, passive: int}
 | 
| 
47404
 
e6e5750f1311
simplified Future.cancel/cancel_group (again) -- running threads only;
 
wenzelm 
parents: 
45642 
diff
changeset
 | 
35  | 
val cancel: queue -> group -> Thread.thread list  | 
| 
44341
 
a93d25fb14fc
purely functional task_queue.ML -- moved actual interrupt_unsynchronized to future.ML;
 
wenzelm 
parents: 
44340 
diff
changeset
 | 
36  | 
val cancel_all: queue -> group list * Thread.thread list  | 
| 
41681
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
37  | 
val finish: task -> queue -> bool * queue  | 
| 
54649
 
99b9249b3e05
more official task context via Task_Queue.enroll, which is required to participate in group cancellation (e.g. to terminate command exec);
 
wenzelm 
parents: 
53190 
diff
changeset
 | 
38  | 
val enroll: Thread.thread -> string -> group -> queue -> task * queue  | 
| 
37854
 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
39  | 
val enqueue_passive: group -> (unit -> bool) -> queue -> task * queue  | 
| 41708 | 40  | 
val enqueue: string -> group -> task list -> int -> (bool -> bool) -> queue -> task * queue  | 
| 29365 | 41  | 
val extend: task -> (bool -> bool) -> queue -> queue option  | 
| 
47423
 
8a179a0493e3
more robust Future.fulfill wrt. duplicate assignment and interrupt;
 
wenzelm 
parents: 
47404 
diff
changeset
 | 
42  | 
val dequeue_passive: Thread.thread -> task -> queue -> bool option * queue  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
43  | 
val dequeue: Thread.thread -> queue -> (task * (bool -> bool) list) option * queue  | 
| 
41695
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
44  | 
val dequeue_deps: Thread.thread -> task list -> queue ->  | 
| 
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
45  | 
(((task * (bool -> bool) list) option * task list) * queue)  | 
| 28165 | 46  | 
end;  | 
47  | 
||
| 
35242
 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 
wenzelm 
parents: 
35012 
diff
changeset
 | 
48  | 
structure Task_Queue: TASK_QUEUE =  | 
| 28165 | 49  | 
struct  | 
50  | 
||
| 52537 | 51  | 
val new_id = Counter.make ();  | 
| 
40450
 
8eab60e1baeb
private counter, to keep externalized ids a bit smaller;
 
wenzelm 
parents: 
39243 
diff
changeset
 | 
52  | 
|
| 
 
8eab60e1baeb
private counter, to keep externalized ids a bit smaller;
 
wenzelm 
parents: 
39243 
diff
changeset
 | 
53  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
54  | 
(** nested groups of tasks **)  | 
| 41670 | 55  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
56  | 
(* groups *)  | 
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
57  | 
|
| 
35242
 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 
wenzelm 
parents: 
35012 
diff
changeset
 | 
58  | 
abstype group = Group of  | 
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
59  | 
 {parent: group option,
 | 
| 
40450
 
8eab60e1baeb
private counter, to keep externalized ids a bit smaller;
 
wenzelm 
parents: 
39243 
diff
changeset
 | 
60  | 
id: int,  | 
| 44247 | 61  | 
status: exn option Synchronized.var}  | 
| 
35242
 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 
wenzelm 
parents: 
35012 
diff
changeset
 | 
62  | 
with  | 
| 29121 | 63  | 
|
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
64  | 
fun make_group (parent, id, status) = Group {parent = parent, id = id, status = status};
 | 
| 32052 | 65  | 
|
| 44247 | 66  | 
fun new_group parent = make_group (parent, new_id (), Synchronized.var "group_status" NONE);  | 
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
67  | 
|
| 
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
68  | 
fun group_id (Group {id, ...}) = id;
 | 
| 
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
69  | 
fun eq_group (group1, group2) = group_id group1 = group_id group2;  | 
| 28551 | 70  | 
|
| 43951 | 71  | 
fun fold_groups f (g as Group {parent = NONE, ...}) a = f g a
 | 
72  | 
  | fold_groups f (g as Group {parent = SOME group, ...}) a = fold_groups f group (f g a);
 | 
|
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
73  | 
|
| 
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
74  | 
|
| 32221 | 75  | 
(* group status *)  | 
76  | 
||
| 32251 | 77  | 
fun cancel_group (Group {status, ...}) exn =
 | 
78  | 
Synchronized.change status  | 
|
| 44247 | 79  | 
(fn exns => SOME (Par_Exn.make (exn :: the_list exns)));  | 
| 29121 | 80  | 
|
| 32251 | 81  | 
fun is_canceled (Group {parent, status, ...}) =
 | 
| 44247 | 82  | 
is_some (Synchronized.value status) orelse  | 
| 32251 | 83  | 
(case parent of NONE => false | SOME group => is_canceled group);  | 
| 32221 | 84  | 
|
| 
53190
 
5d92649a310e
simplified Goal.forked_proofs: status is determined via group instead of dummy future (see also Pure/PIDE/execution.ML);
 
wenzelm 
parents: 
52558 
diff
changeset
 | 
85  | 
fun group_status (Group {parent, status, ...}) =
 | 
| 44247 | 86  | 
the_list (Synchronized.value status) @  | 
| 
53190
 
5d92649a310e
simplified Goal.forked_proofs: status is determined via group instead of dummy future (see also Pure/PIDE/execution.ML);
 
wenzelm 
parents: 
52558 
diff
changeset
 | 
87  | 
(case parent of NONE => [] | SOME group => group_status group);  | 
| 28165 | 88  | 
|
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
89  | 
fun str_of_group group =  | 
| 
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
90  | 
  (is_canceled group ? enclose "(" ")") (string_of_int (group_id group));
 | 
| 28179 | 91  | 
|
| 43951 | 92  | 
fun str_of_groups group =  | 
93  | 
space_implode "/" (map str_of_group (rev (fold_groups cons group [])));  | 
|
94  | 
||
| 
35242
 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 
wenzelm 
parents: 
35012 
diff
changeset
 | 
95  | 
end;  | 
| 
 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 
wenzelm 
parents: 
35012 
diff
changeset
 | 
96  | 
|
| 28165 | 97  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
98  | 
(* tasks *)  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
99  | 
|
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
100  | 
type timing = Time.time * Time.time * string list; (*run, wait, wait dependencies*)  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
101  | 
|
| 
45136
 
2afb928c71ca
static dummy_task (again) to avoid a few extra allocations;
 
wenzelm 
parents: 
44341 
diff
changeset
 | 
102  | 
val timing_start = (Time.zeroTime, Time.zeroTime, []): timing;  | 
| 
 
2afb928c71ca
static dummy_task (again) to avoid a few extra allocations;
 
wenzelm 
parents: 
44341 
diff
changeset
 | 
103  | 
|
| 
45354
 
a2157057024c
optional timing, to avoid redundant allocation of mutable cells;
 
wenzelm 
parents: 
45136 
diff
changeset
 | 
104  | 
fun new_timing () =  | 
| 
 
a2157057024c
optional timing, to avoid redundant allocation of mutable cells;
 
wenzelm 
parents: 
45136 
diff
changeset
 | 
105  | 
if ! Multithreading.trace < 2 then NONE  | 
| 
 
a2157057024c
optional timing, to avoid redundant allocation of mutable cells;
 
wenzelm 
parents: 
45136 
diff
changeset
 | 
106  | 
else SOME (Synchronized.var "timing" timing_start);  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
107  | 
|
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
108  | 
abstype task = Task of  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
109  | 
 {group: group,
 | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
110  | 
name: string,  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
111  | 
id: int,  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
112  | 
pri: int option,  | 
| 50975 | 113  | 
timing: timing Synchronized.var option,  | 
114  | 
pos: Position.T}  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
115  | 
with  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
116  | 
|
| 
45136
 
2afb928c71ca
static dummy_task (again) to avoid a few extra allocations;
 
wenzelm 
parents: 
44341 
diff
changeset
 | 
117  | 
val dummy_task =  | 
| 50975 | 118  | 
  Task {group = new_group NONE, name = "", id = 0, pri = NONE, timing = NONE,
 | 
119  | 
pos = Position.none};  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
120  | 
|
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
121  | 
fun new_task group name pri =  | 
| 50975 | 122  | 
  Task {group = group, name = name, id = new_id (), pri = pri, timing = new_timing (),
 | 
123  | 
pos = Position.thread_data ()};  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
124  | 
|
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
125  | 
fun group_of_task (Task {group, ...}) = group;
 | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
126  | 
fun name_of_task (Task {name, ...}) = name;
 | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
127  | 
fun pri_of_task (Task {pri, ...}) = the_default 0 pri;
 | 
| 43951 | 128  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
129  | 
fun str_of_task (Task {name, id, ...}) =
 | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
130  | 
  if name = "" then string_of_int id else string_of_int id ^ " (" ^ name ^ ")";
 | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
131  | 
|
| 43951 | 132  | 
fun str_of_task_groups task = str_of_task task ^ " in " ^ str_of_groups (group_of_task task);  | 
133  | 
||
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
134  | 
fun update_timing update (Task {timing, ...}) e =
 | 
| 
41702
 
dca4c58c5d57
Task_Queue.update_timing: more precise treatment of interruptibility;
 
wenzelm 
parents: 
41695 
diff
changeset
 | 
135  | 
uninterruptible (fn restore_attributes => fn () =>  | 
| 
 
dca4c58c5d57
Task_Queue.update_timing: more precise treatment of interruptibility;
 
wenzelm 
parents: 
41695 
diff
changeset
 | 
136  | 
let  | 
| 
 
dca4c58c5d57
Task_Queue.update_timing: more precise treatment of interruptibility;
 
wenzelm 
parents: 
41695 
diff
changeset
 | 
137  | 
val start = Time.now ();  | 
| 
 
dca4c58c5d57
Task_Queue.update_timing: more precise treatment of interruptibility;
 
wenzelm 
parents: 
41695 
diff
changeset
 | 
138  | 
val result = Exn.capture (restore_attributes e) ();  | 
| 
 
dca4c58c5d57
Task_Queue.update_timing: more precise treatment of interruptibility;
 
wenzelm 
parents: 
41695 
diff
changeset
 | 
139  | 
val t = Time.- (Time.now (), start);  | 
| 
45136
 
2afb928c71ca
static dummy_task (again) to avoid a few extra allocations;
 
wenzelm 
parents: 
44341 
diff
changeset
 | 
140  | 
val _ = (case timing of NONE => () | SOME var => Synchronized.change var (update t));  | 
| 
41702
 
dca4c58c5d57
Task_Queue.update_timing: more precise treatment of interruptibility;
 
wenzelm 
parents: 
41695 
diff
changeset
 | 
141  | 
in Exn.release result end) ();  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
142  | 
|
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
143  | 
fun task_ord (Task {id = id1, pri = pri1, ...}, Task {id = id2, pri = pri2, ...}) =
 | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
144  | 
prod_ord (rev_order o option_ord int_ord) int_ord ((pri1, id1), (pri2, id2));  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
145  | 
|
| 50975 | 146  | 
fun task_statistics (Task {name, id, timing, pos, ...}) =
 | 
147  | 
let  | 
|
148  | 
val (run, wait, wait_deps) =  | 
|
149  | 
(case timing of NONE => timing_start | SOME var => Synchronized.value var);  | 
|
150  | 
fun micros time = string_of_int (Time.toNanoseconds time div 1000);  | 
|
151  | 
in  | 
|
| 51990 | 152  | 
    [("now", Markup.print_real (Time.toReal (Time.now ()))),
 | 
| 50975 | 153  | 
     ("task_name", name), ("task_id", Markup.print_int id),
 | 
154  | 
     ("run", micros run), ("wait", micros wait), ("wait_deps", commas wait_deps)] @
 | 
|
155  | 
Position.properties_of pos  | 
|
156  | 
end;  | 
|
157  | 
||
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
158  | 
end;  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
159  | 
|
| 
41709
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
160  | 
structure Tasks = Table(type key = task val ord = task_ord);  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
161  | 
structure Task_Graph = Graph(type key = task val ord = task_ord);  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
162  | 
|
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
163  | 
|
| 
41695
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
164  | 
(* timing *)  | 
| 
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
165  | 
|
| 
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
166  | 
fun running task =  | 
| 
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
167  | 
update_timing (fn t => fn (a, b, ds) => (Time.+ (a, t), b, ds)) task;  | 
| 
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
168  | 
|
| 
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
169  | 
fun joining task =  | 
| 
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
170  | 
update_timing (fn t => fn (a, b, ds) => (Time.- (a, t), b, ds)) task;  | 
| 
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
171  | 
|
| 
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
172  | 
fun waiting task deps =  | 
| 
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
173  | 
update_timing (fn t => fn (a, b, ds) =>  | 
| 
41702
 
dca4c58c5d57
Task_Queue.update_timing: more precise treatment of interruptibility;
 
wenzelm 
parents: 
41695 
diff
changeset
 | 
174  | 
(Time.- (a, t), Time.+ (b, t),  | 
| 
 
dca4c58c5d57
Task_Queue.update_timing: more precise treatment of interruptibility;
 
wenzelm 
parents: 
41695 
diff
changeset
 | 
175  | 
if ! Multithreading.trace > 0  | 
| 
 
dca4c58c5d57
Task_Queue.update_timing: more precise treatment of interruptibility;
 
wenzelm 
parents: 
41695 
diff
changeset
 | 
176  | 
then fold (insert (op =) o name_of_task) deps ds else ds)) task;  | 
| 
41695
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
177  | 
|
| 
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
178  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
179  | 
|
| 
41681
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
180  | 
(** queue of jobs and groups **)  | 
| 
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
181  | 
|
| 
41709
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
182  | 
(* known group members *)  | 
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
183  | 
|
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
184  | 
type groups = unit Tasks.table Inttab.table;  | 
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
185  | 
|
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
186  | 
fun get_tasks (groups: groups) gid =  | 
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
187  | 
the_default Tasks.empty (Inttab.lookup groups gid);  | 
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
188  | 
|
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
189  | 
fun add_task (gid, task) groups =  | 
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
190  | 
Inttab.update (gid, Tasks.update (task, ()) (get_tasks groups gid)) groups;  | 
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
191  | 
|
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
192  | 
fun del_task (gid, task) groups =  | 
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
193  | 
let val tasks = Tasks.delete_safe task (get_tasks groups gid) in  | 
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
194  | 
if Tasks.is_empty tasks then Inttab.delete_safe gid groups  | 
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
195  | 
else Inttab.update (gid, tasks) groups  | 
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
196  | 
end;  | 
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
197  | 
|
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
198  | 
|
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
199  | 
(* job dependency graph *)  | 
| 28165 | 200  | 
|
201  | 
datatype job =  | 
|
| 29365 | 202  | 
Job of (bool -> bool) list |  | 
| 
34277
 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 
wenzelm 
parents: 
32814 
diff
changeset
 | 
203  | 
Running of Thread.thread |  | 
| 
37854
 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
204  | 
Passive of unit -> bool;  | 
| 28165 | 205  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
206  | 
type jobs = job Task_Graph.T;  | 
| 
28176
 
01b21886e7f0
job: explicit 'ok' status -- false for canceled jobs;
 
wenzelm 
parents: 
28171 
diff
changeset
 | 
207  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
208  | 
fun get_job (jobs: jobs) task = Task_Graph.get_node jobs task;  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
209  | 
fun set_job task job (jobs: jobs) = Task_Graph.map_node task (K job) jobs;  | 
| 
28202
 
23cb9a974630
added focus, which indicates a particular collection of high-priority tasks;
 
wenzelm 
parents: 
28196 
diff
changeset
 | 
210  | 
|
| 32250 | 211  | 
fun add_job task dep (jobs: jobs) =  | 
212  | 
Task_Graph.add_edge (dep, task) jobs handle Task_Graph.UNDEF _ => jobs;  | 
|
213  | 
||
| 
28176
 
01b21886e7f0
job: explicit 'ok' status -- false for canceled jobs;
 
wenzelm 
parents: 
28171 
diff
changeset
 | 
214  | 
|
| 
41681
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
215  | 
(* queue *)  | 
| 
28176
 
01b21886e7f0
job: explicit 'ok' status -- false for canceled jobs;
 
wenzelm 
parents: 
28171 
diff
changeset
 | 
216  | 
|
| 
41709
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
217  | 
datatype queue = Queue of {groups: groups, jobs: jobs};
 | 
| 28165 | 218  | 
|
| 
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
 | 
219  | 
fun make_queue groups jobs = Queue {groups = groups, jobs = 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
 | 
220  | 
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
 | 
221  | 
|
| 
47404
 
e6e5750f1311
simplified Future.cancel/cancel_group (again) -- running threads only;
 
wenzelm 
parents: 
45642 
diff
changeset
 | 
222  | 
fun group_tasks (Queue {groups, ...}) group = Tasks.keys (get_tasks groups (group_id group));
 | 
| 41708 | 223  | 
fun known_task (Queue {jobs, ...}) task = can (Task_Graph.get_entry jobs) task;
 | 
224  | 
||
| 
37854
 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
225  | 
|
| 
 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
226  | 
(* job status *)  | 
| 
 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
227  | 
|
| 
44338
 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 
wenzelm 
parents: 
44299 
diff
changeset
 | 
228  | 
fun ready_job task (Job list, (deps, _)) =  | 
| 
 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 
wenzelm 
parents: 
44299 
diff
changeset
 | 
229  | 
if Task_Graph.Keys.is_empty deps then SOME (task, rev list) else NONE  | 
| 
 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 
wenzelm 
parents: 
44299 
diff
changeset
 | 
230  | 
| ready_job task (Passive abort, (deps, _)) =  | 
| 
 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 
wenzelm 
parents: 
44299 
diff
changeset
 | 
231  | 
if Task_Graph.Keys.is_empty deps andalso is_canceled (group_of_task task)  | 
| 
 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 
wenzelm 
parents: 
44299 
diff
changeset
 | 
232  | 
then SOME (task, [fn _ => abort ()])  | 
| 
37854
 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
233  | 
else NONE  | 
| 
 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
234  | 
| ready_job _ _ = NONE;  | 
| 
 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
235  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
236  | 
fun active_job (_, (Job _, _)) = SOME ()  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
237  | 
| active_job (_, (Running _, _)) = SOME ()  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
238  | 
| active_job (task, (Passive _, _)) =  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
239  | 
if is_canceled (group_of_task task) then SOME () else NONE;  | 
| 
37854
 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
240  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
241  | 
fun all_passive (Queue {jobs, ...}) = is_none (Task_Graph.get_first active_job jobs);
 | 
| 28165 | 242  | 
|
243  | 
||
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
244  | 
(* queue status *)  | 
| 32052 | 245  | 
|
246  | 
fun status (Queue {jobs, ...}) =
 | 
|
247  | 
let  | 
|
| 
34277
 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 
wenzelm 
parents: 
32814 
diff
changeset
 | 
248  | 
val (x, y, z, w) =  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
249  | 
Task_Graph.fold (fn (_, (job, (deps, _))) => fn (x, y, z, w) =>  | 
| 32052 | 250  | 
(case job of  | 
| 
44338
 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 
wenzelm 
parents: 
44299 
diff
changeset
 | 
251  | 
Job _ => if Task_Graph.Keys.is_empty 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
 | 
252  | 
| 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
 | 
253  | 
| 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
 | 
254  | 
jobs (0, 0, 0, 0);  | 
| 
 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 
wenzelm 
parents: 
32814 
diff
changeset
 | 
255  | 
  in {ready = x, pending = y, running = z, passive = w} end;
 | 
| 32052 | 256  | 
|
257  | 
||
| 
41681
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
258  | 
|
| 
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
259  | 
(** task queue operations **)  | 
| 
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
260  | 
|
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
261  | 
(* cancel -- peers and sub-groups *)  | 
| 
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
262  | 
|
| 
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
 | 
263  | 
fun cancel (Queue {groups, jobs}) group =
 | 
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
264  | 
let  | 
| 
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
265  | 
val _ = cancel_group group Exn.Interrupt;  | 
| 
44341
 
a93d25fb14fc
purely functional task_queue.ML -- moved actual interrupt_unsynchronized to future.ML;
 
wenzelm 
parents: 
44340 
diff
changeset
 | 
266  | 
val running =  | 
| 
47404
 
e6e5750f1311
simplified Future.cancel/cancel_group (again) -- running threads only;
 
wenzelm 
parents: 
45642 
diff
changeset
 | 
267  | 
Tasks.fold (fn (task, _) =>  | 
| 
 
e6e5750f1311
simplified Future.cancel/cancel_group (again) -- running threads only;
 
wenzelm 
parents: 
45642 
diff
changeset
 | 
268  | 
(case get_job jobs task of Running thread => insert Thread.equal thread | _ => I))  | 
| 
 
e6e5750f1311
simplified Future.cancel/cancel_group (again) -- running threads only;
 
wenzelm 
parents: 
45642 
diff
changeset
 | 
269  | 
(get_tasks groups (group_id group)) [];  | 
| 
44341
 
a93d25fb14fc
purely functional task_queue.ML -- moved actual interrupt_unsynchronized to future.ML;
 
wenzelm 
parents: 
44340 
diff
changeset
 | 
270  | 
in running end;  | 
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
271  | 
|
| 
37854
 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
272  | 
fun cancel_all (Queue {jobs, ...}) =
 | 
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
273  | 
let  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
274  | 
fun cancel_job (task, (job, _)) (groups, running) =  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
275  | 
let  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
276  | 
val group = group_of_task task;  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
277  | 
val _ = cancel_group group Exn.Interrupt;  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
278  | 
in  | 
| 
34277
 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 
wenzelm 
parents: 
32814 
diff
changeset
 | 
279  | 
(case job of  | 
| 
 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 
wenzelm 
parents: 
32814 
diff
changeset
 | 
280  | 
Running t => (insert eq_group group groups, insert Thread.equal t running)  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
281  | 
| _ => (groups, running))  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
282  | 
end;  | 
| 
44341
 
a93d25fb14fc
purely functional task_queue.ML -- moved actual interrupt_unsynchronized to future.ML;
 
wenzelm 
parents: 
44340 
diff
changeset
 | 
283  | 
val running = Task_Graph.fold cancel_job jobs ([], []);  | 
| 
 
a93d25fb14fc
purely functional task_queue.ML -- moved actual interrupt_unsynchronized to future.ML;
 
wenzelm 
parents: 
44340 
diff
changeset
 | 
284  | 
in running end;  | 
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
285  | 
|
| 
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
286  | 
|
| 
41681
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
287  | 
(* finish *)  | 
| 
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
288  | 
|
| 
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
289  | 
fun finish task (Queue {groups, jobs}) =
 | 
| 
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
290  | 
let  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
291  | 
val group = group_of_task task;  | 
| 43951 | 292  | 
val groups' = fold_groups (fn g => del_task (group_id g, task)) group groups;  | 
| 
41681
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
293  | 
val jobs' = Task_Graph.del_node task jobs;  | 
| 
44338
 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 
wenzelm 
parents: 
44299 
diff
changeset
 | 
294  | 
val maximal = Task_Graph.is_maximal jobs task;  | 
| 
41681
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
295  | 
in (maximal, make_queue groups' jobs') end;  | 
| 
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
296  | 
|
| 
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
297  | 
|
| 
54649
 
99b9249b3e05
more official task context via Task_Queue.enroll, which is required to participate in group cancellation (e.g. to terminate command exec);
 
wenzelm 
parents: 
53190 
diff
changeset
 | 
298  | 
(* enroll *)  | 
| 
 
99b9249b3e05
more official task context via Task_Queue.enroll, which is required to participate in group cancellation (e.g. to terminate command exec);
 
wenzelm 
parents: 
53190 
diff
changeset
 | 
299  | 
|
| 
 
99b9249b3e05
more official task context via Task_Queue.enroll, which is required to participate in group cancellation (e.g. to terminate command exec);
 
wenzelm 
parents: 
53190 
diff
changeset
 | 
300  | 
fun enroll thread name group (Queue {groups, jobs}) =
 | 
| 
 
99b9249b3e05
more official task context via Task_Queue.enroll, which is required to participate in group cancellation (e.g. to terminate command exec);
 
wenzelm 
parents: 
53190 
diff
changeset
 | 
301  | 
let  | 
| 
 
99b9249b3e05
more official task context via Task_Queue.enroll, which is required to participate in group cancellation (e.g. to terminate command exec);
 
wenzelm 
parents: 
53190 
diff
changeset
 | 
302  | 
val task = new_task group name NONE;  | 
| 
 
99b9249b3e05
more official task context via Task_Queue.enroll, which is required to participate in group cancellation (e.g. to terminate command exec);
 
wenzelm 
parents: 
53190 
diff
changeset
 | 
303  | 
val groups' = fold_groups (fn g => add_task (group_id g, task)) group groups;  | 
| 
 
99b9249b3e05
more official task context via Task_Queue.enroll, which is required to participate in group cancellation (e.g. to terminate command exec);
 
wenzelm 
parents: 
53190 
diff
changeset
 | 
304  | 
val jobs' = jobs |> Task_Graph.new_node (task, Running thread);  | 
| 
 
99b9249b3e05
more official task context via Task_Queue.enroll, which is required to participate in group cancellation (e.g. to terminate command exec);
 
wenzelm 
parents: 
53190 
diff
changeset
 | 
305  | 
in (task, make_queue groups' jobs') end;  | 
| 
 
99b9249b3e05
more official task context via Task_Queue.enroll, which is required to participate in group cancellation (e.g. to terminate command exec);
 
wenzelm 
parents: 
53190 
diff
changeset
 | 
306  | 
|
| 
 
99b9249b3e05
more official task context via Task_Queue.enroll, which is required to participate in group cancellation (e.g. to terminate command exec);
 
wenzelm 
parents: 
53190 
diff
changeset
 | 
307  | 
|
| 
28185
 
0f20cbce4935
simplified dequeue: provide Thread.self internally;
 
wenzelm 
parents: 
28184 
diff
changeset
 | 
308  | 
(* enqueue *)  | 
| 28165 | 309  | 
|
| 
37854
 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
310  | 
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
 | 
311  | 
let  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
312  | 
val task = new_task group "passive" NONE;  | 
| 43951 | 313  | 
val groups' = fold_groups (fn g => add_task (group_id g, task)) group groups;  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
314  | 
val jobs' = jobs |> Task_Graph.new_node (task, 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
 | 
315  | 
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
 | 
316  | 
|
| 41673 | 317  | 
fun enqueue name group deps pri job (Queue {groups, jobs}) =
 | 
| 28165 | 318  | 
let  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
319  | 
val task = new_task group name (SOME pri);  | 
| 43951 | 320  | 
val groups' = fold_groups (fn g => add_task (group_id g, task)) group groups;  | 
| 
28185
 
0f20cbce4935
simplified dequeue: provide Thread.self internally;
 
wenzelm 
parents: 
28184 
diff
changeset
 | 
321  | 
val jobs' = jobs  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
322  | 
|> Task_Graph.new_node (task, Job [job])  | 
| 
41684
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
323  | 
|> fold (add_job task) deps;  | 
| 41708 | 324  | 
in (task, make_queue groups' jobs') end;  | 
| 28165 | 325  | 
|
| 
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
 | 
326  | 
fun extend task job (Queue {groups, jobs}) =
 | 
| 29365 | 327  | 
(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
 | 
328  | 
SOME (Job list) => SOME (make_queue groups (set_job task (Job (job :: list)) jobs))  | 
| 29365 | 329  | 
| _ => NONE);  | 
330  | 
||
| 
28185
 
0f20cbce4935
simplified dequeue: provide Thread.self internally;
 
wenzelm 
parents: 
28184 
diff
changeset
 | 
331  | 
|
| 
 
0f20cbce4935
simplified dequeue: provide Thread.self internally;
 
wenzelm 
parents: 
28184 
diff
changeset
 | 
332  | 
(* dequeue *)  | 
| 
 
0f20cbce4935
simplified dequeue: provide Thread.self internally;
 
wenzelm 
parents: 
28184 
diff
changeset
 | 
333  | 
|
| 
39243
 
307e3d07d19f
Future.promise: more robust treatment of concurrent abort vs. fulfill (amending 047c96f41455);
 
wenzelm 
parents: 
39232 
diff
changeset
 | 
334  | 
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
 | 
335  | 
(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
 | 
336  | 
SOME (Passive _) =>  | 
| 
 
307e3d07d19f
Future.promise: more robust treatment of concurrent abort vs. fulfill (amending 047c96f41455);
 
wenzelm 
parents: 
39232 
diff
changeset
 | 
337  | 
let val jobs' = set_job task (Running thread) jobs  | 
| 
47423
 
8a179a0493e3
more robust Future.fulfill wrt. duplicate assignment and interrupt;
 
wenzelm 
parents: 
47404 
diff
changeset
 | 
338  | 
in (SOME true, make_queue groups jobs') end  | 
| 
 
8a179a0493e3
more robust Future.fulfill wrt. duplicate assignment and interrupt;
 
wenzelm 
parents: 
47404 
diff
changeset
 | 
339  | 
| SOME _ => (SOME false, queue)  | 
| 
 
8a179a0493e3
more robust Future.fulfill wrt. duplicate assignment and interrupt;
 
wenzelm 
parents: 
47404 
diff
changeset
 | 
340  | 
| NONE => (NONE, queue));  | 
| 
39243
 
307e3d07d19f
Future.promise: more robust treatment of concurrent abort vs. fulfill (amending 047c96f41455);
 
wenzelm 
parents: 
39232 
diff
changeset
 | 
341  | 
|
| 
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
 | 
342  | 
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
 | 
343  | 
(case Task_Graph.get_first (uncurry ready_job) jobs of  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
344  | 
SOME (result as (task, _)) =>  | 
| 
37854
 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
345  | 
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
 | 
346  | 
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
 | 
347  | 
| NONE => (NONE, queue));  | 
| 
28202
 
23cb9a974630
added focus, which indicates a particular collection of high-priority tasks;
 
wenzelm 
parents: 
28196 
diff
changeset
 | 
348  | 
|
| 28165 | 349  | 
|
| 
41681
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
350  | 
(* dequeue wrt. dynamic dependencies *)  | 
| 
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
351  | 
|
| 
41695
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
352  | 
fun dequeue_deps thread deps (queue as Queue {groups, jobs}) =
 | 
| 
32055
 
6a46898aa805
recovered a version of dequeue_towards (cf. bb7b5a5942c7);
 
wenzelm 
parents: 
32052 
diff
changeset
 | 
353  | 
let  | 
| 
41684
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
354  | 
fun ready [] rest = (NONE, rev rest)  | 
| 
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
355  | 
| ready (task :: tasks) rest =  | 
| 
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
356  | 
(case try (Task_Graph.get_entry jobs) task of  | 
| 
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
357  | 
NONE => ready tasks rest  | 
| 
43792
 
d5803c3d537a
Table.lookup_key and Graph.get_entry allow to retrieve the original key, which is not necessarily identical to the given one;
 
wenzelm 
parents: 
41709 
diff
changeset
 | 
358  | 
| SOME (_, entry) =>  | 
| 
41684
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
359  | 
(case ready_job task entry of  | 
| 
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
360  | 
NONE => ready tasks (task :: rest)  | 
| 45642 | 361  | 
| some => (some, fold cons rest tasks)));  | 
| 
41684
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
362  | 
|
| 
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
363  | 
fun ready_dep _ [] = NONE  | 
| 
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
364  | 
| ready_dep seen (task :: tasks) =  | 
| 
41709
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
365  | 
if Tasks.defined seen task then ready_dep seen tasks  | 
| 
41684
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
366  | 
else  | 
| 
43792
 
d5803c3d537a
Table.lookup_key and Graph.get_entry allow to retrieve the original key, which is not necessarily identical to the given one;
 
wenzelm 
parents: 
41709 
diff
changeset
 | 
367  | 
let val entry as (_, (ds, _)) = #2 (Task_Graph.get_entry jobs task) in  | 
| 
41684
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
368  | 
(case ready_job task entry of  | 
| 
44338
 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 
wenzelm 
parents: 
44299 
diff
changeset
 | 
369  | 
NONE => ready_dep (Tasks.update (task, ()) seen) (Task_Graph.Keys.dest ds @ tasks)  | 
| 
41684
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
370  | 
| some => some)  | 
| 
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
371  | 
end;  | 
| 
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
372  | 
|
| 
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
373  | 
fun result (res as (task, _)) deps' =  | 
| 
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
 | 
374  | 
let val jobs' = set_job task (Running thread) jobs  | 
| 
41695
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
375  | 
in ((SOME res, deps'), make_queue groups jobs') end;  | 
| 
32055
 
6a46898aa805
recovered a version of dequeue_towards (cf. bb7b5a5942c7);
 
wenzelm 
parents: 
32052 
diff
changeset
 | 
376  | 
in  | 
| 
41684
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
377  | 
(case ready deps [] of  | 
| 
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
378  | 
(SOME res, deps') => result res deps'  | 
| 
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
379  | 
| (NONE, deps') =>  | 
| 
41709
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
380  | 
(case ready_dep Tasks.empty deps' of  | 
| 
41684
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
381  | 
SOME res => result res deps'  | 
| 
41695
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
382  | 
| NONE => ((NONE, deps'), queue)))  | 
| 
32055
 
6a46898aa805
recovered a version of dequeue_towards (cf. bb7b5a5942c7);
 
wenzelm 
parents: 
32052 
diff
changeset
 | 
383  | 
end;  | 
| 
 
6a46898aa805
recovered a version of dequeue_towards (cf. bb7b5a5942c7);
 
wenzelm 
parents: 
32052 
diff
changeset
 | 
384  | 
|
| 
41681
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
385  | 
end;  |