]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_assembly/num/word32_lemmas.ma
Release 0.5.9.
[helm.git] / helm / software / matita / contribs / ng_assembly / num / word32_lemmas.ma
index 84679640a15d79b9b27b57cb58068a61bacec2df..5f42bbc7bc22401e8c01f34dd5cf76b0ed63bd0a 100755 (executable)
@@ -16,7 +16,7 @@
 (*                          Progetto FreeScale                            *)
 (*                                                                        *)
 (*   Sviluppato da: Ing. Cosimo Oliboni, oliboni@cs.unibo.it              *)
-(*   Sviluppo: 2008-2010                                                  *)
+(*   Ultima modifica: 05/08/2009                                          *)
 (*                                                                        *)
 (* ********************************************************************** *)
 
@@ -27,15 +27,351 @@ include "num/word16_lemmas.ma".
 (* DWORD *)
 (* ***** *)
 
-ndefinition word32_destruct_1 ≝ cn_destruct_1 word32.
-ndefinition word32_destruct_2 ≝ cn_destruct_2 word32.
+nlemma word32_destruct_1 :
+∀x1,x2,y1,y2.
+ mk_word32 x1 y1 = mk_word32 x2 y2 → x1 = x2.
+ #x1; #x2; #y1; #y2; #H;
+ nchange with (match mk_word32 x2 y2 with [ mk_word32 a _ ⇒ x1 = a ]);
+ nrewrite < H;
+ nnormalize;
+ napply refl_eq.
+nqed.
 
-ndefinition symmetric_eqw32 ≝ symmetric_eqcn ? eq_w16 symmetric_eqw16.
+nlemma word32_destruct_2 :
+∀x1,x2,y1,y2.
+ mk_word32 x1 y1 = mk_word32 x2 y2 → y1 = y2.
+ #x1; #x2; #y1; #y2; #H;
+ nchange with (match mk_word32 x2 y2 with [ mk_word32 _ b ⇒ y1 = b ]);
+ nrewrite < H;
+ nnormalize;
+ napply refl_eq.
+nqed.
 
-ndefinition eqw32_to_eq ≝ eqcn_to_eq ? eq_w16 eqw16_to_eq.
-ndefinition eq_to_eqw32 ≝ eq_to_eqcn ? eq_w16 eq_to_eqw16.
+nlemma symmetric_eqw32 : symmetricT word32 bool eq_w32.
+ #dw1; #dw2;
+ nelim dw1;
+ nelim dw2;
+ #w1; #w2; #w3; #w4;
+ nchange with (((eq_w16 w3 w1)⊗(eq_w16 w4 w2)) = ((eq_w16 w1 w3)⊗(eq_w16 w2 w4)));
+ nrewrite > (symmetric_eqw16 w1 w3);
+ nrewrite > (symmetric_eqw16 w2 w4);
+ napply refl_eq.
+nqed.
 
-ndefinition decidable_w32 ≝ decidable_cn ? decidable_w16.
+nlemma symmetric_andw32 : symmetricT word32 word32 and_w32.
+ #dw1; #dw2;
+ nelim dw1;
+ nelim dw2;
+ #w1; #w2; #w3; #w4;
+ nchange with ((mk_word32 (and_w16 w3 w1) (and_w16 w4 w2)) = (mk_word32 (and_w16 w1 w3) (and_w16 w2 w4)));
+ nrewrite > (symmetric_andw16 w1 w3);
+ nrewrite > (symmetric_andw16 w2 w4);
+ napply refl_eq.
+nqed.
 
