| author | berghofe | 
| Thu, 10 Oct 2002 14:19:17 +0200 | |
| changeset 13636 | fdf7e9388be7 | 
| parent 13588 | 07b66a557487 | 
| child 14264 | 3d0c6238162a | 
| permissions | -rw-r--r-- | 
| 13577 | 1 | (* Title: Integ/Int.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1998 University of Cambridge | |
| 5 | ||
| 6 | Type "int" is a linear order | |
| 7 | *) | |
| 8 | ||
| 9 | theory Int = IntDef | |
| 13588 
07b66a557487
Renamed Integ/int.ML to Integ/Int_lemmas.ML to prevent confusion with Int.ML
 paulson parents: 
13577diff
changeset | 10 | files ("Int_lemmas.ML"):
 | 
| 13577 | 11 | |
| 12 | instance int :: order | |
| 13 | proof qed (assumption | rule zle_refl zle_trans zle_anti_sym int_less_le)+ | |
| 14 | ||
| 15 | instance int :: plus_ac0 | |
| 16 | proof qed (rule zadd_commute zadd_assoc zadd_0)+ | |
| 17 | ||
| 18 | instance int :: linorder | |
| 19 | proof qed (rule zle_linear) | |
| 20 | ||
| 21 | constdefs | |
| 22 | nat :: "int => nat" | |
| 23 | "nat(Z) == if neg Z then 0 else (THE m. Z = int m)" | |
| 24 | ||
| 25 | defs (overloaded) | |
| 26 | zabs_def: "abs(i::int) == if i < 0 then -i else i" | |
| 27 | ||
| 13588 
07b66a557487
Renamed Integ/int.ML to Integ/Int_lemmas.ML to prevent confusion with Int.ML
 paulson parents: 
13577diff
changeset | 28 | use "Int_lemmas.ML" | 
| 13577 | 29 | |
| 30 | end |