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