author | wenzelm |
Fri, 04 Aug 2000 22:56:31 +0200 | |
changeset 9528 | 95852b4be214 |
parent 3033 | 50e14d6d894f |
permissions | -rw-r--r-- |
2357 | 1 |
(* Title: HOLCF/Lift2.thy |
2 |
ID: $Id$ |
|
3033 | 3 |
Author: Olaf Mueller |
2357 | 4 |
Copyright 1996 Technische Universitaet Muenchen |
5 |
||
6 |
Class Instance lift::(term)po |
|
7 |
*) |
|
8 |
||
2356
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
9 |
Lift2 = Lift1 + |
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
10 |
|
2640 | 11 |
instance "lift" :: (term)po (refl_less_lift,antisym_less_lift,trans_less_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 |
end |
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
14 |
|
2640 | 15 |