| author | paulson |
| Tue, 18 Mar 1997 10:42:08 +0100 | |
| changeset 2802 | f119c3686782 |
| parent 2640 | ee4dfce170a0 |
| child 2838 | 2e908f29bc3d |
| 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 + instance "->" :: (pcpo,pcpo)pcpo (least_cfun,cpo_cfun) consts Istrictify :: "('a->'b)=>'a=>'b" strictify :: "('a->'b)->'a->'b" defs Istrictify_def "Istrictify f x == if x=UU then UU else f`x" strictify_def "strictify == (LAM f x.Istrictify f x)" end