New function read_cterms is a combination of read_def_cterm and
read_cterm. It reads and simultaneously type-checks a list of strings.
cterm_of no longer catches exception TYPE; instead it must be caught later on
and a message printed using Sign.exn_type_msg. More informative messages can
be printed this way.
(* 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]]
*)