X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fdama%2Fgroups.ma;h=f2b4bd01cfc47198e66f7a52130c38410778b7c9;hb=29ba55bc6e228a2f5021b4131075ea3a3ba23458;hp=0df357af57a4a52084a41ad238c2d62fc9c43859;hpb=449e9430dd9857b05443ef133de25fe6455075c0;p=helm.git diff --git a/matita/dama/groups.ma b/matita/dama/groups.ma index 0df357af5..f2b4bd01c 100644 --- a/matita/dama/groups.ma +++ b/matita/dama/groups.ma @@ -182,26 +182,26 @@ intros (G x y z H1 H2); apply eq_sym; apply (feq_oppr ??y); [2:apply eq_sym] assumption; qed. -lemma eq_opp: ∀G:abelian_group.∀x,y:G. x ≈ y → -x ≈ -y. +lemma feq_opp: ∀G:abelian_group.∀x,y:G. x ≈ y → -x ≈ -y. intros (G x y H); apply (feq_oppl ??y ? H); apply eq_reflexive; qed. -coercion cic:/matita/groups/eq_opp.con nocomposites. +coercion cic:/matita/groups/feq_opp.con nocomposites. lemma eq_opp_sym: ∀G:abelian_group.∀x,y:G. y ≈ x → -x ≈ -y. -compose eq_opp with eq_sym (H); apply H; assumption; +compose feq_opp with eq_sym (H); apply H; assumption; qed. coercion cic:/matita/groups/eq_opp_sym.con nocomposites. lemma eq_opp_plusr: ∀G:abelian_group.∀x,y,z:G. x ≈ y → -(x + z) ≈ -(y + z). -compose feq_plusr with eq_opp(H); apply H; assumption; +compose feq_plusr with feq_opp(H); apply H; assumption; qed. coercion cic:/matita/groups/eq_opp_plusr.con nocomposites. lemma eq_opp_plusl: ∀G:abelian_group.∀x,y,z:G. x ≈ y → -(z + x) ≈ -(z + y). -compose feq_plusl with eq_opp(H); apply H; assumption; +compose feq_plusl with feq_opp(H); apply H; assumption; qed. coercion cic:/matita/groups/eq_opp_plusl.con nocomposites.