| author | wenzelm | 
| Tue, 05 Mar 2024 16:06:06 +0100 | |
| changeset 79777 | db9c6be8e236 | 
| parent 79446 | ec7a1603129a | 
| child 80809 | 4a64fc4d1cde | 
| 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  | 
| 
60610
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59333 
diff
changeset
 | 
18  | 
val urgent_pri: int  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
19  | 
type task  | 
| 
45136
 
2afb928c71ca
static dummy_task (again) to avoid a few extra allocations;
 
wenzelm 
parents: 
44341 
diff
changeset
 | 
20  | 
val dummy_task: task  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
21  | 
val group_of_task: task -> group  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
22  | 
val name_of_task: task -> string  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
23  | 
val pri_of_task: task -> int  | 
| 
68354
 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 
wenzelm 
parents: 
68129 
diff
changeset
 | 
24  | 
val eq_task: task * task -> bool  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
25  | 
val str_of_task: task -> string  | 
| 43951 | 26  | 
val str_of_task_groups: task -> string  | 
| 50975 | 27  | 
val task_statistics: task -> Properties.T  | 
| 
41695
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
28  | 
val running: task -> (unit -> 'a) -> 'a  | 
| 
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
29  | 
val joining: task -> (unit -> 'a) -> 'a  | 
| 
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
30  | 
val waiting: task -> task list -> (unit -> 'a) -> 'a  | 
| 28165 | 31  | 
type queue  | 
32  | 
val empty: queue  | 
|
| 66378 | 33  | 
val group_tasks: queue -> group list -> task list  | 
| 41708 | 34  | 
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
 | 
35  | 
val all_passive: queue -> bool  | 
| 68129 | 36  | 
val total_jobs: queue -> int  | 
| 
72038
 
254c324f31fd
clarified user counters: expose tasks to external monitor;
 
wenzelm 
parents: 
68598 
diff
changeset
 | 
37  | 
  val status: queue -> {ready: int, pending: int, running: int, passive: int, urgent: int}
 | 
| 78648 | 38  | 
val cancel: queue -> group -> Isabelle_Thread.T list  | 
39  | 
val cancel_all: queue -> group list * Isabelle_Thread.T list  | 
|
| 
41681
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
40  | 
val finish: task -> queue -> bool * queue  | 
| 78648 | 41  | 
val enroll: Isabelle_Thread.T -> string -> group -> queue -> task * queue  | 
| 66166 | 42  | 
val enqueue_passive: group -> string -> (unit -> bool) -> queue -> task * queue  | 
| 41708 | 43  | 
val enqueue: string -> group -> task list -> int -> (bool -> bool) -> queue -> task * queue  | 
| 29365 | 44  | 
val extend: task -> (bool -> bool) -> queue -> queue option  | 
| 78648 | 45  | 
val dequeue_passive: Isabelle_Thread.T -> task -> queue -> bool option * queue  | 
46  | 
val dequeue: Isabelle_Thread.T -> bool -> queue -> (task * (bool -> bool) list) option * queue  | 
|
47  | 
val dequeue_deps: Isabelle_Thread.T -> task list -> queue ->  | 
|
| 
41695
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
48  | 
(((task * (bool -> bool) list) option * task list) * queue)  | 
| 28165 | 49  | 
end;  | 
50  | 
||
| 
35242
 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 
wenzelm 
parents: 
35012 
diff
changeset
 | 
51  | 
structure Task_Queue: TASK_QUEUE =  | 
| 28165 | 52  | 
struct  | 
53  | 
||
| 52537 | 54  | 
val new_id = Counter.make ();  | 
| 
40450
 
8eab60e1baeb
private counter, to keep externalized ids a bit smaller;
 
wenzelm 
parents: 
39243 
diff
changeset
 | 
55  | 
|
| 
 
8eab60e1baeb
private counter, to keep externalized ids a bit smaller;
 
wenzelm 
parents: 
39243 
diff
changeset
 | 
56  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
57  | 
(** nested groups of tasks **)  | 
| 41670 | 58  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
59  | 
(* groups *)  | 
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
60  | 
|
| 
35242
 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 
wenzelm 
parents: 
35012 
diff
changeset
 | 
61  | 
abstype group = Group of  | 
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
62  | 
 {parent: group option,
 | 
| 
40450
 
8eab60e1baeb
private counter, to keep externalized ids a bit smaller;
 
wenzelm 
parents: 
39243 
diff
changeset
 | 
63  | 
id: int,  | 
| 
68598
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
64  | 
status: exn option Unsynchronized.ref}  | 
| 
35242
 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 
wenzelm 
parents: 
35012 
diff
changeset
 | 
65  | 
with  | 
| 29121 | 66  | 
|
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
67  | 
fun make_group (parent, id, status) = Group {parent = parent, id = id, status = status};
 | 
| 32052 | 68  | 
|
| 
68598
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
69  | 
fun new_group parent =  | 
| 
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
70  | 
make_group (parent, new_id (), Unsynchronized.ref NONE);  | 
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
71  | 
|
| 
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
72  | 
fun group_id (Group {id, ...}) = id;
 | 
| 
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
73  | 
fun eq_group (group1, group2) = group_id group1 = group_id group2;  | 
| 28551 | 74  | 
|
| 43951 | 75  | 
fun fold_groups f (g as Group {parent = NONE, ...}) a = f g a
 | 
76  | 
  | 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
 | 
77  | 
|
| 
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
78  | 
|
| 32221 | 79  | 
(* group status *)  | 
80  | 
||
| 
68598
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
81  | 
local  | 
| 
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
82  | 
|
| 
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
83  | 
fun is_canceled_unsynchronized (Group {parent, status, ...}) =
 | 
| 
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
84  | 
is_some (! status) orelse  | 
| 
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
85  | 
(case parent of NONE => false | SOME group => is_canceled_unsynchronized group);  | 
| 
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
86  | 
|
| 
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
87  | 
fun group_status_unsynchronized (Group {parent, status, ...}) =
 | 
| 
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
88  | 
the_list (! status) @  | 
| 
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
89  | 
(case parent of NONE => [] | SOME group => group_status_unsynchronized group);  | 
| 
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
90  | 
|
| 
78650
 
47d0c333d155
clarified signature: retain original Poly/ML names Thread.Thread, Thread.Mutex, Thread.ConditionVar and de-emphasize them for Isabelle/ML;
 
wenzelm 
parents: 
78648 
diff
changeset
 | 
91  | 
val lock = Thread.Mutex.mutex ();  | 
| 
68598
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
92  | 
fun SYNCHRONIZED e = Multithreading.synchronized "group_status" lock e;  | 
| 29121 | 93  | 
|
| 
68598
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
94  | 
in  | 
| 
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
95  | 
|
| 
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
96  | 
fun cancel_group (Group {status, ...}) exn =
 | 
| 
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
97  | 
SYNCHRONIZED (fn () =>  | 
| 
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
98  | 
Unsynchronized.change status (fn exns => SOME (Par_Exn.make (exn :: the_list exns))));  | 
| 32221 | 99  | 
|
| 
68598
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
100  | 
fun is_canceled group =  | 
| 
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
101  | 
SYNCHRONIZED (fn () => is_canceled_unsynchronized group);  | 
| 
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
102  | 
|
| 
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
103  | 
fun group_status group =  | 
| 
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
104  | 
SYNCHRONIZED (fn () => group_status_unsynchronized group);  | 
| 
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
105  | 
|
| 
 
d465b396ef85
just one global lock for group status: avoid proliferation of mutexes, condvars;
 
wenzelm 
parents: 
68354 
diff
changeset
 | 
106  | 
end;  | 
| 28165 | 107  | 
|
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
108  | 
fun str_of_group group =  | 
| 
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
109  | 
  (is_canceled group ? enclose "(" ")") (string_of_int (group_id group));
 | 
| 28179 | 110  | 
|
| 43951 | 111  | 
fun str_of_groups group =  | 
112  | 
space_implode "/" (map str_of_group (rev (fold_groups cons group [])));  | 
|
113  | 
||
| 
35242
 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 
wenzelm 
parents: 
35012 
diff
changeset
 | 
114  | 
end;  | 
| 
 
1c80c29086d7
eliminated opaque signature matching -- tends to cause problems with toplevel pp for abstract types;
 
wenzelm 
parents: 
35012 
diff
changeset
 | 
115  | 
|
| 28165 | 116  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
117  | 
(* tasks *)  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
118  | 
|
| 
60610
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59333 
diff
changeset
 | 
119  | 
val urgent_pri = 1000;  | 
| 
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59333 
diff
changeset
 | 
120  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
121  | 
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
 | 
122  | 
|
| 
45136
 
2afb928c71ca
static dummy_task (again) to avoid a few extra allocations;
 
wenzelm 
parents: 
44341 
diff
changeset
 | 
123  | 
val timing_start = (Time.zeroTime, Time.zeroTime, []): timing;  | 
| 
 
2afb928c71ca
static dummy_task (again) to avoid a few extra allocations;
 
wenzelm 
parents: 
44341 
diff
changeset
 | 
124  | 
|
| 
45354
 
a2157057024c
optional timing, to avoid redundant allocation of mutable cells;
 
wenzelm 
parents: 
45136 
diff
changeset
 | 
125  | 
fun new_timing () =  | 
| 
 
a2157057024c
optional timing, to avoid redundant allocation of mutable cells;
 
wenzelm 
parents: 
45136 
diff
changeset
 | 
126  | 
if ! Multithreading.trace < 2 then NONE  | 
| 
 
a2157057024c
optional timing, to avoid redundant allocation of mutable cells;
 
wenzelm 
parents: 
45136 
diff
changeset
 | 
127  | 
else SOME (Synchronized.var "timing" timing_start);  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
128  | 
|
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
129  | 
abstype task = Task of  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
130  | 
 {group: group,
 | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
131  | 
name: string,  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
132  | 
id: int,  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
133  | 
pri: int option,  | 
| 50975 | 134  | 
timing: timing Synchronized.var option,  | 
135  | 
pos: Position.T}  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
136  | 
with  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
137  | 
|
| 
45136
 
2afb928c71ca
static dummy_task (again) to avoid a few extra allocations;
 
wenzelm 
parents: 
44341 
diff
changeset
 | 
138  | 
val dummy_task =  | 
| 50975 | 139  | 
  Task {group = new_group NONE, name = "", id = 0, pri = NONE, timing = NONE,
 | 
140  | 
pos = Position.none};  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
141  | 
|
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
142  | 
fun new_task group name pri =  | 
| 50975 | 143  | 
  Task {group = group, name = name, id = new_id (), pri = pri, timing = new_timing (),
 | 
144  | 
pos = Position.thread_data ()};  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
145  | 
|
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
146  | 
fun group_of_task (Task {group, ...}) = group;
 | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
147  | 
fun name_of_task (Task {name, ...}) = name;
 | 
| 
68354
 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 
wenzelm 
parents: 
68129 
diff
changeset
 | 
148  | 
fun id_of_task (Task {id, ...}) = id;
 | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
149  | 
fun pri_of_task (Task {pri, ...}) = the_default 0 pri;
 | 
| 43951 | 150  | 
|
| 
68354
 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 
wenzelm 
parents: 
68129 
diff
changeset
 | 
151  | 
fun eq_task (task1, task2) = id_of_task task1 = id_of_task task2;  | 
| 
 
93d3c967802e
record active execution task and depend on it -- avoid new executions bumping into old ones;
 
wenzelm 
parents: 
68129 
diff
changeset
 | 
152  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
153  | 
fun str_of_task (Task {name, id, ...}) =
 | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
154  | 
  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
 | 
155  | 
|
| 43951 | 156  | 
fun str_of_task_groups task = str_of_task task ^ " in " ^ str_of_groups (group_of_task task);  | 
157  | 
||
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
158  | 
fun update_timing update (Task {timing, ...}) e =
 | 
| 78720 | 159  | 
Thread_Attributes.uninterruptible_body (fn run =>  | 
| 
41702
 
dca4c58c5d57
Task_Queue.update_timing: more precise treatment of interruptibility;
 
wenzelm 
parents: 
41695 
diff
changeset
 | 
160  | 
let  | 
| 
 
dca4c58c5d57
Task_Queue.update_timing: more precise treatment of interruptibility;
 
wenzelm 
parents: 
41695 
diff
changeset
 | 
161  | 
val start = Time.now ();  | 
| 78757 | 162  | 
val result = Exn.capture_body (run e);  | 
| 62826 | 163  | 
val t = Time.now () - start;  | 
| 
45136
 
2afb928c71ca
static dummy_task (again) to avoid a few extra allocations;
 
wenzelm 
parents: 
44341 
diff
changeset
 | 
164  | 
val _ = (case timing of NONE => () | SOME var => Synchronized.change var (update t));  | 
| 78720 | 165  | 
in Exn.release result end);  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
166  | 
|
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
167  | 
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
 | 
168  | 
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
 | 
169  | 
|
| 50975 | 170  | 
fun task_statistics (Task {name, id, timing, pos, ...}) =
 | 
171  | 
let  | 
|
172  | 
val (run, wait, wait_deps) =  | 
|
173  | 
(case timing of NONE => timing_start | SOME var => Synchronized.value var);  | 
|
174  | 
fun micros time = string_of_int (Time.toNanoseconds time div 1000);  | 
|
175  | 
in  | 
|
| 63806 | 176  | 
    [("now", Value.print_real (Time.toReal (Time.now ()))),
 | 
177  | 
     ("task_name", name), ("task_id", Value.print_int id),
 | 
|
| 50975 | 178  | 
     ("run", micros run), ("wait", micros wait), ("wait_deps", commas wait_deps)] @
 | 
179  | 
Position.properties_of pos  | 
|
180  | 
end;  | 
|
181  | 
||
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
182  | 
end;  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
183  | 
|
| 
77723
 
b761c91c2447
performance tuning: prefer functor Set() over Table();
 
wenzelm 
parents: 
72038 
diff
changeset
 | 
184  | 
structure Tasks = Set(type key = task val ord = task_ord);  | 
| 77731 | 185  | 
structure Task_Graph = Graph(Tasks.Key);  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
186  | 
|
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
187  | 
|
| 
41695
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
188  | 
(* timing *)  | 
| 
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
189  | 
|
| 
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
190  | 
fun running task =  | 
| 62826 | 191  | 
update_timing (fn t => fn (a, b, ds) => (a + t, b, ds)) task;  | 
| 
41695
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
192  | 
|
| 
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
193  | 
fun joining task =  | 
| 62826 | 194  | 
update_timing (fn t => fn (a, b, ds) => (a - t, b, ds)) task;  | 
| 
41695
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
195  | 
|
| 
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
196  | 
fun waiting task deps =  | 
| 
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
197  | 
update_timing (fn t => fn (a, b, ds) =>  | 
| 62826 | 198  | 
(a - t, b + t,  | 
| 
41702
 
dca4c58c5d57
Task_Queue.update_timing: more precise treatment of interruptibility;
 
wenzelm 
parents: 
41695 
diff
changeset
 | 
199  | 
if ! Multithreading.trace > 0  | 
| 
 
dca4c58c5d57
Task_Queue.update_timing: more precise treatment of interruptibility;
 
wenzelm 
parents: 
41695 
diff
changeset
 | 
200  | 
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
 | 
201  | 
|
| 
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
202  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
203  | 
|
| 
41681
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
204  | 
(** queue of jobs and groups **)  | 
| 
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
205  | 
|
| 
41709
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
206  | 
(* known group members *)  | 
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
207  | 
|
| 
77723
 
b761c91c2447
performance tuning: prefer functor Set() over Table();
 
wenzelm 
parents: 
72038 
diff
changeset
 | 
208  | 
type groups = Tasks.T Inttab.table;  | 
| 
41709
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
209  | 
|
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
210  | 
fun get_tasks (groups: groups) gid =  | 
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
211  | 
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
 | 
212  | 
|
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
213  | 
fun add_task (gid, task) groups =  | 
| 
77723
 
b761c91c2447
performance tuning: prefer functor Set() over Table();
 
wenzelm 
parents: 
72038 
diff
changeset
 | 
214  | 
Inttab.update (gid, Tasks.insert task (get_tasks groups gid)) groups;  | 
| 
41709
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
215  | 
|
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
216  | 
fun del_task (gid, task) groups =  | 
| 
77723
 
b761c91c2447
performance tuning: prefer functor Set() over Table();
 
wenzelm 
parents: 
72038 
diff
changeset
 | 
217  | 
let val tasks = Tasks.remove task (get_tasks groups gid) in  | 
| 
41709
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
218  | 
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
 | 
219  | 
else Inttab.update (gid, tasks) groups  | 
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
220  | 
end;  | 
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
221  | 
|
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
222  | 
|
| 
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
223  | 
(* job dependency graph *)  | 
| 28165 | 224  | 
|
225  | 
datatype job =  | 
|
| 29365 | 226  | 
Job of (bool -> bool) list |  | 
| 78648 | 227  | 
Running of Isabelle_Thread.T |  | 
| 
37854
 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
228  | 
Passive of unit -> bool;  | 
| 28165 | 229  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
230  | 
type jobs = job Task_Graph.T;  | 
| 
28176
 
01b21886e7f0
job: explicit 'ok' status -- false for canceled jobs;
 
wenzelm 
parents: 
28171 
diff
changeset
 | 
231  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
232  | 
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
 | 
233  | 
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
 | 
234  | 
|
| 32250 | 235  | 
fun add_job task dep (jobs: jobs) =  | 
236  | 
Task_Graph.add_edge (dep, task) jobs handle Task_Graph.UNDEF _ => jobs;  | 
|
237  | 
||
| 
28176
 
01b21886e7f0
job: explicit 'ok' status -- false for canceled jobs;
 
wenzelm 
parents: 
28171 
diff
changeset
 | 
238  | 
|
| 
41681
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
239  | 
(* queue *)  | 
| 
28176
 
01b21886e7f0
job: explicit 'ok' status -- false for canceled jobs;
 
wenzelm 
parents: 
28171 
diff
changeset
 | 
240  | 
|
| 68129 | 241  | 
datatype queue = Queue of {groups: groups, jobs: jobs, urgent: int, total: int};
 | 
| 28165 | 242  | 
|
| 68129 | 243  | 
fun make_queue groups jobs urgent total =  | 
244  | 
  Queue {groups = groups, jobs = jobs, urgent = urgent, total = total};
 | 
|
245  | 
||
246  | 
val empty = make_queue Inttab.empty Task_Graph.empty 0 0;  | 
|
| 
34277
 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 
wenzelm 
parents: 
32814 
diff
changeset
 | 
247  | 
|
| 66378 | 248  | 
fun group_tasks (Queue {groups, ...}) gs =
 | 
| 
77723
 
b761c91c2447
performance tuning: prefer functor Set() over Table();
 
wenzelm 
parents: 
72038 
diff
changeset
 | 
249  | 
fold (fn g => fn tasks => Tasks.merge (tasks, get_tasks groups (group_id g)))  | 
| 66378 | 250  | 
gs Tasks.empty  | 
| 
77723
 
b761c91c2447
performance tuning: prefer functor Set() over Table();
 
wenzelm 
parents: 
72038 
diff
changeset
 | 
251  | 
|> Tasks.dest;  | 
| 66378 | 252  | 
|
| 79446 | 253  | 
fun known_task (Queue {jobs, ...}) task = Task_Graph.defined jobs task;
 | 
| 41708 | 254  | 
|
| 
37854
 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
255  | 
|
| 
 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
256  | 
(* job status *)  | 
| 
 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
257  | 
|
| 59332 | 258  | 
fun ready_job (task, (Job list, (deps, _))) =  | 
| 
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
 | 
259  | 
if Task_Graph.Keys.is_empty deps then SOME (task, rev list) else NONE  | 
| 59332 | 260  | 
| ready_job (task, (Passive abort, (deps, _))) =  | 
| 
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
 | 
261  | 
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
 | 
262  | 
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
 | 
263  | 
else NONE  | 
| 59332 | 264  | 
| ready_job _ = NONE;  | 
| 
37854
 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
265  | 
|
| 
60610
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59333 
diff
changeset
 | 
266  | 
fun ready_job_urgent false = ready_job  | 
| 
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59333 
diff
changeset
 | 
267  | 
| ready_job_urgent true = (fn entry as (task, _) =>  | 
| 
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59333 
diff
changeset
 | 
268  | 
if pri_of_task task >= urgent_pri then ready_job entry else NONE);  | 
| 
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59333 
diff
changeset
 | 
269  | 
|
| 
59333
 
4ef80efc36c8
clarified active_job: take dependencies into account (e.g. future based on promise);
 
wenzelm 
parents: 
59332 
diff
changeset
 | 
270  | 
fun active_job (task, (Running _, _)) = SOME (task, [])  | 
| 
 
4ef80efc36c8
clarified active_job: take dependencies into account (e.g. future based on promise);
 
wenzelm 
parents: 
59332 
diff
changeset
 | 
271  | 
| active_job arg = ready_job arg;  | 
| 
37854
 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
272  | 
|
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
273  | 
fun all_passive (Queue {jobs, ...}) = is_none (Task_Graph.get_first active_job jobs);
 | 
| 28165 | 274  | 
|
| 68129 | 275  | 
fun total_jobs (Queue {total, ...}) = total;
 | 
276  | 
||
| 28165 | 277  | 
|
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
278  | 
(* queue status *)  | 
| 32052 | 279  | 
|
| 
72038
 
254c324f31fd
clarified user counters: expose tasks to external monitor;
 
wenzelm 
parents: 
68598 
diff
changeset
 | 
280  | 
fun status (Queue {jobs, urgent, ...}) =
 | 
| 32052 | 281  | 
let  | 
| 
34277
 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 
wenzelm 
parents: 
32814 
diff
changeset
 | 
282  | 
val (x, y, z, w) =  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
283  | 
Task_Graph.fold (fn (_, (job, (deps, _))) => fn (x, y, z, w) =>  | 
| 32052 | 284  | 
(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
 | 
285  | 
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
 | 
286  | 
| 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
 | 
287  | 
| 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
 | 
288  | 
jobs (0, 0, 0, 0);  | 
| 
72038
 
254c324f31fd
clarified user counters: expose tasks to external monitor;
 
wenzelm 
parents: 
68598 
diff
changeset
 | 
289  | 
  in {ready = x, pending = y, running = z, passive = w, urgent = urgent} end;
 | 
| 32052 | 290  | 
|
291  | 
||
| 
41681
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
292  | 
|
| 
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
293  | 
(** task queue operations **)  | 
| 
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
294  | 
|
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
295  | 
(* cancel -- peers and sub-groups *)  | 
| 
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
296  | 
|
| 
60610
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59333 
diff
changeset
 | 
297  | 
fun cancel (Queue {groups, jobs, ...}) group =
 | 
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
298  | 
let  | 
| 78681 | 299  | 
val _ = cancel_group group Isabelle_Thread.interrupt;  | 
| 
44341
 
a93d25fb14fc
purely functional task_queue.ML -- moved actual interrupt_unsynchronized to future.ML;
 
wenzelm 
parents: 
44340 
diff
changeset
 | 
300  | 
val running =  | 
| 
77723
 
b761c91c2447
performance tuning: prefer functor Set() over Table();
 
wenzelm 
parents: 
72038 
diff
changeset
 | 
301  | 
Tasks.fold (fn task =>  | 
| 78648 | 302  | 
(case get_job jobs task of Running thread => insert Isabelle_Thread.equal thread | _ => I))  | 
| 
47404
 
e6e5750f1311
simplified Future.cancel/cancel_group (again) -- running threads only;
 
wenzelm 
parents: 
45642 
diff
changeset
 | 
303  | 
(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
 | 
304  | 
in running end;  | 
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
305  | 
|
| 
37854
 
047c96f41455
back to more strict dependencies, even for canceled groups (reverting parts of 02936e77a07c);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
306  | 
fun cancel_all (Queue {jobs, ...}) =
 | 
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
307  | 
let  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
308  | 
fun cancel_job (task, (job, _)) (groups, running) =  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
309  | 
let  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
310  | 
val group = group_of_task task;  | 
| 78681 | 311  | 
val _ = cancel_group group Isabelle_Thread.interrupt;  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
312  | 
in  | 
| 
34277
 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 
wenzelm 
parents: 
32814 
diff
changeset
 | 
313  | 
(case job of  | 
| 78648 | 314  | 
Running t => (insert eq_group group groups, insert Isabelle_Thread.equal t running)  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
315  | 
| _ => (groups, running))  | 
| 
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
316  | 
end;  | 
| 
44341
 
a93d25fb14fc
purely functional task_queue.ML -- moved actual interrupt_unsynchronized to future.ML;
 
wenzelm 
parents: 
44340 
diff
changeset
 | 
317  | 
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
 | 
318  | 
in running end;  | 
| 
32101
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
319  | 
|
| 
 
e25107ff4f56
support for nested groups -- cancellation is propagated to peers and subgroups;
 
wenzelm 
parents: 
32099 
diff
changeset
 | 
320  | 
|
| 
41681
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
321  | 
(* finish *)  | 
| 
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
322  | 
|
| 68129 | 323  | 
fun finish task (Queue {groups, jobs, urgent, total}) =
 | 
| 
41681
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
324  | 
let  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
325  | 
val group = group_of_task task;  | 
| 43951 | 326  | 
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
 | 
327  | 
val jobs' = Task_Graph.del_node task jobs;  | 
| 68129 | 328  | 
val total' = total - 1;  | 
| 
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
 | 
329  | 
val maximal = Task_Graph.is_maximal jobs task;  | 
| 68129 | 330  | 
in (maximal, make_queue groups' jobs' urgent total') end;  | 
| 
41681
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
331  | 
|
| 
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
332  | 
|
| 
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
 | 
333  | 
(* 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
 | 
334  | 
|
| 68129 | 335  | 
fun enroll thread name group (Queue {groups, jobs, urgent, total}) =
 | 
| 
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
 | 
336  | 
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
 | 
337  | 
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
 | 
338  | 
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
 | 
339  | 
val jobs' = jobs |> Task_Graph.new_node (task, Running thread);  | 
| 68129 | 340  | 
val total' = total + 1;  | 
341  | 
in (task, make_queue groups' jobs' urgent total') end;  | 
|
| 
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
 | 
342  | 
|
| 
 
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
 | 
343  | 
|
| 
28185
 
0f20cbce4935
simplified dequeue: provide Thread.self internally;
 
wenzelm 
parents: 
28184 
diff
changeset
 | 
344  | 
(* enqueue *)  | 
| 28165 | 345  | 
|
| 68129 | 346  | 
fun enqueue_passive group name abort (Queue {groups, jobs, urgent, total}) =
 | 
| 
34277
 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 
wenzelm 
parents: 
32814 
diff
changeset
 | 
347  | 
let  | 
| 66166 | 348  | 
val task = new_task group name NONE;  | 
| 43951 | 349  | 
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
 | 
350  | 
val jobs' = jobs |> Task_Graph.new_node (task, Passive abort);  | 
| 68129 | 351  | 
val total' = total + 1;  | 
352  | 
in (task, make_queue groups' jobs' urgent total') end;  | 
|
| 
34277
 
7325a5e3587f
added Future.promise/fulfill -- promised futures that are fulfilled by external means;
 
wenzelm 
parents: 
32814 
diff
changeset
 | 
353  | 
|
| 68129 | 354  | 
fun enqueue name group deps pri job (Queue {groups, jobs, urgent, total}) =
 | 
| 28165 | 355  | 
let  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
356  | 
val task = new_task group name (SOME pri);  | 
| 43951 | 357  | 
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
 | 
358  | 
val jobs' = jobs  | 
| 
41683
 
73dde8006820
maintain Task_Queue.group within Task_Queue.task;
 
wenzelm 
parents: 
41682 
diff
changeset
 | 
359  | 
|> Task_Graph.new_node (task, Job [job])  | 
| 
41684
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
360  | 
|> fold (add_job task) deps;  | 
| 
60610
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59333 
diff
changeset
 | 
361  | 
val urgent' = if pri >= urgent_pri then urgent + 1 else urgent;  | 
| 68129 | 362  | 
val total' = total + 1;  | 
363  | 
in (task, make_queue groups' jobs' urgent' total') end;  | 
|
| 28165 | 364  | 
|
| 68129 | 365  | 
fun extend task job (Queue {groups, jobs, urgent, total}) =
 | 
| 29365 | 366  | 
(case try (get_job jobs) task of  | 
| 68129 | 367  | 
SOME (Job list) => SOME (make_queue groups (set_job task (Job (job :: list)) jobs) urgent total)  | 
| 29365 | 368  | 
| _ => NONE);  | 
369  | 
||
| 
28185
 
0f20cbce4935
simplified dequeue: provide Thread.self internally;
 
wenzelm 
parents: 
28184 
diff
changeset
 | 
370  | 
|
| 
 
0f20cbce4935
simplified dequeue: provide Thread.self internally;
 
wenzelm 
parents: 
28184 
diff
changeset
 | 
371  | 
(* dequeue *)  | 
| 
 
0f20cbce4935
simplified dequeue: provide Thread.self internally;
 
wenzelm 
parents: 
28184 
diff
changeset
 | 
372  | 
|
| 68129 | 373  | 
fun dequeue_passive thread task (queue as Queue {groups, jobs, urgent, total}) =
 | 
| 
39243
 
307e3d07d19f
Future.promise: more robust treatment of concurrent abort vs. fulfill (amending 047c96f41455);
 
wenzelm 
parents: 
39232 
diff
changeset
 | 
374  | 
(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
 | 
375  | 
SOME (Passive _) =>  | 
| 
 
307e3d07d19f
Future.promise: more robust treatment of concurrent abort vs. fulfill (amending 047c96f41455);
 
wenzelm 
parents: 
39232 
diff
changeset
 | 
376  | 
let val jobs' = set_job task (Running thread) jobs  | 
| 68129 | 377  | 
in (SOME true, make_queue groups jobs' urgent total) end  | 
| 
47423
 
8a179a0493e3
more robust Future.fulfill wrt. duplicate assignment and interrupt;
 
wenzelm 
parents: 
47404 
diff
changeset
 | 
378  | 
| SOME _ => (SOME false, queue)  | 
| 
 
8a179a0493e3
more robust Future.fulfill wrt. duplicate assignment and interrupt;
 
wenzelm 
parents: 
47404 
diff
changeset
 | 
379  | 
| NONE => (NONE, queue));  | 
| 
39243
 
307e3d07d19f
Future.promise: more robust treatment of concurrent abort vs. fulfill (amending 047c96f41455);
 
wenzelm 
parents: 
39232 
diff
changeset
 | 
380  | 
|
| 68129 | 381  | 
fun dequeue thread urgent_only (queue as Queue {groups, jobs, urgent, total}) =
 | 
| 
60610
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59333 
diff
changeset
 | 
382  | 
if not urgent_only orelse urgent > 0 then  | 
| 
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59333 
diff
changeset
 | 
383  | 
(case Task_Graph.get_first (ready_job_urgent urgent_only) jobs of  | 
| 
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59333 
diff
changeset
 | 
384  | 
SOME (result as (task, _)) =>  | 
| 
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59333 
diff
changeset
 | 
385  | 
let  | 
| 
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59333 
diff
changeset
 | 
386  | 
val jobs' = set_job task (Running thread) jobs;  | 
| 
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59333 
diff
changeset
 | 
387  | 
val urgent' = if pri_of_task task >= urgent_pri then urgent - 1 else urgent;  | 
| 68129 | 388  | 
in (SOME result, make_queue groups jobs' urgent' total) end  | 
| 
60610
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59333 
diff
changeset
 | 
389  | 
| NONE => (NONE, queue))  | 
| 
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59333 
diff
changeset
 | 
390  | 
else (NONE, queue);  | 
| 
28202
 
23cb9a974630
added focus, which indicates a particular collection of high-priority tasks;
 
wenzelm 
parents: 
28196 
diff
changeset
 | 
391  | 
|
| 28165 | 392  | 
|
| 
41681
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
393  | 
(* dequeue wrt. dynamic dependencies *)  | 
| 
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
394  | 
|
| 68129 | 395  | 
fun dequeue_deps thread deps (queue as Queue {groups, jobs, urgent, total}) =
 | 
| 
32055
 
6a46898aa805
recovered a version of dequeue_towards (cf. bb7b5a5942c7);
 
wenzelm 
parents: 
32052 
diff
changeset
 | 
396  | 
let  | 
| 
41684
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
397  | 
fun ready [] rest = (NONE, rev rest)  | 
| 
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
398  | 
| ready (task :: tasks) rest =  | 
| 
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
399  | 
(case try (Task_Graph.get_entry jobs) task of  | 
| 
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
400  | 
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
 | 
401  | 
| SOME (_, entry) =>  | 
| 59332 | 402  | 
(case ready_job (task, entry) of  | 
| 
41684
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
403  | 
NONE => ready tasks (task :: rest)  | 
| 45642 | 404  | 
| some => (some, fold cons rest tasks)));  | 
| 
41684
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
405  | 
|
| 
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
406  | 
fun ready_dep _ [] = NONE  | 
| 
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
407  | 
| ready_dep seen (task :: tasks) =  | 
| 
77723
 
b761c91c2447
performance tuning: prefer functor Set() over Table();
 
wenzelm 
parents: 
72038 
diff
changeset
 | 
408  | 
if Tasks.member seen task then ready_dep seen tasks  | 
| 
41684
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
409  | 
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
 | 
410  | 
let val entry as (_, (ds, _)) = #2 (Task_Graph.get_entry jobs task) in  | 
| 59332 | 411  | 
(case ready_job (task, entry) of  | 
| 
77723
 
b761c91c2447
performance tuning: prefer functor Set() over Table();
 
wenzelm 
parents: 
72038 
diff
changeset
 | 
412  | 
NONE => ready_dep (Tasks.insert task seen) (Task_Graph.Keys.dest ds @ tasks)  | 
| 
41684
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
413  | 
| some => some)  | 
| 
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
414  | 
end;  | 
| 
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
415  | 
|
| 
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
416  | 
fun result (res as (task, _)) deps' =  | 
| 
60610
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59333 
diff
changeset
 | 
417  | 
let  | 
| 
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59333 
diff
changeset
 | 
418  | 
val jobs' = set_job task (Running thread) jobs;  | 
| 
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59333 
diff
changeset
 | 
419  | 
val urgent' = if pri_of_task task >= urgent_pri then urgent - 1 else urgent;  | 
| 68129 | 420  | 
in ((SOME res, deps'), make_queue groups jobs' urgent' total) end;  | 
| 
32055
 
6a46898aa805
recovered a version of dequeue_towards (cf. bb7b5a5942c7);
 
wenzelm 
parents: 
32052 
diff
changeset
 | 
421  | 
in  | 
| 
41684
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
422  | 
(case ready deps [] of  | 
| 
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
423  | 
(SOME res, deps') => result res deps'  | 
| 
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
424  | 
| (NONE, deps') =>  | 
| 
41709
 
2e427f340ad1
more scalable collections of tasks, notably for totality of known group members;
 
wenzelm 
parents: 
41708 
diff
changeset
 | 
425  | 
(case ready_dep Tasks.empty deps' of  | 
| 
41684
 
b828d4b53386
refined Task_Queue.dequeue_deps (more incremental);
 
wenzelm 
parents: 
41683 
diff
changeset
 | 
426  | 
SOME res => result res deps'  | 
| 
41695
 
afdbec23b92b
eliminated slightly odd abstract type Task_Queue.deps;
 
wenzelm 
parents: 
41684 
diff
changeset
 | 
427  | 
| NONE => ((NONE, deps'), queue)))  | 
| 
32055
 
6a46898aa805
recovered a version of dequeue_towards (cf. bb7b5a5942c7);
 
wenzelm 
parents: 
32052 
diff
changeset
 | 
428  | 
end;  | 
| 
 
6a46898aa805
recovered a version of dequeue_towards (cf. bb7b5a5942c7);
 
wenzelm 
parents: 
32052 
diff
changeset
 | 
429  | 
|
| 62663 | 430  | 
|
431  | 
(* toplevel pretty printing *)  | 
|
432  | 
||
| 
62819
 
d3ff367a16a0
careful export of type-dependent functions, without losing their special status;
 
wenzelm 
parents: 
62663 
diff
changeset
 | 
433  | 
val _ = ML_system_pp (fn _ => fn _ => Pretty.to_polyml o Pretty.str o str_of_task);  | 
| 
 
d3ff367a16a0
careful export of type-dependent functions, without losing their special status;
 
wenzelm 
parents: 
62663 
diff
changeset
 | 
434  | 
val _ = ML_system_pp (fn _ => fn _ => Pretty.to_polyml o Pretty.str o str_of_group);  | 
| 62663 | 435  | 
|
| 
41681
 
b5d7b15166bf
Future.join_results: discontinued post-hoc recording of dynamic dependencies;
 
wenzelm 
parents: 
41680 
diff
changeset
 | 
436  | 
end;  |