| author | paulson | 
| Sun, 15 Feb 2004 10:46:37 +0100 | |
| changeset 14387 | e96d5c42c4b0 | 
| parent 12030 | 46d57d0290a2 | 
| child 14981 | e73f8140af78 | 
| permissions | -rw-r--r-- | 
(* Title: HOLCF/Up3.thy ID: $Id$ Author: Franz Regensburger License: GPL (GNU GENERAL PUBLIC LICENSE) Class instance of ('a)u for class pcpo *) Up3 = Up2 + instance u :: (pcpo)pcpo (least_up,cpo_up) constdefs up :: "'a -> ('a)u" "up == (LAM x. Iup(x))" fup :: "('a->'c)-> ('a)u -> 'c" "fup == (LAM f p. Ifup(f)(p))" translations "case l of up$x => t1" == "fup$(LAM x. t1)$l" end