| author | wenzelm | 
| Thu, 01 Feb 2018 15:31:25 +0100 | |
| changeset 67561 | f0b11413f1c9 | 
| parent 62923 | 3a122e1e352a | 
| child 67658 | 67e5deb9e290 | 
| 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
 | 
| 
58993
 
302104d8366b
prefer independent parallel map where user input is processed -- avoid non-deterministic feedback in error situations;
 
wenzelm 
parents: 
52945 
diff
changeset
 | 
21  | 
  val map_independent: ('a -> 'b) -> 'a list -> 'b list
 | 
| 28199 | 22  | 
  val map: ('a -> 'b) -> 'a list -> 'b list
 | 
23  | 
  val get_some: ('a -> 'b option) -> 'a list -> 'b option
 | 
|
24  | 
  val find_some: ('a -> bool) -> 'a list -> 'a option
 | 
|
25  | 
  val exists: ('a -> bool) -> 'a list -> bool
 | 
|
26  | 
  val forall: ('a -> bool) -> 'a list -> bool
 | 
|
27  | 
end;  | 
|
28  | 
||
| 29368 | 29  | 
structure Par_List: PAR_LIST =  | 
| 28199 | 30  | 
struct  | 
31  | 
||
| 41673 | 32  | 
fun managed_results name f xs =  | 
| 
59180
 
c0fa3b3bdabd
discontinued central critical sections: NAMED_CRITICAL / CRITICAL;
 
wenzelm 
parents: 
59137 
diff
changeset
 | 
33  | 
if null xs orelse null (tl xs) orelse not (Multithreading.enabled ())  | 
| 41675 | 34  | 
then map (Exn.capture f) xs  | 
35  | 
else  | 
|
| 62923 | 36  | 
Thread_Attributes.uninterruptible (fn restore_attributes => fn () =>  | 
| 49319 | 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 =>  | 
|
| 62505 | 49  | 
(if Exn.is_interrupt exn then Future.cancel_group group else (); Exn.reraise exn);  | 
| 49319 | 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;  | 
| 
58993
 
302104d8366b
prefer independent parallel map where user input is processed -- avoid non-deterministic feedback in error situations;
 
wenzelm 
parents: 
52945 
diff
changeset
 | 
54  | 
fun map_independent f = map (Exn.interruptible_capture f) #> Par_Exn.release_all;  | 
| 28199 | 55  | 
|
56  | 
fun get_some f xs =  | 
|
57  | 
let  | 
|
| 
59136
 
c2b23cb8a677
added Par_List in Scala, in accordance to ML version;
 
wenzelm 
parents: 
58993 
diff
changeset
 | 
58  | 
exception FOUND of 'b;  | 
| 
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
 | 
59  | 
val results =  | 
| 41673 | 60  | 
managed_results "Par_List.get_some"  | 
| 
59136
 
c2b23cb8a677
added Par_List in Scala, in accordance to ML version;
 
wenzelm 
parents: 
58993 
diff
changeset
 | 
61  | 
(fn x => (case f x of NONE => () | SOME y => raise FOUND y)) xs;  | 
| 28199 | 62  | 
in  | 
| 
59136
 
c2b23cb8a677
added Par_List in Scala, in accordance to ML version;
 
wenzelm 
parents: 
58993 
diff
changeset
 | 
63  | 
(case get_first (fn Exn.Exn (FOUND res) => SOME res | _ => NONE) results of  | 
| 
 
c2b23cb8a677
added Par_List in Scala, in accordance to ML version;
 
wenzelm 
parents: 
58993 
diff
changeset
 | 
64  | 
NONE => (Par_Exn.release_first results; NONE)  | 
| 
 
c2b23cb8a677
added Par_List in Scala, in accordance to ML version;
 
wenzelm 
parents: 
58993 
diff
changeset
 | 
65  | 
| some => some)  | 
| 28199 | 66  | 
end;  | 
67  | 
||
68  | 
fun find_some P = get_some (fn x => if P x then SOME x else NONE);  | 
|
69  | 
||
| 59137 | 70  | 
fun exists P = is_some o find_some P;  | 
| 28199 | 71  | 
fun forall P = not o exists (not o P);  | 
72  | 
||
73  | 
end;  |