author | wenzelm |
Mon, 22 Oct 2001 18:01:15 +0200 | |
changeset 11884 | 341b1fbc6130 |
parent 10834 | a7897aebbffc |
permissions | -rw-r--r-- |
2357 | 1 |
(* Title: HOLCF/Lift3.thy |
2 |
ID: $Id$ |
|
3034 | 3 |
Author: Olaf Mueller |
2357 | 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 |
|
2640 | 11 |
instance lift :: (term)pcpo (cpo_lift,least_lift) |
2356
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 |
consts |
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
14 |
flift1 :: "('a => 'b::pcpo) => ('a lift -> 'b)" |
3034 | 15 |
flift2 :: "('a => 'b) => ('a lift -> 'b lift)" |
16 |
||
17 |
liftpair ::"'a::term lift * 'b::term lift => ('a * 'b) lift" |
|
2356
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
18 |
|
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
19 |
translations |
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
20 |
"UU" <= "Undef" |
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 |
defs |
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
23 |
flift1_def |
3034 | 24 |
"flift1 f == (LAM x. (case x of |
2356
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
25 |
Undef => UU |
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
26 |
| Def y => (f y)))" |
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
27 |
flift2_def |
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
28 |
"flift2 f == (LAM x. (case x of |
3034 | 29 |
Undef => UU |
30 |
| Def y => Def (f y)))" |
|
31 |
liftpair_def |
|
10834 | 32 |
"liftpair x == (case (cfst$x) of |
3034 | 33 |
Undef => UU |
10834 | 34 |
| Def x1 => (case (csnd$x) of |
3034 | 35 |
Undef => UU |
36 |
| Def x2 => Def (x1,x2)))" |
|
2356
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
37 |
|
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
38 |
end |
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
39 |