author | wenzelm |
Mon, 22 Oct 2001 17:58:26 +0200 | |
changeset 11880 | a625de9ad62a |
parent 5192 | 704dd3a6d47d |
permissions | -rw-r--r-- |
2357 | 1 |
(* Title: HOLCF/Lift1.thy |
2 |
ID: $Id$ |
|
3033 | 3 |
Author: Olaf Mueller |
2357 | 4 |
Copyright 1996 Technische Universitaet Muenchen |
5 |
||
6 |
Lifting types of class term to flat pcpo's |
|
7 |
*) |
|
8 |
||
5192 | 9 |
Lift1 = Cprod3 + Datatype + |
2356
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 |
|
2394 | 13 |
datatype 'a lift = Undef | Def 'a |
2356
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
14 |
|
3323
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents:
3033
diff
changeset
|
15 |
instance lift :: (term)sq_ord |
2356
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
16 |
defs |
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
17 |
|
3323
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents:
3033
diff
changeset
|
18 |
less_lift_def "x << y == (x=y | x=Undef)" |
2356
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
19 |
|
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
20 |
end |