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