| author | wenzelm | 
| Sun, 08 Mar 2009 16:53:07 +0100 | |
| changeset 30359 | 3f9b3ff851ca | 
| parent 29897 | 9049a2bfbe6d | 
| child 32169 | fbada8ed12e6 | 
| permissions | -rw-r--r-- | 
| 6118 | 1 | (* Title: Pure/General/seq.ML | 
| 5014 | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 12851 | 3 | Author: Markus Wenzel, TU Munich | 
| 5014 | 4 | |
| 5 | Unbounded sequences implemented by closures. RECOMPUTES if sequence | |
| 6 | is re-inspected. Memoing, using polymorphic refs, was found to be | |
| 7 | slower! (More GCs) | |
| 8 | *) | |
| 9 | ||
| 10 | signature SEQ = | |
| 11 | sig | |
| 12 | type 'a seq | |
| 13 |   val make: (unit -> ('a * 'a seq) option) -> 'a seq
 | |
| 14 |   val pull: 'a seq -> ('a * 'a seq) option
 | |
| 15 | val empty: 'a seq | |
| 19475 | 16 | val cons: 'a -> 'a seq -> 'a seq | 
| 5014 | 17 | val single: 'a -> 'a seq | 
| 11721 | 18 |   val try: ('a -> 'b) -> 'a -> 'b seq
 | 
| 5014 | 19 | val hd: 'a seq -> 'a | 
| 20 | val tl: 'a seq -> 'a seq | |
| 19865 | 21 | val chop: int -> 'a seq -> 'a list * 'a seq | 
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29606diff
changeset | 22 | val take: int -> 'a seq -> 'a seq | 
| 5014 | 23 | val list_of: 'a seq -> 'a list | 
| 24 | val of_list: 'a list -> 'a seq | |
| 19865 | 25 | val append: 'a seq -> 'a seq -> 'a seq | 
| 5014 | 26 |   val mapp: ('a -> 'b) -> 'a seq -> 'b seq -> 'b seq
 | 
