src/Pure/General/lazy.ML
author wenzelm
Thu, 23 Oct 2008 14:22:16 +0200
changeset 28673 d746a8c12c43
child 28971 300ec36a19af
permissions -rw-r--r--
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
28673
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/General/lazy.ML
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
     3
    Author:     Florian Haftmann and Makarius, TU Muenchen
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
     4
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
     5
Lazy evaluation with memoing.  Concurrency may lead to multiple
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
     6
attempts of evaluation -- the first result persists.
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
     7
*)
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
     8
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
     9
signature LAZY =
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    10
sig
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    11
  type 'a T
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    12
  val same: 'a T * 'a T -> bool
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    13
  val lazy: (unit -> 'a) -> 'a T
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    14
  val value: 'a -> 'a T
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    15
  val peek: 'a T -> 'a Exn.result option
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    16
  val force: 'a T -> 'a
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    17
  val map_force: ('a -> 'a) -> 'a T -> 'a T
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    18
end
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    19
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    20
structure Lazy :> LAZY =
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    21
struct
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    22
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    23
(* datatype *)
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    24
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    25
datatype 'a lazy =
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    26
  Lazy of unit -> 'a |
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    27
  Result of 'a Exn.result;
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    28
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    29
type 'a T = 'a lazy ref;
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    30
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    31
fun same (r1: 'a T, r2) = r1 = r2;
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    32
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    33
fun lazy e = ref (Lazy e);
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    34
fun value x = ref (Result (Exn.Result x));
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    35
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    36
fun peek r =
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    37
  (case ! r of
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    38
    Lazy _ => NONE
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    39
  | Result res => SOME res);
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    40
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    41
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    42
(* force result *)
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    43
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    44
fun force r =
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    45
  let
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    46
    (*potentially concurrent evaluation*)
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    47
    val res =
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    48
      (case ! r of
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    49
        Lazy e => Exn.capture e ()
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    50
      | Result res => res);
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    51
    (*assign result -- first one persists*)
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    52
    val res' = NAMED_CRITICAL "lazy" (fn () =>
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    53
      (case ! r of
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    54
        Lazy _ => (r := Result res; res)
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    55
      | Result res' => res'));
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    56
  in Exn.release res' end;
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    57
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    58
fun map_force f = value o f o force;
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    59
d746a8c12c43 renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff changeset
    60
end;