| author | wenzelm | 
| Mon, 09 Mar 1998 16:11:28 +0100 | |
| changeset 4700 | 20ade76722d6 | 
| parent 2841 | c2508f4ab739 | 
| child 12030 | 46d57d0290a2 | 
| 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/Discrete.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 | Discrete = Discrete1 + | 
| 
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 | instance discr :: (term)cpo (discr_cpo) | 
| 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 nipkow parents: diff
changeset | 12 | |
| 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 nipkow parents: diff
changeset | 13 | constdefs | 
| 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 nipkow parents: diff
changeset | 14 |    undiscr :: ('a::term)discr => 'a
 | 
| 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 nipkow parents: diff
changeset | 15 | "undiscr x == (case x of Discr y => y)" | 
| 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 nipkow parents: diff
changeset | 16 | |
| 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 nipkow parents: diff
changeset | 17 | end |