]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/library/Z/times.ma
979cc87aa37f8ddbebe67e2273cc42a48a894f76
[helm.git] / helm / software / matita / library / Z / times.ma
1 (**************************************************************************)
2 (*       __                                                               *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||       A.Asperti, C.Sacerdoti Coen,                          *)
8 (*      ||A||       E.Tassi, S.Zacchiroli                                 *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU Lesser General Public License Version 2.1         *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "nat/lt_arith.ma".
16 include "Z/plus.ma".
17
18 definition Ztimes :Z \to Z \to Z \def
19 \lambda x,y.
20   match x with
21     [ OZ \Rightarrow OZ
22     | (pos m) \Rightarrow
23         match y with
24          [ OZ \Rightarrow OZ
25          | (pos n) \Rightarrow (pos (pred ((S m) * (S n))))
26          | (neg n) \Rightarrow (neg (pred ((S m) * (S n))))]
27     | (neg m) \Rightarrow
28         match y with
29          [ OZ \Rightarrow OZ
30          | (pos n) \Rightarrow (neg (pred ((S m) * (S n))))
31          | (neg n) \Rightarrow (pos (pred ((S m) * (S n))))]].
32          
33 (*CSC: the URI must disappear: there is a bug now *)
34 interpretation "integer times" 'times x y = (cic:/matita/Z/times/Ztimes.con x y).
35
36 theorem Ztimes_z_OZ:  \forall z:Z. z*OZ = OZ.
37 intro.elim z.
38 simplify.reflexivity.
39 simplify.reflexivity.
40 simplify.reflexivity.
41 qed.
42
43 definition Zone \def pos O.
44
45 theorem Ztimes_neg_Zopp: \forall n:nat.\forall x:Z.
46 neg n * x = - (pos n * x).
47 intros.elim x.
48 simplify.reflexivity.
49 simplify.reflexivity.
50 simplify.reflexivity.
51 qed.
52
53 theorem symmetric_Ztimes : symmetric Z Ztimes.
54 change with (\forall x,y:Z. x*y = y*x).
55 intros.elim x.rewrite > Ztimes_z_OZ.reflexivity.
56 elim y.simplify.reflexivity. 
57 change with (pos (pred ((S n) * (S n1))) = pos (pred ((S n1) * (S n)))).
58 rewrite < sym_times.reflexivity.
59 change with (neg (pred ((S n) * (S n1))) = neg (pred ((S n1) * (S n)))).
60 rewrite < sym_times.reflexivity.
61 elim y.simplify.reflexivity.
62 change with (neg (pred ((S n) * (S n1))) = neg (pred ((S n1) * (S n)))).
63 rewrite < sym_times.reflexivity.
64 change with (pos (pred ((S n) * (S n1))) = pos (pred ((S n1) * (S n)))).
65 rewrite < sym_times.reflexivity.
66 qed.
67
68 variant sym_Ztimes : \forall x,y:Z. x*y = y*x
69 \def symmetric_Ztimes.
70
71 theorem Ztimes_Zone_l: \forall z:Z. Ztimes Zone z = z.
72 intro.unfold Zone.simplify.
73 elim z;simplify
74   [reflexivity
75   |rewrite < plus_n_O.reflexivity
76   |rewrite < plus_n_O.reflexivity
77   ]
78 qed.
79
80 theorem Ztimes_Zone_r: \forall z:Z. Ztimes z Zone = z.
81 intro.
82 rewrite < sym_Ztimes.
83 apply Ztimes_Zone_l.
84 qed.
85
86 theorem associative_Ztimes: associative Z Ztimes.
87 unfold associative.
88 intros.elim x.
89   simplify.reflexivity. 
90   elim y.
91     simplify.reflexivity.
92     elim z.
93       simplify.reflexivity.
94       change with 
95        (pos (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
96        pos (pred ((S n) * (S (pred ((S n1) * (S n2))))))).
97         rewrite < S_pred.rewrite < S_pred.rewrite < assoc_times.reflexivity.
98          apply lt_O_times_S_S.apply lt_O_times_S_S.
99       change with 
100        (neg (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
101        neg (pred ((S n) * (S (pred ((S n1) * (S n2))))))).
102         rewrite < S_pred.rewrite < S_pred.rewrite < assoc_times.reflexivity.
103          apply lt_O_times_S_S.apply lt_O_times_S_S.
104     elim z.
105       simplify.reflexivity.
106       change with 
107        (neg (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
108        neg (pred ((S n) * (S (pred ((S n1) * (S n2))))))).
109         rewrite < S_pred.rewrite < S_pred.rewrite < assoc_times.reflexivity.
110          apply lt_O_times_S_S.apply lt_O_times_S_S.
111       change with 
112        (pos (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
113        pos(pred ((S n) * (S (pred ((S n1) * (S n2))))))).
114         rewrite < S_pred.rewrite < S_pred.rewrite < assoc_times.reflexivity.
115         apply lt_O_times_S_S.apply lt_O_times_S_S.
116   elim y.
117     simplify.reflexivity.
118     elim z.
119       simplify.reflexivity.
120       change with 
121        (neg (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
122        neg (pred ((S n) * (S (pred ((S n1) * (S n2))))))).
123         rewrite < S_pred.rewrite < S_pred.rewrite < assoc_times.reflexivity.
124          apply lt_O_times_S_S.apply lt_O_times_S_S.
125       change with 
126        (pos (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
127        pos (pred ((S n) * (S (pred ((S n1) * (S n2))))))).
128         rewrite < S_pred.rewrite < S_pred.rewrite < assoc_times.reflexivity.
129          apply lt_O_times_S_S.apply lt_O_times_S_S.
130     elim z.
131       simplify.reflexivity.
132       change with 
133        (pos (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
134        pos (pred ((S n) * (S (pred ((S n1) * (S n2))))))).
135         rewrite < S_pred.rewrite < S_pred.rewrite < assoc_times.reflexivity.
136          apply lt_O_times_S_S.apply lt_O_times_S_S.
137       change with 
138        (neg (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
139        neg(pred ((S n) * (S (pred ((S n1) * (S n2))))))).
140         rewrite < S_pred.rewrite < S_pred.rewrite < assoc_times.reflexivity.
141          apply lt_O_times_S_S.apply lt_O_times_S_S.
142 qed.
143
144 variant assoc_Ztimes : \forall x,y,z:Z. 
145 (x * y) * z = x * (y * z) \def 
146 associative_Ztimes.
147
148 lemma times_minus1: \forall n,p,q:nat. lt q p \to
149 (S n) * (S (pred ((S p) - (S q)))) =
150 pred ((S n) * (S p)) - pred ((S n) * (S q)).
151 intros.
152 rewrite < S_pred.  
153 rewrite > minus_pred_pred.
154 rewrite < distr_times_minus. 
155 reflexivity.
156 (* we now close all positivity conditions *)
157 apply lt_O_times_S_S.                    
158 apply lt_O_times_S_S.
159 simplify.unfold lt.
160 apply le_SO_minus.  exact H.
161 qed.
162
163 lemma Ztimes_Zplus_pos_neg_pos: \forall n,p,q:nat.
164 (pos n)*((neg p)+(pos q)) = (pos n)*(neg p)+ (pos n)*(pos q). 
165 intros.
166 simplify. 
167 change in match (p + n * (S p)) with (pred ((S n) * (S p))).
168 change in match (q + n * (S q)) with (pred ((S n) * (S q))).
169 rewrite < nat_compare_pred_pred.
170 rewrite < nat_compare_times_l.
171 rewrite < nat_compare_S_S.
172 apply (nat_compare_elim p q).
173 intro.
174 (* uff *)
175 change with (pos (pred ((S n) * (S (pred ((S q) - (S p)))))) =
176             pos (pred ((pred ((S n) * (S q))) - (pred ((S n) * (S p)))))).
177 rewrite < (times_minus1 n q p H).reflexivity.
178 intro.rewrite < H.simplify.reflexivity.
179 intro.
180 change with (neg (pred ((S n) * (S (pred ((S p) - (S q)))))) =
181             neg (pred ((pred ((S n) * (S p))) - (pred ((S n) * (S q)))))). 
182 rewrite < (times_minus1 n p q H).reflexivity.                                 
183 (* two more positivity conditions from nat_compare_pred_pred *)   
184 apply lt_O_times_S_S.  
185 apply lt_O_times_S_S. 
186 qed. 
187
188 lemma Ztimes_Zplus_pos_pos_neg: \forall n,p,q:nat.
189 (pos n)*((pos p)+(neg q)) = (pos n)*(pos p)+ (pos n)*(neg q).
190 intros.
191 rewrite < sym_Zplus.
192 rewrite > Ztimes_Zplus_pos_neg_pos.
193 apply sym_Zplus.
194 qed.
195
196 lemma distributive2_Ztimes_pos_Zplus: 
197 distributive2 nat Z (\lambda n,z. (pos n) * z) Zplus.
198 change with (\forall n,y,z.
199 (pos n) * (y + z) = (pos n) * y + (pos n) * z).  
200 intros.elim y.
201   reflexivity.
202   elim z.
203     reflexivity.
204     change with
205      (pos (pred ((S n) * ((S n1) + (S n2)))) =
206      pos (pred ((S n) * (S n1) + (S n) * (S n2)))).
207       rewrite < distr_times_plus.reflexivity.
208     apply Ztimes_Zplus_pos_pos_neg.
209   elim z.
210     reflexivity.
211     apply Ztimes_Zplus_pos_neg_pos.
212     change with
213      (neg (pred ((S n) * ((S n1) + (S n2)))) =
214      neg (pred ((S n) * (S n1) + (S n) * (S n2)))).
215     rewrite < distr_times_plus.reflexivity.
216 qed.
217
218 variant distr_Ztimes_Zplus_pos: \forall n,y,z.
219 (pos n) * (y + z) = ((pos n) * y + (pos n) * z) \def
220 distributive2_Ztimes_pos_Zplus.
221
222 lemma distributive2_Ztimes_neg_Zplus : 
223 distributive2 nat Z (\lambda n,z. (neg n) * z) Zplus.
224 change with (\forall n,y,z.
225 (neg n) * (y + z) = (neg n) * y + (neg n) * z).  
226 intros.
227 rewrite > Ztimes_neg_Zopp. 
228 rewrite > distr_Ztimes_Zplus_pos.
229 rewrite > Zopp_Zplus.
230 rewrite < Ztimes_neg_Zopp. rewrite < Ztimes_neg_Zopp.
231 reflexivity.
232 qed.
233
234 variant distr_Ztimes_Zplus_neg: \forall n,y,z.
235 (neg n) * (y + z) = (neg n) * y + (neg n) * z \def
236 distributive2_Ztimes_neg_Zplus.
237
238 theorem distributive_Ztimes_Zplus: distributive Z Ztimes Zplus.
239 change with (\forall x,y,z:Z. x * (y + z) = x*y + x*z).
240 intros.elim x.
241 (* case x = OZ *)
242 simplify.reflexivity.
243 (* case x = pos n *)
244 apply distr_Ztimes_Zplus_pos.
245 (* case x = neg n *)
246 apply distr_Ztimes_Zplus_neg.
247 qed.
248
249 variant distr_Ztimes_Zplus: \forall x,y,z.
250 x * (y + z) = x*y + x*z \def
251 distributive_Ztimes_Zplus.
252
253 theorem eq_plus_Zplus: \forall n,m:nat. Z_of_nat (n+m) =
254 Z_of_nat n + Z_of_nat m.
255 intro.cases n;intro
256   [reflexivity
257   |cases m
258     [simplify.rewrite < plus_n_O.reflexivity
259     |simplify.reflexivity.
260     ]]
261 qed.