X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fng_assembly%2Fnum%2Fword32_lemmas.ma;h=7f1c9ddc8d08d769ced18c0faf357fd23a6d7158;hb=29cfb9e2961e62c836cb50217905c0594a074e81;hp=d7f6cc5588da9e94803eb6b9deb90ee6742af88a;hpb=b1c174cffd3c1d10383a52d63a6e662156fb0bb7;p=helm.git diff --git a/helm/software/matita/contribs/ng_assembly/num/word32_lemmas.ma b/helm/software/matita/contribs/ng_assembly/num/word32_lemmas.ma index d7f6cc558..7f1c9ddc8 100755 --- a/helm/software/matita/contribs/ng_assembly/num/word32_lemmas.ma +++ b/helm/software/matita/contribs/ng_assembly/num/word32_lemmas.ma @@ -15,8 +15,8 @@ (* ********************************************************************** *) (* Progetto FreeScale *) (* *) -(* Sviluppato da: Cosimo Oliboni, oliboni@cs.unibo.it *) -(* Cosimo Oliboni, oliboni@cs.unibo.it *) +(* Sviluppato da: Ing. Cosimo Oliboni, oliboni@cs.unibo.it *) +(* Ultima modifica: 05/08/2009 *) (* *) (* ********************************************************************** *) @@ -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 (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_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 (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.