# HG changeset patch # User wenzelm # Date 1220970572 -7200 # Node ID 01b21886e7f04735b29d2ebf12ece98f79f0c867 # Parent 6ab2cab482479461524aa8b295382dd8b6e6afca job: explicit 'ok' status -- false for canceled jobs; group is now non-optional; tuned signature; replaced low-level interrupts by group cancel operation; misc tuning; diff -r 6ab2cab48247 -r 01b21886e7f0 src/Pure/Concurrent/task_queue.ML --- a/src/Pure/Concurrent/task_queue.ML Tue Sep 09 16:17:08 2008 +0200 +++ b/src/Pure/Concurrent/task_queue.ML Tue Sep 09 16:29:32 2008 +0200 @@ -12,12 +12,10 @@ val new_group: unit -> group type queue val empty: queue - val enqueue: group option -> task list -> (unit -> bool) -> queue -> task * queue - val dequeue: Thread.thread -> queue -> (task * group option * (unit -> bool)) option * queue - val finished: task -> queue -> queue - val interrupt_task: queue -> task -> unit - val interrupt_group: queue -> group -> unit - val interrupt_task_group: queue -> task -> unit + val enqueue: group -> task list -> (bool -> bool) -> queue -> task * queue + val dequeue: Thread.thread -> queue -> (task * group * (unit -> bool)) option * queue + val cancel: group -> queue -> Thread.thread list * queue + val finish: task -> queue -> queue end; structure TaskQueue: TASK_QUEUE = @@ -31,83 +29,80 @@ fun new_group () = Group (serial ()); -(* queue of dependent jobs *) +(* jobs *) datatype job = - Job of unit -> bool | + Job of bool * (bool -> bool) | Running of Thread.thread; +type jobs = (group * job) IntGraph.T; + +fun get_group (jobs: jobs) (Task id) = #1 (IntGraph.get_node jobs id); +fun get_job (jobs: jobs) (Task id) = #2 (IntGraph.get_node jobs id); +fun map_job (Task id) f (jobs: jobs) = IntGraph.map_node id (apsnd f) jobs; + + +(* queue of grouped jobs *) + datatype queue = Queue of - {jobs: (group option * job) IntGraph.T, (*job dependency graph*) - cache: (task * group option * (unit -> bool)) Queue.T option, (*cache of ready tasks*) - groups: task list Inttab.table}; (*active group members*) + {groups: task list Inttab.table, (*groups with presently active members*) + jobs: jobs, (*job dependency graph*) + cache: (task * group * (unit -> bool)) Queue.T option}; (*cache of ready tasks*) -fun make_queue jobs cache groups = - Queue {jobs = jobs, cache = cache, groups = groups}; +fun make_queue groups jobs cache = + Queue {groups = groups, jobs = jobs, cache = cache}; + +val empty = make_queue Inttab.empty IntGraph.empty NONE; (* queue operations *) -val empty = make_queue IntGraph.empty NONE Inttab.empty; - -fun enqueue group deps job (Queue {jobs, groups, ...}) = +fun enqueue (group as Group gid) deps job (Queue {groups, jobs, ...}) = let val id = serial (); val task = Task id; + val groups' = Inttab.cons_list (gid, task) groups; fun add_dep (Task dep) G = IntGraph.add_edge_acyclic (dep, id) G handle IntGraph.UNDEF _ => G; - val jobs' = jobs |> IntGraph.new_node (id, (group, Job job)) |> fold add_dep deps; + val jobs' = jobs |> IntGraph.new_node (id, (group, Job (true, job))) |> fold add_dep deps; + in (task, make_queue groups' jobs' NONE) end; - val groups' = - (case group of - NONE => groups - | SOME (Group gid) => Inttab.cons_list (gid, task) groups); - in (task, make_queue jobs' NONE groups') end; - -fun dequeue thread (queue as Queue {jobs, cache, groups}) = +fun dequeue thread (queue as Queue {groups, jobs, cache}) = let + fun add_ready (id, ((group, Job (ok, job)), ([], _))) = + Queue.enqueue (Task id, group, (fn () => job ok)) + | add_ready _ = I; val ready = (case cache of SOME ready => ready - | NONE => - let - fun add (id, ((group, Job job), ([], _))) = Queue.enqueue (Task id, group, job) - | add _ = I; - in IntGraph.fold add jobs Queue.empty end); + | NONE => IntGraph.fold add_ready jobs Queue.empty); in - if Queue.is_empty ready then (NONE, make_queue jobs (SOME ready) groups) - else - let - val (result as (Task id, _, _), ready') = Queue.dequeue ready; - val jobs' = IntGraph.map_node id (fn (group, _) => (group, Running thread)) jobs; - in (SOME result, make_queue jobs' (SOME ready') groups) end + (case try Queue.dequeue ready of + NONE => (NONE, make_queue groups jobs (SOME ready)) + | SOME (result, ready') => + let val jobs' = map_job (#1 result) (K (Running thread)) jobs + in (SOME result, make_queue groups jobs' (SOME ready')) end) end; -fun finished (task as Task id) (Queue {jobs, groups, ...}) = - let - val groups' = - (case #1 (IntGraph.get_node jobs id) of - NONE => groups - | SOME (Group gid) => Inttab.remove_list (op =) (gid, task) groups); - val jobs' = IntGraph.del_nodes [id] jobs; - in make_queue jobs' NONE groups' end; - -(* interrupt *) +(* termination *) -fun interrupt_task (Queue {jobs, ...}) (Task id) = - (case IntGraph.get_node jobs id of - (_, Running thread) => (Thread.interrupt thread handle Thread _ => ()) - | _ => ()) - handle IntGraph.UNDEF _ => (); +fun cancel (group as Group gid) (Queue {groups, jobs, ...}) = + let + val tasks = Inttab.lookup_list groups gid; + val running = fold (get_job jobs #> (fn Running thread => cons thread | _ => I)) tasks []; + val jobs' = fold (fn task => + (case get_job jobs task of + Job (true, job) => map_job task (K (Job (false, job))) + | _ => I)) tasks jobs; + in (running, make_queue groups jobs' NONE) end; -fun interrupt_group (queue as Queue {groups, ...}) (Group gid) = - List.app (interrupt_task queue) (Inttab.lookup_list groups gid); - -fun interrupt_task_group (queue as Queue {jobs, ...}) (task as Task id) = - (case IntGraph.get_node jobs id of - (NONE, _) => interrupt_task queue task - | (SOME group, _) => interrupt_group queue group); +fun finish (task as Task id) (Queue {groups, jobs, ...}) = + let + val Group gid = get_group jobs task; + val groups' = Inttab.remove_list (op =) (gid, task) groups; + val jobs' = IntGraph.del_nodes [id] jobs; + in make_queue groups' jobs' NONE end; end;