src/Pure/Concurrent/task_queue.ML
author wenzelm
Thu, 09 Oct 2008 21:06:08 +0200
changeset 28556 85d2972fe9e6
parent 28551 91eec4012bc5
child 28998 23cbaa9f9834
permissions -rw-r--r--
fixed spelling;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Concurrent/task_queue.ML
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
     3
    Author:     Makarius
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
     4
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
     5
Ordered queue of grouped tasks.
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
     6
*)
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
     7
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
     8
signature TASK_QUEUE =
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
     9
sig
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    10
  eqtype task
28196
wenzelm
parents: 28190
diff changeset
    11
  val str_of_task: task -> string
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    12
  eqtype group
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    13
  val new_group: unit -> group
28551
91eec4012bc5 added invalidate_group;
wenzelm
parents: 28384
diff changeset
    14
  val invalidate_group: group -> unit
28179
8e8313aededc human-readable printing of TaskQueue.task/group;
wenzelm
parents: 28176
diff changeset
    15
  val str_of_group: group -> string
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    16
  type queue
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    17
  val empty: queue
28204
2d93b158ad99 added is_empty;
wenzelm
parents: 28202
diff changeset
    18
  val is_empty: queue -> bool
28304
4b0477452943 future tasks: support boolean priorities (true = high, false = low/irrelevant);
wenzelm
parents: 28204
diff changeset
    19
  val enqueue: group -> task list -> bool -> (bool -> bool) -> queue -> task * queue
28185
0f20cbce4935 simplified dequeue: provide Thread.self internally;
wenzelm
parents: 28184
diff changeset
    20
  val depend: task list -> task -> queue -> queue
28202
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
    21
  val focus: task list -> queue -> queue
28185
0f20cbce4935 simplified dequeue: provide Thread.self internally;
wenzelm
parents: 28184
diff changeset
    22
  val dequeue: queue -> (task * group * (unit -> bool)) option * queue
28384
70abca69247b dequeue_towards: return bound for unfinished tasks;
wenzelm
parents: 28332
diff changeset
    23
  val dequeue_towards: task list -> queue ->
70abca69247b dequeue_towards: return bound for unfinished tasks;
wenzelm
parents: 28332
diff changeset
    24
    (((task * group * (unit -> bool)) * task list) option * queue)
28190
0a2434cf38c9 cancel: invalidate group implicitly, via bool ref;
wenzelm
parents: 28185
diff changeset
    25
  val interrupt: queue -> task -> unit
0a2434cf38c9 cancel: invalidate group implicitly, via bool ref;
wenzelm
parents: 28185
diff changeset
    26
  val interrupt_external: queue -> string -> unit
28176
01b21886e7f0 job: explicit 'ok' status -- false for canceled jobs;
wenzelm
parents: 28171
diff changeset
    27
  val finish: task -> queue -> queue
28190
0a2434cf38c9 cancel: invalidate group implicitly, via bool ref;
wenzelm
parents: 28185
diff changeset
    28
  val cancel: queue -> group -> bool
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    29
end;
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    30
28171
9b2f9cc9ff4b proper signature constraint;
wenzelm
parents: 28168
diff changeset
    31
structure TaskQueue: TASK_QUEUE =
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    32
struct
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    33
28168
ba410235ff04 moved thread data to future.ML (again);
wenzelm
parents: 28165
diff changeset
    34
(* identifiers *)
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    35
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    36
datatype task = Task of serial;
28190
0a2434cf38c9 cancel: invalidate group implicitly, via bool ref;
wenzelm
parents: 28185
diff changeset
    37
fun str_of_task (Task i) = string_of_int i;
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    38
28190
0a2434cf38c9 cancel: invalidate group implicitly, via bool ref;
wenzelm
parents: 28185
diff changeset
    39
datatype group = Group of serial * bool ref;
28551
91eec4012bc5 added invalidate_group;
wenzelm
parents: 28384
diff changeset
    40
28190
0a2434cf38c9 cancel: invalidate group implicitly, via bool ref;
wenzelm
parents: 28185
diff changeset
    41
fun new_group () = Group (serial (), ref true);
28551
91eec4012bc5 added invalidate_group;
wenzelm
parents: 28384
diff changeset
    42
fun invalidate_group (Group (_, ok)) = ok := false;
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    43
28190
0a2434cf38c9 cancel: invalidate group implicitly, via bool ref;
wenzelm
parents: 28185
diff changeset
    44
fun str_of_group (Group (i, ref ok)) =
0a2434cf38c9 cancel: invalidate group implicitly, via bool ref;
wenzelm
parents: 28185
diff changeset
    45
  if ok then string_of_int i else enclose "(" ")" (string_of_int i);
