1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| A.Asperti, C.Sacerdoti Coen, *)
8 (* ||A|| E.Tassi, S.Zacchiroli *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU Lesser General Public License Version 2.1 *)
13 (**************************************************************************)
15 set "baseuri" "cic:/matita/nat/relevant_equations".
17 include "nat/times.ma".
18 include "nat/minus.ma".
20 (* if gcd is compiled before this, the applys will take too much *)
22 theorem times_plus_l: \forall n,m,p:nat. (n+m)*p = n*p + m*p.
24 apply (trans_eq ? ? (p*(n+m))).
26 apply (trans_eq ? ? (p*n+p*m)).
27 apply distr_times_plus.
33 theorem times_minus_l: \forall n,m,p:nat. (n-m)*p = n*p - m*p.
35 apply (trans_eq ? ? (p*(n-m))).
37 apply (trans_eq ? ? (p*n-p*m)).
38 apply distr_times_minus.
44 theorem times_plus_plus: \forall n,m,p,q:nat. (n + m)*(p + q) =
45 n*p + n*q + m*p + m*q.
47 apply (trans_eq nat ? ((n*(p+q) + m*(p+q)))).
49 rewrite > distr_times_plus.
50 rewrite > distr_times_plus.
51 rewrite < assoc_plus.reflexivity.
54 theorem eq_pred_to_eq:
55 ∀n,m. O < n → O < m → pred n = pred m → n = m.
57 generalize in match (eq_f ? ? S ? ? H2);
59 rewrite < S_pred in H3;
60 rewrite < S_pred in H3;