| author | wenzelm | 
| Tue, 05 Aug 1997 16:43:54 +0200 | |
| changeset 3591 | 412c62dd43de | 
| parent 3323 | 194ae2e0c193 | 
| child 5192 | 704dd3a6d47d | 
| permissions | -rw-r--r-- | 
| 2841 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 nipkow parents: diff
changeset | 1 | (* Title: HOLCF/Discrete0.thy | 
| 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 nipkow parents: diff
changeset | 2 | ID: $Id$ | 
| 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 nipkow parents: diff
changeset | 3 | Author: Tobias Nipkow | 
| 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 nipkow parents: diff
changeset | 4 | Copyright 1997 TUM | 
| 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 nipkow parents: diff
changeset | 5 | |
| 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 nipkow parents: diff
changeset | 6 | Discrete CPOs | 
| 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 nipkow parents: diff
changeset | 7 | *) | 
| 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 nipkow parents: diff
changeset | 8 | |
| 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 nipkow parents: diff
changeset | 9 | Discrete0 = Cont + | 
| 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 nipkow parents: diff
changeset | 10 | |
| 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 nipkow parents: diff
changeset | 11 | datatype 'a discr = Discr 'a | 
| 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 nipkow parents: diff
changeset | 12 | |
| 3323 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
2841diff
changeset | 13 | instance discr :: (term)sq_ord | 
| 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
2841diff
changeset | 14 | |
| 2841 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 nipkow parents: diff
changeset | 15 | defs | 
| 3323 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
2841diff
changeset | 16 | less_discr_def "((op <<)::('a::term)discr=>'a discr=>bool)  ==  op ="
 | 
| 2841 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 nipkow parents: diff
changeset | 17 | |
| 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 nipkow parents: diff
changeset | 18 | end |