X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Farithmetics%2Fdiv_and_mod.ma;h=44b796e6b37f5107dd6370e129cf7839c7f09b9e;hb=09d38cb67e92bc6cdfe834cb1524a79643cc13e7;hp=6932a86e04d711cb617bb5715b9f56ff5bed5b33;hpb=ddc80515997a3f56085c6234d4db326141e189aa;p=helm.git diff --git a/matita/matita/lib/arithmetics/div_and_mod.ma b/matita/matita/lib/arithmetics/div_and_mod.ma index 6932a86e0..44b796e6b 100644 --- a/matita/matita/lib/arithmetics/div_and_mod.ma +++ b/matita/matita/lib/arithmetics/div_and_mod.ma @@ -178,7 +178,7 @@ theorem or_div_mod: ∀n,q. O < q → #n #q #posq (elim (le_to_or_lt_eq ?? (lt_mod_m_m n q posq))) #H [%2 % // applyS eq_f // - |%1 % // /demod/ (div_mod n m) {⊢ (? % ?)} >commutative_plus +>(div_mod n m) in ⊢ (? % ?); >commutative_plus @monotonic_lt_plus_l @lt_mod_m_m // qed. theorem le_div: ∀n,m. O < n → m/n ≤ m. #n #m #posn ->(div_mod m n) {⊢ (? ? %)} @(transitive_le ? (m/n*n)) /2/ +>(div_mod m n) in ⊢ (? ? %); @(transitive_le ? (m/n*n)) /2/ qed. theorem le_plus_mod: ∀m,n,q. O < q → @@ -228,9 +228,9 @@ theorem le_plus_mod: ∀m,n,q. O < q → @(div_mod_spec_to_eq2 … (m/q + n/q) ? (div_mod_spec_div_mod … posq)). @div_mod_spec_intro [@not_le_to_lt // - |>(div_mod n q) {⊢ (? ? (? ? %) ?)} + |>(div_mod n q) in ⊢ (? ? (? ? %) ?); (applyS (eq_f … (λx.plus x (n \mod q)))) - >(div_mod m q) {⊢ (? ? (? % ?) ?)} + >(div_mod m q) in ⊢ (? ? (? % ?) ?); (applyS (eq_f … (λx.plus x (m \mod q)))) // ] ] @@ -241,10 +241,10 @@ theorem le_plus_div: ∀m,n,q. O < q → #m #n #q #posq @(le_times_to_le … posq) @(le_plus_to_le_r ((m+n) \mod q)) (* bruttino *) ->commutative_times {⊢ (? ? %)} (div_mod m q) {⊢ (? ? (? % ?))} >(div_mod n q) {⊢ (? ? (? ? %))} ->commutative_plus {⊢ (? ? (? % ?))} >associative_plus {⊢ (? ? %)} -commutative_times in ⊢ (? ? %); (div_mod m q) in ⊢ (? ? (? % ?)); >(div_mod n q) in ⊢ (? ? (? ? %)); +>commutative_plus in ⊢ (? ? (? % ?)); >associative_plus in ⊢ (? ? %); +