| author | clasohm |
| Tue, 07 Nov 1995 13:15:04 +0100 | |
| changeset 1323 | ae24fa249266 |
| parent 1274 | ea0668a1c0ba |
| child 1479 | 21eb5e156d91 |
| permissions | -rw-r--r-- |
| 1274 | 1 |
(* |
2 |
ID: $Id$ |
|
3 |
Author: Franz Regensburger |
|
4 |
Copyright 1993 Technische Universitaet Muenchen |
|
5 |
||
6 |
Additional constants for stream |
|
7 |
*) |
|
8 |
||
9 |
Stream2 = Stream + |
|
10 |
||
11 |
consts |
|
12 |
||
13 |
smap :: "('a -> 'b) -> 'a stream -> 'b stream"
|
|
14 |
||
15 |
defs |
|
16 |
||
17 |
smap_def |
|
18 |
"smap == fix`(LAM h f s. stream_when`(LAM x l.scons `(f`x) `(h`f`l)) `s)" |
|
19 |
||
20 |
||
21 |
end |
|
22 |
||
23 |
||
24 |
(* |
|
25 |
smap`f`UU = UU |
|
26 |
x~=UU --> smap`f`(scons`x`xs) = scons `(f`x) `(smap`f`xs) |
|
27 |
||
28 |
*) |
|
29 |