28179
8e8313aededc human-readable printing of TaskQueue.task/group;
wenzelm
parents: 28176
diff changeset
    46
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    47
28176
01b21886e7f0 job: explicit 'ok' status -- false for canceled jobs;
wenzelm
parents: 28171
diff changeset
    48
(* jobs *)
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    49
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    50
datatype job =
28304
4b0477452943 future tasks: support boolean priorities (true = high, false = low/irrelevant);
wenzelm
parents: 28204
diff changeset
    51
  Job of bool * (bool -> bool) |   (*priority, job: status -> status*)
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    52
  Running of Thread.thread;
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    53
28176
01b21886e7f0 job: explicit 'ok' status -- false for canceled jobs;
wenzelm
parents: 28171
diff changeset
    54
type jobs = (group * job) IntGraph.T;
01b21886e7f0 job: explicit 'ok' status -- false for canceled jobs;
wenzelm
parents: 28171
diff changeset
    55
01b21886e7f0 job: explicit 'ok' status -- false for canceled jobs;
wenzelm
parents: 28171
diff changeset
    56
fun get_group (jobs: jobs) (Task id) = #1 (IntGraph.get_node jobs id);
01b21886e7f0 job: explicit 'ok' status -- false for canceled jobs;
wenzelm
parents: 28171
diff changeset
    57
fun get_job (jobs: jobs) (Task id) = #2 (IntGraph.get_node jobs id);
01b21886e7f0 job: explicit 'ok' status -- false for canceled jobs;
wenzelm
parents: 28171
diff changeset
    58
fun map_job (Task id) f (jobs: jobs) = IntGraph.map_node id (apsnd f) jobs;
28202
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
    59
28185
0f20cbce4935 simplified dequeue: provide Thread.self internally;
wenzelm
parents: 28184
diff changeset
    60
fun add_job (Task id) (Task dep) (jobs: jobs) =
0f20cbce4935 simplified dequeue: provide Thread.self internally;
wenzelm
parents: 28184
diff changeset
    61
  IntGraph.add_edge_acyclic (dep, id) jobs handle IntGraph.UNDEF _ => jobs;
28176
01b21886e7f0 job: explicit 'ok' status -- false for canceled jobs;
wenzelm
parents: 28171
diff changeset
    62
28202
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
    63
fun check_job (jobs: jobs) (task as Task id) =
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
    64
  if can (IntGraph.get_node jobs) id then SOME task else NONE;
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
    65
28176
01b21886e7f0 job: explicit 'ok' status -- false for canceled jobs;
wenzelm
parents: 28171
diff changeset
    66
01b21886e7f0 job: explicit 'ok' status -- false for canceled jobs;
wenzelm
parents: 28171
diff changeset
    67
(* queue of grouped jobs *)
01b21886e7f0 job: explicit 'ok' status -- false for canceled jobs;
wenzelm
parents: 28171
diff changeset
    68
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    69
datatype queue = Queue of
28184
5ed5cb73a2e9 eliminated cache, access queue efficiently via IntGraph.get_first;
wenzelm
parents: 28179
diff changeset
    70
 {groups: task list Inttab.table,   (*groups with presently active members*)
28202
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
    71
  jobs: jobs,                       (*job dependency graph*)
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
    72
  focus: task list};                (*particular collection of high-priority tasks*)
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    73
28202
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
    74
fun make_queue groups jobs focus = Queue {groups = groups, jobs = jobs, focus = focus};
28204
2d93b158ad99 added is_empty;
wenzelm
parents: 28202
diff changeset
    75
28202
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
    76
val empty = make_queue Inttab.empty IntGraph.empty [];
28204
2d93b158ad99 added is_empty;
wenzelm
parents: 28202
diff changeset
    77
fun is_empty (Queue {jobs, ...}) = IntGraph.is_empty jobs;
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    78
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    79
28185
0f20cbce4935 simplified dequeue: provide Thread.self internally;
wenzelm
parents: 28184
diff changeset
    80
(* enqueue *)
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    81
28304
4b0477452943 future tasks: support boolean priorities (true = high, false = low/irrelevant);
wenzelm
parents: 28204
diff changeset
    82
fun enqueue (group as Group (gid, _)) deps pri job (Queue {groups, jobs, focus}) =
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    83
  let
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    84
    val id = serial ();
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    85
    val task = Task id;
28176
01b21886e7f0 job: explicit 'ok' status -- false for canceled jobs;
wenzelm
parents: 28171
diff changeset
    86
    val groups' = Inttab.cons_list (gid, task) groups;
28185
0f20cbce4935 simplified dequeue: provide Thread.self internally;
wenzelm
parents: 28184
diff changeset
    87
    val jobs' = jobs
