| 2278 |      1 | (*  Title:      HOLCF/Up3.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Franz Regensburger
 | 
| 12030 |      4 |     License:    GPL (GNU GENERAL PUBLIC LICENSE)
 | 
| 2278 |      5 | 
 | 
|  |      6 | Class instance of  ('a)u for class pcpo
 | 
|  |      7 | *)
 | 
|  |      8 | 
 | 
|  |      9 | Up3 = Up2 +
 | 
|  |     10 | 
 | 
| 2640 |     11 | instance u :: (pcpo)pcpo      (least_up,cpo_up)
 | 
| 2278 |     12 | 
 | 
| 2640 |     13 | constdefs  
 | 
|  |     14 |         up  :: "'a -> ('a)u"
 | 
| 3842 |     15 |        "up  == (LAM x. Iup(x))"
 | 
| 2640 |     16 |         fup :: "('a->'c)-> ('a)u -> 'c"
 | 
| 3842 |     17 |        "fup == (LAM f p. Ifup(f)(p))"
 | 
| 2278 |     18 | 
 | 
|  |     19 | translations
 | 
| 10834 |     20 | "case l of up$x => t1" == "fup$(LAM x. t1)$l"
 | 
| 2278 |     21 | 
 | 
|  |     22 | end
 | 
|  |     23 | 
 | 
|  |     24 | 
 | 
|  |     25 | 
 |