src/HOLCF/Lift1.thy
author mueller
Thu, 24 Apr 1997 17:38:33 +0200
changeset 3033 50e14d6d894f
parent 2640 ee4dfce170a0
child 3323 194ae2e0c193
permissions -rw-r--r--
only in comments
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2357
dd2e5e655fd2 Headers added
sandnerr
parents: 2356
diff changeset
     1
(*  Title:      HOLCF/Lift1.thy
dd2e5e655fd2 Headers added
sandnerr
parents: 2356
diff changeset
     2
    ID:         $Id$
3033
50e14d6d894f only in comments
mueller
parents: 2640
diff changeset
     3
    Author:     Olaf Mueller
2357
dd2e5e655fd2 Headers added
sandnerr
parents: 2356
diff changeset
     4
    Copyright   1996 Technische Universitaet Muenchen
dd2e5e655fd2 Headers added
sandnerr
parents: 2356
diff changeset
     5
dd2e5e655fd2 Headers added
sandnerr
parents: 2356
diff changeset
     6
Lifting types of class term to flat pcpo's
dd2e5e655fd2 Headers added
sandnerr
parents: 2356
diff changeset
     7
*)
dd2e5e655fd2 Headers added
sandnerr
parents: 2356
diff changeset
     8
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2394
diff changeset
     9
Lift1 = ccc1 + 
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
91d8abf108be adaptions for symbol font
oheimb
parents: 2357
diff changeset
    13
datatype 'a lift = Undef | Def 'a
2356
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    14
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    15
defs 
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    16
 
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2394
diff changeset
    17
 less_lift_def  "less x y == (x=y | x=Undef)"
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
end