]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/ng_assembly/num/word16_lemmas.ma
Release 0.5.9.
[helm.git] / helm / software / matita / contribs / ng_assembly / num / word16_lemmas.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 (* ********************************************************************** *)
16 (*                          Progetto FreeScale                            *)
17 (*                                                                        *)
18 (*   Sviluppato da: Ing. Cosimo Oliboni, oliboni@cs.unibo.it              *)
19 (*   Ultima modifica: 05/08/2009                                          *)
20 (*                                                                        *)
21 (* ********************************************************************** *)
22
23 include "num/word16.ma".
24 include "num/byte8_lemmas.ma".
25
26 (* **** *)
27 (* WORD *)
28 (* **** *)
29
30 nlemma word16_destruct_1 :
31 ∀x1,x2,y1,y2.
32  mk_word16 x1 y1 = mk_word16 x2 y2 → x1 = x2.
33  #x1; #x2; #y1; #y2; #H;
34  nchange with (match mk_word16 x2 y2 with [ mk_word16 a _ ⇒ x1 = a ]);
35  nrewrite < H;
36  nnormalize;
37  napply refl_eq.
38 nqed.
39
40 nlemma word16_destruct_2 :
41 ∀x1,x2,y1,y2.
42  mk_word16 x1 y1 = mk_word16 x2 y2 → y1 = y2.
43  #x1; #x2; #y1; #y2; #H;
44  nchange with (match mk_word16 x2 y2 with [ mk_word16 _ b ⇒ y1 = b ]);
45  nrewrite < H;
46  nnormalize;
47  napply refl_eq.
48 nqed.
49
50 nlemma symmetric_eqw16 : symmetricT word16 bool eq_w16.
51  #w1; #w2;
52  nelim w1;
53  nelim w2;
54  #b1; #b2; #b3; #b4;
55  nchange with (((eq_b8 b3 b1)⊗(eq_b8 b4 b2)) = ((eq_b8 b1 b3)⊗(eq_b8 b2 b4)));
56  nrewrite > (symmetric_eqb8 b1 b3);
57  nrewrite > (symmetric_eqb8 b2 b4);
58  napply refl_eq.
59 nqed.
60
61 nlemma symmetric_andw16 : symmetricT word16 word16 and_w16.
62  #w1; #w2;
63  nelim w1;
64  nelim w2;
65  #b1; #b2; #b3; #b4;
66  nchange with ((mk_word16 (and_b8 b3 b1) (and_b8 b4 b2)) = (mk_word16 (and_b8 b1 b3) (and_b8 b2 b4)));
67  nrewrite > (symmetric_andb8 b1 b3);
68  nrewrite > (symmetric_andb8 b2 b4);
69  napply refl_eq.
70 nqed.
71
72 nlemma associative_andw16 : ∀w1,w2,w3.(and_w16 (and_w16 w1 w2) w3) = (and_w16 w1 (and_w16 w2 w3)).
73  #w1; #w2; #w3;
74  nelim w1;
75  #b1; #b2;
76  nelim w2;
77  #b3; #b4;
78  nelim w3;
79  #b5; #b6;
80  nchange with (mk_word16 (and_b8 (and_b8 b1 b3) b5) (and_b8 (and_b8 b2 b4) b6) =
81   mk_word16 (and_b8 b1 (and_b8 b3 b5)) (and_b8 b2 (and_b8 b4 b6)));
82  nrewrite < (associative_andb8 b1 b3 b5);
83  nrewrite < (associative_andb8 b2 b4 b6);
84  napply refl_eq.
85 nqed.
86
87 nlemma symmetric_orw16 : symmetricT word16 word16 or_w16.
88  #w1; #w2;
89  nelim w1;
90  nelim w2;
91  #b1; #b2; #b3; #b4;
92  nchange with ((mk_word16 (or_b8 b3 b1) (or_b8 b4 b2)) = (mk_word16 (or_b8 b1 b3) (or_b8 b2 b4)));
93  nrewrite > (symmetric_orb8 b1 b3);
94  nrewrite > (symmetric_orb8 b2 b4);
95  napply refl_eq.
96 nqed.
97
98 nlemma associative_orw16 : ∀w1,w2,w3.(or_w16 (or_w16 w1 w2) w3) = (or_w16 w1 (or_w16 w2 w3)).
99  #w1; #w2; #w3;
100  nelim w1;
101  #b1; #b2;
102  nelim w2;
103  #b3; #b4;
104  nelim w3;
105  #b5; #b6;
106  nchange with (mk_word16 (or_b8 (or_b8 b1 b3) b5) (or_b8 (or_b8 b2 b4) b6) =
107   mk_word16 (or_b8 b1 (or_b8 b3 b5)) (or_b8 b2 (or_b8 b4 b6)));
108  nrewrite < (associative_orb8 b1 b3 b5);
109  nrewrite < (associative_orb8 b2 b4 b6);
110  napply refl_eq.
111 nqed.
112
113 nlemma symmetric_xorw16 : symmetricT word16 word16 xor_w16.
114  #w1; #w2;
115  nelim w1;
116  nelim w2;
117  #b1; #b2; #b3; #b4;
118  nchange with ((mk_word16 (xor_b8 b3 b1) (xor_b8 b4 b2)) = (mk_word16 (xor_b8 b1 b3) (xor_b8 b2 b4)));
119  nrewrite > (symmetric_xorb8 b1 b3);
120  nrewrite > (symmetric_xorb8 b2 b4);
121  napply refl_eq.
122 nqed.
123
124 nlemma associative_xorw16 : ∀w1,w2,w3.(xor_w16 (xor_w16 w1 w2) w3) = (xor_w16 w1 (xor_w16 w2 w3)).
125  #w1; #w2; #w3;
126  nelim w1;
127  #b1; #b2;
128  nelim w2;
129  #b3; #b4;
130  nelim w3;
131  #b5; #b6;
132  nchange with (mk_word16 (xor_b8 (xor_b8 b1 b3) b5) (xor_b8 (xor_b8 b2 b4) b6) =
133   mk_word16 (xor_b8 b1 (xor_b8 b3 b5)) (xor_b8 b2 (xor_b8 b4 b6)));
134  nrewrite < (associative_xorb8 b1 b3 b5);
135  nrewrite < (associative_xorb8 b2 b4 b6);
136  napply refl_eq.
137 nqed.
138
139 nlemma symmetric_plusw16_dc_dc : ∀w1,w2,c.plus_w16_dc_dc w1 w2 c = plus_w16_dc_dc w2 w1 c.
140  #w1; #w2; #c;
141  nelim w1;
142  #b1; #b2;
143  nelim w2;
144  #b3; #b4;
145  nchange with (
146   match plus_b8_dc_dc b2 b4 c with [ pair l c ⇒ match plus_b8_dc_dc b1 b3 c with [ pair h c' ⇒ pair … 〈h:l〉 c' ]] =
147   match plus_b8_dc_dc b4 b2 c with [ pair l c ⇒ match plus_b8_dc_dc b3 b1 c with [ pair h c' ⇒ pair … 〈h:l〉 c' ]]);
148  nrewrite > (symmetric_plusb8_dc_dc b4 b2 c);
149  ncases (plus_b8_dc_dc b2 b4 c);
150  #b5; #c1;
151  nchange with (
152   match plus_b8_dc_dc b1 b3 c1 with [ pair h c' ⇒ pair … 〈h:b5〉 c' ] =
153   match plus_b8_dc_dc b3 b1 c1 with [ pair h c' ⇒ pair … 〈h:b5〉 c' ]);
154  nrewrite > (symmetric_plusb8_dc_dc b1 b3 c1);
155  napply refl_eq.
156 nqed.
157
158 nlemma symmetric_plusw16_dc_d : ∀w1,w2,c.plus_w16_dc_d w1 w2 c = plus_w16_dc_d w2 w1 c.
159  #w1; #w2; #c;
160  nelim w1;
161  #b1; #b2;
162  nelim w2;
163  #b3; #b4;
164  nchange with (
165   match plus_b8_dc_dc b2 b4 c with [ pair l c ⇒ 〈plus_b8_dc_d b1 b3 c:l〉 ] =
166   match plus_b8_dc_dc b4 b2 c with [ pair l c ⇒ 〈plus_b8_dc_d b3 b1 c:l〉 ]);
167  nrewrite > (symmetric_plusb8_dc_dc b4 b2 c);
168  ncases (plus_b8_dc_dc b2 b4 c);
169  #b5; #c1;
170  nchange with (〈plus_b8_dc_d b1 b3 c1:b5〉 = 〈plus_b8_dc_d b3 b1 c1:b5〉);
171  nrewrite > (symmetric_plusb8_dc_d b1 b3 c1);
172  napply refl_eq.
173 nqed.
174
175 nlemma symmetric_plusw16_dc_c : ∀w1,w2,c.plus_w16_dc_c w1 w2 c = plus_w16_dc_c w2 w1 c.
176  #w1; #w2; #c;
177  nelim w1;
178  #b1; #b2;
179  nelim w2;
180  #b3; #b4;
181  nchange with (
182   plus_b8_dc_c b1 b3 (plus_b8_dc_c b2 b4 c) =
183   plus_b8_dc_c b3 b1 (plus_b8_dc_c b4 b2 c));
184  nrewrite > (symmetric_plusb8_dc_c b4 b2 c);
185  nrewrite > (symmetric_plusb8_dc_c b3 b1 (plus_b8_dc_c b2 b4 c));
186  napply refl_eq.
187 nqed.
188
189 nlemma symmetric_plusw16_d_dc : ∀w1,w2.plus_w16_d_dc w1 w2 = plus_w16_d_dc w2 w1.
190  #w1; #w2;
191  nelim w1;
192  #b1; #b2;
193  nelim w2;
194  #b3; #b4;
195  nchange with (
196   match plus_b8_d_dc b2 b4 with [ pair l c ⇒ match plus_b8_dc_dc b1 b3 c with [ pair h c' ⇒ pair … 〈h:l〉 c' ]] =
197   match plus_b8_d_dc b4 b2 with [ pair l c ⇒ match plus_b8_dc_dc b3 b1 c with [ pair h c' ⇒ pair … 〈h:l〉 c' ]]);
198  nrewrite > (symmetric_plusb8_d_dc b4 b2);
199  ncases (plus_b8_d_dc b2 b4);
200  #b5; #c;
201  nchange with (
202   match plus_b8_dc_dc b1 b3 c with [ pair h c' ⇒ pair … 〈h:b5〉 c' ] =
203   match plus_b8_dc_dc b3 b1 c with [ pair h c' ⇒ pair … 〈h:b5〉 c' ]);
204  nrewrite > (symmetric_plusb8_dc_dc b1 b3 c);
205  napply refl_eq.
206 nqed.
207
208 nlemma symmetric_plusw16_d_d : ∀w1,w2.plus_w16_d_d w1 w2 = plus_w16_d_d w2 w1.
209  #w1; #w2;
210  nelim w1;
211  #b1; #b2;
212  nelim w2;
213  #b3; #b4;
214  nchange with (
215   match plus_b8_d_dc b2 b4 with [ pair l c ⇒ 〈plus_b8_dc_d b1 b3 c:l〉 ] =
216   match plus_b8_d_dc b4 b2 with [ pair l c ⇒ 〈plus_b8_dc_d b3 b1 c:l〉 ]);
217  nrewrite > (symmetric_plusb8_d_dc b4 b2);
218  ncases (plus_b8_d_dc b2 b4);
219  #b5; #c;
220  nchange with (〈plus_b8_dc_d b1 b3 c:b5〉 = 〈plus_b8_dc_d b3 b1 c:b5〉);
221  nrewrite > (symmetric_plusb8_dc_d b1 b3 c);
222  napply refl_eq.
223 nqed.
224
225 nlemma symmetric_plusw16_d_c : ∀w1,w2.plus_w16_d_c w1 w2 = plus_w16_d_c w2 w1.
226  #w1; #w2;
227  nelim w1;
228  #b1; #b2;
229  nelim w2;
230  #b3; #b4;
231  nchange with (
232   plus_b8_dc_c b1 b3 (plus_b8_d_c b2 b4) =
233   plus_b8_dc_c b3 b1 (plus_b8_d_c b4 b2));
234  nrewrite > (symmetric_plusb8_d_c b4 b2);
235  nrewrite > (symmetric_plusb8_dc_c b3 b1 (plus_b8_d_c b2 b4));
236  napply refl_eq.
237 nqed.
238
239 nlemma symmetric_mulb8 : symmetricT byte8 word16 mul_b8.
240  #b1; #b2;
241  nelim b1;
242  #e1; #e2;
243  nelim b2;
244  #e3; #e4;
245  nchange with (match mul_ex e2 e4 with
246  [ mk_byte8 t1_h t1_l ⇒ match mul_ex e1 e4 with
247  [ mk_byte8 t2_h t2_l ⇒ match mul_ex e3 e2 with
248  [ mk_byte8 t3_h t3_l ⇒ match mul_ex e1 e3 with
249  [ mk_byte8 t4_h t4_l ⇒ plus_w16_d_d (plus_w16_d_d (plus_w16_d_d 〈〈x0,t3_h〉:〈t3_l,x0〉〉 〈〈x0,t2_h〉:〈t2_l,x0〉〉) 〈〈t4_h,t4_l〉:〈x0,x0〉〉)〈〈x0,x0〉:〈t1_h,t1_l〉〉
250  ]]]] = match mul_ex e4 e2 with
251  [ mk_byte8 t1_h t1_l ⇒ match mul_ex e3 e2 with
252  [ mk_byte8 t2_h t2_l ⇒ match mul_ex e1 e4 with
253  [ mk_byte8 t3_h t3_l ⇒ match mul_ex e3 e1 with
254  [ mk_byte8 t4_h t4_l ⇒ plus_w16_d_d (plus_w16_d_d (plus_w16_d_d 〈〈x0,t3_h〉:〈t3_l,x0〉〉 〈〈x0,t2_h〉:〈t2_l,x0〉〉) 〈〈t4_h,t4_l〉:〈x0,x0〉〉)〈〈x0,x0〉:〈t1_h,t1_l〉〉
255  ]]]]);
256  nrewrite < (symmetric_mulex e2 e4);
257  ncases (mul_ex e2 e4);
258  #e5; #e6;
259  nchange with (match mul_ex e1 e4 with
260  [ mk_byte8 t2_h t2_l ⇒ match mul_ex e3 e2 with
261  [ mk_byte8 t3_h t3_l ⇒ match mul_ex e1 e3 with
262  [ mk_byte8 t4_h t4_l ⇒ plus_w16_d_d (plus_w16_d_d (plus_w16_d_d 〈〈x0,t3_h〉:〈t3_l,x0〉〉 〈〈x0,t2_h〉:〈t2_l,x0〉〉) 〈〈t4_h,t4_l〉:〈x0,x0〉〉)〈〈x0,x0〉:〈e5,e6〉〉
263  ]]] = match mul_ex e3 e2 with
264  [ mk_byte8 t2_h t2_l ⇒ match mul_ex e1 e4 with
265  [ mk_byte8 t3_h t3_l ⇒ match mul_ex e3 e1 with
266  [ mk_byte8 t4_h t4_l ⇒ plus_w16_d_d (plus_w16_d_d (plus_w16_d_d 〈〈x0,t3_h〉:〈t3_l,x0〉〉 〈〈x0,t2_h〉:〈t2_l,x0〉〉) 〈〈t4_h,t4_l〉:〈x0,x0〉〉)〈〈x0,x0〉:〈e5,e6〉〉
267  ]]]);
268  ncases (mul_ex e3 e2);
269  #e7; #e8;
270  ncases (mul_ex e1 e4);
271  #e9; #e10;
272  nchange with (match mul_ex e1 e3 with
273  [ mk_byte8 t4_h t4_l ⇒ plus_w16_d_d (plus_w16_d_d (plus_w16_d_d 〈〈x0,e7〉:〈e8,x0〉〉 〈〈x0,e9〉:〈e10,x0〉〉) 〈〈t4_h,t4_l〉:〈x0,x0〉〉)〈〈x0,x0〉:〈e5,e6〉〉
274  ] = match mul_ex e3 e1 with
275  [ mk_byte8 t4_h t4_l ⇒ plus_w16_d_d (plus_w16_d_d (plus_w16_d_d 〈〈x0,e9〉:〈e10,x0〉〉 〈〈x0,e7〉:〈e8,x0〉〉) 〈〈t4_h,t4_l〉:〈x0,x0〉〉)〈〈x0,x0〉:〈e5,e6〉〉
276  ]);
277  nrewrite < (symmetric_mulex e1 e3);
278  ncases (mul_ex e1 e3);
279  #e11; #e12;
280  nchange with ((plus_w16_d_d (plus_w16_d_d (plus_w16_d_d 〈〈x0,e7〉:〈e8,x0〉〉 〈〈x0,e9〉:〈e10,x0〉〉) 〈〈e11,e12〉:〈x0,x0〉〉)〈〈x0,x0〉:〈e5,e6〉〉) =
281   (plus_w16_d_d (plus_w16_d_d (plus_w16_d_d 〈〈x0,e9〉:〈e10,x0〉〉 〈〈x0,e7〉:〈e8,x0〉〉) 〈〈e11,e12〉:〈x0,x0〉〉)〈〈x0,x0〉:〈e5,e6〉〉));
282  nrewrite > (symmetric_plusw16_d_d 〈〈x0,e7〉:〈e8,x0〉〉 〈〈x0,e9〉:〈e10,x0〉〉);
283  napply refl_eq.
284 nqed.
285
286 nlemma eqw16_to_eq : ∀w1,w2:word16.(eq_w16 w1 w2 = true) → (w1 = w2).
287  #w1; #w2;
288  nelim w1;
289  nelim w2;
290  #b1; #b2; #b3; #b4;
291  nchange in ⊢ (% → ?) with (((eq_b8 b3 b1)⊗(eq_b8 b4 b2)) = true);
292  #H;
293  nrewrite < (eqb8_to_eq … (andb_true_true_l … H));
294  nrewrite < (eqb8_to_eq … (andb_true_true_r … H));
295  napply refl_eq.
296 nqed.
297
298 nlemma eq_to_eqw16 : ∀w1,w2.w1 = w2 → eq_w16 w1 w2 = true.
299  #w1; #w2;
300  nelim w1;
301  nelim w2;
302  #b1; #b2; #b3; #b4;
303  #H;
304  nrewrite < (word16_destruct_1 … H);
305  nrewrite < (word16_destruct_2 … H);
306  nchange with (((eq_b8 b3 b3)⊗(eq_b8 b4 b4)) = true);
307  nrewrite > (eq_to_eqb8 b3 b3 (refl_eq …));
308  nrewrite > (eq_to_eqb8 b4 b4 (refl_eq …)); 
309  nnormalize;
310  napply refl_eq.
311 nqed.
312
313 nlemma decidable_w16_aux1 : ∀b1,b2,b3,b4.b1 ≠ b3 → 〈b1:b2〉 ≠ 〈b3:b4〉.
314  #b1; #b2; #b3; #b4;
315  nnormalize; #H; #H1;
316  napply (H (word16_destruct_1 … H1)).
317 nqed.
318
319 nlemma decidable_w16_aux2 : ∀b1,b2,b3,b4.b2 ≠ b4 → 〈b1:b2〉 ≠ 〈b3:b4〉.
320  #b1; #b2; #b3; #b4;
321  nnormalize; #H; #H1;
322  napply (H (word16_destruct_2 … H1)).
323 nqed.
324
325 nlemma decidable_w16 : ∀x,y:word16.decidable (x = y).
326  #x; nelim x; #b1; #b2;
327  #y; nelim y; #b3; #b4;
328  nnormalize;
329  napply (or2_elim (b1 = b3) (b1 ≠ b3) ? (decidable_b8 b1 b3) …);
330  ##[ ##2: #H; napply (or2_intro2 … (decidable_w16_aux1 b1 b2 b3 b4 H))
331  ##| ##1: #H; napply (or2_elim (b2 = b4) (b2 ≠ b4) ? (decidable_b8 b2 b4) …);
332           ##[ ##2: #H1; napply (or2_intro2 … (decidable_w16_aux2 b1 b2 b3 b4 H1))
333           ##| ##1: #H1; nrewrite > H; nrewrite > H1;
334                         napply (or2_intro1 … (refl_eq ? 〈b3:b4〉))
335           ##]
336  ##]
337 nqed.
338
339 nlemma neqw16_to_neq : ∀w1,w2:word16.(eq_w16 w1 w2 = false) → (w1 ≠ w2).
340  #w1; #w2;
341  nelim w1;
342  nelim w2;
343  #b1; #b2; #b3; #b4;
344  nchange with ((((eq_b8 b3 b1) ⊗ (eq_b8 b4 b2)) = false) → ?);
345  #H;
346  napply (or2_elim ((eq_b8 b3 b1) = false) ((eq_b8 b4 b2) = false) ? (andb_false2 … H) …);
347  ##[ ##1: #H1; napply (decidable_w16_aux1 … (neqb8_to_neq … H1))
348  ##| ##2: #H1; napply (decidable_w16_aux2 … (neqb8_to_neq … H1))
349  ##]
350 nqed.
351
352 nlemma word16_destruct : ∀b1,b2,b3,b4.〈b1:b2〉 ≠ 〈b3:b4〉 → b1 ≠ b3 ∨ b2 ≠ b4.
353  #b1; #b2; #b3; #b4;
354  nnormalize; #H;
355  napply (or2_elim (b1 = b3) (b1 ≠ b3) ? (decidable_b8 b1 b3) …);
356  ##[ ##2: #H1; napply (or2_intro1 … H1)
357  ##| ##1: #H1; napply (or2_elim (b2 = b4) (b2 ≠ b4) ? (decidable_b8 b2 b4) …);
358           ##[ ##2: #H2; napply (or2_intro2 … H2)
359           ##| ##1: #H2; nrewrite > H1 in H:(%);
360                    nrewrite > H2;
361                    #H; nelim (H (refl_eq …))
362           ##]
363  ##]
364 nqed.
365
366 nlemma neq_to_neqw16 : ∀w1,w2.w1 ≠ w2 → eq_w16 w1 w2 = false.
367  #w1; #w2;
368  nelim w1; #b1; #b2;
369  nelim w2; #b3; #b4;
370  #H; nchange with (((eq_b8 b1 b3) ⊗ (eq_b8 b2 b4)) = false);
371  napply (or2_elim (b1 ≠ b3) (b2 ≠ b4) ? (word16_destruct … H) …);
372  ##[ ##1: #H1; nrewrite > (neq_to_neqb8 … H1); nnormalize; napply refl_eq
373  ##| ##2: #H1; nrewrite > (neq_to_neqb8 … H1);
374           nrewrite > (symmetric_andbool (eq_b8 b1 b3) false);
375           nnormalize; napply refl_eq
376  ##]
377 nqed.