author | hoelzl |
Wed, 02 Apr 2014 18:35:01 +0200 | |
changeset 56369 | 2704ca85be98 |
parent 52945 | 13687b130f1f |
child 58993 | 302104d8366b |
permissions | -rw-r--r-- |
28199 | 1 |
(* Title: Pure/Concurrent/par_list.ML |
2 |
Author: Makarius |
|
3 |
||
4 |
Parallel list combinators. |
|
5 |
||
6 |
Notes: |
|
7 |
||
8 |
* These combinators only make sense if the operator (function or |
|
9 |
predicate) applied to the list of operands takes considerable |
|
10 |
time. The overhead of scheduling is significantly higher than |
|
11 |
just traversing the list of operands sequentially. |
|
12 |
||
28358 | 13 |
* The order of operator application is non-deterministic. Watch out |
28199 | 14 |
for operators that have side-effects or raise exceptions! |
15 |
*) |
|
16 |
||
17 |
signature PAR_LIST = |
|
18 |
sig |
|
44265
b94951f06e48
export Par_List.managed_results, to enable specific treatment of results apart from default Par_Exn.release_first;
wenzelm
parents:
44247
diff
changeset
|
19 |
val managed_results: string -> ('a -> 'b) -> 'a list -> 'b Exn.result list |
41711 | 20 |
val map_name: string -> ('a -> 'b) -> 'a list -> 'b list |
28199 | 21 |
val map: ('a -> 'b) -> 'a list -> 'b list |
22 |
val get_some: ('a -> 'b option) -> 'a list -> 'b option |
|
23 |
val find_some: ('a -> bool) -> 'a list -> 'a option |
|
24 |
val exists: ('a -> bool) -> 'a list -> bool |
|
25 |
val forall: ('a -> bool) -> 'a list -> bool |
|
26 |
end; |
|
27 |
||
29368 | 28 |
structure Par_List: PAR_LIST = |
28199 | 29 |
struct |
30 |
||
41673 | 31 |
fun managed_results name f xs = |
41675 | 32 |
if null xs orelse null (tl xs) orelse |
33 |
not (Multithreading.enabled ()) orelse Multithreading.self_critical () |
|
34 |
then map (Exn.capture f) xs |
|
35 |
else |
|
49319 | 36 |
uninterruptible (fn restore_attributes => fn () => |
37 |
let |
|
52945
13687b130f1f
retain current task priority, to avoid distortion due to local forks (e.g. sledgehammer query operation);
wenzelm
parents:
49319
diff
changeset
|
38 |
val (group, pri) = |
13687b130f1f
retain current task priority, to avoid distortion due to local forks (e.g. sledgehammer query operation);
wenzelm
parents:
49319
diff
changeset
|
39 |
(case Future.worker_task () of |
13687b130f1f
retain current task priority, to avoid distortion due to local forks (e.g. sledgehammer query operation);
wenzelm
parents:
49319
diff
changeset
|
40 |
SOME task => |
13687b130f1f
retain current task priority, to avoid distortion due to local forks (e.g. sledgehammer query operation);
wenzelm
parents:
49319
diff
changeset
|
41 |
(Future.new_group (SOME (Task_Queue.group_of_task task)), Task_Queue.pri_of_task task) |
13687b130f1f
retain current task priority, to avoid distortion due to local forks (e.g. sledgehammer query operation);
wenzelm
parents:
49319
diff
changeset
|
42 |
| NONE => (Future.new_group NONE, 0)); |
49319 | 43 |
val futures = |
52945
13687b130f1f
retain current task priority, to avoid distortion due to local forks (e.g. sledgehammer query operation);
wenzelm
parents:
49319
diff
changeset
|
44 |
Future.forks {name = name, group = SOME group, deps = [], pri = pri, interrupts = true} |
49319 | 45 |
(map (fn x => fn () => f x) xs); |
46 |
val results = |
|
47 |
restore_attributes Future.join_results futures |
|
48 |
handle exn => |
|
49 |
(if Exn.is_interrupt exn then Future.cancel_group group else (); reraise exn); |
|
50 |
in results end) (); |
|
28199 | 51 |
|
44247 | 52 |
fun map_name name f xs = Par_Exn.release_first (managed_results name f xs); |
41711 | 53 |
fun map f = map_name "Par_List.map" f; |
28199 | 54 |
|
55 |
fun get_some f xs = |
|
56 |
let |
|
57 |
exception FOUND of 'b option; |
|
58 |
fun found (Exn.Exn (FOUND some)) = some |
|
59 |
| found _ = NONE; |
|
40700
4b4dfe05b5d7
clarified Par_List.managed_results, with explicit propagation of outermost physical interrupt to forked futures (e.g. to make timeout apply here as expected and prevent zombies);
wenzelm
parents:
37865
diff
changeset
|
60 |
val results = |
41673 | 61 |
managed_results "Par_List.get_some" |
62 |
(fn x => (case f x of NONE => () | some => raise FOUND some)) xs; |
|
28199 | 63 |
in |
64 |
(case get_first found results of |
|
65 |
SOME y => SOME y |
|
44247 | 66 |
| NONE => (Par_Exn.release_first results; NONE)) |
28199 | 67 |
end; |
68 |
||
69 |
fun find_some P = get_some (fn x => if P x then SOME x else NONE); |
|
70 |
||
71 |
fun exists P = is_some o get_some (fn x => if P x then SOME () else NONE); |
|
72 |
fun forall P = not o exists (not o P); |
|
73 |
||
74 |
end; |