| author | nipkow | 
| Sat, 18 Nov 1995 14:55:44 +0100 | |
| changeset 1344 | f172a7f14e49 | 
| parent 1168 | 74be52691d62 | 
| child 1479 | 21eb5e156d91 | 
| permissions | -rw-r--r-- | 
(* Title: HOLCF/cfun3.thy ID: $Id$ Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Class instance of -> for class pcpo *) Cfun3 = Cfun2 + arities "->" :: (pcpo,pcpo)pcpo (* Witness cfun2.ML *) consts Istrictify :: "('a->'b)=>'a=>'b" strictify :: "('a->'b)->'a->'b" rules inst_cfun_pcpo "(UU::'a->'b) = UU_cfun" defs Istrictify_def "Istrictify f x == if x=UU then UU else f`x" strictify_def "strictify == (LAM f x.Istrictify f x)" end