X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fdama%2Fgroups.ma;h=1b39c1518348e7431de23d10ee1d2e6fc3dcf86f;hb=894c3e4038c93b484896e81132eae55046e47605;hp=8353ea67649b26c14f837451568fd11f47e581f5;hpb=85fce74d6a0cd8e83e59895e81f1bc09d21ef4c2;p=helm.git diff --git a/helm/software/matita/dama/groups.ma b/helm/software/matita/dama/groups.ma index 8353ea676..1b39c1518 100644 --- a/helm/software/matita/dama/groups.ma +++ b/helm/software/matita/dama/groups.ma @@ -17,6 +17,7 @@ set "baseuri" "cic:/matita/groups/". include "higher_order_defs/functions.ma". include "nat/nat.ma". include "nat/orders.ma". +include "constructive_connectives.ma". definition left_neutral \def λC,op.λe:C. ∀x:C. op e x = x. @@ -102,3 +103,27 @@ 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. \ No newline at end of file