author | paulson |
Wed, 05 Aug 1998 10:57:25 +0200 | |
changeset 5253 | 82a5ca6290aa |
parent 5078 | 7b5ea59c0275 |
child 5491 | 22f8331cdf47 |
permissions | -rw-r--r-- |
(* Title: HOL/Integ/IntRingDefs.thy ID: $Id$ Author: Tobias Nipkow and Markus Wenzel Copyright 1996 TU Muenchen Provides the basic defs and thms for showing that int is a commutative ring. Most of it has been defined and shown already. *) IntRingDefs = Integ + Ring + instance int :: zero defs zero_int_def "zero::int == $# 0" end