X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fdama%2Fordered_divisible_group.ma;h=15dd52cdb853ba822b7d35e4a1ab959b0bf6e311;hb=1ba7566bacd8d29e772646b3c86c7d5c944e9a6e;hp=cae552114ca54acb9984530350314d8edf873127;hpb=11f667afbb87725dd5e243d4b3717d19f584a481;p=helm.git diff --git a/helm/software/matita/dama/ordered_divisible_group.ma b/helm/software/matita/dama/ordered_divisible_group.ma index cae552114..15dd52cdb 100644 --- a/helm/software/matita/dama/ordered_divisible_group.ma +++ b/helm/software/matita/dama/ordered_divisible_group.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/ordered_divisible_group/". + include "nat/orders.ma". include "nat/times.ma". @@ -32,44 +32,44 @@ qed. coercion cic:/matita/ordered_divisible_group/todg_division.con. -lemma pow_ge: ∀G:todgroup.∀x:G.∀n.0 ≤ x → 0 ≤ n * x. +lemma mul_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+(n1*x)) (zero_neutral ??)); +apply (Le≪ (0+(n1*x)) (zero_neutral ??)); apply fle_plusr; assumption; qed. -lemma lt_ltpow: ∀G:todgroup.∀x,y:G.∀n. x < y → S n * x < S n * y. +lemma lt_ltmul: ∀G:todgroup.∀x,y:G.∀n. x < y → S n * x < S n * y. intros; elim n; [simplify; apply flt_plusr; assumption] simplify; apply (ltplus); [assumption] assumption; qed. -lemma ltpow_lt: ∀G:todgroup.∀x,y:G.∀n. S n * x < S n * y → x < y. +lemma ltmul_lt: ∀G:todgroup.∀x,y:G.∀n. S n * x < S n * y → x < y. intros 4; elim n; [apply (plus_cancr_lt ??? 0); assumption] simplify in l; cases (ltplus_orlt ????? l); [assumption] apply f; assumption; qed. lemma divide_preserves_lt: ∀G:todgroup.∀e:G.∀n.0 sym_plus; simplify; apply (lt_rewl ??? (0+(y+n*y))); [ + rewrite > sym_plus; simplify; apply (Lt≪ (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+(n1+n)*y)); [apply eq_sym; apply zero_neutral] +apply (Lt≪ (0+(n1+n)*y)); [apply eq_sym; apply zero_neutral] apply flt_plusr; assumption; qed.