| author | paulson | 
| Mon, 10 Feb 1997 12:34:54 +0100 | |
| changeset 2602 | 5ac837d98a85 | 
| parent 2357 | dd2e5e655fd2 | 
| child 2640 | ee4dfce170a0 | 
| permissions | -rw-r--r-- | 
| 2357 | 1 | (* Title: HOLCF/Lift3.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Olaf Mueller, Robert Sandner | |
| 4 | Copyright 1996 Technische Universitaet Muenchen | |
| 5 | ||
| 6 | Class Instance lift::(term)pcpo | |
| 7 | *) | |
| 8 | ||
| 2356 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 9 | Lift3 = Lift2 + | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 10 | |
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 11 | default term | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 12 | |
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 13 | arities | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 14 | "lift" :: (term)pcpo | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 15 | |
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 16 | consts | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 17 |  flift1      :: "('a => 'b::pcpo) => ('a lift -> 'b)"
 | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 18 | blift :: "bool => tr" | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 19 |  plift        :: "('a => bool) => 'a lift -> tr"   
 | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 20 |  flift2      :: "('a => 'b) => ('a lift -> 'b lift)"
 | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 21 | |
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 22 | translations | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 23 | "UU" <= "Undef" | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 24 | |
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 25 | defs | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 26 | blift_def | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 27 | "blift b == (if b then TT else FF)" | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 28 | |
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 29 | flift1_def | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 30 | "flift1 f == (LAM x. (case x of | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 31 | Undef => UU | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 32 | | Def y => (f y)))" | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 33 | |
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 34 | flift2_def | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 35 | "flift2 f == (LAM x. (case x of | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 36 | Undef => Undef | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 37 | | Def y => Def (f y)))" | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 38 | |
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 39 | plift_def | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 40 | "plift p == (LAM x. flift1 (%a. blift (p a))`x)" | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 41 | |
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 42 | |
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 43 | rules | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 44 | inst_lift_pcpo "(UU::'a lift) = Undef" | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 45 | |
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 46 | end | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 47 |