src/HOL/Integ/IntRingDefs.ML
author nipkow
Thu, 24 Apr 1997 17:59:55 +0200
changeset 3037 99ed078e6ae7
parent 2281 e00c13a29eda
child 4423 a129b817b58a
permissions -rw-r--r--
rename_params_rule used to check if the new name clashed with a free name in the whole goal state. Now checks only the subgoal concerned.

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

*)

open IntRingDefs;

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

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