X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fng_assembly%2Fnum%2Fbyte8_lemmas.ma;h=3fa810b46ce5b5b53565e1570f4b70eae914789d;hb=3e373f30525987b7b121ed74c8a1292dc71d185c;hp=1d32793cc3e1b724048df8598278c2572fe18047;hpb=38fccc2b774e493a94eedef76342b56079c0e694;p=helm.git diff --git a/helm/software/matita/contribs/ng_assembly/num/byte8_lemmas.ma b/helm/software/matita/contribs/ng_assembly/num/byte8_lemmas.ma index 1d32793cc..3fa810b46 100755 --- a/helm/software/matita/contribs/ng_assembly/num/byte8_lemmas.ma +++ b/helm/software/matita/contribs/ng_assembly/num/byte8_lemmas.ma @@ -16,273 +16,27 @@ (* Progetto FreeScale *) (* *) (* Sviluppato da: Ing. Cosimo Oliboni, oliboni@cs.unibo.it *) -(* Ultima modifica: 05/08/2009 *) +(* Sviluppo: 2008-2010 *) (* *) (* ********************************************************************** *) include "num/byte8.ma". include "num/exadecim_lemmas.ma". +include "num/comp_num_lemmas.ma". (* **** *) (* BYTE *) (* **** *) -nlemma byte8_destruct_1 : -∀x1,x2,y1,y2. - mk_byte8 x1 y1 = mk_byte8 x2 y2 → x1 = x2. - #x1; #x2; #y1; #y2; #H; - nchange with (match mk_byte8 x2 y2 with [ mk_byte8 a _ ⇒ x1 = a ]); - nrewrite < H; - nnormalize; - napply refl_eq. -nqed. +ndefinition byte8_destruct_1 ≝ cn_destruct_1 byte8. +ndefinition byte8_destruct_2 ≝ cn_destruct_2 byte8. -nlemma byte8_destruct_2 : -∀x1,x2,y1,y2. - mk_byte8 x1 y1 = mk_byte8 x2 y2 → y1 = y2. - #x1; #x2; #y1; #y2; #H; - nchange with (match mk_byte8 x2 y2 with [ mk_byte8 _ b ⇒ y1 = b ]); - nrewrite < H; - nnormalize; - napply refl_eq. -nqed. +ndefinition symmetric_eqb8 ≝ symmetric_eqcn ? eq_ex symmetric_eqex. -nlemma symmetric_eqb8 : symmetricT byte8 bool eq_b8. - #b1; #b2; - nelim b1; - nelim b2; - #e1; #e2; #e3; #e4; - nchange with (((eq_ex e3 e1)⊗(eq_ex e4 e2)) = ((eq_ex e1 e3)⊗(eq_ex e2 e4))); - nrewrite > (symmetric_eqex e1 e3); - nrewrite > (symmetric_eqex e2 e4); - napply refl_eq. -nqed. +ndefinition eqb8_to_eq ≝ eqcn_to_eq ? eq_ex eqex_to_eq. +ndefinition eq_to_eqb8 ≝ eq_to_eqcn ? eq_ex eq_to_eqex. -nlemma symmetric_andb8 : symmetricT byte8 byte8 and_b8. - #b1; #b2; - nelim b1; - nelim b2; - #e1; #e2; #e3; #e4; - nchange with ((mk_byte8 (and_ex e3 e1) (and_ex e4 e2)) = (mk_byte8 (and_ex e1 e3) (and_ex e2 e4))); - nrewrite > (symmetric_andex e1 e3); - nrewrite > (symmetric_andex e2 e4); - napply refl_eq. -nqed. +ndefinition decidable_b8 ≝ decidable_cn ? decidable_ex. -nlemma associative_andb8 : ∀b1,b2,b3.(and_b8 (and_b8 b1 b2) b3) = (and_b8 b1 (and_b8 b2 b3)). - #b1; #b2; #b3; - nelim b1; - #e1; #e2; - nelim b2; - #e3; #e4; - nelim b3; - #e5; #e6; - nchange with (mk_byte8 (and_ex (and_ex e1 e3) e5) (and_ex (and_ex e2 e4) e6) = - mk_byte8 (and_ex e1 (and_ex e3 e5)) (and_ex e2 (and_ex e4 e6))); - nrewrite < (associative_andex e1 e3 e5); - nrewrite < (associative_andex e2 e4 e6); - napply refl_eq. -nqed. - -nlemma symmetric_orb8 : symmetricT byte8 byte8 or_b8. - #b1; #b2; - nelim b1; - nelim b2; - #e1; #e2; #e3; #e4; - nchange with ((mk_byte8 (or_ex e3 e1) (or_ex e4 e2)) = (mk_byte8 (or_ex e1 e3) (or_ex e2 e4))); - nrewrite > (symmetric_orex e1 e3); - nrewrite > (symmetric_orex e2 e4); - napply refl_eq. -nqed. - -nlemma associative_orb8 : ∀b1,b2,b3.(or_b8 (or_b8 b1 b2) b3) = (or_b8 b1 (or_b8 b2 b3)). - #b1; #b2; #b3; - nelim b1; - #e1; #e2; - nelim b2; - #e3; #e4; - nelim b3; - #e5; #e6; - nchange with (mk_byte8 (or_ex (or_ex e1 e3) e5) (or_ex (or_ex e2 e4) e6) = - mk_byte8 (or_ex e1 (or_ex e3 e5)) (or_ex e2 (or_ex e4 e6))); - nrewrite < (associative_orex e1 e3 e5); - nrewrite < (associative_orex e2 e4 e6); - napply refl_eq. -nqed. - -nlemma symmetric_xorb8 : symmetricT byte8 byte8 xor_b8. - #b1; #b2; - nelim b1; - nelim b2; - #e1; #e2; #e3; #e4; - nchange with ((mk_byte8 (xor_ex e3 e1) (xor_ex e4 e2)) = (mk_byte8 (xor_ex e1 e3) (xor_ex e2 e4))); - nrewrite > (symmetric_xorex e1 e3); - nrewrite > (symmetric_xorex e2 e4); - napply refl_eq. -nqed. - -nlemma associative_xorb8 : ∀b1,b2,b3.(xor_b8 (xor_b8 b1 b2) b3) = (xor_b8 b1 (xor_b8 b2 b3)). - #b1; #b2; #b3; - nelim b1; - #e1; #e2; - nelim b2; - #e3; #e4; - nelim b3; - #e5; #e6; - nchange with (mk_byte8 (xor_ex (xor_ex e1 e3) e5) (xor_ex (xor_ex e2 e4) e6) = - mk_byte8 (xor_ex e1 (xor_ex e3 e5)) (xor_ex e2 (xor_ex e4 e6))); - nrewrite < (associative_xorex e1 e3 e5); - nrewrite < (associative_xorex e2 e4 e6); - napply refl_eq. -nqed. - -nlemma symmetric_plusb8_dc_dc : ∀b1,b2,c.plus_b8_dc_dc b1 b2 c = plus_b8_dc_dc b2 b1 c. - #b1; #b2; #c; - nelim b1; - #e1; #e2; - nelim b2; - #e3; #e4; - nchange with ( - match plus_ex_dc_dc e2 e4 c with - [ pair l c ⇒ match plus_ex_dc_dc e1 e3 c with - [ pair h c' ⇒ pair … 〈h,l〉 c' ]] = - match plus_ex_dc_dc e4 e2 c with - [ pair l c ⇒ match plus_ex_dc_dc e3 e1 c with - [ pair h c' ⇒ pair … 〈h,l〉 c' ]]); - nrewrite > (symmetric_plusex_dc_dc e4 e2 c); - ncases (plus_ex_dc_dc e2 e4 c); - #e5; #c1; - nchange with ( - match plus_ex_dc_dc e1 e3 c1 with - [ pair h c' ⇒ pair … 〈h,e5〉 c' ] = - match plus_ex_dc_dc e3 e1 c1 with - [ pair h c' ⇒ pair … 〈h,e5〉 c' ]); - nrewrite > (symmetric_plusex_dc_dc e1 e3 c1); - napply refl_eq. -nqed. - -nlemma symmetric_plusb8_dc_d : ∀b1,b2,c.plus_b8_dc_d b1 b2 c = plus_b8_dc_d b2 b1 c. - #b1; #b2; #c; - nelim b1; - #e1; #e2; - nelim b2; - #e3; #e4; - nchange with ( - match plus_ex_dc_dc e2 e4 c with - [ pair l c ⇒ 〈plus_ex_dc_d e1 e3 c,l〉 ] = - match plus_ex_dc_dc e4 e2 c with - [ pair l c ⇒ 〈plus_ex_dc_d e3 e1 c,l〉 ]); - nrewrite > (symmetric_plusex_dc_dc e4 e2 c); - ncases (plus_ex_dc_dc e2 e4 c); - #e5; #c1; - nchange with (〈plus_ex_dc_d e1 e3 c1,e5〉 = 〈plus_ex_dc_d e3 e1 c1,e5〉); - nrewrite > (symmetric_plusex_dc_d e1 e3 c1); - napply refl_eq. -nqed. - -nlemma symmetric_plusb8_dc_c : ∀b1,b2,c.plus_b8_dc_c b1 b2 c = plus_b8_dc_c b2 b1 c. - #b1; #b2; #c; - nelim b1; - #e1; #e2; - nelim b2; - #e3; #e4; - nchange with ( - plus_ex_dc_c e1 e3 (plus_ex_dc_c e2 e4 c) = - plus_ex_dc_c e3 e1 (plus_ex_dc_c e4 e2 c)); - nrewrite > (symmetric_plusex_dc_c e4 e2 c); - nrewrite > (symmetric_plusex_dc_c e3 e1 (plus_ex_dc_c e2 e4 c)); - napply refl_eq. -nqed. - -nlemma symmetric_plusb8_d_dc : ∀b1,b2.plus_b8_d_dc b1 b2 = plus_b8_d_dc b2 b1. - #b1; #b2; - nelim b1; - #e1; #e2; - nelim b2; - #e3; #e4; - nchange with ( - match plus_ex_d_dc e2 e4 with - [ pair l c ⇒ match plus_ex_dc_dc e1 e3 c with - [ pair h c' ⇒ pair … 〈h,l〉 c' ]] = - match plus_ex_d_dc e4 e2 with - [ pair l c ⇒ match plus_ex_dc_dc e3 e1 c with - [ pair h c' ⇒ pair … 〈h,l〉 c' ]]); - nrewrite > (symmetric_plusex_d_dc e4 e2); - ncases (plus_ex_d_dc e2 e4); - #e5; #c; - nchange with ( - match plus_ex_dc_dc e1 e3 c with - [ pair h c' ⇒ pair … 〈h,e5〉 c' ] = - match plus_ex_dc_dc e3 e1 c with - [ pair h c' ⇒ pair … 〈h,e5〉 c' ]); - nrewrite > (symmetric_plusex_dc_dc e1 e3 c); - napply refl_eq. -nqed. - -nlemma symmetric_plusb8_d_d : ∀b1,b2.plus_b8_d_d b1 b2 = plus_b8_d_d b2 b1. - #b1; #b2; - nelim b1; - #e1; #e2; - nelim b2; - #e3; #e4; - nchange with ( - match plus_ex_d_dc e2 e4 with - [ pair l c ⇒ 〈plus_ex_dc_d e1 e3 c,l〉 ] = - match plus_ex_d_dc e4 e2 with - [ pair l c ⇒ 〈plus_ex_dc_d e3 e1 c,l〉 ]); - nrewrite > (symmetric_plusex_d_dc e4 e2); - ncases (plus_ex_d_dc e2 e4); - #e5; #c; - nchange with (〈plus_ex_dc_d e1 e3 c,e5〉 = 〈plus_ex_dc_d e3 e1 c,e5〉); - nrewrite > (symmetric_plusex_dc_d e1 e3 c); - napply refl_eq. -nqed. - -nlemma symmetric_plusb8_d_c : ∀b1,b2.plus_b8_d_c b1 b2 = plus_b8_d_c b2 b1. - #b1; #b2; - nelim b1; - #e1; #e2; - nelim b2; - #e3; #e4; - nchange with ( - plus_ex_dc_c e1 e3 (plus_ex_d_c e2 e4) = - plus_ex_dc_c e3 e1 (plus_ex_d_c e4 e2)); - nrewrite > (symmetric_plusex_d_c e4 e2); - nrewrite > (symmetric_plusex_dc_c e3 e1 (plus_ex_d_c e2 e4)); - napply refl_eq. -nqed. - -nlemma symmetric_mulex : symmetricT exadecim byte8 mul_ex. - #e1; #e2; - nelim e1; - nelim e2; - nnormalize; - napply refl_eq. -nqed. - -nlemma eqb8_to_eq : ∀b1,b2:byte8.(eq_b8 b1 b2 = true) → (b1 = b2). - #b1; #b2; - nelim b1; - nelim b2; - #e1; #e2; #e3; #e4; - nchange in ⊢ (% → ?) with (((eq_ex e3 e1)⊗(eq_ex e4 e2)) = true); - #H; - nrewrite < (eqex_to_eq … (andb_true_true_l … H)); - nrewrite < (eqex_to_eq … (andb_true_true_r … H)); - napply refl_eq. -nqed. - -nlemma eq_to_eqb8 : ∀b1,b2.b1 = b2 → eq_b8 b1 b2 = true. - #b1; #b2; - nelim b1; - nelim b2; - #e1; #e2; #e3; #e4; - #H; - nrewrite < (byte8_destruct_1 … H); - nrewrite < (byte8_destruct_2 … H); - nchange with (((eq_ex e3 e3)⊗(eq_ex e4 e4)) = true); - nrewrite > (eq_to_eqex e3 e3 (refl_eq …)); - nrewrite > (eq_to_eqex e4 e4 (refl_eq …)); - nnormalize; - napply refl_eq. -nqed. +ndefinition neqb8_to_neq ≝ neqcn_to_neq ? eq_ex neqex_to_neq. +ndefinition neq_to_neqb8 ≝ neq_to_neqcn ? eq_ex neq_to_neqex decidable_ex.