| author | wenzelm | 
| Mon, 06 Oct 1997 18:25:04 +0200 | |
| changeset 3778 | b70c41bc7491 | 
| 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/Discrete1.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  | 
Discrete1 = Discrete0 +  | 
| 
 
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)po  | 
| 
 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 
nipkow 
parents:  
diff
changeset
 | 
12  | 
(less_discr_refl,less_discr_trans,less_discr_antisym)  | 
| 
 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 
nipkow 
parents:  
diff
changeset
 | 
13  | 
|
| 
 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 
nipkow 
parents:  
diff
changeset
 | 
14  | 
end  |