]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/library/Z/times.ma
Preparing for 0.5.9 release.
[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 interpretation "integer times" 'times x y = (Ztimes x y).
34
35 theorem Ztimes_z_OZ:  \forall z:Z. z*OZ = OZ.
36 intro.elim z.
37 simplify.reflexivity.
38 simplify.reflexivity.
39 simplify.reflexivity.
40 qed.
41
42 definition Zone \def pos O.
43
44 theorem Ztimes_neg_Zopp: \forall n:nat.\forall x:Z.
45 neg n * x = - (pos n * x).
46 intros.elim x.
47 simplify.reflexivity.
48 simplify.reflexivity.
49 simplify.reflexivity.
50 qed.
51
52 theorem symmetric_Ztimes : symmetric Z Ztimes.
53 change with (\forall x,y:Z. x*y = y*x).
54 intros.elim x.rewrite > Ztimes_z_OZ.reflexivity.
55 elim y.simplify.reflexivity. 
56 change with (pos (pred ((S n) * (S n1))) = pos (pred ((S n1) * (S n)))).
57 rewrite < sym_times.reflexivity.
58 change with (neg (pred ((S n) * (S n1))) = neg (pred ((S n1) * (S n)))).
59 rewrite < sym_times.reflexivity.
60 elim y.simplify.reflexivity.
61 change with (neg (pred ((S n) * (S n1))) = neg (pred ((S n1) * (S n)))).
62 rewrite < sym_times.reflexivity.
63 change with (pos (pred ((S n) * (S n1))) = pos (pred ((S n1) * (S n)))).
64 rewrite < sym_times.reflexivity.
65 qed.
66
67 variant sym_Ztimes : \forall x,y:Z. x*y = y*x
68 \def symmetric_Ztimes.
69
70 theorem Ztimes_Zone_l: \forall z:Z. Ztimes Zone z = z.
71 intro.unfold Zone.simplify.
72 elim z;simplify
73   [reflexivity
74   |rewrite < plus_n_O.reflexivity
75   |rewrite < plus_n_O.reflexivity
76   ]
77 qed.
78
79 theorem Ztimes_Zone_r: \forall z:Z. Ztimes z Zone = z.
80 intro.
81 rewrite < sym_Ztimes.
82 apply Ztimes_Zone_l.
83 qed.
84
85 theorem associative_Ztimes: associative Z Ztimes.
86 unfold associative.
87 intros.elim x.
88   simplify.reflexivity. 
89   elim y.
90     simplify.reflexivity.
91     elim z.
92       simplify.reflexivity.
93       change with 
94        (pos (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
95        pos (pred ((S n) * (S (pred ((S n1) * (S n2))))))).
96         rewrite < S_pred.rewrite < S_pred.rewrite < assoc_times.reflexivity.
97          apply lt_O_times_S_S.apply lt_O_times_S_S.
98       change with 
99        (neg (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
100        neg (pred ((S n) * (S (pred ((S n1) * (S n2))))))).
101         rewrite < S_pred.rewrite < S_pred.rewrite < assoc_times.reflexivity.
102          apply lt_O_times_S_S.apply lt_O_times_S_S.
103     elim z.
104       simplify.reflexivity.
105       change with 
106        (neg (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
107        neg (pred ((S n) * (S (pred ((S n1) * (S n2))))))).
108         rewrite < S_pred.rewrite < S_pred.rewrite < assoc_times.reflexivity.
109          apply lt_O_times_S_S.apply lt_O_times_S_S.
110       change with 
111        (pos (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
112        pos(pred ((S n) * (S (pred ((S n1) * (S n2))))))).
113         rewrite < S_pred.rewrite < S_pred.rewrite < assoc_times.reflexivity.
114         apply lt_O_times_S_S.apply lt_O_times_S_S.
115   elim y.
116     simplify.reflexivity.
117     elim z.
118       simplify.reflexivity.
119       change with 
120        (neg (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
121        neg (pred ((S n) * (S (pred ((S n1) * (S n2))))))).
122         rewrite < S_pred.rewrite < S_pred.rewrite < assoc_times.reflexivity.
123          apply lt_O_times_S_S.apply lt_O_times_S_S.
124       change with 
125        (pos (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
126        pos (pred ((S n) * (S (pred ((S n1) * (S n2))))))).
127         rewrite < S_pred.rewrite < S_pred.rewrite < assoc_times.reflexivity.
128          apply lt_O_times_S_S.apply lt_O_times_S_S.
129     elim z.
130       simplify.reflexivity.
131       change with 
132        (pos (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
133        pos (pred ((S n) * (S (pred ((S n1) * (S n2))))))).
134         rewrite < S_pred.rewrite < S_pred.rewrite < assoc_times.reflexivity.
135          apply lt_O_times_S_S.apply lt_O_times_S_S.
136       change with 
137        (neg (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
138        neg(pred ((S n) * (S (pred ((S n1) * (S n2))))))).
139         rewrite < S_pred.rewrite < S_pred.rewrite < assoc_times.reflexivity.
140          apply lt_O_times_S_S.apply lt_O_times_S_S.
141 qed.
142
143 variant assoc_Ztimes : \forall x,y,z:Z. 
144 (x * y) * z = x * (y * z) \def 
145 associative_Ztimes.
146
147 lemma times_minus1: \forall n,p,q:nat. lt q p \to
148 (S n) * (S (pred ((S p) - (S q)))) =
149 pred ((S n) * (S p)) - pred ((S n) * (S q)).
150 intros.
151 rewrite < S_pred.  
152 rewrite > minus_pred_pred.
153 rewrite < distr_times_minus. 
154 reflexivity.
155 (* we now close all positivity conditions *)
156 apply lt_O_times_S_S.                    
157 apply lt_O_times_S_S.
158 simplify.unfold lt.
159 apply le_SO_minus.  exact H.
160 qed.
161
162 lemma Ztimes_Zplus_pos_neg_pos: \forall n,p,q:nat.
163 (pos n)*((neg p)+(pos q)) = (pos n)*(neg p)+ (pos n)*(pos q). 
164 intros.
165 simplify. 
166 change in match (p + n * (S p)) with (pred ((S n) * (S p))).
167 change in match (q + n * (S q)) with (pred ((S n) * (S q))).
168 rewrite < nat_compare_pred_pred.
169 rewrite < nat_compare_times_l.
170 rewrite < nat_compare_S_S.
171 apply (nat_compare_elim p q).
172 intro.
173 (* uff *)
174 change with (pos (pred ((S n) * (S (pred ((S q) - (S p)))))) =
175             pos (pred ((pred ((S n) * (S q))) - (pred ((S n) * (S p)))))).
176 rewrite < (times_minus1 n q p H).reflexivity.
177 intro.rewrite < H.simplify.reflexivity.
178 intro.
179 change with (neg (pred ((S n) * (S (pred ((S p) - (S q)))))) =
180             neg (pred ((pred ((S n) * (S p))) - (pred ((S n) * (S q)))))). 
181 rewrite < (times_minus1 n p q H).reflexivity.                                 
182 (* two more positivity conditions from nat_compare_pred_pred *)   
183 apply lt_O_times_S_S.  
184 apply lt_O_times_S_S. 
185 qed. 
186
187 lemma Ztimes_Zplus_pos_pos_neg: \forall n,p,q:nat.
188 (pos n)*((pos p)+(neg q)) = (pos n)*(pos p)+ (pos n)*(neg q).
189 intros.
190 rewrite < sym_Zplus.
191 rewrite > Ztimes_Zplus_pos_neg_pos.
192 apply sym_Zplus.
193 qed.
194
195 lemma distributive2_Ztimes_pos_Zplus: 
196 distributive2 nat Z (\lambda n,z. (pos n) * z) Zplus.
197 change with (\forall n,y,z.
198 (pos n) * (y + z) = (pos n) * y + (pos n) * z).  
199 intros.elim y.
200   reflexivity.
201   elim z.
202     reflexivity.
203     change with
204      (pos (pred ((S n) * ((S n1) + (S n2)))) =
205      pos (pred ((S n) * (S n1) + (S n) * (S n2)))).
206       rewrite < distr_times_plus.reflexivity.
207     apply Ztimes_Zplus_pos_pos_neg.
208   elim z.
209     reflexivity.
210     apply Ztimes_Zplus_pos_neg_pos.
211     change with
212      (neg (pred ((S n) * ((S n1) + (S n2)))) =
213      neg (pred ((S n) * (S n1) + (S n) * (S n2)))).
214     rewrite < distr_times_plus.reflexivity.
215 qed.
216
217 variant distr_Ztimes_Zplus_pos: \forall n,y,z.
218 (pos n) * (y + z) = ((pos n) * y + (pos n) * z) \def
219 distributive2_Ztimes_pos_Zplus.
220
221 lemma distributive2_Ztimes_neg_Zplus : 
222 distributive2 nat Z (\lambda n,z. (neg n) * z) Zplus.
223 change with (\forall n,y,z.
224 (neg n) * (y + z) = (neg n) * y + (neg n) * z).  
225 intros.
226 rewrite > Ztimes_neg_Zopp. 
227 rewrite > distr_Ztimes_Zplus_pos.
228 rewrite > Zopp_Zplus.
229 rewrite < Ztimes_neg_Zopp. rewrite < Ztimes_neg_Zopp.
230 reflexivity.
231 qed.
232
233 variant distr_Ztimes_Zplus_neg: \forall n,y,z.
234 (neg n) * (y + z) = (neg n) * y + (neg n) * z \def
235 distributive2_Ztimes_neg_Zplus.
236
237 theorem distributive_Ztimes_Zplus: distributive Z Ztimes Zplus.
238 change with (\forall x,y,z:Z. x * (y + z) = x*y + x*z).
239 intros.elim x.
240 (* case x = OZ *)
241 simplify.reflexivity.
242 (* case x = pos n *)
243 apply distr_Ztimes_Zplus_pos.
244 (* case x = neg n *)
245 apply distr_Ztimes_Zplus_neg.
246 qed.
247
248 variant distr_Ztimes_Zplus: \forall x,y,z.
249 x * (y + z) = x*y + x*z \def
250 distributive_Ztimes_Zplus.