src/HOL/ex/IntRingDefs.ML
author paulson
Thu, 25 Jun 1998 13:57:34 +0200
changeset 5078 7b5ea59c0275
child 5227 e5a6ace920a0
permissions -rw-r--r--
Installation of target HOL-Real

(*  Title:      HOL/Integ/IntRingDefs.thy
    ID:         $Id$
    Author:     Tobias Nipkow and Markus Wenzel
    Copyright   1996 TU Muenchen

*)

open IntRingDefs;

Goalw [zero_int_def,zdiff_def] "(zero-x)+(x::int) = zero";
by (Simp_tac 1);
qed "left_inv_int";

Goalw [zero_int_def,zdiff_def] "x-y = (x::int) + (zero-y)";
by (Simp_tac 1);
qed "minus_inv_int";