author | paulson |
Tue, 20 May 1997 11:38:50 +0200 | |
changeset 3235 | 351565b7321b |
parent 2281 | e00c13a29eda |
permissions | -rw-r--r-- |
2281 | 1 |
(* Title: HOL/Integ/IntRingDefs.thy |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow and Markus Wenzel |
|
4 |
Copyright 1996 TU Muenchen |
|
5 |
||
6 |
Provides the basic defs and thms for showing that int is a commutative ring. |
|
7 |
Most of it has been defined and shown already. |
|
8 |
*) |
|
9 |
||
10 |
IntRingDefs = Integ + Ring + |
|
11 |
||
12 |
instance int :: zero |
|
13 |
defs zero_int_def "zero::int == $# 0" |
|
14 |
||
15 |
end |