changeset 23957 | 54fab60ddc97 |
parent 23350 | 50c5b0912a0c |
child 26202 | 51f8a696cd8d |
--- a/src/HOL/Algebra/Ring.thy Tue Jul 24 15:21:54 2007 +0200 +++ b/src/HOL/Algebra/Ring.thy Tue Jul 24 15:29:57 2007 +0200 @@ -471,8 +471,8 @@ with R show ?thesis by (simp add: a_assoc r_neg ) qed -lemma (in ring) minus_eq: - "[| x \<in> carrier R; y \<in> carrier R |] ==> x \<ominus> y = x \<oplus> \<ominus> y" +lemma (in abelian_group) minus_eq: + "[| x \<in> carrier G; y \<in> carrier G |] ==> x \<ominus> y = x \<oplus> \<ominus> y" by (simp only: a_minus_def) text {* Setup algebra method: