src/Pure/Concurrent/par_exn.ML
author wenzelm
Wed, 17 Aug 2011 22:14:22 +0200
changeset 44247 270366301bd7
child 44248 6a3541412b23
permissions -rw-r--r--
more systematic handling of parallel exceptions; distinguish exception concatanation EXCEPTIONS from parallel Par_Exn;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
44247
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Concurrent/par_exn.ML
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
     3
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
     4
Parallel exceptions as flattened results from acyclic graph of
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
     5
evaluations.  Interrupt counts as neutral element.
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
     6
*)
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
     7
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
     8
signature PAR_EXN =
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
     9
sig
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    10
  val make: exn list -> exn
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    11
  val dest: exn -> (serial * exn) list option
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    12
  val release_all: 'a Exn.result list -> 'a list
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    13
  val release_first: 'a Exn.result list -> 'a list
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    14
end;
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    15
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    16
structure Par_Exn =
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    17
struct
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    18
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    19
(* parallel exceptions *)
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    20
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    21
exception Par_Exn of (serial * exn) Ord_List.T;
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    22
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    23
fun exn_ord ((i, _), (j, _)) = int_ord (j, i);
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    24
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    25
fun par_exns (Par_Exn exns) = SOME exns
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    26
  | par_exns exn = if Exn.is_interrupt exn then NONE else SOME [(serial (), exn)];
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    27
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    28
fun make exns =
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    29
  (case map_filter par_exns exns of
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    30
    [] => Exn.Interrupt
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    31
  | e :: es => Par_Exn (Library.foldl (Ord_List.merge exn_ord) (e, es)));
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    32
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    33
fun dest (Par_Exn exns) = SOME (rev exns)
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    34
  | dest _ = NONE;
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    35
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    36
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    37
(* parallel results *)
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    38
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    39
fun all_res results = forall (fn Exn.Res _ => true | _ => false) results;
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    40
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    41
fun release_all results =
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    42
  if all_res results then map Exn.release results
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    43
  else raise make (map_filter Exn.get_exn results);
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    44
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    45
fun release_first results = release_all results
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    46
  handle Par_Exn ((serial, exn) :: _) => reraise exn;  (* FIXME preserve serial in location (?!?) *)
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    47
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    48
end;
270366301bd7 more systematic handling of parallel exceptions;
wenzelm
parents:
diff changeset
    49