| author | paulson | 
| Wed, 10 Jan 2001 11:05:13 +0100 | |
| changeset 10841 | 2fb8089ab6cd | 
| parent 9043 | ca761fe227d8 | 
| child 14265 | 95b42e69436c | 
| permissions | -rw-r--r-- | 
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 1 | (* Title: Real/RealOrd.thy | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 2 | ID: $Id$ | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 3 | Author: Jacques D. Fleuriot and Lawrence C. Paulson | 
| 7334 | 4 | Copyright: 1998 University of Cambridge | 
| 9043 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
9013diff
changeset | 5 | Description: Type "real" is a linear order and also | 
| 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
9013diff
changeset | 6 | satisfies plus_ac0: + is an AC-operator with zero | 
| 7334 | 7 | *) | 
| 8 | ||
| 9 | RealOrd = RealDef + | |
| 10 | ||
| 11 | instance real :: order (real_le_refl,real_le_trans,real_le_anti_sym,real_less_le) | |
| 12 | instance real :: linorder (real_le_linear) | |
| 13 | ||
| 9043 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
9013diff
changeset | 14 | instance real :: plus_ac0 (real_add_commute,real_add_assoc,real_add_zero_left) | 
| 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
9013diff
changeset | 15 | |
| 7334 | 16 | end |