| author | oheimb | 
| Thu, 31 May 2001 16:52:20 +0200 | |
| changeset 11345 | cd605c85e421 | 
| parent 8936 | a1c426541757 | 
| child 11701 | 3d51fbf81c17 | 
| permissions | -rw-r--r-- | 
| 5078 | 1 | (* Title: HOL/Integ/IntRing.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Tobias Nipkow and Markus Wenzel | |
| 4 | Copyright 1996 TU Muenchen | |
| 5 | ||
| 6 | The integers form a commutative ring. | |
| 7 | With an application of Lagrange's lemma. | |
| 8 | *) | |
| 9 | ||
| 8936 
a1c426541757
Now that 0 is overloaded, constant "zero" and its type class "zero" are
 paulson parents: 
6921diff
changeset | 10 | IntRing = Ring + Lagrange + | 
| 5078 | 11 | |
| 12 | instance int :: add_semigroup (zadd_assoc) | |
| 8936 
a1c426541757
Now that 0 is overloaded, constant "zero" and its type class "zero" are
 paulson parents: 
6921diff
changeset | 13 | instance int :: add_monoid (IntDef.Zero_def,zadd_int0,zadd_int0_right) | 
| 
a1c426541757
Now that 0 is overloaded, constant "zero" and its type class "zero" are
 paulson parents: 
6921diff
changeset | 14 | instance int :: add_group {|Auto_tac|}
 | 
| 5078 | 15 | instance int :: add_agroup (zadd_commute) | 
| 16 | instance int :: ring (zmult_assoc,zadd_zmult_distrib2,zadd_zmult_distrib) | |
| 17 | instance int :: cring (zmult_commute) | |
| 18 | ||
| 19 | end |