| 11348 |      1 | (*  Title: 	HOLCF/ex/Stream.thy
 | 
| 9169 |      2 |     ID:         $Id$
 | 
| 2570 |      3 |     Author: 	Franz Regensburger, David von Oheimb
 | 
| 12036 |      4 |     License:    GPL (GNU GENERAL PUBLIC LICENSE)
 | 
| 2570 |      5 | 
 | 
| 12036 |      6 | General Stream domain.
 | 
| 2570 |      7 | *)
 | 
|  |      8 | 
 | 
| 11348 |      9 | Stream = HOLCF + Nat_Infinity +
 | 
| 2570 |     10 | 
 | 
| 4122 |     11 | domain 'a stream = "&&" (ft::'a) (lazy rt::'a stream) (infixr 65)
 | 
| 2570 |     12 | 
 | 
| 11348 |     13 | consts
 | 
|  |     14 | 
 | 
| 12036 |     15 |   smap		:: "('a -> 'b) -> 'a stream -> 'b stream"
 | 
|  |     16 |   sfilter	:: "('a -> tr) -> 'a stream -> 'a stream"
 | 
|  |     17 |   slen		:: "'a stream => inat"			("#_" [1000] 1000)
 | 
| 11348 |     18 | 
 | 
|  |     19 | defs
 | 
|  |     20 | 
 | 
| 12036 |     21 |   smap_def	"smap	 == fix\\<cdot>(\\<Lambda>h f s. case s of x && xs => f\\<cdot>x && h\\<cdot>f\\<cdot>xs)"
 | 
|  |     22 |   sfilter_def	"sfilter == fix\\<cdot>(\\<Lambda>h p s. case s of x && xs => 
 | 
| 11348 |     23 | 				     If p\\<cdot>x then x && h\\<cdot>p\\<cdot>xs else h\\<cdot>p\\<cdot>xs fi)"
 | 
| 12036 |     24 |   slen_def	"#s == if stream_finite s 
 | 
| 11348 |     25 | 		      then Fin (LEAST n. stream_take n\\<cdot>s = s) else \\<infinity>"
 | 
|  |     26 | 
 | 
| 2570 |     27 | end
 | 
|  |     28 | 
 | 
|  |     29 | 
 |