]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_assembly/num/word32_lemmas.ma
freescale porting, work in progress
[helm.git] / helm / software / matita / contribs / ng_assembly / num / word32_lemmas.ma
index 98639a64257f3eb82d09654b46ee5b7dab359851..2bb6c0b24c00df969d64e8307699d6c3e64a01f2 100755 (executable)
@@ -325,3 +325,69 @@ nlemma eq_to_eqw32 : ∀dw1,dw2.dw1 = dw2 → eq_w32 dw1 dw2 = true.
  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 (or_elim (w1 = w3) (w1 ≠ w3) ? (decidable_w16 w1 w3) …);
+ ##[ ##2: #H; napply (or_intror … (decidable_w32_aux1 w1 w2 w3 w4 H))
+ ##| ##1: #H; napply (or_elim (w2 = w4) (w2 ≠ w4) ? (decidable_w16 w2 w4) …);
+          ##[ ##2: #H1; napply (or_intror … (decidable_w32_aux2 w1 w2 w3 w4 H1))
+          ##| ##1: #H1; nrewrite > H; nrewrite > H1;
+                        napply (or_introl … (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 (or_elim ((eq_w16 w3 w1) = false) ((eq_w16 w4 w2) = false) ? (andb_false … 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 (or_elim (w1 = w3) (w1 ≠ w3) ? (decidable_w16 w1 w3) …);
+ ##[ ##2: #H1; napply (or_introl … H1)
+ ##| ##1: #H1; napply (or_elim (w2 = w4) (w2 ≠ w4) ? (decidable_w16 w2 w4) …);
+          ##[ ##2: #H2; napply (or_intror … 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 (or_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.