##]
nqed.
-nlemma symmetric_foldrightlist2_aux
- : ∀T1,T2:Type.∀f:T1 → T1 → T2 → T2.∀acc:T2.∀l1,l2:list T1.
- ∀H1:len_list T1 l1 = len_list T1 l2.∀H2:len_list T1 l2 = len_list T1 l1.
- (∀x,y,z.f x y z = f y x z) →
- fold_right_list2 T1 T2 f acc l1 l2 H1 = fold_right_list2 T1 T2 f acc l2 l1 H2.
- #T1; #T2; #f; #acc; #l1;
+nlemma symmetric_foldrightlist2_aux :
+∀T1,T2:Type.∀f:T1 → T1 → T2 → T2.
+ (∀x,y,z.f x y z = f y x z) →
+ (∀acc:T2.∀l1,l2:list T1.
+ ∀H1:(len_list T1 l1 = len_list T1 l2).
+ ∀H2:(len_list T1 l2 = len_list T1 l1).
+ (fold_right_list2 T1 T2 f acc l1 l2 H1 = fold_right_list2 T1 T2 f acc l2 l1 H2)).
+ #T1; #T2; #f; #H; #acc; #l1;
nelim l1;
##[ ##1: #l2; ncases l2;
- ##[ ##1: nnormalize; #H1; #H2; #H3; napply refl_eq
- ##| ##2: #h; #l; #H1; #H2; #H3;
+ ##[ ##1: nnormalize; #H1; #H2; napply refl_eq
+ ##| ##2: #h; #l; #H1; #H2;
nchange in H1:(%) with (O = (S (len_list ? l)));
ndestruct (*nelim (nat_destruct_0_S ? H1)*)
##]
- ##| ##2: #h3; #l3; #H; #l2; ncases l2;
- ##[ ##1: #H1; #H2; #H3; nchange in H1:(%) with ((S (len_list ? l3)) = O);
+ ##| ##2: #h3; #l3; #H1; #l2; ncases l2;
+ ##[ ##1: #H2; #H3; nchange in H2:(%) with ((S (len_list ? l3)) = O);
ndestruct (*nelim (nat_destruct_S_0 ? H1)*)
- ##| ##2: #h4; #l4; #H1; #H2; #H3;
- nchange in H1:(%) with ((S (len_list ? l3)) = (S (len_list ? l4)));
- nchange in H2:(%) with ((S (len_list ? l4)) = (S (len_list ? l3)));
+ ##| ##2: #h4; #l4; #H2; #H3;
+ nchange in H2:(%) with ((S (len_list ? l3)) = (S (len_list ? l4)));
+ nchange in H3:(%) with ((S (len_list ? l4)) = (S (len_list ? l3)));
nchange with ((f h3 h4 (fold_right_list2 T1 T2 f acc l3 l4 (fold_right_list2_aux3 T1 h3 h4 l3 l4 ?))) =
(f h4 h3 (fold_right_list2 T1 T2 f acc l4 l3 (fold_right_list2_aux3 T1 h4 h3 l4 l3 ?))));
- nrewrite < (H l4 (fold_right_list2_aux3 T1 h3 h4 l3 l4 H1) (fold_right_list2_aux3 T1 h4 h3 l4 l3 H2) H3);
- nrewrite > (H3 h3 h4 (fold_right_list2 T1 T2 f acc l3 l4 ?));
+ nrewrite < (H1 l4 (fold_right_list2_aux3 T1 h3 h4 l3 l4 H2) (fold_right_list2_aux3 T1 h4 h3 l4 l3 H3));
+ nrewrite > (H h3 h4 (fold_right_list2 T1 T2 f acc l3 l4 ?));
napply refl_eq
##]
##]
nqed.
-nlemma symmetric_foldrightlist2
- : ∀T1,T2:Type.∀f:T1 → T1 → T2 → T2.∀acc:T2.∀l1,l2:list T1.∀H:len_list T1 l1 = len_list T1 l2.
- (∀x,y,z.f x y z = f y x z) →
- fold_right_list2 T1 T2 f acc l1 l2 H = fold_right_list2 T1 T2 f acc l2 l1 (symmetric_lenlist T1 l1 l2 H).
- #T1; #T2; #f; #acc; #l1; #l2; #H; #H1;
- nrewrite > (symmetric_foldrightlist2_aux T1 T2 f acc l1 l2 H (symmetric_lenlist T1 l1 l2 H) H1);
+nlemma symmetric_foldrightlist2 :
+∀T1,T2:Type.∀f:T1 → T1 → T2 → T2.
+ (∀x,y,z.f x y z = f y x z) →
+ (∀acc:T2.∀l1,l2:list T1.∀H:len_list T1 l1 = len_list T1 l2.
+ fold_right_list2 T1 T2 f acc l1 l2 H = fold_right_list2 T1 T2 f acc l2 l1 (symmetric_lenlist T1 l1 l2 H)).
+ #T1; #T2; #f; #H; #acc; #l1; #l2; #H1;
+ nrewrite > (symmetric_foldrightlist2_aux T1 T2 f H acc l1 l2 H1 (symmetric_lenlist T1 l1 l2 H1));
napply refl_eq.
nqed.
-nlemma symmetric_bfoldrightlist2
- : ∀T1:Type.∀f:T1 → T1 → bool.∀l1,l2:list T1.
- (∀x,y.f x y = f y x) →
- bfold_right_list2 T1 f l1 l2 = bfold_right_list2 T1 f l2 l1.
- #T; #f; #l1;
+nlemma symmetric_bfoldrightlist2 :
+∀T1:Type.∀f:T1 → T1 → bool.
+ (∀x,y.f x y = f y x) →
+ (∀l1,l2:list T1.
+ bfold_right_list2 T1 f l1 l2 = bfold_right_list2 T1 f l2 l1).
+ #T; #f; #H; #l1;
nelim l1;
##[ ##1: #l2; ncases l2;
- ##[ ##1: #H; nnormalize; napply refl_eq
- ##| ##2: #hh2; #ll2; #H; nnormalize; napply refl_eq
+ ##[ ##1: nnormalize; napply refl_eq
+ ##| ##2: #hh2; #ll2; nnormalize; napply refl_eq
##]
- ##| ##2: #hh1; #ll1; #H; #l2; ncases l2;
- ##[ ##1: #H1; nnormalize; napply refl_eq
- ##| ##2: #hh2; #ll2; #H1; nnormalize;
- nrewrite > (H ll2 H1);
- nrewrite > (H1 hh1 hh2);
+ ##| ##2: #hh1; #ll1; #H1; #l2; ncases l2;
+ ##[ ##1: nnormalize; napply refl_eq
+ ##| ##2: #hh2; #ll2; nnormalize;
+ nrewrite > (H1 ll2);
+ nrewrite > (H hh1 hh2);
napply refl_eq
##]
##]
nqed.
-nlemma bfoldrightlist2_to_eq
- : ∀T1:Type.∀f:T1 → T1 → bool.∀l1,l2:list T1.
- (∀x,y.(f x y = true → x = y)) →
- (bfold_right_list2 T1 f l1 l2 = true → l1 = l2).
- #T; #f; #l1;
+nlemma bfoldrightlist2_to_eq :
+∀T1:Type.∀f:T1 → T1 → bool.
+ (∀x,y.(f x y = true → x = y)) →
+ (∀l1,l2:list T1.
+ (bfold_right_list2 T1 f l1 l2 = true → l1 = l2)).
+ #T; #f; #H; #l1;
nelim l1;
##[ ##1: #l2; ncases l2;
- ##[ ##1: #H; #H1; napply refl_eq
- ##| ##2: #hh2; #ll2; #H; nnormalize; #H1;
+ ##[ ##1: #H1; napply refl_eq
+ ##| ##2: #hh2; #ll2; nnormalize; #H1;
ndestruct (*napply (bool_destruct … H1)*)
##]
- ##| ##2: #hh1; #ll1; #H; #l2; ncases l2;
- ##[ ##1: #H1; nnormalize; #H2;
+ ##| ##2: #hh1; #ll1; #H1; #l2; ncases l2;
+ ##[ ##1: nnormalize; #H2;
ndestruct (*napply (bool_destruct … H2)*)
- ##| ##2: #hh2; #ll2; #H1; #H2;
+ ##| ##2: #hh2; #ll2; #H2;
nchange in H2:(%) with (((f hh1 hh2)⊗(bfold_right_list2 T f ll1 ll2)) = true);
- nrewrite > (H1 hh1 hh2 (andb_true_true_l … H2));
- nrewrite > (H ll2 H1 (andb_true_true_r … H2));
+ nrewrite > (H hh1 hh2 (andb_true_true_l … H2));
+ nrewrite > (H1 ll2 (andb_true_true_r … H2));
napply refl_eq
##]
##]
nqed.
-nlemma eq_to_bfoldrightlist2
- : ∀T1:Type.∀f:T1 → T1 → bool.∀l1,l2:list T1.
- (∀x,y.(x = y → f x y = true)) →
- (l1 = l2 → bfold_right_list2 T1 f l1 l2 = true).
- #T; #f; #l1;
+nlemma eq_to_bfoldrightlist2 :
+∀T1:Type.∀f:T1 → T1 → bool.
+ (∀x,y.(x = y → f x y = true)) →
+ (∀l1,l2:list T1.
+ (l1 = l2 → bfold_right_list2 T1 f l1 l2 = true)).
+ #T; #f; #H; #l1;
nelim l1;
##[ ##1: #l2; ncases l2;
- ##[ ##1: #H; #H1; nnormalize; napply refl_eq
- ##| ##2: #hh2; #ll2; #H; #H1;
+ ##[ ##1: #H1; nnormalize; napply refl_eq
+ ##| ##2: #hh2; #ll2; #H1;
(* !!! ndestruct: assert false *)
nelim (list_destruct_nil_cons ??? H1)
##]
- ##| ##2: #hh1; #ll1; #H; #l2; ncases l2;
- ##[ ##1: #H1; #H2;
+ ##| ##2: #hh1; #ll1; #H1; #l2; ncases l2;
+ ##[ ##1: #H2;
(* !!! ndestruct: assert false *)
nelim (list_destruct_cons_nil ??? H2)
- ##| ##2: #hh2; #ll2; #H1; #H2; nnormalize;
+ ##| ##2: #hh2; #ll2; #H2; nnormalize;
nrewrite > (list_destruct_1 … H2);
- nrewrite > (H1 hh2 hh2 (refl_eq …));
+ nrewrite > (H hh2 hh2 (refl_eq …));
nnormalize;
- nrewrite > (H ll2 H1 (list_destruct_2 … H2));
+ nrewrite > (H1 ll2 (list_destruct_2 … H2));
napply refl_eq
##]
##]
nqed.
-nlemma bfoldrightlist2_to_lenlist
- : ∀T.∀f:T → T → bool.∀l1,l2:list T.bfold_right_list2 T f l1 l2 = true → len_list T l1 = len_list T l2.
+nlemma bfoldrightlist2_to_lenlist :
+∀T.∀f:T → T → bool.
+ (∀l1,l2:list T.bfold_right_list2 T f l1 l2 = true → len_list T l1 = len_list T l2).
#T; #f; #l1;
nelim l1;
##[ ##1: #l2; ncases l2;
##]
nqed.
-nlemma decidable_list : ∀T.∀H:(Πx,y:T.decidable (x = y)).∀x,y:list T.decidable (x = y).
+nlemma decidable_list :
+∀T.(∀x,y:T.decidable (x = y)) →
+ (∀x,y:list T.decidable (x = y)).
#T; #H; #x; nelim x;
##[ ##1: #y; ncases y;
##[ ##1: nnormalize; napply (or2_intro1 (? = ?) (? ≠ ?) (refl_eq …))
##]
nqed.
-nlemma nbfoldrightlist2_to_neq
- : ∀T1:Type.∀f:T1 → T1 → bool.∀l1,l2:list T1.
- (∀x,y.(f x y = false → x ≠ y)) →
- (bfold_right_list2 T1 f l1 l2 = false → l1 ≠ l2).
- #T; #f; #l1;
+nlemma nbfoldrightlist2_to_neq :
+∀T1:Type.∀f:T1 → T1 → bool.
+ (∀x,y.(f x y = false → x ≠ y)) →
+ (∀l1,l2:list T1.
+ (bfold_right_list2 T1 f l1 l2 = false → l1 ≠ l2)).
+ #T; #f; #H; #l1;
nelim l1;
##[ ##1: #l2; ncases l2;
- ##[ ##1: #H; nnormalize; #H1;
+ ##[ ##1: nnormalize; #H1;
ndestruct (*napply (bool_destruct … H1)*)
- ##| ##2: #hh2; #ll2; #H; #H1; nnormalize; #H2;
+ ##| ##2: #hh2; #ll2; #H1; nnormalize; #H2;
(* !!! ndestruct: assert false *)
napply (list_destruct_nil_cons T … H2)
##]
- ##| ##2: #hh1; #ll1; #H; #l2; ncases l2;
- ##[ ##1: #H1; #H2; nnormalize; #H3;
+ ##| ##2: #hh1; #ll1; #H1; #l2; ncases l2;
+ ##[ ##1: #H2; nnormalize; #H3;
(* !!! ndestruct: assert false *)
napply (list_destruct_cons_nil T … H3)
- ##| ##2: #hh2; #ll2; #H1; #H2; nnormalize; #H3;
+ ##| ##2: #hh2; #ll2; #H2; nnormalize; #H3;
nchange in H2:(%) with (((f hh1 hh2)⊗(bfold_right_list2 T f ll1 ll2)) = false);
- napply (H ll2 H1 ? (list_destruct_2 T … H3));
+ napply (H1 ll2 ? (list_destruct_2 T … H3));
napply (or2_elim ??? (andb_false2 … H2) );
##[ ##1: #H4; napply (absurd (hh1 = hh2) …);
##[ ##1: nrewrite > (list_destruct_1 T … H3); napply refl_eq
- ##| ##2: napply (H1 … H4)
+ ##| ##2: napply (H … H4)
##]
##| ##2: #H4; napply H4
##]
##]
nqed.
-nlemma list_destruct
- : ∀T.∀H:(Πx,y:T.decidable (x = y)).∀h1,h2:T.∀l1,l2:list T.(h1::l1) ≠ (h2::l2) → h1 ≠ h2 ∨ l1 ≠ l2.
+nlemma list_destruct :
+∀T.(∀x,y:T.decidable (x = y)) →
+ (∀h1,h2:T.∀l1,l2:list T.
+ (h1::l1) ≠ (h2::l2) → h1 ≠ h2 ∨ l1 ≠ l2).
#T; #H; #h1; #h2; #l1; nelim l1;
##[ ##1: #l2; ncases l2;
##[ ##1: #H1; napply (or2_intro1 (h1 ≠ h2) ([] ≠ []) …);
##]
nqed.
-nlemma neq_to_nbfoldrightlist2
- : ∀T:Type.∀f:T → T → bool.∀l1,l2:list T.
- (∀x,y:T.decidable (x = y)) →
- (∀x,y.(x ≠ y → f x y = false)) →
- (l1 ≠ l2 → bfold_right_list2 T f l1 l2 = false).
- #T; #f; #l1;
+nlemma neq_to_nbfoldrightlist2 :
+∀T:Type.∀f:T → T → bool.
+ (∀x,y:T.decidable (x = y)) →
+ (∀x,y.(x ≠ y → f x y = false)) →
+ (∀l1,l2:list T.
+ (l1 ≠ l2 → bfold_right_list2 T f l1 l2 = false)).
+ #T; #f; #H; #H1; #l1;
nelim l1;
##[ ##1: #l2; ncases l2;
- ##[ ##1: #H; #H1; nnormalize; #H2; nelim (H2 (refl_eq …))
- ##| ##2: #hh2; #ll2; #H; nnormalize; #H1; #H2; napply refl_eq
+ ##[ ##1: nnormalize; #H2; nelim (H2 (refl_eq …))
+ ##| ##2: #hh2; #ll2; nnormalize; #H2; napply refl_eq
##]
- ##| ##2: #hh1; #ll1; #H; #l2; ncases l2;
- ##[ ##1: #H1; #H2; nnormalize; #H3; napply refl_eq
- ##| ##2: #hh2; #ll2; #H1; #H2; #H3;
+ ##| ##2: #hh1; #ll1; #H2; #l2; ncases l2;
+ ##[ ##1: nnormalize; #H3; napply refl_eq
+ ##| ##2: #hh2; #ll2; #H3;
nchange with (((f hh1 hh2)⊗(bfold_right_list2 T f ll1 ll2)) = false);
- napply (or2_elim (hh1 ≠ hh2) (ll1 ≠ ll2) ? (list_destruct T H1 … H3) …);
- ##[ ##1: #H4; nrewrite > (H2 hh1 hh2 H4); nnormalize; napply refl_eq
- ##| ##2: #H4; nrewrite > (H ll2 H1 H2 H4);
+ napply (or2_elim (hh1 ≠ hh2) (ll1 ≠ ll2) ? (list_destruct T H … H3) …);
+ ##[ ##1: #H4; nrewrite > (H1 hh1 hh2 H4); nnormalize; napply refl_eq
+ ##| ##2: #H4; nrewrite > (H2 ll2 H4);
nrewrite > (symmetric_andbool (f hh1 hh2) false);
nnormalize; napply refl_eq
##]
##]
nqed.
-nlemma symmetric_foldrightnelist2_aux
- : ∀T1,T2:Type.∀f:T1 → T1 → T2 → T2.∀acc:T2.∀l1,l2:ne_list T1.
- ∀H1:len_neList T1 l1 = len_neList T1 l2.∀H2:len_neList T1 l2 = len_neList T1 l1.
- (∀x,y,z.f x y z = f y x z) →
- fold_right_neList2 T1 T2 f acc l1 l2 H1 = fold_right_neList2 T1 T2 f acc l2 l1 H2.
- #T1; #T2; #f; #acc; #l1;
+nlemma symmetric_foldrightnelist2_aux :
+∀T1,T2:Type.∀f:T1 → T1 → T2 → T2.
+ (∀x,y,z.f x y z = f y x z) →
+ (∀acc:T2.∀l1,l2:ne_list T1.
+ ∀H1:len_neList T1 l1 = len_neList T1 l2.∀H2:len_neList T1 l2 = len_neList T1 l1.
+ fold_right_neList2 T1 T2 f acc l1 l2 H1 = fold_right_neList2 T1 T2 f acc l2 l1 H2).
+ #T1; #T2; #f; #H; #acc; #l1;
nelim l1;
##[ ##1: #h; #l2; ncases l2;
- ##[ ##1: #h1; nnormalize; #H1; #H2; #H3; nrewrite > (H3 h h1 acc); napply refl_eq
+ ##[ ##1: #h1; nnormalize; #H1; #H2; nrewrite > (H h h1 acc); napply refl_eq
##| ##2: #h1; #l; ncases l;
- ##[ ##1: #h3; #H1; #H2; #H3;
+ ##[ ##1: #h3; #H1; #H2;
nchange in H1:(%) with ((S O) = (S (S O)));
(* !!! ndestruct: si inceppa su un'ipotesi che non e' H1 *)
nelim (nat_destruct_0_S ? (nat_destruct_S_S … H1))
- ##| ##2: #h3; #l3; #H1; #H2; #H3;
+ ##| ##2: #h3; #l3; #H1; #H2;
nchange in H1:(%) with ((S O) = (S (S (len_neList ? l3))));
(* !!! ndestruct: si inceppa su un'ipotesi che non e' H1 *)
nelim (nat_destruct_0_S ? (nat_destruct_S_S … H1))
##]
##]
- ##| ##2: #h3; #l3; #H; #l2; ncases l2;
+ ##| ##2: #h3; #l3; #H1; #l2; ncases l2;
##[ ##1: #h4; ncases l3;
- ##[ ##1: #h5; #H1; #H2; #H3;
- nchange in H1:(%) with ((S (S O)) = (S O));
+ ##[ ##1: #h5; #H2; #H3;
+ nchange in H2:(%) with ((S (S O)) = (S O));
(* !!! ndestruct: si inceppa su un'ipotesi che non e' H1 *)
- nelim (nat_destruct_S_0 ? (nat_destruct_S_S … H1))
- ##| ##2: #h5; #l5; #H1; #H2; #H3;
- nchange in H1:(%) with ((S (S (len_neList ? l5))) = (S O));
+ nelim (nat_destruct_S_0 ? (nat_destruct_S_S … H2))
+ ##| ##2: #h5; #l5; #H2; #H3;
+ nchange in H2:(%) with ((S (S (len_neList ? l5))) = (S O));
(* !!! ndestruct: si inceppa su un'ipotesi che non e' H1 *)
- nelim (nat_destruct_S_0 ? (nat_destruct_S_S … H1))
+ nelim (nat_destruct_S_0 ? (nat_destruct_S_S … H2))
##]
- ##| ##2: #h4; #l4; #H1; #H2; #H3;
- nchange in H1:(%) with ((S (len_neList ? l3)) = (S (len_neList ? l4)));
- nchange in H2:(%) with ((S (len_neList ? l4)) = (S (len_neList ? l3)));
+ ##| ##2: #h4; #l4; #H2; #H3;
+ nchange in H2:(%) with ((S (len_neList ? l3)) = (S (len_neList ? l4)));
+ nchange in H3:(%) with ((S (len_neList ? l4)) = (S (len_neList ? l3)));
nchange with ((f h3 h4 (fold_right_neList2 T1 T2 f acc l3 l4 (fold_right_neList2_aux3 T1 h3 h4 l3 l4 ?))) =
(f h4 h3 (fold_right_neList2 T1 T2 f acc l4 l3 (fold_right_neList2_aux3 T1 h4 h3 l4 l3 ?))));
- nrewrite < (H l4 (fold_right_neList2_aux3 T1 h3 h4 l3 l4 H1) (fold_right_neList2_aux3 T1 h4 h3 l4 l3 H2) H3);
- nrewrite > (H3 h3 h4 (fold_right_neList2 T1 T2 f acc l3 l4 ?));
+ nrewrite < (H1 l4 (fold_right_neList2_aux3 T1 h3 h4 l3 l4 H2) (fold_right_neList2_aux3 T1 h4 h3 l4 l3 H3));
+ nrewrite > (H h3 h4 (fold_right_neList2 T1 T2 f acc l3 l4 ?));
napply refl_eq
##]
##]
nqed.
-nlemma symmetric_foldrightnelist2
- : ∀T1,T2:Type.∀f:T1 → T1 → T2 → T2.∀acc:T2.∀l1,l2:ne_list T1.∀H:len_neList T1 l1 = len_neList T1 l2.
- (∀x,y,z.f x y z = f y x z) →
- fold_right_neList2 T1 T2 f acc l1 l2 H = fold_right_neList2 T1 T2 f acc l2 l1 (symmetric_lennelist T1 l1 l2 H).
- #T1; #T2; #f; #acc; #l1; #l2; #H; #H1;
- nrewrite > (symmetric_foldrightnelist2_aux T1 T2 f acc l1 l2 H (symmetric_lennelist T1 l1 l2 H) H1);
+nlemma symmetric_foldrightnelist2 :
+∀T1,T2:Type.∀f:T1 → T1 → T2 → T2.
+ (∀x,y,z.f x y z = f y x z) →
+ (∀acc:T2.∀l1,l2:ne_list T1.∀H:len_neList T1 l1 = len_neList T1 l2.
+ fold_right_neList2 T1 T2 f acc l1 l2 H = fold_right_neList2 T1 T2 f acc l2 l1 (symmetric_lennelist T1 l1 l2 H)).
+ #T1; #T2; #f; #H; #acc; #l1; #l2; #H1;
+ nrewrite > (symmetric_foldrightnelist2_aux T1 T2 f H acc l1 l2 H1 (symmetric_lennelist T1 l1 l2 H1));
napply refl_eq.
nqed.
-nlemma symmetric_bfoldrightnelist2
- : ∀T1:Type.∀f:T1 → T1 → bool.∀l1,l2:ne_list T1.
- (∀x,y.f x y = f y x) →
- bfold_right_neList2 T1 f l1 l2 = bfold_right_neList2 T1 f l2 l1.
- #T; #f; #l1;
+nlemma symmetric_bfoldrightnelist2 :
+∀T1:Type.∀f:T1 → T1 → bool.
+ (∀x,y.f x y = f y x) →
+ (∀l1,l2:ne_list T1.
+ bfold_right_neList2 T1 f l1 l2 = bfold_right_neList2 T1 f l2 l1).
+ #T; #f; #H; #l1;
nelim l1;
##[ ##1: #hh1; #l2; ncases l2;
- ##[ ##1: #hh2; #H; nnormalize; nrewrite > (H hh1 hh2); napply refl_eq
- ##| ##2: #hh2; #ll2; #H; nnormalize; napply refl_eq
- ##]
- ##| ##2: #hh1; #ll1; #H; #l2; ncases l2;
- ##[ ##1: #hh2; #H1; nnormalize; napply refl_eq
- ##| ##2: #hh2; #ll2; #H1; nnormalize;
- nrewrite > (H ll2 H1);
- nrewrite > (H1 hh1 hh2);
+ ##[ ##1: #hh2; nnormalize; nrewrite > (H hh1 hh2); napply refl_eq
+ ##| ##2: #hh2; #ll2; nnormalize; napply refl_eq
+ ##]
+ ##| ##2: #hh1; #ll1; #H1; #l2; ncases l2;
+ ##[ ##1: #hh2; nnormalize; napply refl_eq
+ ##| ##2: #hh2; #ll2; nnormalize;
+ nrewrite > (H1 ll2);
+ nrewrite > (H hh1 hh2);
napply refl_eq
##]
##]
nqed.
-nlemma bfoldrightnelist2_to_eq
- : ∀T1:Type.∀f:T1 → T1 → bool.∀l1,l2:ne_list T1.
- (∀x,y.(f x y = true → x = y)) →
- (bfold_right_neList2 T1 f l1 l2 = true → l1 = l2).
- #T; #f; #l1;
+nlemma bfoldrightnelist2_to_eq :
+∀T1:Type.∀f:T1 → T1 → bool.
+ (∀x,y.(f x y = true → x = y)) →
+ (∀l1,l2:ne_list T1.
+ (bfold_right_neList2 T1 f l1 l2 = true → l1 = l2)).
+ #T; #f; #H; #l1;
nelim l1;
##[ ##1: #hh1; #l2; ncases l2;
- ##[ ##1: #hh2; #H; #H1; nnormalize in H1:(%); nrewrite > (H hh1 hh2 H1); napply refl_eq
- ##| ##2: #hh2; #ll2; #H; nnormalize; #H1; ndestruct (*napply (bool_destruct … H1)*)
+ ##[ ##1: #hh2; #H1; nnormalize in H1:(%); nrewrite > (H hh1 hh2 H1); napply refl_eq
+ ##| ##2: #hh2; #ll2; nnormalize; #H1; ndestruct (*napply (bool_destruct … H1)*)
##]
- ##| ##2: #hh1; #ll1; #H; #l2; ncases l2;
- ##[ ##1: #hh2; #H1; nnormalize; #H2; ndestruct (*napply (bool_destruct … H2)*)
- ##| ##2: #hh2; #ll2; #H1; #H2;
+ ##| ##2: #hh1; #ll1; #H1; #l2; ncases l2;
+ ##[ ##1: #hh2; nnormalize; #H2; ndestruct (*napply (bool_destruct … H2)*)
+ ##| ##2: #hh2; #ll2; #H2;
nchange in H2:(%) with (((f hh1 hh2)⊗(bfold_right_neList2 T f ll1 ll2)) = true);
- nrewrite > (H1 hh1 hh2 (andb_true_true_l … H2));
- nrewrite > (H ll2 H1 (andb_true_true_r … H2));
+ nrewrite > (H hh1 hh2 (andb_true_true_l … H2));
+ nrewrite > (H1 ll2 (andb_true_true_r … H2));
napply refl_eq
##]
##]
nqed.
-nlemma eq_to_bfoldrightnelist2
- : ∀T1:Type.∀f:T1 → T1 → bool.∀l1,l2:ne_list T1.
- (∀x,y.(x = y → f x y = true)) →
- (l1 = l2 → bfold_right_neList2 T1 f l1 l2 = true).
- #T; #f; #l1;
+nlemma eq_to_bfoldrightnelist2 :
+∀T1:Type.∀f:T1 → T1 → bool.
+ (∀x,y.(x = y → f x y = true)) →
+ (∀l1,l2:ne_list T1.
+ (l1 = l2 → bfold_right_neList2 T1 f l1 l2 = true)).
+ #T; #f; #H; #l1;
nelim l1;
##[ ##1: #hh1; #l2; ncases l2;
- ##[ ##1: #hh2; #H; #H1; nnormalize;
+ ##[ ##1: #hh2; #H1; nnormalize;
nrewrite > (H hh1 hh2 (nelist_destruct_nil_nil … H1));
napply refl_eq
- ##| ##2: #hh2; #ll2; #H; #H1;
+ ##| ##2: #hh2; #ll2; #H1;
(* !!! ndestruct: assert false *)
nelim (nelist_destruct_nil_cons ???? H1)
##]
- ##| ##2: #hh1; #ll1; #H; #l2; ncases l2;
- ##[ ##1: #hh2; #H1; #H2;
+ ##| ##2: #hh1; #ll1; #H1; #l2; ncases l2;
+ ##[ ##1: #hh2; #H2;
(* !!! ndestruct: assert false *)
nelim (nelist_destruct_cons_nil ???? H2)
- ##| ##2: #hh2; #ll2; #H1; #H2; nnormalize;
+ ##| ##2: #hh2; #ll2; #H2; nnormalize;
nrewrite > (nelist_destruct_cons_cons_1 … H2);
- nrewrite > (H1 hh2 hh2 (refl_eq …));
+ nrewrite > (H hh2 hh2 (refl_eq …));
nnormalize;
- nrewrite > (H ll2 H1 (nelist_destruct_cons_cons_2 … H2));
+ nrewrite > (H1 ll2 (nelist_destruct_cons_cons_2 … H2));
napply refl_eq
##]
##]
nqed.
-nlemma bfoldrightnelist2_to_lennelist
- : ∀T.∀f:T → T → bool.∀l1,l2:ne_list T.bfold_right_neList2 T f l1 l2 = true → len_neList T l1 = len_neList T l2.
+nlemma bfoldrightnelist2_to_lennelist :
+∀T.∀f:T → T → bool.
+ (∀l1,l2:ne_list T.bfold_right_neList2 T f l1 l2 = true → len_neList T l1 = len_neList T l2).
#T; #f; #l1;
nelim l1;
##[ ##1: #hh1; #l2; ncases l2;
##]
nqed.
-nlemma decidable_nelist : ∀T.∀H:(Πx,y:T.decidable (x = y)).∀x,y:ne_list T.decidable (x = y).
+nlemma decidable_nelist :
+∀T.(∀x,y:T.decidable (x = y)) →
+ (∀x,y:ne_list T.decidable (x = y)).
#T; #H; #x; nelim x;
##[ ##1: #hh1; #y; ncases y;
##[ ##1: #hh2; nnormalize; napply (or2_elim (hh1 = hh2) (hh1 ≠ hh2) ? (H hh1 hh2));
##]
nqed.
-nlemma nbfoldrightnelist2_to_neq
- : ∀T1:Type.∀f:T1 → T1 → bool.∀l1,l2:ne_list T1.
- (∀x,y.(f x y = false → x ≠ y)) →
- (bfold_right_neList2 T1 f l1 l2 = false → l1 ≠ l2).
- #T; #f; #l1;
+nlemma nbfoldrightnelist2_to_neq :
+∀T1:Type.∀f:T1 → T1 → bool.
+ (∀x,y.(f x y = false → x ≠ y)) →
+ (∀l1,l2:ne_list T1.
+ (bfold_right_neList2 T1 f l1 l2 = false → l1 ≠ l2)).
+ #T; #f; #H; #l1;
nelim l1;
##[ ##1: #hh1; #l2; ncases l2;
- ##[ ##1: #hh2; #H; nnormalize; #H1; #H2; napply (H hh1 hh2 H1 (nelist_destruct_nil_nil T … H2))
- ##| ##2: #hh2; #ll2; #H; #H1; nnormalize; #H2;
+ ##[ ##1: #hh2; nnormalize; #H1; #H2; napply (H hh1 hh2 H1 (nelist_destruct_nil_nil T … H2))
+ ##| ##2: #hh2; #ll2; #H1; nnormalize; #H2;
(* !!! ndestruct: assert false *)
napply (nelist_destruct_nil_cons T … H2)
##]
- ##| ##2: #hh1; #ll1; #H; #l2; ncases l2;
- ##[ ##1: #hh2; #H1; #H2; nnormalize; #H3;
+ ##| ##2: #hh1; #ll1; #H1; #l2; ncases l2;
+ ##[ ##1: #hh2; #H2; nnormalize; #H3;
(* !!! ndestruct: assert false *)
napply (nelist_destruct_cons_nil T … H3)
- ##| ##2: #hh2; #ll2; #H1; #H2; nnormalize; #H3;
+ ##| ##2: #hh2; #ll2; #H2; nnormalize; #H3;
nchange in H2:(%) with (((f hh1 hh2)⊗(bfold_right_neList2 T f ll1 ll2)) = false);
- napply (H ll2 H1 ? (nelist_destruct_cons_cons_2 T … H3));
+ napply (H1 ll2 ? (nelist_destruct_cons_cons_2 T … H3));
napply (or2_elim ??? (andb_false2 … H2) );
##[ ##1: #H4; napply (absurd (hh1 = hh2) …);
##[ ##1: nrewrite > (nelist_destruct_cons_cons_1 T … H3); napply refl_eq
- ##| ##2: napply (H1 … H4)
+ ##| ##2: napply (H … H4)
##]
##| ##2: #H4; napply H4
##]
##]
nqed.
-nlemma nelist_destruct
- : ∀T.∀H:(Πx,y:T.decidable (x = y)).∀h1,h2:T.∀l1,l2:ne_list T.(h1§§l1) ≠ (h2§§l2) → h1 ≠ h2 ∨ l1 ≠ l2.
+nlemma nelist_destruct :
+∀T.(∀x,y:T.decidable (x = y)) →
+ (∀h1,h2:T.∀l1,l2:ne_list T.
+ (h1§§l1) ≠ (h2§§l2) → h1 ≠ h2 ∨ l1 ≠ l2).
#T; #H; #h1; #h2; #l1; nelim l1;
##[ ##1: #hh1; #l2; ncases l2;
##[ ##1: #hh2; #H1; napply (or2_elim (h1 = h2) (h1 ≠ h2) ? (H …) …);
##]
nqed.
-nlemma neq_to_nbfoldrightnelist2
- : ∀T:Type.∀f:T → T → bool.∀l1,l2:ne_list T.
- (∀x,y:T.decidable (x = y)) →
- (∀x,y.(x ≠ y → f x y = false)) →
- (l1 ≠ l2 → bfold_right_neList2 T f l1 l2 = false).
- #T; #f; #l1;
+nlemma neq_to_nbfoldrightnelist2 :
+∀T:Type.∀f:T → T → bool.
+ (∀x,y:T.decidable (x = y)) →
+ (∀x,y.(x ≠ y → f x y = false)) →
+ (∀l1,l2:ne_list T.
+ (l1 ≠ l2 → bfold_right_neList2 T f l1 l2 = false)).
+ #T; #f; #H; #H1; #l1;
nelim l1;
##[ ##1: #hh1; #l2; ncases l2;
- ##[ ##1: #hh2; #H; #H1; nnormalize; #H2; napply (H1 hh1 hh2 ?);
+ ##[ ##1: #hh2; nnormalize; #H2; napply (H1 hh1 hh2 ?);
nnormalize; #H3; nrewrite > H3 in H2:(%); #H2; napply (H2 (refl_eq …))
- ##| ##2: #hh2; #ll2; #H; nnormalize; #H1; #H2; napply refl_eq
+ ##| ##2: #hh2; #ll2; nnormalize; #H2; napply refl_eq
##]
- ##| ##2: #hh1; #ll1; #H; #l2; ncases l2;
- ##[ ##1: #hh2; #H1; #H2; nnormalize; #H3; napply refl_eq
- ##| ##2: #hh2; #ll2; #H1; #H2; #H3;
+ ##| ##2: #hh1; #ll1; #H2; #l2; ncases l2;
+ ##[ ##1: #hh2; nnormalize; #H3; napply refl_eq
+ ##| ##2: #hh2; #ll2; #H3;
nchange with (((f hh1 hh2)⊗(bfold_right_neList2 T f ll1 ll2)) = false);
- napply (or2_elim (hh1 ≠ hh2) (ll1 ≠ ll2) ? (nelist_destruct T H1 … H3) …);
- ##[ ##1: #H4; nrewrite > (H2 hh1 hh2 H4); nnormalize; napply refl_eq
- ##| ##2: #H4; nrewrite > (H ll2 H1 H2 H4);
+ napply (or2_elim (hh1 ≠ hh2) (ll1 ≠ ll2) ? (nelist_destruct T H … H3) …);
+ ##[ ##1: #H4; nrewrite > (H1 hh1 hh2 H4); nnormalize; napply refl_eq
+ ##| ##2: #H4; nrewrite > (H2 ll2 H4);
nrewrite > (symmetric_andbool (f hh1 hh2) false);
nnormalize; napply refl_eq
##]
alias num (instance 0) = "natural number".
*)
+nlet rec nat_it (T:Type) (f:T → T) (arg:T) (n:nat) on n ≝
+ match n with
+ [ O ⇒ arg
+ | S n' ⇒ nat_it T f (f arg) n'
+ ].
+
ndefinition nat1 ≝ S O.
ndefinition nat2 ≝ S nat1.
ndefinition nat3 ≝ S nat2.
ndefinition pred ≝ λn.match n with [ O ⇒ O | S n ⇒ n ].
-ndefinition nat128 ≝ nat64 * nat2.
-ndefinition nat256 ≝ nat128 * nat2.
-ndefinition nat512 ≝ nat256 * nat2.
-ndefinition nat1024 ≝ nat512 * nat2.
-ndefinition nat2048 ≝ nat1024 * nat2.
-ndefinition nat4096 ≝ nat2048 * nat2.
-ndefinition nat8192 ≝ nat4096 * nat2.
-ndefinition nat16384 ≝ nat8192 * nat2.
-ndefinition nat32768 ≝ nat16384 * nat2.
-ndefinition nat65536 ≝ nat32768 * nat2.
+ndefinition nat128 ≝ nat64 + nat64.
+ndefinition nat256 ≝ nat128 + nat128.
+ndefinition nat512 ≝ nat256 + nat256.
+ndefinition nat1024 ≝ nat512 + nat512.
+ndefinition nat2048 ≝ nat1024 + nat1024.
+ndefinition nat4096 ≝ nat2048 + nat2048.
+ndefinition nat8192 ≝ nat4096 + nat4096.
+ndefinition nat16384 ≝ nat8192 + nat8192.
+ndefinition nat32768 ≝ nat16384 + nat16384.
+ndefinition nat65536 ≝ nat32768 + nat32768.
ndefinition nat_of_w16 : word16 → nat ≝ λn:word16.nat_of_w16_aux n (w16_to_recw16 n).
-ndefinition nat_of_w32 : word32 → nat ≝ λn:word32.(nat65536 * (nat_of_w16 (w32h n))) + (nat_of_w16 (w32l n)).
+ndefinition nat_of_w32 : word32 → nat ≝ λn:word32.(nat65536 * (nat_of_w16 (cnH ? n))) + (nat_of_w16 (cnL ? n)).
nqed.
nlemma symmetric_eqoption :
-∀T:Type.∀f:T → T → bool.∀op1,op2:option T.
+∀T:Type.∀f:T → T → bool.
(symmetricT T bool f) →
- (eq_option T f op1 op2 = eq_option T f op2 op1).
- #T; #f; #op1; #op2; #H;
- nelim op1;
- nelim op2;
+ (∀op1,op2:option T.
+ (eq_option T f op1 op2 = eq_option T f op2 op1)).
+ #T; #f; #H;
+ #op1; #op2; nelim op1; nelim op2;
nnormalize;
##[ ##1: napply refl_eq
##| ##2,3: #H; napply refl_eq
nqed.
nlemma eq_to_eqoption :
-∀T.∀f:T → T → bool.∀op1,op2:option T.
+∀T.∀f:T → T → bool.
(∀x1,x2:T.x1 = x2 → f x1 x2 = true) →
- (op1 = op2 → eq_option T f op1 op2 = true).
- #T; #f; #op1; #op2; #H;
- nelim op1;
- nelim op2;
+ (∀op1,op2:option T.
+ (op1 = op2 → eq_option T f op1 op2 = true)).
+ #T; #f; #H;
+ #op1; #op2; nelim op1; nelim op2;
nnormalize;
##[ ##1: #H1; napply refl_eq
##| ##2: #a; #H1;
nqed.
nlemma eqoption_to_eq :
-∀T.∀f:T → T → bool.∀op1,op2:option T.
+∀T.∀f:T → T → bool.
(∀x1,x2:T.f x1 x2 = true → x1 = x2) →
- (eq_option T f op1 op2 = true → op1 = op2).
- #T; #f; #op1; #op2; #H;
- nelim op1;
- nelim op2;
+ (∀op1,op2:option T.
+ (eq_option T f op1 op2 = true → op1 = op2)).
+ #T; #f; #H;
+ #op1; #op2; nelim op1; nelim op2;
nnormalize;
##[ ##1: #H1; napply refl_eq
##| ##2,3: #a; #H1; ndestruct (*napply (bool_destruct … H1)*)
##]
nqed.
-nlemma decidable_option : ∀T.∀H:(Πx,y:T.decidable (x = y)).∀x,y:option T.decidable (x = y).
+nlemma decidable_option :
+∀T.(Πx,y:T.decidable (x = y)) →
+ (∀x,y:option T.decidable (x = y)).
#T; #H; #x; nelim x;
##[ ##1: #y; ncases y;
##[ ##1: nnormalize; napply (or2_intro1 (? = ?) (? ≠ ?) (refl_eq …))
nqed.
nlemma neq_to_neqoption :
-∀T.∀f:T → T → bool.∀op1,op2:option T.
+∀T.∀f:T → T → bool.
(∀x1,x2:T.x1 ≠ x2 → f x1 x2 = false) →
- (op1 ≠ op2 → eq_option T f op1 op2 = false).
- #T; #f; #op1; nelim op1;
+ (∀op1,op2:option T.
+ (op1 ≠ op2 → eq_option T f op1 op2 = false)).
+ #T; #f; #H; #op1; nelim op1;
##[ ##1: #op2; ncases op2;
- ##[ ##1: nnormalize; #H; #H1; nelim (H1 (refl_eq …))
- ##| ##2: #yy; #H; nnormalize; #H1; napply refl_eq
+ ##[ ##1: nnormalize; #H1; nelim (H1 (refl_eq …))
+ ##| ##2: #yy; nnormalize; #H1; napply refl_eq
##]
##| ##2: #xx; #op2; ncases op2;
- ##[ ##1: #H; nnormalize; #H1; napply refl_eq
- ##| ##2: #yy; #H; nnormalize; #H1; napply (H xx yy …);
+ ##[ ##1: nnormalize; #H1; napply refl_eq
+ ##| ##2: #yy; nnormalize; #H1; napply (H xx yy …);
nnormalize; #H2; nrewrite > H2 in H1:(%); #H1;
napply (H1 (refl_eq …))
##]
nqed.
nlemma neqoption_to_neq :
-∀T.∀f:T → T → bool.∀op1,op2:option T.
+∀T.∀f:T → T → bool.
(∀x1,x2:T.f x1 x2 = false → x1 ≠ x2) →
- (eq_option T f op1 op2 = false → op1 ≠ op2).
- #T; #f; #op1; nelim op1;
+ (∀op1,op2:option T.
+ (eq_option T f op1 op2 = false → op1 ≠ op2)).
+ #T; #f; #H; #op1; nelim op1;
##[ ##1: #op2; ncases op2;
- ##[ ##1: nnormalize; #H; #H1;
+ ##[ ##1: nnormalize; #H1;
ndestruct (*napply (bool_destruct … H1)*)
- ##| ##2: #yy; #H; nnormalize; #H1; #H2;
+ ##| ##2: #yy; nnormalize; #H1; #H2;
(* !!! ndestruct: assert false *)
napply (option_destruct_none_some T … H2)
##]
##| ##2: #xx; #op2; ncases op2;
- ##[ ##1: nnormalize; #H; #H1; #H2;
+ ##[ ##1: nnormalize; #H1; #H2;
(* !!! ndestruct: assert false *)
napply (option_destruct_some_none T … H2)
- ##| ##2: #yy; #H; nnormalize; #H1; #H2; napply (H xx yy H1 ?);
+ ##| ##2: #yy; nnormalize; #H1; #H2; napply (H xx yy H1 ?);
napply (option_destruct_some_some T … H2)
##]
##]
nlemma symmetric_eqpair :
∀T1,T2:Type.
∀f1:T1 → T1 → bool.∀f2:T2 → T2 → bool.
-∀p1,p2:ProdT T1 T2.
(symmetricT T1 bool f1) →
(symmetricT T2 bool f2) →
- (eq_pair T1 T2 f1 f2 p1 p2 = eq_pair T1 T2 f1 f2 p1 p2).
- #T1; #T2; #f1; #f2; #p1; #p2; #H; #H1;
- nelim p1;
- #x1; #y1;
- nelim p2;
- #x2; #y2;
+ (∀p1,p2:ProdT T1 T2.
+ (eq_pair T1 T2 f1 f2 p1 p2 = eq_pair T1 T2 f1 f2 p1 p2)).
+ #T1; #T2; #f1; #f2; #H; #H1;
+ #p1; nelim p1; #x1; #y1;
+ #p2; nelim p2; #x2; #y2;
nnormalize;
nrewrite > (H x1 x2);
ncases (f1 x2 x1);
nlemma eq_to_eqpair :
∀T1,T2.
∀f1:T1 → T1 → bool.∀f2:T2 → T2 → bool.
-∀p1,p2:ProdT T1 T2.
(∀x,y:T1.x = y → f1 x y = true) →
(∀x,y:T2.x = y → f2 x y = true) →
- (p1 = p2 → eq_pair T1 T2 f1 f2 p1 p2 = true).
- #T1; #T2; #f1; #f2; #p1; #p2; #H1; #H2;
- nelim p1;
- #x1; #y1;
- nelim p2;
- #x2; #y2; #H;
+ (∀p1,p2:ProdT T1 T2.
+ (p1 = p2 → eq_pair T1 T2 f1 f2 p1 p2 = true)).
+ #T1; #T2; #f1; #f2; #H1; #H2;
+ #p1; nelim p1; #x1; #y1;
+ #p2; nelim p2; #x2; #y2; #H;
nnormalize;
nrewrite > (H1 … (pair_destruct_1 … H));
nnormalize;
nlemma eqpair_to_eq :
∀T1,T2.
∀f1:T1 → T1 → bool.∀f2:T2 → T2 → bool.
-∀p1,p2:ProdT T1 T2.
(∀x,y:T1.f1 x y = true → x = y) →
(∀x,y:T2.f2 x y = true → x = y) →
- (eq_pair T1 T2 f1 f2 p1 p2 = true → p1 = p2).
- #T1; #T2; #f1; #f2; #p1; #p2; #H1; #H2;
- nelim p1;
- #x1; #y1;
- nelim p2;
- #x2; #y2; #H;
+ (∀p1,p2:ProdT T1 T2.
+ (eq_pair T1 T2 f1 f2 p1 p2 = true → p1 = p2)).
+ #T1; #T2; #f1; #f2; #H1; #H2;
+ #p1; nelim p1; #x1; #y1;
+ #p2; nelim p2; #x2; #y2; #H;
nnormalize in H:(%);
nletin K ≝ (H1 x1 x2);
ncases (f1 x1 x2) in H:(%) K:(%);
napply refl_eq.
nqed.
-nlemma decidable_pair
- : ∀T1,T2.
- (∀x,y:T1.decidable (x = y)) →
- (∀x,y:T2.decidable (x = y)) →
- ∀x,y:ProdT T1 T2.decidable (x = y).
+nlemma decidable_pair :
+∀T1,T2.
+ (∀x,y:T1.decidable (x = y)) →
+ (∀x,y:T2.decidable (x = y)) →
+ (∀x,y:ProdT T1 T2.decidable (x = y)).
#T1; #T2; #H; #H1;
#x; nelim x; #xx1; #xx2;
#y; nelim y; #yy1; #yy2;
nlemma neqpair_to_neq :
∀T1,T2.
∀f1:T1 → T1 → bool.∀f2:T2 → T2 → bool.
-∀p1,p2:ProdT T1 T2.
(∀x,y:T1.f1 x y = false → x ≠ y) →
(∀x,y:T2.f2 x y = false → x ≠ y) →
- (eq_pair T1 T2 f1 f2 p1 p2 = false → p1 ≠ p2).
- #T1; #T2; #f1; #f2; #p1; #p2; #H1; #H2;
- nelim p1;
- #x1; #y1;
- nelim p2;
- #x2; #y2;
+ (∀p1,p2:ProdT T1 T2.
+ (eq_pair T1 T2 f1 f2 p1 p2 = false → p1 ≠ p2)).
+ #T1; #T2; #f1; #f2; #H1; #H2;
+ #p1; nelim p1; #x1; #y1;
+ #p2; nelim p2; #x2; #y2;
nchange with ((((f1 x1 x2) ⊗ (f2 y1 y2)) = false) → ?); #H;
nnormalize; #H3;
napply (or2_elim ((f1 x1 x2) = false) ((f2 y1 y2) = false) ? (andb_false2 … H) ?);
##]
nqed.
-nlemma pair_destruct
- : ∀T1,T2.
- (∀x,y:T1.decidable (x = y)) →
- (∀x,y:T2.decidable (x = y)) →
- ∀x1,x2:T1.∀y1,y2:T2.(pair T1 T2 x1 y1) ≠ (pair T1 T2 x2 y2) → x1 ≠ x2 ∨ y1 ≠ y2.
+nlemma pair_destruct :
+∀T1,T2.
+ (∀x,y:T1.decidable (x = y)) →
+ (∀x,y:T2.decidable (x = y)) →
+ (∀x1,x2:T1.∀y1,y2:T2.
+ (pair T1 T2 x1 y1) ≠ (pair T1 T2 x2 y2) → x1 ≠ x2 ∨ y1 ≠ y2).
#T1; #T2; #H1; #H2; #x1; #x2; #y1; #y2;
nnormalize; #H;
napply (or2_elim (x1 = x2) (x1 ≠ x2) ? (H1 x1 x2) ?);
nlemma neq_to_neqpair :
∀T1,T2.
∀f1:T1 → T1 → bool.∀f2:T2 → T2 → bool.
-∀p1,p2:ProdT T1 T2.
(∀x,y:T1.decidable (x = y)) →
(∀x,y:T2.decidable (x = y)) →
(∀x,y:T1.x ≠ y → f1 x y = false) →
(∀x,y:T2.x ≠ y → f2 x y = false) →
- (p1 ≠ p2 → eq_pair T1 T2 f1 f2 p1 p2 = false).
- #T1; #T2; #f1; #f2; #p1; #p2; #H1; #H2; #H3; #H4;
- nelim p1;
- #x1; #y1;
- nelim p2;
- #x2; #y2; #H;
+ (∀p1,p2:ProdT T1 T2.
+ (p1 ≠ p2 → eq_pair T1 T2 f1 f2 p1 p2 = false)).
+ #T1; #T2; #f1; #f2; #H1; #H2; #H3; #H4;
+ #p1; nelim p1; #x1; #y1;
+ #p2; nelim p2; #x2; #y2; #H;
nchange with (((f1 x1 x2) ⊗ (f2 y1 y2)) = false);
napply (or2_elim (x1 ≠ x2) (y1 ≠ y2) ? (pair_destruct T1 T2 H1 H2 … H) ?);
##[ ##2: #H5; nrewrite > (H4 … H5); nrewrite > (andb_false2_2 (f1 x1 x2)); napply refl_eq
nlemma symmetric_eqtriple :
∀T1,T2,T3:Type.
∀f1:T1 → T1 → bool.∀f2:T2 → T2 → bool.∀f3:T3 → T3 → bool.
-∀p1,p2:Prod3T T1 T2 T3.
(symmetricT T1 bool f1) →
(symmetricT T2 bool f2) →
(symmetricT T3 bool f3) →
- (eq_triple T1 T2 T3 f1 f2 f3 p1 p2 = eq_triple T1 T2 T3 f1 f2 f3 p2 p1).
- #T1; #T2; #T3; #f1; #f2; #f3; #p1; #p2; #H; #H1; #H2;
- nelim p1;
- #x1; #y1; #z1;
- nelim p2;
- #x2; #y2; #z2;
+ (∀p1,p2:Prod3T T1 T2 T3.
+ (eq_triple T1 T2 T3 f1 f2 f3 p1 p2 = eq_triple T1 T2 T3 f1 f2 f3 p2 p1)).
+ #T1; #T2; #T3; #f1; #f2; #f3; #H; #H1; #H2;
+ #p1; nelim p1; #x1; #y1; #z1;
+ #p2; nelim p2; #x2; #y2; #z2;
nnormalize;
nrewrite > (H x1 x2);
ncases (f1 x2 x1);
nlemma eq_to_eqtriple :
∀T1,T2,T3.
∀f1:T1 → T1 → bool.∀f2:T2 → T2 → bool.∀f3:T3 → T3 → bool.
-∀p1,p2:Prod3T T1 T2 T3.
(∀x1,x2:T1.x1 = x2 → f1 x1 x2 = true) →
(∀y1,y2:T2.y1 = y2 → f2 y1 y2 = true) →
(∀z1,z2:T3.z1 = z2 → f3 z1 z2 = true) →
- (p1 = p2 → eq_triple T1 T2 T3 f1 f2 f3 p1 p2 = true).
- #T1; #T2; #T3; #f1; #f2; #f3; #p1; #p2; #H1; #H2; #H3;
- nelim p1;
- #x1; #y1; #z1;
- nelim p2;
- #x2; #y2; #z2; #H;
+ (∀p1,p2:Prod3T T1 T2 T3.
+ (p1 = p2 → eq_triple T1 T2 T3 f1 f2 f3 p1 p2 = true)).
+ #T1; #T2; #T3; #f1; #f2; #f3; #H1; #H2; #H3;
+ #p1; nelim p1; #x1; #y1; #z1;
+ #p2; nelim p2; #x2; #y2; #z2; #H;
nnormalize;
nrewrite > (H1 … (triple_destruct_1 … H));
nnormalize;
nlemma eqtriple_to_eq :
∀T1,T2,T3.
∀f1:T1 → T1 → bool.∀f2:T2 → T2 → bool.∀f3:T3 → T3 → bool.
-∀p1,p2:Prod3T T1 T2 T3.
(∀x1,x2:T1.f1 x1 x2 = true → x1 = x2) →
(∀y1,y2:T2.f2 y1 y2 = true → y1 = y2) →
(∀z1,z2:T3.f3 z1 z2 = true → z1 = z2) →
- (eq_triple T1 T2 T3 f1 f2 f3 p1 p2 = true → p1 = p2).
- #T1; #T2; #T3; #f1; #f2; #f3; #p1; #p2; #H1; #H2; #H3;
- nelim p1;
- #x1; #y1; #z1;
- nelim p2;
- #x2; #y2; #z2; #H;
+ (∀p1,p2:Prod3T T1 T2 T3.
+ (eq_triple T1 T2 T3 f1 f2 f3 p1 p2 = true → p1 = p2)).
+ #T1; #T2; #T3; #f1; #f2; #f3; #H1; #H2; #H3;
+ #p1; nelim p1; #x1; #y1; #z1;
+ #p2; nelim p2; #x2; #y2; #z2; #H;
nnormalize in H:(%);
nletin K ≝ (H1 x1 x2);
ncases (f1 x1 x2) in H:(%) K:(%);
napply refl_eq.
nqed.
-nlemma decidable_triple
- : ∀T1,T2,T3.
- (∀x,y:T1.decidable (x = y)) →
- (∀x,y:T2.decidable (x = y)) →
- (∀x,y:T3.decidable (x = y)) →
- ∀x,y:Prod3T T1 T2 T3.decidable (x = y).
+nlemma decidable_triple :
+∀T1,T2,T3.
+ (∀x,y:T1.decidable (x = y)) →
+ (∀x,y:T2.decidable (x = y)) →
+ (∀x,y:T3.decidable (x = y)) →
+ (∀x,y:Prod3T T1 T2 T3.decidable (x = y)).
#T1; #T2; #T3; #H; #H1; #H2;
#x; nelim x; #xx1; #xx2; #xx3;
#y; nelim y; #yy1; #yy2; #yy3;
nlemma neqtriple_to_neq :
∀T1,T2,T3.
∀f1:T1 → T1 → bool.∀f2:T2 → T2 → bool.∀f3:T3 → T3 → bool.
-∀p1,p2:Prod3T T1 T2 T3.
(∀x,y:T1.f1 x y = false → x ≠ y) →
(∀x,y:T2.f2 x y = false → x ≠ y) →
(∀x,y:T3.f3 x y = false → x ≠ y) →
- (eq_triple T1 T2 T3 f1 f2 f3 p1 p2 = false → p1 ≠ p2).
- #T1; #T2; #T3; #f1; #f2; #f3; #p1; #p2; #H1; #H2; #H3;
- nelim p1;
- #x1; #y1; #z1;
- nelim p2;
- #x2; #y2; #z2;
+ (∀p1,p2:Prod3T T1 T2 T3.
+ (eq_triple T1 T2 T3 f1 f2 f3 p1 p2 = false → p1 ≠ p2)).
+ #T1; #T2; #T3; #f1; #f2; #f3; #H1; #H2; #H3;
+ #p1; nelim p1; #x1; #y1; #z1;
+ #p2; nelim p2; #x2; #y2; #z2;
nchange with ((((f1 x1 x2) ⊗ (f2 y1 y2) ⊗ (f3 z1 z2)) = false) → ?); #H;
nnormalize; #H4;
napply (or3_elim ((f1 x1 x2) = false) ((f2 y1 y2) = false) ((f3 z1 z2) = false) ? (andb_false3 … H) ?);
##]
nqed.
-nlemma triple_destruct
- : ∀T1,T2,T3.
- (∀x,y:T1.decidable (x = y)) →
- (∀x,y:T2.decidable (x = y)) →
- (∀x,y:T3.decidable (x = y)) →
- ∀x1,x2:T1.∀y1,y2:T2.∀z1,z2:T3.(triple T1 T2 T3 x1 y1 z1) ≠ (triple T1 T2 T3 x2 y2 z2) →
- Or3 (x1 ≠ x2) (y1 ≠ y2) (z1 ≠ z2).
+nlemma triple_destruct :
+∀T1,T2,T3.
+ (∀x,y:T1.decidable (x = y)) →
+ (∀x,y:T2.decidable (x = y)) →
+ (∀x,y:T3.decidable (x = y)) →
+ (∀x1,x2:T1.∀y1,y2:T2.∀z1,z2:T3.
+ (triple T1 T2 T3 x1 y1 z1) ≠ (triple T1 T2 T3 x2 y2 z2) →
+ Or3 (x1 ≠ x2) (y1 ≠ y2) (z1 ≠ z2)).
#T1; #T2; #T3; #H1; #H2; #H3; #x1; #x2; #y1; #y2; #z1; #z2;
nnormalize; #H;
napply (or2_elim (x1 = x2) (x1 ≠ x2) ? (H1 x1 x2) ?);
nlemma neq_to_neqtriple :
∀T1,T2,T3.
∀f1:T1 → T1 → bool.∀f2:T2 → T2 → bool.∀f3:T3 → T3 → bool.
-∀p1,p2:Prod3T T1 T2 T3.
(∀x,y:T1.decidable (x = y)) →
(∀x,y:T2.decidable (x = y)) →
(∀x,y:T3.decidable (x = y)) →
(∀x,y:T1.x ≠ y → f1 x y = false) →
(∀x,y:T2.x ≠ y → f2 x y = false) →
(∀x,y:T3.x ≠ y → f3 x y = false) →
- (p1 ≠ p2 → eq_triple T1 T2 T3 f1 f2 f3 p1 p2 = false).
- #T1; #T2; #T3; #f1; #f2; #f3; #p1; #p2; #H1; #H2; #H3; #H4; #H5; #H6;
- nelim p1;
- #x1; #y1; #z1;
- nelim p2;
- #x2; #y2; #z2; #H;
+ (∀p1,p2:Prod3T T1 T2 T3.
+ (p1 ≠ p2 → eq_triple T1 T2 T3 f1 f2 f3 p1 p2 = false)).
+ #T1; #T2; #T3; #f1; #f2; #f3; #H1; #H2; #H3; #H4; #H5; #H6;
+ #p1; nelim p1; #x1; #y1; #z1;
+ #p2; nelim p2; #x2; #y2; #z2; #H;
nchange with (((f1 x1 x2) ⊗ (f2 y1 y2) ⊗ (f3 z1 z2)) = false);
napply (or3_elim (x1 ≠ x2) (y1 ≠ y2) (z1 ≠ z2) ? (triple_destruct T1 T2 T3 H1 H2 H3 … H) ?);
##[ ##1: #H7; nrewrite > (H4 … H7); nrewrite > (andb_false3_1 (f2 y1 y2) (f3 z1 z2)); napply refl_eq
nlemma symmetric_eqquadruple :
∀T1,T2,T3,T4:Type.
∀f1:T1 → T1 → bool.∀f2:T2 → T2 → bool.∀f3:T3 → T3 → bool.∀f4:T4 → T4 → bool.
-∀p1,p2:Prod4T T1 T2 T3 T4.
(symmetricT T1 bool f1) →
(symmetricT T2 bool f2) →
(symmetricT T3 bool f3) →
(symmetricT T4 bool f4) →
- (eq_quadruple T1 T2 T3 T4 f1 f2 f3 f4 p1 p2 = eq_quadruple T1 T2 T3 T4 f1 f2 f3 f4 p2 p1).
- #T1; #T2; #T3; #T4; #f1; #f2; #f3; #f4; #p1; #p2; #H; #H1; #H2; #H3;
- nelim p1;
- #x1; #y1; #z1; #v1;
- nelim p2;
- #x2; #y2; #z2; #v2;
+ (∀p1,p2:Prod4T T1 T2 T3 T4.
+ (eq_quadruple T1 T2 T3 T4 f1 f2 f3 f4 p1 p2 = eq_quadruple T1 T2 T3 T4 f1 f2 f3 f4 p2 p1)).
+ #T1; #T2; #T3; #T4; #f1; #f2; #f3; #f4; #H; #H1; #H2; #H3;
+ #p1; nelim p1; #x1; #y1; #z1; #v1;
+ #p2; nelim p2; #x2; #y2; #z2; #v2;
nnormalize;
nrewrite > (H x1 x2);
ncases (f1 x2 x1);
nlemma eq_to_eqquadruple :
∀T1,T2,T3,T4.
∀f1:T1 → T1 → bool.∀f2:T2 → T2 → bool.∀f3:T3 → T3 → bool.∀f4:T4 → T4 → bool.
-∀p1,p2:Prod4T T1 T2 T3 T4.
(∀x,y:T1.x = y → f1 x y = true) →
(∀x,y:T2.x = y → f2 x y = true) →
(∀x,y:T3.x = y → f3 x y = true) →
(∀x,y:T4.x = y → f4 x y = true) →
- (p1 = p2 → eq_quadruple T1 T2 T3 T4 f1 f2 f3 f4 p1 p2 = true).
- #T1; #T2; #T3; #T4; #f1; #f2; #f3; #f4; #p1; #p2; #H1; #H2; #H3; #H4;
- nelim p1;
- #x1; #y1; #z1; #v1;
- nelim p2;
- #x2; #y2; #z2; #v2; #H;
+ (∀p1,p2:Prod4T T1 T2 T3 T4.
+ (p1 = p2 → eq_quadruple T1 T2 T3 T4 f1 f2 f3 f4 p1 p2 = true)).
+ #T1; #T2; #T3; #T4; #f1; #f2; #f3; #f4; #H1; #H2; #H3; #H4;
+ #p1; nelim p1; #x1; #y1; #z1; #v1;
+ #p2; nelim p2; #x2; #y2; #z2; #v2; #H;
nnormalize;
nrewrite > (H1 … (quadruple_destruct_1 … H));
nnormalize;
nlemma eqquadruple_to_eq :
∀T1,T2,T3,T4.
∀f1:T1 → T1 → bool.∀f2:T2 → T2 → bool.∀f3:T3 → T3 → bool.∀f4:T4 → T4 → bool.
-∀p1,p2:Prod4T T1 T2 T3 T4.
(∀x,y:T1.f1 x y = true → x = y) →
(∀x,y:T2.f2 x y = true → x = y) →
(∀x,y:T3.f3 x y = true → x = y) →
(∀x,y:T4.f4 x y = true → x = y) →
- (eq_quadruple T1 T2 T3 T4 f1 f2 f3 f4 p1 p2 = true → p1 = p2).
- #T1; #T2; #T3; #T4; #f1; #f2; #f3; #f4; #p1; #p2; #H1; #H2; #H3; #H4;
- nelim p1;
- #x1; #y1; #z1; #v1;
- nelim p2;
- #x2; #y2; #z2; #v2; #H;
+ (∀p1,p2:Prod4T T1 T2 T3 T4.
+ (eq_quadruple T1 T2 T3 T4 f1 f2 f3 f4 p1 p2 = true → p1 = p2)).
+ #T1; #T2; #T3; #T4; #f1; #f2; #f3; #f4; #H1; #H2; #H3; #H4;
+ #p1; nelim p1; #x1; #y1; #z1; #v1;
+ #p2; nelim p2; #x2; #y2; #z2; #v2; #H;
nnormalize in H:(%);
nletin K ≝ (H1 x1 x2);
ncases (f1 x1 x2) in H:(%) K:(%);
napply refl_eq.
nqed.
-nlemma decidable_quadruple
- : ∀T1,T2,T3,T4.
- (∀x,y:T1.decidable (x = y)) →
- (∀x,y:T2.decidable (x = y)) →
- (∀x,y:T3.decidable (x = y)) →
- (∀x,y:T4.decidable (x = y)) →
- ∀x,y:Prod4T T1 T2 T3 T4.decidable (x = y).
+nlemma decidable_quadruple :
+∀T1,T2,T3,T4.
+ (∀x,y:T1.decidable (x = y)) →
+ (∀x,y:T2.decidable (x = y)) →
+ (∀x,y:T3.decidable (x = y)) →
+ (∀x,y:T4.decidable (x = y)) →
+ (∀x,y:Prod4T T1 T2 T3 T4.decidable (x = y)).
#T1; #T2; #T3; #T4; #H; #H1; #H2; #H3;
#x; nelim x; #xx1; #xx2; #xx3; #xx4;
#y; nelim y; #yy1; #yy2; #yy3; #yy4;
nlemma neqquadruple_to_neq :
∀T1,T2,T3,T4.
∀f1:T1 → T1 → bool.∀f2:T2 → T2 → bool.∀f3:T3 → T3 → bool.∀f4:T4 → T4 → bool.
-∀p1,p2:Prod4T T1 T2 T3 T4.
(∀x,y:T1.f1 x y = false → x ≠ y) →
(∀x,y:T2.f2 x y = false → x ≠ y) →
(∀x,y:T3.f3 x y = false → x ≠ y) →
(∀x,y:T4.f4 x y = false → x ≠ y) →
- (eq_quadruple T1 T2 T3 T4 f1 f2 f3 f4 p1 p2 = false → p1 ≠ p2).
- #T1; #T2; #T3; #T4; #f1; #f2; #f3; #f4; #p1; #p2; #H1; #H2; #H3; #H4;
- nelim p1;
- #x1; #y1; #z1; #v1;
- nelim p2;
- #x2; #y2; #z2; #v2;
+ (∀p1,p2:Prod4T T1 T2 T3 T4.
+ (eq_quadruple T1 T2 T3 T4 f1 f2 f3 f4 p1 p2 = false → p1 ≠ p2)).
+ #T1; #T2; #T3; #T4; #f1; #f2; #f3; #f4; #H1; #H2; #H3; #H4;
+ #p1; nelim p1; #x1; #y1; #z1; #v1;
+ #p2; nelim p2; #x2; #y2; #z2; #v2;
nchange with ((((f1 x1 x2) ⊗ (f2 y1 y2) ⊗ (f3 z1 z2) ⊗ (f4 v1 v2)) = false) → ?); #H;
nnormalize; #H5;
napply (or4_elim ((f1 x1 x2) = false) ((f2 y1 y2) = false) ((f3 z1 z2) = false) ((f4 v1 v2) = false) ? (andb_false4 … H) ?);
##]
nqed.
-nlemma quadruple_destruct
- : ∀T1,T2,T3,T4.
- (∀x,y:T1.decidable (x = y)) →
- (∀x,y:T2.decidable (x = y)) →
- (∀x,y:T3.decidable (x = y)) →
- (∀x,y:T4.decidable (x = y)) →
- ∀x1,x2:T1.∀y1,y2:T2.∀z1,z2:T3.∀v1,v2:T4.
- (quadruple T1 T2 T3 T4 x1 y1 z1 v1) ≠ (quadruple T1 T2 T3 T4 x2 y2 z2 v2) →
- Or4 (x1 ≠ x2) (y1 ≠ y2) (z1 ≠ z2) (v1 ≠ v2).
+nlemma quadruple_destruct :
+∀T1,T2,T3,T4.
+ (∀x,y:T1.decidable (x = y)) →
+ (∀x,y:T2.decidable (x = y)) →
+ (∀x,y:T3.decidable (x = y)) →
+ (∀x,y:T4.decidable (x = y)) →
+ (∀x1,x2:T1.∀y1,y2:T2.∀z1,z2:T3.∀v1,v2:T4.
+ (quadruple T1 T2 T3 T4 x1 y1 z1 v1) ≠ (quadruple T1 T2 T3 T4 x2 y2 z2 v2) →
+ Or4 (x1 ≠ x2) (y1 ≠ y2) (z1 ≠ z2) (v1 ≠ v2)).
#T1; #T2; #T3; #T4; #H1; #H2; #H3; #H4;
#x1; #x2; #y1; #y2; #z1; #z2; #v1; #v2;
nnormalize; #H;
nlemma neq_to_neqquadruple :
∀T1,T2,T3,T4.
∀f1:T1 → T1 → bool.∀f2:T2 → T2 → bool.∀f3:T3 → T3 → bool.∀f4:T4 → T4 → bool.
-∀p1,p2:Prod4T T1 T2 T3 T4.
(∀x,y:T1.decidable (x = y)) →
(∀x,y:T2.decidable (x = y)) →
(∀x,y:T3.decidable (x = y)) →
(∀x,y:T2.x ≠ y → f2 x y = false) →
(∀x,y:T3.x ≠ y → f3 x y = false) →
(∀x,y:T4.x ≠ y → f4 x y = false) →
- (p1 ≠ p2 → eq_quadruple T1 T2 T3 T4 f1 f2 f3 f4 p1 p2 = false).
- #T1; #T2; #T3; #T4; #f1; #f2; #f3; #f4; #p1; #p2; #H1; #H2; #H3; #H4; #H5; #H6; #H7; #H8;
- nelim p1;
- #x1; #y1; #z1; #v1;
- nelim p2;
- #x2; #y2; #z2; #v2; #H;
+ (∀p1,p2:Prod4T T1 T2 T3 T4.
+ (p1 ≠ p2 → eq_quadruple T1 T2 T3 T4 f1 f2 f3 f4 p1 p2 = false)).
+ #T1; #T2; #T3; #T4; #f1; #f2; #f3; #f4; #H1; #H2; #H3; #H4; #H5; #H6; #H7; #H8;
+ #p1; nelim p1; #x1; #y1; #z1; #v1;
+ #p2; nelim p2; #x2; #y2; #z2; #v2; #H;
nchange with (((f1 x1 x2) ⊗ (f2 y1 y2) ⊗ (f3 z1 z2) ⊗ (f4 v1 v2)) = false);
napply (or4_elim (x1 ≠ x2) (y1 ≠ y2) (z1 ≠ z2) (v1 ≠ v2) ? (quadruple_destruct T1 T2 T3 T4 H1 H2 H3 H4 … H) ?);
##[ ##1: #H9; nrewrite > (H5 … H9); nrewrite > (andb_false4_1 (f2 y1 y2) (f3 z1 z2) (f4 v1 v2)); napply refl_eq
nlemma symmetric_eqquintuple :
∀T1,T2,T3,T4,T5:Type.
∀f1:T1 → T1 → bool.∀f2:T2 → T2 → bool.∀f3:T3 → T3 → bool.∀f4:T4 → T4 → bool.∀f5:T5 → T5 → bool.
-∀p1,p2:Prod5T T1 T2 T3 T4 T5.
(symmetricT T1 bool f1) →
(symmetricT T2 bool f2) →
(symmetricT T3 bool f3) →
(symmetricT T4 bool f4) →
(symmetricT T5 bool f5) →
- (eq_quintuple T1 T2 T3 T4 T5 f1 f2 f3 f4 f5 p1 p2 = eq_quintuple T1 T2 T3 T4 T5 f1 f2 f3 f4 f5 p2 p1).
- #T1; #T2; #T3; #T4; #T5; #f1; #f2; #f3; #f4; #f5; #p1; #p2; #H; #H1; #H2; #H3; #H4;
- nelim p1;
- #x1; #y1; #z1; #v1; #w1;
- nelim p2;
- #x2; #y2; #z2; #v2; #w2;
+ (∀p1,p2:Prod5T T1 T2 T3 T4 T5.
+ (eq_quintuple T1 T2 T3 T4 T5 f1 f2 f3 f4 f5 p1 p2 = eq_quintuple T1 T2 T3 T4 T5 f1 f2 f3 f4 f5 p2 p1)).
+ #T1; #T2; #T3; #T4; #T5; #f1; #f2; #f3; #f4; #f5; #H; #H1; #H2; #H3; #H4;
+ #p1; nelim p1; #x1; #y1; #z1; #v1; #w1;
+ #p2; nelim p2; #x2; #y2; #z2; #v2; #w2;
nnormalize;
nrewrite > (H x1 x2);
ncases (f1 x2 x1);
nlemma eq_to_eqquintuple :
∀T1,T2,T3,T4,T5.
∀f1:T1 → T1 → bool.∀f2:T2 → T2 → bool.∀f3:T3 → T3 → bool.∀f4:T4 → T4 → bool.∀f5:T5 → T5 → bool.
-∀p1,p2:Prod5T T1 T2 T3 T4 T5.
(∀x,y:T1.x = y → f1 x y = true) →
(∀x,y:T2.x = y → f2 x y = true) →
(∀x,y:T3.x = y → f3 x y = true) →
(∀x,y:T4.x = y → f4 x y = true) →
(∀x,y:T5.x = y → f5 x y = true) →
- (p1 = p2 → eq_quintuple T1 T2 T3 T4 T5 f1 f2 f3 f4 f5 p1 p2 = true).
- #T1; #T2; #T3; #T4; #T5; #f1; #f2; #f3; #f4; #f5; #p1; #p2; #H1; #H2; #H3; #H4; #H5;
- nelim p1;
- #x1; #y1; #z1; #v1; #w1;
- nelim p2;
- #x2; #y2; #z2; #v2; #w2; #H;
+ (∀p1,p2:Prod5T T1 T2 T3 T4 T5.
+ (p1 = p2 → eq_quintuple T1 T2 T3 T4 T5 f1 f2 f3 f4 f5 p1 p2 = true)).
+ #T1; #T2; #T3; #T4; #T5; #f1; #f2; #f3; #f4; #f5; #H1; #H2; #H3; #H4; #H5;
+ #p1; nelim p1; #x1; #y1; #z1; #v1; #w1;
+ #p2; nelim p2; #x2; #y2; #z2; #v2; #w2; #H;
nnormalize;
nrewrite > (H1 … (quintuple_destruct_1 … H));
nnormalize;
nlemma eqquintuple_to_eq :
∀T1,T2,T3,T4,T5.
∀f1:T1 → T1 → bool.∀f2:T2 → T2 → bool.∀f3:T3 → T3 → bool.∀f4:T4 → T4 → bool.∀f5:T5 → T5 → bool.
-∀p1,p2:Prod5T T1 T2 T3 T4 T5.
(∀x,y:T1.f1 x y = true → x = y) →
(∀x,y:T2.f2 x y = true → x = y) →
(∀x,y:T3.f3 x y = true → x = y) →
(∀x,y:T4.f4 x y = true → x = y) →
(∀x,y:T5.f5 x y = true → x = y) →
- (eq_quintuple T1 T2 T3 T4 T5 f1 f2 f3 f4 f5 p1 p2 = true → p1 = p2).
- #T1; #T2; #T3; #T4; #T5; #f1; #f2; #f3; #f4; #f5; #p1; #p2; #H1; #H2; #H3; #H4; #H5;
- nelim p1;
- #x1; #y1; #z1; #v1; #w1;
- nelim p2;
- #x2; #y2; #z2; #v2; #w2; #H;
+ (∀p1,p2:Prod5T T1 T2 T3 T4 T5.
+ (eq_quintuple T1 T2 T3 T4 T5 f1 f2 f3 f4 f5 p1 p2 = true → p1 = p2)).
+ #T1; #T2; #T3; #T4; #T5; #f1; #f2; #f3; #f4; #f5; #H1; #H2; #H3; #H4; #H5;
+ #p1; nelim p1; #x1; #y1; #z1; #v1; #w1;
+ #p2; nelim p2; #x2; #y2; #z2; #v2; #w2; #H;
nnormalize in H:(%);
nletin K ≝ (H1 x1 x2);
ncases (f1 x1 x2) in H:(%) K:(%);
napply refl_eq.
nqed.
-nlemma decidable_quintuple
- : ∀T1,T2,T3,T4,T5.
- (∀x,y:T1.decidable (x = y)) →
- (∀x,y:T2.decidable (x = y)) →
- (∀x,y:T3.decidable (x = y)) →
- (∀x,y:T4.decidable (x = y)) →
- (∀x,y:T5.decidable (x = y)) →
- ∀x,y:Prod5T T1 T2 T3 T4 T5.decidable (x = y).
+nlemma decidable_quintuple :
+∀T1,T2,T3,T4,T5.
+ (∀x,y:T1.decidable (x = y)) →
+ (∀x,y:T2.decidable (x = y)) →
+ (∀x,y:T3.decidable (x = y)) →
+ (∀x,y:T4.decidable (x = y)) →
+ (∀x,y:T5.decidable (x = y)) →
+ (∀x,y:Prod5T T1 T2 T3 T4 T5.decidable (x = y)).
#T1; #T2; #T3; #T4; #T5; #H; #H1; #H2; #H3; #H4;
#x; nelim x; #xx1; #xx2; #xx3; #xx4; #xx5;
#y; nelim y; #yy1; #yy2; #yy3; #yy4; #yy5;
nlemma neqquintuple_to_neq :
∀T1,T2,T3,T4,T5.
∀f1:T1 → T1 → bool.∀f2:T2 → T2 → bool.∀f3:T3 → T3 → bool.∀f4:T4 → T4 → bool.∀f5:T5 → T5 → bool.
-∀p1,p2:Prod5T T1 T2 T3 T4 T5.
(∀x,y:T1.f1 x y = false → x ≠ y) →
(∀x,y:T2.f2 x y = false → x ≠ y) →
(∀x,y:T3.f3 x y = false → x ≠ y) →
(∀x,y:T4.f4 x y = false → x ≠ y) →
(∀x,y:T5.f5 x y = false → x ≠ y) →
- (eq_quintuple T1 T2 T3 T4 T5 f1 f2 f3 f4 f5 p1 p2 = false → p1 ≠ p2).
- #T1; #T2; #T3; #T4; #T5; #f1; #f2; #f3; #f4; #f5; #p1; #p2; #H1; #H2; #H3; #H4; #H5;
- nelim p1;
- #x1; #y1; #z1; #v1; #w1;
- nelim p2;
- #x2; #y2; #z2; #v2; #w2;
+ (∀p1,p2:Prod5T T1 T2 T3 T4 T5.
+ (eq_quintuple T1 T2 T3 T4 T5 f1 f2 f3 f4 f5 p1 p2 = false → p1 ≠ p2)).
+ #T1; #T2; #T3; #T4; #T5; #f1; #f2; #f3; #f4; #f5; #H1; #H2; #H3; #H4; #H5;
+ #p1; nelim p1; #x1; #y1; #z1; #v1; #w1;
+ #p2; nelim p2; #x2; #y2; #z2; #v2; #w2;
nchange with ((((f1 x1 x2) ⊗ (f2 y1 y2) ⊗ (f3 z1 z2) ⊗ (f4 v1 v2) ⊗ (f5 w1 w2)) = false) → ?); #H;
nnormalize; #H6;
napply (or5_elim ((f1 x1 x2) = false) ((f2 y1 y2) = false) ((f3 z1 z2) = false) ((f4 v1 v2) = false) ((f5 w1 w2) = false) ? (andb_false5 … H) ?);
##]
nqed.
-nlemma quintuple_destruct
- : ∀T1,T2,T3,T4,T5.
- (∀x,y:T1.decidable (x = y)) →
- (∀x,y:T2.decidable (x = y)) →
- (∀x,y:T3.decidable (x = y)) →
- (∀x,y:T4.decidable (x = y)) →
- (∀x,y:T5.decidable (x = y)) →
- ∀x1,x2:T1.∀y1,y2:T2.∀z1,z2:T3.∀v1,v2:T4.∀w1,w2:T5.
- (quintuple T1 T2 T3 T4 T5 x1 y1 z1 v1 w1) ≠ (quintuple T1 T2 T3 T4 T5 x2 y2 z2 v2 w2) →
- Or5 (x1 ≠ x2) (y1 ≠ y2) (z1 ≠ z2) (v1 ≠ v2) (w1 ≠ w2).
+nlemma quintuple_destruct :
+∀T1,T2,T3,T4,T5.
+ (∀x,y:T1.decidable (x = y)) →
+ (∀x,y:T2.decidable (x = y)) →
+ (∀x,y:T3.decidable (x = y)) →
+ (∀x,y:T4.decidable (x = y)) →
+ (∀x,y:T5.decidable (x = y)) →
+ (∀x1,x2:T1.∀y1,y2:T2.∀z1,z2:T3.∀v1,v2:T4.∀w1,w2:T5.
+ (quintuple T1 T2 T3 T4 T5 x1 y1 z1 v1 w1) ≠ (quintuple T1 T2 T3 T4 T5 x2 y2 z2 v2 w2) →
+ Or5 (x1 ≠ x2) (y1 ≠ y2) (z1 ≠ z2) (v1 ≠ v2) (w1 ≠ w2)).
#T1; #T2; #T3; #T4; #T5; #H1; #H2; #H3; #H4; #H5;
#x1; #x2; #y1; #y2; #z1; #z2; #v1; #v2; #w1; #w2;
nnormalize; #H;
nlemma neq_to_neqquintuple :
∀T1,T2,T3,T4,T5.
∀f1:T1 → T1 → bool.∀f2:T2 → T2 → bool.∀f3:T3 → T3 → bool.∀f4:T4 → T4 → bool.∀f5:T5 → T5 → bool.
-∀p1,p2:Prod5T T1 T2 T3 T4 T5.
(∀x,y:T1.decidable (x = y)) →
(∀x,y:T2.decidable (x = y)) →
(∀x,y:T3.decidable (x = y)) →
(∀x,y:T3.x ≠ y → f3 x y = false) →
(∀x,y:T4.x ≠ y → f4 x y = false) →
(∀x,y:T5.x ≠ y → f5 x y = false) →
- (p1 ≠ p2 → eq_quintuple T1 T2 T3 T4 T5 f1 f2 f3 f4 f5 p1 p2 = false).
- #T1; #T2; #T3; #T4; #T5; #f1; #f2; #f3; #f4; #f5; #p1; #p2;
+ (∀p1,p2:Prod5T T1 T2 T3 T4 T5.
+ (p1 ≠ p2 → eq_quintuple T1 T2 T3 T4 T5 f1 f2 f3 f4 f5 p1 p2 = false)).
+ #T1; #T2; #T3; #T4; #T5; #f1; #f2; #f3; #f4; #f5;
#H1; #H2; #H3; #H4; #H5; #H6; #H7; #H8; #H9; #H10;
- nelim p1;
- #x1; #y1; #z1; #v1; #w1;
- nelim p2;
- #x2; #y2; #z2; #v2; #w2; #H;
+ #p1; nelim p1; #x1; #y1; #z1; #v1; #w1;
+ #p2; nelim p2; #x2; #y2; #z2; #v2; #w2; #H;
nchange with (((f1 x1 x2) ⊗ (f2 y1 y2) ⊗ (f3 z1 z2) ⊗ (f4 v1 v2) ⊗ (f5 w1 w2)) = false);
napply (or5_elim (x1 ≠ x2) (y1 ≠ y2) (z1 ≠ z2) (v1 ≠ v2) (w1 ≠ w2) ? (quintuple_destruct T1 T2 T3 T4 T5 H1 H2 H3 H4 H5 … H) ?);
##[ ##1: #H11; nrewrite > (H6 … H11); nrewrite > (andb_false5_1 (f2 y1 y2) (f3 z1 z2) (f4 v1 v2) (f5 w1 w2)); napply refl_eq
nlemma symmetric_eqstr : symmetricT (list ascii) bool eq_str.
#s1; #s2;
- napply (symmetric_bfoldrightlist2 ascii eq_ascii s1 s2 symmetric_eqascii).
+ napply (symmetric_bfoldrightlist2 ascii eq_ascii symmetric_eqascii s1 s2).
nqed.
nlemma eqstr_to_eq : ∀s,s'.eq_str s s' = true → s = s'.
#s1; #s2;
- napply (bfoldrightlist2_to_eq ascii eq_ascii s1 s2 eqascii_to_eq).
+ napply (bfoldrightlist2_to_eq ascii eq_ascii eqascii_to_eq s1 s2).
nqed.
nlemma eq_to_eqstr : ∀s,s'.s = s' → eq_str s s' = true.
#s1; #s2;
- napply (eq_to_bfoldrightlist2 ascii eq_ascii s1 s2 eq_to_eqascii).
+ napply (eq_to_bfoldrightlist2 ascii eq_ascii eq_to_eqascii s1 s2).
nqed.
nlemma decidable_str : ∀x,y:list ascii.decidable (x = y).
nlemma neqstr_to_neq : ∀s,s'.eq_str s s' = false → s ≠ s'.
#s1; #s2;
- napply (nbfoldrightlist2_to_neq ascii eq_ascii s1 s2 …);
+ napply (nbfoldrightlist2_to_neq ascii eq_ascii ? s1 s2 …);
napply neqascii_to_neq.
nqed.
nlemma neq_to_neqstr : ∀s,s'.s ≠ s' → eq_str s s' = false.
#s1; #s2;
- napply (neq_to_nbfoldrightlist2 ascii eq_ascii s1 s2 …);
+ napply (neq_to_nbfoldrightlist2 ascii eq_ascii ?? s1 s2 …);
##[ ##1: napply decidable_ascii
##| ##2: napply neq_to_neqascii
##]
nrewrite > (symmetric_eqstr (str_elem si1) (str_elem si2));
nrewrite > (symmetric_eqnat (id_elem si1) (id_elem si2));
napply refl_eq.
-nqed.
+nqed.
nlemma eqstrid_to_eq : ∀s,s'.eq_strId s s' = true → s = s'.
#si1; #si2;
common/string_lemmas.ma common/ascii_lemmas.ma common/list_utility_lemmas.ma common/string.ma
compiler/ast_type_lemmas.ma common/list_utility_lemmas.ma compiler/ast_type.ma
num/quatern.ma num/bool.ma
-num/exadecim.ma common/nat.ma common/prod.ma num/bool.ma num/oct.ma num/quatern.ma
+num/exadecim.ma common/prod.ma num/bool.ma num/oct.ma num/quatern.ma
num/bitrigesim_lemmas.ma num/bitrigesim.ma num/bool_lemmas.ma
emulator/status/status.ma emulator/memory/memory_abs.ma emulator/opcodes/opcode.ma emulator/status/HC05_status.ma emulator/status/HC08_status.ma emulator/status/IP2022_status.ma emulator/status/RS08_status.ma
-num/byte8.ma num/bitrigesim.ma num/exadecim.ma
+num/byte8.ma common/nat.ma num/bitrigesim.ma num/comp_num.ma num/exadecim.ma
emulator/model/model.ma emulator/model/HC05_model.ma emulator/model/HC08_model.ma emulator/model/HCS08_model.ma emulator/model/IP2022_model.ma emulator/model/RS08_model.ma
emulator/opcodes/HC08_table.ma common/list.ma emulator/opcodes/HC08_instr_mode.ma emulator/opcodes/HC08_opcode.ma emulator/opcodes/byte_or_word.ma
common/nat_lemmas.ma common/nat.ma num/bool_lemmas.ma
num/quatern_lemmas.ma num/bool_lemmas.ma num/quatern.ma
emulator/memory/memory_func.ma common/list.ma common/option.ma emulator/memory/memory_struct.ma num/word32.ma
emulator/opcodes/RS08_instr_mode.ma num/word16.ma
+num/comp_num_lemmas.ma num/bool_lemmas.ma num/comp_num.ma
emulator/status/RS08_status.ma num/word16.ma
emulator/opcodes/HC08_instr_mode.ma num/word16.ma
emulator/opcodes/opcode.ma common/list.ma emulator/opcodes/HC05_instr_mode.ma emulator/opcodes/HC05_opcode.ma emulator/opcodes/HC08_instr_mode.ma emulator/opcodes/HC08_opcode.ma emulator/opcodes/HCS08_opcode.ma emulator/opcodes/IP2022_instr_mode.ma emulator/opcodes/IP2022_opcode.ma emulator/opcodes/RS08_instr_mode.ma emulator/opcodes/RS08_opcode.ma emulator/opcodes/byte_or_word.ma
common/option.ma num/bool.ma
common/option_lemmas.ma common/option.ma num/bool_lemmas.ma
-num/byte8_lemmas.ma num/byte8.ma num/exadecim_lemmas.ma
+num/byte8_lemmas.ma num/byte8.ma num/comp_num_lemmas.ma num/exadecim_lemmas.ma
emulator/opcodes/RS08_table.ma common/list.ma emulator/opcodes/RS08_instr_mode.ma emulator/opcodes/RS08_opcode.ma emulator/opcodes/byte_or_word.ma
emulator/model/RS08_model.ma emulator/status/status.ma
+num/comp_num.ma num/exadecim.ma
common/sigma.ma common/theory.ma
common/list_lemmas.ma common/list.ma
emulator/opcodes/HCS08_table.ma common/list.ma emulator/opcodes/HC08_instr_mode.ma emulator/opcodes/HCS08_opcode.ma emulator/opcodes/byte_or_word.ma
| xE ⇒ quadruple … true true true false
| xF ⇒ quadruple … true true true true
].
+
ndefinition bits_of_byte8 ≝
λb:byte8.
- mk_Array8T ? (fst4T … (bits_of_exadecim (b8h b)))
- (snd4T … (bits_of_exadecim (b8h b)))
- (thd4T … (bits_of_exadecim (b8h b)))
- (fth4T … (bits_of_exadecim (b8h b)))
- (fst4T … (bits_of_exadecim (b8l b)))
- (snd4T … (bits_of_exadecim (b8l b)))
- (thd4T … (bits_of_exadecim (b8l b)))
- (fth4T … (bits_of_exadecim (b8l b))).
+ mk_Array8T ? (fst4T … (bits_of_exadecim (cnH ? b)))
+ (snd4T … (bits_of_exadecim (cnH ? b)))
+ (thd4T … (bits_of_exadecim (cnH ? b)))
+ (fth4T … (bits_of_exadecim (cnH ? b)))
+ (fst4T … (bits_of_exadecim (cnL ? b)))
+ (snd4T … (bits_of_exadecim (cnL ? b)))
+ (thd4T … (bits_of_exadecim (cnL ? b)))
+ (fth4T … (bits_of_exadecim (cnL ? b))).
(* visita di un albero da 4GB di elementi: ln16(4GB)=8 passaggi *)
ndefinition mt_visit ≝
λT:Type.λdata:aux_32B_type T.λaddr:word32.
- getn_array16T (b8l (w16l (w32l addr))) ?
- (getn_array16T (b8h (w16l (w32l addr))) ?
- (getn_array16T (b8l (w16h (w32l addr))) ?
- (getn_array16T (b8h (w16h (w32l addr))) ?
- (getn_array16T (b8l (w16l (w32h addr))) ?
- (getn_array16T (b8h (w16l (w32h addr))) ?
- (getn_array16T (b8l (w16h (w32h addr))) ?
- (getn_array16T (b8h (w16h (w32h addr))) ? data))))))).
+ getn_array16T (cnL ? (cnL ? (cnL ? addr))) ?
+ (getn_array16T (cnH ? (cnL ? (cnL ? addr))) ?
+ (getn_array16T (cnL ? (cnH ? (cnL ? addr))) ?
+ (getn_array16T (cnH ? (cnH ? (cnL ? addr))) ?
+ (getn_array16T (cnL ? (cnL ? (cnH ? addr))) ?
+ (getn_array16T (cnH ? (cnL ? (cnH ? addr))) ?
+ (getn_array16T (cnL ? (cnH ? (cnH ? addr))) ?
+ (getn_array16T (cnH ? (cnH ? (cnH ? addr))) ? data))))))).
(* scrittura di un elemento in un albero da 4GB *)
ndefinition mt_update ≝
λT:Type.λdata:aux_32B_type T.λaddr:word32.λv:T.
- let lev2 ≝ getn_array16T (b8h (w16h (w32h addr))) ? data in
- let lev3 ≝ getn_array16T (b8l (w16h (w32h addr))) ? lev2 in
- let lev4 ≝ getn_array16T (b8h (w16l (w32h addr))) ? lev3 in
- let lev5 ≝ getn_array16T (b8l (w16l (w32h addr))) ? lev4 in
- let lev6 ≝ getn_array16T (b8h (w16h (w32l addr))) ? lev5 in
- let lev7 ≝ getn_array16T (b8l (w16h (w32l addr))) ? lev6 in
- let lev8 ≝ getn_array16T (b8h (w16l (w32l addr))) ? lev7 in
+ let lev2 ≝ getn_array16T (cnH ? (cnH ? (cnH ? addr))) ? data in
+ let lev3 ≝ getn_array16T (cnL ? (cnH ? (cnH ? addr))) ? lev2 in
+ let lev4 ≝ getn_array16T (cnH ? (cnL ? (cnH ? addr))) ? lev3 in
+ let lev5 ≝ getn_array16T (cnL ? (cnL ? (cnH ? addr))) ? lev4 in
+ let lev6 ≝ getn_array16T (cnH ? (cnH ? (cnL ? addr))) ? lev5 in
+ let lev7 ≝ getn_array16T (cnL ? (cnH ? (cnL ? addr))) ? lev6 in
+ let lev8 ≝ getn_array16T (cnH ? (cnL ? (cnL ? addr))) ? lev7 in
- setn_array16T (b8h (w16h (w32h addr))) ? data
- (setn_array16T (b8l (w16h (w32h addr))) ? lev2
- (setn_array16T (b8h (w16l (w32h addr))) ? lev3
- (setn_array16T (b8l (w16l (w32h addr))) ? lev4
- (setn_array16T (b8h (w16h (w32l addr))) ? lev5
- (setn_array16T (b8l (w16h (w32l addr))) ? lev6
- (setn_array16T (b8h (w16l (w32l addr))) ? lev7
- (setn_array16T (b8l (w16l (w32l addr))) T lev8 v))))))).
+ setn_array16T (cnH ? (cnH ? (cnH ? addr))) ? data
+ (setn_array16T (cnL ? (cnH ? (cnH ? addr))) ? lev2
+ (setn_array16T (cnH ? (cnL ? (cnH ? addr))) ? lev3
+ (setn_array16T (cnL ? (cnL ? (cnH ? addr))) ? lev4
+ (setn_array16T (cnH ? (cnH ? (cnL ? addr))) ? lev5
+ (setn_array16T (cnL ? (cnH ? (cnL ? addr))) ? lev6
+ (setn_array16T (cnH ? (cnL ? (cnL ? addr))) ? lev7
+ (setn_array16T (cnL ? (cnL ? (cnL ? addr))) T lev8 v))))))).
(* scrittura di un range (max 64KB) in un albero da 4GB *)
nlet rec mt_update_ranged (T:Type) (data:aux_32B_type T) (addr:word32) (w:word16) (rw:rec_word16 w) (v:T) on rw ≝
λalu.λindx16:word16.
mk_alu_HC08
(acc_low_reg_HC08 alu)
- (w16l indx16)
- (w16h indx16)
+ (cnL ? indx16)
+ (cnH ? indx16)
(sp_reg_HC08 alu)
(pc_reg_HC08 alu)
(v_flag_HC08 alu)
ndefinition get_low3bits ≝
λaddrsel:byte8.
- match b8l addrsel with
+ match cnL ? addrsel with
[ x0 ⇒ o0 | x1 ⇒ o1 | x2 ⇒ o2 | x3 ⇒ o3 | x4 ⇒ o4 | x5 ⇒ o5 | x6 ⇒ o6 | x7 ⇒ o7
| x8 ⇒ o0 | x9 ⇒ o1 | xA ⇒ o2 | xB ⇒ o3 | xC ⇒ o4 | xD ⇒ o5 | xE ⇒ o6 | xF ⇒ o7 ].
(* inherent address: legale se nessun operando/1 byte/1 word *)
| maINHX0ADD ⇒ nil ?
| maINHX1ADD b ⇒ [ TByte HC05 b ]
- | maINHX2ADD w ⇒ [ TByte HC05 (w16h w); TByte HC05 (w16l w) ]
+ | maINHX2ADD w ⇒ [ TByte HC05 (cnH ? w); TByte HC05 (cnL ? w) ]
(* _0/1/2: legale se nessun operando/1 byte/1 word *)
| maIMM1 b ⇒ [ TByte HC05 b ]
| maIMM1EXT b ⇒ [ TByte HC05 b ]
- | maIMM2 w ⇒ [ TByte HC05 (w16h w); TByte HC05 (w16l w) ]
+ | maIMM2 w ⇒ [ TByte HC05 (cnH ? w); TByte HC05 (cnL ? w) ]
| maDIR1 b ⇒ [ TByte HC05 b ]
- | maDIR2 w ⇒ [ TByte HC05 (w16h w); TByte HC05 (w16l w) ]
+ | maDIR2 w ⇒ [ TByte HC05 (cnH ? w); TByte HC05 (cnL ? w) ]
| maIX0 ⇒ nil ?
| maIX1 b ⇒ [ TByte HC05 b ]
- | maIX2 w ⇒ [ TByte HC05 (w16h w); TByte HC05 (w16l w) ]
+ | maIX2 w ⇒ [ TByte HC05 (cnH ? w); TByte HC05 (cnL ? w) ]
(* DIRn: legale se 1 operando byte *)
| maDIRn _ b ⇒ [ TByte HC05 b]
(* DIRn_and_IMM1: legale se 2 operandi byte *)
(* inherent address: legale se nessun operando/1 byte/1 word *)
| maINHX0ADD ⇒ nil ?
| maINHX1ADD b ⇒ [ TByte HC08 b ]
- | maINHX2ADD w ⇒ [ TByte HC08 (w16h w); TByte HC08 (w16l w) ]
+ | maINHX2ADD w ⇒ [ TByte HC08 (cnH ? w); TByte HC08 (cnL ? w) ]
(* _0/1/2: legale se nessun operando/1 byte/1 word *)
| maIMM1 b ⇒ [ TByte HC08 b ]
| maIMM1EXT b ⇒ [ TByte HC08 b ]
- | maIMM2 w ⇒ [ TByte HC08 (w16h w); TByte HC08 (w16l w) ]
+ | maIMM2 w ⇒ [ TByte HC08 (cnH ? w); TByte HC08 (cnL ? w) ]
| maDIR1 b ⇒ [ TByte HC08 b ]
- | maDIR2 w ⇒ [ TByte HC08 (w16h w); TByte HC08 (w16l w) ]
+ | maDIR2 w ⇒ [ TByte HC08 (cnH ? w); TByte HC08 (cnL ? w) ]
| maIX0 ⇒ nil ?
| maIX1 b ⇒ [ TByte HC08 b ]
- | maIX2 w ⇒ [ TByte HC08 (w16h w); TByte HC08 (w16l w) ]
+ | maIX2 w ⇒ [ TByte HC08 (cnH ? w); TByte HC08 (cnL ? w) ]
| maSP1 b ⇒ [ TByte HC08 b ]
- | maSP2 w ⇒ [ TByte HC08 (w16h w); TByte HC08 (w16l w) ]
+ | maSP2 w ⇒ [ TByte HC08 (cnH ? w); TByte HC08 (cnL ? w) ]
(* movimento: legale se 2 operandi byte *)
| maDIR1_to_DIR1 b1 b2 ⇒ [ TByte HC08 b1 ; TByte HC08 b2 ]
| maIMM1_to_DIR1 b1 b2 ⇒ [ TByte HC08 b1 ; TByte HC08 b2 ]
(* inherent address: legale se nessun operando/1 byte/1 word *)
| maINHX0ADD ⇒ nil ?
| maINHX1ADD b ⇒ [ TByte HCS08 b ]
- | maINHX2ADD w ⇒ [ TByte HCS08 (w16h w); TByte HCS08 (w16l w) ]
+ | maINHX2ADD w ⇒ [ TByte HCS08 (cnH ? w); TByte HCS08 (cnL ? w) ]
(* _0/1/2: legale se nessun operando/1 byte/1 word *)
| maIMM1 b ⇒ [ TByte HCS08 b ]
| maIMM1EXT b ⇒ [ TByte HCS08 b ]
- | maIMM2 w ⇒ [ TByte HCS08 (w16h w); TByte HCS08 (w16l w) ]
+ | maIMM2 w ⇒ [ TByte HCS08 (cnH ? w); TByte HCS08 (cnL ? w) ]
| maDIR1 b ⇒ [ TByte HCS08 b ]
- | maDIR2 w ⇒ [ TByte HCS08 (w16h w); TByte HCS08 (w16l w) ]
+ | maDIR2 w ⇒ [ TByte HCS08 (cnH ? w); TByte HCS08 (cnL ? w) ]
| maIX0 ⇒ nil ?
| maIX1 b ⇒ [ TByte HCS08 b ]
- | maIX2 w ⇒ [ TByte HCS08 (w16h w); TByte HCS08 (w16l w) ]
+ | maIX2 w ⇒ [ TByte HCS08 (cnH ? w); TByte HCS08 (cnL ? w) ]
| maSP1 b ⇒ [ TByte HCS08 b ]
- | maSP2 w ⇒ [ TByte HCS08 (w16h w); TByte HCS08 (w16l w) ]
+ | maSP2 w ⇒ [ TByte HCS08 (cnH ? w); TByte HCS08 (cnL ? w) ]
(* movimento: legale se 2 operandi byte *)
| maDIR1_to_DIR1 b1 b2 ⇒ [ TByte HCS08 b1 ; TByte HCS08 b2 ]
| maIMM1_to_DIR1 b1 b2 ⇒ [ TByte HCS08 b1 ; TByte HCS08 b2 ]
| maINHA ⇒ nil ?
(* _0/1/2: legale se nessun operando/1 byte/1 word *)
| maIMM1 b ⇒ [ TByte RS08 b ]
- | maIMM2 w ⇒ [ TByte RS08 (w16h w); TByte RS08 (w16l w) ]
+ | maIMM2 w ⇒ [ TByte RS08 (cnH ? w); TByte RS08 (cnL ? w) ]
| maDIR1 b ⇒ [ TByte RS08 b ]
(* movimento: legale se 2 operandi byte *)
| maDIR1_to_DIR1 b1 b2 ⇒ [ TByte RS08 b1 ; TByte RS08 b2 ]
[ Byte isab ⇒
Some ? ([ (TByte m isab) ]@(args_picker m i p))
| Word isaw ⇒
- Some ? ([ (TByte m (w16h isaw)) ; (TByte m (w16l isaw)) ]@(args_picker m i p))
+ Some ? ([ (TByte m (cnH ? isaw)) ; (TByte m (cnL ? isaw)) ]@(args_picker m i p))
]
| false ⇒ bytes_of_pseudo_instr_mode_param_aux m o i p tl ]].
(* ********************************************************************** *)
include "num/exadecim.ma".
+include "num/comp_num.ma".
include "num/bitrigesim.ma".
+include "common/nat.ma".
(* **** *)
(* BYTE *)
(* **** *)
-nrecord byte8 : Type ≝
- {
- b8h: exadecim;
- b8l: exadecim
- }.
+ndefinition byte8 ≝ comp_num exadecim.
+ndefinition mk_byte8 ≝ λe1,e2.mk_comp_num exadecim e1 e2.
(* \langle \rangle *)
notation "〈x,y〉" non associative with precedence 80
- for @{ 'mk_byte8 $x $y }.
-interpretation "mk_byte8" 'mk_byte8 x y = (mk_byte8 x y).
+ for @{ mk_comp_num exadecim $x $y }.
+
+(* iteratore sui byte *)
+ndefinition forall_b8 ≝ forall_cn ? forall_ex.
(* operatore = *)
-ndefinition eq_b8 ≝ λb1,b2:byte8.(eq_ex (b8h b1) (b8h b2)) ⊗ (eq_ex (b8l b1) (b8l b2)).
+ndefinition eq_b8 ≝ eq2_cn ? eq_ex.
(* operatore < *)
-ndefinition lt_b8 ≝
-λb1,b2:byte8.
- (lt_ex (b8h b1) (b8h b2)) ⊕
- ((eq_ex (b8h b1) (b8h b2)) ⊗ (lt_ex (b8l b1) (b8l b2))).
+ndefinition lt_b8 ≝ ltgt_cn ? eq_ex lt_ex.
(* operatore ≤ *)
-ndefinition le_b8 ≝
-λb1,b2:byte8.
- (lt_ex (b8h b1) (b8h b2)) ⊕
- ((eq_ex (b8h b1) (b8h b2)) ⊗ (le_ex (b8l b1) (b8l b2))).
+ndefinition le_b8 ≝ lege_cn ? eq_ex lt_ex le_ex.
(* operatore > *)
-ndefinition gt_b8 ≝
-λb1,b2:byte8.
- (gt_ex (b8h b1) (b8h b2)) ⊕
- ((eq_ex (b8h b1) (b8h b2)) ⊗ (gt_ex (b8l b1) (b8l b2))).
+ndefinition gt_b8 ≝ ltgt_cn ? eq_ex gt_ex.
(* operatore ≥ *)
-ndefinition ge_b8 ≝
-λb1,b2:byte8.
- (gt_ex (b8h b1) (b8h b2)) ⊕
- ((eq_ex (b8h b1) (b8h b2)) ⊗ (ge_ex (b8l b1) (b8l b2))).
+ndefinition ge_b8 ≝ lege_cn ? eq_ex gt_ex ge_ex.
(* operatore and *)
-ndefinition and_b8 ≝
-λb1,b2:byte8.mk_byte8 (and_ex (b8h b1) (b8h b2)) (and_ex (b8l b1) (b8l b2)).
+ndefinition and_b8 ≝ fop2_cn ? and_ex.
(* operatore or *)
-ndefinition or_b8 ≝
-λb1,b2:byte8.mk_byte8 (or_ex (b8h b1) (b8h b2)) (or_ex (b8l b1) (b8l b2)).
+ndefinition or_b8 ≝ fop2_cn ? or_ex.
(* operatore xor *)
-ndefinition xor_b8 ≝
-λb1,b2:byte8.mk_byte8 (xor_ex (b8h b1) (b8h b2)) (xor_ex (b8l b1) (b8l b2)).
+ndefinition xor_b8 ≝ fop2_cn ? xor_ex.
+
+(* operatore Most Significant Bit *)
+ndefinition getMSB_b8 ≝ getOPH_cn ? getMSB_ex.
+ndefinition setMSB_b8 ≝ setOPH_cn ? setMSB_ex.
+
+(* operatore Least Significant Bit *)
+ndefinition getLSB_b8 ≝ getOPL_cn ? getLSB_ex.
+ndefinition setLSB_b8 ≝ setOPL_cn ? setLSB_ex.
(* operatore rotazione destra con carry *)
-ndefinition rcr_b8 ≝
-λb:byte8.λc:bool.match rcr_ex (b8h b) c with
- [ pair bh' c' ⇒ match rcr_ex (b8l b) c' with
- [ pair bl' c'' ⇒ pair … (mk_byte8 bh' bl') c'' ]].
+ndefinition rcr_b8 ≝ opcr_cn ? rcr_ex.
(* operatore shift destro *)
-ndefinition shr_b8 ≝
-λb:byte8.match shr_ex (b8h b) with
- [ pair bh' c' ⇒ match rcr_ex (b8l b) c' with
- [ pair bl' c'' ⇒ pair … (mk_byte8 bh' bl') c'' ]].
+ndefinition shr_b8 ≝ opcr_cn ? rcr_ex false.
(* operatore rotazione destra *)
ndefinition ror_b8 ≝
-λb:byte8.match shr_ex (b8h b) with
- [ pair bh' c' ⇒ match rcr_ex (b8l b) c' with
- [ pair bl' c'' ⇒ match c'' with
- [ true ⇒ mk_byte8 (or_ex x8 bh') bl'
- | false ⇒ mk_byte8 bh' bl' ]]].
-
-(* operatore rotazione destra n-volte *)
-nlet rec ror_b8_n (b:byte8) (o:oct) (ro:rec_oct o) on ro ≝
- match ro with
- [ oc_O ⇒ b
- | oc_S o' n' ⇒ ror_b8_n (ror_b8 b) o' n' ].
+λb.match shr_b8 b with
+ [ pair c b' ⇒ match c with
+ [ true ⇒ setMSB_b8 b' | false ⇒ b' ]].
(* operatore rotazione sinistra con carry *)
-ndefinition rcl_b8 ≝
-λb:byte8.λc:bool.match rcl_ex (b8l b) c with
- [ pair bl' c' ⇒ match rcl_ex (b8h b) c' with
- [ pair bh' c'' ⇒ pair … (mk_byte8 bh' bl') c'' ]].
+ndefinition rcl_b8 ≝ opcl_cn ? rcl_ex.
(* operatore shift sinistro *)
-ndefinition shl_b8 ≝
-λb:byte8.match shl_ex (b8l b) with
- [ pair bl' c' ⇒ match rcl_ex (b8h b) c' with
- [ pair bh' c'' ⇒ pair … (mk_byte8 bh' bl') c'' ]].
+ndefinition shl_b8 ≝ opcl_cn ? rcl_ex false.
(* operatore rotazione sinistra *)
ndefinition rol_b8 ≝
-λb:byte8.match shl_ex (b8l b) with
- [ pair bl' c' ⇒ match rcl_ex (b8h b) c' with
- [ pair bh' c'' ⇒ match c'' with
- [ true ⇒ mk_byte8 bh' (or_ex x1 bl')
- | false ⇒ mk_byte8 bh' bl' ]]].
-
-(* operatore rotazione sinistra n-volte *)
-nlet rec rol_b8_n (b:byte8) (o:oct) (ro:rec_oct o) on ro ≝
- match ro with
- [ oc_O ⇒ b
- | oc_S o' n' ⇒ rol_b8_n (rol_b8 b) o' n' ].
+λb.match shl_b8 b with
+ [ pair c b' ⇒ match c with
+ [ true ⇒ setLSB_b8 b' | false ⇒ b' ]].
(* operatore not/complemento a 1 *)
-ndefinition not_b8 ≝
-λb:byte8.mk_byte8 (not_ex (b8h b)) (not_ex (b8l b)).
+ndefinition not_b8 ≝ fop_cn ? not_ex.
(* operatore somma con data+carry → data+carry *)
-ndefinition plus_b8_dc_dc ≝
-λb1,b2:byte8.λc:bool.
- match plus_ex_dc_dc (b8l b1) (b8l b2) c with
- [ pair l c' ⇒ match plus_ex_dc_dc (b8h b1) (b8h b2) c' with
- [ pair h c'' ⇒ pair … 〈h,l〉 c'' ]].
+ndefinition plus_b8_dc_dc ≝ opcl2_cn ? plus_ex_dc_dc.
(* operatore somma con data+carry → data *)
-ndefinition plus_b8_dc_d ≝
-λb1,b2:byte8.λc:bool.
- match plus_ex_dc_dc (b8l b1) (b8l b2) c with
- [ pair l c' ⇒ 〈plus_ex_dc_d (b8h b1) (b8h b2) c',l〉 ].
+ndefinition plus_b8_dc_d ≝ λc,b1,b2.snd … (plus_b8_dc_dc c b1 b2).
(* operatore somma con data+carry → c *)
-ndefinition plus_b8_dc_c ≝
-λb1,b2:byte8.λc:bool.
- plus_ex_dc_c (b8h b1) (b8h b2) (plus_ex_dc_c (b8l b1) (b8l b2) c).
+ndefinition plus_b8_dc_c ≝ λc,b1,b2.fst … (plus_b8_dc_dc c b1 b2).
(* operatore somma con data → data+carry *)
-ndefinition plus_b8_d_dc ≝
-λb1,b2:byte8.
- match plus_ex_d_dc (b8l b1) (b8l b2) with
- [ pair l c ⇒ match plus_ex_dc_dc (b8h b1) (b8h b2) c with
- [ pair h c' ⇒ pair … 〈h,l〉 c' ]].
+ndefinition plus_b8_d_dc ≝ opcl2_cn ? plus_ex_dc_dc false.
(* operatore somma con data → data *)
-ndefinition plus_b8_d_d ≝
-λb1,b2:byte8.
- match plus_ex_d_dc (b8l b1) (b8l b2) with
- [ pair l c ⇒ 〈plus_ex_dc_d (b8h b1) (b8h b2) c,l〉 ].
+ndefinition plus_b8_d_d ≝ λb1,b2.snd … (plus_b8_d_dc b1 b2).
(* operatore somma con data → c *)
-ndefinition plus_b8_d_c ≝
-λb1,b2:byte8.
- plus_ex_dc_c (b8h b1) (b8h b2) (plus_ex_d_c (b8l b1) (b8l b2)).
-
-(* operatore Most Significant Bit *)
-ndefinition MSB_b8 ≝ λb:byte8.eq_ex x8 (and_ex x8 (b8h b)).
+ndefinition plus_b8_d_c ≝ λb1,b2.fst … (plus_b8_d_dc b1 b2).
(* operatore predecessore *)
-ndefinition pred_b8 ≝
-λb:byte8.match eq_ex (b8l b) x0 with
- [ true ⇒ mk_byte8 (pred_ex (b8h b)) (pred_ex (b8l b))
- | false ⇒ mk_byte8 (b8h b) (pred_ex (b8l b)) ].
+ndefinition pred_b8 ≝ predsucc_cn ? (eq_ex x0) pred_ex.
(* operatore successore *)
-ndefinition succ_b8 ≝
-λb:byte8.match eq_ex (b8l b) xF with
- [ true ⇒ mk_byte8 (succ_ex (b8h b)) (succ_ex (b8l b))
- | false ⇒ mk_byte8 (b8h b) (succ_ex (b8l b)) ].
+ndefinition succ_b8 ≝ predsucc_cn ? (eq_ex xF) succ_ex.
(* operatore neg/complemento a 2 *)
ndefinition compl_b8 ≝
-λb:byte8.match MSB_b8 b with
+λb:byte8.match getMSB_b8 b with
[ true ⇒ succ_b8 (not_b8 b)
| false ⇒ not_b8 (pred_b8 b) ].
+(* operatore x in [inf,sup] o in sup],[inf *)
+ndefinition inrange_b8 ≝
+λx,inf,sup:byte8.
+ match le_b8 inf sup with
+ [ true ⇒ and_bool | false ⇒ or_bool ]
+ (le_b8 inf x) (le_b8 x sup).
+
(* operatore moltiplicazione senza segno: e*e=[0x00,0xE1] *)
ndefinition mul_ex ≝
λe1,e2:exadecim.match e1 with
(* X' = [(b16l X):0x0-0x9] X + [h=1 ? 0x06 : 0x00] + [c=1 ? 0x60 : 0x00]
[(b16l X):0xA-0xF] X + 0x06 + [c=1 ? 0x60 : 0x00] *)
[ true ⇒
- let X' ≝ match (lt_ex (b8l X) xA) ⊗ (⊖h) with
+ let X' ≝ match (lt_ex (cnL ? X) xA) ⊗ (⊖h) with
[ true ⇒ X
| false ⇒ plus_b8_d_d X 〈x0,x6〉 ] in
let X'' ≝ match c with
[(b16l X):0x0-0x9] X + [h=1 ? 0x06 : 0x00] + 0x60
[(b16l X):0xA-0xF] X + 0x6 + 0x60 *)
| false ⇒
- let X' ≝ match (lt_ex (b8l X) xA) ⊗ (⊖h) with
+ let X' ≝ match (lt_ex (cnL ? X) xA) ⊗ (⊖h) with
[ true ⇒ X
| false ⇒ plus_b8_d_d X 〈x0,x6〉 ] in
let X'' ≝ plus_b8_d_d X' 〈x6,x0〉 in
].
(* divisione senza segno (secondo la logica delle ALU): (quoziente resto) overflow *)
-nlet rec div_b8_ex_aux (divd:byte8) (divs:byte8) (molt:exadecim) (q:exadecim) (qu:quatern) (rqu:rec_quatern qu) on rqu ≝
+nlet rec div_b8_ex_aux (divd:byte8) (divs:byte8) (molt:exadecim) (q:exadecim) (n:nat) on n ≝
let w' ≝ plus_b8_d_d divd (compl_b8 divs) in
- match rqu with
- [ qu_O ⇒ match le_b8 divs divd with
- [ true ⇒ triple … (or_ex molt q) (b8l w') (⊖ (eq_ex (b8h w') x0))
- | false ⇒ triple … q (b8l divd) (⊖ (eq_ex (b8h divd) x0)) ]
- | qu_S qu' rqu' ⇒ match le_b8 divs divd with
- [ true ⇒ div_b8_ex_aux w' (ror_b8 divs) (ror_ex molt) (or_ex molt q) qu' rqu'
- | false ⇒ div_b8_ex_aux divd (ror_b8 divs) (ror_ex molt) q qu' rqu' ]].
+ match n with
+ [ O ⇒ match le_b8 divs divd with
+ [ true ⇒ triple … (or_ex molt q) (cnL ? w') (⊖ (eq_ex (cnH ? w') x0))
+ | false ⇒ triple … q (cnL ? divd) (⊖ (eq_ex (cnH ? divd) x0)) ]
+ | S n' ⇒ match le_b8 divs divd with
+ [ true ⇒ div_b8_ex_aux w' (ror_b8 divs) (ror_ex molt) (or_ex molt q) n'
+ | false ⇒ div_b8_ex_aux divd (ror_b8 divs) (ror_ex molt) q n' ]].
ndefinition div_b8_ex ≝
λb:byte8.λe:exadecim.match eq_ex e x0 with
-(*
- la combinazione n/0 e' illegale, segnala solo overflow senza dare risultato
-*)
- [ true ⇒ triple … xF (b8l b) true
+(* la combinazione n/0 e' illegale, segnala solo overflow senza dare risultato *)
+ [ true ⇒ triple … xF (cnL ? b) true
| false ⇒ match eq_b8 b 〈x0,x0〉 with
(* 0 diviso qualsiasi cosa diverso da 0 da' q=0 r=0 o=false *)
[ true ⇒ triple … x0 x0 false
(* 3) ad ogni ciclo il divisore e il moltiplicatore vengono scalati di 1 a dx *)
(* 4) il moltiplicatore e' la quantita' aggiunta al quoziente se il divisore *)
(* puo' essere sottratto al dividendo *)
- | false ⇒ div_b8_ex_aux b (rol_b8_n 〈x0,e〉 ? (oct_to_recoct o3)) x8 x0 ? (qu_to_recqu q3) ]].
-
-(* operatore x in [inf,sup] o in sup],[inf *)
-ndefinition inrange_b8 ≝
-λx,inf,sup:byte8.
- match le_b8 inf sup with
- [ true ⇒ and_bool | false ⇒ or_bool ]
- (le_b8 inf x) (le_b8 x sup).
-
-(* iteratore sui byte *)
-ndefinition forall_b8 ≝
- λP.
- forall_ex (λbh.
- forall_ex (λbl.
- P (mk_byte8 bh bl))).
+ | false ⇒ div_b8_ex_aux b (nat_it ? rol_b8 〈x0,e〉 nat3) x8 x0 nat3 ]].
(* byte ricorsivi *)
ninductive rec_byte8 : byte8 → Type ≝
*)
ndefinition b8_to_recb8 : Πb.rec_byte8 b ≝
-λb.match b with [ mk_byte8 h l ⇒ b8_to_recb8_aux3 h l (b8_to_recb8_aux2 h (ex_to_recex h)) ].
+λb.match b with
+ [ mk_comp_num h l ⇒ b8_to_recb8_aux3 h l (b8_to_recb8_aux2 h (ex_to_recex h)) ].
(* ottali → esadecimali *)
ndefinition b8_of_bit ≝
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.
-
-nlemma decidable_b8_aux1 : ∀e1,e2,e3,e4.e1 ≠ e3 → 〈e1,e2〉 ≠ 〈e3,e4〉.
- #e1; #e2; #e3; #e4;
- nnormalize; #H; #H1;
- napply (H (byte8_destruct_1 … H1)).
-nqed.
-
-nlemma decidable_b8_aux2 : ∀e1,e2,e3,e4.e2 ≠ e4 → 〈e1,e2〉 ≠ 〈e3,e4〉.
- #e1; #e2; #e3; #e4;
- nnormalize; #H; #H1;
- napply (H (byte8_destruct_2 … H1)).
-nqed.
-
-nlemma decidable_b8 : ∀x,y:byte8.decidable (x = y).
- #x; nelim x; #e1; #e2;
- #y; nelim y; #e3; #e4;
- nnormalize;
- napply (or2_elim (e1 = e3) (e1 ≠ e3) ? (decidable_ex e1 e3) …);
- ##[ ##2: #H; napply (or2_intro2 … (decidable_b8_aux1 e1 e2 e3 e4 H))
- ##| ##1: #H; napply (or2_elim (e2 = e4) (e2 ≠ e4) ? (decidable_ex e2 e4) …);
- ##[ ##2: #H1; napply (or2_intro2 … (decidable_b8_aux2 e1 e2 e3 e4 H1))
- ##| ##1: #H1; nrewrite > H; nrewrite > H1;
- napply (or2_intro1 … (refl_eq ? 〈e3,e4〉))
- ##]
- ##]
-nqed.
-
-nlemma neqb8_to_neq : ∀b1,b2:byte8.(eq_b8 b1 b2 = false) → (b1 ≠ b2).
- #b1; #b2;
- nelim b1;
- nelim b2;
- #e1; #e2; #e3; #e4;
- nchange with ((((eq_ex e3 e1) ⊗ (eq_ex e4 e2)) = false) → ?);
- #H;
- napply (or2_elim ((eq_ex e3 e1) = false) ((eq_ex e4 e2) = false) ? (andb_false2 … H) …);
- ##[ ##1: #H1; napply (decidable_b8_aux1 … (neqex_to_neq … H1))
- ##| ##2: #H1; napply (decidable_b8_aux2 … (neqex_to_neq … H1))
- ##]
-nqed.
-
-nlemma byte8_destruct : ∀e1,e2,e3,e4.〈e1,e2〉 ≠ 〈e3,e4〉 → e1 ≠ e3 ∨ e2 ≠ e4.
- #e1; #e2; #e3; #e4;
- nnormalize; #H;
- napply (or2_elim (e1 = e3) (e1 ≠ e3) ? (decidable_ex e1 e3) …);
- ##[ ##2: #H1; napply (or2_intro1 … H1)
- ##| ##1: #H1; napply (or2_elim (e2 = e4) (e2 ≠ e4) ? (decidable_ex e2 e4) …);
- ##[ ##2: #H2; napply (or2_intro2 … H2)
- ##| ##1: #H2; nrewrite > H1 in H:(%);
- nrewrite > H2;
- #H; nelim (H (refl_eq …))
- ##]
- ##]
-nqed.
-
-nlemma neq_to_neqb8 : ∀b1,b2.b1 ≠ b2 → eq_b8 b1 b2 = false.
- #b1; #b2;
- nelim b1; #e1; #e2;
- nelim b2; #e3; #e4;
- #H; nchange with (((eq_ex e1 e3) ⊗ (eq_ex e2 e4)) = false);
- napply (or2_elim (e1 ≠ e3) (e2 ≠ e4) ? (byte8_destruct … H) …);
- ##[ ##1: #H1; nrewrite > (neq_to_neqex … H1); nnormalize; napply refl_eq
- ##| ##2: #H1; nrewrite > (neq_to_neqex … H1);
- nrewrite > (symmetric_andbool (eq_ex e1 e3) false);
- 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.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(* Progetto FreeScale *)
+(* *)
+(* Sviluppato da: Ing. Cosimo Oliboni, oliboni@cs.unibo.it *)
+(* Sviluppo: 2008-2010 *)
+(* *)
+(* ********************************************************************** *)
+
+include "num/exadecim.ma".
+
+(* *********** *)
+(* ESADECIMALI *)
+(* *********** *)
+
+nrecord comp_num (T:Type) : Type ≝
+ {
+ cnH: T;
+ cnL: T
+ }.
+
+ndefinition forall_cn ≝
+λT.λf:(T → bool) → bool.λP.
+ f (λh.
+ f (λl.
+ P (mk_comp_num T h l))).
+
+(* operatore = *)
+ndefinition eq2_cn ≝
+λT.λfeq:T → T → bool.
+λcn1,cn2:comp_num T.
+ (feq (cnH ? cn1) (cnH ? cn2)) ⊗ (feq (cnL ? cn1) (cnL ? cn2)).
+
+ndefinition eq_cn ≝
+λT.λfeq:T → bool.
+λcn:comp_num T.
+ (feq (cnH ? cn)) ⊗ (feq (cnL ? cn)).
+
+(* operatore < > *)
+ndefinition ltgt_cn ≝
+λT.λfeq:T → T → bool.λfltgt:T → T → bool.
+λcn1,cn2:comp_num T.
+ (fltgt (cnH ? cn1) (cnH ? cn2)) ⊕
+ ((feq (cnH ? cn1) (cnH ? cn2)) ⊗ (fltgt (cnL ? cn1) (cnL ? cn2))).
+
+(* operatore ≤ ≥ *)
+ndefinition lege_cn ≝
+λT.λfeq:T → T → bool.λfltgt:T → T → bool.λflege:T → T → bool.
+λcn1,cn2:comp_num T.
+ (fltgt (cnH ? cn1) (cnH ? cn2)) ⊕
+ ((feq (cnH ? cn1) (cnH ? cn2)) ⊗ (flege (cnL ? cn1) (cnL ? cn2))).
+
+(* operatore cn1 op cn2 *)
+ndefinition fop2_cn ≝
+λT.λfop:T → T → T.
+λcn1,cn2:comp_num T.
+ mk_comp_num ? (fop (cnH ? cn1) (cnH ? cn2)) (fop (cnL ? cn1) (cnL ? cn2)).
+
+ndefinition fop_cn ≝
+λT.λfop:T → T.
+λcn:comp_num T.
+ mk_comp_num ? (fop (cnH ? cn)) (fop (cnL ? cn)).
+
+(* operatore su parte high *)
+ndefinition getOPH_cn ≝
+λT.λf:T → bool.
+λcn:comp_num T.
+ f (cnH ? cn).
+
+ndefinition setOPH_cn ≝
+λT.λf:T → T.
+λcn:comp_num T.
+ mk_comp_num ? (f (cnH ? cn)) (cnL ? cn).
+
+(* operatore su parte low *)
+ndefinition getOPL_cn ≝
+λT.λf:T → bool.
+λcn:comp_num T.
+ f (cnL ? cn).
+
+ndefinition setOPL_cn ≝
+λT.λf:T → T.
+λcn:comp_num T.
+ mk_comp_num ? (cnH ? cn) (f (cnL ? cn)).
+
+(* operatore pred/succ *)
+ndefinition predsucc_cn ≝
+λT.λfz:T → bool.λfps:T → T.
+λcn:comp_num T.
+ match fz (cnL ? cn) with
+ [ true ⇒ mk_comp_num ? (fps (cnH ? cn)) (fps (cnL ? cn))
+ | false ⇒ mk_comp_num ? (cnH ? cn) (fps (cnL ? cn)) ].
+
+(* operatore con carry applicato DX → SX *)
+ndefinition opcr2_cn ≝
+λT.λfopcr:bool → T → T → (ProdT bool T).
+λc:bool.λcn1,cn2:comp_num T.
+ match fopcr c (cnH ? cn1) (cnH ? cn2) with
+ [ pair c' cnh' ⇒ match fopcr c' (cnL ? cn1) (cnL ? cn2) with
+ [ pair c'' cnl' ⇒ pair … c'' (mk_comp_num ? cnh' cnl') ]].
+
+ndefinition opcr_cn ≝
+λT.λfopcr:bool → T → (ProdT bool T).
+λc:bool.λcn:comp_num T.
+ match fopcr c (cnH ? cn) with
+ [ pair c' cnh' ⇒ match fopcr c' (cnL ? cn) with
+ [ pair c'' cnl' ⇒ pair … c'' (mk_comp_num ? cnh' cnl') ]].
+
+(* operatore con carry applicato SX → DX *)
+ndefinition opcl2_cn ≝
+λT.λfopcl:bool → T → T → (ProdT bool T).
+λc:bool.λcn1,cn2:comp_num T.
+ match fopcl c (cnL ? cn1) (cnL ? cn2) with
+ [ pair c' cnl' ⇒ match fopcl c' (cnH ? cn1) (cnH ? cn2) with
+ [ pair c'' cnh' ⇒ pair … c'' (mk_comp_num ? cnh' cnl') ]].
+
+ndefinition opcl_cn ≝
+λT.λfopcl:bool → T → (ProdT bool T).
+λc:bool.λcn:comp_num T.
+ match fopcl c (cnL ? cn) with
+ [ pair c' cnl' ⇒ match fopcl c' (cnH ? cn) with
+ [ pair c'' cnh' ⇒ pair … c'' (mk_comp_num ? cnh' cnl') ]].
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(* Progetto FreeScale *)
+(* *)
+(* Sviluppato da: Ing. Cosimo Oliboni, oliboni@cs.unibo.it *)
+(* Sviluppo: 2008-2010 *)
+(* *)
+(* ********************************************************************** *)
+
+include "num/comp_num.ma".
+include "num/bool_lemmas.ma".
+
+(* **** *)
+(* BYTE *)
+(* **** *)
+
+nlemma cn_destruct_1 :
+∀T.∀x1,x2,y1,y2:T.
+ mk_comp_num T x1 y1 = mk_comp_num T x2 y2 → x1 = x2.
+ #T; #x1; #x2; #y1; #y2; #H;
+ nchange with (match mk_comp_num ? x2 y2 with [ mk_comp_num a _ ⇒ x1 = a ]);
+ nrewrite < H;
+ nnormalize;
+ napply refl_eq.
+nqed.
+
+nlemma cn_destruct_2 :
+∀T.∀x1,x2,y1,y2:T.
+ mk_comp_num T x1 y1 = mk_comp_num T x2 y2 → y1 = y2.
+ #T; #x1; #x2; #y1; #y2; #H;
+ nchange with (match mk_comp_num ? x2 y2 with [ mk_comp_num _ b ⇒ y1 = b ]);
+ nrewrite < H;
+ nnormalize;
+ napply refl_eq.
+nqed.
+
+nlemma symmetric_eqcn :
+∀T.∀feq:T → T → bool.
+ (symmetricT T bool feq) →
+ (symmetricT (comp_num T) bool (eq2_cn T feq)).
+ #T; #feq; #H;
+ #b1; nelim b1; #e1; #e2;
+ #b2; nelim b2; #e3; #e4;
+ nchange with (((feq e1 e3)⊗(feq e2 e4)) = ((feq e3 e1)⊗(feq e4 e2)));
+ nrewrite > (H e1 e3);
+ nrewrite > (H e2 e4);
+ napply refl_eq.
+nqed.
+
+nlemma eqcn_to_eq :
+∀T.∀feq:T → T → bool.
+ (∀x,y:T.(feq x y = true) → (x = y)) →
+ (∀b1,b2:comp_num T.
+ ((eq2_cn T feq b1 b2 = true) → (b1 = b2))).
+ #T; #feq; #H; #b1; #b2;
+ nelim b1; #e1; #e2;
+ nelim b2; #e3; #e4;
+ nchange in ⊢ (% → ?) with (((feq e1 e3)⊗(feq e2 e4)) = true);
+ #H1;
+ nrewrite < (H … (andb_true_true_l … H1));
+ nrewrite < (H … (andb_true_true_r … H1));
+ napply refl_eq.
+nqed.
+
+nlemma eq_to_eqcn :
+∀T.∀feq:T → T → bool.
+ (∀x,y:T.(x = y) → (feq x y = true)) →
+ (∀b1,b2:comp_num T.
+ ((b1 = b2) → (eq2_cn T feq b1 b2 = true))).
+ #T; #feq; #H; #b1; #b2;
+ nelim b1; #e1; #e2;
+ nelim b2; #e3; #e4;
+ #H1;
+ nrewrite < (cn_destruct_1 … H1);
+ nrewrite < (cn_destruct_2 … H1);
+ nchange with (((feq e1 e1)⊗(feq e2 e2)) = true);
+ nrewrite > (H e1 e1 (refl_eq …));
+ nrewrite > (H e2 e2 (refl_eq …));
+ nnormalize;
+ napply refl_eq.
+nqed.
+
+nlemma decidable_cn_aux1 :
+∀T.∀e1,e2,e3,e4:T.e1 ≠ e3 → (mk_comp_num T e1 e2) ≠ (mk_comp_num T e3 e4).
+ #T; #e1; #e2; #e3; #e4;
+ nnormalize; #H; #H1;
+ napply (H (cn_destruct_1 … H1)).
+nqed.
+
+nlemma decidable_cn_aux2 :
+∀T.∀e1,e2,e3,e4:T.e2 ≠ e4 → (mk_comp_num T e1 e2) ≠ (mk_comp_num T e3 e4).
+ #T; #e1; #e2; #e3; #e4;
+ nnormalize; #H; #H1;
+ napply (H (cn_destruct_2 … H1)).
+nqed.
+
+nlemma decidable_cn :
+∀T.(∀x,y:T.decidable (x = y)) →
+ (∀b1,b2:comp_num T.
+ (decidable (b1 = b2))).
+ #T; #H;
+ #b1; nelim b1; #e1; #e2;
+ #b2; nelim b2; #e3; #e4;
+ nnormalize;
+ napply (or2_elim (e1 = e3) (e1 ≠ e3) ? (H e1 e3) …);
+ ##[ ##2: #H1; napply (or2_intro2 … (decidable_cn_aux1 T e1 e2 e3 e4 H1))
+ ##| ##1: #H1; napply (or2_elim (e2 = e4) (e2 ≠ e4) ? (H e2 e4) …);
+ ##[ ##2: #H2; napply (or2_intro2 … (decidable_cn_aux2 T e1 e2 e3 e4 H2))
+ ##| ##1: #H2; nrewrite > H1; nrewrite > H2;
+ napply (or2_intro1 … (refl_eq ? (mk_comp_num T e3 e4)))
+ ##]
+ ##]
+nqed.
+
+nlemma neqcn_to_neq :
+∀T.∀feq:T → T → bool.
+ (∀x,y:T.(feq x y = false) → (x ≠ y)) →
+ (∀b1,b2:comp_num T.
+ ((eq2_cn T feq b1 b2 = false) → (b1 ≠ b2))).
+ #T; #feq; #H; #b1; #b2;
+ nelim b1; #e1; #e2;
+ nelim b2; #e3; #e4;
+ nchange with ((((feq e1 e3) ⊗ (feq e2 e4)) = false) → ?);
+ #H1;
+ napply (or2_elim ((feq e1 e3) = false) ((feq e2 e4) = false) ? (andb_false2 … H1) …);
+ ##[ ##1: #H2; napply (decidable_cn_aux1 … (H … H2))
+ ##| ##2: #H2; napply (decidable_cn_aux2 … (H … H2))
+ ##]
+nqed.
+
+nlemma cn_destruct :
+∀T.(∀x,y:T.decidable (x = y)) →
+ (∀e1,e2,e3,e4:T.
+ ((mk_comp_num T e1 e2) ≠ (mk_comp_num T e3 e4)) →
+ ((e1 ≠ e3) ∨ (e2 ≠ e4))).
+ #T; #H; #e1; #e2; #e3; #e4;
+ nnormalize; #H1;
+ napply (or2_elim (e1 = e3) (e1 ≠ e3) ? (H e1 e3) …);
+ ##[ ##2: #H2; napply (or2_intro1 … H2)
+ ##| ##1: #H2; napply (or2_elim (e2 = e4) (e2 ≠ e4) ? (H e2 e4) …);
+ ##[ ##2: #H3; napply (or2_intro2 … H3)
+ ##| ##1: #H3; nrewrite > H2 in H1:(%);
+ nrewrite > H3;
+ #H1; nelim (H1 (refl_eq …))
+ ##]
+ ##]
+nqed.
+
+nlemma neq_to_neqcn :
+∀T.∀feq:T → T → bool.
+ (∀x,y:T.(x ≠ y) → (feq x y = false)) →
+ (∀x,y:T.decidable (x = y)) →
+ (∀b1,b2:comp_num T.
+ ((b1 ≠ b2) → (eq2_cn T feq b1 b2 = false))).
+ #T; #feq; #H; #H1; #b1; #b2;
+ nelim b1; #e1; #e2;
+ nelim b2; #e3; #e4;
+ #H2; nchange with (((feq e1 e3) ⊗ (feq e2 e4)) = false);
+ napply (or2_elim (e1 ≠ e3) (e2 ≠ e4) ? (cn_destruct T H1 e1 e2 e3 e4 … H2) …);
+ ##[ ##1: #H3; nrewrite > (H … H3); nnormalize; napply refl_eq
+ ##| ##2: #H3; nrewrite > (H … H3);
+ nrewrite > (symmetric_andbool (feq e1 e3) false);
+ nnormalize; napply refl_eq
+ ##]
+nqed.
include "num/quatern.ma".
include "num/oct.ma".
include "common/prod.ma".
-include "common/nat.ma".
(* *********** *)
(* ESADECIMALI *)
| xE: exadecim
| xF: exadecim.
+(* iteratore sugli esadecimali *)
+ndefinition forall_ex ≝ λP.
+ P x0 ⊗ P x1 ⊗ P x2 ⊗ P x3 ⊗ P x4 ⊗ P x5 ⊗ P x6 ⊗ P x7 ⊗
+ P x8 ⊗ P x9 ⊗ P xA ⊗ P xB ⊗ P xC ⊗ P xD ⊗ P xE ⊗ P xF.
+
(* operatore = *)
ndefinition eq_ex ≝
λe1,e2:exadecim.
| xC ⇒ x3 | xD ⇒ x2 | xE ⇒ x1 | xF ⇒ x0 ]
].
+(* operatore Most Significant Bit *)
+ndefinition getMSB_ex ≝
+λe:exadecim.match e with
+ [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
+ | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+ | x8 ⇒ true | x9 ⇒ true | xA ⇒ true | xB ⇒ true
+ | xC ⇒ true | xD ⇒ true | xE ⇒ true | xF ⇒ true ].
+
+ndefinition setMSB_ex ≝
+λe:exadecim.match e with
+ [ x0 ⇒ x8 | x1 ⇒ x9 | x2 ⇒ xA | x3 ⇒ xB
+ | x4 ⇒ xC | x5 ⇒ xD | x6 ⇒ xE | x7 ⇒ xF
+ | x8 ⇒ x8 | x9 ⇒ x9 | xA ⇒ xA | xB ⇒ xB
+ | xC ⇒ xC | xD ⇒ xD | xE ⇒ xE | xF ⇒ xF ].
+
+(* operatore Least Significant Bit *)
+ndefinition getLSB_ex ≝
+λe:exadecim.match e with
+ [ x0 ⇒ false | x1 ⇒ true | x2 ⇒ false | x3 ⇒ true
+ | x4 ⇒ false | x5 ⇒ true | x6 ⇒ false | x7 ⇒ true
+ | x8 ⇒ false | x9 ⇒ true | xA ⇒ false | xB ⇒ true
+ | xC ⇒ false | xD ⇒ true | xE ⇒ false | xF ⇒ true ].
+
+ndefinition setLSB_ex ≝
+λe:exadecim.match e with
+ [ x0 ⇒ x1 | x1 ⇒ x1 | x2 ⇒ x3 | x3 ⇒ x3
+ | x4 ⇒ x5 | x5 ⇒ x5 | x6 ⇒ x7 | x7 ⇒ x7
+ | x8 ⇒ x9 | x9 ⇒ x9 | xA ⇒ xB | xB ⇒ xB
+ | xC ⇒ xD | xD ⇒ xD | xE ⇒ xF | xF ⇒ xF ].
+
(* operatore rotazione destra con carry *)
ndefinition rcr_ex ≝
-λe:exadecim.λc:bool.match c with
+λc:bool.λe:exadecim.match c with
[ true ⇒ match e with
- [ x0 ⇒ pair exadecim bool x8 false | x1 ⇒ pair exadecim bool x8 true
- | x2 ⇒ pair exadecim bool x9 false | x3 ⇒ pair exadecim bool x9 true
- | x4 ⇒ pair exadecim bool xA false | x5 ⇒ pair exadecim bool xA true
- | x6 ⇒ pair exadecim bool xB false | x7 ⇒ pair exadecim bool xB true
- | x8 ⇒ pair exadecim bool xC false | x9 ⇒ pair exadecim bool xC true
- | xA ⇒ pair exadecim bool xD false | xB ⇒ pair exadecim bool xD true
- | xC ⇒ pair exadecim bool xE false | xD ⇒ pair exadecim bool xE true
- | xE ⇒ pair exadecim bool xF false | xF ⇒ pair exadecim bool xF true ]
+ [ x0 ⇒ pair … false x8 | x1 ⇒ pair … true x8
+ | x2 ⇒ pair … false x9 | x3 ⇒ pair … true x9
+ | x4 ⇒ pair … false xA | x5 ⇒ pair … true xA
+ | x6 ⇒ pair … false xB | x7 ⇒ pair … true xB
+ | x8 ⇒ pair … false xC | x9 ⇒ pair … true xC
+ | xA ⇒ pair … false xD | xB ⇒ pair … true xD
+ | xC ⇒ pair … false xE | xD ⇒ pair … true xE
+ | xE ⇒ pair … false xF | xF ⇒ pair … true xF ]
| false ⇒ match e with
- [ x0 ⇒ pair exadecim bool x0 false | x1 ⇒ pair exadecim bool x0 true
- | x2 ⇒ pair exadecim bool x1 false | x3 ⇒ pair exadecim bool x1 true
- | x4 ⇒ pair exadecim bool x2 false | x5 ⇒ pair exadecim bool x2 true
- | x6 ⇒ pair exadecim bool x3 false | x7 ⇒ pair exadecim bool x3 true
- | x8 ⇒ pair exadecim bool x4 false | x9 ⇒ pair exadecim bool x4 true
- | xA ⇒ pair exadecim bool x5 false | xB ⇒ pair exadecim bool x5 true
- | xC ⇒ pair exadecim bool x6 false | xD ⇒ pair exadecim bool x6 true
- | xE ⇒ pair exadecim bool x7 false | xF ⇒ pair exadecim bool x7 true ]
+ [ x0 ⇒ pair … false x0 | x1 ⇒ pair … true x0
+ | x2 ⇒ pair … false x1 | x3 ⇒ pair … true x1
+ | x4 ⇒ pair … false x2 | x5 ⇒ pair … true x2
+ | x6 ⇒ pair … false x3 | x7 ⇒ pair … true x3
+ | x8 ⇒ pair … false x4 | x9 ⇒ pair … true x4
+ | xA ⇒ pair … false x5 | xB ⇒ pair … true x5
+ | xC ⇒ pair … false x6 | xD ⇒ pair … true x6
+ | xE ⇒ pair … false x7 | xF ⇒ pair … true x7 ]
].
(* operatore shift destro *)
ndefinition shr_ex ≝
λe:exadecim.match e with
- [ x0 ⇒ pair exadecim bool x0 false | x1 ⇒ pair exadecim bool x0 true
- | x2 ⇒ pair exadecim bool x1 false | x3 ⇒ pair exadecim bool x1 true
- | x4 ⇒ pair exadecim bool x2 false | x5 ⇒ pair exadecim bool x2 true
- | x6 ⇒ pair exadecim bool x3 false | x7 ⇒ pair exadecim bool x3 true
- | x8 ⇒ pair exadecim bool x4 false | x9 ⇒ pair exadecim bool x4 true
- | xA ⇒ pair exadecim bool x5 false | xB ⇒ pair exadecim bool x5 true
- | xC ⇒ pair exadecim bool x6 false | xD ⇒ pair exadecim bool x6 true
- | xE ⇒ pair exadecim bool x7 false | xF ⇒ pair exadecim bool x7 true ].
+ [ x0 ⇒ pair … false x0 | x1 ⇒ pair … true x0
+ | x2 ⇒ pair … false x1 | x3 ⇒ pair … true x1
+ | x4 ⇒ pair … false x2 | x5 ⇒ pair … true x2
+ | x6 ⇒ pair … false x3 | x7 ⇒ pair … true x3
+ | x8 ⇒ pair … false x4 | x9 ⇒ pair … true x4
+ | xA ⇒ pair … false x5 | xB ⇒ pair … true x5
+ | xC ⇒ pair … false x6 | xD ⇒ pair … true x6
+ | xE ⇒ pair … false x7 | xF ⇒ pair … true x7 ].
(* operatore rotazione destra *)
ndefinition ror_ex ≝
| x8 ⇒ x4 | x9 ⇒ xC | xA ⇒ x5 | xB ⇒ xD
| xC ⇒ x6 | xD ⇒ xE | xE ⇒ x7 | xF ⇒ xF ].
-(* operatore rotazione destra n-volte *)
-nlet rec ror_ex_n (e:exadecim) (q:quatern) (rq:rec_quatern q) on rq ≝
- match rq with
- [ qu_O ⇒ e
- | qu_S q' n' ⇒ ror_ex_n (ror_ex e) q' n' ].
-
(* operatore rotazione sinistra con carry *)
ndefinition rcl_ex ≝
-λe:exadecim.λc:bool.match c with
+λc:bool.λe:exadecim.match c with
[ true ⇒ match e with
- [ x0 ⇒ pair exadecim bool x1 false | x1 ⇒ pair exadecim bool x3 false
- | x2 ⇒ pair exadecim bool x5 false | x3 ⇒ pair exadecim bool x7 false
- | x4 ⇒ pair exadecim bool x9 false | x5 ⇒ pair exadecim bool xB false
- | x6 ⇒ pair exadecim bool xD false | x7 ⇒ pair exadecim bool xF false
- | x8 ⇒ pair exadecim bool x1 true | x9 ⇒ pair exadecim bool x3 true
- | xA ⇒ pair exadecim bool x5 true | xB ⇒ pair exadecim bool x7 true
- | xC ⇒ pair exadecim bool x9 true | xD ⇒ pair exadecim bool xB true
- | xE ⇒ pair exadecim bool xD true | xF ⇒ pair exadecim bool xF true ]
+ [ x0 ⇒ pair … false x1 | x1 ⇒ pair … false x3
+ | x2 ⇒ pair … false x5 | x3 ⇒ pair … false x7
+ | x4 ⇒ pair … false x9 | x5 ⇒ pair … false xB
+ | x6 ⇒ pair … false xD | x7 ⇒ pair … false xF
+ | x8 ⇒ pair … true x1 | x9 ⇒ pair … true x3
+ | xA ⇒ pair … true x5 | xB ⇒ pair … true x7
+ | xC ⇒ pair … true x9 | xD ⇒ pair … true xB
+ | xE ⇒ pair … true xD | xF ⇒ pair … true xF ]
| false ⇒ match e with
- [ x0 ⇒ pair exadecim bool x0 false | x1 ⇒ pair exadecim bool x2 false
- | x2 ⇒ pair exadecim bool x4 false | x3 ⇒ pair exadecim bool x6 false
- | x4 ⇒ pair exadecim bool x8 false | x5 ⇒ pair exadecim bool xA false
- | x6 ⇒ pair exadecim bool xC false | x7 ⇒ pair exadecim bool xE false
- | x8 ⇒ pair exadecim bool x0 true | x9 ⇒ pair exadecim bool x2 true
- | xA ⇒ pair exadecim bool x4 true | xB ⇒ pair exadecim bool x6 true
- | xC ⇒ pair exadecim bool x8 true | xD ⇒ pair exadecim bool xA true
- | xE ⇒ pair exadecim bool xC true | xF ⇒ pair exadecim bool xE true ]
+ [ x0 ⇒ pair … false x0 | x1 ⇒ pair … false x2
+ | x2 ⇒ pair … false x4 | x3 ⇒ pair … false x6
+ | x4 ⇒ pair … false x8 | x5 ⇒ pair … false xA
+ | x6 ⇒ pair … false xC | x7 ⇒ pair … false xE
+ | x8 ⇒ pair … true x0 | x9 ⇒ pair … true x2
+ | xA ⇒ pair … true x4 | xB ⇒ pair … true x6
+ | xC ⇒ pair … true x8 | xD ⇒ pair … true xA
+ | xE ⇒ pair … true xC | xF ⇒ pair … true xE ]
].
(* operatore shift sinistro *)
ndefinition shl_ex ≝
λe:exadecim.match e with
- [ x0 ⇒ pair exadecim bool x0 false | x1 ⇒ pair exadecim bool x2 false
- | x2 ⇒ pair exadecim bool x4 false | x3 ⇒ pair exadecim bool x6 false
- | x4 ⇒ pair exadecim bool x8 false | x5 ⇒ pair exadecim bool xA false
- | x6 ⇒ pair exadecim bool xC false | x7 ⇒ pair exadecim bool xE false
- | x8 ⇒ pair exadecim bool x0 true | x9 ⇒ pair exadecim bool x2 true
- | xA ⇒ pair exadecim bool x4 true | xB ⇒ pair exadecim bool x6 true
- | xC ⇒ pair exadecim bool x8 true | xD ⇒ pair exadecim bool xA true
- | xE ⇒ pair exadecim bool xC true | xF ⇒ pair exadecim bool xE true ].
+ [ x0 ⇒ pair … false x0 | x1 ⇒ pair … false x2
+ | x2 ⇒ pair … false x4 | x3 ⇒ pair … false x6
+ | x4 ⇒ pair … false x8 | x5 ⇒ pair … false xA
+ | x6 ⇒ pair … false xC | x7 ⇒ pair … false xE
+ | x8 ⇒ pair … true x0 | x9 ⇒ pair … true x2
+ | xA ⇒ pair … true x4 | xB ⇒ pair … true x6
+ | xC ⇒ pair … true x8 | xD ⇒ pair … true xA
+ | xE ⇒ pair … true xC | xF ⇒ pair … true xE ].
(* operatore rotazione sinistra *)
ndefinition rol_ex ≝
| x8 ⇒ x1 | x9 ⇒ x3 | xA ⇒ x5 | xB ⇒ x7
| xC ⇒ x9 | xD ⇒ xB | xE ⇒ xD | xF ⇒ xF ].
-(* operatore rotazione sinistra n-volte *)
-nlet rec rol_ex_n (e:exadecim) (q:quatern) (rq:rec_quatern q) on rq ≝
- match rq with
- [ qu_O ⇒ e
- | qu_S q' n' ⇒ rol_ex_n (rol_ex e) q' n' ].
-
(* operatore not/complemento a 1 *)
ndefinition not_ex ≝
λe:exadecim.match e with
(* operatore somma con data+carry → data+carry *)
ndefinition plus_ex_dc_dc ≝
-λe1,e2:exadecim.λc:bool.
+λc:bool.λe1,e2:exadecim.
match c with
[ true ⇒ match e1 with
[ x0 ⇒ match e2 with
- [ x0 ⇒ pair … x1 false | x1 ⇒ pair … x2 false | x2 ⇒ pair … x3 false | x3 ⇒ pair … x4 false
- | x4 ⇒ pair … x5 false | x5 ⇒ pair … x6 false | x6 ⇒ pair … x7 false | x7 ⇒ pair … x8 false
- | x8 ⇒ pair … x9 false | x9 ⇒ pair … xA false | xA ⇒ pair … xB false | xB ⇒ pair … xC false
- | xC ⇒ pair … xD false | xD ⇒ pair … xE false | xE ⇒ pair … xF false | xF ⇒ pair … x0 true ]
+ [ x0 ⇒ pair … false x1 | x1 ⇒ pair … false x2 | x2 ⇒ pair … false x3 | x3 ⇒ pair … false x4
+ | x4 ⇒ pair … false x5 | x5 ⇒ pair … false x6 | x6 ⇒ pair … false x7 | x7 ⇒ pair … false x8
+ | x8 ⇒ pair … false x9 | x9 ⇒ pair … false xA | xA ⇒ pair … false xB | xB ⇒ pair … false xC
+ | xC ⇒ pair … false xD | xD ⇒ pair … false xE | xE ⇒ pair … false xF | xF ⇒ pair … true x0 ]
| x1 ⇒ match e2 with
- [ x0 ⇒ pair … x2 false | x1 ⇒ pair … x3 false | x2 ⇒ pair … x4 false | x3 ⇒ pair … x5 false
- | x4 ⇒ pair … x6 false | x5 ⇒ pair … x7 false | x6 ⇒ pair … x8 false | x7 ⇒ pair … x9 false
- | x8 ⇒ pair … xA false | x9 ⇒ pair … xB false | xA ⇒ pair … xC false | xB ⇒ pair … xD false
- | xC ⇒ pair … xE false | xD ⇒ pair … xF false | xE ⇒ pair … x0 true | xF ⇒ pair … x1 true ]
+ [ x0 ⇒ pair … false x2 | x1 ⇒ pair … false x3 | x2 ⇒ pair … false x4 | x3 ⇒ pair … false x5
+ | x4 ⇒ pair … false x6 | x5 ⇒ pair … false x7 | x6 ⇒ pair … false x8 | x7 ⇒ pair … false x9
+ | x8 ⇒ pair … false xA | x9 ⇒ pair … false xB | xA ⇒ pair … false xC | xB ⇒ pair … false xD
+ | xC ⇒ pair … false xE | xD ⇒ pair … false xF | xE ⇒ pair … true x0 | xF ⇒ pair … true x1 ]
| x2 ⇒ match e2 with
- [ x0 ⇒ pair … x3 false | x1 ⇒ pair … x4 false | x2 ⇒ pair … x5 false | x3 ⇒ pair … x6 false
- | x4 ⇒ pair … x7 false | x5 ⇒ pair … x8 false | x6 ⇒ pair … x9 false | x7 ⇒ pair … xA false
- | x8 ⇒ pair … xB false | x9 ⇒ pair … xC false | xA ⇒ pair … xD false | xB ⇒ pair … xE false
- | xC ⇒ pair … xF false | xD ⇒ pair … x0 true | xE ⇒ pair … x1 true | xF ⇒ pair … x2 true ]
+ [ x0 ⇒ pair … false x3 | x1 ⇒ pair … false x4 | x2 ⇒ pair … false x5 | x3 ⇒ pair … false x6
+ | x4 ⇒ pair … false x7 | x5 ⇒ pair … false x8 | x6 ⇒ pair … false x9 | x7 ⇒ pair … false xA
+ | x8 ⇒ pair … false xB | x9 ⇒ pair … false xC | xA ⇒ pair … false xD | xB ⇒ pair … false xE
+ | xC ⇒ pair … false xF | xD ⇒ pair … true x0 | xE ⇒ pair … true x1 | xF ⇒ pair … true x2 ]
| x3 ⇒ match e2 with
- [ x0 ⇒ pair … x4 false | x1 ⇒ pair … x5 false | x2 ⇒ pair … x6 false | x3 ⇒ pair … x7 false
- | x4 ⇒ pair … x8 false | x5 ⇒ pair … x9 false | x6 ⇒ pair … xA false | x7 ⇒ pair … xB false
- | x8 ⇒ pair … xC false | x9 ⇒ pair … xD false | xA ⇒ pair … xE false | xB ⇒ pair … xF false
- | xC ⇒ pair … x0 true | xD ⇒ pair … x1 true | xE ⇒ pair … x2 true | xF ⇒ pair … x3 true ]
+ [ x0 ⇒ pair … false x4 | x1 ⇒ pair … false x5 | x2 ⇒ pair … false x6 | x3 ⇒ pair … false x7
+ | x4 ⇒ pair … false x8 | x5 ⇒ pair … false x9 | x6 ⇒ pair … false xA | x7 ⇒ pair … false xB
+ | x8 ⇒ pair … false xC | x9 ⇒ pair … false xD | xA ⇒ pair … false xE | xB ⇒ pair … false xF
+ | xC ⇒ pair … true x0 | xD ⇒ pair … true x1 | xE ⇒ pair … true x2 | xF ⇒ pair … true x3 ]
| x4 ⇒ match e2 with
- [ x0 ⇒ pair … x5 false | x1 ⇒ pair … x6 false | x2 ⇒ pair … x7 false | x3 ⇒ pair … x8 false
- | x4 ⇒ pair … x9 false | x5 ⇒ pair … xA false | x6 ⇒ pair … xB false | x7 ⇒ pair … xC false
- | x8 ⇒ pair … xD false | x9 ⇒ pair … xE false | xA ⇒ pair … xF false | xB ⇒ pair … x0 true
- | xC ⇒ pair … x1 true | xD ⇒ pair … x2 true | xE ⇒ pair … x3 true | xF ⇒ pair … x4 true ]
+ [ x0 ⇒ pair … false x5 | x1 ⇒ pair … false x6 | x2 ⇒ pair … false x7 | x3 ⇒ pair … false x8
+ | x4 ⇒ pair … false x9 | x5 ⇒ pair … false xA | x6 ⇒ pair … false xB | x7 ⇒ pair … false xC
+ | x8 ⇒ pair … false xD | x9 ⇒ pair … false xE | xA ⇒ pair … false xF | xB ⇒ pair … true x0
+ | xC ⇒ pair … true x1 | xD ⇒ pair … true x2 | xE ⇒ pair … true x3 | xF ⇒ pair … true x4 ]
| x5 ⇒ match e2 with
- [ x0 ⇒ pair … x6 false | x1 ⇒ pair … x7 false | x2 ⇒ pair … x8 false | x3 ⇒ pair … x9 false
- | x4 ⇒ pair … xA false | x5 ⇒ pair … xB false | x6 ⇒ pair … xC false | x7 ⇒ pair … xD false
- | x8 ⇒ pair … xE false | x9 ⇒ pair … xF false | xA ⇒ pair … x0 true | xB ⇒ pair … x1 true
- | xC ⇒ pair … x2 true | xD ⇒ pair … x3 true | xE ⇒ pair … x4 true | xF ⇒ pair … x5 true ]
+ [ x0 ⇒ pair … false x6 | x1 ⇒ pair … false x7 | x2 ⇒ pair … false x8 | x3 ⇒ pair … false x9
+ | x4 ⇒ pair … false xA | x5 ⇒ pair … false xB | x6 ⇒ pair … false xC | x7 ⇒ pair … false xD
+ | x8 ⇒ pair … false xE | x9 ⇒ pair … false xF | xA ⇒ pair … true x0 | xB ⇒ pair … true x1
+ | xC ⇒ pair … true x2 | xD ⇒ pair … true x3 | xE ⇒ pair … true x4 | xF ⇒ pair … true x5 ]
| x6 ⇒ match e2 with
- [ x0 ⇒ pair … x7 false | x1 ⇒ pair … x8 false | x2 ⇒ pair … x9 false | x3 ⇒ pair … xA false
- | x4 ⇒ pair … xB false | x5 ⇒ pair … xC false | x6 ⇒ pair … xD false | x7 ⇒ pair … xE false
- | x8 ⇒ pair … xF false | x9 ⇒ pair … x0 true | xA ⇒ pair … x1 true | xB ⇒ pair … x2 true
- | xC ⇒ pair … x3 true | xD ⇒ pair … x4 true | xE ⇒ pair … x5 true | xF ⇒ pair … x6 true ]
+ [ x0 ⇒ pair … false x7 | x1 ⇒ pair … false x8 | x2 ⇒ pair … false x9 | x3 ⇒ pair … false xA
+ | x4 ⇒ pair … false xB | x5 ⇒ pair … false xC | x6 ⇒ pair … false xD | x7 ⇒ pair … false xE
+ | x8 ⇒ pair … false xF | x9 ⇒ pair … true x0 | xA ⇒ pair … true x1 | xB ⇒ pair … true x2
+ | xC ⇒ pair … true x3 | xD ⇒ pair … true x4 | xE ⇒ pair … true x5 | xF ⇒ pair … true x6 ]
| x7 ⇒ match e2 with
- [ x0 ⇒ pair … x8 false | x1 ⇒ pair … x9 false | x2 ⇒ pair … xA false | x3 ⇒ pair … xB false
- | x4 ⇒ pair … xC false | x5 ⇒ pair … xD false | x6 ⇒ pair … xE false | x7 ⇒ pair … xF false
- | x8 ⇒ pair … x0 true | x9 ⇒ pair … x1 true | xA ⇒ pair … x2 true | xB ⇒ pair … x3 true
- | xC ⇒ pair … x4 true | xD ⇒ pair … x5 true | xE ⇒ pair … x6 true | xF ⇒ pair … x7 true ]
+ [ x0 ⇒ pair … false x8 | x1 ⇒ pair … false x9 | x2 ⇒ pair … false xA | x3 ⇒ pair … false xB
+ | x4 ⇒ pair … false xC | x5 ⇒ pair … false xD | x6 ⇒ pair … false xE | x7 ⇒ pair … false xF
+ | x8 ⇒ pair … true x0 | x9 ⇒ pair … true x1 | xA ⇒ pair … true x2 | xB ⇒ pair … true x3
+ | xC ⇒ pair … true x4 | xD ⇒ pair … true x5 | xE ⇒ pair … true x6 | xF ⇒ pair … true x7 ]
| x8 ⇒ match e2 with
- [ x0 ⇒ pair … x9 false | x1 ⇒ pair … xA false | x2 ⇒ pair … xB false | x3 ⇒ pair … xC false
- | x4 ⇒ pair … xD false | x5 ⇒ pair … xE false | x6 ⇒ pair … xF false | x7 ⇒ pair … x0 true
- | x8 ⇒ pair … x1 true | x9 ⇒ pair … x2 true | xA ⇒ pair … x3 true | xB ⇒ pair … x4 true
- | xC ⇒ pair … x5 true | xD ⇒ pair … x6 true | xE ⇒ pair … x7 true | xF ⇒ pair … x8 true ]
+ [ x0 ⇒ pair … false x9 | x1 ⇒ pair … false xA | x2 ⇒ pair … false xB | x3 ⇒ pair … false xC
+ | x4 ⇒ pair … false xD | x5 ⇒ pair … false xE | x6 ⇒ pair … false xF | x7 ⇒ pair … true x0
+ | x8 ⇒ pair … true x1 | x9 ⇒ pair … true x2 | xA ⇒ pair … true x3 | xB ⇒ pair … true x4
+ | xC ⇒ pair … true x5 | xD ⇒ pair … true x6 | xE ⇒ pair … true x7 | xF ⇒ pair … true x8 ]
| x9 ⇒ match e2 with
- [ x0 ⇒ pair … xA false | x1 ⇒ pair … xB false | x2 ⇒ pair … xC false | x3 ⇒ pair … xD false
- | x4 ⇒ pair … xE false | x5 ⇒ pair … xF false | x6 ⇒ pair … x0 true | x7 ⇒ pair … x1 true
- | x8 ⇒ pair … x2 true | x9 ⇒ pair … x3 true | xA ⇒ pair … x4 true | xB ⇒ pair … x5 true
- | xC ⇒ pair … x6 true | xD ⇒ pair … x7 true | xE ⇒ pair … x8 true | xF ⇒ pair … x9 true ]
+ [ x0 ⇒ pair … false xA | x1 ⇒ pair … false xB | x2 ⇒ pair … false xC | x3 ⇒ pair … false xD
+ | x4 ⇒ pair … false xE | x5 ⇒ pair … false xF | x6 ⇒ pair … true x0 | x7 ⇒ pair … true x1
+ | x8 ⇒ pair … true x2 | x9 ⇒ pair … true x3 | xA ⇒ pair … true x4 | xB ⇒ pair … true x5
+ | xC ⇒ pair … true x6 | xD ⇒ pair … true x7 | xE ⇒ pair … true x8 | xF ⇒ pair … true x9 ]
| xA ⇒ match e2 with
- [ x0 ⇒ pair … xB false | x1 ⇒ pair … xC false | x2 ⇒ pair … xD false | x3 ⇒ pair … xE false
- | x4 ⇒ pair … xF false | x5 ⇒ pair … x0 true | x6 ⇒ pair … x1 true | x7 ⇒ pair … x2 true
- | x8 ⇒ pair … x3 true | x9 ⇒ pair … x4 true | xA ⇒ pair … x5 true | xB ⇒ pair … x6 true
- | xC ⇒ pair … x7 true | xD ⇒ pair … x8 true | xE ⇒ pair … x9 true | xF ⇒ pair … xA true ]
+ [ x0 ⇒ pair … false xB | x1 ⇒ pair … false xC | x2 ⇒ pair … false xD | x3 ⇒ pair … false xE
+ | x4 ⇒ pair … false xF | x5 ⇒ pair … true x0 | x6 ⇒ pair … true x1 | x7 ⇒ pair … true x2
+ | x8 ⇒ pair … true x3 | x9 ⇒ pair … true x4 | xA ⇒ pair … true x5 | xB ⇒ pair … true x6
+ | xC ⇒ pair … true x7 | xD ⇒ pair … true x8 | xE ⇒ pair … true x9 | xF ⇒ pair … true xA ]
| xB ⇒ match e2 with
- [ x0 ⇒ pair … xC false | x1 ⇒ pair … xD false | x2 ⇒ pair … xE false | x3 ⇒ pair … xF false
- | x4 ⇒ pair … x0 true | x5 ⇒ pair … x1 true | x6 ⇒ pair … x2 true | x7 ⇒ pair … x3 true
- | x8 ⇒ pair … x4 true | x9 ⇒ pair … x5 true | xA ⇒ pair … x6 true | xB ⇒ pair … x7 true
- | xC ⇒ pair … x8 true | xD ⇒ pair … x9 true | xE ⇒ pair … xA true | xF ⇒ pair … xB true ]
+ [ x0 ⇒ pair … false xC | x1 ⇒ pair … false xD | x2 ⇒ pair … false xE | x3 ⇒ pair … false xF
+ | x4 ⇒ pair … true x0 | x5 ⇒ pair … true x1 | x6 ⇒ pair … true x2 | x7 ⇒ pair … true x3
+ | x8 ⇒ pair … true x4 | x9 ⇒ pair … true x5 | xA ⇒ pair … true x6 | xB ⇒ pair … true x7
+ | xC ⇒ pair … true x8 | xD ⇒ pair … true x9 | xE ⇒ pair … true xA | xF ⇒ pair … true xB ]
| xC ⇒ match e2 with
- [ x0 ⇒ pair … xD false | x1 ⇒ pair … xE false | x2 ⇒ pair … xF false | x3 ⇒ pair … x0 true
- | x4 ⇒ pair … x1 true | x5 ⇒ pair … x2 true | x6 ⇒ pair … x3 true | x7 ⇒ pair … x4 true
- | x8 ⇒ pair … x5 true | x9 ⇒ pair … x6 true | xA ⇒ pair … x7 true | xB ⇒ pair … x8 true
- | xC ⇒ pair … x9 true | xD ⇒ pair … xA true | xE ⇒ pair … xB true | xF ⇒ pair … xC true ]
+ [ x0 ⇒ pair … false xD | x1 ⇒ pair … false xE | x2 ⇒ pair … false xF | x3 ⇒ pair … true x0
+ | x4 ⇒ pair … true x1 | x5 ⇒ pair … true x2 | x6 ⇒ pair … true x3 | x7 ⇒ pair … true x4
+ | x8 ⇒ pair … true x5 | x9 ⇒ pair … true x6 | xA ⇒ pair … true x7 | xB ⇒ pair … true x8
+ | xC ⇒ pair … true x9 | xD ⇒ pair … true xA | xE ⇒ pair … true xB | xF ⇒ pair … true xC ]
| xD ⇒ match e2 with
- [ x0 ⇒ pair … xE false | x1 ⇒ pair … xF false | x2 ⇒ pair … x0 true | x3 ⇒ pair … x1 true
- | x4 ⇒ pair … x2 true | x5 ⇒ pair … x3 true | x6 ⇒ pair … x4 true | x7 ⇒ pair … x5 true
- | x8 ⇒ pair … x6 true | x9 ⇒ pair … x7 true | xA ⇒ pair … x8 true | xB ⇒ pair … x9 true
- | xC ⇒ pair … xA true | xD ⇒ pair … xB true | xE ⇒ pair … xC true | xF ⇒ pair … xD true ]
+ [ x0 ⇒ pair … false xE | x1 ⇒ pair … false xF | x2 ⇒ pair … true x0 | x3 ⇒ pair … true x1
+ | x4 ⇒ pair … true x2 | x5 ⇒ pair … true x3 | x6 ⇒ pair … true x4 | x7 ⇒ pair … true x5
+ | x8 ⇒ pair … true x6 | x9 ⇒ pair … true x7 | xA ⇒ pair … true x8 | xB ⇒ pair … true x9
+ | xC ⇒ pair … true xA | xD ⇒ pair … true xB | xE ⇒ pair … true xC | xF ⇒ pair … true xD ]
| xE ⇒ match e2 with
- [ x0 ⇒ pair … xF false | x1 ⇒ pair … x0 true | x2 ⇒ pair … x1 true | x3 ⇒ pair … x2 true
- | x4 ⇒ pair … x3 true | x5 ⇒ pair … x4 true | x6 ⇒ pair … x5 true | x7 ⇒ pair … x6 true
- | x8 ⇒ pair … x7 true | x9 ⇒ pair … x8 true | xA ⇒ pair … x9 true | xB ⇒ pair … xA true
- | xC ⇒ pair … xB true | xD ⇒ pair … xC true | xE ⇒ pair … xD true | xF ⇒ pair … xE true ]
+ [ x0 ⇒ pair … false xF | x1 ⇒ pair … true x0 | x2 ⇒ pair … true x1 | x3 ⇒ pair … true x2
+ | x4 ⇒ pair … true x3 | x5 ⇒ pair … true x4 | x6 ⇒ pair … true x5 | x7 ⇒ pair … true x6
+ | x8 ⇒ pair … true x7 | x9 ⇒ pair … true x8 | xA ⇒ pair … true x9 | xB ⇒ pair … true xA
+ | xC ⇒ pair … true xB | xD ⇒ pair … true xC | xE ⇒ pair … true xD | xF ⇒ pair … true xE ]
| xF ⇒ match e2 with
- [ x0 ⇒ pair … x0 true | x1 ⇒ pair … x1 true | x2 ⇒ pair … x2 true | x3 ⇒ pair … x3 true
- | x4 ⇒ pair … x4 true | x5 ⇒ pair … x5 true | x6 ⇒ pair … x6 true | x7 ⇒ pair … x7 true
- | x8 ⇒ pair … x8 true | x9 ⇒ pair … x9 true | xA ⇒ pair … xA true | xB ⇒ pair … xB true
- | xC ⇒ pair … xC true | xD ⇒ pair … xD true | xE ⇒ pair … xE true | xF ⇒ pair … xF true ]
+ [ x0 ⇒ pair … true x0 | x1 ⇒ pair … true x1 | x2 ⇒ pair … true x2 | x3 ⇒ pair … true x3
+ | x4 ⇒ pair … true x4 | x5 ⇒ pair … true x5 | x6 ⇒ pair … true x6 | x7 ⇒ pair … true x7
+ | x8 ⇒ pair … true x8 | x9 ⇒ pair … true x9 | xA ⇒ pair … true xA | xB ⇒ pair … true xB
+ | xC ⇒ pair … true xC | xD ⇒ pair … true xD | xE ⇒ pair … true xE | xF ⇒ pair … true xF ]
]
| false ⇒ match e1 with
[ x0 ⇒ match e2 with
- [ x0 ⇒ pair … x0 false | x1 ⇒ pair … x1 false | x2 ⇒ pair … x2 false | x3 ⇒ pair … x3 false
- | x4 ⇒ pair … x4 false | x5 ⇒ pair … x5 false | x6 ⇒ pair … x6 false | x7 ⇒ pair … x7 false
- | x8 ⇒ pair … x8 false | x9 ⇒ pair … x9 false | xA ⇒ pair … xA false | xB ⇒ pair … xB false
- | xC ⇒ pair … xC false | xD ⇒ pair … xD false | xE ⇒ pair … xE false | xF ⇒ pair … xF false ]
+ [ x0 ⇒ pair … false x0 | x1 ⇒ pair … false x1 | x2 ⇒ pair … false x2 | x3 ⇒ pair … false x3
+ | x4 ⇒ pair … false x4 | x5 ⇒ pair … false x5 | x6 ⇒ pair … false x6 | x7 ⇒ pair … false x7
+ | x8 ⇒ pair … false x8 | x9 ⇒ pair … false x9 | xA ⇒ pair … false xA | xB ⇒ pair … false xB
+ | xC ⇒ pair … false xC | xD ⇒ pair … false xD | xE ⇒ pair … false xE | xF ⇒ pair … false xF ]
| x1 ⇒ match e2 with
- [ x0 ⇒ pair … x1 false | x1 ⇒ pair … x2 false | x2 ⇒ pair … x3 false | x3 ⇒ pair … x4 false
- | x4 ⇒ pair … x5 false | x5 ⇒ pair … x6 false | x6 ⇒ pair … x7 false | x7 ⇒ pair … x8 false
- | x8 ⇒ pair … x9 false | x9 ⇒ pair … xA false | xA ⇒ pair … xB false | xB ⇒ pair … xC false
- | xC ⇒ pair … xD false | xD ⇒ pair … xE false | xE ⇒ pair … xF false | xF ⇒ pair … x0 true ]
+ [ x0 ⇒ pair … false x1 | x1 ⇒ pair … false x2 | x2 ⇒ pair … false x3 | x3 ⇒ pair … false x4
+ | x4 ⇒ pair … false x5 | x5 ⇒ pair … false x6 | x6 ⇒ pair … false x7 | x7 ⇒ pair … false x8
+ | x8 ⇒ pair … false x9 | x9 ⇒ pair … false xA | xA ⇒ pair … false xB | xB ⇒ pair … false xC
+ | xC ⇒ pair … false xD | xD ⇒ pair … false xE | xE ⇒ pair … false xF | xF ⇒ pair … true x0 ]
| x2 ⇒ match e2 with
- [ x0 ⇒ pair … x2 false | x1 ⇒ pair … x3 false | x2 ⇒ pair … x4 false | x3 ⇒ pair … x5 false
- | x4 ⇒ pair … x6 false | x5 ⇒ pair … x7 false | x6 ⇒ pair … x8 false | x7 ⇒ pair … x9 false
- | x8 ⇒ pair … xA false | x9 ⇒ pair … xB false | xA ⇒ pair … xC false | xB ⇒ pair … xD false
- | xC ⇒ pair … xE false | xD ⇒ pair … xF false | xE ⇒ pair … x0 true | xF ⇒ pair … x1 true ]
+ [ x0 ⇒ pair … false x2 | x1 ⇒ pair … false x3 | x2 ⇒ pair … false x4 | x3 ⇒ pair … false x5
+ | x4 ⇒ pair … false x6 | x5 ⇒ pair … false x7 | x6 ⇒ pair … false x8 | x7 ⇒ pair … false x9
+ | x8 ⇒ pair … false xA | x9 ⇒ pair … false xB | xA ⇒ pair … false xC | xB ⇒ pair … false xD
+ | xC ⇒ pair … false xE | xD ⇒ pair … false xF | xE ⇒ pair … true x0 | xF ⇒ pair … true x1 ]
| x3 ⇒ match e2 with
- [ x0 ⇒ pair … x3 false | x1 ⇒ pair … x4 false | x2 ⇒ pair … x5 false | x3 ⇒ pair … x6 false
- | x4 ⇒ pair … x7 false | x5 ⇒ pair … x8 false | x6 ⇒ pair … x9 false | x7 ⇒ pair … xA false
- | x8 ⇒ pair … xB false | x9 ⇒ pair … xC false | xA ⇒ pair … xD false | xB ⇒ pair … xE false
- | xC ⇒ pair … xF false | xD ⇒ pair … x0 true | xE ⇒ pair … x1 true | xF ⇒ pair … x2 true ]
+ [ x0 ⇒ pair … false x3 | x1 ⇒ pair … false x4 | x2 ⇒ pair … false x5 | x3 ⇒ pair … false x6
+ | x4 ⇒ pair … false x7 | x5 ⇒ pair … false x8 | x6 ⇒ pair … false x9 | x7 ⇒ pair … false xA
+ | x8 ⇒ pair … false xB | x9 ⇒ pair … false xC | xA ⇒ pair … false xD | xB ⇒ pair … false xE
+ | xC ⇒ pair … false xF | xD ⇒ pair … true x0 | xE ⇒ pair … true x1 | xF ⇒ pair … true x2 ]
| x4 ⇒ match e2 with
- [ x0 ⇒ pair … x4 false | x1 ⇒ pair … x5 false | x2 ⇒ pair … x6 false | x3 ⇒ pair … x7 false
- | x4 ⇒ pair … x8 false | x5 ⇒ pair … x9 false | x6 ⇒ pair … xA false | x7 ⇒ pair … xB false
- | x8 ⇒ pair … xC false | x9 ⇒ pair … xD false | xA ⇒ pair … xE false | xB ⇒ pair … xF false
- | xC ⇒ pair … x0 true | xD ⇒ pair … x1 true | xE ⇒ pair … x2 true | xF ⇒ pair … x3 true ]
+ [ x0 ⇒ pair … false x4 | x1 ⇒ pair … false x5 | x2 ⇒ pair … false x6 | x3 ⇒ pair … false x7
+ | x4 ⇒ pair … false x8 | x5 ⇒ pair … false x9 | x6 ⇒ pair … false xA | x7 ⇒ pair … false xB
+ | x8 ⇒ pair … false xC | x9 ⇒ pair … false xD | xA ⇒ pair … false xE | xB ⇒ pair … false xF
+ | xC ⇒ pair … true x0 | xD ⇒ pair … true x1 | xE ⇒ pair … true x2 | xF ⇒ pair … true x3 ]
| x5 ⇒ match e2 with
- [ x0 ⇒ pair … x5 false | x1 ⇒ pair … x6 false | x2 ⇒ pair … x7 false | x3 ⇒ pair … x8 false
- | x4 ⇒ pair … x9 false | x5 ⇒ pair … xA false | x6 ⇒ pair … xB false | x7 ⇒ pair … xC false
- | x8 ⇒ pair … xD false | x9 ⇒ pair … xE false | xA ⇒ pair … xF false | xB ⇒ pair … x0 true
- | xC ⇒ pair … x1 true | xD ⇒ pair … x2 true | xE ⇒ pair … x3 true | xF ⇒ pair … x4 true ]
+ [ x0 ⇒ pair … false x5 | x1 ⇒ pair … false x6 | x2 ⇒ pair … false x7 | x3 ⇒ pair … false x8
+ | x4 ⇒ pair … false x9 | x5 ⇒ pair … false xA | x6 ⇒ pair … false xB | x7 ⇒ pair … false xC
+ | x8 ⇒ pair … false xD | x9 ⇒ pair … false xE | xA ⇒ pair … false xF | xB ⇒ pair … true x0
+ | xC ⇒ pair … true x1 | xD ⇒ pair … true x2 | xE ⇒ pair … true x3 | xF ⇒ pair … true x4 ]
| x6 ⇒ match e2 with
- [ x0 ⇒ pair … x6 false | x1 ⇒ pair … x7 false | x2 ⇒ pair … x8 false | x3 ⇒ pair … x9 false
- | x4 ⇒ pair … xA false | x5 ⇒ pair … xB false | x6 ⇒ pair … xC false | x7 ⇒ pair … xD false
- | x8 ⇒ pair … xE false | x9 ⇒ pair … xF false | xA ⇒ pair … x0 true | xB ⇒ pair … x1 true
- | xC ⇒ pair … x2 true | xD ⇒ pair … x3 true | xE ⇒ pair … x4 true | xF ⇒ pair … x5 true ]
+ [ x0 ⇒ pair … false x6 | x1 ⇒ pair … false x7 | x2 ⇒ pair … false x8 | x3 ⇒ pair … false x9
+ | x4 ⇒ pair … false xA | x5 ⇒ pair … false xB | x6 ⇒ pair … false xC | x7 ⇒ pair … false xD
+ | x8 ⇒ pair … false xE | x9 ⇒ pair … false xF | xA ⇒ pair … true x0 | xB ⇒ pair … true x1
+ | xC ⇒ pair … true x2 | xD ⇒ pair … true x3 | xE ⇒ pair … true x4 | xF ⇒ pair … true x5 ]
| x7 ⇒ match e2 with
- [ x0 ⇒ pair … x7 false | x1 ⇒ pair … x8 false | x2 ⇒ pair … x9 false | x3 ⇒ pair … xA false
- | x4 ⇒ pair … xB false | x5 ⇒ pair … xC false | x6 ⇒ pair … xD false | x7 ⇒ pair … xE false
- | x8 ⇒ pair … xF false | x9 ⇒ pair … x0 true | xA ⇒ pair … x1 true | xB ⇒ pair … x2 true
- | xC ⇒ pair … x3 true | xD ⇒ pair … x4 true | xE ⇒ pair … x5 true | xF ⇒ pair … x6 true ]
+ [ x0 ⇒ pair … false x7 | x1 ⇒ pair … false x8 | x2 ⇒ pair … false x9 | x3 ⇒ pair … false xA
+ | x4 ⇒ pair … false xB | x5 ⇒ pair … false xC | x6 ⇒ pair … false xD | x7 ⇒ pair … false xE
+ | x8 ⇒ pair … false xF | x9 ⇒ pair … true x0 | xA ⇒ pair … true x1 | xB ⇒ pair … true x2
+ | xC ⇒ pair … true x3 | xD ⇒ pair … true x4 | xE ⇒ pair … true x5 | xF ⇒ pair … true x6 ]
| x8 ⇒ match e2 with
- [ x0 ⇒ pair … x8 false | x1 ⇒ pair … x9 false | x2 ⇒ pair … xA false | x3 ⇒ pair … xB false
- | x4 ⇒ pair … xC false | x5 ⇒ pair … xD false | x6 ⇒ pair … xE false | x7 ⇒ pair … xF false
- | x8 ⇒ pair … x0 true | x9 ⇒ pair … x1 true | xA ⇒ pair … x2 true | xB ⇒ pair … x3 true
- | xC ⇒ pair … x4 true | xD ⇒ pair … x5 true | xE ⇒ pair … x6 true | xF ⇒ pair … x7 true ]
+ [ x0 ⇒ pair … false x8 | x1 ⇒ pair … false x9 | x2 ⇒ pair … false xA | x3 ⇒ pair … false xB
+ | x4 ⇒ pair … false xC | x5 ⇒ pair … false xD | x6 ⇒ pair … false xE | x7 ⇒ pair … false xF
+ | x8 ⇒ pair … true x0 | x9 ⇒ pair … true x1 | xA ⇒ pair … true x2 | xB ⇒ pair … true x3
+ | xC ⇒ pair … true x4 | xD ⇒ pair … true x5 | xE ⇒ pair … true x6 | xF ⇒ pair … true x7 ]
| x9 ⇒ match e2 with
- [ x0 ⇒ pair … x9 false | x1 ⇒ pair … xA false | x2 ⇒ pair … xB false | x3 ⇒ pair … xC false
- | x4 ⇒ pair … xD false | x5 ⇒ pair … xE false | x6 ⇒ pair … xF false | x7 ⇒ pair … x0 true
- | x8 ⇒ pair … x1 true | x9 ⇒ pair … x2 true | xA ⇒ pair … x3 true | xB ⇒ pair … x4 true
- | xC ⇒ pair … x5 true | xD ⇒ pair … x6 true | xE ⇒ pair … x7 true | xF ⇒ pair … x8 true ]
+ [ x0 ⇒ pair … false x9 | x1 ⇒ pair … false xA | x2 ⇒ pair … false xB | x3 ⇒ pair … false xC
+ | x4 ⇒ pair … false xD | x5 ⇒ pair … false xE | x6 ⇒ pair … false xF | x7 ⇒ pair … true x0
+ | x8 ⇒ pair … true x1 | x9 ⇒ pair … true x2 | xA ⇒ pair … true x3 | xB ⇒ pair … true x4
+ | xC ⇒ pair … true x5 | xD ⇒ pair … true x6 | xE ⇒ pair … true x7 | xF ⇒ pair … true x8 ]
| xA ⇒ match e2 with
- [ x0 ⇒ pair … xA false | x1 ⇒ pair … xB false | x2 ⇒ pair … xC false | x3 ⇒ pair … xD false
- | x4 ⇒ pair … xE false | x5 ⇒ pair … xF false | x6 ⇒ pair … x0 true | x7 ⇒ pair … x1 true
- | x8 ⇒ pair … x2 true | x9 ⇒ pair … x3 true | xA ⇒ pair … x4 true | xB ⇒ pair … x5 true
- | xC ⇒ pair … x6 true | xD ⇒ pair … x7 true | xE ⇒ pair … x8 true | xF ⇒ pair … x9 true ]
+ [ x0 ⇒ pair … false xA | x1 ⇒ pair … false xB | x2 ⇒ pair … false xC | x3 ⇒ pair … false xD
+ | x4 ⇒ pair … false xE | x5 ⇒ pair … false xF | x6 ⇒ pair … true x0 | x7 ⇒ pair … true x1
+ | x8 ⇒ pair … true x2 | x9 ⇒ pair … true x3 | xA ⇒ pair … true x4 | xB ⇒ pair … true x5
+ | xC ⇒ pair … true x6 | xD ⇒ pair … true x7 | xE ⇒ pair … true x8 | xF ⇒ pair … true x9 ]
| xB ⇒ match e2 with
- [ x0 ⇒ pair … xB false | x1 ⇒ pair … xC false | x2 ⇒ pair … xD false | x3 ⇒ pair … xE false
- | x4 ⇒ pair … xF false | x5 ⇒ pair … x0 true | x6 ⇒ pair … x1 true | x7 ⇒ pair … x2 true
- | x8 ⇒ pair … x3 true | x9 ⇒ pair … x4 true | xA ⇒ pair … x5 true | xB ⇒ pair … x6 true
- | xC ⇒ pair … x7 true | xD ⇒ pair … x8 true | xE ⇒ pair … x9 true | xF ⇒ pair … xA true ]
+ [ x0 ⇒ pair … false xB | x1 ⇒ pair … false xC | x2 ⇒ pair … false xD | x3 ⇒ pair … false xE
+ | x4 ⇒ pair … false xF | x5 ⇒ pair … true x0 | x6 ⇒ pair … true x1 | x7 ⇒ pair … true x2
+ | x8 ⇒ pair … true x3 | x9 ⇒ pair … true x4 | xA ⇒ pair … true x5 | xB ⇒ pair … true x6
+ | xC ⇒ pair … true x7 | xD ⇒ pair … true x8 | xE ⇒ pair … true x9 | xF ⇒ pair … true xA ]
| xC ⇒ match e2 with
- [ x0 ⇒ pair … xC false | x1 ⇒ pair … xD false | x2 ⇒ pair … xE false | x3 ⇒ pair … xF false
- | x4 ⇒ pair … x0 true | x5 ⇒ pair … x1 true | x6 ⇒ pair … x2 true | x7 ⇒ pair … x3 true
- | x8 ⇒ pair … x4 true | x9 ⇒ pair … x5 true | xA ⇒ pair … x6 true | xB ⇒ pair … x7 true
- | xC ⇒ pair … x8 true | xD ⇒ pair … x9 true | xE ⇒ pair … xA true | xF ⇒ pair … xB true ]
+ [ x0 ⇒ pair … false xC | x1 ⇒ pair … false xD | x2 ⇒ pair … false xE | x3 ⇒ pair … false xF
+ | x4 ⇒ pair … true x0 | x5 ⇒ pair … true x1 | x6 ⇒ pair … true x2 | x7 ⇒ pair … true x3
+ | x8 ⇒ pair … true x4 | x9 ⇒ pair … true x5 | xA ⇒ pair … true x6 | xB ⇒ pair … true x7
+ | xC ⇒ pair … true x8 | xD ⇒ pair … true x9 | xE ⇒ pair … true xA | xF ⇒ pair … true xB ]
| xD ⇒ match e2 with
- [ x0 ⇒ pair … xD false | x1 ⇒ pair … xE false | x2 ⇒ pair … xF false | x3 ⇒ pair … x0 true
- | x4 ⇒ pair … x1 true | x5 ⇒ pair … x2 true | x6 ⇒ pair … x3 true | x7 ⇒ pair … x4 true
- | x8 ⇒ pair … x5 true | x9 ⇒ pair … x6 true | xA ⇒ pair … x7 true | xB ⇒ pair … x8 true
- | xC ⇒ pair … x9 true | xD ⇒ pair … xA true | xE ⇒ pair … xB true | xF ⇒ pair … xC true ]
+ [ x0 ⇒ pair … false xD | x1 ⇒ pair … false xE | x2 ⇒ pair … false xF | x3 ⇒ pair … true x0
+ | x4 ⇒ pair … true x1 | x5 ⇒ pair … true x2 | x6 ⇒ pair … true x3 | x7 ⇒ pair … true x4
+ | x8 ⇒ pair … true x5 | x9 ⇒ pair … true x6 | xA ⇒ pair … true x7 | xB ⇒ pair … true x8
+ | xC ⇒ pair … true x9 | xD ⇒ pair … true xA | xE ⇒ pair … true xB | xF ⇒ pair … true xC ]
| xE ⇒ match e2 with
- [ x0 ⇒ pair … xE false | x1 ⇒ pair … xF false | x2 ⇒ pair … x0 true | x3 ⇒ pair … x1 true
- | x4 ⇒ pair … x2 true | x5 ⇒ pair … x3 true | x6 ⇒ pair … x4 true | x7 ⇒ pair … x5 true
- | x8 ⇒ pair … x6 true | x9 ⇒ pair … x7 true | xA ⇒ pair … x8 true | xB ⇒ pair … x9 true
- | xC ⇒ pair … xA true | xD ⇒ pair … xB true | xE ⇒ pair … xC true | xF ⇒ pair … xD true ]
+ [ x0 ⇒ pair … false xE | x1 ⇒ pair … false xF | x2 ⇒ pair … true x0 | x3 ⇒ pair … true x1
+ | x4 ⇒ pair … true x2 | x5 ⇒ pair … true x3 | x6 ⇒ pair … true x4 | x7 ⇒ pair … true x5
+ | x8 ⇒ pair … true x6 | x9 ⇒ pair … true x7 | xA ⇒ pair … true x8 | xB ⇒ pair … true x9
+ | xC ⇒ pair … true xA | xD ⇒ pair … true xB | xE ⇒ pair … true xC | xF ⇒ pair … true xD ]
| xF ⇒ match e2 with
- [ x0 ⇒ pair … xF false | x1 ⇒ pair … x0 true | x2 ⇒ pair … x1 true | x3 ⇒ pair … x2 true
- | x4 ⇒ pair … x3 true | x5 ⇒ pair … x4 true | x6 ⇒ pair … x5 true | x7 ⇒ pair … x6 true
- | x8 ⇒ pair … x7 true | x9 ⇒ pair … x8 true | xA ⇒ pair … x9 true | xB ⇒ pair … xA true
- | xC ⇒ pair … xB true | xD ⇒ pair … xC true | xE ⇒ pair … xD true | xF ⇒ pair … xE true ]
+ [ x0 ⇒ pair … false xF | x1 ⇒ pair … true x0 | x2 ⇒ pair … true x1 | x3 ⇒ pair … true x2
+ | x4 ⇒ pair … true x3 | x5 ⇒ pair … true x4 | x6 ⇒ pair … true x5 | x7 ⇒ pair … true x6
+ | x8 ⇒ pair … true x7 | x9 ⇒ pair … true x8 | xA ⇒ pair … true x9 | xB ⇒ pair … true xA
+ | xC ⇒ pair … true xB | xD ⇒ pair … true xC | xE ⇒ pair … true xD | xF ⇒ pair … true xE ]
]].
(* operatore somma con data → data+carry *)
λe1,e2:exadecim.
match e1 with
[ x0 ⇒ match e2 with
- [ x0 ⇒ pair … x0 false | x1 ⇒ pair … x1 false | x2 ⇒ pair … x2 false | x3 ⇒ pair … x3 false
- | x4 ⇒ pair … x4 false | x5 ⇒ pair … x5 false | x6 ⇒ pair … x6 false | x7 ⇒ pair … x7 false
- | x8 ⇒ pair … x8 false | x9 ⇒ pair … x9 false | xA ⇒ pair … xA false | xB ⇒ pair … xB false
- | xC ⇒ pair … xC false | xD ⇒ pair … xD false | xE ⇒ pair … xE false | xF ⇒ pair … xF false ]
+ [ x0 ⇒ pair … false x0 | x1 ⇒ pair … false x1 | x2 ⇒ pair … false x2 | x3 ⇒ pair … false x3
+ | x4 ⇒ pair … false x4 | x5 ⇒ pair … false x5 | x6 ⇒ pair … false x6 | x7 ⇒ pair … false x7
+ | x8 ⇒ pair … false x8 | x9 ⇒ pair … false x9 | xA ⇒ pair … false xA | xB ⇒ pair … false xB
+ | xC ⇒ pair … false xC | xD ⇒ pair … false xD | xE ⇒ pair … false xE | xF ⇒ pair … false xF ]
| x1 ⇒ match e2 with
- [ x0 ⇒ pair … x1 false | x1 ⇒ pair … x2 false | x2 ⇒ pair … x3 false | x3 ⇒ pair … x4 false
- | x4 ⇒ pair … x5 false | x5 ⇒ pair … x6 false | x6 ⇒ pair … x7 false | x7 ⇒ pair … x8 false
- | x8 ⇒ pair … x9 false | x9 ⇒ pair … xA false | xA ⇒ pair … xB false | xB ⇒ pair … xC false
- | xC ⇒ pair … xD false | xD ⇒ pair … xE false | xE ⇒ pair … xF false | xF ⇒ pair … x0 true ]
+ [ x0 ⇒ pair … false x1 | x1 ⇒ pair … false x2 | x2 ⇒ pair … false x3 | x3 ⇒ pair … false x4
+ | x4 ⇒ pair … false x5 | x5 ⇒ pair … false x6 | x6 ⇒ pair … false x7 | x7 ⇒ pair … false x8
+ | x8 ⇒ pair … false x9 | x9 ⇒ pair … false xA | xA ⇒ pair … false xB | xB ⇒ pair … false xC
+ | xC ⇒ pair … false xD | xD ⇒ pair … false xE | xE ⇒ pair … false xF | xF ⇒ pair … true x0 ]
| x2 ⇒ match e2 with
- [ x0 ⇒ pair … x2 false | x1 ⇒ pair … x3 false | x2 ⇒ pair … x4 false | x3 ⇒ pair … x5 false
- | x4 ⇒ pair … x6 false | x5 ⇒ pair … x7 false | x6 ⇒ pair … x8 false | x7 ⇒ pair … x9 false
- | x8 ⇒ pair … xA false | x9 ⇒ pair … xB false | xA ⇒ pair … xC false | xB ⇒ pair … xD false
- | xC ⇒ pair … xE false | xD ⇒ pair … xF false | xE ⇒ pair … x0 true | xF ⇒ pair … x1 true ]
+ [ x0 ⇒ pair … false x2 | x1 ⇒ pair … false x3 | x2 ⇒ pair … false x4 | x3 ⇒ pair … false x5
+ | x4 ⇒ pair … false x6 | x5 ⇒ pair … false x7 | x6 ⇒ pair … false x8 | x7 ⇒ pair … false x9
+ | x8 ⇒ pair … false xA | x9 ⇒ pair … false xB | xA ⇒ pair … false xC | xB ⇒ pair … false xD
+ | xC ⇒ pair … false xE | xD ⇒ pair … false xF | xE ⇒ pair … true x0 | xF ⇒ pair … true x1 ]
| x3 ⇒ match e2 with
- [ x0 ⇒ pair … x3 false | x1 ⇒ pair … x4 false | x2 ⇒ pair … x5 false | x3 ⇒ pair … x6 false
- | x4 ⇒ pair … x7 false | x5 ⇒ pair … x8 false | x6 ⇒ pair … x9 false | x7 ⇒ pair … xA false
- | x8 ⇒ pair … xB false | x9 ⇒ pair … xC false | xA ⇒ pair … xD false | xB ⇒ pair … xE false
- | xC ⇒ pair … xF false | xD ⇒ pair … x0 true | xE ⇒ pair … x1 true | xF ⇒ pair … x2 true ]
+ [ x0 ⇒ pair … false x3 | x1 ⇒ pair … false x4 | x2 ⇒ pair … false x5 | x3 ⇒ pair … false x6
+ | x4 ⇒ pair … false x7 | x5 ⇒ pair … false x8 | x6 ⇒ pair … false x9 | x7 ⇒ pair … false xA
+ | x8 ⇒ pair … false xB | x9 ⇒ pair … false xC | xA ⇒ pair … false xD | xB ⇒ pair … false xE
+ | xC ⇒ pair … false xF | xD ⇒ pair … true x0 | xE ⇒ pair … true x1 | xF ⇒ pair … true x2 ]
| x4 ⇒ match e2 with
- [ x0 ⇒ pair … x4 false | x1 ⇒ pair … x5 false | x2 ⇒ pair … x6 false | x3 ⇒ pair … x7 false
- | x4 ⇒ pair … x8 false | x5 ⇒ pair … x9 false | x6 ⇒ pair … xA false | x7 ⇒ pair … xB false
- | x8 ⇒ pair … xC false | x9 ⇒ pair … xD false | xA ⇒ pair … xE false | xB ⇒ pair … xF false
- | xC ⇒ pair … x0 true | xD ⇒ pair … x1 true | xE ⇒ pair … x2 true | xF ⇒ pair … x3 true ]
+ [ x0 ⇒ pair … false x4 | x1 ⇒ pair … false x5 | x2 ⇒ pair … false x6 | x3 ⇒ pair … false x7
+ | x4 ⇒ pair … false x8 | x5 ⇒ pair … false x9 | x6 ⇒ pair … false xA | x7 ⇒ pair … false xB
+ | x8 ⇒ pair … false xC | x9 ⇒ pair … false xD | xA ⇒ pair … false xE | xB ⇒ pair … false xF
+ | xC ⇒ pair … true x0 | xD ⇒ pair … true x1 | xE ⇒ pair … true x2 | xF ⇒ pair … true x3 ]
| x5 ⇒ match e2 with
- [ x0 ⇒ pair … x5 false | x1 ⇒ pair … x6 false | x2 ⇒ pair … x7 false | x3 ⇒ pair … x8 false
- | x4 ⇒ pair … x9 false | x5 ⇒ pair … xA false | x6 ⇒ pair … xB false | x7 ⇒ pair … xC false
- | x8 ⇒ pair … xD false | x9 ⇒ pair … xE false | xA ⇒ pair … xF false | xB ⇒ pair … x0 true
- | xC ⇒ pair … x1 true | xD ⇒ pair … x2 true | xE ⇒ pair … x3 true | xF ⇒ pair … x4 true ]
+ [ x0 ⇒ pair … false x5 | x1 ⇒ pair … false x6 | x2 ⇒ pair … false x7 | x3 ⇒ pair … false x8
+ | x4 ⇒ pair … false x9 | x5 ⇒ pair … false xA | x6 ⇒ pair … false xB | x7 ⇒ pair … false xC
+ | x8 ⇒ pair … false xD | x9 ⇒ pair … false xE | xA ⇒ pair … false xF | xB ⇒ pair … true x0
+ | xC ⇒ pair … true x1 | xD ⇒ pair … true x2 | xE ⇒ pair … true x3 | xF ⇒ pair … true x4 ]
| x6 ⇒ match e2 with
- [ x0 ⇒ pair … x6 false | x1 ⇒ pair … x7 false | x2 ⇒ pair … x8 false | x3 ⇒ pair … x9 false
- | x4 ⇒ pair … xA false | x5 ⇒ pair … xB false | x6 ⇒ pair … xC false | x7 ⇒ pair … xD false
- | x8 ⇒ pair … xE false | x9 ⇒ pair … xF false | xA ⇒ pair … x0 true | xB ⇒ pair … x1 true
- | xC ⇒ pair … x2 true | xD ⇒ pair … x3 true | xE ⇒ pair … x4 true | xF ⇒ pair … x5 true ]
+ [ x0 ⇒ pair … false x6 | x1 ⇒ pair … false x7 | x2 ⇒ pair … false x8 | x3 ⇒ pair … false x9
+ | x4 ⇒ pair … false xA | x5 ⇒ pair … false xB | x6 ⇒ pair … false xC | x7 ⇒ pair … false xD
+ | x8 ⇒ pair … false xE | x9 ⇒ pair … false xF | xA ⇒ pair … true x0 | xB ⇒ pair … true x1
+ | xC ⇒ pair … true x2 | xD ⇒ pair … true x3 | xE ⇒ pair … true x4 | xF ⇒ pair … true x5 ]
| x7 ⇒ match e2 with
- [ x0 ⇒ pair … x7 false | x1 ⇒ pair … x8 false | x2 ⇒ pair … x9 false | x3 ⇒ pair … xA false
- | x4 ⇒ pair … xB false | x5 ⇒ pair … xC false | x6 ⇒ pair … xD false | x7 ⇒ pair … xE false
- | x8 ⇒ pair … xF false | x9 ⇒ pair … x0 true | xA ⇒ pair … x1 true | xB ⇒ pair … x2 true
- | xC ⇒ pair … x3 true | xD ⇒ pair … x4 true | xE ⇒ pair … x5 true | xF ⇒ pair … x6 true ]
+ [ x0 ⇒ pair … false x7 | x1 ⇒ pair … false x8 | x2 ⇒ pair … false x9 | x3 ⇒ pair … false xA
+ | x4 ⇒ pair … false xB | x5 ⇒ pair … false xC | x6 ⇒ pair … false xD | x7 ⇒ pair … false xE
+ | x8 ⇒ pair … false xF | x9 ⇒ pair … true x0 | xA ⇒ pair … true x1 | xB ⇒ pair … true x2
+ | xC ⇒ pair … true x3 | xD ⇒ pair … true x4 | xE ⇒ pair … true x5 | xF ⇒ pair … true x6 ]
| x8 ⇒ match e2 with
- [ x0 ⇒ pair … x8 false | x1 ⇒ pair … x9 false | x2 ⇒ pair … xA false | x3 ⇒ pair … xB false
- | x4 ⇒ pair … xC false | x5 ⇒ pair … xD false | x6 ⇒ pair … xE false | x7 ⇒ pair … xF false
- | x8 ⇒ pair … x0 true | x9 ⇒ pair … x1 true | xA ⇒ pair … x2 true | xB ⇒ pair … x3 true
- | xC ⇒ pair … x4 true | xD ⇒ pair … x5 true | xE ⇒ pair … x6 true | xF ⇒ pair … x7 true ]
+ [ x0 ⇒ pair … false x8 | x1 ⇒ pair … false x9 | x2 ⇒ pair … false xA | x3 ⇒ pair … false xB
+ | x4 ⇒ pair … false xC | x5 ⇒ pair … false xD | x6 ⇒ pair … false xE | x7 ⇒ pair … false xF
+ | x8 ⇒ pair … true x0 | x9 ⇒ pair … true x1 | xA ⇒ pair … true x2 | xB ⇒ pair … true x3
+ | xC ⇒ pair … true x4 | xD ⇒ pair … true x5 | xE ⇒ pair … true x6 | xF ⇒ pair … true x7 ]
| x9 ⇒ match e2 with
- [ x0 ⇒ pair … x9 false | x1 ⇒ pair … xA false | x2 ⇒ pair … xB false | x3 ⇒ pair … xC false
- | x4 ⇒ pair … xD false | x5 ⇒ pair … xE false | x6 ⇒ pair … xF false | x7 ⇒ pair … x0 true
- | x8 ⇒ pair … x1 true | x9 ⇒ pair … x2 true | xA ⇒ pair … x3 true | xB ⇒ pair … x4 true
- | xC ⇒ pair … x5 true | xD ⇒ pair … x6 true | xE ⇒ pair … x7 true | xF ⇒ pair … x8 true ]
+ [ x0 ⇒ pair … false x9 | x1 ⇒ pair … false xA | x2 ⇒ pair … false xB | x3 ⇒ pair … false xC
+ | x4 ⇒ pair … false xD | x5 ⇒ pair … false xE | x6 ⇒ pair … false xF | x7 ⇒ pair … true x0
+ | x8 ⇒ pair … true x1 | x9 ⇒ pair … true x2 | xA ⇒ pair … true x3 | xB ⇒ pair … true x4
+ | xC ⇒ pair … true x5 | xD ⇒ pair … true x6 | xE ⇒ pair … true x7 | xF ⇒ pair … true x8 ]
| xA ⇒ match e2 with
- [ x0 ⇒ pair … xA false | x1 ⇒ pair … xB false | x2 ⇒ pair … xC false | x3 ⇒ pair … xD false
- | x4 ⇒ pair … xE false | x5 ⇒ pair … xF false | x6 ⇒ pair … x0 true | x7 ⇒ pair … x1 true
- | x8 ⇒ pair … x2 true | x9 ⇒ pair … x3 true | xA ⇒ pair … x4 true | xB ⇒ pair … x5 true
- | xC ⇒ pair … x6 true | xD ⇒ pair … x7 true | xE ⇒ pair … x8 true | xF ⇒ pair … x9 true ]
+ [ x0 ⇒ pair … false xA | x1 ⇒ pair … false xB | x2 ⇒ pair … false xC | x3 ⇒ pair … false xD
+ | x4 ⇒ pair … false xE | x5 ⇒ pair … false xF | x6 ⇒ pair … true x0 | x7 ⇒ pair … true x1
+ | x8 ⇒ pair … true x2 | x9 ⇒ pair … true x3 | xA ⇒ pair … true x4 | xB ⇒ pair … true x5
+ | xC ⇒ pair … true x6 | xD ⇒ pair … true x7 | xE ⇒ pair … true x8 | xF ⇒ pair … true x9 ]
| xB ⇒ match e2 with
- [ x0 ⇒ pair … xB false | x1 ⇒ pair … xC false | x2 ⇒ pair … xD false | x3 ⇒ pair … xE false
- | x4 ⇒ pair … xF false | x5 ⇒ pair … x0 true | x6 ⇒ pair … x1 true | x7 ⇒ pair … x2 true
- | x8 ⇒ pair … x3 true | x9 ⇒ pair … x4 true | xA ⇒ pair … x5 true | xB ⇒ pair … x6 true
- | xC ⇒ pair … x7 true | xD ⇒ pair … x8 true | xE ⇒ pair … x9 true | xF ⇒ pair … xA true ]
+ [ x0 ⇒ pair … false xB | x1 ⇒ pair … false xC | x2 ⇒ pair … false xD | x3 ⇒ pair … false xE
+ | x4 ⇒ pair … false xF | x5 ⇒ pair … true x0 | x6 ⇒ pair … true x1 | x7 ⇒ pair … true x2
+ | x8 ⇒ pair … true x3 | x9 ⇒ pair … true x4 | xA ⇒ pair … true x5 | xB ⇒ pair … true x6
+ | xC ⇒ pair … true x7 | xD ⇒ pair … true x8 | xE ⇒ pair … true x9 | xF ⇒ pair … true xA ]
| xC ⇒ match e2 with
- [ x0 ⇒ pair … xC false | x1 ⇒ pair … xD false | x2 ⇒ pair … xE false | x3 ⇒ pair … xF false
- | x4 ⇒ pair … x0 true | x5 ⇒ pair … x1 true | x6 ⇒ pair … x2 true | x7 ⇒ pair … x3 true
- | x8 ⇒ pair … x4 true | x9 ⇒ pair … x5 true | xA ⇒ pair … x6 true | xB ⇒ pair … x7 true
- | xC ⇒ pair … x8 true | xD ⇒ pair … x9 true | xE ⇒ pair … xA true | xF ⇒ pair … xB true ]
+ [ x0 ⇒ pair … false xC | x1 ⇒ pair … false xD | x2 ⇒ pair … false xE | x3 ⇒ pair … false xF
+ | x4 ⇒ pair … true x0 | x5 ⇒ pair … true x1 | x6 ⇒ pair … true x2 | x7 ⇒ pair … true x3
+ | x8 ⇒ pair … true x4 | x9 ⇒ pair … true x5 | xA ⇒ pair … true x6 | xB ⇒ pair … true x7
+ | xC ⇒ pair … true x8 | xD ⇒ pair … true x9 | xE ⇒ pair … true xA | xF ⇒ pair … true xB ]
| xD ⇒ match e2 with
- [ x0 ⇒ pair … xD false | x1 ⇒ pair … xE false | x2 ⇒ pair … xF false | x3 ⇒ pair … x0 true
- | x4 ⇒ pair … x1 true | x5 ⇒ pair … x2 true | x6 ⇒ pair … x3 true | x7 ⇒ pair … x4 true
- | x8 ⇒ pair … x5 true | x9 ⇒ pair … x6 true | xA ⇒ pair … x7 true | xB ⇒ pair … x8 true
- | xC ⇒ pair … x9 true | xD ⇒ pair … xA true | xE ⇒ pair … xB true | xF ⇒ pair … xC true ]
+ [ x0 ⇒ pair … false xD | x1 ⇒ pair … false xE | x2 ⇒ pair … false xF | x3 ⇒ pair … true x0
+ | x4 ⇒ pair … true x1 | x5 ⇒ pair … true x2 | x6 ⇒ pair … true x3 | x7 ⇒ pair … true x4
+ | x8 ⇒ pair … true x5 | x9 ⇒ pair … true x6 | xA ⇒ pair … true x7 | xB ⇒ pair … true x8
+ | xC ⇒ pair … true x9 | xD ⇒ pair … true xA | xE ⇒ pair … true xB | xF ⇒ pair … true xC ]
| xE ⇒ match e2 with
- [ x0 ⇒ pair … xE false | x1 ⇒ pair … xF false | x2 ⇒ pair … x0 true | x3 ⇒ pair … x1 true
- | x4 ⇒ pair … x2 true | x5 ⇒ pair … x3 true | x6 ⇒ pair … x4 true | x7 ⇒ pair … x5 true
- | x8 ⇒ pair … x6 true | x9 ⇒ pair … x7 true | xA ⇒ pair … x8 true | xB ⇒ pair … x9 true
- | xC ⇒ pair … xA true | xD ⇒ pair … xB true | xE ⇒ pair … xC true | xF ⇒ pair … xD true ]
+ [ x0 ⇒ pair … false xE | x1 ⇒ pair … false xF | x2 ⇒ pair … true x0 | x3 ⇒ pair … true x1
+ | x4 ⇒ pair … true x2 | x5 ⇒ pair … true x3 | x6 ⇒ pair … true x4 | x7 ⇒ pair … true x5
+ | x8 ⇒ pair … true x6 | x9 ⇒ pair … true x7 | xA ⇒ pair … true x8 | xB ⇒ pair … true x9
+ | xC ⇒ pair … true xA | xD ⇒ pair … true xB | xE ⇒ pair … true xC | xF ⇒ pair … true xD ]
| xF ⇒ match e2 with
- [ x0 ⇒ pair … xF false | x1 ⇒ pair … x0 true | x2 ⇒ pair … x1 true | x3 ⇒ pair … x2 true
- | x4 ⇒ pair … x3 true | x5 ⇒ pair … x4 true | x6 ⇒ pair … x5 true | x7 ⇒ pair … x6 true
- | x8 ⇒ pair … x7 true | x9 ⇒ pair … x8 true | xA ⇒ pair … x9 true | xB ⇒ pair … xA true
- | xC ⇒ pair … xB true | xD ⇒ pair … xC true | xE ⇒ pair … xD true | xF ⇒ pair … xE true ]
+ [ x0 ⇒ pair … false xF | x1 ⇒ pair … true x0 | x2 ⇒ pair … true x1 | x3 ⇒ pair … true x2
+ | x4 ⇒ pair … true x3 | x5 ⇒ pair … true x4 | x6 ⇒ pair … true x5 | x7 ⇒ pair … true x6
+ | x8 ⇒ pair … true x7 | x9 ⇒ pair … true x8 | xA ⇒ pair … true x9 | xB ⇒ pair … true xA
+ | xC ⇒ pair … true xB | xD ⇒ pair … true xC | xE ⇒ pair … true xD | xF ⇒ pair … true xE ]
].
(* operatore somma con data+carry → data *)
ndefinition plus_ex_dc_d ≝
-λe1,e2:exadecim.λc:bool.
+λc:bool.λe1,e2:exadecim.
match c with
[ true ⇒ match e1 with
[ x0 ⇒ match e2 with
(* operatore somma con data+carry → carry *)
ndefinition plus_ex_dc_c ≝
-λe1,e2:exadecim.λc:bool.
+λc:bool.λe1,e2:exadecim.
match c with
[ true ⇒ match e1 with
[ x0 ⇒ match e2 with
| x8 ⇒ true | x9 ⇒ true | xA ⇒ true | xB ⇒ true | xC ⇒ true | xD ⇒ true | xE ⇒ true | xF ⇒ true ]
].
-(* operatore Most Significant Bit *)
-ndefinition MSB_ex ≝
-λe:exadecim.match e with
- [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
- | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
- | x8 ⇒ true | x9 ⇒ true | xA ⇒ true | xB ⇒ true
- | xC ⇒ true | xD ⇒ true | xE ⇒ true | xF ⇒ true ].
-
(* operatore predecessore *)
ndefinition pred_ex ≝
λe:exadecim.
[ true ⇒ and_bool | false ⇒ or_bool ]
(le_ex inf x) (le_ex x sup).
-(* iteratore sugli esadecimali *)
-ndefinition forall_ex ≝ λP.
- P x0 ⊗ P x1 ⊗ P x2 ⊗ P x3 ⊗ P x4 ⊗ P x5 ⊗ P x6 ⊗ P x7 ⊗
- P x8 ⊗ P x9 ⊗ P xA ⊗ P xB ⊗ P xC ⊗ P xD ⊗ P xE ⊗ P xF.
-
(* esadecimali ricorsivi *)
ninductive rec_exadecim : exadecim → Type ≝
ex_O : rec_exadecim x0
nqed.
*)
-nlemma symmetric_andex : symmetricT exadecim exadecim and_ex.
- #e1; #e2;
- nelim e1;
- nelim e2;
- nnormalize;
- napply refl_eq.
-nqed.
-
-nlemma associative_andex1 : ∀e2,e3.(and_ex (and_ex x0 e2) e3) = (and_ex x0 (and_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_andex2 : ∀e2,e3.(and_ex (and_ex x1 e2) e3) = (and_ex x1 (and_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_andex3 : ∀e2,e3.(and_ex (and_ex x2 e2) e3) = (and_ex x2 (and_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_andex4 : ∀e2,e3.(and_ex (and_ex x3 e2) e3) = (and_ex x3 (and_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_andex5 : ∀e2,e3.(and_ex (and_ex x4 e2) e3) = (and_ex x4 (and_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_andex6 : ∀e2,e3.(and_ex (and_ex x5 e2) e3) = (and_ex x5 (and_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_andex7 : ∀e2,e3.(and_ex (and_ex x6 e2) e3) = (and_ex x6 (and_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_andex8 : ∀e2,e3.(and_ex (and_ex x7 e2) e3) = (and_ex x7 (and_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_andex9 : ∀e2,e3.(and_ex (and_ex x8 e2) e3) = (and_ex x8 (and_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_andex10 : ∀e2,e3.(and_ex (and_ex x9 e2) e3) = (and_ex x9 (and_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_andex11 : ∀e2,e3.(and_ex (and_ex xA e2) e3) = (and_ex xA (and_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_andex12 : ∀e2,e3.(and_ex (and_ex xB e2) e3) = (and_ex xB (and_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_andex13 : ∀e2,e3.(and_ex (and_ex xC e2) e3) = (and_ex xC (and_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_andex14 : ∀e2,e3.(and_ex (and_ex xD e2) e3) = (and_ex xD (and_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_andex15 : ∀e2,e3.(and_ex (and_ex xE e2) e3) = (and_ex xE (and_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_andex16 : ∀e2,e3.(and_ex (and_ex xF e2) e3) = (and_ex xF (and_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-
-nlemma associative_andex : ∀e1,e2,e3.(and_ex (and_ex e1 e2) e3) = (and_ex e1 (and_ex e2 e3)).
- #e1; nelim e1;
- ##[ ##1: napply associative_andex1 ##| ##2: napply associative_andex2
- ##| ##3: napply associative_andex3 ##| ##4: napply associative_andex4
- ##| ##5: napply associative_andex5 ##| ##6: napply associative_andex6
- ##| ##7: napply associative_andex7 ##| ##8: napply associative_andex8
- ##| ##9: napply associative_andex9 ##| ##10: napply associative_andex10
- ##| ##11: napply associative_andex11 ##| ##12: napply associative_andex12
- ##| ##13: napply associative_andex13 ##| ##14: napply associative_andex14
- ##| ##15: napply associative_andex15 ##| ##16: napply associative_andex16
- ##]
-nqed.
-
-nlemma symmetric_orex : symmetricT exadecim exadecim or_ex.
- #e1; #e2;
- nelim e1;
- nelim e2;
- nnormalize;
- napply refl_eq.
-nqed.
-
-nlemma associative_orex1 : ∀e2,e3.(or_ex (or_ex x0 e2) e3) = (or_ex x0 (or_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_orex2 : ∀e2,e3.(or_ex (or_ex x1 e2) e3) = (or_ex x1 (or_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_orex3 : ∀e2,e3.(or_ex (or_ex x2 e2) e3) = (or_ex x2 (or_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_orex4 : ∀e2,e3.(or_ex (or_ex x3 e2) e3) = (or_ex x3 (or_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_orex5 : ∀e2,e3.(or_ex (or_ex x4 e2) e3) = (or_ex x4 (or_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_orex6 : ∀e2,e3.(or_ex (or_ex x5 e2) e3) = (or_ex x5 (or_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_orex7 : ∀e2,e3.(or_ex (or_ex x6 e2) e3) = (or_ex x6 (or_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_orex8 : ∀e2,e3.(or_ex (or_ex x7 e2) e3) = (or_ex x7 (or_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_orex9 : ∀e2,e3.(or_ex (or_ex x8 e2) e3) = (or_ex x8 (or_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_orex10 : ∀e2,e3.(or_ex (or_ex x9 e2) e3) = (or_ex x9 (or_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_orex11 : ∀e2,e3.(or_ex (or_ex xA e2) e3) = (or_ex xA (or_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_orex12 : ∀e2,e3.(or_ex (or_ex xB e2) e3) = (or_ex xB (or_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_orex13 : ∀e2,e3.(or_ex (or_ex xC e2) e3) = (or_ex xC (or_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_orex14 : ∀e2,e3.(or_ex (or_ex xD e2) e3) = (or_ex xD (or_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_orex15 : ∀e2,e3.(or_ex (or_ex xE e2) e3) = (or_ex xE (or_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_orex16 : ∀e2,e3.(or_ex (or_ex xF e2) e3) = (or_ex xF (or_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-
-nlemma associative_orex : ∀e1,e2,e3.(or_ex (or_ex e1 e2) e3) = (or_ex e1 (or_ex e2 e3)).
- #e1; nelim e1;
- ##[ ##1: napply associative_orex1 ##| ##2: napply associative_orex2
- ##| ##3: napply associative_orex3 ##| ##4: napply associative_orex4
- ##| ##5: napply associative_orex5 ##| ##6: napply associative_orex6
- ##| ##7: napply associative_orex7 ##| ##8: napply associative_orex8
- ##| ##9: napply associative_orex9 ##| ##10: napply associative_orex10
- ##| ##11: napply associative_orex11 ##| ##12: napply associative_orex12
- ##| ##13: napply associative_orex13 ##| ##14: napply associative_orex14
- ##| ##15: napply associative_orex15 ##| ##16: napply associative_orex16
- ##]
-nqed.
-
-nlemma symmetric_xorex : symmetricT exadecim exadecim xor_ex.
- #e1; #e2;
- nelim e1;
- nelim e2;
- nnormalize;
- napply refl_eq.
-nqed.
-
-nlemma associative_xorex1 : ∀e2,e3.(xor_ex (xor_ex x0 e2) e3) = (xor_ex x0 (xor_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_xorex2 : ∀e2,e3.(xor_ex (xor_ex x1 e2) e3) = (xor_ex x1 (xor_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_xorex3 : ∀e2,e3.(xor_ex (xor_ex x2 e2) e3) = (xor_ex x2 (xor_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_xorex4 : ∀e2,e3.(xor_ex (xor_ex x3 e2) e3) = (xor_ex x3 (xor_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_xorex5 : ∀e2,e3.(xor_ex (xor_ex x4 e2) e3) = (xor_ex x4 (xor_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_xorex6 : ∀e2,e3.(xor_ex (xor_ex x5 e2) e3) = (xor_ex x5 (xor_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_xorex7 : ∀e2,e3.(xor_ex (xor_ex x6 e2) e3) = (xor_ex x6 (xor_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_xorex8 : ∀e2,e3.(xor_ex (xor_ex x7 e2) e3) = (xor_ex x7 (xor_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_xorex9 : ∀e2,e3.(xor_ex (xor_ex x8 e2) e3) = (xor_ex x8 (xor_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_xorex10 : ∀e2,e3.(xor_ex (xor_ex x9 e2) e3) = (xor_ex x9 (xor_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_xorex11 : ∀e2,e3.(xor_ex (xor_ex xA e2) e3) = (xor_ex xA (xor_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_xorex12 : ∀e2,e3.(xor_ex (xor_ex xB e2) e3) = (xor_ex xB (xor_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_xorex13 : ∀e2,e3.(xor_ex (xor_ex xC e2) e3) = (xor_ex xC (xor_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_xorex14 : ∀e2,e3.(xor_ex (xor_ex xD e2) e3) = (xor_ex xD (xor_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_xorex15 : ∀e2,e3.(xor_ex (xor_ex xE e2) e3) = (xor_ex xE (xor_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_xorex16 : ∀e2,e3.(xor_ex (xor_ex xF e2) e3) = (xor_ex xF (xor_ex e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-
-nlemma associative_xorex : ∀e1,e2,e3.(xor_ex (xor_ex e1 e2) e3) = (xor_ex e1 (xor_ex e2 e3)).
- #e1; nelim e1;
- ##[ ##1: napply associative_xorex1 ##| ##2: napply associative_xorex2
- ##| ##3: napply associative_xorex3 ##| ##4: napply associative_xorex4
- ##| ##5: napply associative_xorex5 ##| ##6: napply associative_xorex6
- ##| ##7: napply associative_xorex7 ##| ##8: napply associative_xorex8
- ##| ##9: napply associative_xorex9 ##| ##10: napply associative_xorex10
- ##| ##11: napply associative_xorex11 ##| ##12: napply associative_xorex12
- ##| ##13: napply associative_xorex13 ##| ##14: napply associative_xorex14
- ##| ##15: napply associative_xorex15 ##| ##16: napply associative_xorex16
- ##]
-nqed.
-
-nlemma symmetric_plusex_dc_dc : ∀e1,e2,c.plus_ex_dc_dc e1 e2 c = plus_ex_dc_dc e2 e1 c.
- #e1; #e2; #c;
- nelim e1;
- nelim e2;
- nelim c;
- nnormalize;
- napply refl_eq.
-nqed.
-
-nlemma plusex_dc_dc_to_dc_d : ∀e1,e2,c.fst … (plus_ex_dc_dc e1 e2 c) = plus_ex_dc_d e1 e2 c.
- #e1; #e2; #c;
- nelim e1;
- nelim e2;
- nelim c;
- nnormalize;
- napply refl_eq.
-nqed.
-
-nlemma plusex_dc_dc_to_dc_c : ∀e1,e2,c.snd … (plus_ex_dc_dc e1 e2 c) = plus_ex_dc_c e1 e2 c.
- #e1; #e2; #c;
- nelim e1;
- nelim e2;
- nelim c;
- nnormalize;
- napply refl_eq.
-nqed.
-
-nlemma plusex_dc_dc_to_d_dc : ∀e1,e2.plus_ex_dc_dc e1 e2 false = plus_ex_d_dc e1 e2.
- #e1; #e2;
- nelim e1;
- nelim e2;
- nnormalize;
- napply refl_eq.
-nqed.
-
-nlemma plusex_dc_dc_to_d_d : ∀e1,e2.fst … (plus_ex_dc_dc e1 e2 false) = plus_ex_d_d e1 e2.
- #e1; #e2;
- nelim e1;
- nelim e2;
- nnormalize;
- napply refl_eq.
-nqed.
-
-nlemma plusex_dc_dc_to_d_c : ∀e1,e2.snd … (plus_ex_dc_dc e1 e2 false) = plus_ex_d_c e1 e2.
- #e1; #e2;
- nelim e1;
- nelim e2;
- nnormalize;
- napply refl_eq.
-nqed.
-
-nlemma symmetric_plusex_dc_d : ∀e1,e2,c.plus_ex_dc_d e1 e2 c = plus_ex_dc_d e2 e1 c.
- #e1; #e2; #c;
- nelim e1;
- nelim e2;
- nelim c;
- nnormalize;
- napply refl_eq.
-nqed.
-
-nlemma symmetric_plusex_dc_c : ∀e1,e2,c.plus_ex_dc_c e1 e2 c = plus_ex_dc_c e2 e1 c.
- #e1; #e2; #c;
- nelim e1;
- nelim e2;
- nelim c;
- nnormalize;
- napply refl_eq.
-nqed.
-
-nlemma symmetric_plusex_d_dc : ∀e1,e2.plus_ex_d_dc e1 e2 = plus_ex_d_dc e2 e1.
- #e1; #e2;
- nelim e1;
- nelim e2;
- nnormalize;
- napply refl_eq.
-nqed.
-
-nlemma plusex_d_dc_to_d_d : ∀e1,e2.fst … (plus_ex_d_dc e1 e2) = plus_ex_d_d e1 e2.
- #e1; #e2;
- nelim e1;
- nelim e2;
- nnormalize;
- napply refl_eq.
-nqed.
-
-nlemma plusex_d_dc_to_d_c : ∀e1,e2.snd … (plus_ex_d_dc e1 e2) = plus_ex_d_c e1 e2.
- #e1; #e2;
- nelim e1;
- nelim e2;
- nnormalize;
- napply refl_eq.
-nqed.
-
-nlemma symmetric_plusex_d_d : ∀e1,e2.plus_ex_d_d e1 e2 = plus_ex_d_d e2 e1.
- #e1; #e2;
- nelim e1;
- nelim e2;
- nnormalize;
- napply refl_eq.
-nqed.
-
-nlemma associative_plusex_d_d1 : ∀e2,e3.(plus_ex_d_d (plus_ex_d_d x0 e2) e3) = (plus_ex_d_d x0 (plus_ex_d_d e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_plusex_d_d2 : ∀e2,e3.(plus_ex_d_d (plus_ex_d_d x1 e2) e3) = (plus_ex_d_d x1 (plus_ex_d_d e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_plusex_d_d3 : ∀e2,e3.(plus_ex_d_d (plus_ex_d_d x2 e2) e3) = (plus_ex_d_d x2 (plus_ex_d_d e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_plusex_d_d4 : ∀e2,e3.(plus_ex_d_d (plus_ex_d_d x3 e2) e3) = (plus_ex_d_d x3 (plus_ex_d_d e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_plusex_d_d5 : ∀e2,e3.(plus_ex_d_d (plus_ex_d_d x4 e2) e3) = (plus_ex_d_d x4 (plus_ex_d_d e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_plusex_d_d6 : ∀e2,e3.(plus_ex_d_d (plus_ex_d_d x5 e2) e3) = (plus_ex_d_d x5 (plus_ex_d_d e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_plusex_d_d7 : ∀e2,e3.(plus_ex_d_d (plus_ex_d_d x6 e2) e3) = (plus_ex_d_d x6 (plus_ex_d_d e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_plusex_d_d8 : ∀e2,e3.(plus_ex_d_d (plus_ex_d_d x7 e2) e3) = (plus_ex_d_d x7 (plus_ex_d_d e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_plusex_d_d9 : ∀e2,e3.(plus_ex_d_d (plus_ex_d_d x8 e2) e3) = (plus_ex_d_d x8 (plus_ex_d_d e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_plusex_d_d10 : ∀e2,e3.(plus_ex_d_d (plus_ex_d_d x9 e2) e3) = (plus_ex_d_d x9 (plus_ex_d_d e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_plusex_d_d11 : ∀e2,e3.(plus_ex_d_d (plus_ex_d_d xA e2) e3) = (plus_ex_d_d xA (plus_ex_d_d e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_plusex_d_d12 : ∀e2,e3.(plus_ex_d_d (plus_ex_d_d xB e2) e3) = (plus_ex_d_d xB (plus_ex_d_d e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_plusex_d_d13 : ∀e2,e3.(plus_ex_d_d (plus_ex_d_d xC e2) e3) = (plus_ex_d_d xC (plus_ex_d_d e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_plusex_d_d14 : ∀e2,e3.(plus_ex_d_d (plus_ex_d_d xD e2) e3) = (plus_ex_d_d xD (plus_ex_d_d e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_plusex_d_d15 : ∀e2,e3.(plus_ex_d_d (plus_ex_d_d xE e2) e3) = (plus_ex_d_d xE (plus_ex_d_d e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-nlemma associative_plusex_d_d16 : ∀e2,e3.(plus_ex_d_d (plus_ex_d_d xF e2) e3) = (plus_ex_d_d xF (plus_ex_d_d e2 e3)). #e2; #e3; nelim e2; nelim e3; nnormalize; napply refl_eq. nqed.
-
-nlemma associative_plusex_d_d : ∀e1,e2,e3.(plus_ex_d_d (plus_ex_d_d e1 e2) e3) = (plus_ex_d_d e1 (plus_ex_d_d e2 e3)).
- #e1; nelim e1;
- ##[ ##1: napply associative_plusex_d_d1 ##| ##2: napply associative_plusex_d_d2
- ##| ##3: napply associative_plusex_d_d3 ##| ##4: napply associative_plusex_d_d4
- ##| ##5: napply associative_plusex_d_d5 ##| ##6: napply associative_plusex_d_d6
- ##| ##7: napply associative_plusex_d_d7 ##| ##8: napply associative_plusex_d_d8
- ##| ##9: napply associative_plusex_d_d9 ##| ##10: napply associative_plusex_d_d10
- ##| ##11: napply associative_plusex_d_d11 ##| ##12: napply associative_plusex_d_d12
- ##| ##13: napply associative_plusex_d_d13 ##| ##14: napply associative_plusex_d_d14
- ##| ##15: napply associative_plusex_d_d15 ##| ##16: napply associative_plusex_d_d16
- ##]
-nqed.
-
-nlemma symmetric_plusex_d_c : ∀e1,e2.plus_ex_d_c e1 e2 = plus_ex_d_c e2 e1.
- #e1; #e2;
- nelim e1;
- nelim e2;
- nnormalize;
- napply refl_eq.
-nqed.
-
nlemma eq_to_eqex : ∀n1,n2.n1 = n2 → eq_ex n1 n2 = true.
#n1; #n2; #H;
nrewrite > H;
(* WORD *)
(* **** *)
-nrecord word16 : Type ≝
- {
- w16h: byte8;
- w16l: byte8
- }.
+ndefinition word16 ≝ comp_num byte8.
+ndefinition mk_word16 ≝ λb1,b2.mk_comp_num byte8 b1 b2.
(* \langle \rangle *)
notation "〈x:y〉" non associative with precedence 80
- for @{ 'mk_word16 $x $y }.
-interpretation "mk_word16" 'mk_word16 x y = (mk_word16 x y).
+ for @{ mk_comp_num byte8 $x $y }.
+
+(* iteratore sulle word *)
+ndefinition forall_w16 ≝ forall_cn ? forall_b8.
(* operatore = *)
-ndefinition eq_w16 ≝ λw1,w2.(eq_b8 (w16h w1) (w16h w2)) ⊗ (eq_b8 (w16l w1) (w16l w2)).
+ndefinition eq_w16 ≝ eq2_cn ? eq_b8.
(* operatore < *)
-ndefinition lt_w16 ≝
-λw1,w2:word16.
- (lt_b8 (w16h w1) (w16h w2)) ⊕
- ((eq_b8 (w16h w1) (w16h w2)) ⊗ (lt_b8 (w16l w1) (w16l w2))).
+ndefinition lt_w16 ≝ ltgt_cn ? eq_b8 lt_b8.
(* operatore ≤ *)
-ndefinition le_w16 ≝
-λw1,w2:word16.
- (lt_b8 (w16h w1) (w16h w2)) ⊕
- ((eq_b8 (w16h w1) (w16h w2)) ⊗ (le_b8 (w16l w1) (w16l w2))).
+ndefinition le_w16 ≝ lege_cn ? eq_b8 lt_b8 le_b8.
(* operatore > *)
-ndefinition gt_w16 ≝
-λw1,w2:word16.
- (gt_b8 (w16h w1) (w16h w2)) ⊕
- ((eq_b8 (w16h w1) (w16h w2)) ⊗ (gt_b8 (w16l w1) (w16l w2))).
+ndefinition gt_w16 ≝ ltgt_cn ? eq_b8 gt_b8.
(* operatore ≥ *)
-ndefinition ge_w16 ≝
-λw1,w2:word16.
- (gt_b8 (w16h w1) (w16h w2)) ⊕
- ((eq_b8 (w16h w1) (w16h w2)) ⊗ (ge_b8 (w16l w1) (w16l w2))).
+ndefinition ge_w16 ≝ lege_cn ? eq_b8 gt_b8 ge_b8.
(* operatore and *)
-ndefinition and_w16 ≝
-λw1,w2:word16.mk_word16 (and_b8 (w16h w1) (w16h w2)) (and_b8 (w16l w1) (w16l w2)).
+ndefinition and_w16 ≝ fop2_cn ? and_b8.
(* operatore or *)
-ndefinition or_w16 ≝
-λw1,w2:word16.mk_word16 (or_b8 (w16h w1) (w16h w2)) (or_b8 (w16l w1) (w16l w2)).
+ndefinition or_w16 ≝ fop2_cn ? or_b8.
(* operatore xor *)
-ndefinition xor_w16 ≝
-λw1,w2:word16.mk_word16 (xor_b8 (w16h w1) (w16h w2)) (xor_b8 (w16l w1) (w16l w2)).
+ndefinition xor_w16 ≝ fop2_cn ? xor_b8.
+
+(* operatore Most Significant Bit *)
+ndefinition getMSB_w16 ≝ getOPH_cn ? getMSB_b8.
+ndefinition setMSB_w16 ≝ setOPH_cn ? setMSB_b8.
+
+(* operatore Least Significant Bit *)
+ndefinition getLSB_w16 ≝ getOPL_cn ? getLSB_b8.
+ndefinition setLSB_w16 ≝ setOPL_cn ? setLSB_b8.
(* operatore rotazione destra con carry *)
-ndefinition rcr_w16 ≝
-λw:word16.λc:bool.match rcr_b8 (w16h w) c with
- [ pair wh' c' ⇒ match rcr_b8 (w16l w) c' with
- [ pair wl' c'' ⇒ pair … (mk_word16 wh' wl') c'' ]].
+ndefinition rcr_w16 ≝ opcr_cn ? rcr_b8.
(* operatore shift destro *)
-ndefinition shr_w16 ≝
-λw:word16.match shr_b8 (w16h w) with
- [ pair wh' c' ⇒ match rcr_b8 (w16l w) c' with
- [ pair wl' c'' ⇒ pair … (mk_word16 wh' wl') c'' ]].
+ndefinition shr_w16 ≝ opcr_cn ? rcr_b8 false.
(* operatore rotazione destra *)
ndefinition ror_w16 ≝
-λw:word16.match shr_b8 (w16h w) with
- [ pair wh' c' ⇒ match rcr_b8 (w16l w) c' with
- [ pair wl' c'' ⇒ match c'' with
- [ true ⇒ mk_word16 (or_b8 (mk_byte8 x8 x0) wh') wl'
- | false ⇒ mk_word16 wh' wl' ]]].
-
-(* operatore rotazione destra n-volte *)
-nlet rec ror_w16_n (w:word16) (e:exadecim) (re:rec_exadecim e) on re ≝
- match re with
- [ ex_O ⇒ w
- | ex_S e' n' ⇒ ror_w16_n (ror_w16 w) e' n' ].
+λw.match shr_w16 w with
+ [ pair c w' ⇒ match c with
+ [ true ⇒ setMSB_w16 w' | false ⇒ w' ]].
(* operatore rotazione sinistra con carry *)
-ndefinition rcl_w16 ≝
-λw:word16.λc:bool.match rcl_b8 (w16l w) c with
- [ pair wl' c' ⇒ match rcl_b8 (w16h w) c' with
- [ pair wh' c'' ⇒ pair … (mk_word16 wh' wl') c'' ]].
+ndefinition rcl_w16 ≝ opcl_cn ? rcl_b8.
(* operatore shift sinistro *)
-ndefinition shl_w16 ≝
-λw:word16.match shl_b8 (w16l w) with
- [ pair wl' c' ⇒ match rcl_b8 (w16h w) c' with
- [ pair wh' c'' ⇒ pair … (mk_word16 wh' wl') c'' ]].
+ndefinition shl_w16 ≝ opcl_cn ? rcl_b8 false.
(* operatore rotazione sinistra *)
ndefinition rol_w16 ≝
-λw:word16.match shl_b8 (w16l w) with
- [ pair wl' c' ⇒ match rcl_b8 (w16h w) c' with
- [ pair wh' c'' ⇒ match c'' with
- [ true ⇒ mk_word16 wh' (or_b8 (mk_byte8 x0 x1) wl')
- | false ⇒ mk_word16 wh' wl' ]]].
-
-(* operatore rotazione sinistra n-volte *)
-nlet rec rol_w16_n (w:word16) (e:exadecim) (re:rec_exadecim e) on re ≝
- match re with
- [ ex_O ⇒ w
- | ex_S e' n' ⇒ rol_w16_n (rol_w16 w) e' n' ].
+λw.match shl_w16 w with
+ [ pair c w' ⇒ match c with
+ [ true ⇒ setLSB_w16 w' | false ⇒ w' ]].
(* operatore not/complemento a 1 *)
-ndefinition not_w16 ≝
-λw:word16.mk_word16 (not_b8 (w16h w)) (not_b8 (w16l w)).
+ndefinition not_w16 ≝ fop_cn ? not_b8.
(* operatore somma con data+carry → data+carry *)
-ndefinition plus_w16_dc_dc ≝
-λw1,w2:word16.λc:bool.
- match plus_b8_dc_dc (w16l w1) (w16l w2) c with
- [ pair l c' ⇒ match plus_b8_dc_dc (w16h w1) (w16h w2) c' with
- [ pair h c'' ⇒ pair … 〈h:l〉 c'' ]].
+ndefinition plus_w16_dc_dc ≝ opcl2_cn ? plus_b8_dc_dc.
(* operatore somma con data+carry → data *)
-ndefinition plus_w16_dc_d ≝
-λw1,w2:word16.λc:bool.
- match plus_b8_dc_dc (w16l w1) (w16l w2) c with
- [ pair l c' ⇒ 〈plus_b8_dc_d (w16h w1) (w16h w2) c':l〉 ].
+ndefinition plus_w16_dc_d ≝ λc,w1,w2.snd … (plus_w16_dc_dc c w1 w2).
(* operatore somma con data+carry → c *)
-ndefinition plus_w16_dc_c ≝
-λw1,w2:word16.λc:bool.
- plus_b8_dc_c (w16h w1) (w16h w2) (plus_b8_dc_c (w16l w1) (w16l w2) c).
+ndefinition plus_w16_dc_c ≝ λc,w1,w2.fst … (plus_w16_dc_dc c w1 w2).
(* operatore somma con data → data+carry *)
-ndefinition plus_w16_d_dc ≝
-λw1,w2:word16.
- match plus_b8_d_dc (w16l w1) (w16l w2) with
- [ pair l c ⇒ match plus_b8_dc_dc (w16h w1) (w16h w2) c with
- [ pair h c' ⇒ pair … 〈h:l〉 c' ]].
+ndefinition plus_w16_d_dc ≝ opcl2_cn ? plus_b8_dc_dc false.
(* operatore somma con data → data *)
-ndefinition plus_w16_d_d ≝
-λw1,w2:word16.
- match plus_b8_d_dc (w16l w1) (w16l w2) with
- [ pair l c ⇒ 〈plus_b8_dc_d (w16h w1) (w16h w2) c:l〉 ].
+ndefinition plus_w16_d_d ≝ λw1,w2.snd … (plus_w16_d_dc w1 w2).
(* operatore somma con data → c *)
-ndefinition plus_w16_d_c ≝
-λw1,w2:word16.
- plus_b8_dc_c (w16h w1) (w16h w2) (plus_b8_d_c (w16l w1) (w16l w2)).
-
-(* operatore Most Significant Bit *)
-ndefinition MSB_w16 ≝ λw:word16.eq_ex x8 (and_ex x8 (b8h (w16h w))).
+ndefinition plus_w16_d_c ≝ λw1,w2.fst … (plus_w16_d_dc w1 w2).
(* operatore predecessore *)
-ndefinition pred_w16 ≝
-λw:word16.match eq_b8 (w16l w) (mk_byte8 x0 x0) with
- [ true ⇒ mk_word16 (pred_b8 (w16h w)) (pred_b8 (w16l w))
- | false ⇒ mk_word16 (w16h w) (pred_b8 (w16l w)) ].
+ndefinition pred_w16 ≝ predsucc_cn ? (eq_b8 〈x0,x0〉) pred_b8.
(* operatore successore *)
-ndefinition succ_w16 ≝
-λw:word16.match eq_b8 (w16l w) (mk_byte8 xF xF) with
- [ true ⇒ mk_word16 (succ_b8 (w16h w)) (succ_b8 (w16l w))
- | false ⇒ mk_word16 (w16h w) (succ_b8 (w16l w)) ].
+ndefinition succ_w16 ≝ predsucc_cn ? (eq_b8 〈xF,xF〉) succ_b8.
(* operatore neg/complemento a 2 *)
ndefinition compl_w16 ≝
-λw:word16.match MSB_w16 w with
+λw:word16.match getMSB_w16 w with
[ true ⇒ succ_w16 (not_w16 w)
| false ⇒ not_w16 (pred_w16 w) ].
-(*
- operatore moltiplicazione senza segno: b*b=[0x0000,0xFE01]
- ... in pratica (〈a,b〉*〈c,d〉) = (a*c)<<8+(a*d)<<4+(b*c)<<4+(b*d)
-*)
-ndefinition mul_b8 ≝
-λb1,b2:byte8.match b1 with
-[ mk_byte8 b1h b1l ⇒ match b2 with
-[ mk_byte8 b2h b2l ⇒ match mul_ex b1l b2l with
-[ mk_byte8 t1_h t1_l ⇒ match mul_ex b1h b2l with
-[ mk_byte8 t2_h t2_l ⇒ match mul_ex b2h b1l with
-[ mk_byte8 t3_h t3_l ⇒ match mul_ex b1h b2h with
-[ 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〉〉
-]]]]]].
-
-(* divisione senza segno (secondo la logica delle ALU): (quoziente resto) overflow *)
-nlet rec div_w16_b8_aux (divd:word16) (divs:word16) (molt:byte8) (q:byte8) (o:oct) (ro:rec_oct o) on ro ≝
- let w' ≝ plus_w16_d_d divd (compl_w16 divs) in
- match ro with
- [ oc_O ⇒ match le_w16 divs divd with
- [ true ⇒ triple … (or_b8 molt q) (w16l w') (⊖ (eq_b8 (w16h w') 〈x0,x0〉))
- | false ⇒ triple … q (w16l divd) (⊖ (eq_b8 (w16h divd) 〈x0,x0〉)) ]
- | oc_S o' ro' ⇒ match le_w16 divs divd with
- [ true ⇒ div_w16_b8_aux w' (ror_w16 divs) (ror_b8 molt) (or_b8 molt q) o' ro'
- | false ⇒ div_w16_b8_aux divd (ror_w16 divs) (ror_b8 molt) q o' ro' ]].
-
-ndefinition div_w16_b8 ≝
-λw:word16.λb:byte8.match eq_b8 b 〈x0,x0〉 with
-(*
- la combinazione n/0 e' illegale, segnala solo overflow senza dare risultato
-*)
- [ true ⇒ triple … 〈xF,xF〉 (w16l w) true
- | false ⇒ match eq_w16 w 〈〈x0,x0〉:〈x0,x0〉〉 with
-(* 0 diviso qualsiasi cosa diverso da 0 da' q=0 r=0 o=false *)
- [ true ⇒ triple … 〈x0,x0〉 〈x0,x0〉 false
-(* 1) e' una divisione sensata che produrra' overflow/risultato *)
-(* 2) parametri: dividendo, divisore, moltiplicatore, quoziente, contatore *)
-(* 3) ad ogni ciclo il divisore e il moltiplicatore vengono scalati di 1 a dx *)
-(* 4) il moltiplicatore e' la quantita' aggiunta al quoziente se il divisore *)
-(* puo' essere sottratto al dividendo *)
- | false ⇒ div_w16_b8_aux w (rol_w16_n 〈〈x0,x0〉:b〉 ? (ex_to_recex x7)) 〈x8,x0〉 〈x0,x0〉 ? (oct_to_recoct o7) ]].
-
(* operatore x in [inf,sup] o in sup],[inf *)
ndefinition inrange_w16 ≝
λx,inf,sup:word16.
[ true ⇒ and_bool | false ⇒ or_bool ]
(le_w16 inf x) (le_w16 x sup).
-(* iteratore sulle word *)
-ndefinition forall_w16 ≝
- λP.
- forall_b8 (λbh.
- forall_b8 (λbl.
- P (mk_word16 bh bl ))).
-
(* word16 ricorsive *)
ninductive rec_word16 : word16 → Type ≝
w16_O : rec_word16 〈〈x0,x0〉:〈x0,x0〉〉
ndefinition w16_to_recw16 : Πw.rec_word16 w ≝
λw.
- match w with [ mk_word16 h l ⇒
- match l with [ mk_byte8 lh ll ⇒
+ match w with [ mk_comp_num h l ⇒
+ match l with [ mk_comp_num lh ll ⇒
w16_to_recw16_aux4 h lh ll (w16_to_recw16_aux3 h lh (w16_to_recw16_aux2 h (b8_to_recb8 h))) ]].
(* WORD *)
(* **** *)
-nlemma word16_destruct_1 :
-∀x1,x2,y1,y2.
- mk_word16 x1 y1 = mk_word16 x2 y2 → x1 = x2.
- #x1; #x2; #y1; #y2; #H;
- nchange with (match mk_word16 x2 y2 with [ mk_word16 a _ ⇒ x1 = a ]);
- nrewrite < H;
- nnormalize;
- napply refl_eq.
-nqed.
+ndefinition word16_destruct_1 ≝ cn_destruct_1 word16.
+ndefinition word16_destruct_2 ≝ cn_destruct_2 word16.
-nlemma word16_destruct_2 :
-∀x1,x2,y1,y2.
- mk_word16 x1 y1 = mk_word16 x2 y2 → y1 = y2.
- #x1; #x2; #y1; #y2; #H;
- nchange with (match mk_word16 x2 y2 with [ mk_word16 _ b ⇒ y1 = b ]);
- nrewrite < H;
- nnormalize;
- napply refl_eq.
-nqed.
+ndefinition symmetric_eqw16 ≝ symmetric_eqcn ? eq_b8 symmetric_eqb8.
-nlemma symmetric_eqw16 : symmetricT word16 bool eq_w16.
- #w1; #w2;
- nelim w1;
- nelim w2;
- #b1; #b2; #b3; #b4;
- nchange with (((eq_b8 b3 b1)⊗(eq_b8 b4 b2)) = ((eq_b8 b1 b3)⊗(eq_b8 b2 b4)));
- nrewrite > (symmetric_eqb8 b1 b3);
- nrewrite > (symmetric_eqb8 b2 b4);
- napply refl_eq.
-nqed.
+ndefinition eqw16_to_eq ≝ eqcn_to_eq ? eq_b8 eqb8_to_eq.
+ndefinition eq_to_eqw16 ≝ eq_to_eqcn ? eq_b8 eq_to_eqb8.
-nlemma symmetric_andw16 : symmetricT word16 word16 and_w16.
- #w1; #w2;
- nelim w1;
- nelim w2;
- #b1; #b2; #b3; #b4;
- nchange with ((mk_word16 (and_b8 b3 b1) (and_b8 b4 b2)) = (mk_word16 (and_b8 b1 b3) (and_b8 b2 b4)));
- nrewrite > (symmetric_andb8 b1 b3);
- nrewrite > (symmetric_andb8 b2 b4);
- napply refl_eq.
-nqed.
+ndefinition decidable_w16 ≝ decidable_cn ? decidable_b8.
-nlemma associative_andw16 : ∀w1,w2,w3.(and_w16 (and_w16 w1 w2) w3) = (and_w16 w1 (and_w16 w2 w3)).
- #w1; #w2; #w3;
- nelim w1;
- #b1; #b2;
- nelim w2;
- #b3; #b4;
- nelim w3;
- #b5; #b6;
- nchange with (mk_word16 (and_b8 (and_b8 b1 b3) b5) (and_b8 (and_b8 b2 b4) b6) =
- mk_word16 (and_b8 b1 (and_b8 b3 b5)) (and_b8 b2 (and_b8 b4 b6)));
- nrewrite < (associative_andb8 b1 b3 b5);
- nrewrite < (associative_andb8 b2 b4 b6);
- napply refl_eq.
-nqed.
-
-nlemma symmetric_orw16 : symmetricT word16 word16 or_w16.
- #w1; #w2;
- nelim w1;
- nelim w2;
- #b1; #b2; #b3; #b4;
- nchange with ((mk_word16 (or_b8 b3 b1) (or_b8 b4 b2)) = (mk_word16 (or_b8 b1 b3) (or_b8 b2 b4)));
- nrewrite > (symmetric_orb8 b1 b3);
- nrewrite > (symmetric_orb8 b2 b4);
- napply refl_eq.
-nqed.
-
-nlemma associative_orw16 : ∀w1,w2,w3.(or_w16 (or_w16 w1 w2) w3) = (or_w16 w1 (or_w16 w2 w3)).
- #w1; #w2; #w3;
- nelim w1;
- #b1; #b2;
- nelim w2;
- #b3; #b4;
- nelim w3;
- #b5; #b6;
- nchange with (mk_word16 (or_b8 (or_b8 b1 b3) b5) (or_b8 (or_b8 b2 b4) b6) =
- mk_word16 (or_b8 b1 (or_b8 b3 b5)) (or_b8 b2 (or_b8 b4 b6)));
- nrewrite < (associative_orb8 b1 b3 b5);
- nrewrite < (associative_orb8 b2 b4 b6);
- napply refl_eq.
-nqed.
-
-nlemma symmetric_xorw16 : symmetricT word16 word16 xor_w16.
- #w1; #w2;
- nelim w1;
- nelim w2;
- #b1; #b2; #b3; #b4;
- nchange with ((mk_word16 (xor_b8 b3 b1) (xor_b8 b4 b2)) = (mk_word16 (xor_b8 b1 b3) (xor_b8 b2 b4)));
- nrewrite > (symmetric_xorb8 b1 b3);
- nrewrite > (symmetric_xorb8 b2 b4);
- napply refl_eq.
-nqed.
-
-nlemma associative_xorw16 : ∀w1,w2,w3.(xor_w16 (xor_w16 w1 w2) w3) = (xor_w16 w1 (xor_w16 w2 w3)).
- #w1; #w2; #w3;
- nelim w1;
- #b1; #b2;
- nelim w2;
- #b3; #b4;
- nelim w3;
- #b5; #b6;
- nchange with (mk_word16 (xor_b8 (xor_b8 b1 b3) b5) (xor_b8 (xor_b8 b2 b4) b6) =
- mk_word16 (xor_b8 b1 (xor_b8 b3 b5)) (xor_b8 b2 (xor_b8 b4 b6)));
- nrewrite < (associative_xorb8 b1 b3 b5);
- nrewrite < (associative_xorb8 b2 b4 b6);
- napply refl_eq.
-nqed.
-
-nlemma symmetric_plusw16_dc_dc : ∀w1,w2,c.plus_w16_dc_dc w1 w2 c = plus_w16_dc_dc w2 w1 c.
- #w1; #w2; #c;
- nelim w1;
- #b1; #b2;
- nelim w2;
- #b3; #b4;
- nchange with (
- 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' ]] =
- 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' ]]);
- nrewrite > (symmetric_plusb8_dc_dc b4 b2 c);
- ncases (plus_b8_dc_dc b2 b4 c);
- #b5; #c1;
- nchange with (
- match plus_b8_dc_dc b1 b3 c1 with [ pair h c' ⇒ pair … 〈h:b5〉 c' ] =
- match plus_b8_dc_dc b3 b1 c1 with [ pair h c' ⇒ pair … 〈h:b5〉 c' ]);
- nrewrite > (symmetric_plusb8_dc_dc b1 b3 c1);
- napply refl_eq.
-nqed.
-
-nlemma symmetric_plusw16_dc_d : ∀w1,w2,c.plus_w16_dc_d w1 w2 c = plus_w16_dc_d w2 w1 c.
- #w1; #w2; #c;
- nelim w1;
- #b1; #b2;
- nelim w2;
- #b3; #b4;
- nchange with (
- match plus_b8_dc_dc b2 b4 c with [ pair l c ⇒ 〈plus_b8_dc_d b1 b3 c:l〉 ] =
- match plus_b8_dc_dc b4 b2 c with [ pair l c ⇒ 〈plus_b8_dc_d b3 b1 c:l〉 ]);
- nrewrite > (symmetric_plusb8_dc_dc b4 b2 c);
- ncases (plus_b8_dc_dc b2 b4 c);
- #b5; #c1;
- nchange with (〈plus_b8_dc_d b1 b3 c1:b5〉 = 〈plus_b8_dc_d b3 b1 c1:b5〉);
- nrewrite > (symmetric_plusb8_dc_d b1 b3 c1);
- napply refl_eq.
-nqed.
-
-nlemma symmetric_plusw16_dc_c : ∀w1,w2,c.plus_w16_dc_c w1 w2 c = plus_w16_dc_c w2 w1 c.
- #w1; #w2; #c;
- nelim w1;
- #b1; #b2;
- nelim w2;
- #b3; #b4;
- nchange with (
- plus_b8_dc_c b1 b3 (plus_b8_dc_c b2 b4 c) =
- plus_b8_dc_c b3 b1 (plus_b8_dc_c b4 b2 c));
- nrewrite > (symmetric_plusb8_dc_c b4 b2 c);
- nrewrite > (symmetric_plusb8_dc_c b3 b1 (plus_b8_dc_c b2 b4 c));
- napply refl_eq.
-nqed.
-
-nlemma symmetric_plusw16_d_dc : ∀w1,w2.plus_w16_d_dc w1 w2 = plus_w16_d_dc w2 w1.
- #w1; #w2;
- nelim w1;
- #b1; #b2;
- nelim w2;
- #b3; #b4;
- nchange with (
- 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' ]] =
- 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' ]]);
- nrewrite > (symmetric_plusb8_d_dc b4 b2);
- ncases (plus_b8_d_dc b2 b4);
- #b5; #c;
- nchange with (
- match plus_b8_dc_dc b1 b3 c with [ pair h c' ⇒ pair … 〈h:b5〉 c' ] =
- match plus_b8_dc_dc b3 b1 c with [ pair h c' ⇒ pair … 〈h:b5〉 c' ]);
- nrewrite > (symmetric_plusb8_dc_dc b1 b3 c);
- napply refl_eq.
-nqed.
-
-nlemma symmetric_plusw16_d_d : ∀w1,w2.plus_w16_d_d w1 w2 = plus_w16_d_d w2 w1.
- #w1; #w2;
- nelim w1;
- #b1; #b2;
- nelim w2;
- #b3; #b4;
- nchange with (
- match plus_b8_d_dc b2 b4 with [ pair l c ⇒ 〈plus_b8_dc_d b1 b3 c:l〉 ] =
- match plus_b8_d_dc b4 b2 with [ pair l c ⇒ 〈plus_b8_dc_d b3 b1 c:l〉 ]);
- nrewrite > (symmetric_plusb8_d_dc b4 b2);
- ncases (plus_b8_d_dc b2 b4);
- #b5; #c;
- nchange with (〈plus_b8_dc_d b1 b3 c:b5〉 = 〈plus_b8_dc_d b3 b1 c:b5〉);
- nrewrite > (symmetric_plusb8_dc_d b1 b3 c);
- napply refl_eq.
-nqed.
-
-nlemma symmetric_plusw16_d_c : ∀w1,w2.plus_w16_d_c w1 w2 = plus_w16_d_c w2 w1.
- #w1; #w2;
- nelim w1;
- #b1; #b2;
- nelim w2;
- #b3; #b4;
- nchange with (
- plus_b8_dc_c b1 b3 (plus_b8_d_c b2 b4) =
- plus_b8_dc_c b3 b1 (plus_b8_d_c b4 b2));
- nrewrite > (symmetric_plusb8_d_c b4 b2);
- nrewrite > (symmetric_plusb8_dc_c b3 b1 (plus_b8_d_c b2 b4));
- napply refl_eq.
-nqed.
-
-nlemma symmetric_mulb8 : symmetricT byte8 word16 mul_b8.
- #b1; #b2;
- nelim b1;
- #e1; #e2;
- nelim b2;
- #e3; #e4;
- nchange with (match mul_ex e2 e4 with
- [ mk_byte8 t1_h t1_l ⇒ match mul_ex e1 e4 with
- [ mk_byte8 t2_h t2_l ⇒ match mul_ex e3 e2 with
- [ mk_byte8 t3_h t3_l ⇒ match mul_ex e1 e3 with
- [ 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〉〉
- ]]]] = match mul_ex e4 e2 with
- [ mk_byte8 t1_h t1_l ⇒ match mul_ex e3 e2 with
- [ mk_byte8 t2_h t2_l ⇒ match mul_ex e1 e4 with
- [ mk_byte8 t3_h t3_l ⇒ match mul_ex e3 e1 with
- [ 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〉〉
- ]]]]);
- nrewrite < (symmetric_mulex e2 e4);
- ncases (mul_ex e2 e4);
- #e5; #e6;
- nchange with (match mul_ex e1 e4 with
- [ mk_byte8 t2_h t2_l ⇒ match mul_ex e3 e2 with
- [ mk_byte8 t3_h t3_l ⇒ match mul_ex e1 e3 with
- [ 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〉〉
- ]]] = match mul_ex e3 e2 with
- [ mk_byte8 t2_h t2_l ⇒ match mul_ex e1 e4 with
- [ mk_byte8 t3_h t3_l ⇒ match mul_ex e3 e1 with
- [ 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〉〉
- ]]]);
- ncases (mul_ex e3 e2);
- #e7; #e8;
- ncases (mul_ex e1 e4);
- #e9; #e10;
- nchange with (match mul_ex e1 e3 with
- [ 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〉〉
- ] = match mul_ex e3 e1 with
- [ 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〉〉
- ]);
- nrewrite < (symmetric_mulex e1 e3);
- ncases (mul_ex e1 e3);
- #e11; #e12;
- 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〉〉) =
- (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〉〉));
- nrewrite > (symmetric_plusw16_d_d 〈〈x0,e7〉:〈e8,x0〉〉 〈〈x0,e9〉:〈e10,x0〉〉);
- napply refl_eq.
-nqed.
-
-nlemma eqw16_to_eq : ∀w1,w2:word16.(eq_w16 w1 w2 = true) → (w1 = w2).
- #w1; #w2;
- nelim w1;
- nelim w2;
- #b1; #b2; #b3; #b4;
- nchange in ⊢ (% → ?) with (((eq_b8 b3 b1)⊗(eq_b8 b4 b2)) = true);
- #H;
- nrewrite < (eqb8_to_eq … (andb_true_true_l … H));
- nrewrite < (eqb8_to_eq … (andb_true_true_r … H));
- napply refl_eq.
-nqed.
-
-nlemma eq_to_eqw16 : ∀w1,w2.w1 = w2 → eq_w16 w1 w2 = true.
- #w1; #w2;
- nelim w1;
- nelim w2;
- #b1; #b2; #b3; #b4;
- #H;
- nrewrite < (word16_destruct_1 … H);
- nrewrite < (word16_destruct_2 … H);
- nchange with (((eq_b8 b3 b3)⊗(eq_b8 b4 b4)) = true);
- nrewrite > (eq_to_eqb8 b3 b3 (refl_eq …));
- nrewrite > (eq_to_eqb8 b4 b4 (refl_eq …));
- nnormalize;
- napply refl_eq.
-nqed.
-
-nlemma decidable_w16_aux1 : ∀b1,b2,b3,b4.b1 ≠ b3 → 〈b1:b2〉 ≠ 〈b3:b4〉.
- #b1; #b2; #b3; #b4;
- nnormalize; #H; #H1;
- napply (H (word16_destruct_1 … H1)).
-nqed.
-
-nlemma decidable_w16_aux2 : ∀b1,b2,b3,b4.b2 ≠ b4 → 〈b1:b2〉 ≠ 〈b3:b4〉.
- #b1; #b2; #b3; #b4;
- nnormalize; #H; #H1;
- napply (H (word16_destruct_2 … H1)).
-nqed.
-
-nlemma decidable_w16 : ∀x,y:word16.decidable (x = y).
- #x; nelim x; #b1; #b2;
- #y; nelim y; #b3; #b4;
- nnormalize;
- napply (or2_elim (b1 = b3) (b1 ≠ b3) ? (decidable_b8 b1 b3) …);
- ##[ ##2: #H; napply (or2_intro2 … (decidable_w16_aux1 b1 b2 b3 b4 H))
- ##| ##1: #H; napply (or2_elim (b2 = b4) (b2 ≠ b4) ? (decidable_b8 b2 b4) …);
- ##[ ##2: #H1; napply (or2_intro2 … (decidable_w16_aux2 b1 b2 b3 b4 H1))
- ##| ##1: #H1; nrewrite > H; nrewrite > H1;
- napply (or2_intro1 … (refl_eq ? 〈b3:b4〉))
- ##]
- ##]
-nqed.
-
-nlemma neqw16_to_neq : ∀w1,w2:word16.(eq_w16 w1 w2 = false) → (w1 ≠ w2).
- #w1; #w2;
- nelim w1;
- nelim w2;
- #b1; #b2; #b3; #b4;
- nchange with ((((eq_b8 b3 b1) ⊗ (eq_b8 b4 b2)) = false) → ?);
- #H;
- napply (or2_elim ((eq_b8 b3 b1) = false) ((eq_b8 b4 b2) = false) ? (andb_false2 … H) …);
- ##[ ##1: #H1; napply (decidable_w16_aux1 … (neqb8_to_neq … H1))
- ##| ##2: #H1; napply (decidable_w16_aux2 … (neqb8_to_neq … H1))
- ##]
-nqed.
-
-nlemma word16_destruct : ∀b1,b2,b3,b4.〈b1:b2〉 ≠ 〈b3:b4〉 → b1 ≠ b3 ∨ b2 ≠ b4.
- #b1; #b2; #b3; #b4;
- nnormalize; #H;
- napply (or2_elim (b1 = b3) (b1 ≠ b3) ? (decidable_b8 b1 b3) …);
- ##[ ##2: #H1; napply (or2_intro1 … H1)
- ##| ##1: #H1; napply (or2_elim (b2 = b4) (b2 ≠ b4) ? (decidable_b8 b2 b4) …);
- ##[ ##2: #H2; napply (or2_intro2 … H2)
- ##| ##1: #H2; nrewrite > H1 in H:(%);
- nrewrite > H2;
- #H; nelim (H (refl_eq …))
- ##]
- ##]
-nqed.
-
-nlemma neq_to_neqw16 : ∀w1,w2.w1 ≠ w2 → eq_w16 w1 w2 = false.
- #w1; #w2;
- nelim w1; #b1; #b2;
- nelim w2; #b3; #b4;
- #H; nchange with (((eq_b8 b1 b3) ⊗ (eq_b8 b2 b4)) = false);
- napply (or2_elim (b1 ≠ b3) (b2 ≠ b4) ? (word16_destruct … H) …);
- ##[ ##1: #H1; nrewrite > (neq_to_neqb8 … H1); nnormalize; napply refl_eq
- ##| ##2: #H1; nrewrite > (neq_to_neqb8 … H1);
- nrewrite > (symmetric_andbool (eq_b8 b1 b3) false);
- nnormalize; napply refl_eq
- ##]
-nqed.
+ndefinition neqw16_to_neq ≝ neqcn_to_neq ? eq_b8 neqb8_to_neq.
+ndefinition neq_to_neqw16 ≝ neq_to_neqcn ? eq_b8 neq_to_neqb8 decidable_b8.
for @{ 'mk_word24 $x $y $z }.
interpretation "mk_word24" 'mk_word24 x y z = (mk_word24 x y z).
+(* iteratore sulle word *)
+ndefinition forall_w24 ≝
+ λP.
+ forall_b8 (λbx.
+ forall_b8 (λbh.
+ forall_b8 (λbl.
+ P (mk_word24 bx bh bl )))).
+
(* operatore = *)
ndefinition eq_w24 ≝
λw1,w2.(eq_b8 (w24x w1) (w24x w2)) ⊗
(xor_b8 (w24h w1) (w24h w2))
(xor_b8 (w24l w1) (w24l w2)).
+(* operatore Most Significant Bit *)
+ndefinition getMSB_w24 ≝ λw:word24.getMSB_b8 (w24x w).
+ndefinition setMSB_w24 ≝ λw:word24.mk_word24 (setMSB_b8 (w24x w)) (w24h w) (w24l w).
+
+(* operatore Least Significant Bit *)
+ndefinition getLSB_w24 ≝ λw:word24.getLSB_b8 (w24l w).
+ndefinition setLSB_w24 ≝ λw:word24.mk_word24 (w24x w) (w24h w) (setLSB_b8 (w24l w)).
+
(* operatore rotazione destra con carry *)
ndefinition rcr_w24 ≝
-λw:word24.λc:bool.match rcr_b8 (w24x w) c with
- [ pair wx' c' ⇒ match rcr_b8 (w24h w) c' with
- [ pair wh' c'' ⇒ match rcr_b8 (w24l w) c'' with
- [ pair wl' c''' ⇒ pair … (mk_word24 wx' wh' wl') c''' ]]].
+λc:bool.λw:word24.match rcr_b8 c (w24x w) with
+ [ pair c' wx' ⇒ match rcr_b8 c' (w24h w) with
+ [ pair c'' wh' ⇒ match rcr_b8 c'' (w24l w) with
+ [ pair c''' wl' ⇒ pair … c''' (mk_word24 wx' wh' wl') ]]].
(* operatore shift destro *)
-ndefinition shr_w24 ≝
-λw:word24.match shr_b8 (w24x w) with
- [ pair wx' c' ⇒ match rcr_b8 (w24h w) c' with
- [ pair wh' c'' ⇒ match rcr_b8 (w24l w) c'' with
- [ pair wl' c''' ⇒ pair … (mk_word24 wx' wh' wl') c''' ]]].
+ndefinition shr_w24 ≝ rcr_w24 false.
(* operatore rotazione destra *)
ndefinition ror_w24 ≝
-λw:word24.match shr_b8 (w24x w) with
- [ pair wx' c' ⇒ match rcr_b8 (w24h w) c' with
- [ pair wh' c'' ⇒ match rcr_b8 (w24l w) c'' with
- [ pair wl' c''' ⇒ match c''' with
- [ true ⇒ mk_word24 (or_b8 (mk_byte8 x8 x0) wx') wh' wl'
- | false ⇒ mk_word24 wx' wh' wl' ]]]].
-
-(* operatore rotazione destra n-volte *)
-nlet rec ror_w24_n (w:word24) (t:bitrigesim) (rt:rec_bitrigesim t) on rt ≝
- match rt with
- [ bi_O ⇒ w
- | bi_S t' n' ⇒ ror_w24_n (ror_w24 w) t' n' ].
+λw.match shr_w24 w with
+ [ pair c w' ⇒ match c with
+ [ true ⇒ setMSB_w24 w' | false ⇒ w' ]].
(* operatore rotazione sinistra con carry *)
ndefinition rcl_w24 ≝
-λw:word24.λc:bool.match rcl_b8 (w24l w) c with
- [ pair wl' c' ⇒ match rcl_b8 (w24h w) c' with
- [ pair wh' c'' ⇒ match rcl_b8 (w24x w) c'' with
- [ pair wx' c''' ⇒ pair … (mk_word24 wx' wh' wl') c''' ]]].
+λc:bool.λw:word24.match rcl_b8 c (w24l w) with
+ [ pair c' wl' ⇒ match rcl_b8 c' (w24h w) with
+ [ pair c'' wh' ⇒ match rcl_b8 c'' (w24x w) with
+ [ pair c''' wx' ⇒ pair … c''' (mk_word24 wx' wh' wl') ]]].
(* operatore shift sinistro *)
-ndefinition shl_w24 ≝
-λw:word24.match shl_b8 (w24l w) with
- [ pair wl' c' ⇒ match rcl_b8 (w24h w) c' with
- [ pair wh' c'' ⇒ match rcl_b8 (w24x w) c'' with
- [ pair wx' c''' ⇒ pair … (mk_word24 wx' wh' wl') c''' ]]].
+ndefinition shl_w24 ≝ rcl_w24 false.
(* operatore rotazione sinistra *)
ndefinition rol_w24 ≝
-λw:word24.match shl_b8 (w24l w) with
- [ pair wl' c' ⇒ match rcl_b8 (w24h w) c' with
- [ pair wh' c'' ⇒ match rcl_b8 (w24x w) c'' with
- [ pair wx' c''' ⇒ match c''' with
- [ true ⇒ mk_word24 wx' wh' (or_b8 (mk_byte8 x0 x1) wl')
- | false ⇒ mk_word24 wx' wh' wl' ]]]].
-
-(* operatore rotazione sinistra n-volte *)
-nlet rec rol_w24_n (w:word24) (t:bitrigesim) (rt:rec_bitrigesim t) on rt ≝
- match rt with
- [ bi_O ⇒ w
- | bi_S t' n' ⇒ rol_w24_n (rol_w24 w) t' n' ].
+λw.match shl_w24 w with
+ [ pair c w' ⇒ match c with
+ [ true ⇒ setLSB_w24 w' | false ⇒ w' ]].
(* operatore not/complemento a 1 *)
ndefinition not_w24 ≝
(* operatore somma con data+carry → data+carry *)
ndefinition plus_w24_dc_dc ≝
-λw1,w2:word24.λc:bool.
- match plus_b8_dc_dc (w24l w1) (w24l w2) c with
- [ pair l c' ⇒ match plus_b8_dc_dc (w24h w1) (w24h w2) c' with
- [ pair h c'' ⇒ match plus_b8_dc_dc (w24x w1) (w24x w2) c'' with
- [ pair x c''' ⇒ pair … 〈x;h;l〉 c''' ]]].
+λc:bool.λw1,w2:word24.
+ match plus_b8_dc_dc c (w24l w1) (w24l w2) with
+ [ pair c' l ⇒ match plus_b8_dc_dc c' (w24h w1) (w24h w2) with
+ [ pair c'' h ⇒ match plus_b8_dc_dc c'' (w24x w1) (w24x w2) with
+ [ pair c''' x ⇒ pair … c''' 〈x;h;l〉 ]]].
(* operatore somma con data+carry → data *)
-ndefinition plus_w24_dc_d ≝
-λw1,w2:word24.λc:bool.
- match plus_b8_dc_dc (w24l w1) (w24l w2) c with
- [ pair l c' ⇒ match plus_b8_dc_dc (w24h w1) (w24h w2) c' with
- [ pair h c'' ⇒ 〈plus_b8_dc_d (w24x w1) (w24x w2) c'';h;l〉 ]].
+ndefinition plus_w24_dc_d ≝ λc,w1,w2.snd … (plus_w24_dc_dc c w1 w2).
(* operatore somma con data+carry → c *)
-ndefinition plus_w24_dc_c ≝
-λw1,w2:word24.λc:bool.
- plus_b8_dc_c (w24x w1) (w24x w2) (plus_b8_dc_c (w24h w1) (w24h w2) (plus_b8_dc_c (w24l w1) (w24l w2) c)).
+ndefinition plus_w24_dc_c ≝ λc,w1,w2.fst … (plus_w24_dc_dc c w1 w2).
(* operatore somma con data → data+carry *)
-ndefinition plus_w24_d_dc ≝
-λw1,w2:word24.
- match plus_b8_d_dc (w24l w1) (w24l w2) with
- [ pair l c ⇒ match plus_b8_dc_dc (w24h w1) (w24h w2) c with
- [ pair h c' ⇒ match plus_b8_dc_dc (w24x w1) (w24x w2) c' with
- [ pair x c'' ⇒ pair … 〈x;h;l〉 c'' ]]].
+ndefinition plus_w24_d_dc ≝ plus_w24_dc_dc false.
(* operatore somma con data → data *)
-ndefinition plus_w24_d_d ≝
-λw1,w2:word24.
- match plus_b8_d_dc (w24l w1) (w24l w2) with
- [ pair l c ⇒ match plus_b8_dc_dc (w24h w1) (w24h w2) c with
- [ pair h c' ⇒ 〈plus_b8_dc_d (w24x w1) (w24x w2) c';h;l〉 ]].
+ndefinition plus_w24_d_d ≝ λw1,w2.snd … (plus_w24_d_dc w1 w2).
(* operatore somma con data → c *)
-ndefinition plus_w24_d_c ≝
-λw1,w2:word24.
- plus_b8_dc_c (w24x w1) (w24x w2) (plus_b8_dc_c (w24h w1) (w24h w2) (plus_b8_d_c (w24l w1) (w24l w2))).
-
-(* operatore Most Significant Bit *)
-ndefinition MSB_w24 ≝ λw:word24.eq_ex x8 (and_ex x8 (b8h (w24x w))).
+ndefinition plus_w24_d_c ≝ λw1,w2.fst … (plus_w24_d_dc w1 w2).
(* operatore predecessore *)
ndefinition pred_w24 ≝
-λw:word24.match eq_b8 (w24l w) (mk_byte8 x0 x0) with
- [ true ⇒ match eq_b8 (w24h w) (mk_byte8 x0 x0) with
+λw:word24.match eq_b8 (w24l w) 〈x0,x0〉 with
+ [ true ⇒ match eq_b8 (w24h w) 〈x0,x0〉 with
[ true ⇒ mk_word24 (pred_b8 (w24x w)) (pred_b8 (w24h w)) (pred_b8 (w24l w))
| false ⇒ mk_word24 (w24x w) (pred_b8 (w24h w)) (pred_b8 (w24l w)) ]
| false ⇒ mk_word24 (w24x w) (w24h w) (pred_b8 (w24l w)) ].
(* operatore successore *)
ndefinition succ_w24 ≝
-λw:word24.match eq_b8 (w24l w) (mk_byte8 xF xF) with
- [ true ⇒ match eq_b8 (w24h w) (mk_byte8 xF xF) with
+λw:word24.match eq_b8 (w24l w) 〈xF,xF〉 with
+ [ true ⇒ match eq_b8 (w24h w) 〈xF,xF〉 with
[ true ⇒ mk_word24 (succ_b8 (w24x w)) (succ_b8 (w24h w)) (succ_b8 (w24l w))
| false ⇒ mk_word24 (w24x w) (succ_b8 (w24h w)) (succ_b8 (w24l w)) ]
| false ⇒ mk_word24 (w24x w) (w24h w) (succ_b8 (w24l w)) ].
(* operatore neg/complemento a 2 *)
ndefinition compl_w24 ≝
-λw:word24.match MSB_w24 w with
+λw:word24.match getMSB_w24 w with
[ true ⇒ succ_w24 (not_w24 w)
| false ⇒ not_w24 (pred_w24 w) ].
match le_w24 inf sup with
[ true ⇒ and_bool | false ⇒ or_bool ]
(le_w24 inf x) (le_w24 x sup).
-
-(* iteratore sulle word *)
-ndefinition forall_w24 ≝
- λP.
- forall_b8 (λbx.
- forall_b8 (λbh.
- forall_b8 (λbl.
- P (mk_word24 bx bh bl )))).
(* DWORD *)
(* ***** *)
-nrecord word32 : Type ≝
- {
- w32h: word16;
- w32l: word16
- }.
+ndefinition word32 ≝ comp_num word16.
+ndefinition mk_word32 ≝ λw1,w2.mk_comp_num word16 w1 w2.
(* \langle \rangle *)
notation "〈x.y〉" non associative with precedence 80
- for @{ 'mk_word32 $x $y }.
-interpretation "mk_word32" 'mk_word32 x y = (mk_word32 x y).
+ for @{ mk_comp_num word16 $x $y }.
+
+(* iteratore sulle dword *)
+ndefinition forall_w32 ≝ forall_cn ? forall_w16.
(* operatore = *)
-ndefinition eq_w32 ≝ λdw1,dw2.(eq_w16 (w32h dw1) (w32h dw2)) ⊗ (eq_w16 (w32l dw1) (w32l dw2)).
+ndefinition eq_w32 ≝ eq2_cn ? eq_w16.
(* operatore < *)
-ndefinition lt_w32 ≝
-λdw1,dw2:word32.
- (lt_w16 (w32h dw1) (w32h dw2)) ⊕
- ((eq_w16 (w32h dw1) (w32h dw2)) ⊗ (lt_w16 (w32l dw1) (w32l dw2))).
+ndefinition lt_w32 ≝ ltgt_cn ? eq_w16 lt_w16.
(* operatore ≤ *)
-ndefinition le_w32 ≝
-λdw1,dw2:word32.
- (lt_w16 (w32h dw1) (w32h dw2)) ⊕
- ((eq_w16 (w32h dw1) (w32h dw2)) ⊗ (le_w16 (w32l dw1) (w32l dw2))).
+ndefinition le_w32 ≝ lege_cn ? eq_w16 lt_w16 le_w16.
(* operatore > *)
-ndefinition gt_w32 ≝
-λdw1,dw2:word32.
- (gt_w16 (w32h dw1) (w32h dw2)) ⊕
- ((eq_w16 (w32h dw1) (w32h dw2)) ⊗ (gt_w16 (w32l dw1) (w32l dw2))).
+ndefinition gt_w32 ≝ ltgt_cn ? eq_w16 gt_w16.
(* operatore ≥ *)
-ndefinition ge_w32 ≝
-λdw1,dw2:word32.
- (gt_w16 (w32h dw1) (w32h dw2)) ⊕
- ((eq_w16 (w32h dw1) (w32h dw2)) ⊗ (ge_w16 (w32l dw1) (w32l dw2))).
+ndefinition ge_w32 ≝ lege_cn ? eq_w16 gt_w16 ge_w16.
(* operatore and *)
-ndefinition and_w32 ≝
-λdw1,dw2:word32.mk_word32 (and_w16 (w32h dw1) (w32h dw2)) (and_w16 (w32l dw1) (w32l dw2)).
+ndefinition and_w32 ≝ fop2_cn ? and_w16.
(* operatore or *)
-ndefinition or_w32 ≝
-λdw1,dw2:word32.mk_word32 (or_w16 (w32h dw1) (w32h dw2)) (or_w16 (w32l dw1) (w32l dw2)).
+ndefinition or_w32 ≝ fop2_cn ? or_w16.
(* operatore xor *)
-ndefinition xor_w32 ≝
-λdw1,dw2:word32.mk_word32 (xor_w16 (w32h dw1) (w32h dw2)) (xor_w16 (w32l dw1) (w32l dw2)).
+ndefinition xor_w32 ≝ fop2_cn ? xor_w16.
+
+(* operatore Most Significant Bit *)
+ndefinition getMSB_w32 ≝ getOPH_cn ? getMSB_w16.
+ndefinition setMSB_w32 ≝ setOPH_cn ? setMSB_w16.
+
+(* operatore Least Significant Bit *)
+ndefinition getLSB_w32 ≝ getOPL_cn ? getLSB_w16.
+ndefinition setLSB_w32 ≝ setOPL_cn ? setLSB_w16.
(* operatore rotazione destra con carry *)
-ndefinition rcr_w32 ≝
-λdw:word32.λc:bool.match rcr_w16 (w32h dw) c with
- [ pair wh' c' ⇒ match rcr_w16 (w32l dw) c' with
- [ pair wl' c'' ⇒ pair … (mk_word32 wh' wl') c'' ]].
+ndefinition rcr_w32 ≝ opcr_cn ? rcr_w16.
(* operatore shift destro *)
-ndefinition shr_w32 ≝
-λdw:word32.match rcr_w16 (w32h dw) false with
- [ pair wh' c' ⇒ match rcr_w16 (w32l dw) c' with
- [ pair wl' c'' ⇒ pair … (mk_word32 wh' wl') c'' ]].
+ndefinition shr_w32 ≝ opcr_cn ? rcr_w16 false.
(* operatore rotazione destra *)
ndefinition ror_w32 ≝
-λdw:word32.match rcr_w16 (w32h dw) false with
- [ pair wh' c' ⇒ match rcr_w16 (w32l dw) c' with
- [ pair wl' c'' ⇒ match c'' with
- [ true ⇒ mk_word32 (or_w16 (mk_word16 (mk_byte8 x8 x0) (mk_byte8 x0 x0)) wh') wl'
- | false ⇒ mk_word32 wh' wl' ]]].
-
-(* operatore rotazione destra n-volte *)
-nlet rec ror_w32_n (dw:word32) (t:bitrigesim) (rt:rec_bitrigesim t) on rt ≝
- match rt with
- [ bi_O ⇒ dw
- | bi_S t' n' ⇒ ror_w32_n (ror_w32 dw) t' n' ].
+λw.match shr_w32 w with
+ [ pair c w' ⇒ match c with
+ [ true ⇒ setMSB_w32 w' | false ⇒ w' ]].
(* operatore rotazione sinistra con carry *)
-ndefinition rcl_w32 ≝
-λdw:word32.λc:bool.match rcl_w16 (w32l dw) c with
- [ pair wl' c' ⇒ match rcl_w16 (w32h dw) c' with
- [ pair wh' c'' ⇒ pair … (mk_word32 wh' wl') c'' ]].
+ndefinition rcl_w32 ≝ opcl_cn ? rcl_w16.
(* operatore shift sinistro *)
-ndefinition shl_w32 ≝
-λdw:word32.match rcl_w16 (w32l dw) false with
- [ pair wl' c' ⇒ match rcl_w16 (w32h dw) c' with
- [ pair wh' c'' ⇒ pair … (mk_word32 wh' wl') c'' ]].
+ndefinition shl_w32 ≝ opcl_cn ? rcl_w16 false.
(* operatore rotazione sinistra *)
ndefinition rol_w32 ≝
-λdw:word32.match rcl_w16 (w32l dw) false with
- [ pair wl' c' ⇒ match rcl_w16 (w32h dw) c' with
- [ pair wh' c'' ⇒ match c'' with
- [ true ⇒ mk_word32 wh' (or_w16 (mk_word16 (mk_byte8 x0 x0) (mk_byte8 x0 x1)) wl')
- | false ⇒ mk_word32 wh' wl' ]]].
-
-(* operatore rotazione sinistra n-volte *)
-nlet rec rol_w32_n (dw:word32) (t:bitrigesim) (rt:rec_bitrigesim t) on rt ≝
- match rt with
- [ bi_O ⇒ dw
- | bi_S t' n' ⇒ rol_w32_n (rol_w32 dw) t' n' ].
+λw.match shl_w32 w with
+ [ pair c w' ⇒ match c with
+ [ true ⇒ setLSB_w32 w' | false ⇒ w' ]].
(* operatore not/complemento a 1 *)
-ndefinition not_w32 ≝
-λdw:word32.mk_word32 (not_w16 (w32h dw)) (not_w16 (w32l dw)).
+ndefinition not_w32 ≝ fop_cn ? not_w16.
(* operatore somma con data+carry → data+carry *)
-ndefinition plus_w32_dc_dc ≝
-λdw1,dw2:word32.λc:bool.
- match plus_w16_dc_dc (w32l dw1) (w32l dw2) c with
- [ pair l c' ⇒ match plus_w16_dc_dc (w32h dw1) (w32h dw2) c' with
- [ pair h c'' ⇒ pair … 〈h.l〉 c'' ]].
+ndefinition plus_w32_dc_dc ≝ opcl2_cn ? plus_w16_dc_dc.
(* operatore somma con data+carry → data *)
-ndefinition plus_w32_dc_d ≝
-λdw1,dw2:word32.λc:bool.
- match plus_w16_dc_dc (w32l dw1) (w32l dw2) c with
- [ pair l c' ⇒ 〈plus_w16_dc_d (w32h dw1) (w32h dw2) c'.l〉 ].
+ndefinition plus_w32_dc_d ≝ λc,w1,w2.snd … (plus_w32_dc_dc c w1 w2).
(* operatore somma con data+carry → c *)
-ndefinition plus_w32_dc_c ≝
-λdw1,dw2:word32.λc:bool.
- plus_w16_dc_c (w32h dw1) (w32h dw2) (plus_w16_dc_c (w32l dw1) (w32l dw2) c).
+ndefinition plus_w32_dc_c ≝ λc,w1,w2.fst … (plus_w32_dc_dc c w1 w2).
(* operatore somma con data → data+carry *)
-ndefinition plus_w32_d_dc ≝
-λdw1,dw2:word32.
- match plus_w16_d_dc (w32l dw1) (w32l dw2) with
- [ pair l c ⇒ match plus_w16_dc_dc (w32h dw1) (w32h dw2) c with
- [ pair h c' ⇒ pair … 〈h.l〉 c' ]].
+ndefinition plus_w32_d_dc ≝ opcl2_cn ? plus_w16_dc_dc false.
(* operatore somma con data → data *)
-ndefinition plus_w32_d_d ≝
-λdw1,dw2:word32.
- match plus_w16_d_dc (w32l dw1) (w32l dw2) with
- [ pair l c ⇒ 〈plus_w16_dc_d (w32h dw1) (w32h dw2) c.l〉 ].
+ndefinition plus_w32_d_d ≝ λw1,w2.snd … (plus_w32_d_dc w1 w2).
(* operatore somma con data → c *)
-ndefinition plus_w32_d_c ≝
-λdw1,dw2:word32.
- plus_w16_dc_c (w32h dw1) (w32h dw2) (plus_w16_d_c (w32l dw1) (w32l dw2)).
-
-(* operatore Most Significant Bit *)
-ndefinition MSB_w32 ≝ λdw:word32.eq_ex x8 (and_ex x8 (b8h (w16h (w32h dw)))).
+ndefinition plus_w32_d_c ≝ λw1,w2.fst … (plus_w32_d_dc w1 w2).
(* operatore predecessore *)
-ndefinition pred_w32 ≝
-λdw:word32.match eq_w16 (w32l dw) (mk_word16 (mk_byte8 x0 x0) (mk_byte8 x0 x0)) with
- [ true ⇒ mk_word32 (pred_w16 (w32h dw)) (pred_w16 (w32l dw))
- | false ⇒ mk_word32 (w32h dw) (pred_w16 (w32l dw)) ].
+ndefinition pred_w32 ≝ predsucc_cn ? (eq_w16 〈〈x0,x0〉:〈x0,x0〉〉) pred_w16.
(* operatore successore *)
-ndefinition succ_w32 ≝
-λdw:word32.match eq_w16 (w32l dw) (mk_word16 (mk_byte8 xF xF) (mk_byte8 xF xF)) with
- [ true ⇒ mk_word32 (succ_w16 (w32h dw)) (succ_w16 (w32l dw))
- | false ⇒ mk_word32 (w32h dw) (succ_w16 (w32l dw)) ].
+ndefinition succ_w32 ≝ predsucc_cn ? (eq_w16 〈〈xF,xF〉:〈xF,xF〉〉) succ_w16.
(* operatore neg/complemento a 2 *)
ndefinition compl_w32 ≝
-λdw:word32.match MSB_w32 dw with
- [ true ⇒ succ_w32 (not_w32 dw)
- | false ⇒ not_w32 (pred_w32 dw) ].
-
-(*
- operatore moltiplicazione senza segno: b*b=[0x00000000,0xFFFE0001]
- ... in pratica (〈a:b〉*〈c:d〉) = (a*c)<<16+(a*d)<<8+(b*c)<<8+(b*d)
-*)
-ndefinition mul_w16 ≝
-λw1,w2:word16.match w1 with
-[ mk_word16 b1h b1l ⇒ match w2 with
-[ mk_word16 b2h b2l ⇒ match mul_b8 b1l b2l with
-[ mk_word16 t1_h t1_l ⇒ match mul_b8 b1h b2l with
-[ mk_word16 t2_h t2_l ⇒ match mul_b8 b2h b1l with
-[ mk_word16 t3_h t3_l ⇒ match mul_b8 b1h b2h 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〉〉
-]]]]]].
-
-(* divisione senza segno (secondo la logica delle ALU): (quoziente resto) overflow *)
-nlet rec div_w32_w16_aux (divd:word32) (divs:word32) (molt:word16) (q:word16) (e:exadecim) (re:rec_exadecim e) on re ≝
- let w' ≝ plus_w32_d_d divd (compl_w32 divs) in
- match re with
- [ ex_O ⇒ match le_w32 divs divd with
- [ true ⇒ triple … (or_w16 molt q) (w32l w') (⊖ (eq_w16 (w32h w') 〈〈x0,x0〉:〈x0,x0〉〉))
- | false ⇒ triple … q (w32l divd) (⊖ (eq_w16 (w32h divd) 〈〈x0,x0〉:〈x0,x0〉〉)) ]
- | ex_S e' re' ⇒ match le_w32 divs divd with
- [ true ⇒ div_w32_w16_aux w' (ror_w32 divs) (ror_w16 molt) (or_w16 molt q) e' re'
- | false ⇒ div_w32_w16_aux divd (ror_w32 divs) (ror_w16 molt) q e' re' ]].
-
-ndefinition div_w32_w16 ≝
-λw:word32.λb:word16.match eq_w16 b 〈〈x0,x0〉:〈x0,x0〉〉 with
-(*
- la combinazione n/0 e' illegale, segnala solo overflow senza dare risultato
-*)
- [ true ⇒ triple … 〈〈xF,xF〉:〈xF,xF〉〉 (w32l w) true
- | false ⇒ match eq_w32 w 〈〈〈x0,x0〉:〈x0,x0〉〉.〈〈x0,x0〉:〈x0,x0〉〉〉 with
-(* 0 diviso qualsiasi cosa diverso da 0 da' q=0 r=0 o=false *)
- [ true ⇒ triple … 〈〈x0,x0〉:〈x0,x0〉〉 〈〈x0,x0〉:〈x0,x0〉〉 false
-(* 1) e' una divisione sensata che produrra' overflow/risultato *)
-(* 2) parametri: dividendo, divisore, moltiplicatore, quoziente, contatore *)
-(* 3) ad ogni ciclo il divisore e il moltiplicatore vengono scalati di 1 a dx *)
-(* 4) il moltiplicatore e' la quantita' aggiunta al quoziente se il divisore *)
-(* puo' essere sottratto al dividendo *)
- | false ⇒ div_w32_w16_aux w (rol_w32_n 〈〈〈x0,x0〉:〈x0,x0〉〉.b〉 ? (bit_to_recbit t0F)) 〈〈x8,x0〉:〈x0,x0〉〉 〈〈x0,x0〉:〈x0,x0〉〉 ? (ex_to_recex xF) ]].
+λw:word32.match getMSB_w32 w with
+ [ true ⇒ succ_w32 (not_w32 w)
+ | false ⇒ not_w32 (pred_w32 w) ].
(* operatore x in [inf,sup] o in sup],[inf *)
ndefinition inrange_w32 ≝
match le_w32 inf sup with
[ true ⇒ and_bool | false ⇒ or_bool ]
(le_w32 inf x) (le_w32 x sup).
-
-(* iteratore sulle word *)
-ndefinition forall_w32 ≝
- λP.
- forall_w16 (λwh.
- forall_w16 (λwl.
- P (mk_word32 wh wl ))).
(* DWORD *)
(* ***** *)
-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 word32_destruct_1 ≝ cn_destruct_1 word32.
+ndefinition word32_destruct_2 ≝ cn_destruct_2 word32.
-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 symmetric_eqw32 ≝ symmetric_eqcn ? eq_w16 symmetric_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 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_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 decidable_w32 ≝ decidable_cn ? 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.
+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.