X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fdama%2Fordered_divisible_group.ma;h=cae552114ca54acb9984530350314d8edf873127;hb=9689ab3488bfd89eed1fd188c0f279a11737a20a;hp=805cce9568ad029a230fc4efc79e6081cec6115f;hpb=855b41852b429f211c579b3a4b2095d370a3c983;p=helm.git diff --git a/helm/software/matita/dama/ordered_divisible_group.ma b/helm/software/matita/dama/ordered_divisible_group.ma index 805cce956..cae552114 100644 --- a/helm/software/matita/dama/ordered_divisible_group.ma +++ b/helm/software/matita/dama/ordered_divisible_group.ma @@ -32,123 +32,44 @@ qed. coercion cic:/matita/ordered_divisible_group/todg_division.con. -lemma pow_lt: ∀G:todgroup.∀x:G.∀n.0 < x → 0 < x + pow ? x n. -intros (G x n H); elim n; [ - simplify; apply (lt_rewr ???? (plus_comm ???)); - apply (lt_rewr ???x (zero_neutral ??)); assumption] -simplify; apply (lt_transitive ?? (x+(x)\sup(n1))); [assumption] -apply flt_plusl; apply (lt_rewr ???? (plus_comm ???)); -apply (lt_rewl ??? (0 + (x \sup n1)) (eq_sym ??? (zero_neutral ??))); -apply (lt_rewl ???? (plus_comm ???)); -apply flt_plusl; assumption; -qed. - -lemma pow_ge: ∀G:todgroup.∀x:G.∀n.0 ≤ x → 0 ≤ pow ? x n. +lemma pow_ge: ∀G:todgroup.∀x:G.∀n.0 ≤ x → 0 ≤ n * x. intros (G x n); elim n; simplify; [apply le_reflexive] apply (le_transitive ???? H1); -apply (le_rewl ??? (0+(x\sup n1)) (zero_neutral ??)); +apply (le_rewl ??? (0+(n1*x)) (zero_neutral ??)); apply fle_plusr; assumption; qed. -lemma gt_pow: ∀G:todgroup.∀x:G.∀n.0 < pow ? x n → 0 < x. -intros 3; elim n; [ - simplify in l; cases (lt_coreflexive ?? l);] -simplify in l; -cut (0+0 sym_plus; simplify; apply (lt_rewl ??? (0+(y+y\sup n))); [ + rewrite > sym_plus; simplify; apply (lt_rewl ??? (0+(y+n*y))); [ apply eq_sym; apply zero_neutral] apply flt_plusr; assumption;] apply (lt_transitive ???? l); rewrite > sym_plus; simplify; rewrite > (sym_plus n); simplify; repeat apply flt_plusl; -apply (lt_rewl ???(0+y\sup(n1+n))); [apply eq_sym; apply zero_neutral] +apply (lt_rewl ???(0+(n1+n)*y)); [apply eq_sym; apply zero_neutral] apply flt_plusr; assumption; qed. -alias num (instance 0) = "natural number". -lemma core1: ∀G:todgroup.∀e:G.0