src/Tools/Metis/src/Stream.sml
changeset 39348 6f9c9899f99f
child 39349 2d0a4361c3ef
equal deleted inserted replaced
39347:50dec19e682b 39348:6f9c9899f99f
       
     1 (* ========================================================================= *)
       
     2 (* A POSSIBLY-INFINITE STREAM DATATYPE FOR ML                                *)
       
     3 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the GNU GPL version 2 *)
       
     4 (* ========================================================================= *)
       
     5 
       
     6 structure Stream :> Stream =
       
     7 struct
       
     8 
       
     9 val K = Useful.K;
       
    10 
       
    11 val pair = Useful.pair;
       
    12 
       
    13 val funpow = Useful.funpow;
       
    14 
       
    15 (* ------------------------------------------------------------------------- *)
       
    16 (* The stream type.                                                          *)
       
    17 (* ------------------------------------------------------------------------- *)
       
    18 
       
    19 datatype 'a stream =
       
    20     Nil
       
    21   | Cons of 'a * (unit -> 'a stream);
       
    22 
       
    23 (* ------------------------------------------------------------------------- *)
       
    24 (* Stream constructors.                                                      *)
       
    25 (* ------------------------------------------------------------------------- *)
       
    26 
       
    27 fun repeat x = let fun rep () = Cons (x,rep) in rep () end;
       
    28 
       
    29 fun count n = Cons (n, fn () => count (n + 1));
       
    30 
       
    31 fun funpows f x = Cons (x, fn () => funpows f (f x));
       
    32 
       
    33 (* ------------------------------------------------------------------------- *)
       
    34 (* Stream versions of standard list operations: these should all terminate.  *)
       
    35 (* ------------------------------------------------------------------------- *)
       
    36 
       
    37 fun cons h t = Cons (h,t);
       
    38 
       
    39 fun null Nil = true
       
    40   | null (Cons _) = false;
       
    41 
       
    42 fun hd Nil = raise Empty
       
    43   | hd (Cons (h,_)) = h;
       
    44 
       
    45 fun tl Nil = raise Empty
       
    46   | tl (Cons (_,t)) = t ();
       
    47 
       
    48 fun hdTl s = (hd s, tl s);
       
    49 
       
    50 fun singleton s = Cons (s, K Nil);
       
    51 
       
    52 fun append Nil s = s ()
       
    53   | append (Cons (h,t)) s = Cons (h, fn () => append (t ()) s);
       
    54 
       
    55 fun map f =
       
    56     let
       
    57       fun m Nil = Nil
       
    58         | m (Cons (h,t)) = Cons (f h, m o t)
       
    59     in
       
    60       m
       
    61     end;
       
    62 
       
    63 fun maps f g =
       
    64     let
       
    65       fun mm s Nil = g s
       
    66         | mm s (Cons (x,xs)) =
       
    67           let
       
    68             val (y,s') = f x s
       
    69           in
       
    70             Cons (y, mm s' o xs)
       
    71           end
       
    72     in
       
    73       mm
       
    74     end;
       
    75 
       
    76 fun zipwith f =
       
    77     let
       
    78       fun z Nil _ = Nil
       
    79         | z _ Nil = Nil
       
    80         | z (Cons (x,xs)) (Cons (y,ys)) =
       
    81           Cons (f x y, fn () => z (xs ()) (ys ()))
       
    82     in
       
    83       z
       
    84     end;
       
    85 
       
    86 fun zip s t = zipwith pair s t;
       
    87 
       
    88 fun take 0 _ = Nil
       
    89   | take n Nil = raise Subscript
       
    90   | take 1 (Cons (x,_)) = Cons (x, K Nil)
       
    91   | take n (Cons (x,xs)) = Cons (x, fn () => take (n - 1) (xs ()));
       
    92 
       
    93 fun drop n s = funpow n tl s handle Empty => raise Subscript;
       
    94 
       
    95 (* ------------------------------------------------------------------------- *)
       
    96 (* Stream versions of standard list operations: these might not terminate.   *)
       
    97 (* ------------------------------------------------------------------------- *)
       
    98 
       
    99 local
       
   100   fun len n Nil = n
       
   101     | len n (Cons (_,t)) = len (n + 1) (t ());
       
   102 in
       
   103   fun length s = len 0 s;
       
   104 end;
       
   105 
       
   106 fun exists pred =
       
   107     let
       
   108       fun f Nil = false
       
   109         | f (Cons (h,t)) = pred h orelse f (t ())
       
   110     in
       
   111       f
       
   112     end;
       
   113 
       
   114 fun all pred = not o exists (not o pred);
       
   115 
       
   116 fun filter p Nil = Nil
       
   117   | filter p (Cons (x,xs)) =
       
   118     if p x then Cons (x, fn () => filter p (xs ())) else filter p (xs ());
       
   119 
       
   120 fun foldl f =
       
   121     let
       
   122       fun fold b Nil = b
       
   123         | fold b (Cons (h,t)) = fold (f (h,b)) (t ())
       
   124     in
       
   125       fold
       
   126     end;
       
   127 
       
   128 fun concat Nil = Nil
       
   129   | concat (Cons (Nil, ss)) = concat (ss ())
       
   130   | concat (Cons (Cons (x, xs), ss)) =
       
   131     Cons (x, fn () => concat (Cons (xs (), ss)));
       
   132 
       
   133 fun mapPartial f =
       
   134     let
       
   135       fun mp Nil = Nil
       
   136         | mp (Cons (h,t)) =
       
   137           case f h of
       
   138             NONE => mp (t ())
       
   139           | SOME h' => Cons (h', fn () => mp (t ()))
       
   140     in
       
   141       mp
       
   142     end;
       
   143 
       
   144 fun mapsPartial f g =
       
   145     let
       
   146       fun mp s Nil = g s
       
   147         | mp s (Cons (h,t)) =
       
   148           let
       
   149             val (h,s) = f h s
       
   150           in
       
   151             case h of
       
   152               NONE => mp s (t ())
       
   153             | SOME h => Cons (h, fn () => mp s (t ()))
       
   154           end
       
   155     in
       
   156       mp
       
   157     end;
       
   158 
       
   159 fun mapConcat f =
       
   160     let
       
   161       fun mc Nil = Nil
       
   162         | mc (Cons (h,t)) = append (f h) (fn () => mc (t ()))
       
   163     in
       
   164       mc
       
   165     end;
       
   166 
       
   167 fun mapsConcat f g =
       
   168     let
       
   169       fun mc s Nil = g s
       
   170         | mc s (Cons (h,t)) =
       
   171           let
       
   172             val (l,s) = f h s
       
   173           in
       
   174             append l (fn () => mc s (t ()))
       
   175           end
       
   176     in
       
   177       mc
       
   178     end;
       
   179 
       
   180 (* ------------------------------------------------------------------------- *)
       
   181 (* Stream operations.                                                        *)
       
   182 (* ------------------------------------------------------------------------- *)
       
   183 
       
   184 fun memoize Nil = Nil
       
   185   | memoize (Cons (h,t)) = Cons (h, Lazy.memoize (fn () => memoize (t ())));
       
   186 
       
   187 fun concatList [] = Nil
       
   188   | concatList (h :: t) = append h (fn () => concatList t);
       
   189 
       
   190 local
       
   191   fun toLst res Nil = rev res
       
   192     | toLst res (Cons (x, xs)) = toLst (x :: res) (xs ());
       
   193 in
       
   194   fun toList s = toLst [] s;
       
   195 end;
       
   196 
       
   197 fun fromList [] = Nil
       
   198   | fromList (x :: xs) = Cons (x, fn () => fromList xs);
       
   199 
       
   200 fun listConcat s = concat (map fromList s);
       
   201 
       
   202 fun toString s = implode (toList s);
       
   203 
       
   204 fun fromString s = fromList (explode s);
       
   205 
       
   206 fun toTextFile {filename = f} s =
       
   207     let
       
   208       val (h,close) =
       
   209           if f = "-" then (TextIO.stdOut, K ())
       
   210           else (TextIO.openOut f, TextIO.closeOut)
       
   211 
       
   212       fun toFile Nil = ()
       
   213         | toFile (Cons (x,y)) = (TextIO.output (h,x); toFile (y ()))
       
   214 
       
   215       val () = toFile s
       
   216     in
       
   217       close h
       
   218     end;
       
   219 
       
   220 fun fromTextFile {filename = f} =
       
   221     let
       
   222       val (h,close) =
       
   223           if f = "-" then (TextIO.stdIn, K ())
       
   224           else (TextIO.openIn f, TextIO.closeIn)
       
   225 
       
   226       fun strm () =
       
   227           case TextIO.inputLine h of
       
   228             NONE => (close h; Nil)
       
   229           | SOME s => Cons (s,strm)
       
   230     in
       
   231       memoize (strm ())
       
   232     end;
       
   233 
       
   234 end