author | nipkow |
Wed, 19 Jan 1994 17:35:01 +0100 | |
changeset 243 | c22b85994e17 |
child 297 | 5ef75ff3baeb |
permissions | -rw-r--r-- |
(* Title: HOLCF/stream2.thy ID: $Id$ Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Additional constants for stream *) Stream2 = Stream + consts smap :: "('a -> 'b) -> 'a stream -> 'b stream" rules smap_def "smap = fix[LAM h f s. stream_when[LAM x l.scons[f[x]][h[f][l]]][s]]" end (* smap[f][UU] = UU x~=UU --> smap[f][scons[x][xs]] = scons[f[x]][smap[f][xs]] *)