-ndefinition neqw32_to_neq ≝ neqcn_to_neq ? eq_w16 neqw16_to_neq.
-ndefinition neq_to_neqw32 ≝ neq_to_neqcn ? eq_w16 neq_to_neqw16 decidable_w16.
+nlemma associative_andw32 : ∀dw1,dw2,dw3.(and_w32 (and_w32 dw1 dw2) dw3) = (and_w32 dw1 (and_w32 dw2 dw3)).
+ #dw1; #dw2; #dw3;
+ nelim dw1;
+ #w1; #w2;
+ nelim dw2;
+ #w3; #w4;
+ nelim dw3;
+ #w5; #w6;
+ nchange with (mk_word32 (and_w16 (and_w16 w1 w3) w5) (and_w16 (and_w16 w2 w4) w6) =
+  mk_word32 (and_w16 w1 (and_w16 w3 w5)) (and_w16 w2 (and_w16 w4 w6)));
+ nrewrite < (associative_andw16 w1 w3 w5);
+ nrewrite < (associative_andw16 w2 w4 w6);
+ napply refl_eq.
+nqed.
+
+nlemma symmetric_orw32 : symmetricT word32 word32 or_w32.
+ #dw1; #dw2;
+ nelim dw1;
+ nelim dw2;
+ #w1; #w2; #w3; #w4;
+ nchange with ((mk_word32 (or_w16 w3 w1) (or_w16 w4 w2)) = (mk_word32 (or_w16 w1 w3) (or_w16 w2 w4)));
+ nrewrite > (symmetric_orw16 w1 w3);
+ nrewrite > (symmetric_orw16 w2 w4);
+ napply refl_eq.
+nqed.
+
+nlemma associative_orw32 : ∀dw1,dw2,dw3.(or_w32 (or_w32 dw1 dw2) dw3) = (or_w32 dw1 (or_w32 dw2 dw3)).
+ #dw1; #dw2; #dw3;
+ nelim dw1;
+ #w1; #w2;
+ nelim dw2;
+ #w3; #w4;
+ nelim dw3;
+ #w5; #w6;
+ nchange with (mk_word32 (or_w16 (or_w16 w1 w3) w5) (or_w16 (or_w16 w2 w4) w6) =
+  mk_word32 (or_w16 w1 (or_w16 w3 w5)) (or_w16 w2 (or_w16 w4 w6)));
+ nrewrite < (associative_orw16 w1 w3 w5);
+ nrewrite < (associative_orw16 w2 w4 w6);
+ napply refl_eq.
+nqed.
+
+nlemma symmetric_xorw32 : symmetricT word32 word32 xor_w32.
+ #dw1; #dw2;
+ nelim dw1;
+ nelim dw2;
+ #w1; #w2; #w3; #w4;
+ nchange with ((mk_word32 (xor_w16 w3 w1) (xor_w16 w4 w2)) = (mk_word32 (xor_w16 w1 w3) (xor_w16 w2 w4)));
+ nrewrite > (symmetric_xorw16 w1 w3);
+ nrewrite > (symmetric_xorw16 w2 w4);
+ napply refl_eq.
+nqed.
+
+nlemma associative_xorw32 : ∀dw1,dw2,dw3.(xor_w32 (xor_w32 dw1 dw2) dw3) = (xor_w32 dw1 (xor_w32 dw2 dw3)).
+ #dw1; #dw2; #dw3;
+ nelim dw1;
+ #w1; #w2;
+ nelim dw2;
+ #w3; #w4;
+ nelim dw3;
+ #w5; #w6;
+ nchange with (mk_word32 (xor_w16 (xor_w16 w1 w3) w5) (xor_w16 (xor_w16 w2 w4) w6) =
+  mk_word32 (xor_w16 w1 (xor_w16 w3 w5)) (xor_w16 w2 (xor_w16 w4 w6)));
+ nrewrite < (associative_xorw16 w1 w3 w5);
+ nrewrite < (associative_xorw16 w2 w4 w6);
+ napply refl_eq.
+nqed.
+
+nlemma symmetric_plusw32_dc_dc : ∀dw1,dw2,c.plus_w32_dc_dc dw1 dw2 c = plus_w32_dc_dc dw2 dw1 c.
+ #dw1; #dw2; #c;
+ nelim dw1;
+ #w1; #w2;
+ nelim dw2;
+ #w3; #w4;
+ nchange with (
+  match plus_w16_dc_dc w2 w4 c with [ pair l c ⇒ match plus_w16_dc_dc w1 w3 c with [ pair h c' ⇒ pair … 〈h.l〉 c' ]] =
+  match plus_w16_dc_dc w4 w2 c with [ pair l c ⇒ match plus_w16_dc_dc w3 w1 c with [ pair h c' ⇒ pair … 〈h.l〉 c' ]]);
+ nrewrite > (symmetric_plusw16_dc_dc w4 w2 c);
+ ncases (plus_w16_dc_dc w2 w4 c);
+ #w5; #c1;
+ nchange with (
+  match plus_w16_dc_dc w1 w3 c1 with [ pair h c' ⇒ pair … 〈h.w5〉 c' ] =
+  match plus_w16_dc_dc w3 w1 c1 with [ pair h c' ⇒ pair … 〈h.w5〉 c' ]);
+ nrewrite > (symmetric_plusw16_dc_dc w1 w3 c1);
+ napply refl_eq.
+nqed.
+
+nlemma symmetric_plusw32_dc_d : ∀dw1,dw2,c.plus_w32_dc_d dw1 dw2 c = plus_w32_dc_d dw2 dw1 c.
+ #dw1; #dw2; #c;
+ nelim dw1;
+ #w1; #w2;
+ nelim dw2;
+ #w3; #w4;
+ nchange with (
+  match plus_w16_dc_dc w2 w4 c with [ pair l c ⇒ 〈plus_w16_dc_d w1 w3 c.l〉 ] =
+  match plus_w16_dc_dc w4 w2 c with [ pair l c ⇒ 〈plus_w16_dc_d w3 w1 c.l〉 ]);
+ nrewrite > (symmetric_plusw16_dc_dc w4 w2 c);
+ ncases (plus_w16_dc_dc w2 w4 c);
+ #w5; #c1;
+ nchange with (〈plus_w16_dc_d w1 w3 c1.w5〉 = 〈plus_w16_dc_d w3 w1 c1.w5〉);
+ nrewrite > (symmetric_plusw16_dc_d w1 w3 c1);
+ napply refl_eq.
+nqed.
+
+nlemma symmetric_plusw32_dc_c : ∀dw1,dw2,c.plus_w32_dc_c dw1 dw2 c = plus_w32_dc_c dw2 dw1 c.
+ #dw1; #dw2; #c;
+ nelim dw1;
+ #w1; #w2;
+ nelim dw2;
+ #w3; #w4;
+ nchange with (
+  plus_w16_dc_c w1 w3 (plus_w16_dc_c w2 w4 c) =
+  plus_w16_dc_c w3 w1 (plus_w16_dc_c w4 w2 c));
+ nrewrite > (symmetric_plusw16_dc_c w4 w2 c);
+ nrewrite > (symmetric_plusw16_dc_c w3 w1 (plus_w16_dc_c w2 w4 c));
+ napply refl_eq.
+nqed.
+
+nlemma symmetric_plusw32_d_dc : ∀dw1,dw2.plus_w32_d_dc dw1 dw2 = plus_w32_d_dc dw2 dw1.
+ #dw1; #dw2;
+ nelim dw1;
+ #w1; #w2;
+ nelim dw2;
+ #w3; #w4;
+ nchange with (
+  match plus_w16_d_dc w2 w4 with [ pair l c ⇒ match plus_w16_dc_dc w1 w3 c with [ pair h c' ⇒ pair … 〈h.l〉 c' ]] =
+  match plus_w16_d_dc w4 w2 with [ pair l c ⇒ match plus_w16_dc_dc w3 w1 c with [ pair h c' ⇒ pair … 〈h.l〉 c' ]]);
+ nrewrite > (symmetric_plusw16_d_dc w4 w2);
+ ncases (plus_w16_d_dc w2 w4);
+ #w5; #c;
+ nchange with (
+  match plus_w16_dc_dc w1 w3 c with [ pair h c' ⇒ pair … 〈h.w5〉 c' ] =
+  match plus_w16_dc_dc w3 w1 c with [ pair h c' ⇒ pair … 〈h.w5〉 c' ]);
+ nrewrite > (symmetric_plusw16_dc_dc w1 w3 c);
+ napply refl_eq.
+nqed.
+
+nlemma symmetric_plusw32_d_d : ∀dw1,dw2.plus_w32_d_d dw1 dw2 = plus_w32_d_d dw2 dw1.
+ #dw1; #dw2;
+ nelim dw1;
+ #w1; #w2;
+ nelim dw2;
+ #w3; #w4;
+ nchange with (
+  match plus_w16_d_dc w2 w4 with [ pair l c ⇒ 〈plus_w16_dc_d w1 w3 c.l〉 ] =
+  match plus_w16_d_dc w4 w2 with [ pair l c ⇒ 〈plus_w16_dc_d w3 w1 c.l〉 ]);
+ nrewrite > (symmetric_plusw16_d_dc w4 w2);
+ ncases (plus_w16_d_dc w2 w4);
+ #w5; #c;
+ nchange with (〈plus_w16_dc_d w1 w3 c.w5〉 = 〈plus_w16_dc_d w3 w1 c.w5〉);
+ nrewrite > (symmetric_plusw16_dc_d w1 w3 c);
+ napply refl_eq.
+nqed.
+
+nlemma symmetric_plusw32_d_c : ∀dw1,dw2.plus_w32_d_c dw1 dw2 = plus_w32_d_c dw2 dw1.
+ #dw1; #dw2;
+ nelim dw1;
+ #w1; #w2;
+ nelim dw2;
+ #w3; #w4;
+ nchange with (
+  plus_w16_dc_c w1 w3 (plus_w16_d_c w2 w4) =
+  plus_w16_dc_c w3 w1 (plus_w16_d_c w4 w2));
+ nrewrite > (symmetric_plusw16_d_c w4 w2);
+ nrewrite > (symmetric_plusw16_dc_c w3 w1 (plus_w16_d_c w2 w4));
+ napply refl_eq.
+nqed.
+
+nlemma symmetric_mulw16 : symmetricT word16 word32 mul_w16.
+ #w1; #w2;
+ nelim w1;
+ #b1; #b2;
+ nelim w2;
+ #b3; #b4;
+ nchange with (match mul_b8 b2 b4 with
+ [ mk_word16 t1_h t1_l ⇒ match mul_b8 b1 b4 with
+ [ mk_word16 t2_h t2_l ⇒ match mul_b8 b3 b2 with
+ [ mk_word16 t3_h t3_l ⇒ match mul_b8 b1 b3 with
+ [ mk_word16 t4_h t4_l ⇒ plus_w32_d_d (plus_w32_d_d (plus_w32_d_d 〈〈〈x0,x0〉:t3_h〉.〈t3_l:〈x0,x0〉〉〉 〈〈〈x0,x0〉:t2_h〉.〈t2_l:〈x0,x0〉〉〉) 〈〈t4_h:t4_l〉.〈〈x0,x0〉:〈x0,x0〉〉〉)〈〈〈x0,x0〉:〈x0,x0〉〉.〈t1_h:t1_l〉〉 
+ ]]]] = match mul_b8 b4 b2 with
+ [ mk_word16 t1_h t1_l ⇒ match mul_b8 b3 b2 with
+ [ mk_word16 t2_h t2_l ⇒ match mul_b8 b1 b4 with
+ [ mk_word16 t3_h t3_l ⇒ match mul_b8 b3 b1 with
+ [ mk_word16 t4_h t4_l ⇒ plus_w32_d_d (plus_w32_d_d (plus_w32_d_d 〈〈〈x0,x0〉:t3_h〉.〈t3_l:〈x0,x0〉〉〉 〈〈〈x0,x0〉:t2_h〉.〈t2_l:〈x0,x0〉〉〉) 〈〈t4_h:t4_l〉.〈〈x0,x0〉:〈x0,x0〉〉〉)〈〈〈x0,x0〉:〈x0,x0〉〉.〈t1_h:t1_l〉〉 
+ ]]]]);
+ nrewrite < (symmetric_mulb8 b2 b4);
+ ncases (mul_b8 b2 b4);
+ #b5; #b6;
+ nchange with (match mul_b8 b1 b4 with
+ [ mk_word16 t2_h t2_l ⇒ match mul_b8 b3 b2 with
+ [ mk_word16 t3_h t3_l ⇒ match mul_b8 b1 b3 with
+ [ mk_word16 t4_h t4_l ⇒ plus_w32_d_d (plus_w32_d_d (plus_w32_d_d 〈〈〈x0,x0〉:t3_h〉.〈t3_l:〈x0,x0〉〉〉 〈〈〈x0,x0〉:t2_h〉.〈t2_l:〈x0,x0〉〉〉) 〈〈t4_h:t4_l〉.〈〈x0,x0〉:〈x0,x0〉〉〉)〈〈〈x0,x0〉:〈x0,x0〉〉.〈b5:b6〉〉 
+ ]]] = match mul_b8 b3 b2 with
+ [ mk_word16 t2_h t2_l ⇒ match mul_b8 b1 b4 with
+ [ mk_word16 t3_h t3_l ⇒ match mul_b8 b3 b1 with
+ [ mk_word16 t4_h t4_l ⇒ plus_w32_d_d (plus_w32_d_d (plus_w32_d_d 〈〈〈x0,x0〉:t3_h〉.〈t3_l:〈x0,x0〉〉〉 〈〈〈x0,x0〉:t2_h〉.〈t2_l:〈x0,x0〉〉〉) 〈〈t4_h:t4_l〉.〈〈x0,x0〉:〈x0,x0〉〉〉)〈〈〈x0,x0〉:〈x0,x0〉〉.〈b5:b6〉〉 
+ ]]]);
+ ncases (mul_b8 b3 b2);
+ #b7; #b8;
+ ncases (mul_b8 b1 b4);
+ #b9; #b10;
+ nchange with (match mul_b8 b1 b3 with
+ [ mk_word16 t4_h t4_l ⇒ plus_w32_d_d (plus_w32_d_d (plus_w32_d_d 〈〈〈x0,x0〉:b7〉.〈b8:〈x0,x0〉〉〉 〈〈〈x0,x0〉:b9〉.〈b10:〈x0,x0〉〉〉) 〈〈t4_h:t4_l〉.〈〈x0,x0〉:〈x0,x0〉〉〉)〈〈〈x0,x0〉:〈x0,x0〉〉.〈b5:b6〉〉 
+ ] = match mul_b8 b3 b1 with
+ [ mk_word16 t4_h t4_l ⇒ plus_w32_d_d (plus_w32_d_d (plus_w32_d_d 〈〈〈x0,x0〉:b9〉.〈b10:〈x0,x0〉〉〉 〈〈〈x0,x0〉:b7〉.〈b8:〈x0,x0〉〉〉) 〈〈t4_h:t4_l〉.〈〈x0,x0〉:〈x0,x0〉〉〉)〈〈〈x0,x0〉:〈x0,x0〉〉.〈b5:b6〉〉 
+ ]);
+ nrewrite < (symmetric_mulb8 b1 b3);
+ ncases (mul_b8 b1 b3);
+ #b11; #b12;
+ nchange with ((plus_w32_d_d (plus_w32_d_d (plus_w32_d_d 〈〈〈x0,x0〉:b7〉.〈b8:〈x0,x0〉〉〉 〈〈〈x0,x0〉:b9〉.〈b10:〈x0,x0〉〉〉) 〈〈b11:b12〉.〈〈x0,x0〉:〈x0,x0〉〉〉)〈〈〈x0,x0〉:〈x0,x0〉〉.〈b5:b6〉〉) =
+  (plus_w32_d_d (plus_w32_d_d (plus_w32_d_d 〈〈〈x0,x0〉:b9〉.〈b10:〈x0,x0〉〉〉 〈〈〈x0,x0〉:b7〉.〈b8:〈x0,x0〉〉〉) 〈〈b11:b12〉.〈〈x0,x0〉:〈x0,x0〉〉〉)〈〈〈x0,x0〉:〈x0,x0〉〉.〈b5:b6〉〉));
+ nrewrite > (symmetric_plusw32_d_d 〈〈〈x0,x0〉:b7〉.〈b8:〈x0,x0〉〉〉 〈〈〈x0,x0〉:b9〉.〈b10:〈x0,x0〉〉〉);
+ napply refl_eq.
+nqed.
+
+nlemma eqw32_to_eq : ∀dw1,dw2:word32.(eq_w32 dw1 dw2 = true) → (dw1 = dw2).
+ #dw1; #dw2;
+ nelim dw1;
+ nelim dw2;
+ #w1; #w2; #w3; #w4;
+ nchange in ⊢ (% → ?) with (((eq_w16 w3 w1)⊗(eq_w16 w4 w2)) = true);
+ #H;
+ nrewrite < (eqw16_to_eq … (andb_true_true_l … H));
+ nrewrite < (eqw16_to_eq … (andb_true_true_r … H));
+ napply refl_eq.
+nqed.
+
+nlemma eq_to_eqw32 : ∀dw1,dw2.dw1 = dw2 → eq_w32 dw1 dw2 = true.
+ #dw1; #dw2;
+ nelim dw1;
+ nelim dw2;
+ #w1; #w2; #w3; #w4;
+ #H;
+ nrewrite < (word32_destruct_1 … H);
+ nrewrite < (word32_destruct_2 … H);
+ nchange with (((eq_w16 w3 w3)⊗(eq_w16 w4 w4)) = true);
+ nrewrite > (eq_to_eqw16 w3 w3 (refl_eq …));
+ nrewrite > (eq_to_eqw16 w4 w4 (refl_eq …)); 
+ nnormalize;
+ napply refl_eq.
+nqed.
+
+nlemma decidable_w32_aux1 : ∀w1,w2,w3,w4.w1 ≠ w3 → 〈w1.w2〉 ≠ 〈w3.w4〉.
+ #w1; #w2; #w3; #w4;
+ nnormalize; #H; #H1;
+ napply (H (word32_destruct_1 … H1)).
+nqed.
+
+nlemma decidable_w32_aux2 : ∀w1,w2,w3,w4.w2 ≠ w4 → 〈w1.w2〉 ≠ 〈w3.w4〉.
+ #w1; #w2; #w3; #w4;
+ nnormalize; #H; #H1;
+ napply (H (word32_destruct_2 … H1)).
+nqed.
+
+nlemma decidable_w32 : ∀x,y:word32.decidable (x = y).
+ #x; nelim x; #w1; #w2;
+ #y; nelim y; #w3; #w4;
+ nnormalize;
+ napply (or2_elim (w1 = w3) (w1 ≠ w3) ? (decidable_w16 w1 w3) …);
+ ##[ ##2: #H; napply (or2_intro2 … (decidable_w32_aux1 w1 w2 w3 w4 H))
+ ##| ##1: #H; napply (or2_elim (w2 = w4) (w2 ≠ w4) ? (decidable_w16 w2 w4) …);
+          ##[ ##2: #H1; napply (or2_intro2 … (decidable_w32_aux2 w1 w2 w3 w4 H1))
+          ##| ##1: #H1; nrewrite > H; nrewrite > H1;
+                        napply (or2_intro1 … (refl_eq ? 〈w3.w4〉))
+          ##]
+ ##]
+nqed.
+
+nlemma neqw32_to_neq : ∀dw1,dw2:word32.(eq_w32 dw1 dw2 = false) → (dw1 ≠ dw2).
+ #dw1; #dw2;
+ nelim dw1;
+ nelim dw2;
+ #w1; #w2; #w3; #w4;
+ nchange with ((((eq_w16 w3 w1) ⊗ (eq_w16 w4 w2)) = false) → ?);
+ #H;
+ napply (or2_elim ((eq_w16 w3 w1) = false) ((eq_w16 w4 w2) = false) ? (andb_false2 … H) …);
+ ##[ ##1: #H1; napply (decidable_w32_aux1 … (neqw16_to_neq … H1))
+ ##| ##2: #H1; napply (decidable_w32_aux2 … (neqw16_to_neq … H1))
+ ##]
+nqed.
+
+nlemma word32_destruct : ∀w1,w2,w3,w4.〈w1.w2〉 ≠ 〈w3.w4〉 → w1 ≠ w3 ∨ w2 ≠ w4.
+ #w1; #w2; #w3; #w4;
+ nnormalize; #H;
+ napply (or2_elim (w1 = w3) (w1 ≠ w3) ? (decidable_w16 w1 w3) …);
+ ##[ ##2: #H1; napply (or2_intro1 … H1)
+ ##| ##1: #H1; napply (or2_elim (w2 = w4) (w2 ≠ w4) ? (decidable_w16 w2 w4) …);
+          ##[ ##2: #H2; napply (or2_intro2 … H2)
+          ##| ##1: #H2; nrewrite > H1 in H:(%);
+                   nrewrite > H2;
+                   #H; nelim (H (refl_eq …))
+          ##]
+ ##]
+nqed.
+
+nlemma neq_to_neqw32 : ∀dw1,dw2.dw1 ≠ dw2 → eq_w32 dw1 dw2 = false.
+ #dw1; #dw2;
+ nelim dw1; #w1; #w2;
+ nelim dw2; #w3; #w4;
+ #H; nchange with (((eq_w16 w1 w3) ⊗ (eq_w16 w2 w4)) = false);
+ napply (or2_elim (w1 ≠ w3) (w2 ≠ w4) ? (word32_destruct … H) …);
+ ##[ ##1: #H1; nrewrite > (neq_to_neqw16 … H1); nnormalize; napply refl_eq
+ ##| ##2: #H1; nrewrite > (neq_to_neqw16 … H1);
+          nrewrite > (symmetric_andbool (eq_w16 w1 w3) false);
+          nnormalize; napply refl_eq
+ ##]
+nqed.