28304
4b0477452943 future tasks: support boolean priorities (true = high, false = low/irrelevant);
wenzelm
parents: 28204
diff changeset
    88
      |> IntGraph.new_node (id, (group, Job (pri, job))) |> fold (add_job task) deps;
28202
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
    89
  in (task, make_queue groups' jobs' focus) end;
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
    90
28202
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
    91
fun depend deps task (Queue {groups, jobs, focus}) =
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
    92
  make_queue groups (fold (add_job task) deps jobs) focus;
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
    93
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
    94
fun focus tasks (Queue {groups, jobs, ...}) =
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
    95
  make_queue groups jobs (map_filter (check_job jobs) tasks);
28185
0f20cbce4935 simplified dequeue: provide Thread.self internally;
wenzelm
parents: 28184
diff changeset
    96
0f20cbce4935 simplified dequeue: provide Thread.self internally;
wenzelm
parents: 28184
diff changeset
    97
0f20cbce4935 simplified dequeue: provide Thread.self internally;
wenzelm
parents: 28184
diff changeset
    98
(* dequeue *)
0f20cbce4935 simplified dequeue: provide Thread.self internally;
wenzelm
parents: 28184
diff changeset
    99
28202
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   100
local
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   101
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   102
fun dequeue_result NONE queue = (NONE, queue)
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   103
  | dequeue_result (SOME (result as (task, _, _))) (Queue {groups, jobs, focus}) =
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   104
      (SOME result, make_queue groups (map_job task (K (Running (Thread.self ()))) jobs) focus);
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   105
28304
4b0477452943 future tasks: support boolean priorities (true = high, false = low/irrelevant);
wenzelm
parents: 28204
diff changeset
   106
fun dequeue_global req_pri (queue as Queue {jobs, ...}) =
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
   107
  let
28304
4b0477452943 future tasks: support boolean priorities (true = high, false = low/irrelevant);
wenzelm
parents: 28204
diff changeset
   108
    fun ready (id, ((group as Group (_, ref ok), Job (pri, job)), ([], _))) =
4b0477452943 future tasks: support boolean priorities (true = high, false = low/irrelevant);
wenzelm
parents: 28204
diff changeset
   109
          if pri = req_pri then SOME (Task id, group, (fn () => job ok)) else NONE
28184
5ed5cb73a2e9 eliminated cache, access queue efficiently via IntGraph.get_first;
wenzelm
parents: 28179
diff changeset
   110
      | ready _ = NONE;
28202
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   111
  in dequeue_result (IntGraph.get_first ready jobs) queue end;
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
   112
28202
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   113
fun dequeue_local focus (queue as Queue {jobs, ...}) =
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   114
  let
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   115
    fun ready id =
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   116
      (case IntGraph.get_node jobs id of
28304
4b0477452943 future tasks: support boolean priorities (true = high, false = low/irrelevant);
wenzelm
parents: 28204
diff changeset
   117
        (group as Group (_, ref ok), Job (_, job)) =>
28202
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   118
          if null (IntGraph.imm_preds jobs id) then SOME (Task id, group, (fn () => job ok))
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   119
          else NONE
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   120
      | _ => NONE);
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   121
    val ids = map (fn Task id => id) focus;
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   122
  in dequeue_result (get_first ready (IntGraph.all_preds jobs ids)) queue end;
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   123
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   124
in
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   125
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   126
fun dequeue (queue as Queue {focus, ...}) =
28304
4b0477452943 future tasks: support boolean priorities (true = high, false = low/irrelevant);
wenzelm
parents: 28204
diff changeset
   127
  (case dequeue_local focus queue of
4b0477452943 future tasks: support boolean priorities (true = high, false = low/irrelevant);
wenzelm
parents: 28204
diff changeset
   128
    (NONE, _) =>
4b0477452943 future tasks: support boolean priorities (true = high, false = low/irrelevant);
wenzelm
parents: 28204
diff changeset
   129
      (case dequeue_global true queue of (NONE, _) => dequeue_global false queue | res => res)
4b0477452943 future tasks: support boolean priorities (true = high, false = low/irrelevant);
wenzelm
parents: 28204
diff changeset
   130
  | res => res);
28185
0f20cbce4935 simplified dequeue: provide Thread.self internally;
wenzelm
parents: 28184
diff changeset
   131
0f20cbce4935 simplified dequeue: provide Thread.self internally;
wenzelm
parents: 28184
diff changeset
   132
fun dequeue_towards tasks (queue as Queue {jobs, ...}) =
28384
70abca69247b dequeue_towards: return bound for unfinished tasks;
wenzelm
parents: 28332
diff changeset
   133
  let val tasks' = map_filter (check_job jobs) tasks in
