| author | wenzelm | 
| Sat, 18 Mar 2000 19:10:02 +0100 | |
| changeset 8516 | f5f6a97ee43f | 
| parent 7588 | 26384af93359 | 
| child 8838 | 4eaa99f0d223 | 
| permissions | -rw-r--r-- | 
(* Title: Real/RealOrd.thy ID: $Id$ Author: Lawrence C. Paulson Jacques D. Fleuriot Copyright: 1998 University of Cambridge Description: Type "real" is a linear order *) RealOrd = RealDef + instance real :: order (real_le_refl,real_le_trans,real_le_anti_sym,real_less_le) instance real :: linorder (real_le_linear) constdefs rabs :: real => real "rabs r == if 0r<=r then r else -r" end