X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fdama%2Fgroups.ma;h=699bd73fc3286190a3dbcf9ee3c643b0d8b9ce38;hb=f497a1884ebadb42e63018e16a29337bc5166421;hp=6d832c208272231be0be1276410fddf5e67c4c88;hpb=348acb421355c345a9af0754c1c16508a43eeea5;p=helm.git diff --git a/matita/dama/groups.ma b/matita/dama/groups.ma index 6d832c208..699bd73fc 100644 --- a/matita/dama/groups.ma +++ b/matita/dama/groups.ma @@ -103,3 +103,38 @@ rewrite > zero_neutral in H1; assumption. qed. +theorem eq_opp_plus_plus_opp_opp: ∀G:abelian_group.∀x,y:G. -(x+y) = -x + -y. + intros; + apply (cancellationlaw ? (x+y)); + rewrite < plus_comm; + rewrite > opp_inverse; + rewrite > plus_assoc; + rewrite > plus_comm in ⊢ (? ? ? (? ? ? (? ? ? %))); + rewrite < plus_assoc in ⊢ (? ? ? (? ? ? %)); + rewrite > plus_comm; + rewrite > plus_comm in ⊢ (? ? ? (? ? (? ? % ?) ?)); + rewrite > opp_inverse; + rewrite > zero_neutral; + rewrite > opp_inverse; + reflexivity. +qed. + +theorem eq_opp_opp_x_x: ∀G:abelian_group.∀x:G.--x=x. + intros; + apply (cancellationlaw ? (-x)); + rewrite > opp_inverse; + rewrite > plus_comm; + rewrite > opp_inverse; + reflexivity. +qed. + +theorem eq_zero_opp_zero: ∀G:abelian_group.0=-0. + [ assumption + | intros; + apply (cancellationlaw ? 0); + rewrite < plus_comm in ⊢ (? ? ? %); + rewrite > opp_inverse; + rewrite > zero_neutral; + reflexivity + ]. +qed. \ No newline at end of file