| 17347 | 27 | val interleave: 'a seq * 'a seq -> 'a seq | 
| 5014 | 28 |   val filter: ('a -> bool) -> 'a seq -> 'a seq
 | 
| 29 | val flat: 'a seq seq -> 'a seq | |
| 17347 | 30 |   val map: ('a -> 'b) -> 'a seq -> 'b seq
 | 
| 31 |   val maps: ('a -> 'b seq) -> 'a seq -> 'b seq
 | |
| 19484 | 32 |   val map_filter: ('a -> 'b option) -> 'a seq -> 'b seq
 | 
| 17347 | 33 |   val lift: ('a -> 'b -> 'c) -> 'a seq -> 'b -> 'c seq
 | 
| 34 |   val lifts: ('a -> 'b -> 'c seq) -> 'a seq -> 'b -> 'c seq
 | |
| 19912 | 35 |   val singleton: ('a list -> 'b list seq) -> 'a -> 'b seq
 | 
| 25955 | 36 |   val wrap: ((unit -> ('a * 'a seq) option) -> ('a * 'a seq) option) -> 'a seq -> 'a seq
 | 
| 5014 | 37 | val print: (int -> 'a -> unit) -> int -> 'a seq -> unit | 
| 38 |   val it_right : ('a * 'b seq -> 'b seq) -> 'a seq * 'b seq -> 'b seq
 | |
| 39 | val succeed: 'a -> 'a seq | |
| 40 | val fail: 'a -> 'b seq | |
| 41 |   val THEN: ('a -> 'b seq) * ('b -> 'c seq) -> 'a -> 'c seq
 | |
| 42 |   val ORELSE: ('a -> 'b seq) * ('a -> 'b seq) -> 'a -> 'b seq
 | |
| 43 |   val APPEND: ('a -> 'b seq) * ('a -> 'b seq) -> 'a -> 'b seq
 | |
| 44 |   val EVERY: ('a -> 'a seq) list -> 'a -> 'a seq
 | |
| 45 |   val FIRST: ('a -> 'b seq) list -> 'a -> 'b seq
 | |
| 46 |   val TRY: ('a -> 'a seq) -> 'a -> 'a seq
 | |
| 47 |   val REPEAT: ('a -> 'a seq) -> 'a -> 'a seq
 | |
| 5558 | 48 |   val REPEAT1: ('a -> 'a seq) -> 'a -> 'a seq
 | 
| 8535 | 49 | val INTERVAL: (int -> 'a -> 'a seq) -> int -> int -> 'a -> 'a seq | 
| 12851 | 50 |   val DETERM: ('a -> 'b seq) -> 'a -> 'b seq
 | 
| 5014 | 51 | end; | 
| 52 | ||
| 53 | structure Seq: SEQ = | |
| 54 | struct | |
| 55 | ||
| 56 | ||
| 57 | (** lazy sequences **) | |
| 58 | ||
| 59 | datatype 'a seq = Seq of unit -> ('a * 'a seq) option;
 | |
| 60 | ||
| 61 | (*the abstraction for making a sequence*) | |
| 62 | val make = Seq; | |
| 63 | ||
| 15531 | 64 | (*return next sequence element as NONE or SOME (x, xq)*) | 
| 5014 | 65 | fun pull (Seq f) = f (); | 
| 66 | ||
| 67 | ||
| 68 | (*the empty sequence*) | |
| 15531 | 69 | val empty = Seq (fn () => NONE); | 
| 5014 | 70 | |
| 71 | (*prefix an element to the sequence -- use cons (x, xq) only if | |
| 72 | evaluation of xq need not be delayed, otherwise use | |
| 15531 | 73 | make (fn () => SOME (x, xq))*) | 
| 19475 | 74 | fun cons x xq = make (fn () => SOME (x, xq)); | 
| 5014 | 75 | |
| 19475 | 76 | fun single x = cons x empty; | 
| 5014 | 77 | |
| 78 | (*head and tail -- beware of calling the sequence function twice!!*) | |
| 16002 | 79 | fun hd xq = #1 (the (pull xq)) | 
| 80 | and tl xq = #2 (the (pull xq)); | |
| 5014 | 81 | |
| 11721 | 82 | (*partial function as procedure*) | 
| 83 | fun try f x = | |
| 21395 
f34ac19659ae
moved some fundamental concepts to General/basics.ML;
 wenzelm parents: 
21270diff
changeset | 84 | (case Basics.try f x of | 
| 15531 | 85 | SOME y => single y | 
| 86 | | NONE => empty); | |
| 11721 | 87 | |
| 5014 | 88 | |
| 89 | (*the list of the first n elements, paired with rest of sequence; | |
| 90 | if length of list is less than n, then sequence had less than n elements*) | |
| 19865 | 91 | fun chop n xq = | 
| 29897 | 92 | if n <= (0 : int) then ([], xq) | 
| 5014 | 93 | else | 
| 94 | (case pull xq of | |
| 15531 | 95 | NONE => ([], xq) | 
| 21395 
f34ac19659ae
moved some fundamental concepts to General/basics.ML;
 wenzelm parents: 
21270diff
changeset | 96 | | SOME (x, xq') => apfst (Basics.cons x) (chop (n - 1) xq')); | 
| 5014 | 97 | |
| 29897 | 98 | (*truncate the sequence after n elements*) | 
| 99 | fun take n xq = | |
| 100 | if n <= (0 : int) then empty | |
| 101 | else make (fn () => | |
| 102 | (Option.map o apsnd) (take (n - 1)) (pull xq)); | |
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29606diff
changeset | 103 | |
| 5014 | 104 | (*conversion from sequence to list*) | 
| 105 | fun list_of xq = | |
| 106 | (case pull xq of | |
| 15531 | 107 | NONE => [] | 
| 108 | | SOME (x, xq') => x :: list_of xq'); | |
| 5014 | 109 | |
| 110 | (*conversion from list to sequence*) | |
| 19475 | 111 | fun of_list xs = fold_rev cons xs empty; | 
| 5014 | 112 | |
| 113 | ||
| 17347 | 114 | (*sequence append: put the elements of xq in front of those of yq*) | 
| 19865 | 115 | fun append xq yq = | 
| 17347 | 116 | let | 
| 117 | fun copy s = | |
| 118 | make (fn () => | |
| 119 | (case pull s of | |
| 120 | NONE => pull yq | |
| 121 | | SOME (x, s') => SOME (x, copy s'))) | |
| 122 | in copy xq end; | |
| 5014 | 123 | |
| 124 | (*map over a sequence xq, append the sequence yq*) | |
| 125 | fun mapp f xq yq = | |
| 126 | let | |
| 127 | fun copy s = | |
| 128 | make (fn () => | |
| 129 | (case pull s of | |
| 15531 | 130 | NONE => pull yq | 
| 131 | | SOME (x, s') => SOME (f x, copy s'))) | |
| 5014 | 132 | in copy xq end; | 
| 133 | ||
| 17347 | 134 | (*interleave elements of xq with those of yq -- fairer than append*) | 
| 135 | fun interleave (xq, yq) = | |
| 136 | make (fn () => | |
| 137 | (case pull xq of | |
| 138 | NONE => pull yq | |
| 139 | | SOME (x, xq') => SOME (x, interleave (yq, xq')))); | |
| 5014 | 140 | |
| 141 | (*filter sequence by predicate*) | |
| 142 | fun filter pred xq = | |
| 143 | let | |
| 144 | fun copy s = | |
| 145 | make (fn () => | |
| 146 | (case pull s of | |
| 15531 | 147 | NONE => NONE | 
| 148 | | SOME (x, s') => if pred x then SOME (x, copy s') else pull (copy s'))); | |
| 5014 | 149 | in copy xq end; | 
| 150 | ||
| 151 | (*flatten a sequence of sequences to a single sequence*) | |
| 152 | fun flat xqq = | |
| 153 | make (fn () => | |
| 154 | (case pull xqq of | |
| 15531 | 155 | NONE => NONE | 
| 19865 | 156 | | SOME (xq, xqq') => pull (append xq (flat xqq')))); | 
| 5014 | 157 | |
| 17347 | 158 | (*map the function f over the sequence, making a new sequence*) | 
| 159 | fun map f xq = | |
| 5014 | 160 | make (fn () => | 
| 161 | (case pull xq of | |
| 17347 | 162 | NONE => NONE | 
| 163 | | SOME (x, xq') => SOME (f x, map f xq'))); | |
| 164 | ||
| 165 | fun maps f = flat o map f; | |
| 19484 | 166 | fun map_filter f = maps (fn x => (case f x of NONE => empty | SOME y => single y)); | 
| 17347 | 167 | |
| 168 | fun lift f xq y = map (fn x => f x y) xq; | |
| 169 | fun lifts f xq y = maps (fn x => f x y) xq; | |
| 5014 | 170 | |
| 19912 | 171 | fun singleton f x = f [x] |> map (fn [y] => y | _ => raise Empty); | 
| 172 | ||
| 25955 | 173 | (*wrapped lazy evaluation*) | 
| 174 | fun wrap f xq = | |
| 175 | make (fn () => | |
| 176 | (case f (fn () => pull xq) of | |
| 177 | NONE => NONE | |
| 178 | | SOME (x, xq') => SOME (x, wrap f xq'))); | |
| 19912 | 179 | |
| 16534 | 180 | (*print a sequence, up to "count" elements*) | 
| 181 | fun print print_elem count = | |
| 5014 | 182 | let | 
| 26887 | 183 | fun prnt (k: int) xq = | 
| 5014 | 184 | if k > count then () | 
| 185 | else | |
| 186 | (case pull xq of | |
| 15531 | 187 | NONE => () | 
| 16534 | 188 | | SOME (x, xq') => (print_elem k x; writeln ""; prnt (k + 1) xq')); | 
| 189 | in prnt 1 end; | |
| 5014 | 190 | |
| 191 | (*accumulating a function over a sequence; this is lazy*) | |
| 192 | fun it_right f (xq, yq) = | |
| 193 | let | |
| 194 | fun its s = | |
| 195 | make (fn () => | |
| 196 | (case pull s of | |
| 15531 | 197 | NONE => pull yq | 
| 198 | | SOME (a, s') => pull (f (a, its s')))) | |
| 5014 | 199 | in its xq end; | 
| 200 | ||
| 201 | ||
| 202 | ||
| 21270 | 203 | (** sequence functions **) (*cf. Pure/tctical.ML*) | 
| 5014 | 204 | |
| 205 | fun succeed x = single x; | |
| 206 | fun fail _ = empty; | |
| 207 | ||
| 21270 | 208 | fun op THEN (f, g) x = maps g (f x); | 
| 5014 | 209 | |
| 5864 | 210 | fun op ORELSE (f, g) x = | 
| 5014 | 211 | (case pull (f x) of | 
| 15531 | 212 | NONE => g x | 
| 5014 | 213 | | some => make (fn () => some)); | 
| 214 | ||
| 5864 | 215 | fun op APPEND (f, g) x = | 
| 19865 | 216 | append (f x) (make (fn () => pull (g x))); | 
| 5014 | 217 | |
| 23178 | 218 | fun EVERY fs = fold_rev (curry op THEN) fs succeed; | 
| 219 | fun FIRST fs = fold_rev (curry op ORELSE) fs fail; | |
| 5014 | 220 | |
| 221 | fun TRY f = ORELSE (f, succeed); | |
| 222 | ||
| 223 | fun REPEAT f = | |
| 224 | let | |
| 225 | fun rep qs x = | |
| 226 | (case pull (f x) of | |
| 15531 | 227 | NONE => SOME (x, make (fn () => repq qs)) | 
| 228 | | SOME (x', q) => rep (q :: qs) x') | |
| 229 | and repq [] = NONE | |
| 5014 | 230 | | repq (q :: qs) = | 
| 231 | (case pull q of | |
| 15531 | 232 | NONE => repq qs | 
| 233 | | SOME (x, q) => rep (q :: qs) x); | |
| 5014 | 234 | in fn x => make (fn () => rep [] x) end; | 
| 235 | ||
| 5558 | 236 | fun REPEAT1 f = THEN (f, REPEAT f); | 
| 237 | ||
| 26887 | 238 | fun INTERVAL f (i: int) j x = | 
| 8535 | 239 | if i > j then single x | 
| 240 | else op THEN (f j, INTERVAL f i (j - 1)) x; | |
| 241 | ||
| 12851 | 242 | fun DETERM f x = | 
| 243 | (case pull (f x) of | |
| 15531 | 244 | NONE => empty | 
| 19475 | 245 | | SOME (x', _) => cons x' empty); | 
| 8535 | 246 | |
| 5014 | 247 | end; |