70abca69247b dequeue_towards: return bound for unfinished tasks;
wenzelm
parents: 28332
diff changeset
   134
    (case dequeue_local tasks' queue of
70abca69247b dequeue_towards: return bound for unfinished tasks;
wenzelm
parents: 28332
diff changeset
   135
      (NONE, queue') => (NONE, queue')
70abca69247b dequeue_towards: return bound for unfinished tasks;
wenzelm
parents: 28332
diff changeset
   136
    | (SOME work, queue') => (SOME (work, tasks'), queue'))
70abca69247b dequeue_towards: return bound for unfinished tasks;
wenzelm
parents: 28332
diff changeset
   137
  end;
28202
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   138
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   139
end;
28185
0f20cbce4935 simplified dequeue: provide Thread.self internally;
wenzelm
parents: 28184
diff changeset
   140
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
   141
28190
0a2434cf38c9 cancel: invalidate group implicitly, via bool ref;
wenzelm
parents: 28185
diff changeset
   142
(* sporadic interrupts *)
0a2434cf38c9 cancel: invalidate group implicitly, via bool ref;
wenzelm
parents: 28185
diff changeset
   143
0a2434cf38c9 cancel: invalidate group implicitly, via bool ref;
wenzelm
parents: 28185
diff changeset
   144
fun interrupt (Queue {jobs, ...}) task =
28551
91eec4012bc5 added invalidate_group;
wenzelm
parents: 28384
diff changeset
   145
  (case try (get_job jobs) task of SOME (Running thread) => SimpleThread.interrupt thread | _ => ());
28190
0a2434cf38c9 cancel: invalidate group implicitly, via bool ref;
wenzelm
parents: 28185
diff changeset
   146
0a2434cf38c9 cancel: invalidate group implicitly, via bool ref;
wenzelm
parents: 28185
diff changeset
   147
fun interrupt_external queue str =
0a2434cf38c9 cancel: invalidate group implicitly, via bool ref;
wenzelm
parents: 28185
diff changeset
   148
  (case Int.fromString str of SOME id => interrupt queue (Task id) | NONE => ());
0a2434cf38c9 cancel: invalidate group implicitly, via bool ref;
wenzelm
parents: 28185
diff changeset
   149
0a2434cf38c9 cancel: invalidate group implicitly, via bool ref;
wenzelm
parents: 28185
diff changeset
   150
28202
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   151
(* misc operations *)
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
   152
28551
91eec4012bc5 added invalidate_group;
wenzelm
parents: 28384
diff changeset
   153
fun cancel (Queue {groups, jobs, ...}) (group as Group (gid, _)) =
28176
01b21886e7f0 job: explicit 'ok' status -- false for canceled jobs;
wenzelm
parents: 28171
diff changeset
   154
  let
28551
91eec4012bc5 added invalidate_group;
wenzelm
parents: 28384
diff changeset
   155
    val _ = invalidate_group group;
28176
01b21886e7f0 job: explicit 'ok' status -- false for canceled jobs;
wenzelm
parents: 28171
diff changeset
   156
    val tasks = Inttab.lookup_list groups gid;
01b21886e7f0 job: explicit 'ok' status -- false for canceled jobs;
wenzelm
parents: 28171
diff changeset
   157
    val running = fold (get_job jobs #> (fn Running thread => cons thread | _ => I)) tasks [];
28551
91eec4012bc5 added invalidate_group;
wenzelm
parents: 28384
diff changeset
   158
    val _ = List.app SimpleThread.interrupt running;
28190
0a2434cf38c9 cancel: invalidate group implicitly, via bool ref;
wenzelm
parents: 28185
diff changeset
   159
  in null running end;
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
   160
28202
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   161
fun finish (task as Task id) (Queue {groups, jobs, focus}) =
28176
01b21886e7f0 job: explicit 'ok' status -- false for canceled jobs;
wenzelm
parents: 28171
diff changeset
   162
  let
28190
0a2434cf38c9 cancel: invalidate group implicitly, via bool ref;
wenzelm
parents: 28185
diff changeset
   163
    val Group (gid, _) = get_group jobs task;
28176
01b21886e7f0 job: explicit 'ok' status -- false for canceled jobs;
wenzelm
parents: 28171
diff changeset
   164
    val groups' = Inttab.remove_list (op =) (gid, task) groups;
28332
c33c8ad8de70 IntGraph.del_node;
wenzelm
parents: 28304
diff changeset
   165
    val jobs' = IntGraph.del_node id jobs;
28202
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   166
    val focus' = remove (op =) task focus;
23cb9a974630 added focus, which indicates a particular collection of high-priority tasks;
wenzelm
parents: 28196
diff changeset
   167
  in make_queue groups' jobs' focus' end;
28165
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
   168
26bb048f463c Ordered queue of grouped tasks.
wenzelm
parents:
diff changeset
   169
end;