match ast with
[ AST_TYPE_BASE b ⇒ eval_size_base_type b
| AST_TYPE_ARRAY sub_ast dim ⇒ (dim+1)*(eval_size_type sub_ast)
- | AST_TYPE_STRUCT nel_ast ⇒ fold_right_neList ?? (λt,x.(eval_size_type t)+x) O nel_ast
+ | AST_TYPE_STRUCT nel_ast ⇒ fold_right_neList … (λt,x.(eval_size_type t)+x) O nel_ast
].
nqed.
nlemma symmetric_eqastbasetype : symmetricT ast_base_type bool eq_ast_base_type.
- #b1; #b2; ncases b1; ncases b2; nnormalize; napply (refl_eq ??). nqed.
+ #b1; #b2; ncases b1; ncases b2; nnormalize; napply refl_eq. nqed.
nlemma eqastbasetype_to_eq : ∀b1,b2.eq_ast_base_type b1 b2 = true → b1 = b2.
#b1; #b2; ncases b1; ncases b2; nnormalize;
- ##[ ##1,5,9: #H; napply (refl_eq ??)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##1,5,9: #H; napply refl_eq
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eq_to_eqastbasetype : ∀b1,b2.b1 = b2 → eq_ast_base_type b1 b2 = true.
#b1; #b2; ncases b1; ncases b2; nnormalize;
- ##[ ##1,5,9: #H; napply (refl_eq ??)
- ##| ##*: #H; napply (astbasetype_destruct ??? H)
+ ##[ ##1,5,9: #H; napply refl_eq
+ ##| ##*: #H; napply (astbasetype_destruct … H)
##]
nqed.
nchange with (match AST_TYPE_BASE b2 with [ AST_TYPE_BASE a ⇒ b1 = a | _ ⇒ False ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma asttype_destruct_array_array_1 : ∀x1,x2,y1,y2.AST_TYPE_ARRAY x1 y1 = AST_TYPE_ARRAY x2 y2 → x1 = x2.
nchange with (match AST_TYPE_ARRAY x2 y2 with [ AST_TYPE_ARRAY a _ ⇒ x1 = a | _ ⇒ False ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma asttype_destruct_array_array_2 : ∀x1,x2,y1,y2.AST_TYPE_ARRAY x1 y1 = AST_TYPE_ARRAY x2 y2 → y1 = y2.
nchange with (match AST_TYPE_ARRAY x2 y2 with [ AST_TYPE_ARRAY _ b ⇒ y1 = b | _ ⇒ False ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma asttype_destruct_struct_struct : ∀b1,b2.AST_TYPE_STRUCT b1 = AST_TYPE_STRUCT b2 → b1 = b2.
nchange with (match AST_TYPE_STRUCT b2 with [ AST_TYPE_STRUCT a ⇒ b1 = a | _ ⇒ False ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
ndefinition asttype_destruct_aux ≝
##[ ##1: ncases b2;
##[ ##1: nnormalize; #s1; #s2; ncases s1; ncases s2; nnormalize;
##[ ##1,5,9: #H; napply (λx:P.x)
- ##| ##*: #H; napply (astbasetype_destruct ??? (asttype_destruct_base_base ?? H))
+ ##| ##*: #H; napply (astbasetype_destruct … (asttype_destruct_base_base … H))
##]
##| ##2: #t; #n; #b; nnormalize; #H
##| ##3: #l; #b; nnormalize; #H
nqed.
nlemma symmetric_eqasttype : symmetricT ast_type bool eq_ast_type.
- #t1; napply (ast_type_index ????? t1);
+ #t1; napply (ast_type_index … t1);
##[ ##1: #b1; #t2; ncases t2;
##[ ##1: #b2; nchange with ((eq_ast_base_type b1 b2) = (eq_ast_base_type b2 b1));
nrewrite > (symmetric_eqastbasetype b1 b2);
- napply (refl_eq ??)
- ##| ##2: #st2; #n2; nnormalize; napply (refl_eq ??)
- ##| ##3: #nl2; nnormalize; napply (refl_eq ??)
+ napply refl_eq
+ ##| ##2: #st2; #n2; nnormalize; napply refl_eq
+ ##| ##3: #nl2; nnormalize; napply refl_eq
##]
##| ##2: #st1; #n1; #H; #t2; ncases t2;
##[ ##2: #st2; #n2; nchange with (((eq_ast_type st1 st2)⊗(eq_nat n1 n2)) = ((eq_ast_type st2 st1)⊗(eq_nat n2 n1)));
nrewrite > (symmetric_eqnat n1 n2);
nrewrite > (H st2);
- napply (refl_eq ??)
- ##| ##1: #b2; nnormalize; napply (refl_eq ??)
- ##| ##3: #nl2; nnormalize; napply (refl_eq ??)
+ napply refl_eq
+ ##| ##1: #b2; nnormalize; napply refl_eq
+ ##| ##3: #nl2; nnormalize; napply refl_eq
##]
##| ##3: #hh1; #H; #t2; ncases t2;
##[ ##3: #nl2; ncases nl2;
##[ ##1: #hh2; nchange with ((eq_ast_type hh1 hh2) = (eq_ast_type hh2 hh1));
nrewrite > (H hh2);
- napply (refl_eq ??)
- ##| ##2: #hh2; #ll2; nnormalize; napply (refl_eq ??)
+ napply refl_eq
+ ##| ##2: #hh2; #ll2; nnormalize; napply refl_eq
##]
- ##| ##1: #b2; nnormalize; napply (refl_eq ??)
- ##| ##2: #st2; #n2; nnormalize; napply (refl_eq ??)
+ ##| ##1: #b2; nnormalize; napply refl_eq
+ ##| ##2: #st2; #n2; nnormalize; napply refl_eq
##]
##| ##4: #hh1; #ll1; #H; #H1; #t2; ncases t2;
##[ ##3: #nl2; ncases nl2;
- ##[ ##1: #hh2; nnormalize; napply (refl_eq ??)
+ ##[ ##1: #hh2; nnormalize; napply refl_eq
##| ##2: #hh2; #ll2; nnormalize;
nrewrite > (H hh2);
nrewrite > (symmetric_eqasttype_aux1 ll1 ll2 (H1 (AST_TYPE_STRUCT ll2)));
- napply (refl_eq ??)
+ napply refl_eq
##]
- ##| ##1: #b2; nnormalize; napply (refl_eq ??)
- ##| ##2: #st2; #n2; nnormalize; napply (refl_eq ??)
+ ##| ##1: #b2; nnormalize; napply refl_eq
+ ##| ##2: #st2; #n2; nnormalize; napply refl_eq
##]
##]
nqed.
nlemma eqasttype_to_eq : ∀t1,t2.eq_ast_type t1 t2 = true → t1 = t2.
#t1;
- napply (ast_type_index ????? t1);
+ napply (ast_type_index … t1);
##[ ##1: #b1; #t2; ncases t2;
##[ ##1: #b2; #H; nchange in H:(%) with ((eq_ast_base_type b1 b2) = true);
nrewrite > (eqastbasetype_to_eq b1 b2 H);
- napply (refl_eq ??)
- ##| ##2: #st2; #n2; nnormalize; #H; napply (bool_destruct ??? H)
- ##| ##3: #nl2; nnormalize; #H; napply (bool_destruct ??? H)
+ napply refl_eq
+ ##| ##2: #st2; #n2; nnormalize; #H; napply (bool_destruct … H)
+ ##| ##3: #nl2; nnormalize; #H; napply (bool_destruct … H)
##]
##| ##2: #st1; #n1; #H; #t2; ncases t2;
##[ ##2: #st2; #n2; #H1; nchange in H1:(%) with (((eq_ast_type st1 st2)⊗(eq_nat n1 n2)) = true);
- nrewrite > (H st2 (andb_true_true_l ?? H1));
- nrewrite > (eqnat_to_eq n1 n2 (andb_true_true_r ?? H1));
- napply (refl_eq ??)
- ##| ##1: #b2; nnormalize; #H1; napply (bool_destruct ??? H1)
- ##| ##3: #nl2; nnormalize; #H1; napply (bool_destruct ??? H1)
+ nrewrite > (H st2 (andb_true_true_l … H1));
+ nrewrite > (eqnat_to_eq n1 n2 (andb_true_true_r … H1));
+ napply refl_eq
+ ##| ##1: #b2; nnormalize; #H1; napply (bool_destruct … H1)
+ ##| ##3: #nl2; nnormalize; #H1; napply (bool_destruct … H1)
##]
##| ##3: #hh1; #H; #t2; ncases t2;
##[ ##3: #nl2; ncases nl2;
##[ ##1: #hh2; #H1; nchange in H1:(%) with ((eq_ast_type hh1 hh2) = true);
nrewrite > (H hh2 H1);
- napply (refl_eq ??)
- ##| ##2: #hh2; #ll2; nnormalize; #H1; napply (bool_destruct ??? H1)
+ napply refl_eq
+ ##| ##2: #hh2; #ll2; nnormalize; #H1; napply (bool_destruct … H1)
##]
- ##| ##1: #b2; nnormalize; #H1; napply (bool_destruct ??? H1)
- ##| ##2: #st2; #n2; nnormalize; #H1; napply (bool_destruct ??? H1)
+ ##| ##1: #b2; nnormalize; #H1; napply (bool_destruct … H1)
+ ##| ##2: #st2; #n2; nnormalize; #H1; napply (bool_destruct … H1)
##]
##| ##4: #hh1; #ll1; #H; #H1; #t2; ncases t2;
##[ ##3: #nl2; ncases nl2;
- ##[ ##1: #hh2; nnormalize; #H2; napply (bool_destruct ??? H2)
+ ##[ ##1: #hh2; nnormalize; #H2; napply (bool_destruct … H2)
##| ##2: #hh2; #ll2; #H2; nchange in H2:(%) with (((eq_ast_type hh1 hh2)⊗(bfold_right_neList2 ? (λx,y.eq_ast_type x y) ll1 ll2)) = true);
- nrewrite > (H hh2 (andb_true_true_l ?? H2));
- nrewrite > (asttype_destruct_struct_struct ll1 ll2 (H1 (AST_TYPE_STRUCT ll2) (andb_true_true_r ?? H2)));
- napply (refl_eq ??)
+ nrewrite > (H hh2 (andb_true_true_l … H2));
+ nrewrite > (asttype_destruct_struct_struct ll1 ll2 (H1 (AST_TYPE_STRUCT ll2) (andb_true_true_r … H2)));
+ napply refl_eq
##]
- ##| ##1: #b2; nnormalize; #H2; napply (bool_destruct ??? H2)
- ##| ##2: #st2; #n2; nnormalize; #H2; napply (bool_destruct ??? H2)
+ ##| ##1: #b2; nnormalize; #H2; napply (bool_destruct … H2)
+ ##| ##2: #st2; #n2; nnormalize; #H2; napply (bool_destruct … H2)
##]
##]
nqed.
nlemma eq_to_eqasttype : ∀t1,t2.t1 = t2 → eq_ast_type t1 t2 = true.
#t1;
- napply (ast_type_index ????? t1);
+ napply (ast_type_index … t1);
##[ ##1: #b1; #t2; ncases t2;
- ##[ ##1: #b2; #H; nrewrite > (asttype_destruct_base_base ?? H);
+ ##[ ##1: #b2; #H; nrewrite > (asttype_destruct_base_base … H);
nchange with ((eq_ast_base_type b2 b2) = true);
- nrewrite > (eq_to_eqastbasetype b2 b2 (refl_eq ??));
- napply (refl_eq ??)
- ##| ##2: #st2; #n2; #H; napply (asttype_destruct ??? H)
- ##| ##3: #nl2; #H; napply (asttype_destruct ??? H)
+ nrewrite > (eq_to_eqastbasetype b2 b2 (refl_eq …));
+ napply refl_eq
+ ##| ##2: #st2; #n2; #H; napply (asttype_destruct … H)
+ ##| ##3: #nl2; #H; napply (asttype_destruct … H)
##]
##| ##2: #st1; #n1; #H; #t2; ncases t2;
##[ ##2: #st2; #n2; #H1; nchange with (((eq_ast_type st1 st2)⊗(eq_nat n1 n2)) = true);
- nrewrite > (H st2 (asttype_destruct_array_array_1 ???? H1));
- nrewrite > (eq_to_eqnat n1 n2 (asttype_destruct_array_array_2 ???? H1));
+ nrewrite > (H st2 (asttype_destruct_array_array_1 … H1));
+ nrewrite > (eq_to_eqnat n1 n2 (asttype_destruct_array_array_2 … H1));
nnormalize;
- napply (refl_eq ??)
- ##| ##1: #b2; #H1; napply (asttype_destruct ??? H1)
- ##| ##3: #nl2; #H1; napply (asttype_destruct ??? H1)
+ napply refl_eq
+ ##| ##1: #b2; #H1; napply (asttype_destruct … H1)
+ ##| ##3: #nl2; #H1; napply (asttype_destruct … H1)
##]
##| ##3: #hh1; #H; #t2; ncases t2;
##[ ##3: #nl2; ncases nl2;
##[ ##1: #hh2; #H1; nchange with ((eq_ast_type hh1 hh2) = true);
- nrewrite > (H hh2 (nelist_destruct_nil_nil ? hh1 hh2 (asttype_destruct_struct_struct ?? H1)));
- napply (refl_eq ??)
- ##| ##2: #hh2; #ll2; #H1; nelim (nelist_destruct_nil_cons ? hh1 hh2 ll2 (asttype_destruct_struct_struct ?? H1))
+ nrewrite > (H hh2 (nelist_destruct_nil_nil ? hh1 hh2 (asttype_destruct_struct_struct … H1)));
+ napply refl_eq
+ ##| ##2: #hh2; #ll2; #H1; nelim (nelist_destruct_nil_cons ? hh1 hh2 ll2 (asttype_destruct_struct_struct … H1))
##]
- ##| ##1: #b2; #H1; napply (asttype_destruct ??? H1)
- ##| ##2: #st2; #n2; #H1; napply (asttype_destruct ??? H1)
+ ##| ##1: #b2; #H1; napply (asttype_destruct … H1)
+ ##| ##2: #st2; #n2; #H1; napply (asttype_destruct … H1)
##]
##| ##4: #hh1; #ll1; #H; #H1; #t2; ncases t2;
##[ ##3: #nl2; ncases nl2;
- ##[ ##1: #hh2; #H2; nelim (nelist_destruct_cons_nil ? hh1 hh2 ll1 (asttype_destruct_struct_struct ?? H2))
+ ##[ ##1: #hh2; #H2; nelim (nelist_destruct_cons_nil ? hh1 hh2 ll1 (asttype_destruct_struct_struct … H2))
##| ##2: #hh2; #ll2; #H2; nchange with (((eq_ast_type hh1 hh2)⊗(bfold_right_neList2 ? (λx,y.eq_ast_type x y) ll1 ll2)) = true);
- nrewrite > (H hh2 (nelist_destruct_cons_cons_1 ????? (asttype_destruct_struct_struct ?? H2)));
+ nrewrite > (H hh2 (nelist_destruct_cons_cons_1 … (asttype_destruct_struct_struct … H2)));
nrewrite > (eq_to_eqasttype_aux1 ll1 ll2 (H1 (AST_TYPE_STRUCT ll2) ?));
- ##[ ##1: nnormalize; napply (refl_eq ??)
- ##| ##2: nrewrite > (nelist_destruct_cons_cons_2 ????? (asttype_destruct_struct_struct ?? H2));
- napply (refl_eq ??)
+ ##[ ##1: nnormalize; napply refl_eq
+ ##| ##2: nrewrite > (nelist_destruct_cons_cons_2 … (asttype_destruct_struct_struct … H2));
+ napply refl_eq
##]
##]
- ##| ##1: #b2; #H2; napply (asttype_destruct ??? H2)
- ##| ##2: #st2; #n2; #H2; napply (asttype_destruct ??? H2)
+ ##| ##1: #b2; #H2; napply (asttype_destruct … H2)
+ ##| ##2: #st2; #n2; #H2; napply (asttype_destruct … H2)
##]
##]
nqed.
ncases ast;
nnormalize;
##[ ##1: #t; #H; napply I
- ##| ##2: #t; #n; #H; napply (bool_destruct ??? H)
- ##| ##3: #t; #H; napply (bool_destruct ??? H)
+ ##| ##2: #t; #n; #H; napply (bool_destruct … H)
+ ##| ##3: #t; #H; napply (bool_destruct … H)
##]
nqed.
#ast;
ncases ast;
nnormalize;
- ##[ ##1: #t; #H; napply (bool_destruct ??? H)
+ ##[ ##1: #t; #H; napply (bool_destruct … H)
##| ##2: #t; #n; #H; napply I
##| ##3: #l; #H; napply I
##]
nelim n1;
##[ ##1: nelim n2; nnormalize; #H;
##[ ##1: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match o0 with [ o0 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
##| ##2: nelim n2; nnormalize; #H;
##[ ##2: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match o1 with [ o1 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
##| ##3: nelim n2; nnormalize; #H;
##[ ##3: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match o2 with [ o2 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
##| ##4: nelim n2; nnormalize; #H;
##[ ##4: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match o3 with [ o3 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
##| ##5: nelim n2; nnormalize; #H;
##[ ##5: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match o4 with [ o4 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
##| ##6: nelim n2; nnormalize; #H;
##[ ##6: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match o5 with [ o5 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
##| ##7: nelim n2; nnormalize; #H;
##[ ##7: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match o6 with [ o6 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
##| ##8: nelim n2; nnormalize; #H;
##[ ##8: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match o7 with [ o7 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
nelim n1;
nelim n2;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma eqoct_to_eq : ∀n1,n2.eq_oct n1 n2 = true → n1 = n2.
ncases n1;
ncases n2;
nnormalize;
- ##[ ##1,10,19,28,37,46,55,64: #H; napply (refl_eq ??)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##1,10,19,28,37,46,55,64: #H; napply refl_eq
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
ncases n1;
ncases n2;
nnormalize;
- ##[ ##1,10,19,28,37,46,55,64: #H; napply (refl_eq ??)
- ##| ##*: #H; napply (oct_destruct ??? H)
+ ##[ ##1,10,19,28,37,46,55,64: #H; napply refl_eq
+ ##| ##*: #H; napply (oct_destruct … H)
##]
nqed.
ncases t2;
nnormalize; #H;
##[ ##1: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t00 with [ t00 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##2: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t01 with [ t01 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##3: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t02 with [ t02 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##4: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t03 with [ t03 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##5: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t04 with [ t04 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##6: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t05 with [ t05 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##7: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t06 with [ t06 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##8: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t07 with [ t07 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##9: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t08 with [ t08 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##10: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t09 with [ t09 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##11: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t0A with [ t0A ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##12: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t0B with [ t0B ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##13: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t0C with [ t0C ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##14: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t0D with [ t0D ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##15: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t0E with [ t0E ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##16: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t0F with [ t0F ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##17: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t10 with [ t10 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##18: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t11 with [ t11 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##19: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t12 with [ t12 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##20: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t13 with [ t13 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##21: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t14 with [ t14 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##22: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t15 with [ t15 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##23: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t16 with [ t16 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##24: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t17 with [ t17 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##25: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t18 with [ t18 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##26: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t19 with [ t19 ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##27: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t1A with [ t1A ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##28: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t1B with [ t1B ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##29: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t1C with [ t1C ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##30: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t1D with [ t1D ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##31: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t1E with [ t1E ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
ncases t2;
nnormalize; #H;
##[ ##32: napply (λx:P.x)
- ##| ##*: napply (False_ind (λ_.?) ?);
+ ##| ##*: napply (False_ind …);
nchange with (match t1F with [ t1F ⇒ False | _ ⇒ True ]);
nrewrite > H; nnormalize; napply I
##]
nlemma symmetric_eqbitrig : symmetricT bitrigesim bool eq_bitrig.
#t1;
nelim t1;
- ##[ ##1: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##2: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##3: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##4: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##5: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##6: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##7: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##8: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##9: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##10: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##11: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##12: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##13: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##14: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##15: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##16: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##17: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##18: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##19: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##20: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##21: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##22: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##23: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##24: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##25: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##26: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##27: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##28: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##29: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##30: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##31: #t2; nelim t2; nnormalize; napply (refl_eq ??)
- ##| ##32: #t2; nelim t2; nnormalize; napply (refl_eq ??)
+ ##[ ##1: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##2: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##3: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##4: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##5: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##6: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##7: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##8: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##9: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##10: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##11: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##12: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##13: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##14: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##15: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##16: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##17: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##18: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##19: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##20: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##21: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##22: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##23: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##24: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##25: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##26: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##27: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##28: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##29: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##30: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##31: #t2; nelim t2; nnormalize; napply refl_eq
+ ##| ##32: #t2; nelim t2; nnormalize; napply refl_eq
##]
nqed.
nlemma eqbitrig_to_eq1 : ∀t2.eq_bitrig t00 t2 = true → t00 = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##1: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##1: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq2 : ∀t2.eq_bitrig t01 t2 = true → t01 = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##2: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##2: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq3 : ∀t2.eq_bitrig t02 t2 = true → t02 = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##3: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##3: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq4 : ∀t2.eq_bitrig t03 t2 = true → t03 = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##4: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##4: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq5 : ∀t2.eq_bitrig t04 t2 = true → t04 = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##5: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##5: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq6 : ∀t2.eq_bitrig t05 t2 = true → t05 = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##6: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##6: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq7 : ∀t2.eq_bitrig t06 t2 = true → t06 = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##7: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##7: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq8 : ∀t2.eq_bitrig t07 t2 = true → t07 = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##8: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##8: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq9 : ∀t2.eq_bitrig t08 t2 = true → t08 = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##9: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##9: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq10 : ∀t2.eq_bitrig t09 t2 = true → t09 = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##10: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##10: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq11 : ∀t2.eq_bitrig t0A t2 = true → t0A = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##11: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##11: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq12 : ∀t2.eq_bitrig t0B t2 = true → t0B = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##12: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##12: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq13 : ∀t2.eq_bitrig t0C t2 = true → t0C = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##13: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##13: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq14 : ∀t2.eq_bitrig t0D t2 = true → t0D = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##14: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##14: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq15 : ∀t2.eq_bitrig t0E t2 = true → t0E = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##15: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##15: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq16 : ∀t2.eq_bitrig t0F t2 = true → t0F = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##16: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##16: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq17 : ∀t2.eq_bitrig t10 t2 = true → t10 = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##17: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##17: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq18 : ∀t2.eq_bitrig t11 t2 = true → t11 = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##18: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##18: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq19 : ∀t2.eq_bitrig t12 t2 = true → t12 = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##19: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##19: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq20 : ∀t2.eq_bitrig t13 t2 = true → t13 = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##20: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##20: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq21 : ∀t2.eq_bitrig t14 t2 = true → t14 = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##21: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##21: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq22 : ∀t2.eq_bitrig t15 t2 = true → t15 = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##22: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##22: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq23 : ∀t2.eq_bitrig t16 t2 = true → t16 = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##23: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##23: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq24 : ∀t2.eq_bitrig t17 t2 = true → t17 = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##24: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##24: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq25 : ∀t2.eq_bitrig t18 t2 = true → t18 = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##25: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##25: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq26 : ∀t2.eq_bitrig t19 t2 = true → t19 = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##26: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##26: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq27 : ∀t2.eq_bitrig t1A t2 = true → t1A = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##27: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##27: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq28 : ∀t2.eq_bitrig t1B t2 = true → t1B = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##28: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##28: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq29 : ∀t2.eq_bitrig t1C t2 = true → t1C = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##29: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##29: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq30 : ∀t2.eq_bitrig t1D t2 = true → t1D = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##30: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##30: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq31 : ∀t2.eq_bitrig t1E t2 = true → t1E = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##31: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##31: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq32 : ∀t2.eq_bitrig t1F t2 = true → t1F = t2.
- #t2; ncases t2; nnormalize; #H; ##[ ##32: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##32: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]
nqed.
nlemma eqbitrig_to_eq : ∀t1,t2.eq_bitrig t1 t2 = true → t1 = t2.
nqed.
nlemma eq_to_eqbitrig1 : ∀t2.t00 = t2 → eq_bitrig t00 t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##1: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##1: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig2 : ∀t2.t01 = t2 → eq_bitrig t01 t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##2: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##2: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig3 : ∀t2.t02 = t2 → eq_bitrig t02 t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##3: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##3: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig4 : ∀t2.t03 = t2 → eq_bitrig t03 t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##4: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##4: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig5 : ∀t2.t04 = t2 → eq_bitrig t04 t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##5: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##5: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig6 : ∀t2.t05 = t2 → eq_bitrig t05 t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##6: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##6: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig7 : ∀t2.t06 = t2 → eq_bitrig t06 t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##7: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##7: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig8 : ∀t2.t07 = t2 → eq_bitrig t07 t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##8: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##8: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig9 : ∀t2.t08 = t2 → eq_bitrig t08 t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##9: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##9: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig10 : ∀t2.t09 = t2 → eq_bitrig t09 t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##10: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##10: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig11 : ∀t2.t0A = t2 → eq_bitrig t0A t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##11: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##11: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig12 : ∀t2.t0B = t2 → eq_bitrig t0B t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##12: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##12: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig13 : ∀t2.t0C = t2 → eq_bitrig t0C t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##13: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##13: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig14 : ∀t2.t0D = t2 → eq_bitrig t0D t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##14: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##14: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig15 : ∀t2.t0E = t2 → eq_bitrig t0E t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##15: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##15: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig16 : ∀t2.t0F = t2 → eq_bitrig t0F t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##16: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##16: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig17 : ∀t2.t10 = t2 → eq_bitrig t10 t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##17: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##17: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig18 : ∀t2.t11 = t2 → eq_bitrig t11 t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##18: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##18: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig19 : ∀t2.t12 = t2 → eq_bitrig t12 t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##19: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##19: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig20 : ∀t2.t13 = t2 → eq_bitrig t13 t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##20: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##20: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig21 : ∀t2.t14 = t2 → eq_bitrig t14 t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##21: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##21: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig22 : ∀t2.t15 = t2 → eq_bitrig t15 t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##22: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##22: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig23 : ∀t2.t16 = t2 → eq_bitrig t16 t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##23: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##23: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig24 : ∀t2.t17 = t2 → eq_bitrig t17 t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##24: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##24: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig25 : ∀t2.t18 = t2 → eq_bitrig t18 t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##25: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##25: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig26 : ∀t2.t19 = t2 → eq_bitrig t19 t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##26: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##26: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig27 : ∀t2.t1A = t2 → eq_bitrig t1A t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##27: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##27: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig28 : ∀t2.t1B = t2 → eq_bitrig t1B t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##28: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##28: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig29 : ∀t2.t1C = t2 → eq_bitrig t1C t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##29: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##29: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig30 : ∀t2.t1D = t2 → eq_bitrig t1D t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##30: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##30: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig31 : ∀t2.t1E = t2 → eq_bitrig t1E t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##31: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##31: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig32 : ∀t2.t1F = t2 → eq_bitrig t1F t2 = true.
- #t2; ncases t2; nnormalize; #H; ##[ ##32: napply (refl_eq ??) ##| ##*: napply (bitrigesim_destruct ??? H) ##]
+ #t2; ncases t2; nnormalize; #H; ##[ ##32: napply refl_eq ##| ##*: napply (bitrigesim_destruct … H) ##]
nqed.
nlemma eq_to_eqbitrig : ∀t1,t2.t1 = t2 → eq_bitrig t1 t2 = true.
nelim b1;
nelim b2;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_andbool : symmetricT bool bool and_bool.
nelim b1;
nelim b2;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma associative_andbool : ∀b1,b2,b3.((b1 ⊗ b2) ⊗ b3) = (b1 ⊗ (b2 ⊗ b3)).
nelim b2;
nelim b3;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_orbool : symmetricT bool bool or_bool.
nelim b1;
nelim b2;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma associative_orbool : ∀b1,b2,b3.((b1 ⊕ b2) ⊕ b3) = (b1 ⊕ (b2 ⊕ b3)).
nelim b2;
nelim b3;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_xorbool : symmetricT bool bool xor_bool.
nelim b1;
nelim b2;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma associative_xorbool : ∀b1,b2,b3.((b1 ⊙ b2) ⊙ b3) = (b1 ⊙ (b2 ⊙ b3)).
nelim b2;
nelim b3;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma eqbool_to_eq : ∀b1,b2:bool.(eq_bool b1 b2 = true) → (b1 = b2).
ncases b1;
ncases b2;
nnormalize;
- ##[ ##1,4: #H; napply (refl_eq ??)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##1,4: #H; napply refl_eq
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
ncases b1;
ncases b2;
nnormalize;
- ##[ ##1,4: #H; napply (refl_eq ??)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##1,4: #H; napply refl_eq
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
ncases b1;
ncases b2;
nnormalize;
- ##[ ##1,2: #H; napply (refl_eq ??)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##1,2: #H; napply refl_eq
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
ncases b1;
ncases b2;
nnormalize;
- ##[ ##1,3: #H; napply (refl_eq ??)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##1,3: #H; napply refl_eq
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
ncases b1;
ncases b2;
nnormalize;
- ##[ ##4: #H; napply (refl_eq ??)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##4: #H; napply refl_eq
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
ncases b1;
ncases b2;
nnormalize;
- ##[ ##4: #H; napply (refl_eq ??)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##4: #H; napply refl_eq
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
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'' ]].
+ [ pair bl' c'' ⇒ pair … (mk_byte8 bh' bl') c'' ]].
(* operatore shift destro *)
ndefinition shr_b8 ≝
λb:byte8.match rcr_ex (b8h b) false with
[ pair bh' c' ⇒ match rcr_ex (b8l b) c' with
- [ pair bl' c'' ⇒ pair ?? (mk_byte8 bh' bl') c'' ]].
+ [ pair bl' c'' ⇒ pair … (mk_byte8 bh' bl') c'' ]].
(* operatore rotazione destra *)
ndefinition ror_b8 ≝
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'' ]].
+ [ pair bh' c'' ⇒ pair … (mk_byte8 bh' bl') c'' ]].
(* operatore shift sinistro *)
ndefinition shl_b8 ≝
λb:byte8.match rcl_ex (b8l b) false with
[ pair bl' c' ⇒ match rcl_ex (b8h b) c' with
- [ pair bh' c'' ⇒ pair ?? (mk_byte8 bh' bl') c'' ]].
+ [ pair bh' c'' ⇒ pair … (mk_byte8 bh' bl') c'' ]].
(* operatore rotazione sinistra *)
ndefinition rol_b8 ≝
λ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' ]].
+ [ pair h c' ⇒ pair … 〈h,l〉 c' ]].
(* operatore somma con data+carry → data *)
ndefinition plus_b8_dc_d ≝
λ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' ]].
+ [ pair h c' ⇒ pair … 〈h,l〉 c' ]].
(* operatore somma con data → data *)
ndefinition plus_b8_d_d ≝
let X'' ≝ match c with
[ true ⇒ plus_b8_d_d X' 〈x6,x0〉
| false ⇒ X' ] in
- pair ?? X'' c
+ pair … X'' c
(* [X:0x9A-0xFF] *)
(* c' = 1 *)
(* X' = [X:0x9A-0xFF]
[ true ⇒ X
| false ⇒ plus_b8_d_d X 〈x0,x6〉 ] in
let X'' ≝ plus_b8_d_d X' 〈x6,x0〉 in
- pair ?? X'' true
+ pair … X'' true
].
(* iteratore sui byte *)
nchange with (match mk_byte8 x2 y2 with [ mk_byte8 a _ ⇒ x1 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma byte8_destruct_2 :
nchange with (match mk_byte8 x2 y2 with [ mk_byte8 _ b ⇒ y1 = b ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_eqb8 : symmetricT byte8 bool eq_b8.
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 ??).
+ napply refl_eq.
nqed.
nlemma symmetric_andb8 : symmetricT byte8 byte8 and_b8.
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 ??).
+ napply refl_eq.
nqed.
nlemma associative_andb8 : ∀b1,b2,b3.(and_b8 (and_b8 b1 b2) b3) = (and_b8 b1 (and_b8 b2 b3)).
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 ??).
+ napply refl_eq.
nqed.
nlemma symmetric_orb8 : symmetricT byte8 byte8 or_b8.
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 ??).
+ napply refl_eq.
nqed.
nlemma associative_orb8 : ∀b1,b2,b3.(or_b8 (or_b8 b1 b2) b3) = (or_b8 b1 (or_b8 b2 b3)).
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 ??).
+ napply refl_eq.
nqed.
nlemma symmetric_xorb8 : symmetricT byte8 byte8 xor_b8.
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 ??).
+ napply refl_eq.
nqed.
nlemma associative_xorb8 : ∀b1,b2,b3.(xor_b8 (xor_b8 b1 b2) b3) = (xor_b8 b1 (xor_b8 b2 b3)).
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 ??).
+ 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.
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' ]] =
+ [ 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' ]]);
+ [ 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' ] =
+ [ pair h c' ⇒ pair … 〈h,e5〉 c' ] =
match plus_ex_dc_dc e3 e1 c1 with
- [ pair h c' ⇒ pair ?? 〈h,e5〉 c' ]);
+ [ pair h c' ⇒ pair … 〈h,e5〉 c' ]);
nrewrite > (symmetric_plusex_dc_dc e1 e3 c1);
- napply (refl_eq ??).
+ 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.
#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 ??).
+ 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.
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 ??).
+ napply refl_eq.
nqed.
nlemma symmetric_plusb8_d_dc : ∀b1,b2.plus_b8_d_dc b1 b2 = plus_b8_d_dc b2 b1.
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' ]] =
+ [ 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' ]]);
+ [ 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' ] =
+ [ pair h c' ⇒ pair … 〈h,e5〉 c' ] =
match plus_ex_dc_dc e3 e1 c with
- [ pair h c' ⇒ pair ?? 〈h,e5〉 c' ]);
+ [ pair h c' ⇒ pair … 〈h,e5〉 c' ]);
nrewrite > (symmetric_plusex_dc_dc e1 e3 c);
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_plusb8_d_d : ∀b1,b2.plus_b8_d_d b1 b2 = plus_b8_d_d b2 b1.
#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 ??).
+ napply refl_eq.
nqed.
nlemma symmetric_plusb8_d_c : ∀b1,b2.plus_b8_d_c b1 b2 = plus_b8_d_c b2 b1.
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 ??).
+ napply refl_eq.
nqed.
nlemma symmetric_mulex : symmetricT exadecim byte8 mul_ex.
nelim e1;
nelim e2;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma eqb8_to_eq : ∀b1,b2:byte8.(eq_b8 b1 b2 = true) → (b1 = 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 ??).
+ 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.
nelim b2;
#e1; #e2; #e3; #e4;
#H;
- nrewrite < (byte8_destruct_1 ???? H);
- nrewrite < (byte8_destruct_2 ???? 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 ??));
+ nrewrite > (eq_to_eqex e3 e3 (refl_eq …));
+ nrewrite > (eq_to_eqex e4 e4 (refl_eq …));
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
]].
(* 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
| 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 … 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 ]
].
(* operatore somma con data+carry → data *)
nelim e1;
nelim e2;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_andex : symmetricT exadecim exadecim and_ex.
nelim e1;
nelim e2;
nnormalize;
- napply (refl_eq ??).
+ 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; #e2; #e3;
nelim e1;
- ##[ ##1: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##2: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##3: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##4: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##5: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##6: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##7: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##8: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##9: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##10: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##11: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##12: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##13: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##14: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##15: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##16: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
+ ##[ ##1: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##2: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##3: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##4: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##5: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##6: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##7: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##8: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##9: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##10: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##11: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##12: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##13: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##14: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##15: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##16: nelim e2; nelim e3; nnormalize; napply refl_eq
##]
nqed.
nelim e1;
nelim e2;
nnormalize;
- napply (refl_eq ??).
+ 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; #e2; #e3;
nelim e1;
- ##[ ##1: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##2: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##3: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##4: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##5: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##6: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##7: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##8: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##9: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##10: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##11: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##12: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##13: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##14: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##15: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##16: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
+ ##[ ##1: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##2: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##3: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##4: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##5: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##6: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##7: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##8: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##9: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##10: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##11: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##12: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##13: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##14: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##15: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##16: nelim e2; nelim e3; nnormalize; napply refl_eq
##]
nqed.
nelim e1;
nelim e2;
nnormalize;
- napply (refl_eq ??).
+ 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; #e2; #e3;
nelim e1;
- ##[ ##1: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##2: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##3: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##4: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##5: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##6: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##7: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##8: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##9: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##10: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##11: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##12: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##13: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##14: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##15: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##16: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
+ ##[ ##1: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##2: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##3: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##4: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##5: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##6: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##7: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##8: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##9: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##10: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##11: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##12: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##13: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##14: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##15: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##16: nelim e2; nelim e3; nnormalize; napply refl_eq
##]
nqed.
nelim e2;
nelim c;
nnormalize;
- napply (refl_eq ??).
+ 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.
+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 ??).
+ 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.
+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 ??).
+ 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.
nelim e1;
nelim e2;
nnormalize;
- napply (refl_eq ??).
+ 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.
+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 ??).
+ 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.
+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 ??).
+ 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.
nelim e2;
nelim c;
nnormalize;
- napply (refl_eq ??).
+ 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.
nelim e2;
nelim c;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_plusex_d_dc : ∀e1,e2.plus_ex_d_dc e1 e2 = plus_ex_d_dc e2 e1.
nelim e1;
nelim e2;
nnormalize;
- napply (refl_eq ??).
+ 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.
+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 ??).
+ 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.
+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 ??).
+ napply refl_eq.
nqed.
nlemma symmetric_plusex_d_d : ∀e1,e2.plus_ex_d_d e1 e2 = plus_ex_d_d e2 e1.
nelim e1;
nelim e2;
nnormalize;
- napply (refl_eq ??).
+ 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; #e2; #e3;
nelim e1;
- ##[ ##1: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##2: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##3: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##4: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##5: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##6: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##7: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##8: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##9: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##10: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##11: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##12: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##13: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##14: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##15: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
- ##| ##16: nelim e2; nelim e3; nnormalize; napply (refl_eq ??)
+ ##[ ##1: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##2: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##3: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##4: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##5: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##6: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##7: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##8: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##9: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##10: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##11: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##12: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##13: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##14: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##15: nelim e2; nelim e3; nnormalize; napply refl_eq
+ ##| ##16: nelim e2; nelim e3; nnormalize; napply refl_eq
##]
nqed.
nelim e1;
nelim e2;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma eqex_to_eq : ∀e1,e2:exadecim.(eq_ex e1 e2 = true) → (e1 = e2).
ncases e1;
ncases e2;
nnormalize;
- ##[ ##1,18,35,52,69,86,103,120,137,154,171,188,205,222,239,256: #H; napply (refl_eq ??)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##1,18,35,52,69,86,103,120,137,154,171,188,205,222,239,256: #H; napply refl_eq
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
ncases m1;
ncases m2;
nnormalize;
- ##[ ##1,18,35,52,69,86,103,120,137,154,171,188,205,222,239,256: #H; napply (refl_eq ??)
- ##| ##*: #H; napply (exadecim_destruct ??? H)
+ ##[ ##1,18,35,52,69,86,103,120,137,154,171,188,205,222,239,256: #H; napply refl_eq
+ ##| ##*: #H; napply (exadecim_destruct … H)
##]
nqed.
accesso al paging: [00pp pppp ppxx xxxx] con p=PS x=addr
*)
match in_range addr 〈〈x0,x0〉:〈xC,x0〉〉 〈〈x0,x0〉:〈xF,xF〉〉 with
- [ true ⇒ fMEM mem chk (or_w16 (fst ?? (shr_w16 (fst ?? (shr_w16 〈psm:〈x0,x0〉〉))))
+ [ true ⇒ fMEM mem chk (or_w16 (fst … (shr_w16 (fst … (shr_w16 〈psm:〈x0,x0〉〉))))
(and_w16 addr 〈〈x0,x0〉:〈x3,xF〉〉))
(*
accesso normale
altrimenti sarebbero 2 indirezioni
*)
match eq_w16 addr 〈〈x0,x0〉:〈x0,xE〉〉 with
- [ true ⇒ opt_map ?? (fMEM mem chk 〈〈x0,x0〉:xm〉)
+ [ true ⇒ opt_map … (fMEM mem chk 〈〈x0,x0〉:xm〉)
(λmem'.Some ? (mk_any_status RS08 t (mk_alu_RS08 acclow pc pcm spc xm psm zfl cfl) mem' chk clk))
| false ⇒
accesso al paging: [00pp pppp ppxx xxxx] con p=PS x=addr
*)
match in_range addr 〈〈x0,x0〉:〈xC,x0〉〉 〈〈x0,x0〉:〈xF,xF〉〉 with
- [ true ⇒ opt_map ?? (fMEM mem chk (or_w16 (fst ?? (shr_w16 (fst ?? (shr_w16 〈psm:〈x0,x0〉〉))))
+ [ true ⇒ opt_map … (fMEM mem chk (or_w16 (fst … (shr_w16 (fst … (shr_w16 〈psm:〈x0,x0〉〉))))
(and_w16 addr 〈〈x0,x0〉:〈x3,xF〉〉)))
(λmem'.Some ? (mk_any_status RS08 t (mk_alu_RS08 acclow pc pcm spc xm psm zfl cfl) mem' chk clk))
(*
accesso normale
*)
- | false ⇒ opt_map ?? (fMEM mem chk addr)
+ | false ⇒ opt_map … (fMEM mem chk addr)
(λmem'.Some ? (mk_any_status RS08 t (mk_alu_RS08 acclow pc pcm spc xm psm zfl cfl) mem' chk clk)) ]]]]]].
(* scrittura RS08 di un byte *)
λm:mcu_type.λt:memory_impl.match m
return λm:mcu_type.any_status m t → word16 → byte8 → option (any_status m t) with
[ HC05 ⇒ λs:any_status HC05 t.λaddr:word16.λval:byte8.
- opt_map ?? (mem_update t (mem_desc ? t s) (chk_desc ? t s) addr val)
+ opt_map … (mem_update t (mem_desc ? t s) (chk_desc ? t s) addr val)
(λmem.Some ? (set_mem_desc ? t s mem))
| HC08 ⇒ λs:any_status HC08 t.λaddr:word16.λval:byte8.
- opt_map ?? (mem_update t (mem_desc ? t s) (chk_desc ? t s) addr val)
+ opt_map … (mem_update t (mem_desc ? t s) (chk_desc ? t s) addr val)
(λmem.Some ? (set_mem_desc ? t s mem))
| HCS08 ⇒ λs:any_status HCS08 t.λaddr:word16.λval:byte8.
- opt_map ?? (mem_update t (mem_desc ? t s) (chk_desc ? t s) addr val)
+ opt_map … (mem_update t (mem_desc ? t s) (chk_desc ? t s) addr val)
(λmem.Some ? (set_mem_desc ? t s mem))
| RS08 ⇒ λs:any_status RS08 t.λaddr:word16.λval:byte8.
RS08_memory_filter_write t s addr val
λm:mcu_type.λt:memory_impl.match m
return λm:mcu_type.any_status m t → word16 → oct → bool → option (any_status m t) with
[ HC05 ⇒ λs:any_status HC05 t.λaddr:word16.λsub:oct.λval:bool.
- opt_map ?? (mem_update_bit t (mem_desc ? t s) (chk_desc ? t s) addr sub val)
+ opt_map … (mem_update_bit t (mem_desc ? t s) (chk_desc ? t s) addr sub val)
(λmem.Some ? (set_mem_desc ? t s mem))
| HC08 ⇒ λs:any_status HC08 t.λaddr:word16.λsub:oct.λval:bool.
- opt_map ?? (mem_update_bit t (mem_desc ? t s) (chk_desc ? t s) addr sub val)
+ opt_map … (mem_update_bit t (mem_desc ? t s) (chk_desc ? t s) addr sub val)
(λmem.Some ? (set_mem_desc ? t s mem))
| HCS08 ⇒ λs:any_status HCS08 t.λaddr:word16.λsub:oct.λval:bool.
- opt_map ?? (mem_update_bit t (mem_desc ? t s) (chk_desc ? t s) addr sub val)
+ opt_map … (mem_update_bit t (mem_desc ? t s) (chk_desc ? t s) addr sub val)
(λmem.Some ? (set_mem_desc ? t s mem))
| RS08 ⇒ λs:any_status RS08 t.λaddr:word16.λsub:oct.λval:bool.
RS08_memory_filter_write_bit t s addr sub val
(* lettura byte da addr *)
ndefinition loadb_from ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λaddr:word16.λcur_pc:word16.λfetched:nat.
- opt_map ?? (memory_filter_read m t s addr)
- (λb.Some ? (triple ??? s b (filtered_plus_w16 m t s cur_pc fetched))).
+ opt_map … (memory_filter_read m t s addr)
+ (λb.Some ? (triple … s b (filtered_plus_w16 m t s cur_pc fetched))).
(* lettura bit da addr *)
ndefinition loadbit_from ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λaddr:word16.λsub:oct.λcur_pc:word16.λfetched:nat.
- opt_map ?? (memory_filter_read_bit m t s addr sub)
- (λb.Some ? (triple ??? s b (filtered_plus_w16 m t s cur_pc fetched))).
+ opt_map … (memory_filter_read_bit m t s addr sub)
+ (λb.Some ? (triple … s b (filtered_plus_w16 m t s cur_pc fetched))).
(* lettura word da addr *)
ndefinition loadw_from ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λaddr:word16.λcur_pc:word16.λfetched:nat.
- opt_map ?? (memory_filter_read m t s addr)
- (λbh.opt_map ?? (memory_filter_read m t s (succ_w16 addr))
- (λbl.Some ? (triple ??? s (mk_word16 bh bl) (filtered_plus_w16 m t s cur_pc fetched)))).
+ opt_map … (memory_filter_read m t s addr)
+ (λbh.opt_map … (memory_filter_read m t s (succ_w16 addr))
+ (λbl.Some ? (triple … s (mk_word16 bh bl) (filtered_plus_w16 m t s cur_pc fetched)))).
(* scrittura byte su addr *)
ndefinition writeb_to ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λaddr:word16.λcur_pc:word16.λfetched:nat.λwriteb:byte8.
- opt_map ?? (memory_filter_write m t s addr writeb)
- (λtmps.Some ? (pair ?? tmps (filtered_plus_w16 m t s cur_pc fetched))).
+ opt_map … (memory_filter_write m t s addr writeb)
+ (λtmps.Some ? (pair … tmps (filtered_plus_w16 m t s cur_pc fetched))).
(* scrittura bit su addr *)
ndefinition writebit_to ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λaddr:word16.λsub:oct.λcur_pc:word16.λfetched:nat.λwriteb:bool.
- opt_map ?? (memory_filter_write_bit m t s addr sub writeb)
- (λtmps.Some ? (pair ?? tmps (filtered_plus_w16 m t s cur_pc fetched))).
+ opt_map … (memory_filter_write_bit m t s addr sub writeb)
+ (λtmps.Some ? (pair … tmps (filtered_plus_w16 m t s cur_pc fetched))).
(* scrittura word su addr *)
ndefinition writew_to ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λaddr:word16.λcur_pc:word16.λfetched:nat.λwritew:word16.
- opt_map ?? (memory_filter_write m t s addr (w16h writew))
- (λtmps1.opt_map ?? (memory_filter_write m t tmps1 (succ_w16 addr) (w16l writew))
- (λtmps2.Some ? (pair ?? tmps2 (filtered_plus_w16 m t tmps2 cur_pc fetched)))).
+ opt_map … (memory_filter_write m t s addr (w16h writew))
+ (λtmps1.opt_map … (memory_filter_write m t tmps1 (succ_w16 addr) (w16l writew))
+ (λtmps2.Some ? (pair … tmps2 (filtered_plus_w16 m t tmps2 cur_pc fetched)))).
(* ausiliari per definire i tipi e la lettura/scrittura *)
(* lettura da [curpc]: IMM1 comportamento asimmetrico, quindi non si appoggia a loadb *)
ndefinition mode_IMM1_load ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
- opt_map ?? (memory_filter_read m t s cur_pc)
- (λb.Some ? (triple ??? s b (filtered_inc_w16 m t s cur_pc))).
+ opt_map … (memory_filter_read m t s cur_pc)
+ (λb.Some ? (triple … s b (filtered_inc_w16 m t s cur_pc))).
(* lettura da [curpc]: IMM1 + estensione a word *)
ndefinition mode_IMM1EXT_load ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
- opt_map ?? (memory_filter_read m t s cur_pc)
- (λb.Some ? (triple ??? s 〈〈x0,x0〉:b〉 (filtered_inc_w16 m t s cur_pc))).
+ opt_map … (memory_filter_read m t s cur_pc)
+ (λb.Some ? (triple … s 〈〈x0,x0〉:b〉 (filtered_inc_w16 m t s cur_pc))).
(* lettura da [curpc]: IMM2 comportamento asimmetrico, quindi non si appoggia a loadw *)
ndefinition mode_IMM2_load ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
- opt_map ?? (memory_filter_read m t s cur_pc)
- (λbh.opt_map ?? (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
- (λbl.Some ? (triple ??? s (mk_word16 bh bl) (filtered_plus_w16 m t s cur_pc 2)))).
+ opt_map … (memory_filter_read m t s cur_pc)
+ (λbh.opt_map … (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
+ (λbl.Some ? (triple … s (mk_word16 bh bl) (filtered_plus_w16 m t s cur_pc 2)))).
(* lettura da [byte [curpc]]: true=DIR1 loadb, false=DIR1 loadw *)
ndefinition mode_DIR1_load ≝
λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
- opt_map ?? (memory_filter_read m t s cur_pc)
+ opt_map … (memory_filter_read m t s cur_pc)
(λaddr.(aux_load m t byteflag) s 〈〈x0,x0〉:addr〉 cur_pc 1).
(* lettura da [byte [curpc]]: loadbit *)
ndefinition mode_DIR1n_load ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.λsub:oct.
- opt_map ?? (memory_filter_read m t s cur_pc)
+ opt_map … (memory_filter_read m t s cur_pc)
(λaddr.loadbit_from m t s 〈〈x0,x0〉:addr〉 sub cur_pc 1).
(* scrittura su [byte [curpc]]: true=DIR1 writeb, false=DIR1 writew *)
ndefinition mode_DIR1_write ≝
λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
λwritebw:match byteflag with [ true ⇒ byte8 | false ⇒ word16 ].
- opt_map ?? (memory_filter_read m t s cur_pc)
+ opt_map … (memory_filter_read m t s cur_pc)
(λaddr.(aux_write m t byteflag) s 〈〈x0,x0〉:addr〉 cur_pc 1 writebw).
(* scrittura su [byte [curpc]]: writebit *)
ndefinition mode_DIR1n_write ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.λsub:oct.λwriteb:bool.
- opt_map ?? (memory_filter_read m t s cur_pc)
+ opt_map … (memory_filter_read m t s cur_pc)
(λaddr.writebit_to m t s 〈〈x0,x0〉:addr〉 sub cur_pc 1 writeb).
(* lettura da [word [curpc]]: true=DIR2 loadb, false=DIR2 loadw *)
ndefinition mode_DIR2_load ≝
λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
- opt_map ?? (memory_filter_read m t s cur_pc)
- (λaddrh.opt_map ?? (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
+ opt_map … (memory_filter_read m t s cur_pc)
+ (λaddrh.opt_map … (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
(λaddrl.(aux_load m t byteflag) s 〈addrh:addrl〉 cur_pc 2)).
(* scrittura su [word [curpc]]: true=DIR2 writeb, false=DIR2 writew *)
ndefinition mode_DIR2_write ≝
λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
λwritebw:match byteflag with [ true ⇒ byte8 | false ⇒ word16 ].
- opt_map ?? (memory_filter_read m t s cur_pc)
- (λaddrh.opt_map ?? (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
+ opt_map … (memory_filter_read m t s cur_pc)
+ (λaddrh.opt_map … (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
(λaddrl.(aux_write m t byteflag) s 〈addrh:addrl〉 cur_pc 2 writebw)).
ndefinition get_IX ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.
match m with
- [ HC05 ⇒ opt_map ?? (get_indX_8_low_reg m t s) (λindx.Some ? 〈〈x0,x0〉:indx〉)
- | HC08 ⇒ opt_map ?? (get_indX_16_reg m t s) (λindx.Some ? indx)
- | HCS08 ⇒ opt_map ?? (get_indX_16_reg m t s) (λindx.Some ? indx)
+ [ HC05 ⇒ opt_map … (get_indX_8_low_reg m t s) (λindx.Some ? 〈〈x0,x0〉:indx〉)
+ | HC08 ⇒ opt_map … (get_indX_16_reg m t s) (λindx.Some ? indx)
+ | HCS08 ⇒ opt_map … (get_indX_16_reg m t s) (λindx.Some ? indx)
| RS08 ⇒ None ? ].
(* lettura da [IX]: true=IX0 loadb, false=IX0 loadw *)
ndefinition mode_IX0_load ≝
λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
- opt_map ?? (get_IX m t s)
+ opt_map … (get_IX m t s)
(λaddr.(aux_load m t byteflag) s addr cur_pc 0).
(* scrittura su [IX]: true=IX0 writeb, false=IX0 writew *)
ndefinition mode_IX0_write ≝
λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
λwritebw:match byteflag with [ true ⇒ byte8 | false ⇒ word16 ].
- opt_map ?? (get_IX m t s)
+ opt_map … (get_IX m t s)
(λaddr.(aux_write m t byteflag) s addr cur_pc 0 writebw).
(* lettura da [IX+byte [pc]]: true=IX1 loadb, false=IX1 loadw *)
ndefinition mode_IX1_load ≝
λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
- opt_map ?? (get_IX m t s)
- (λaddr.opt_map ?? (memory_filter_read m t s cur_pc)
+ opt_map … (get_IX m t s)
+ (λaddr.opt_map … (memory_filter_read m t s cur_pc)
(λoffs.(aux_load m t byteflag) s (plus_w16_d_d addr 〈〈x0,x0〉:offs〉) cur_pc 1)).
(* lettura da X+[byte curpc] *)
ndefinition mode_IX1ADD_load ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
- opt_map ?? (memory_filter_read m t s cur_pc)
- (λb.opt_map ?? (get_IX m t s)
- (λaddr.Some ? (triple ??? s (plus_w16_d_d addr 〈〈x0,x0〉:b〉) (filtered_inc_w16 m t s cur_pc)))).
+ opt_map … (memory_filter_read m t s cur_pc)
+ (λb.opt_map … (get_IX m t s)
+ (λaddr.Some ? (triple … s (plus_w16_d_d addr 〈〈x0,x0〉:b〉) (filtered_inc_w16 m t s cur_pc)))).
(* scrittura su [IX+byte [pc]]: true=IX1 writeb, false=IX1 writew *)
ndefinition mode_IX1_write ≝
λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
λwritebw:match byteflag with [ true ⇒ byte8 | false ⇒ word16 ].
- opt_map ?? (get_IX m t s)
- (λaddr.opt_map ?? (memory_filter_read m t s cur_pc)
+ opt_map … (get_IX m t s)
+ (λaddr.opt_map … (memory_filter_read m t s cur_pc)
(λoffs.(aux_write m t byteflag) s (plus_w16_d_d addr 〈〈x0,x0〉:offs〉) cur_pc 1 writebw)).
(* lettura da [IX+word [pc]]: true=IX2 loadb, false=IX2 loadw *)
ndefinition mode_IX2_load ≝
λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
- opt_map ?? (get_IX m t s)
- (λaddr.opt_map ?? (memory_filter_read m t s cur_pc)
- (λoffsh.opt_map ?? (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
+ opt_map … (get_IX m t s)
+ (λaddr.opt_map … (memory_filter_read m t s cur_pc)
+ (λoffsh.opt_map … (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
(λoffsl.(aux_load m t byteflag) s (plus_w16_d_d addr 〈offsh:offsl〉) cur_pc 2))).
(* lettura da X+[word curpc] *)
ndefinition mode_IX2ADD_load ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
- opt_map ?? (memory_filter_read m t s cur_pc)
- (λbh.opt_map ?? (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
- (λbl.opt_map ?? (get_IX m t s)
- (λaddr.Some ? (triple ??? s (plus_w16_d_d addr 〈bh:bl〉) (filtered_plus_w16 m t s cur_pc 2))))).
+ opt_map … (memory_filter_read m t s cur_pc)
+ (λbh.opt_map … (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
+ (λbl.opt_map … (get_IX m t s)
+ (λaddr.Some ? (triple … s (plus_w16_d_d addr 〈bh:bl〉) (filtered_plus_w16 m t s cur_pc 2))))).
(* scrittura su [IX+word [pc]]: true=IX2 writeb, false=IX2 writew *)
ndefinition mode_IX2_write ≝
λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
λwritebw:match byteflag with [ true ⇒ byte8 | false ⇒ word16 ].
- opt_map ?? (get_IX m t s)
- (λaddr.opt_map ?? (memory_filter_read m t s cur_pc)
- (λoffsh.opt_map ?? (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
+ opt_map … (get_IX m t s)
+ (λaddr.opt_map … (memory_filter_read m t s cur_pc)
+ (λoffsh.opt_map … (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
(λoffsl.(aux_write m t byteflag) s (plus_w16_d_d addr 〈offsh:offsl〉) cur_pc 2 writebw))).
(* lettura da [SP+byte [pc]]: true=SP1 loadb, false=SP1 loadw *)
ndefinition mode_SP1_load ≝
λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
- opt_map ?? (get_sp_reg m t s)
- (λaddr.opt_map ?? (memory_filter_read m t s cur_pc)
+ opt_map … (get_sp_reg m t s)
+ (λaddr.opt_map … (memory_filter_read m t s cur_pc)
(λoffs.(aux_load m t byteflag) s (plus_w16_d_d addr 〈〈x0,x0〉:offs〉) cur_pc 1)).
(* scrittura su [SP+byte [pc]]: true=SP1 writeb, false=SP1 writew *)
ndefinition mode_SP1_write ≝
λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
λwritebw:match byteflag with [ true ⇒ byte8 | false ⇒ word16 ].
- opt_map ?? (get_sp_reg m t s)
- (λaddr.opt_map ?? (memory_filter_read m t s cur_pc)
+ opt_map … (get_sp_reg m t s)
+ (λaddr.opt_map … (memory_filter_read m t s cur_pc)
(λoffs.(aux_write m t byteflag) s (plus_w16_d_d addr 〈〈x0,x0〉:offs〉) cur_pc 1 writebw)).
(* lettura da [SP+word [pc]]: true=SP2 loadb, false=SP2 loadw *)
ndefinition mode_SP2_load ≝
λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
- opt_map ?? (get_sp_reg m t s)
- (λaddr.opt_map ?? (memory_filter_read m t s cur_pc)
- (λoffsh.opt_map ?? (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
+ opt_map … (get_sp_reg m t s)
+ (λaddr.opt_map … (memory_filter_read m t s cur_pc)
+ (λoffsh.opt_map … (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
(λoffsl.(aux_load m t byteflag) s (plus_w16_d_d addr 〈offsh:offsl〉) cur_pc 2))).
(* scrittura su [SP+word [pc]]: true=SP2 writeb, false=SP2 writew *)
ndefinition mode_SP2_write ≝
λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
λwritebw:match byteflag with [ true ⇒ byte8 | false ⇒ word16 ].
- opt_map ?? (get_sp_reg m t s)
- (λaddr.opt_map ?? (memory_filter_read m t s cur_pc)
- (λoffsh.opt_map ?? (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
+ opt_map … (get_sp_reg m t s)
+ (λaddr.opt_map … (memory_filter_read m t s cur_pc)
+ (λoffsh.opt_map … (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
(λoffsl.(aux_write m t byteflag) s (plus_w16_d_d addr 〈offsh:offsl〉) cur_pc 2 writebw))).
(* ************************************** *)
(* H:X++ *)
ndefinition aux_inc_indX_16 ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.
- opt_map ?? (get_indX_16_reg m t s)
- (λX_op.opt_map ?? (set_indX_16_reg m t s (succ_w16 X_op))
+ opt_map … (get_indX_16_reg m t s)
+ (λX_op.opt_map … (set_indX_16_reg m t s (succ_w16 X_op))
(λs_tmp.Some ? s_tmp)).
(* tutte le modalita' di lettura: false=loadb true=loadw *)
(* NO: non ci sono indicazioni *)
[ MODE_INH ⇒ None ?
(* restituisce A *)
- | MODE_INHA ⇒ Some ? (triple ??? s (get_acc_8_low_reg m t s) cur_pc)
+ | MODE_INHA ⇒ Some ? (triple … s (get_acc_8_low_reg m t s) cur_pc)
(* restituisce X *)
- | MODE_INHX ⇒ opt_map ?? (get_indX_8_low_reg m t s)
- (λindx.Some ? (triple ??? s indx cur_pc))
+ | MODE_INHX ⇒ opt_map … (get_indX_8_low_reg m t s)
+ (λindx.Some ? (triple … s indx cur_pc))
(* restituisce H *)
- | MODE_INHH ⇒ opt_map ?? (get_indX_8_high_reg m t s)
- (λindx.Some ? (triple ??? s indx cur_pc))
+ | MODE_INHH ⇒ opt_map … (get_indX_8_high_reg m t s)
+ (λindx.Some ? (triple … s indx cur_pc))
(* NO: solo lettura word *)
| MODE_INHX0ADD ⇒ None ?
(* NO: solo lettura word *)
| MODE_DIRn_and_IMM1 _ ⇒ None ?
(* preleva 1 byte da 0000 0000 0000 xxxxb *)
- | MODE_TNY e ⇒ opt_map ?? (memory_filter_read m t s 〈〈x0,x0〉:〈x0,e〉〉)
- (λb.Some ? (triple ??? s b cur_pc))
+ | MODE_TNY e ⇒ opt_map … (memory_filter_read m t s 〈〈x0,x0〉:〈x0,e〉〉)
+ (λb.Some ? (triple … s b cur_pc))
(* preleva 1 byte da 0000 0000 000x xxxxb *)
- | MODE_SRT e ⇒ opt_map ?? (memory_filter_read m t s 〈〈x0,x0〉:(byte8_of_bitrigesim e)〉)
- (λb.Some ? (triple ??? s b cur_pc))
+ | MODE_SRT e ⇒ opt_map … (memory_filter_read m t s 〈〈x0,x0〉:(byte8_of_bitrigesim e)〉)
+ (λb.Some ? (triple … s b cur_pc))
]
(* lettura di una word *)
| false ⇒ λs:any_status m t.λcur_pc:word16.λi:instr_mode.match i with
| MODE_INHH ⇒ None ?
(* preleva 1 word immediato *)
- | MODE_INHX0ADD ⇒ opt_map ?? (get_IX m t s)
- (λw.Some ? (triple ??? s w cur_pc))
+ | MODE_INHX0ADD ⇒ opt_map … (get_IX m t s)
+ (λw.Some ? (triple … s w cur_pc))
(* preleva 1 word immediato *)
| MODE_INHX1ADD ⇒ mode_IX1ADD_load m t s cur_pc
(* preleva 1 word immediato *)
| MODE_DIR1_to_IX0p ⇒ None ?
(* preleva 2 byte, possibilita' modificare Io argomento *)
- | MODE_INHA_and_IMM1 ⇒ opt_map ?? (mode_IMM1_load m t s cur_pc)
+ | MODE_INHA_and_IMM1 ⇒ opt_map … (mode_IMM1_load m t s cur_pc)
(λS_immb_and_PC.match S_immb_and_PC with
[ triple _ immb cur_pc' ⇒
- Some ? (triple ??? s 〈(get_acc_8_low_reg m t s):immb〉 cur_pc')])
+ Some ? (triple … s 〈(get_acc_8_low_reg m t s):immb〉 cur_pc')])
(* preleva 2 byte, possibilita' modificare Io argomento *)
- | MODE_INHX_and_IMM1 ⇒ opt_map ?? (get_indX_8_low_reg m t s)
- (λX_op.opt_map ?? (mode_IMM1_load m t s cur_pc)
+ | MODE_INHX_and_IMM1 ⇒ opt_map … (get_indX_8_low_reg m t s)
+ (λX_op.opt_map … (mode_IMM1_load m t s cur_pc)
(λS_immb_and_PC.match S_immb_and_PC with
[ triple _ immb cur_pc' ⇒
- Some ? (triple ??? s 〈X_op:immb〉 cur_pc')]))
+ Some ? (triple … s 〈X_op:immb〉 cur_pc')]))
(* preleva 2 byte, NO possibilita' modificare Io argomento *)
- | MODE_IMM1_and_IMM1 ⇒ opt_map ?? (mode_IMM1_load m t s cur_pc)
+ | MODE_IMM1_and_IMM1 ⇒ opt_map … (mode_IMM1_load m t s cur_pc)
(λS_immb1_and_PC.match S_immb1_and_PC with
[ triple _ immb1 cur_pc' ⇒
- opt_map ?? (mode_IMM1_load m t s cur_pc')
+ opt_map … (mode_IMM1_load m t s cur_pc')
(λS_immb2_and_PC.match S_immb2_and_PC with
[ triple _ immb2 cur_pc'' ⇒
- Some ? (triple ??? s 〈immb1:immb2〉 cur_pc'')])])
+ Some ? (triple … s 〈immb1:immb2〉 cur_pc'')])])
(* preleva 2 byte, possibilita' modificare Io argomento *)
- | MODE_DIR1_and_IMM1 ⇒ opt_map ?? (mode_DIR1_load true m t s cur_pc)
+ | MODE_DIR1_and_IMM1 ⇒ opt_map … (mode_DIR1_load true m t s cur_pc)
(λS_dirb_and_PC.match S_dirb_and_PC with
[ triple _ dirb cur_pc' ⇒
- opt_map ?? (mode_IMM1_load m t s cur_pc')
+ opt_map … (mode_IMM1_load m t s cur_pc')
(λS_immb_and_PC.match S_immb_and_PC with
[ triple _ immb cur_pc'' ⇒
- Some ? (triple ??? s 〈dirb:immb〉 cur_pc'')])])
+ Some ? (triple … s 〈dirb:immb〉 cur_pc'')])])
(* preleva 2 byte, possibilita' modificare Io argomento *)
- | MODE_IX0_and_IMM1 ⇒ opt_map ?? (mode_IX0_load true m t s cur_pc)
+ | MODE_IX0_and_IMM1 ⇒ opt_map … (mode_IX0_load true m t s cur_pc)
(λS_ixb_and_PC.match S_ixb_and_PC with
[ triple _ ixb cur_pc' ⇒
- opt_map ?? (mode_IMM1_load m t s cur_pc')
+ opt_map … (mode_IMM1_load m t s cur_pc')
(λS_immb_and_PC.match S_immb_and_PC with
[ triple _ immb cur_pc'' ⇒
- Some ? (triple ??? s 〈ixb:immb〉 cur_pc'')])])
+ Some ? (triple … s 〈ixb:immb〉 cur_pc'')])])
(* preleva 2 byte, H:X++, NO possibilita' modificare Io argomento *)
- | MODE_IX0p_and_IMM1 ⇒ opt_map ?? (mode_IX0_load true m t s cur_pc)
+ | MODE_IX0p_and_IMM1 ⇒ opt_map … (mode_IX0_load true m t s cur_pc)
(λS_ixb_and_PC.match S_ixb_and_PC with
[ triple _ ixb cur_pc' ⇒
- opt_map ?? (mode_IMM1_load m t s cur_pc')
+ opt_map … (mode_IMM1_load m t s cur_pc')
(λS_immb_and_PC.match S_immb_and_PC with
[ triple _ immb cur_pc'' ⇒
(* H:X++ *)
- opt_map ?? (aux_inc_indX_16 m t s)
- (λs'.Some ? (triple ??? s' 〈ixb:immb〉 cur_pc''))])])
+ opt_map … (aux_inc_indX_16 m t s)
+ (λs'.Some ? (triple … s' 〈ixb:immb〉 cur_pc''))])])
(* preleva 2 byte, possibilita' modificare Io argomento *)
- | MODE_IX1_and_IMM1 ⇒ opt_map ?? (mode_IX1_load true m t s cur_pc)
+ | MODE_IX1_and_IMM1 ⇒ opt_map … (mode_IX1_load true m t s cur_pc)
(λS_ixb_and_PC.match S_ixb_and_PC with
[ triple _ ixb cur_pc' ⇒
- opt_map ?? (mode_IMM1_load m t s cur_pc')
+ opt_map … (mode_IMM1_load m t s cur_pc')
(λS_immb_and_PC.match S_immb_and_PC with
[ triple _ immb cur_pc'' ⇒
- Some ? (triple ??? s 〈ixb:immb〉 cur_pc'')])])
+ Some ? (triple … s 〈ixb:immb〉 cur_pc'')])])
(* preleva 2 byte, H:X++, NO possibilita' modificare Io argomento *)
- | MODE_IX1p_and_IMM1 ⇒ opt_map ?? (mode_IX1_load true m t s cur_pc)
+ | MODE_IX1p_and_IMM1 ⇒ opt_map … (mode_IX1_load true m t s cur_pc)
(λS_ixb_and_PC.match S_ixb_and_PC with
[ triple _ ixb cur_pc' ⇒
- opt_map ?? (mode_IMM1_load m t s cur_pc')
+ opt_map … (mode_IMM1_load m t s cur_pc')
(λS_immb_and_PC.match S_immb_and_PC with
[ triple _ immb cur_pc'' ⇒
(* H:X++ *)
- opt_map ?? (aux_inc_indX_16 m t s)
- (λs'.Some ? (triple ??? s' 〈ixb:immb〉 cur_pc''))])])
+ opt_map … (aux_inc_indX_16 m t s)
+ (λs'.Some ? (triple … s' 〈ixb:immb〉 cur_pc''))])])
(* preleva 2 byte, possibilita' modificare Io argomento *)
- | MODE_SP1_and_IMM1 ⇒ opt_map ?? (mode_SP1_load true m t s cur_pc)
+ | MODE_SP1_and_IMM1 ⇒ opt_map … (mode_SP1_load true m t s cur_pc)
(λS_spb_and_PC.match S_spb_and_PC with
[ triple _ spb cur_pc' ⇒
- opt_map ?? (mode_IMM1_load m t s cur_pc')
+ opt_map … (mode_IMM1_load m t s cur_pc')
(λS_immb_and_PC.match S_immb_and_PC with
[ triple _ immb cur_pc'' ⇒
- Some ? (triple ??? s 〈spb:immb〉 cur_pc'')])])
+ Some ? (triple … s 〈spb:immb〉 cur_pc'')])])
(* NO: solo scrittura byte *)
| MODE_DIRn _ ⇒ None ?
(* preleva 2 byte, il primo e' filtrato per azzerare tutti i bit tranne n-simo *)
- | MODE_DIRn_and_IMM1 msk ⇒ opt_map ?? (mode_DIR1n_load m t s cur_pc msk)
+ | MODE_DIRn_and_IMM1 msk ⇒ opt_map … (mode_DIR1n_load m t s cur_pc msk)
(λS_dirbn_and_PC.match S_dirbn_and_PC with
[ triple _ dirbn cur_pc' ⇒
- opt_map ?? (mode_IMM1_load m t s cur_pc')
+ opt_map … (mode_IMM1_load m t s cur_pc')
(λS_immb_and_PC.match S_immb_and_PC with
[ triple _ immb cur_pc'' ⇒
- Some ? (triple ??? s 〈〈x0,match dirbn with [ true ⇒ x1 | false ⇒ x0 ]〉:immb〉 cur_pc'') ])])
+ Some ? (triple … s 〈〈x0,match dirbn with [ true ⇒ x1 | false ⇒ x0 ]〉:immb〉 cur_pc'') ])])
(* NO: solo lettura/scrittura byte *)
| MODE_TNY _ ⇒ None ?
(* NO: solo lettura/scrittura byte *)
(* NO: non ci sono indicazioni *)
[ MODE_INH ⇒ None ?
(* scrive A *)
- | MODE_INHA ⇒ Some ? (pair ?? (set_acc_8_low_reg m t s writeb) cur_pc)
+ | MODE_INHA ⇒ Some ? (pair … (set_acc_8_low_reg m t s writeb) cur_pc)
(* scrive X *)
- | MODE_INHX ⇒ opt_map ?? (set_indX_8_low_reg m t s writeb)
- (λtmps.Some ? (pair ?? tmps cur_pc))
+ | MODE_INHX ⇒ opt_map … (set_indX_8_low_reg m t s writeb)
+ (λtmps.Some ? (pair … tmps cur_pc))
(* scrive H *)
- | MODE_INHH ⇒ opt_map ?? (set_indX_8_high_reg m t s writeb)
- (λtmps.Some ? (pair ?? tmps cur_pc))
+ | MODE_INHH ⇒ opt_map … (set_indX_8_high_reg m t s writeb)
+ (λtmps.Some ? (pair … tmps cur_pc))
(* NO: solo lettura word *)
| MODE_INHX0ADD ⇒ None ?
(* passo2: scrittura su DIR1, passo1: lettura da IMM1 *)
| MODE_IMM1_to_DIR1 ⇒ mode_DIR1_write true m t s cur_pc writeb
(* passo2: scrittura su DIR1 e X++, passo1: lettura da IX0 *)
- | MODE_IX0p_to_DIR1 ⇒ opt_map ?? (mode_DIR1_write true m t s cur_pc writeb)
+ | MODE_IX0p_to_DIR1 ⇒ opt_map … (mode_DIR1_write true m t s cur_pc writeb)
(λS_and_PC.match S_and_PC with [ pair S_op PC_op ⇒
(* H:X++ *)
- opt_map ?? (aux_inc_indX_16 m t S_op)
- (λS_op'.Some ? (pair ?? S_op' PC_op))])
+ opt_map … (aux_inc_indX_16 m t S_op)
+ (λS_op'.Some ? (pair … S_op' PC_op))])
(* passo2: scrittura su IX0 e X++, passo1: lettura da DIR1 *)
- | MODE_DIR1_to_IX0p ⇒ opt_map ?? (mode_IX0_write true m t s cur_pc writeb)
+ | MODE_DIR1_to_IX0p ⇒ opt_map … (mode_IX0_write true m t s cur_pc writeb)
(λS_and_PC.match S_and_PC with [ pair S_op PC_op ⇒
(* H:X++ *)
- opt_map ?? (aux_inc_indX_16 m t S_op)
- (λS_op'.Some ? (pair ?? S_op' PC_op))])
+ opt_map … (aux_inc_indX_16 m t S_op)
+ (λS_op'.Some ? (pair … S_op' PC_op))])
(* dopo aver prelevato 2 byte la possibilita' modificare Io argomento = INHA *)
- | MODE_INHA_and_IMM1 ⇒ Some ? (pair ?? (set_acc_8_low_reg m t s writeb) cur_pc)
+ | MODE_INHA_and_IMM1 ⇒ Some ? (pair … (set_acc_8_low_reg m t s writeb) cur_pc)
(* dopo aver prelevato 2 byte la possibilita' modificare Io argomento = INHX *)
- | MODE_INHX_and_IMM1 ⇒ opt_map ?? (set_indX_8_low_reg m t s writeb)
- (λtmps.Some ? (pair ?? tmps cur_pc))
+ | MODE_INHX_and_IMM1 ⇒ opt_map … (set_indX_8_low_reg m t s writeb)
+ (λtmps.Some ? (pair … tmps cur_pc))
(* NO: solo lettura word *)
| MODE_IMM1_and_IMM1 ⇒ None ?
(* dopo aver prelevato 2 byte la possibilita' modificare Io argomento = DIR1 *)
(* NO: solo lettura word *)
| MODE_DIRn_and_IMM1 _ ⇒ None ?
(* scrive 1 byte su 0000 0000 0000 xxxxb *)
- | MODE_TNY e ⇒ opt_map ?? (memory_filter_write m t s 〈〈x0,x0〉:〈x0,e〉〉 writeb)
- (λtmps.Some ? (pair ?? tmps cur_pc))
+ | MODE_TNY e ⇒ opt_map … (memory_filter_write m t s 〈〈x0,x0〉:〈x0,e〉〉 writeb)
+ (λtmps.Some ? (pair … tmps cur_pc))
(* scrive 1 byte su 0000 0000 000x xxxxb *)
- | MODE_SRT e ⇒ opt_map ?? (memory_filter_write m t s 〈〈x0,x0〉:(byte8_of_bitrigesim e)〉 writeb)
- (λtmps.Some ? (pair ?? tmps cur_pc)) ]
+ | MODE_SRT e ⇒ opt_map … (memory_filter_write m t s 〈〈x0,x0〉:(byte8_of_bitrigesim e)〉 writeb)
+ (λtmps.Some ? (pair … tmps cur_pc)) ]
(* scrittura di una word *)
| false ⇒ λs:any_status m t.λcur_pc:word16.λi:instr_mode.λwritew:word16.match i with
(* NO: non ci sono indicazioni *)
[ TickERR s _ ⇒ None ?
(* azzeramento tutta RAM tranne dati *)
| TickSUSP s _ ⇒ None ?
- | TickOK s ⇒ Some ? (byte8_hexdump t (get_mem_desc HCS08 t s) dTest_HCS08_RAM (nat_of_word16 (byte8_strlen string)))
+ | TickOK s ⇒ Some ? (byte8_hexdump t (mem_desc HCS08 t s) dTest_HCS08_RAM (nat_of_word16 (byte8_strlen string)))
] =
Some ? (reverse_list ? string)).
[ TickERR s _ ⇒ None ?
(* azzeramento tutta RAM tranne dati *)
| TickSUSP s _ ⇒ None ?
- | TickOK s ⇒ Some ? (set_mem_desc HCS08 t s (load_from_source_at t (get_mem_desc HCS08 t s) dTest_zeros 〈〈x0,xD〉:〈x0,x0〉〉))
+ | TickOK s ⇒ Some ? (set_mem_desc HCS08 t s (load_from_source_at t (mem_desc HCS08 t s) dTest_zeros 〈〈x0,xD〉:〈x0,x0〉〉))
] =
Some ? (set_pc_reg HCS08 t
- (dTest_HCS08_sReverse_status t (fst ?? (shr_b8 (w16h (byte8_strlen string)))) (fst ?? (shr_w16 (byte8_strlen string))) (byte8_strlen string) (reverse_list ? string))
+ (dTest_HCS08_sReverse_status t (fst … (shr_b8 (w16h (byte8_strlen string)))) (fst … (shr_w16 (byte8_strlen string))) (byte8_strlen string) (reverse_list ? string))
(mk_word16 (mk_byte8 x1 x9) (mk_byte8 x2 xB)))).
(*
(42+79*(nat_of_word16 (byte8_strlen string))+5*((nat_of_word16 (byte8_strlen string))/512)) with
[ TickERR s _ ⇒ None ?
| TickSUSP s _ ⇒ None ?
- | TickOK s ⇒ Some ? (set_mem_desc HCS08 MEM_TREE s (load_from_source_at MEM_TREE (get_mem_desc HCS08 MEM_TREE s) dTest_zeros 〈〈x0,xD〉:〈x0,x0〉〉))
+ | TickOK s ⇒ Some ? (set_mem_desc HCS08 MEM_TREE s (load_from_source_at MEM_TREE (mem_desc HCS08 MEM_TREE s) dTest_zeros 〈〈x0,xD〉:〈x0,x0〉〉))
].
ndefinition sReverseNoCalc ≝
λstring:list byte8.
Some ? (set_pc_reg HCS08 MEM_TREE
- (dTest_HCS08_sReverse_status MEM_TREE (fst ?? (shr_b8 (w16h (byte8_strlen string))))
- (fst ?? (shr_w16 (byte8_strlen string)))
+ (dTest_HCS08_sReverse_status MEM_TREE (fst … (shr_b8 (w16h (byte8_strlen string))))
+ (fst … (shr_w16 (byte8_strlen string)))
(byte8_strlen string) (reverse_list ? string))
(mk_word16 (mk_byte8 x1 x9) (mk_byte8 x2 xB))).
((nat_of_word16 〈〈x6,x4〉:〈x6,x4〉〉)+(nat_of_byte8 〈x9,x6〉)*(nat_of_word16 (byte8_strlen string))) with
[ TickERR s _ ⇒ None ?
(* azzeramento tutta RAM tranne dati *)
- | TickSUSP s _ ⇒ Some ? (set_mem_desc HCS08 t s (load_from_source_at t (get_mem_desc HCS08 t s) dTest_zeros 〈〈x0,xD〉:〈x0,x0〉〉))
+ | TickSUSP s _ ⇒ Some ? (set_mem_desc HCS08 t s (load_from_source_at t (mem_desc HCS08 t s) dTest_zeros 〈〈x0,xD〉:〈x0,x0〉〉))
| TickOK s ⇒ None ?
] =
Some ? (set_pc_reg HCS08 t
(TickOK ? (dTest_HCS08_cSort_status MEM_TREE true 〈x0,x0〉 〈〈x0,xF〉:〈x4,xC〉〉 (byte8_strlen string) string))
((nat_of_word16 〈〈x6,x4〉:〈x6,x4〉〉)+(nat_of_byte8 〈x9,x6〉)*(nat_of_word16 (byte8_strlen string))) with
[ TickERR s _ ⇒ None ?
- | TickSUSP s _ ⇒ Some ? (set_mem_desc HCS08 MEM_TREE s (load_from_source_at MEM_TREE (get_mem_desc HCS08 MEM_TREE s) dTest_zeros 〈〈x0,xD〉:〈x0,x0〉〉))
+ | TickSUSP s _ ⇒ Some ? (set_mem_desc HCS08 MEM_TREE s (load_from_source_at MEM_TREE (mem_desc HCS08 MEM_TREE s) dTest_zeros 〈〈x0,xD〉:〈x0,x0〉〉))
| TickOK s ⇒ None ?
].
(nat_of_word16 num) with
[ TickERR s _ ⇒ None ?
(* azzeramento tutta RAM tranne dati *)
- | TickSUSP s _ ⇒ Some ? (set_mem_desc HCS08 t s (load_from_source_at t (get_mem_desc HCS08 t s) dTest_zeros3K 〈〈x0,x1〉:〈x2,x0〉〉))
+ | TickSUSP s _ ⇒ Some ? (set_mem_desc HCS08 t s (load_from_source_at t (mem_desc HCS08 t s) dTest_zeros3K 〈〈x0,x1〉:〈x2,x0〉〉))
| TickOK s ⇒ None ?
] =
Some ? (dTest_HCS08_gNum_status t false 〈x0,x0〉 num 〈〈x1,x9〉:〈x5,x1〉〉 〈〈x1,x8〉:〈xB,xE〉〉 num (dTest_HCS08_gNum_aurei num)).
(TickOK ? (dTest_HCS08_gNum_status MEM_TREE true 〈x0,x0〉 〈〈x1,xA〉:〈x0,x0〉〉 〈〈x1,x8〉:〈xB,xE〉〉 〈〈x1,x8〉:〈xB,xE〉〉 num dTest_zeros))
(nat_of_word16 num) with
[ TickERR s _ ⇒ None ?
- | TickSUSP s _ ⇒ Some ? (set_mem_desc HCS08 MEM_TREE s (load_from_source_at MEM_TREE (get_mem_desc HCS08 MEM_TREE s) dTest_zeros3K 〈〈x0,x1〉:〈x2,x0〉〉))
+ | TickSUSP s _ ⇒ Some ? (set_mem_desc HCS08 MEM_TREE s (load_from_source_at MEM_TREE (mem_desc HCS08 MEM_TREE s) dTest_zeros3K 〈〈x0,x1〉:〈x2,x0〉〉))
| TickOK s ⇒ None ?
].
λc:aux_chk_type MEM_FUNC.
λaddr:word16.
λo:oct.
- opt_map ?? (mf_mem_read m c addr)
+ opt_map … (mf_mem_read m c addr)
(λb.Some ? (getn_array8T o bool (bits_of_byte8 b)))
| MEM_TREE ⇒ λm:aux_mem_type MEM_TREE.
λc:aux_chk_type MEM_TREE.
λaddr:word16.
λo:oct.
- opt_map ?? (mt_mem_read m c addr)
+ opt_map … (mt_mem_read m c addr)
(λb.Some ? (getn_array8T o bool (bits_of_byte8 b)))
| MEM_BITS ⇒ λm:aux_mem_type MEM_BITS.
λc:aux_chk_type MEM_BITS.
λaddr:word16.
λo:oct.
λv:bool.
- opt_map ?? (mf_mem_read m c addr)
+ opt_map … (mf_mem_read m c addr)
(λb.mf_mem_update m (chk_get MEM_FUNC c addr) addr (byte8_of_bits (setn_array8T o bool (bits_of_byte8 b) v)))
| MEM_TREE ⇒ λm:aux_mem_type MEM_TREE.
λc:aux_chk_type MEM_TREE.
λaddr:word16.
λo:oct.
λv:bool.
- opt_map ?? (mt_mem_read m c addr)
+ opt_map … (mt_mem_read m c addr)
(λb.mt_mem_update m (chk_get MEM_TREE c addr) addr (byte8_of_bits (setn_array8T o bool (bits_of_byte8 b) v)))
| MEM_BITS ⇒ λm:aux_mem_type MEM_BITS.
λc:aux_chk_type MEM_BITS.
[ MEM_READ_ONLY ⇒ Some bool (getn_array8T o7 bool old_value)
| MEM_READ_WRITE ⇒ Some bool (getn_array8T o7 bool new_value)
| MEM_OUT_OF_BOUND ⇒ None bool ] in
- opt_map ?? newbit0
- (λnb0.opt_map ?? newbit1
- (λnb1.opt_map ?? newbit2
- (λnb2.opt_map ?? newbit3
- (λnb3.opt_map ?? newbit4
- (λnb4.opt_map ?? newbit5
- (λnb5.opt_map ?? newbit6
- (λnb6.opt_map ?? newbit7
+ opt_map … newbit0
+ (λnb0.opt_map … newbit1
+ (λnb1.opt_map … newbit2
+ (λnb2.opt_map … newbit3
+ (λnb3.opt_map … newbit4
+ (λnb4.opt_map … newbit5
+ (λnb5.opt_map … newbit6
+ (λnb6.opt_map … newbit7
(λnb7.Some ? (mt_update (Array8T bool) mem addr (array_8T bool nb7 nb6 nb5 nb4 nb3 nb2 nb1 nb0)))))))))).
(* leggi controllando il tipo di memoria *)
[ MEM_READ_ONLY ⇒ Some bool (getn_array8T o7 bool value)
| MEM_READ_WRITE ⇒ Some bool (getn_array8T o7 bool value)
| MEM_OUT_OF_BOUND ⇒ None bool ] in
- opt_map ?? newbit0
- (λnb0.opt_map ?? newbit1
- (λnb1.opt_map ?? newbit2
- (λnb2.opt_map ?? newbit3
- (λnb3.opt_map ?? newbit4
- (λnb4.opt_map ?? newbit5
- (λnb5.opt_map ?? newbit6
- (λnb6.opt_map ?? newbit7
+ opt_map … newbit0
+ (λnb0.opt_map … newbit1
+ (λnb1.opt_map … newbit2
+ (λnb2.opt_map … newbit3
+ (λnb3.opt_map … newbit4
+ (λnb4.opt_map … newbit5
+ (λnb5.opt_map … newbit6
+ (λnb6.opt_map … newbit7
(λnb7.Some ? (byte8_of_bits (array_8T bool nb7 nb6 nb5 nb4 nb3 nb2 nb1 nb0)))))))))).
(* ************************** *)
| TickOK s ⇒ get_alu HCS08 MEM_FUNC s ]);
normalize in AFTER_ALU8:(%); *)
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
(* ********* *)
(* astraggo molto *)
(* 0x0000-0x001F,0x0FF0: sarebbe memory mapped IO *)
- triple ??? 〈〈x0,x0〉:〈x8,x0〉〉 〈〈x0,x0〉:〈xF,xF〉〉 MEM_READ_WRITE (* 128B RAM+STACK *)
- ; triple ??? 〈〈x0,x3〉:〈x0,x0〉〉 〈〈x0,xC〉:〈xF,xF〉〉 MEM_READ_ONLY (* 2560B USER ROM *)
- ; triple ??? 〈〈x0,xE〉:〈x0,x0〉〉 〈〈x0,xF〉:〈xF,xF〉〉 MEM_READ_ONLY (* 512B INTERNAL ROM *)
+ triple … 〈〈x0,x0〉:〈x8,x0〉〉 〈〈x0,x0〉:〈xF,xF〉〉 MEM_READ_WRITE (* 128B RAM+STACK *)
+ ; triple … 〈〈x0,x3〉:〈x0,x0〉〉 〈〈x0,xC〉:〈xF,xF〉〉 MEM_READ_ONLY (* 2560B USER ROM *)
+ ; triple … 〈〈x0,xE〉:〈x0,x0〉〉 〈〈x0,xF〉:〈xF,xF〉〉 MEM_READ_ONLY (* 512B INTERNAL ROM *)
]
(*..*)
].
(* 0x0500-0x057F,0xFE02,0xFE04-0xFE07,
0xFE09-0xFE0B,0xFE12-0xFE19,0xFE1E,0xFFC0-0xFFCF : sarebbe reserved *)
- triple ??? 〈〈x0,x0〉:〈x5,x0〉〉 〈〈x0,x2〉:〈x4,xF〉〉 MEM_READ_WRITE (* 512B RAM *)
- ; triple ??? 〈〈x0,x8〉:〈x0,x0〉〉 〈〈x0,x9〉:〈xF,xF〉〉 MEM_READ_ONLY (* 512B EEPROM *)
- ; triple ??? 〈〈xB,xE〉:〈x0,x0〉〉 〈〈xF,xD〉:〈xF,xF〉〉 MEM_READ_ONLY (* 16384B ROM *)
- ; triple ??? 〈〈xF,xE〉:〈x2,x0〉〉 〈〈xF,xF〉:〈x5,x2〉〉 MEM_READ_ONLY (* 307B ROM *)
- ; triple ??? 〈〈xF,xF〉:〈xD,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY (* 48B ROM *) ]
+ triple … 〈〈x0,x0〉:〈x5,x0〉〉 〈〈x0,x2〉:〈x4,xF〉〉 MEM_READ_WRITE (* 512B RAM *)
+ ; triple … 〈〈x0,x8〉:〈x0,x0〉〉 〈〈x0,x9〉:〈xF,xF〉〉 MEM_READ_ONLY (* 512B EEPROM *)
+ ; triple … 〈〈xB,xE〉:〈x0,x0〉〉 〈〈xF,xD〉:〈xF,xF〉〉 MEM_READ_ONLY (* 16384B ROM *)
+ ; triple … 〈〈xF,xE〉:〈x2,x0〉〉 〈〈xF,xF〉:〈x5,x2〉〉 MEM_READ_ONLY (* 307B ROM *)
+ ; triple … 〈〈xF,xF〉:〈xD,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY (* 48B ROM *) ]
(*..*)
].
(* astraggo molto *)
(* 0x0000-0x006F,0x1800-0x185F: sarebbe memory mapped IO *)
- triple ??? 〈〈x0,x0〉:〈x7,x0〉〉 〈〈x0,x8〉:〈x6,xF〉〉 MEM_READ_WRITE (* 2048B RAM *)
- ; triple ??? 〈〈x0,x8〉:〈x7,x0〉〉 〈〈x1,x7〉:〈xF,xF〉〉 MEM_READ_ONLY (* 3984B FLASH *)
- ; triple ??? 〈〈x1,x8〉:〈x6,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY (* 59296B FLASH *) ]
+ triple … 〈〈x0,x0〉:〈x7,x0〉〉 〈〈x0,x8〉:〈x6,xF〉〉 MEM_READ_WRITE (* 2048B RAM *)
+ ; triple … 〈〈x0,x8〉:〈x7,x0〉〉 〈〈x1,x7〉:〈xF,xF〉〉 MEM_READ_ONLY (* 3984B FLASH *)
+ ; triple … 〈〈x1,x8〉:〈x6,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY (* 59296B FLASH *) ]
| MC9S08GB60 ⇒
[
(* astraggo molto *)
(* 0x0000-0x006F,0x1800-0x185F: sarebbe memory mapped IO *)
- triple ??? 〈〈x0,x0〉:〈x8,x0〉〉 〈〈x1,x0〉:〈x7,xF〉〉 MEM_READ_WRITE (* 4096B RAM *)
- ; triple ??? 〈〈x1,x0〉:〈x8,x0〉〉 〈〈x1,x7〉:〈xF,xF〉〉 MEM_READ_ONLY (* 1920B FLASH *)
- ; triple ??? 〈〈x1,x8〉:〈x2,xC〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY (* 59348B FLASH *) ]
+ triple … 〈〈x0,x0〉:〈x8,x0〉〉 〈〈x1,x0〉:〈x7,xF〉〉 MEM_READ_WRITE (* 4096B RAM *)
+ ; triple … 〈〈x1,x0〉:〈x8,x0〉〉 〈〈x1,x7〉:〈xF,xF〉〉 MEM_READ_ONLY (* 1920B FLASH *)
+ ; triple … 〈〈x1,x8〉:〈x2,xC〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY (* 59348B FLASH *) ]
].
(* memoria dei RS08 *)
λm:RS08_mcu_model.match m with
[ MC9RS08KA1 ⇒
[
- triple ??? 〈〈x0,x0〉:〈x0,x0〉〉 〈〈x0,x0〉:〈x0,xE〉〉 MEM_READ_WRITE (* 15B RAM *)
+ triple … 〈〈x0,x0〉:〈x0,x0〉〉 〈〈x0,x0〉:〈x0,xE〉〉 MEM_READ_WRITE (* 15B RAM *)
(* [000F] e' il registro X *)
(* 0x0010-0x001E sarebbe memory mapped IO, proviamo per completezza con RAM *)
- ; triple ??? 〈〈x0,x0〉:〈x1,x0〉〉 〈〈x0,x0〉:〈x1,xE〉〉 MEM_READ_WRITE (* 15B MEMORY MAPPED IO *)
+ ; triple … 〈〈x0,x0〉:〈x1,x0〉〉 〈〈x0,x0〉:〈x1,xE〉〉 MEM_READ_WRITE (* 15B MEMORY MAPPED IO *)
(* [001F] e' il registro PAGESEL *)
- ; triple ??? 〈〈x0,x0〉:〈x2,x0〉〉 〈〈x0,x0〉:〈x4,xF〉〉 MEM_READ_WRITE (* 48B RAM *)
+ ; triple … 〈〈x0,x0〉:〈x2,x0〉〉 〈〈x0,x0〉:〈x4,xF〉〉 MEM_READ_WRITE (* 48B RAM *)
(* [00C0-00FF] mappato da PAGESEL su [00pp pppp ppxx xxxx] *)
- ; triple ??? 〈〈x0,x0〉:〈xC,x0〉〉 〈〈x0,x0〉:〈xF,xF〉〉 MEM_READ_WRITE (* 64B RAM PAGING *)
+ ; triple … 〈〈x0,x0〉:〈xC,x0〉〉 〈〈x0,x0〉:〈xF,xF〉〉 MEM_READ_WRITE (* 64B RAM PAGING *)
(* 0x0200-0x023F sarebbe memory mapped IO, proviamo per completezza con RAM *)
- ; triple ??? 〈〈x0,x2〉:〈x0,x0〉〉 〈〈x0,x2〉:〈x3,xF〉〉 MEM_READ_WRITE (* 64B MEMORY MAPPED IO *)
- ; triple ??? 〈〈x3,xC〉:〈x0,x0〉〉 〈〈x3,xF〉:〈xF,xF〉〉 MEM_READ_ONLY (* 1024B FLASH *) ]
+ ; triple … 〈〈x0,x2〉:〈x0,x0〉〉 〈〈x0,x2〉:〈x3,xF〉〉 MEM_READ_WRITE (* 64B MEMORY MAPPED IO *)
+ ; triple … 〈〈x3,xC〉:〈x0,x0〉〉 〈〈x3,xF〉:〈xF,xF〉〉 MEM_READ_ONLY (* 1024B FLASH *) ]
| MC9RS08KA2 ⇒
[
- triple ??? 〈〈x0,x0〉:〈x0,x0〉〉 〈〈x0,x0〉:〈x0,xE〉〉 MEM_READ_WRITE (* 15B RAM *)
+ triple … 〈〈x0,x0〉:〈x0,x0〉〉 〈〈x0,x0〉:〈x0,xE〉〉 MEM_READ_WRITE (* 15B RAM *)
(* [000F] e' il registro X *)
(* 0x0010-0x001E sarebbe memory mapped IO, proviamo per completezza con RAM *)
- ; triple ??? 〈〈x0,x0〉:〈x1,x0〉〉 〈〈x0,x0〉:〈x1,xE〉〉 MEM_READ_WRITE (* 15B MEMORY MAPPED IO *)
+ ; triple … 〈〈x0,x0〉:〈x1,x0〉〉 〈〈x0,x0〉:〈x1,xE〉〉 MEM_READ_WRITE (* 15B MEMORY MAPPED IO *)
(* [001F] e' il registro PAGESEL *)
- ; triple ??? 〈〈x0,x0〉:〈x2,x0〉〉 〈〈x0,x0〉:〈x4,xF〉〉 MEM_READ_WRITE (* 48B RAM *)
+ ; triple … 〈〈x0,x0〉:〈x2,x0〉〉 〈〈x0,x0〉:〈x4,xF〉〉 MEM_READ_WRITE (* 48B RAM *)
(* [00C0-00FF] mappato da PAGESEL su [00pp pppp ppxx xxxx] *)
- ; triple ??? 〈〈x0,x0〉:〈xC,x0〉〉 〈〈x0,x0〉:〈xF,xF〉〉 MEM_READ_WRITE (* 64B RAM PAGING *)
+ ; triple … 〈〈x0,x0〉:〈xC,x0〉〉 〈〈x0,x0〉:〈xF,xF〉〉 MEM_READ_WRITE (* 64B RAM PAGING *)
(* 0x0200-0x023F sarebbe memory mapped IO, proviamo per completezza con RAM *)
- ; triple ??? 〈〈x0,x2〉:〈x0,x0〉〉 〈〈x0,x2〉:〈x3,xF〉〉 MEM_READ_WRITE (* 64B MEMORY MAPPED IO *)
- ; triple ??? 〈〈x3,x8〉:〈x0,x0〉〉 〈〈x3,xF〉:〈xF,xF〉〉 MEM_READ_ONLY (* 2048B FLASH *) ]
+ ; triple … 〈〈x0,x2〉:〈x0,x0〉〉 〈〈x0,x2〉:〈x3,xF〉〉 MEM_READ_WRITE (* 64B MEMORY MAPPED IO *)
+ ; triple … 〈〈x3,x8〉:〈x0,x0〉〉 〈〈x3,xF〉:〈xF,xF〉〉 MEM_READ_ONLY (* 2048B FLASH *) ]
].
(* ∀modello.descrizione della memoria installata *)
ndefinition execute_ADC_ADD_aux ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.λsetflag:bool.
λfAMC:byte8 → byte8 → bool → ProdT byte8 bool.
- opt_map ?? (multi_mode_load true m t s cur_pc i)
+ opt_map … (multi_mode_load true m t s cur_pc i)
(λS_M_PC.match S_M_PC with
[ triple s_tmp1 M_op new_pc ⇒
let A_op ≝ get_acc_8_low_reg m t s_tmp1 in
(* V = A7&M7&nR7 | nA7&nM7&R7 *)
let s_tmp7 ≝ setweak_v_flag m t s_tmp6 ((A7⊗M7⊗(⊖R7)) ⊕ ((⊖A7)⊗(⊖M7)⊗R7)) in
(* newpc = nextpc *)
- Some ? (pair ?? s_tmp7 new_pc) ]]).
+ Some ? (pair … s_tmp7 new_pc) ]]).
(* A = [true] fAM(A,M), [false] A *)
(* cioe' in caso di false l'operazione viene eseguita ma modifica solo i flag *)
ndefinition execute_AND_BIT_EOR_ORA_aux ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.λsetflag:bool.
λfAM:byte8 → byte8 → byte8.
- opt_map ?? (multi_mode_load true m t s cur_pc i)
+ opt_map … (multi_mode_load true m t s cur_pc i)
(λS_M_PC.match S_M_PC with
[ triple s_tmp1 M_op new_pc ⇒
let R_op ≝ fAM (get_acc_8_low_reg m t s_tmp1) M_op in
(* V = 0 *)
let s_tmp5 ≝ setweak_v_flag m t s_tmp4 false in
(* newpc = nextpc *)
- Some ? (pair ?? s_tmp5 new_pc) ]).
+ Some ? (pair … s_tmp5 new_pc) ]).
(* M = fMC(M,C) *)
(* fMC e' la logica da applicare: rc_/ro_/sh_ *)
ndefinition execute_ASL_ASR_LSR_ROL_ROR_aux ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
λfMC:byte8 → bool → ProdT byte8 bool.
- opt_map ?? (multi_mode_load true m t s cur_pc i)
+ opt_map … (multi_mode_load true m t s cur_pc i)
(λS_M_PC.match S_M_PC with
[ triple s_tmp1 M_op _ ⇒
match fMC M_op (get_c_flag m t s_tmp1) with [ pair R_op carry ⇒
(* M = fMC(M,C) *)
- opt_map ?? (multi_mode_write true m t s_tmp1 cur_pc i R_op)
+ opt_map … (multi_mode_write true m t s_tmp1 cur_pc i R_op)
(λS_PC.match S_PC with
[ pair s_tmp2 new_pc ⇒
(* C = carry *)
(* V = R7 ⊙ carry *)
let s_tmp6 ≝ setweak_v_flag m t s_tmp5 ((MSB_b8 R_op) ⊙ carry) in
(* newpc = nextpc *)
- Some ? (pair ?? s_tmp6 new_pc) ])]]).
+ Some ? (pair … s_tmp6 new_pc) ])]]).
(* estensione del segno byte → word *)
ndefinition byte_extension ≝
(* tutti i branch calcoleranno la condizione e la passeranno qui *)
ndefinition execute_any_BRANCH ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.λfCOND:bool.
- opt_map ?? (multi_mode_load true m t s cur_pc i)
+ opt_map … (multi_mode_load true m t s cur_pc i)
(λS_M_PC.match S_M_PC with
[ triple s_tmp1 M_op new_pc ⇒
(* if true, branch *)
match fCOND with
(* newpc = nextpc + rel *)
- [ true ⇒ Some ? (pair ?? s_tmp1 (branched_pc m t s_tmp1 new_pc M_op))
+ [ true ⇒ Some ? (pair … s_tmp1 (branched_pc m t s_tmp1 new_pc M_op))
(* newpc = nextpc *)
- | false ⇒ Some ? (pair ?? s_tmp1 new_pc) ]]).
+ | false ⇒ Some ? (pair … s_tmp1 new_pc) ]]).
(* Mn = filtered optval *)
(* il chiamante passa 0x00 per azzerare, 0xFF per impostare il bit di M *)
ndefinition execute_BCLRn_BSETn_aux ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.λoptval:byte8.
(* Mn = filtered optval *)
- opt_map ?? (multi_mode_write true m t s cur_pc i optval)
+ opt_map … (multi_mode_write true m t s cur_pc i optval)
(λS_PC.match S_PC with
(* newpc = nextpc *)
- [ pair s_tmp1 new_pc ⇒ Some ? (pair ?? s_tmp1 new_pc) ]).
+ [ pair s_tmp1 new_pc ⇒ Some ? (pair … s_tmp1 new_pc) ]).
(* if COND(Mn) branch *)
(* il chiamante passa la logica da testare (0x00,¬0x00) e poi si salta *)
ndefinition execute_BRCLRn_BRSETn_aux ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.λfCOND:byte8 → bool.
- opt_map ?? (multi_mode_load false m t s cur_pc i)
+ opt_map … (multi_mode_load false m t s cur_pc i)
(λS_M_PC.match S_M_PC with
[ triple s_tmp1 M_op new_pc ⇒ match M_op with
[ mk_word16 MH_op ML_op ⇒
(* if COND(Mn) branch *)
match fCOND MH_op with
(* newpc = nextpc + rel *)
- [ true ⇒ Some ? (pair ?? s_tmp1 (branched_pc m t s_tmp1 new_pc ML_op))
+ [ true ⇒ Some ? (pair … s_tmp1 (branched_pc m t s_tmp1 new_pc ML_op))
(* newpc = nextpc *)
- | false ⇒ Some ? (pair ?? s_tmp1 new_pc) ]]]).
+ | false ⇒ Some ? (pair … s_tmp1 new_pc) ]]]).
(* M = fM(M) *)
(* fM e' la logica da applicare: not/neg/++/-- *)
ndefinition execute_COM_DEC_INC_NEG_aux ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
λfM:byte8 → byte8.λfV:bool → bool → bool.λfC:bool → byte8 → bool.
- opt_map ?? (multi_mode_load true m t s cur_pc i)
+ opt_map … (multi_mode_load true m t s cur_pc i)
(λS_M_PC.match S_M_PC with
[ triple s_tmp1 M_op _ ⇒
let R_op ≝ fM M_op in
(* M = fM(M) *)
- opt_map ?? (multi_mode_write true m t s_tmp1 cur_pc i R_op)
+ opt_map … (multi_mode_write true m t s_tmp1 cur_pc i R_op)
(λS_PC.match S_PC with
[ pair s_tmp2 new_pc ⇒
(* C = fCR (C,R) *)
(* V = fV (M7,R7) *)
let s_tmp6 ≝ setweak_v_flag m t s_tmp5 (fV (MSB_b8 M_op) (MSB_b8 R_op)) in
(* newpc = nextpc *)
- Some ? (pair ?? s_tmp6 new_pc) ])]).
+ Some ? (pair … s_tmp6 new_pc) ])]).
(* A = [true] fAMC(A,M,C), [false] A *)
(* cioe' in caso di false l'operazione viene eseguita ma modifica solo i flag *)
ndefinition execute_SBC_SUB_aux ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.λsetflag:bool.
λfAMC:byte8 → byte8 → bool → ProdT byte8 bool.
- opt_map ?? (multi_mode_load true m t s cur_pc i)
+ opt_map … (multi_mode_load true m t s cur_pc i)
(λS_M_PC.match S_M_PC with
[ triple s_tmp1 M_op new_pc ⇒
let A_op ≝ get_acc_8_low_reg m t s_tmp1 in
(* V = A7&nM7&nR7 | nA7&M7&R7 *)
let s_tmp6 ≝ setweak_v_flag m t s_tmp5 ((A7⊗(⊖M7)⊗(⊖R7)) ⊕ ((⊖A7)⊗M7⊗R7)) in
(* newpc = nextpc *)
- Some ? (pair ?? s_tmp6 new_pc) ]]).
+ Some ? (pair … s_tmp6 new_pc) ]]).
(* il classico push *)
ndefinition aux_push ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λval:byte8.
- opt_map ?? (get_sp_reg m t s)
+ opt_map … (get_sp_reg m t s)
(* [SP] = val *)
- (λSP_op.opt_map ?? (memory_filter_write m t s SP_op val)
+ (λSP_op.opt_map … (memory_filter_write m t s SP_op val)
(* SP -- *)
- (λs_tmp1.opt_map ?? (set_sp_reg m t s_tmp1 (pred_w16 SP_op))
+ (λs_tmp1.opt_map … (set_sp_reg m t s_tmp1 (pred_w16 SP_op))
(λs_tmp2.Some ? s_tmp2))).
(* il classico pop *)
(* NB: l'incremento di SP deve essere filtrato dalla ALU, quindi get(set(SP)) *)
ndefinition aux_pop ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.
- opt_map ?? (get_sp_reg m t s)
+ opt_map … (get_sp_reg m t s)
(* SP ++ *)
- (λSP_op.opt_map ?? (set_sp_reg m t s (succ_w16 SP_op))
- (λs_tmp1.opt_map ?? (get_sp_reg m t s_tmp1)
+ (λSP_op.opt_map … (set_sp_reg m t s (succ_w16 SP_op))
+ (λs_tmp1.opt_map … (get_sp_reg m t s_tmp1)
(* val = [SP] *)
- (λSP_op'.opt_map ?? (memory_filter_read m t s_tmp1 SP_op')
- (λval.Some ? (pair ?? s_tmp1 val))))).
+ (λSP_op'.opt_map … (memory_filter_read m t s_tmp1 SP_op')
+ (λval.Some ? (pair … s_tmp1 val))))).
(* CCR corrisponde a V11HINZC e cmq 1 se un flag non esiste *)
(* i flag mantengono posizione costante nelle varie ALU, e se non sono
(* SP += extended M *)
ndefinition execute_AIS ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (multi_mode_load true m t s cur_pc i)
+ opt_map … (multi_mode_load true m t s cur_pc i)
(λS_M_PC.match S_M_PC with
[ triple s_tmp1 M_op new_pc ⇒
- opt_map ?? (get_sp_reg m t s_tmp1)
+ opt_map … (get_sp_reg m t s_tmp1)
(* SP += extended M *)
- (λSP_op.opt_map ?? (set_sp_reg m t s_tmp1 (plus_w16_d_d SP_op (byte_extension M_op)))
- (λs_tmp2.Some ? (pair ?? s_tmp2 new_pc))) ]).
+ (λSP_op.opt_map … (set_sp_reg m t s_tmp1 (plus_w16_d_d SP_op (byte_extension M_op)))
+ (λs_tmp2.Some ? (pair … s_tmp2 new_pc))) ]).
(* H:X += extended M *)
ndefinition execute_AIX ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (multi_mode_load true m t s cur_pc i)
+ opt_map … (multi_mode_load true m t s cur_pc i)
(λS_M_PC.match S_M_PC with
[ triple s_tmp1 M_op new_pc ⇒
- opt_map ?? (get_indX_16_reg m t s_tmp1)
+ opt_map … (get_indX_16_reg m t s_tmp1)
(* H:X += extended M *)
- (λHX_op.opt_map ?? (set_indX_16_reg m t s_tmp1 (plus_w16_d_d HX_op (byte_extension M_op)))
- (λs_tmp2.Some ? (pair ?? s_tmp2 new_pc))) ]).
+ (λHX_op.opt_map … (set_indX_16_reg m t s_tmp1 (plus_w16_d_d HX_op (byte_extension M_op)))
+ (λs_tmp2.Some ? (pair … s_tmp2 new_pc))) ]).
(* A = A & M *)
ndefinition execute_AND ≝
(* if N⊙V=0, branch *)
ndefinition execute_BGE ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_n_flag m t s)
- (λN_op.opt_map ?? (get_v_flag m t s)
+ opt_map … (get_n_flag m t s)
+ (λN_op.opt_map … (get_v_flag m t s)
(λV_op.execute_any_BRANCH m t s i cur_pc (⊖(N_op ⊙ V_op)))).
(* BGND mode *)
ndefinition execute_BGND ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- Some ? (pair ?? s cur_pc).
+ Some ? (pair … s cur_pc).
(* if Z|N⊙V=0, branch *)
ndefinition execute_BGT ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_n_flag m t s)
- (λN_op.opt_map ?? (get_v_flag m t s)
+ opt_map … (get_n_flag m t s)
+ (λN_op.opt_map … (get_v_flag m t s)
(λV_op.execute_any_BRANCH m t s i cur_pc (⊖((get_z_flag m t s) ⊕ (N_op ⊙ V_op))))).
(* if H=0, branch *)
ndefinition execute_BHCC ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_h_flag m t s)
+ opt_map … (get_h_flag m t s)
(λH_op.execute_any_BRANCH m t s i cur_pc (⊖H_op)).
(* if H=1, branch *)
ndefinition execute_BHCS ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_h_flag m t s)
+ opt_map … (get_h_flag m t s)
(λH_op.execute_any_BRANCH m t s i cur_pc H_op).
(* if C|Z=0, branch *)
(* if nIRQ=1, branch NB: irqflag e' un negato del pin *)
ndefinition execute_BIH ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_irq_flag m t s)
+ opt_map … (get_irq_flag m t s)
(λIRQ_op.execute_any_BRANCH m t s i cur_pc (⊖IRQ_op)).
(* if nIRQ=0, branch NB: irqflag e' un negato del pin *)
ndefinition execute_BIL ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_irq_flag m t s)
+ opt_map … (get_irq_flag m t s)
(λIRQ_op.execute_any_BRANCH m t s i cur_pc IRQ_op).
(* flags = A & M *)
(* if Z|N⊙V=1, branch *)
ndefinition execute_BLE ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_n_flag m t s)
- (λN_op.opt_map ?? (get_v_flag m t s)
+ opt_map … (get_n_flag m t s)
+ (λN_op.opt_map … (get_v_flag m t s)
(λV_op.execute_any_BRANCH m t s i cur_pc ((get_z_flag m t s) ⊕ (N_op ⊙ V_op)))).
(* if C|Z=1, branch *)
(* if N⊙V=1, branch *)
ndefinition execute_BLT ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_n_flag m t s)
- (λN_op.opt_map ?? (get_v_flag m t s)
+ opt_map … (get_n_flag m t s)
+ (λN_op.opt_map … (get_v_flag m t s)
(λV_op.execute_any_BRANCH m t s i cur_pc (N_op ⊙ V_op))).
(* if I=0, branch *)
ndefinition execute_BMC ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_i_flag m t s)
+ opt_map … (get_i_flag m t s)
(λI_op.execute_any_BRANCH m t s i cur_pc (⊖I_op)).
(* if N=1, branch *)
ndefinition execute_BMI ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_n_flag m t s)
+ opt_map … (get_n_flag m t s)
(λN_op.execute_any_BRANCH m t s i cur_pc N_op).
(* if I=1, branch *)
ndefinition execute_BMS ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_i_flag m t s)
+ opt_map … (get_i_flag m t s)
(λI_op.execute_any_BRANCH m t s i cur_pc I_op).
(* if Z=0, branch *)
(* if N=0, branch *)
ndefinition execute_BPL ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_n_flag m t s)
+ opt_map … (get_n_flag m t s)
(λN_op.execute_any_BRANCH m t s i cur_pc (⊖N_op)).
(* branch always *)
(* HC05/HC08/HCS08 si appoggiano allo stack, RS08 a SPC *)
ndefinition execute_BSR ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t .λi:instr_mode.λcur_pc:word16.
- opt_map ?? (multi_mode_load true m t s cur_pc i)
+ opt_map … (multi_mode_load true m t s cur_pc i)
(λS_M_PC.match S_M_PC with
[ triple s_tmp1 M_op new_pc ⇒ let aux ≝
(* push (new_pc low) *)
- opt_map ?? (aux_push m t s_tmp1 (w16l new_pc))
+ opt_map … (aux_push m t s_tmp1 (w16l new_pc))
(* push (new_pc high) *)
- (λs_tmp2.opt_map ?? (aux_push m t s_tmp2 (w16h new_pc))
+ (λs_tmp2.opt_map … (aux_push m t s_tmp2 (w16h new_pc))
(* new_pc = new_pc + rel *)
- (λs_tmp3.Some ? (pair ?? s_tmp3 (branched_pc m t s_tmp3 new_pc M_op))))
+ (λs_tmp3.Some ? (pair … s_tmp3 (branched_pc m t s_tmp3 new_pc M_op))))
in match m with
[ HC05 ⇒ aux | HC08 ⇒ aux | HCS08 ⇒ aux
| RS08 ⇒
(* SPC = new_pc *)
- opt_map ?? (set_spc_reg m t s_tmp1 new_pc)
+ opt_map … (set_spc_reg m t s_tmp1 new_pc)
(* new_pc = new_pc + rel *)
- (λs_tmp2.Some ? (pair ?? s_tmp2 (branched_pc m t s_tmp2 new_pc M_op)))
+ (λs_tmp2.Some ? (pair … s_tmp2 (branched_pc m t s_tmp2 new_pc M_op)))
]]).
(* if A=M, branch *)
ndefinition execute_CBEQA ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (multi_mode_load false m t s cur_pc i)
+ opt_map … (multi_mode_load false m t s cur_pc i)
(λS_M_PC.match S_M_PC with
[ triple s_tmp1 M_op new_pc ⇒
match M_op with
(* if A=M, branch *)
match eq_b8 (get_acc_8_low_reg m t s_tmp1) MH_op with
(* new_pc = new_pc + rel *)
- [ true ⇒ Some ? (pair ?? s_tmp1 (branched_pc m t s_tmp1 new_pc ML_op))
+ [ true ⇒ Some ? (pair … s_tmp1 (branched_pc m t s_tmp1 new_pc ML_op))
(* new_pc = new_pc *)
- | false ⇒ Some ? (pair ?? s_tmp1 new_pc)
+ | false ⇒ Some ? (pair … s_tmp1 new_pc)
]]]).
(* if X=M, branch *)
ndefinition execute_CBEQX ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (multi_mode_load false m t s cur_pc i)
+ opt_map … (multi_mode_load false m t s cur_pc i)
(λS_M_PC.match S_M_PC with
[ triple s_tmp1 M_op new_pc ⇒
match M_op with
[ mk_word16 MH_op ML_op ⇒
- opt_map ?? (get_indX_8_low_reg m t s_tmp1)
+ opt_map … (get_indX_8_low_reg m t s_tmp1)
(* if X=M, branch *)
(λX_op.match eq_b8 X_op MH_op with
(* new_pc = new_pc + rel *)
- [ true ⇒ Some ? (pair ?? s_tmp1 (branched_pc m t s_tmp1 new_pc ML_op))
+ [ true ⇒ Some ? (pair … s_tmp1 (branched_pc m t s_tmp1 new_pc ML_op))
(* new_pc = new_pc *)
- | false ⇒ Some ? (pair ?? s_tmp1 new_pc)
+ | false ⇒ Some ? (pair … s_tmp1 new_pc)
])]]).
(* C = 0 *)
ndefinition execute_CLC ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- Some ? (pair ?? (set_c_flag m t s false) cur_pc).
+ Some ? (pair … (set_c_flag m t s false) cur_pc).
(* I = 0 *)
ndefinition execute_CLI ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (set_i_flag m t s false)
- (λs_tmp.Some ? (pair ?? s_tmp cur_pc)).
+ opt_map … (set_i_flag m t s false)
+ (λs_tmp.Some ? (pair … s_tmp cur_pc)).
(* M = 0 *)
ndefinition execute_CLR ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
(* M = 0 *)
- opt_map ?? (multi_mode_write true m t s cur_pc i 〈x0,x0〉)
+ opt_map … (multi_mode_write true m t s cur_pc i 〈x0,x0〉)
(λS_PC.match S_PC with
[ pair s_tmp1 new_pc ⇒
(* Z = 1 *)
(* V = 0 *)
let s_tmp4 ≝ setweak_v_flag m t s_tmp3 false in
(* newpc = nextpc *)
- Some ? (pair ?? s_tmp4 new_pc) ]).
+ Some ? (pair … s_tmp4 new_pc) ]).
(* flags = A - M *)
ndefinition execute_CMP ≝
(* flags = H:X - M *)
ndefinition execute_CPHX ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (multi_mode_load false m t s cur_pc i)
+ opt_map … (multi_mode_load false m t s cur_pc i)
(λS_M_PC.match S_M_PC with
[ triple s_tmp1 M_op new_pc ⇒
- opt_map ?? (get_indX_16_reg m t s_tmp1)
+ opt_map … (get_indX_16_reg m t s_tmp1)
(λX_op.
match plus_w16_dc_dc X_op (compl_w16 M_op) false with
[ pair R_op carry ⇒
(* V = X15&nM15&nR15 | nX15&M15&R15 *)
let s_tmp5 ≝ setweak_v_flag m t s_tmp4 ((X15⊗(⊖M15)⊗(⊖R15)) ⊕ ((⊖X15)⊗M15⊗R15)) in
(* newpc = nextpc *)
- Some ? (pair ?? s_tmp5 new_pc) ] ) ]).
+ Some ? (pair … s_tmp5 new_pc) ] ) ]).
(* flags = X - M *)
ndefinition execute_CPX ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (multi_mode_load true m t s cur_pc i)
+ opt_map … (multi_mode_load true m t s cur_pc i)
(λS_M_PC.match S_M_PC with
[ triple s_tmp1 M_op new_pc ⇒
- opt_map ?? (get_indX_8_low_reg m t s_tmp1)
+ opt_map … (get_indX_8_low_reg m t s_tmp1)
(λX_op.
match plus_b8_dc_dc X_op (compl_b8 M_op) false with
[ pair R_op carry ⇒
(* V = X7&nM7&nR7 | nX7&M7&R7 *)
let s_tmp5 ≝ setweak_v_flag m t s_tmp4 ((X7⊗(⊖M7)⊗(⊖R7)) ⊕ ((⊖X7)⊗M7⊗R7)) in
(* newpc = nextpc *)
- Some ? (pair ?? s_tmp5 new_pc) ] ) ]).
+ Some ? (pair … s_tmp5 new_pc) ] ) ]).
(* decimal adjiust A *)
(* per i dettagli vedere daa_b8 (modulo byte8) *)
ndefinition execute_DAA ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_h_flag m t s)
+ opt_map … (get_h_flag m t s)
(λH.
let M_op ≝ get_acc_8_low_reg m t s in
match daa_b8 H (get_c_flag m t s) M_op with
(* V = M7 ⊙ R7 *)
let s_tmp5 ≝ setweak_v_flag m t s_tmp4 ((MSB_b8 M_op) ⊙ (MSB_b8 R_op)) in
(* newpc = curpc *)
- Some ? (pair ?? s_tmp5 cur_pc) ]).
+ Some ? (pair … s_tmp5 cur_pc) ]).
(* if (--M)≠0, branch *)
ndefinition execute_DBNZ ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (multi_mode_load false m t s cur_pc i)
+ opt_map … (multi_mode_load false m t s cur_pc i)
(λS_M_PC.match S_M_PC with
[ triple s_tmp1 M_op new_pc ⇒
match M_op with
[ mk_word16 MH_op ML_op ⇒
(* --M *)
let MH_op' ≝ pred_b8 MH_op in
- opt_map ?? (multi_mode_write true m t s_tmp1 cur_pc i MH_op')
+ opt_map … (multi_mode_write true m t s_tmp1 cur_pc i MH_op')
(λS_PC.match S_PC with
[ pair s_tmp2 _ ⇒
(* if (--M)≠0, branch *)
match eq_b8 MH_op' 〈x0,x0〉 with
(* new_pc = new_pc *)
- [ true ⇒ Some ? (pair ?? s_tmp2 new_pc)
+ [ true ⇒ Some ? (pair … s_tmp2 new_pc)
(* new_pc = new_pc + rel *)
- | false ⇒ Some ? (pair ?? s_tmp2 (branched_pc m t s_tmp2 new_pc ML_op)) ]])]]).
+ | false ⇒ Some ? (pair … s_tmp2 (branched_pc m t s_tmp2 new_pc ML_op)) ]])]]).
(* M = M - 1 *)
ndefinition execute_DEC ≝
(* per i dettagli vedere div_b8 (modulo word16) *)
ndefinition execute_DIV ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_indX_8_high_reg m t s)
- (λH_op.opt_map ?? (get_indX_8_low_reg m t s)
+ opt_map … (get_indX_8_high_reg m t s)
+ (λH_op.opt_map … (get_indX_8_low_reg m t s)
(λX_op.match div_b8 〈H_op:(get_acc_8_low_reg m t s)〉 X_op with
[ triple quoz rest overflow ⇒
(* C = overflow *)
(* NB: che A sia cambiato o no, lo testa *)
let s_tmp3 ≝ set_z_flag m t s_tmp2 (eq_b8 (get_acc_8_low_reg m t s_tmp2) 〈x0,x0〉) in
(* H = H o H:AmodX *)
- opt_map ?? (match overflow with
+ opt_map … (match overflow with
[ true ⇒ Some ? s_tmp3
| false ⇒ set_indX_8_high_reg m t s_tmp3 rest])
- (λs_tmp4.Some ? (pair ?? s_tmp4 cur_pc)) ])).
+ (λs_tmp4.Some ? (pair … s_tmp4 cur_pc)) ])).
(* A = A ⊙ M *)
ndefinition execute_EOR ≝
(* jmp, il nuovo indirizzo e' una WORD *)
ndefinition execute_JMP ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (multi_mode_load false m t s cur_pc i)
+ opt_map … (multi_mode_load false m t s cur_pc i)
(λS_M_PC.
(* newpc = M_op *)
- Some ? (pair ?? (fst3T ??? S_M_PC) (snd3T ??? S_M_PC))).
+ Some ? (pair … (fst3T ??? S_M_PC) (snd3T ??? S_M_PC))).
(* jump to subroutine *)
(* HC05/HC08/HCS08 si appoggiano allo stack, RS08 a SPC *)
ndefinition execute_JSR ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (multi_mode_load false m t s cur_pc i)
+ opt_map … (multi_mode_load false m t s cur_pc i)
(λS_M_PC.match S_M_PC with
[ triple s_tmp1 M_op new_pc ⇒ let aux ≝
(* push (new_pc low) *)
- opt_map ?? (aux_push m t s_tmp1 (w16l new_pc))
+ opt_map … (aux_push m t s_tmp1 (w16l new_pc))
(* push (new_pc high) *)
- (λs_tmp2.opt_map ?? (aux_push m t s_tmp2 (w16h new_pc))
+ (λs_tmp2.opt_map … (aux_push m t s_tmp2 (w16h new_pc))
(* newpc = M_op *)
- (λs_tmp3.Some ? (pair ?? s_tmp3 M_op)))
+ (λs_tmp3.Some ? (pair … s_tmp3 M_op)))
in match m with
[ HC05 ⇒ aux | HC08 ⇒ aux | HCS08 ⇒ aux
| RS08 ⇒
(* SPC = new_pc *)
- opt_map ?? (set_spc_reg m t s_tmp1 new_pc)
+ opt_map … (set_spc_reg m t s_tmp1 new_pc)
(* newpc = M_op *)
- (λs_tmp2.Some ? (pair ?? s_tmp2 M_op))
+ (λs_tmp2.Some ? (pair … s_tmp2 M_op))
]]).
(* A = M *)
ndefinition execute_LDA ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (multi_mode_load true m t s cur_pc i)
+ opt_map … (multi_mode_load true m t s cur_pc i)
(λS_M_PC.match S_M_PC with
[ triple s_tmp1 M_op new_pc ⇒
(* A = M *)
(* V = 0 *)
let s_tmp5 ≝ setweak_v_flag m t s_tmp4 false in
(* newpc = nextpc *)
- Some ? (pair ?? s_tmp5 new_pc) ]).
+ Some ? (pair … s_tmp5 new_pc) ]).
(* H:X = M *)
ndefinition execute_LDHX ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (multi_mode_load false m t s cur_pc i)
+ opt_map … (multi_mode_load false m t s cur_pc i)
(λS_M_PC.match S_M_PC with
[ triple s_tmp1 M_op new_pc ⇒
- opt_map ?? (set_indX_16_reg m t s_tmp1 M_op)
+ opt_map … (set_indX_16_reg m t s_tmp1 M_op)
(λs_tmp2.
(* Z = nR15&nR14&nR13nR12&nR11&nR10&nR9&nR8nR7&nR6&nR5&nR4&nR3&nR2&nR1&nR0 *)
let s_tmp3 ≝ set_z_flag m t s_tmp2 (eq_w16 M_op 〈〈x0,x0〉:〈x0,x0〉〉) in
(* V = 0 *)
let s_tmp5 ≝ setweak_v_flag m t s_tmp4 false in
(* newpc = nextpc *)
- Some ? (pair ?? s_tmp5 new_pc)) ]).
+ Some ? (pair … s_tmp5 new_pc)) ]).
(* X = M *)
ndefinition execute_LDX ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (multi_mode_load true m t s cur_pc i)
+ opt_map … (multi_mode_load true m t s cur_pc i)
(λS_M_PC.match S_M_PC with
[ triple s_tmp1 M_op new_pc ⇒
- opt_map ?? (set_indX_8_low_reg m t s_tmp1 M_op)
+ opt_map … (set_indX_8_low_reg m t s_tmp1 M_op)
(λs_tmp2.
(* Z = nR7&nR6&nR5&nR4&nR3&nR2&nR1&nR0 *)
let s_tmp3 ≝ set_z_flag m t s_tmp2 (eq_b8 M_op 〈x0,x0〉) in
(* V = 0 *)
let s_tmp5 ≝ setweak_v_flag m t s_tmp4 false in
(* newpc = nextpc *)
- Some ? (pair ?? s_tmp5 new_pc)) ]).
+ Some ? (pair … s_tmp5 new_pc)) ]).
(* M = 0 -> rcr M -> C' *)
ndefinition execute_LSR ≝
ndefinition execute_MOV ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
(* R_op = M1 *)
- opt_map ?? (multi_mode_load true m t s cur_pc i)
+ opt_map … (multi_mode_load true m t s cur_pc i)
(λS_R_PC.match S_R_PC with
[ triple s_tmp1 R_op tmp_pc ⇒
(* M2 = R_op *)
- opt_map ?? (multi_mode_write true m t s_tmp1 tmp_pc i R_op)
+ opt_map … (multi_mode_write true m t s_tmp1 tmp_pc i R_op)
(λS_PC.match S_PC with
[ pair s_tmp2 new_pc ⇒
(* Z = nR7&nR6&nR5&nR4&nR3&nR2&nR1&nR0 *)
(* V = 0 *)
let s_tmp5 ≝ setweak_v_flag m t s_tmp4 false in
(* newpc = nextpc *)
- Some ? (pair ?? s_tmp5 new_pc)])]).
+ Some ? (pair … s_tmp5 new_pc)])]).
(* X:A = X * A *)
ndefinition execute_MUL ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_indX_8_low_reg m t s)
+ opt_map … (get_indX_8_low_reg m t s)
(λX_op.let R_op ≝ mul_b8 X_op (get_acc_8_low_reg m t s) in
- opt_map ?? (set_indX_8_low_reg m t s (w16h R_op))
- (λs_tmp.Some ? (pair ?? (set_acc_8_low_reg m t s_tmp (w16l R_op)) cur_pc))).
+ opt_map … (set_indX_8_low_reg m t s (w16h R_op))
+ (λs_tmp.Some ? (pair … (set_acc_8_low_reg m t s_tmp (w16l R_op)) cur_pc))).
(* M = compl M *)
ndefinition execute_NEG ≝
(* nulla *)
ndefinition execute_NOP ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- Some ? (pair ?? s cur_pc).
+ Some ? (pair … s cur_pc).
(* A = (mk_byte8 (b8l A) (b8h A)) *)
(* cioe' swap del nibble alto/nibble basso di A *)
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
match get_acc_8_low_reg m t s with [ mk_byte8 ah al ⇒
(* A = (mk_byte8 (b8l A) (b8h A)) *)
- Some ? (pair ?? (set_acc_8_low_reg m t s 〈al,ah〉) cur_pc) ].
+ Some ? (pair … (set_acc_8_low_reg m t s 〈al,ah〉) cur_pc) ].
(* A = A | M *)
ndefinition execute_ORA ≝
(* push A *)
ndefinition execute_PSHA ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (aux_push m t s (get_acc_8_low_reg m t s))
- (λs_tmp1.Some ? (pair ?? s_tmp1 cur_pc)).
+ opt_map … (aux_push m t s (get_acc_8_low_reg m t s))
+ (λs_tmp1.Some ? (pair … s_tmp1 cur_pc)).
(* push H *)
ndefinition execute_PSHH ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_indX_8_high_reg m t s)
- (λH_op.opt_map ?? (aux_push m t s H_op)
- (λs_tmp1.Some ? (pair ?? s_tmp1 cur_pc))).
+ opt_map … (get_indX_8_high_reg m t s)
+ (λH_op.opt_map … (aux_push m t s H_op)
+ (λs_tmp1.Some ? (pair … s_tmp1 cur_pc))).
(* push X *)
ndefinition execute_PSHX ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_indX_8_low_reg m t s)
- (λH_op.opt_map ?? (aux_push m t s H_op)
- (λs_tmp1.Some ? (pair ?? s_tmp1 cur_pc))).
+ opt_map … (get_indX_8_low_reg m t s)
+ (λH_op.opt_map … (aux_push m t s H_op)
+ (λs_tmp1.Some ? (pair … s_tmp1 cur_pc))).
(* pop A *)
ndefinition execute_PULA ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (aux_pop m t s)
+ opt_map … (aux_pop m t s)
(λS_and_A.match S_and_A with [ pair s_tmp1 A_op ⇒
- Some ? (pair ?? (set_acc_8_low_reg m t s_tmp1 A_op) cur_pc) ]).
+ Some ? (pair … (set_acc_8_low_reg m t s_tmp1 A_op) cur_pc) ]).
(* pop H *)
ndefinition execute_PULH ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (aux_pop m t s)
+ opt_map … (aux_pop m t s)
(λS_and_H.match S_and_H with [ pair s_tmp1 H_op ⇒
- opt_map ?? (set_indX_8_high_reg m t s_tmp1 H_op)
- (λs_tmp2.Some ? (pair ?? s_tmp2 cur_pc))]).
+ opt_map … (set_indX_8_high_reg m t s_tmp1 H_op)
+ (λs_tmp2.Some ? (pair … s_tmp2 cur_pc))]).
(* pop X *)
ndefinition execute_PULX ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (aux_pop m t s)
+ opt_map … (aux_pop m t s)
(λS_and_X.match S_and_X with [ pair s_tmp1 X_op ⇒
- opt_map ?? (set_indX_8_low_reg m t s_tmp1 X_op)
- (λs_tmp2.Some ? (pair ?? s_tmp2 cur_pc))]).
+ opt_map … (set_indX_8_low_reg m t s_tmp1 X_op)
+ (λs_tmp2.Some ? (pair … s_tmp2 cur_pc))]).
(* M = C' <- rcl M <- C *)
ndefinition execute_ROL ≝
(* lascia inalterato il byte superiore di SP *)
ndefinition execute_RSP ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_sp_reg m t s)
+ opt_map … (get_sp_reg m t s)
(λSP_op.match SP_op with [ mk_word16 sph spl ⇒
- opt_map ?? (set_sp_reg m t s 〈sph:〈xF,xF〉〉)
- (λs_tmp.Some ? (pair ?? s_tmp cur_pc))]).
+ opt_map … (set_sp_reg m t s 〈sph:〈xF,xF〉〉)
+ (λs_tmp.Some ? (pair … s_tmp cur_pc))]).
(* return from interrupt *)
ndefinition execute_RTI ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
(* pop (CCR) *)
- opt_map ?? (aux_pop m t s)
+ opt_map … (aux_pop m t s)
(λS_and_CCR.match S_and_CCR with [ pair s_tmp1 CCR_op ⇒
let s_tmp2 ≝ aux_set_CCR m t s_tmp1 CCR_op in
(* pop (A) *)
- opt_map ?? (aux_pop m t s_tmp2)
+ opt_map … (aux_pop m t s_tmp2)
(λS_and_A.match S_and_A with [ pair s_tmp3 A_op ⇒
let s_tmp4 ≝ set_acc_8_low_reg m t s_tmp3 A_op in
(* pop (X) *)
- opt_map ?? (aux_pop m t s_tmp4)
+ opt_map … (aux_pop m t s_tmp4)
(λS_and_X.match S_and_X with [ pair s_tmp5 X_op ⇒
- opt_map ?? (set_indX_8_low_reg m t s_tmp5 X_op)
+ opt_map … (set_indX_8_low_reg m t s_tmp5 X_op)
(* pop (PC high) *)
- (λs_tmp6.opt_map ?? (aux_pop m t s_tmp6)
+ (λs_tmp6.opt_map … (aux_pop m t s_tmp6)
(λS_and_PCH.match S_and_PCH with [ pair s_tmp7 PCH_op ⇒
(* pop (PC low) *)
- opt_map ?? (aux_pop m t s_tmp7)
+ opt_map … (aux_pop m t s_tmp7)
(λS_and_PCL.match S_and_PCL with [ pair s_tmp8 PCL_op ⇒
- Some ? (pair ?? s_tmp8 〈PCH_op:PCL_op〉)])]))])])]).
+ Some ? (pair … s_tmp8 〈PCH_op:PCL_op〉)])]))])])]).
(* return from subroutine *)
(* HC05/HC08/HCS08 si appoggia allo stack, RS08 si appoggia a SPC *)
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
let aux ≝
(* pop (PC high) *)
- opt_map ?? (aux_pop m t s)
+ opt_map … (aux_pop m t s)
(λS_and_PCH.match S_and_PCH with [ pair s_tmp1 PCH_op ⇒
(* pop (PC low) *)
- opt_map ?? (aux_pop m t s_tmp1)
+ opt_map … (aux_pop m t s_tmp1)
(λS_and_PCL.match S_and_PCL with [ pair s_tmp2 PCL_op ⇒
- Some ? (pair ?? s_tmp2 〈PCH_op:PCL_op〉)])])
+ Some ? (pair … s_tmp2 〈PCH_op:PCL_op〉)])])
in match m with
[ HC05 ⇒ aux | HC08 ⇒ aux | HCS08 ⇒ aux
| RS08 ⇒
(* new_pc = SPC *)
- opt_map ?? (get_spc_reg m t s)
- (λSPC_op.Some ? (pair ?? s SPC_op))
+ opt_map … (get_spc_reg m t s)
+ (λSPC_op.Some ? (pair … s SPC_op))
].
(* A = A - M - C *)
(λA_op.λM_op.λC_op.match plus_b8_dc_dc A_op (compl_b8 M_op) false with
[ pair resb resc ⇒ match C_op with
[ true ⇒ plus_b8_dc_dc resb 〈xF,xF〉 false
- | false ⇒ pair ?? resb resc ]]).
+ | false ⇒ pair … resb resc ]]).
(* C = 1 *)
ndefinition execute_SEC ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- Some ? (pair ?? (set_c_flag m t s true) cur_pc).
+ Some ? (pair … (set_c_flag m t s true) cur_pc).
(* I = 1 *)
ndefinition execute_SEI ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (set_i_flag m t s true)
- (λs_tmp.Some ? (pair ?? s_tmp cur_pc)).
+ opt_map … (set_i_flag m t s true)
+ (λs_tmp.Some ? (pair … s_tmp cur_pc)).
(* swap SPCh,A *)
(* senso: nell'RS08 SPC non e' accessibile direttamente e come si possono
occore accedere a SPC e salvarne il contenuto *)
ndefinition execute_SHA ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_spc_reg m t s)
- (λSPC_op.opt_map ?? (set_spc_reg m t s 〈(get_acc_8_low_reg m t s):(w16l SPC_op)〉)
- (λs_tmp1.Some ? (pair ?? (set_acc_8_low_reg m t s_tmp1 (w16h SPC_op)) cur_pc))).
+ opt_map … (get_spc_reg m t s)
+ (λSPC_op.opt_map … (set_spc_reg m t s 〈(get_acc_8_low_reg m t s):(w16l SPC_op)〉)
+ (λs_tmp1.Some ? (pair … (set_acc_8_low_reg m t s_tmp1 (w16h SPC_op)) cur_pc))).
(* swap SPCl,A *)
(* senso: nell'RS08 SPC non e' accessibile direttamente e come si possono
occore accedere a SPC e salvarne il contenuto *)
ndefinition execute_SLA ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_spc_reg m t s)
- (λSPC_op.opt_map ?? (set_spc_reg m t s 〈(w16h SPC_op):(get_acc_8_low_reg m t s)〉)
- (λs_tmp1.Some ? (pair ?? (set_acc_8_low_reg m t s_tmp1 (w16l SPC_op)) cur_pc))).
+ opt_map … (get_spc_reg m t s)
+ (λSPC_op.opt_map … (set_spc_reg m t s 〈(w16h SPC_op):(get_acc_8_low_reg m t s)〉)
+ (λs_tmp1.Some ? (pair … (set_acc_8_low_reg m t s_tmp1 (w16l SPC_op)) cur_pc))).
(* M = A *)
ndefinition execute_STA ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
(* M = A *)
let A_op ≝ (get_acc_8_low_reg m t s) in
- opt_map ?? (multi_mode_write true m t s cur_pc i A_op)
+ opt_map … (multi_mode_write true m t s cur_pc i A_op)
(λS_op_and_PC.match S_op_and_PC with
[ pair s_tmp1 new_pc ⇒
(* Z = nA7&nA6&nA5&nA4&nA3&nA2&nA1&nA0 *)
(* V = 0 *)
let s_tmp4 ≝ setweak_v_flag m t s_tmp3 false in
(* newpc = nextpc *)
- Some ? (pair ?? s_tmp4 new_pc) ]).
+ Some ? (pair … s_tmp4 new_pc) ]).
(* M = H:X *)
ndefinition execute_STHX ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
(* M = H:X *)
- opt_map ?? (get_indX_16_reg m t s)
- (λX_op.opt_map ?? (multi_mode_write false m t s cur_pc i X_op)
+ opt_map … (get_indX_16_reg m t s)
+ (λX_op.opt_map … (multi_mode_write false m t s cur_pc i X_op)
(λS_op_and_PC.match S_op_and_PC with
[ pair s_tmp1 new_pc ⇒
(* Z = nR15&nR14&nR13nR12&nR11&nR10&nR9&nR8nR7&nR6&nR5&nR4&nR3&nR2&nR1&nR0 *)
(* V = 0 *)
let s_tmp4 ≝ setweak_v_flag m t s_tmp3 false in
(* newpc = nextpc *)
- Some ? (pair ?? s_tmp4 new_pc) ])).
+ Some ? (pair … s_tmp4 new_pc) ])).
(* I = 0 *)
ndefinition execute_STOP ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- Some ? (pair ?? (setweak_i_flag m t s false) cur_pc).
+ Some ? (pair … (setweak_i_flag m t s false) cur_pc).
(* M = X *)
ndefinition execute_STX ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
(* M = X *)
- opt_map ?? (get_indX_8_low_reg m t s)
- (λX_op.opt_map ?? (multi_mode_write true m t s cur_pc i X_op)
+ opt_map … (get_indX_8_low_reg m t s)
+ (λX_op.opt_map … (multi_mode_write true m t s cur_pc i X_op)
(λS_op_and_PC.match S_op_and_PC with
[ pair s_tmp1 new_pc ⇒
(* Z = nR7&nR6&nR5&nR4&nR3&nR2&nR1&nR0 *)
(* V = 0 *)
let s_tmp4 ≝ setweak_v_flag m t s_tmp3 false in
(* newpc = nextpc *)
- Some ? (pair ?? s_tmp4 new_pc) ])).
+ Some ? (pair … s_tmp4 new_pc) ])).
(* A = A - M *)
ndefinition execute_SUB ≝
(* indirizzo da cui caricare il nuovo pc *)
let vector ≝ get_pc_reg m t (set_pc_reg m t s 〈〈xF,xF〉:〈xF,xC〉〉) in
(* push (cur_pc low) *)
- opt_map ?? (aux_push m t s (w16l cur_pc))
+ opt_map … (aux_push m t s (w16l cur_pc))
(* push (cur_pc high *)
- (λs_tmp1.opt_map ?? (aux_push m t s_tmp1 (w16h cur_pc))
- (λs_tmp2.opt_map ?? (get_indX_8_low_reg m t s_tmp2)
+ (λs_tmp1.opt_map … (aux_push m t s_tmp1 (w16h cur_pc))
+ (λs_tmp2.opt_map … (get_indX_8_low_reg m t s_tmp2)
(* push (X) *)
- (λX_op.opt_map ?? (aux_push m t s_tmp2 X_op)
+ (λX_op.opt_map … (aux_push m t s_tmp2 X_op)
(* push (A) *)
- (λs_tmp3.opt_map ?? (aux_push m t s_tmp3 (get_acc_8_low_reg m t s_tmp3))
+ (λs_tmp3.opt_map … (aux_push m t s_tmp3 (get_acc_8_low_reg m t s_tmp3))
(* push (CCR) *)
- (λs_tmp4.opt_map ?? (aux_push m t s_tmp4 (aux_get_CCR m t s_tmp4))
+ (λs_tmp4.opt_map … (aux_push m t s_tmp4 (aux_get_CCR m t s_tmp4))
(* I = 1 *)
- (λs_tmp5.opt_map ?? (set_i_flag m t s_tmp5 true)
+ (λs_tmp5.opt_map … (set_i_flag m t s_tmp5 true)
(* load from vector high *)
- (λs_tmp6.opt_map ?? (memory_filter_read m t s_tmp6 vector)
+ (λs_tmp6.opt_map … (memory_filter_read m t s_tmp6 vector)
(* load from vector low *)
- (λaddrh.opt_map ?? (memory_filter_read m t s_tmp6 (succ_w16 vector))
+ (λaddrh.opt_map … (memory_filter_read m t s_tmp6 (succ_w16 vector))
(* newpc = [vector] *)
- (λaddrl.Some ? (pair ?? s_tmp6 〈addrh:addrl〉)))))))))).
+ (λaddrl.Some ? (pair … s_tmp6 〈addrh:addrl〉)))))))))).
(* flags = A *)
ndefinition execute_TAP ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- Some ? (pair ?? (aux_set_CCR m t s (get_acc_8_low_reg m t s)) cur_pc).
+ Some ? (pair … (aux_set_CCR m t s (get_acc_8_low_reg m t s)) cur_pc).
(* X = A *)
ndefinition execute_TAX ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (set_indX_8_low_reg m t s (get_acc_8_low_reg m t s))
- (λs_tmp.Some ? (pair ?? s_tmp cur_pc)).
+ opt_map … (set_indX_8_low_reg m t s (get_acc_8_low_reg m t s))
+ (λs_tmp.Some ? (pair … s_tmp cur_pc)).
(* A = flags *)
ndefinition execute_TPA ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- Some ? (pair ?? (set_acc_8_low_reg m t s (aux_get_CCR m t s)) cur_pc).
+ Some ? (pair … (set_acc_8_low_reg m t s (aux_get_CCR m t s)) cur_pc).
(* flags = M - 0 *)
(* implementata senza richiamare la sottrazione, la modifica dei flag
e' immediata *)
ndefinition execute_TST ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (multi_mode_load true m t s cur_pc i)
+ opt_map … (multi_mode_load true m t s cur_pc i)
(λS_M_PC.match S_M_PC with
[ triple s_tmp1 M_op new_pc ⇒
(* Z = nR7&nR6&nR5&nR4&nR3&nR2&nR1&nR0 *)
(* V = 0 *)
let s_tmp4 ≝ setweak_v_flag m t s_tmp3 false in
(* newpc = nextpc *)
- Some ? (pair ?? s_tmp4 new_pc) ]).
+ Some ? (pair … s_tmp4 new_pc) ]).
(* H:X = SP + 1 *)
ndefinition execute_TSX ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_sp_reg m t s )
- (λSP_op.opt_map ?? (set_indX_16_reg m t s (succ_w16 SP_op))
+ opt_map … (get_sp_reg m t s )
+ (λSP_op.opt_map … (set_indX_16_reg m t s (succ_w16 SP_op))
(* H:X = SP + 1 *)
- (λs_tmp.Some ? (pair ?? s_tmp cur_pc))).
+ (λs_tmp.Some ? (pair … s_tmp cur_pc))).
(* A = X *)
ndefinition execute_TXA ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_indX_8_low_reg m t s)
- (λX_op.Some ? (pair ?? (set_acc_8_low_reg m t s X_op) cur_pc)).
+ opt_map … (get_indX_8_low_reg m t s)
+ (λX_op.Some ? (pair … (set_acc_8_low_reg m t s X_op) cur_pc)).
(* SP = H:X - 1 *)
ndefinition execute_TXS ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- opt_map ?? (get_indX_16_reg m t s )
- (λX_op.opt_map ?? (set_sp_reg m t s (pred_w16 X_op))
+ opt_map … (get_indX_16_reg m t s )
+ (λX_op.opt_map … (set_sp_reg m t s (pred_w16 X_op))
(* SP = H:X - 1 *)
- (λs_tmp.Some ? (pair ?? s_tmp cur_pc))).
+ (λs_tmp.Some ? (pair … s_tmp cur_pc))).
(* I = 0 *)
ndefinition execute_WAIT ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
- Some ? (pair ?? (setweak_i_flag m t s false) cur_pc).
+ Some ? (pair … (setweak_i_flag m t s false) cur_pc).
(* **** *)
(* TICK *)
| WAIT_MODE: susp_type.
(* un tipo opzione ad hoc
- - errore: errore+stato (seguira' reset o ??, cmq lo stato non va buttato)
- - sospensione: sospensione+stato (seguira' resume o ??)
+ - errore: errore+stato (seguira' reset o …, cmq lo stato non va buttato)
+ - sospensione: sospensione+stato (seguira' resume o …)
- ok: stato
*)
ninductive tick_result (A:Type) : Type ≝
(* un solo clk, execute subito *)
[ true ⇒ tick_execute m t s pseudo mode cur_pc
(* piu' clk, execute rimandata *)
- | false ⇒ TickOK ? (set_clk_desc m t s (Some ? (quintuple ????? 〈x0,x1〉 pseudo mode tot_clk cur_pc)))
+ | false ⇒ TickOK ? (set_clk_desc m t s (Some ? (quintuple … 〈x0,x1〉 pseudo mode tot_clk cur_pc)))
]
]
]
(* si *)
[ true ⇒ tick_execute m t s pseudo mode cur_pc
(* no, avanzamento cur_clk *)
- | false ⇒ TickOK ? (set_clk_desc m t s (Some ? (quintuple ????? (succ_b8 cur_clk) pseudo mode tot_clk cur_pc)))
+ | false ⇒ TickOK ? (set_clk_desc m t s (Some ? (quintuple … (succ_b8 cur_clk) pseudo mode tot_clk cur_pc)))
]
]
].
ncases n;
##[ ##2: #n1 ##]
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma breakpoint_susp : ∀m,t,s,susp,n.execute m t (TickSUSP ? s susp) n = TickSUSP ? s susp.
ncases n;
##[ ##2: #n1 ##]
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma breakpoint :
∀m,t,n1,n2,s. execute m t s (n1 + n2) = execute m t (execute m t s n1) n2.
#m; #t; #n1;
nelim n1;
- ##[ ##1: nnormalize; #n2; #s; ncases s; nnormalize; ##[ ##1,2: #x ##] #y; napply (refl_eq ??)
+ ##[ ##1: nnormalize; #n2; #s; ncases s; nnormalize; ##[ ##1,2: #x ##] #y; napply refl_eq
##| ##2: #n3; #H; #n2; #s; ncases s;
- ##[ ##1: #x; #y; nnormalize; nrewrite > (breakpoint_err m t x y n2); napply (refl_eq ??)
- ##| ##2: #x; #y; nnormalize; nrewrite > (breakpoint_susp m t x y n2); napply (refl_eq ??)
+ ##[ ##1: #x; #y; nnormalize; nrewrite > (breakpoint_err m t x y n2); napply refl_eq
+ ##| ##2: #x; #y; nnormalize; nrewrite > (breakpoint_susp m t x y n2); napply refl_eq
##| ##3: #x; nrewrite > (Sn_p_n_to_S_npn n3 n2);
nchange with ((execute m t (tick m t x) (n3+n2)) =
(execute m t (execute m t (tick m t x) n3) n2));
nrewrite > (H n2 (tick m t x));
- napply (refl_eq ??)
+ napply refl_eq
##]
##]
nqed.
nchange with (match S n2 with [ O ⇒ False | S a ⇒ n1 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma nat_destruct_0_S : ∀n:nat.O = S n → False.
##[ ##1: #n2;
nelim n2;
nnormalize;
- ##[ ##1: napply (refl_eq ??)
- ##| ##2: #n3; #H; napply (refl_eq ??)
+ ##[ ##1: napply refl_eq
+ ##| ##2: #n3; #H; napply refl_eq
##]
##| ##2: #n2; #H; #n3;
nnormalize;
ncases n3;
nnormalize;
- ##[ ##1: napply (refl_eq ??)
+ ##[ ##1: napply refl_eq
##| ##2: #n4; napply (H n4)
##]
##]
##[ ##1: #n2;
nelim n2;
nnormalize;
- ##[ ##1: #H; napply (refl_eq ??)
+ ##[ ##1: #H; napply refl_eq
##| ##2: #n3; #H; #H1; nelim (nat_destruct_0_S ? H1)
##]
##| ##2: #n2; #H; #n3; #H1;
nnormalize;
##[ ##1: #H1; nelim (nat_destruct_S_0 ? H1)
##| ##2: #n4; #H1;
- napply (H n4 (nat_destruct_S_S ?? H1))
+ napply (H n4 (nat_destruct_S_S … H1))
##]
##]
nqed.
##[ ##1: #n2;
nelim n2;
nnormalize;
- ##[ ##1: #H; napply (refl_eq ??)
- ##| ##2: #n3; #H; #H1; napply (bool_destruct ?? (O = S n3) H1)
+ ##[ ##1: #H; napply refl_eq
+ ##| ##2: #n3; #H; #H1; napply (bool_destruct … (O = S n3) H1)
##]
##| ##2: #n2; #H; #n3; #H1;
ncases n3 in H1:(%) ⊢ %;
nnormalize;
- ##[ ##1: #H1; napply (bool_destruct ?? (S n2 = O) H1)
+ ##[ ##1: #H1; napply (bool_destruct … (S n2 = O) H1)
##| ##2: #n4; #H1;
nrewrite > (H n4 H1);
- napply (refl_eq ??)
+ napply refl_eq
##]
##]
nqed.
nlemma Sn_p_n_to_S_npn : ∀n1,n2.(S n1) + n2 = S (n1 + n2).
#n1;
nelim n1;
- ##[ ##1: nnormalize; #n2; napply (refl_eq ??)
+ ##[ ##1: nnormalize; #n2; napply refl_eq
##| ##2: #n; #H; #n2; nrewrite > (H n2);
ncases n in H:(%) ⊢ %;
- ##[ ##1: nnormalize; #H; napply (refl_eq ??)
- ##| ##2: #n3; nnormalize; #H; napply (refl_eq ??)
+ ##[ ##1: nnormalize; #H; napply refl_eq
+ ##| ##2: #n3; nnormalize; #H; napply refl_eq
##]
##]
nqed.
nlemma n_p_Sn_to_S_npn : ∀n1,n2.n1 + (S n2) = S (n1 + n2).
#n1;
nelim n1;
- ##[ ##1: nnormalize; #n2; napply (refl_eq ??)
+ ##[ ##1: nnormalize; #n2; napply refl_eq
##| ##2: #n; #H; #n2;
nrewrite > (Sn_p_n_to_S_npn n (S n2));
nrewrite > (H n2);
- napply (refl_eq ??)
+ napply refl_eq
##]
nqed.
nlemma Opn_to_n : ∀n.O + n = n.
- #n; nnormalize; napply (refl_eq ??).
+ #n; nnormalize; napply refl_eq.
nqed.
nlemma npO_to_n : ∀n.n + O = n.
#n;
nelim n;
- ##[ ##1: nnormalize; napply (refl_eq ??)
+ ##[ ##1: nnormalize; napply refl_eq
##| ##2: #n1; #H;
nrewrite > (Sn_p_n_to_S_npn n1 O);
nrewrite > H;
- napply (refl_eq ??)
+ napply refl_eq
##]
nqed.
nlemma symmetric_plusnat : symmetricT nat nat plus.
#n1;
nelim n1;
- ##[ ##1: #n2; nrewrite > (npO_to_n n2); nnormalize; napply (refl_eq ??)
+ ##[ ##1: #n2; nrewrite > (npO_to_n n2); nnormalize; napply refl_eq
##| ##2: #n2; #H; #n3;
nrewrite > (Sn_p_n_to_S_npn n2 n3);
nrewrite > (n_p_Sn_to_S_npn n3 n2);
nrewrite > (H n3);
- napply (refl_eq ??)
+ napply refl_eq
##]
nqed.
(l:list (Prod4T (any_opcode m) instr_mode byte8_or_word16 byte8)) on l ≝
match l with
[ nil ⇒ c
- | cons hd tl ⇒ match thd4T ???? hd with
+ | cons hd tl ⇒ match thd4T … hd with
[ Byte b' ⇒ match eq_b8 b b' with
[ true ⇒ get_byte_count m b (S c) tl
| false ⇒ get_byte_count m b c tl
(l:list (Prod4T (any_opcode m) instr_mode byte8_or_word16 byte8)) on l ≝
match l with
[ nil ⇒ c
- | cons hd tl ⇒ match thd4T ???? hd with
+ | cons hd tl ⇒ match thd4T … hd with
[ Byte _ ⇒ get_word_count m b c tl
| Word w ⇒ match eq_w16 〈〈x9,xE〉:b〉 w with
[ true ⇒ get_word_count m b (S c) tl
(l:list (Prod4T (any_opcode m) instr_mode byte8_or_word16 byte8)) on l ≝
match l with
[ nil ⇒ c
- | cons hd tl ⇒ match fst4T ???? hd with
+ | cons hd tl ⇒ match fst4T … hd with
[ anyOP o' ⇒ match eq_op o o' with
[ true ⇒ get_pseudo_count m o (S c) tl
| false ⇒ get_pseudo_count m o c tl
(l:list (Prod4T (any_opcode m) instr_mode byte8_or_word16 byte8)) on l ≝
match l with
[ nil ⇒ c
- | cons hd tl ⇒ match eq_instrmode (snd4T ???? hd) i with
+ | cons hd tl ⇒ match eq_instrmode (snd4T … hd) i with
[ true ⇒ get_mode_count m i (S c) tl
| false ⇒ get_mode_count m i c tl
]
match l with
[ nil ⇒ c
| cons hd tl ⇒
- match (eq_anyop m o (fst4T ???? hd)) ⊗
- (eq_instrmode i (snd4T ???? hd)) with
+ match (eq_anyop m o (fst4T … hd)) ⊗
+ (eq_instrmode i (snd4T … hd)) with
[ true ⇒ get_OpIm_count m o i (S c) tl
| false ⇒ get_OpIm_count m o i c tl
]
nelim m1;
nelim m2;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma eqmcutype_to_eq : ∀m1,m2:mcu_type.(eq_mcutype m1 m2 = true) → (m1 = m2).
ncases m1;
ncases m2;
nnormalize;
- ##[ ##1,6,11,16: #H; napply (refl_eq ??)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##1,6,11,16: #H; napply refl_eq
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
ncases m1;
ncases m2;
nnormalize;
- ##[ ##1,6,11,16: #H; napply (refl_eq ??)
- ##| ##*: #H; napply (mcu_type_destruct ??? H)
+ ##[ ##1,6,11,16: #H; napply refl_eq
+ ##| ##*: #H; napply (mcu_type_destruct … H)
##]
nqed.
nchange with (match MODE_DIRn n2 with [ MODE_DIRn a ⇒ n1 = a | _ ⇒ False ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma instr_mode_destruct_MODE_DIRn_and_IMM1 : ∀n1,n2.MODE_DIRn_and_IMM1 n1 = MODE_DIRn_and_IMM1 n2 → n1 = n2.
nchange with (match MODE_DIRn_and_IMM1 n2 with [ MODE_DIRn_and_IMM1 a ⇒ n1 = a | _ ⇒ False ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma instr_mode_destruct_MODE_TNY : ∀e1,e2.MODE_TNY e1 = MODE_TNY e2 → e1 = e2.
nchange with (match MODE_TNY e2 with [ MODE_TNY a ⇒ e1 = a | _ ⇒ False ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma instr_mode_destruct_MODE_SRT : ∀t1,t2.MODE_SRT t1 = MODE_SRT t2 → t1 = t2.
nchange with (match MODE_SRT t2 with [ MODE_SRT a ⇒ t1 = a | _ ⇒ False ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nchange with (match anyOP m x2 with [ anyOP a ⇒ x1 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_eqanyop : ∀m.∀op1,op2:any_opcode m.eq_anyop m op1 op2 = eq_anyop m op2 op1.
#x2;
nchange with (eq_op x1 x2 = eq_op x2 x1);
nrewrite > (symmetric_eqop x1 x2);
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma eqanyop_to_eq : ∀m.∀op1,op2:any_opcode m.eq_anyop m op1 op2 = true → op1 = op2.
#x2;
nchange with ((eq_op x1 x2 = true) → ?);
#H;
- nrewrite > (eqop_to_eq ?? H);
- napply (refl_eq ??).
+ nrewrite > (eqop_to_eq … H);
+ napply refl_eq.
nqed.
nlemma eq_to_eqanyop : ∀m.∀op1,op2:any_opcode m.op1 = op2 → eq_anyop m op1 op2 = true.
#p1;
ncases op2;
#p2; #H;
- nrewrite > (anyop_destruct ??? H);
+ nrewrite > (anyop_destruct … H);
nchange with (eq_op p2 p2 = true);
nrewrite > (eq_to_eqop p2 p2 (refl_eq opcode p2));
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma b8w16_destruct_b8_b8 : ∀x1,x2.Byte x1 = Byte x2 → x1 = x2.
nchange with (match Byte x2 with [ Byte a ⇒ x1 = a | Word _ ⇒ False ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma b8w16_destruct_w16_w16 : ∀x1,x2.Word x1 = Word x2 → x1 = x2.
nchange with (match Word x2 with [ Word a ⇒ x1 = a | Byte _ ⇒ False ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma b8w16_destruct_b8_w16 : ∀x1,x2.Byte x1 = Word x2 → False.
#x2;
##[ ##1: nchange with (eq_b8 x1 x2 = eq_b8 x2 x1);
nrewrite > (symmetric_eqb8 x1 x2);
- napply (refl_eq ??)
- ##| ##2,3: nnormalize; napply (refl_eq ??)
+ napply refl_eq
+ ##| ##2,3: nnormalize; napply refl_eq
##| ##4: nchange with (eq_w16 x1 x2 = eq_w16 x2 x1);
nrewrite > (symmetric_eqw16 x1 x2);
- napply (refl_eq ??)
+ napply refl_eq
##]
nqed.
nchange with (match MODE_DIRn n2 with [ MODE_DIRn a ⇒ n1 = a | _ ⇒ False ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
ndefinition instr_mode_destruct31 :
ncases n2;
nnormalize;
##[ ##1,10,19,28,37,46,55,64: #H; napply (λx:P.x)
- ##| ##*: #H; napply (oct_destruct ??? (instr_mode_destruct_MODE_DIRn ?? H))
+ ##| ##*: #H; napply (oct_destruct … (instr_mode_destruct_MODE_DIRn … H))
##]
nqed.
nchange with (match MODE_DIRn_and_IMM1 n2 with [ MODE_DIRn_and_IMM1 a ⇒ n1 = a | _ ⇒ False ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
ndefinition instr_mode_destruct32 :
ncases n2;
nnormalize;
##[ ##1,10,19,28,37,46,55,64: #H; napply (λx:P.x)
- ##| ##*: #H; napply (oct_destruct ??? (instr_mode_destruct_MODE_DIRn_and_IMM1 ?? H))
+ ##| ##*: #H; napply (oct_destruct … (instr_mode_destruct_MODE_DIRn_and_IMM1 … H))
##]
nqed.
nchange with (match MODE_TNY n2 with [ MODE_TNY a ⇒ n1 = a | _ ⇒ False ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
ndefinition instr_mode_destruct33 :
ncases n2;
nnormalize;
##[ ##1,18,35,52,69,86,103,120,137,154,171,188,205,222,239,256: #H; napply (λx:P.x)
- ##| ##*: #H; napply (exadecim_destruct ??? (instr_mode_destruct_MODE_TNY ?? H))
+ ##| ##*: #H; napply (exadecim_destruct … (instr_mode_destruct_MODE_TNY … H))
##]
nqed.
nchange with (match MODE_SRT n2 with [ MODE_SRT a ⇒ n1 = a | _ ⇒ False ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
ndefinition instr_mode_destruct34 :
ncases n1;
##[ ##1: ncases n2; nnormalize;
##[ ##1: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##2: ncases n2; nnormalize;
##[ ##2: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##3: ncases n2; nnormalize;
##[ ##3: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##4: ncases n2; nnormalize;
##[ ##4: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##5: ncases n2; nnormalize;
##[ ##5: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##6: ncases n2; nnormalize;
##[ ##6: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##7: ncases n2; nnormalize;
##[ ##7: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##8: ncases n2; nnormalize;
##[ ##8: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##9: ncases n2; nnormalize;
##[ ##9: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##10: ncases n2; nnormalize;
##[ ##10: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##11: ncases n2; nnormalize;
##[ ##11: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##12: ncases n2; nnormalize;
##[ ##12: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##13: ncases n2; nnormalize;
##[ ##13: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##14: ncases n2; nnormalize;
##[ ##14: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##15: ncases n2; nnormalize;
##[ ##15: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##16: ncases n2; nnormalize;
##[ ##16: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##17: ncases n2; nnormalize;
##[ ##17: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##18: ncases n2; nnormalize;
##[ ##18: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##19: ncases n2; nnormalize;
##[ ##19: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##20: ncases n2; nnormalize;
##[ ##20: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##21: ncases n2; nnormalize;
##[ ##21: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##22: ncases n2; nnormalize;
##[ ##22: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##23: ncases n2; nnormalize;
##[ ##23: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##24: ncases n2; nnormalize;
##[ ##24: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##25: ncases n2; nnormalize;
##[ ##25: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##26: ncases n2; nnormalize;
##[ ##26: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##27: ncases n2; nnormalize;
##[ ##27: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##28: ncases n2; nnormalize;
##[ ##28: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##29: ncases n2; nnormalize;
##[ ##29: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##30: ncases n2; nnormalize;
##[ ##30: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##31: ncases n2; nnormalize;
##[ ##31: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##| ##32: ncases n2; nnormalize;
##[ ##32: #H; napply (λx:P.x)
- ##| ##*: #H; napply (bitrigesim_destruct ??? (instr_mode_destruct_MODE_SRT ?? H))
+ ##| ##*: #H; napply (bitrigesim_destruct … (instr_mode_destruct_MODE_SRT … H))
##]
##]
##]
nlemma symmetric_eqinstrmode : symmetricT instr_mode bool eq_instrmode.
#i1; #i2;
ncases i1;
- ##[ ##1: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##2: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##3: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##4: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##5: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##6: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##7: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##8: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##9: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##10: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##11: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##12: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##13: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##14: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##15: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##16: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##17: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##18: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##19: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##20: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##21: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##22: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##23: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##24: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##25: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##26: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##27: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##28: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##29: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
- ##| ##30: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply (refl_eq ??)
+ ##[ ##1: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##2: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##3: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##4: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##5: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##6: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##7: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##8: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##9: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##10: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##11: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##12: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##13: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##14: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##15: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##16: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##17: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##18: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##19: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##20: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##21: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##22: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##23: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##24: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##25: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##26: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##27: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##28: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##29: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
+ ##| ##30: ncases i2; nnormalize; ##[ ##31,32,33,34: #n ##] napply refl_eq
##| ##31: ncases i2; #n1;
##[ ##31: #n2;
nchange with (eq_oct n2 n1 = eq_oct n1 n2);
nrewrite > (symmetric_eqoct n1 n2);
##| ##32,33,34: #n2; nnormalize
##]
- nnormalize; napply (refl_eq ??)
+ nnormalize; napply refl_eq
##| ##32: ncases i2; #n1;
##[ ##32: #n2;
nchange with (eq_oct n2 n1 = eq_oct n1 n2);
nrewrite > (symmetric_eqoct n1 n2);
##| ##31,33,34: #n2; nnormalize
##]
- nnormalize; napply (refl_eq ??)
+ nnormalize; napply refl_eq
##| ##33: ncases i2; #n1;
##[ ##33: #n2;
nchange with (eq_ex n2 n1 = eq_ex n1 n2);
nrewrite > (symmetric_eqex n1 n2);
##| ##31,32,34: #n2; nnormalize
##]
- nnormalize; napply (refl_eq ??)
+ nnormalize; napply refl_eq
##| ##34: ncases i2; #n1;
##[ ##34: #n2;
nchange with (eq_bitrig n2 n1 = eq_bitrig n1 n2);
nrewrite > (symmetric_eqbitrig n1 n2);
##| ##31,32,33: #n2; nnormalize
##]
- nnormalize; napply (refl_eq ??)
+ nnormalize; napply refl_eq
##]
nqed.
nlemma eqinstrmode_to_eq1 : ∀i2.eq_instrmode MODE_INH i2 = true → MODE_INH = i2.
#i2; ncases i2; nnormalize;
- ##[ ##1: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##1: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq2 : ∀i2.eq_instrmode MODE_INHA i2 = true → MODE_INHA = i2.
#i2; ncases i2; nnormalize;
- ##[ ##2: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##2: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq3 : ∀i2.eq_instrmode MODE_INHX i2 = true → MODE_INHX = i2.
#i2; ncases i2; nnormalize;
- ##[ ##3: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##3: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq4 : ∀i2.eq_instrmode MODE_INHH i2 = true → MODE_INHH = i2.
#i2; ncases i2; nnormalize;
- ##[ ##4: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##4: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq5 : ∀i2.eq_instrmode MODE_INHX0ADD i2 = true → MODE_INHX0ADD = i2.
#i2; ncases i2; nnormalize;
- ##[ ##5: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##5: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq6 : ∀i2.eq_instrmode MODE_INHX1ADD i2 = true → MODE_INHX1ADD = i2.
#i2; ncases i2; nnormalize;
- ##[ ##6: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##6: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq7 : ∀i2.eq_instrmode MODE_INHX2ADD i2 = true → MODE_INHX2ADD = i2.
#i2; ncases i2; nnormalize;
- ##[ ##7: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##7: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq8 : ∀i2.eq_instrmode MODE_IMM1 i2 = true → MODE_IMM1 = i2.
#i2; ncases i2; nnormalize;
- ##[ ##8: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##8: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq9 : ∀i2.eq_instrmode MODE_IMM1EXT i2 = true → MODE_IMM1EXT = i2.
#i2; ncases i2; nnormalize;
- ##[ ##9: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##9: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq10 : ∀i2.eq_instrmode MODE_IMM2 i2 = true → MODE_IMM2 = i2.
#i2; ncases i2; nnormalize;
- ##[ ##10: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##10: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq11 : ∀i2.eq_instrmode MODE_DIR1 i2 = true → MODE_DIR1 = i2.
#i2; ncases i2; nnormalize;
- ##[ ##11: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##11: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq12 : ∀i2.eq_instrmode MODE_DIR2 i2 = true → MODE_DIR2 = i2.
#i2; ncases i2; nnormalize;
- ##[ ##12: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##12: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq13 : ∀i2.eq_instrmode MODE_IX0 i2 = true → MODE_IX0 = i2.
#i2; ncases i2; nnormalize;
- ##[ ##13: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##13: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq14 : ∀i2.eq_instrmode MODE_IX1 i2 = true → MODE_IX1 = i2.
#i2; ncases i2; nnormalize;
- ##[ ##14: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##14: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq15 : ∀i2.eq_instrmode MODE_IX2 i2 = true → MODE_IX2 = i2.
#i2; ncases i2; nnormalize;
- ##[ ##15: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##15: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq16 : ∀i2.eq_instrmode MODE_SP1 i2 = true → MODE_SP1 = i2.
#i2; ncases i2; nnormalize;
- ##[ ##16: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##16: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq17 : ∀i2.eq_instrmode MODE_SP2 i2 = true → MODE_SP2 = i2.
#i2; ncases i2; nnormalize;
- ##[ ##17: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##17: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq18 : ∀i2.eq_instrmode MODE_DIR1_to_DIR1 i2 = true → MODE_DIR1_to_DIR1 = i2.
#i2; ncases i2; nnormalize;
- ##[ ##18: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##18: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq19 : ∀i2.eq_instrmode MODE_IMM1_to_DIR1 i2 = true → MODE_IMM1_to_DIR1 = i2.
#i2; ncases i2; nnormalize;
- ##[ ##19: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##19: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq20 : ∀i2.eq_instrmode MODE_IX0p_to_DIR1 i2 = true → MODE_IX0p_to_DIR1 = i2.
#i2; ncases i2; nnormalize;
- ##[ ##20: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##20: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq21 : ∀i2.eq_instrmode MODE_DIR1_to_IX0p i2 = true → MODE_DIR1_to_IX0p = i2.
#i2; ncases i2; nnormalize;
- ##[ ##21: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##21: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq22 : ∀i2.eq_instrmode MODE_INHA_and_IMM1 i2 = true → MODE_INHA_and_IMM1 = i2.
#i2; ncases i2; nnormalize;
- ##[ ##22: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##22: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq23 : ∀i2.eq_instrmode MODE_INHX_and_IMM1 i2 = true → MODE_INHX_and_IMM1 = i2.
#i2; ncases i2; nnormalize;
- ##[ ##23: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##23: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq24 : ∀i2.eq_instrmode MODE_IMM1_and_IMM1 i2 = true → MODE_IMM1_and_IMM1 = i2.
#i2; ncases i2; nnormalize;
- ##[ ##24: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##24: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq25 : ∀i2.eq_instrmode MODE_DIR1_and_IMM1 i2 = true → MODE_DIR1_and_IMM1 = i2.
#i2; ncases i2; nnormalize;
- ##[ ##25: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##25: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq26 : ∀i2.eq_instrmode MODE_IX0_and_IMM1 i2 = true → MODE_IX0_and_IMM1 = i2.
#i2; ncases i2; nnormalize;
- ##[ ##26: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##26: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq27 : ∀i2.eq_instrmode MODE_IX0p_and_IMM1 i2 = true → MODE_IX0p_and_IMM1 = i2.
#i2; ncases i2; nnormalize;
- ##[ ##27: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##27: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq28 : ∀i2.eq_instrmode MODE_IX1_and_IMM1 i2 = true → MODE_IX1_and_IMM1 = i2.
#i2; ncases i2; nnormalize;
- ##[ ##28: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##28: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq29 : ∀i2.eq_instrmode MODE_IX1p_and_IMM1 i2 = true → MODE_IX1p_and_IMM1 = i2.
#i2; ncases i2; nnormalize;
- ##[ ##29: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##29: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
nlemma eqinstrmode_to_eq30 : ∀i2.eq_instrmode MODE_SP1_and_IMM1 i2 = true → MODE_SP1_and_IMM1 = i2.
#i2; ncases i2; nnormalize;
- ##[ ##30: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (bool_destruct ??? H)
- ##| ##*: #H; napply (bool_destruct ??? H)
+ ##[ ##30: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (bool_destruct … H)
+ ##| ##*: #H; napply (bool_destruct … H)
##]
nqed.
#n1; #i2; ncases i2;
##[ ##31: #n2; #H;
nchange in H:(%) with (eq_oct n1 n2 = true);
- nrewrite > (eqoct_to_eq ?? H);
- napply (refl_eq ??)
- ##| ##32,33,34: nnormalize; #n2; #H; napply (bool_destruct ??? H)
- ##| ##*: nnormalize; #H; napply (bool_destruct ??? H)
+ nrewrite > (eqoct_to_eq … H);
+ napply refl_eq
+ ##| ##32,33,34: nnormalize; #n2; #H; napply (bool_destruct … H)
+ ##| ##*: nnormalize; #H; napply (bool_destruct … H)
##]
nqed.
#n1; #i2; ncases i2;
##[ ##32: #n2; #H;
nchange in H:(%) with (eq_oct n1 n2 = true);
- nrewrite > (eqoct_to_eq ?? H);
- napply (refl_eq ??)
- ##| ##31,33,34: nnormalize; #n2; #H; napply (bool_destruct ??? H)
- ##| ##*: nnormalize; #H; napply (bool_destruct ??? H)
+ nrewrite > (eqoct_to_eq … H);
+ napply refl_eq
+ ##| ##31,33,34: nnormalize; #n2; #H; napply (bool_destruct … H)
+ ##| ##*: nnormalize; #H; napply (bool_destruct … H)
##]
nqed.
#n1; #i2; ncases i2;
##[ ##33: #n2; #H;
nchange in H:(%) with (eq_ex n1 n2 = true);
- nrewrite > (eqex_to_eq ?? H);
- napply (refl_eq ??)
- ##| ##31,32,34: nnormalize; #n2; #H; napply (bool_destruct ??? H)
- ##| ##*: nnormalize; #H; napply (bool_destruct ??? H)
+ nrewrite > (eqex_to_eq … H);
+ napply refl_eq
+ ##| ##31,32,34: nnormalize; #n2; #H; napply (bool_destruct … H)
+ ##| ##*: nnormalize; #H; napply (bool_destruct … H)
##]
nqed.
#n1; #i2; ncases i2;
##[ ##34: #n2; #H;
nchange in H:(%) with (eq_bitrig n1 n2 = true);
- nrewrite > (eqbitrig_to_eq ?? H);
- napply (refl_eq ??)
- ##| ##31,32,33: nnormalize; #n2; #H; napply (bool_destruct ??? H)
- ##| ##*: nnormalize; #H; napply (bool_destruct ??? H)
+ nrewrite > (eqbitrig_to_eq … H);
+ napply refl_eq
+ ##| ##31,32,33: nnormalize; #n2; #H; napply (bool_destruct … H)
+ ##| ##*: nnormalize; #H; napply (bool_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode1 : ∀i2.MODE_INH = i2 → eq_instrmode MODE_INH i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##1: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##1: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode2 : ∀i2.MODE_INHA = i2 → eq_instrmode MODE_INHA i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##2: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##2: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode3 : ∀i2.MODE_INHX = i2 → eq_instrmode MODE_INHX i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##3: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##3: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode4 : ∀i2.MODE_INHH = i2 → eq_instrmode MODE_INHH i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##4: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##4: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode5 : ∀i2.MODE_INHX0ADD = i2 → eq_instrmode MODE_INHX0ADD i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##5: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##5: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode6 : ∀i2.MODE_INHX1ADD = i2 → eq_instrmode MODE_INHX1ADD i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##6: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##6: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode7 : ∀i2.MODE_INHX2ADD = i2 → eq_instrmode MODE_INHX2ADD i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##7: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##7: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode8 : ∀i2.MODE_IMM1 = i2 → eq_instrmode MODE_IMM1 i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##8: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##8: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode9 : ∀i2.MODE_IMM1EXT = i2 → eq_instrmode MODE_IMM1EXT i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##9: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##9: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode10 : ∀i2.MODE_IMM2 = i2 → eq_instrmode MODE_IMM2 i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##10: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##10: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode11 : ∀i2.MODE_DIR1 = i2 → eq_instrmode MODE_DIR1 i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##11: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##11: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode12 : ∀i2.MODE_DIR2 = i2 → eq_instrmode MODE_DIR2 i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##12: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##12: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode13 : ∀i2.MODE_IX0 = i2 → eq_instrmode MODE_IX0 i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##13: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##13: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode14 : ∀i2.MODE_IX1 = i2 → eq_instrmode MODE_IX1 i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##14: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##14: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode15 : ∀i2.MODE_IX2 = i2 → eq_instrmode MODE_IX2 i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##15: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##15: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode16 : ∀i2.MODE_SP1 = i2 → eq_instrmode MODE_SP1 i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##16: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##16: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode17 : ∀i2.MODE_SP2 = i2 → eq_instrmode MODE_SP2 i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##17: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##17: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode18 : ∀i2.MODE_DIR1_to_DIR1 = i2 → eq_instrmode MODE_DIR1_to_DIR1 i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##18: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##18: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode19 : ∀i2.MODE_IMM1_to_DIR1 = i2 → eq_instrmode MODE_IMM1_to_DIR1 i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##19: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##19: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode20 : ∀i2.MODE_IX0p_to_DIR1 = i2 → eq_instrmode MODE_IX0p_to_DIR1 i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##20: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##20: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode21 : ∀i2.MODE_DIR1_to_IX0p = i2 → eq_instrmode MODE_DIR1_to_IX0p i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##21: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##21: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode22 : ∀i2.MODE_INHA_and_IMM1 = i2 → eq_instrmode MODE_INHA_and_IMM1 i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##22: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##22: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode23 : ∀i2.MODE_INHX_and_IMM1 = i2 → eq_instrmode MODE_INHX_and_IMM1 i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##23: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##23: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode24 : ∀i2.MODE_IMM1_and_IMM1 = i2 → eq_instrmode MODE_IMM1_and_IMM1 i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##24: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##24: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode25 : ∀i2.MODE_DIR1_and_IMM1 = i2 → eq_instrmode MODE_DIR1_and_IMM1 i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##25: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##25: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode26 : ∀i2.MODE_IX0_and_IMM1 = i2 → eq_instrmode MODE_IX0_and_IMM1 i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##26: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##26: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode27 : ∀i2.MODE_IX0p_and_IMM1 = i2 → eq_instrmode MODE_IX0p_and_IMM1 i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##27: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##27: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode28 : ∀i2.MODE_IX1_and_IMM1 = i2 → eq_instrmode MODE_IX1_and_IMM1 i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##28: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##28: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode29 : ∀i2.MODE_IX1p_and_IMM1 = i2 → eq_instrmode MODE_IX1p_and_IMM1 i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##29: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##29: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
nlemma eq_to_eqinstrmode30 : ∀i2.MODE_SP1_and_IMM1 = i2 → eq_instrmode MODE_SP1_and_IMM1 i2 = true.
#t2; ncases t2; nnormalize;
- ##[ ##30: #H; napply (refl_eq ??)
- ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; napply (instr_mode_destruct ??? H)
+ ##[ ##30: #H; napply refl_eq
+ ##| ##31,32,33,34: #n; #H; napply (instr_mode_destruct … H)
+ ##| ##*: #H; napply (instr_mode_destruct … H)
##]
nqed.
#n1; #t2; ncases t2;
##[ ##31: #n2; #H;
nchange with (eq_oct n1 n2 = true);
- nrewrite > (instr_mode_destruct_MODE_DIRn ?? H);
- nrewrite > (eq_to_eqoct n2 n2 (refl_eq ??));
- napply (refl_eq ??)
- ##| ##32,33,34: #n; #H; nnormalize; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; nnormalize; napply (instr_mode_destruct ??? H)
+ nrewrite > (instr_mode_destruct_MODE_DIRn … H);
+ nrewrite > (eq_to_eqoct n2 n2 (refl_eq …));
+ napply refl_eq
+ ##| ##32,33,34: #n; #H; nnormalize; napply (instr_mode_destruct … H)
+ ##| ##*: #H; nnormalize; napply (instr_mode_destruct … H)
##]
nqed.
#n1; #t2; ncases t2;
##[ ##32: #n2; #H;
nchange with (eq_oct n1 n2 = true);
- nrewrite > (instr_mode_destruct_MODE_DIRn_and_IMM1 ?? H);
- nrewrite > (eq_to_eqoct n2 n2 (refl_eq ??));
- napply (refl_eq ??)
- ##| ##31,33,34: #n; #H; nnormalize; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; nnormalize; napply (instr_mode_destruct ??? H)
+ nrewrite > (instr_mode_destruct_MODE_DIRn_and_IMM1 … H);
+ nrewrite > (eq_to_eqoct n2 n2 (refl_eq …));
+ napply refl_eq
+ ##| ##31,33,34: #n; #H; nnormalize; napply (instr_mode_destruct … H)
+ ##| ##*: #H; nnormalize; napply (instr_mode_destruct … H)
##]
nqed.
#n1; #t2; ncases t2;
##[ ##33: #n2; #H;
nchange with (eq_ex n1 n2 = true);
- nrewrite > (instr_mode_destruct_MODE_TNY ?? H);
- nrewrite > (eq_to_eqex n2 n2 (refl_eq ??));
- napply (refl_eq ??)
- ##| ##31,32,34: #n; #H; nnormalize; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; nnormalize; napply (instr_mode_destruct ??? H)
+ nrewrite > (instr_mode_destruct_MODE_TNY … H);
+ nrewrite > (eq_to_eqex n2 n2 (refl_eq …));
+ napply refl_eq
+ ##| ##31,32,34: #n; #H; nnormalize; napply (instr_mode_destruct … H)
+ ##| ##*: #H; nnormalize; napply (instr_mode_destruct … H)
##]
nqed.
#n1; #t2; ncases t2;
##[ ##34: #n2; #H;
nchange with (eq_bitrig n1 n2 = true);
- nrewrite > (instr_mode_destruct_MODE_SRT ?? H);
- nrewrite > (eq_to_eqbitrig n2 n2 (refl_eq ??));
- napply (refl_eq ??)
- ##| ##31,32,33: #n; #H; nnormalize; napply (instr_mode_destruct ??? H)
- ##| ##*: #H; nnormalize; napply (instr_mode_destruct ??? H)
+ nrewrite > (instr_mode_destruct_MODE_SRT … H);
+ nrewrite > (eq_to_eqbitrig n2 n2 (refl_eq …));
+ napply refl_eq
+ ##| ##31,32,33: #n; #H; nnormalize; napply (instr_mode_destruct … H)
+ ##| ##*: #H; nnormalize; napply (instr_mode_destruct … H)
##]
nqed.
(* MATTONI BASE PER DEFINIRE LE TABELLE DELLE MCU *)
(* ********************************************** *)
-nlemma symmetric_eqop1 : ∀op2.eq_op ADC op2 = eq_op op2 ADC. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop2 : ∀op2.eq_op ADD op2 = eq_op op2 ADD. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop3 : ∀op2.eq_op AIS op2 = eq_op op2 AIS. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop4 : ∀op2.eq_op AIX op2 = eq_op op2 AIX. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop5 : ∀op2.eq_op AND op2 = eq_op op2 AND. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop6 : ∀op2.eq_op ASL op2 = eq_op op2 ASL. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop7 : ∀op2.eq_op ASR op2 = eq_op op2 ASR. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop8 : ∀op2.eq_op BCC op2 = eq_op op2 BCC. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop9 : ∀op2.eq_op BCLRn op2 = eq_op op2 BCLRn. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop10 : ∀op2.eq_op BCS op2 = eq_op op2 BCS. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop11 : ∀op2.eq_op BEQ op2 = eq_op op2 BEQ. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop12 : ∀op2.eq_op BGE op2 = eq_op op2 BGE. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop13 : ∀op2.eq_op BGND op2 = eq_op op2 BGND. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop14 : ∀op2.eq_op BGT op2 = eq_op op2 BGT. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop15 : ∀op2.eq_op BHCC op2 = eq_op op2 BHCC. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop16 : ∀op2.eq_op BHCS op2 = eq_op op2 BHCS. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop17 : ∀op2.eq_op BHI op2 = eq_op op2 BHI. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop18 : ∀op2.eq_op BIH op2 = eq_op op2 BIH. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop19 : ∀op2.eq_op BIL op2 = eq_op op2 BIL. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop20 : ∀op2.eq_op BIT op2 = eq_op op2 BIT. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop21 : ∀op2.eq_op BLE op2 = eq_op op2 BLE. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop22 : ∀op2.eq_op BLS op2 = eq_op op2 BLS. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop23 : ∀op2.eq_op BLT op2 = eq_op op2 BLT. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop24 : ∀op2.eq_op BMC op2 = eq_op op2 BMC. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop25 : ∀op2.eq_op BMI op2 = eq_op op2 BMI. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop26 : ∀op2.eq_op BMS op2 = eq_op op2 BMS. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop27 : ∀op2.eq_op BNE op2 = eq_op op2 BNE. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop28 : ∀op2.eq_op BPL op2 = eq_op op2 BPL. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop29 : ∀op2.eq_op BRA op2 = eq_op op2 BRA. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop30 : ∀op2.eq_op BRCLRn op2 = eq_op op2 BRCLRn. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop31 : ∀op2.eq_op BRN op2 = eq_op op2 BRN. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop32 : ∀op2.eq_op BRSETn op2 = eq_op op2 BRSETn. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop33 : ∀op2.eq_op BSETn op2 = eq_op op2 BSETn. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop34 : ∀op2.eq_op BSR op2 = eq_op op2 BSR. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop35 : ∀op2.eq_op CBEQA op2 = eq_op op2 CBEQA. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop36 : ∀op2.eq_op CBEQX op2 = eq_op op2 CBEQX. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop37 : ∀op2.eq_op CLC op2 = eq_op op2 CLC. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop38 : ∀op2.eq_op CLI op2 = eq_op op2 CLI. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop39 : ∀op2.eq_op CLR op2 = eq_op op2 CLR. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop40 : ∀op2.eq_op CMP op2 = eq_op op2 CMP. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop41 : ∀op2.eq_op COM op2 = eq_op op2 COM. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop42 : ∀op2.eq_op CPHX op2 = eq_op op2 CPHX. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop43 : ∀op2.eq_op CPX op2 = eq_op op2 CPX. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop44 : ∀op2.eq_op DAA op2 = eq_op op2 DAA. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop45 : ∀op2.eq_op DBNZ op2 = eq_op op2 DBNZ. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop46 : ∀op2.eq_op DEC op2 = eq_op op2 DEC. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop47 : ∀op2.eq_op DIV op2 = eq_op op2 DIV. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop48 : ∀op2.eq_op EOR op2 = eq_op op2 EOR. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop49 : ∀op2.eq_op INC op2 = eq_op op2 INC. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop50 : ∀op2.eq_op JMP op2 = eq_op op2 JMP. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop51 : ∀op2.eq_op JSR op2 = eq_op op2 JSR. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop52 : ∀op2.eq_op LDA op2 = eq_op op2 LDA. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop53 : ∀op2.eq_op LDHX op2 = eq_op op2 LDHX. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop54 : ∀op2.eq_op LDX op2 = eq_op op2 LDX. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop55 : ∀op2.eq_op LSR op2 = eq_op op2 LSR. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop56 : ∀op2.eq_op MOV op2 = eq_op op2 MOV. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop57 : ∀op2.eq_op MUL op2 = eq_op op2 MUL. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop58 : ∀op2.eq_op NEG op2 = eq_op op2 NEG. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop59 : ∀op2.eq_op NOP op2 = eq_op op2 NOP. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop60 : ∀op2.eq_op NSA op2 = eq_op op2 NSA. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop61 : ∀op2.eq_op ORA op2 = eq_op op2 ORA. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop62 : ∀op2.eq_op PSHA op2 = eq_op op2 PSHA. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop63 : ∀op2.eq_op PSHH op2 = eq_op op2 PSHH. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop64 : ∀op2.eq_op PSHX op2 = eq_op op2 PSHX. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop65 : ∀op2.eq_op PULA op2 = eq_op op2 PULA. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop66 : ∀op2.eq_op PULH op2 = eq_op op2 PULH. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop67 : ∀op2.eq_op PULX op2 = eq_op op2 PULX. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop68 : ∀op2.eq_op ROL op2 = eq_op op2 ROL. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop69 : ∀op2.eq_op ROR op2 = eq_op op2 ROR. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop70 : ∀op2.eq_op RSP op2 = eq_op op2 RSP. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop71 : ∀op2.eq_op RTI op2 = eq_op op2 RTI. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop72 : ∀op2.eq_op RTS op2 = eq_op op2 RTS. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop73 : ∀op2.eq_op SBC op2 = eq_op op2 SBC. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop74 : ∀op2.eq_op SEC op2 = eq_op op2 SEC. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop75 : ∀op2.eq_op SEI op2 = eq_op op2 SEI. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop76 : ∀op2.eq_op SHA op2 = eq_op op2 SHA. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop77 : ∀op2.eq_op SLA op2 = eq_op op2 SLA. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop78 : ∀op2.eq_op STA op2 = eq_op op2 STA. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop79 : ∀op2.eq_op STHX op2 = eq_op op2 STHX. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop80 : ∀op2.eq_op STOP op2 = eq_op op2 STOP. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop81 : ∀op2.eq_op STX op2 = eq_op op2 STX. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop82 : ∀op2.eq_op SUB op2 = eq_op op2 SUB. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop83 : ∀op2.eq_op SWI op2 = eq_op op2 SWI. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop84 : ∀op2.eq_op TAP op2 = eq_op op2 TAP. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop85 : ∀op2.eq_op TAX op2 = eq_op op2 TAX. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop86 : ∀op2.eq_op TPA op2 = eq_op op2 TPA. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop87 : ∀op2.eq_op TST op2 = eq_op op2 TST. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop88 : ∀op2.eq_op TSX op2 = eq_op op2 TSX. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop89 : ∀op2.eq_op TXA op2 = eq_op op2 TXA. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop90 : ∀op2.eq_op TXS op2 = eq_op op2 TXS. #op2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqop91 : ∀op2.eq_op WAIT op2 = eq_op op2 WAIT. #op2; nnormalize; napply (refl_eq ??).nqed.
+nlemma symmetric_eqop1 : ∀op2.eq_op ADC op2 = eq_op op2 ADC. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop2 : ∀op2.eq_op ADD op2 = eq_op op2 ADD. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop3 : ∀op2.eq_op AIS op2 = eq_op op2 AIS. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop4 : ∀op2.eq_op AIX op2 = eq_op op2 AIX. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop5 : ∀op2.eq_op AND op2 = eq_op op2 AND. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop6 : ∀op2.eq_op ASL op2 = eq_op op2 ASL. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop7 : ∀op2.eq_op ASR op2 = eq_op op2 ASR. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop8 : ∀op2.eq_op BCC op2 = eq_op op2 BCC. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop9 : ∀op2.eq_op BCLRn op2 = eq_op op2 BCLRn. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop10 : ∀op2.eq_op BCS op2 = eq_op op2 BCS. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop11 : ∀op2.eq_op BEQ op2 = eq_op op2 BEQ. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop12 : ∀op2.eq_op BGE op2 = eq_op op2 BGE. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop13 : ∀op2.eq_op BGND op2 = eq_op op2 BGND. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop14 : ∀op2.eq_op BGT op2 = eq_op op2 BGT. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop15 : ∀op2.eq_op BHCC op2 = eq_op op2 BHCC. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop16 : ∀op2.eq_op BHCS op2 = eq_op op2 BHCS. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop17 : ∀op2.eq_op BHI op2 = eq_op op2 BHI. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop18 : ∀op2.eq_op BIH op2 = eq_op op2 BIH. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop19 : ∀op2.eq_op BIL op2 = eq_op op2 BIL. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop20 : ∀op2.eq_op BIT op2 = eq_op op2 BIT. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop21 : ∀op2.eq_op BLE op2 = eq_op op2 BLE. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop22 : ∀op2.eq_op BLS op2 = eq_op op2 BLS. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop23 : ∀op2.eq_op BLT op2 = eq_op op2 BLT. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop24 : ∀op2.eq_op BMC op2 = eq_op op2 BMC. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop25 : ∀op2.eq_op BMI op2 = eq_op op2 BMI. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop26 : ∀op2.eq_op BMS op2 = eq_op op2 BMS. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop27 : ∀op2.eq_op BNE op2 = eq_op op2 BNE. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop28 : ∀op2.eq_op BPL op2 = eq_op op2 BPL. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop29 : ∀op2.eq_op BRA op2 = eq_op op2 BRA. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop30 : ∀op2.eq_op BRCLRn op2 = eq_op op2 BRCLRn. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop31 : ∀op2.eq_op BRN op2 = eq_op op2 BRN. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop32 : ∀op2.eq_op BRSETn op2 = eq_op op2 BRSETn. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop33 : ∀op2.eq_op BSETn op2 = eq_op op2 BSETn. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop34 : ∀op2.eq_op BSR op2 = eq_op op2 BSR. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop35 : ∀op2.eq_op CBEQA op2 = eq_op op2 CBEQA. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop36 : ∀op2.eq_op CBEQX op2 = eq_op op2 CBEQX. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop37 : ∀op2.eq_op CLC op2 = eq_op op2 CLC. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop38 : ∀op2.eq_op CLI op2 = eq_op op2 CLI. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop39 : ∀op2.eq_op CLR op2 = eq_op op2 CLR. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop40 : ∀op2.eq_op CMP op2 = eq_op op2 CMP. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop41 : ∀op2.eq_op COM op2 = eq_op op2 COM. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop42 : ∀op2.eq_op CPHX op2 = eq_op op2 CPHX. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop43 : ∀op2.eq_op CPX op2 = eq_op op2 CPX. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop44 : ∀op2.eq_op DAA op2 = eq_op op2 DAA. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop45 : ∀op2.eq_op DBNZ op2 = eq_op op2 DBNZ. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop46 : ∀op2.eq_op DEC op2 = eq_op op2 DEC. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop47 : ∀op2.eq_op DIV op2 = eq_op op2 DIV. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop48 : ∀op2.eq_op EOR op2 = eq_op op2 EOR. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop49 : ∀op2.eq_op INC op2 = eq_op op2 INC. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop50 : ∀op2.eq_op JMP op2 = eq_op op2 JMP. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop51 : ∀op2.eq_op JSR op2 = eq_op op2 JSR. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop52 : ∀op2.eq_op LDA op2 = eq_op op2 LDA. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop53 : ∀op2.eq_op LDHX op2 = eq_op op2 LDHX. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop54 : ∀op2.eq_op LDX op2 = eq_op op2 LDX. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop55 : ∀op2.eq_op LSR op2 = eq_op op2 LSR. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop56 : ∀op2.eq_op MOV op2 = eq_op op2 MOV. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop57 : ∀op2.eq_op MUL op2 = eq_op op2 MUL. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop58 : ∀op2.eq_op NEG op2 = eq_op op2 NEG. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop59 : ∀op2.eq_op NOP op2 = eq_op op2 NOP. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop60 : ∀op2.eq_op NSA op2 = eq_op op2 NSA. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop61 : ∀op2.eq_op ORA op2 = eq_op op2 ORA. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop62 : ∀op2.eq_op PSHA op2 = eq_op op2 PSHA. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop63 : ∀op2.eq_op PSHH op2 = eq_op op2 PSHH. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop64 : ∀op2.eq_op PSHX op2 = eq_op op2 PSHX. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop65 : ∀op2.eq_op PULA op2 = eq_op op2 PULA. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop66 : ∀op2.eq_op PULH op2 = eq_op op2 PULH. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop67 : ∀op2.eq_op PULX op2 = eq_op op2 PULX. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop68 : ∀op2.eq_op ROL op2 = eq_op op2 ROL. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop69 : ∀op2.eq_op ROR op2 = eq_op op2 ROR. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop70 : ∀op2.eq_op RSP op2 = eq_op op2 RSP. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop71 : ∀op2.eq_op RTI op2 = eq_op op2 RTI. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop72 : ∀op2.eq_op RTS op2 = eq_op op2 RTS. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop73 : ∀op2.eq_op SBC op2 = eq_op op2 SBC. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop74 : ∀op2.eq_op SEC op2 = eq_op op2 SEC. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop75 : ∀op2.eq_op SEI op2 = eq_op op2 SEI. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop76 : ∀op2.eq_op SHA op2 = eq_op op2 SHA. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop77 : ∀op2.eq_op SLA op2 = eq_op op2 SLA. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop78 : ∀op2.eq_op STA op2 = eq_op op2 STA. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop79 : ∀op2.eq_op STHX op2 = eq_op op2 STHX. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop80 : ∀op2.eq_op STOP op2 = eq_op op2 STOP. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop81 : ∀op2.eq_op STX op2 = eq_op op2 STX. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop82 : ∀op2.eq_op SUB op2 = eq_op op2 SUB. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop83 : ∀op2.eq_op SWI op2 = eq_op op2 SWI. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop84 : ∀op2.eq_op TAP op2 = eq_op op2 TAP. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop85 : ∀op2.eq_op TAX op2 = eq_op op2 TAX. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop86 : ∀op2.eq_op TPA op2 = eq_op op2 TPA. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop87 : ∀op2.eq_op TST op2 = eq_op op2 TST. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop88 : ∀op2.eq_op TSX op2 = eq_op op2 TSX. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop89 : ∀op2.eq_op TXA op2 = eq_op op2 TXA. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop90 : ∀op2.eq_op TXS op2 = eq_op op2 TXS. #op2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqop91 : ∀op2.eq_op WAIT op2 = eq_op op2 WAIT. #op2; nnormalize; napply refl_eq.nqed.
nlemma symmetric_eqop : symmetricT opcode bool eq_op.
#op1; ncases op1;
##| ##89: napply symmetric_eqop89 ##| ##90: napply symmetric_eqop90 ##| ##91: napply symmetric_eqop91 ##]
nqed.
-nlemma eqop_to_eq1 : ∀op2.eq_op ADC op2 = true → ADC = op2. #op2; ncases op2; nnormalize; #H; ##[ ##1: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq2 : ∀op2.eq_op ADD op2 = true → ADD = op2. #op2; ncases op2; nnormalize; #H; ##[ ##2: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq3 : ∀op2.eq_op AIS op2 = true → AIS = op2. #op2; ncases op2; nnormalize; #H; ##[ ##3: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq4 : ∀op2.eq_op AIX op2 = true → AIX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##4: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq5 : ∀op2.eq_op AND op2 = true → AND = op2. #op2; ncases op2; nnormalize; #H; ##[ ##5: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq6 : ∀op2.eq_op ASL op2 = true → ASL = op2. #op2; ncases op2; nnormalize; #H; ##[ ##6: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq7 : ∀op2.eq_op ASR op2 = true → ASR = op2. #op2; ncases op2; nnormalize; #H; ##[ ##7: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq8 : ∀op2.eq_op BCC op2 = true → BCC = op2. #op2; ncases op2; nnormalize; #H; ##[ ##8: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq9 : ∀op2.eq_op BCLRn op2 = true → BCLRn = op2. #op2; ncases op2; nnormalize; #H; ##[ ##9: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq10 : ∀op2.eq_op BCS op2 = true → BCS = op2. #op2; ncases op2; nnormalize; #H; ##[ ##10: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq11 : ∀op2.eq_op BEQ op2 = true → BEQ = op2. #op2; ncases op2; nnormalize; #H; ##[ ##11: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq12 : ∀op2.eq_op BGE op2 = true → BGE = op2. #op2; ncases op2; nnormalize; #H; ##[ ##12: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq13 : ∀op2.eq_op BGND op2 = true → BGND = op2. #op2; ncases op2; nnormalize; #H; ##[ ##13: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq14 : ∀op2.eq_op BGT op2 = true → BGT = op2. #op2; ncases op2; nnormalize; #H; ##[ ##14: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq15 : ∀op2.eq_op BHCC op2 = true → BHCC = op2. #op2; ncases op2; nnormalize; #H; ##[ ##15: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq16 : ∀op2.eq_op BHCS op2 = true → BHCS = op2. #op2; ncases op2; nnormalize; #H; ##[ ##16: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq17 : ∀op2.eq_op BHI op2 = true → BHI = op2. #op2; ncases op2; nnormalize; #H; ##[ ##17: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq18 : ∀op2.eq_op BIH op2 = true → BIH = op2. #op2; ncases op2; nnormalize; #H; ##[ ##18: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq19 : ∀op2.eq_op BIL op2 = true → BIL = op2. #op2; ncases op2; nnormalize; #H; ##[ ##19: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq20 : ∀op2.eq_op BIT op2 = true → BIT = op2. #op2; ncases op2; nnormalize; #H; ##[ ##20: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq21 : ∀op2.eq_op BLE op2 = true → BLE = op2. #op2; ncases op2; nnormalize; #H; ##[ ##21: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq22 : ∀op2.eq_op BLS op2 = true → BLS = op2. #op2; ncases op2; nnormalize; #H; ##[ ##22: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq23 : ∀op2.eq_op BLT op2 = true → BLT = op2. #op2; ncases op2; nnormalize; #H; ##[ ##23: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq24 : ∀op2.eq_op BMC op2 = true → BMC = op2. #op2; ncases op2; nnormalize; #H; ##[ ##24: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq25 : ∀op2.eq_op BMI op2 = true → BMI = op2. #op2; ncases op2; nnormalize; #H; ##[ ##25: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq26 : ∀op2.eq_op BMS op2 = true → BMS = op2. #op2; ncases op2; nnormalize; #H; ##[ ##26: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq27 : ∀op2.eq_op BNE op2 = true → BNE = op2. #op2; ncases op2; nnormalize; #H; ##[ ##27: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq28 : ∀op2.eq_op BPL op2 = true → BPL = op2. #op2; ncases op2; nnormalize; #H; ##[ ##28: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq29 : ∀op2.eq_op BRA op2 = true → BRA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##29: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq30 : ∀op2.eq_op BRCLRn op2 = true → BRCLRn = op2. #op2; ncases op2; nnormalize; #H; ##[ ##30: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq31 : ∀op2.eq_op BRN op2 = true → BRN = op2. #op2; ncases op2; nnormalize; #H; ##[ ##31: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq32 : ∀op2.eq_op BRSETn op2 = true → BRSETn = op2. #op2; ncases op2; nnormalize; #H; ##[ ##32: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq33 : ∀op2.eq_op BSETn op2 = true → BSETn = op2. #op2; ncases op2; nnormalize; #H; ##[ ##33: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq34 : ∀op2.eq_op BSR op2 = true → BSR = op2. #op2; ncases op2; nnormalize; #H; ##[ ##34: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq35 : ∀op2.eq_op CBEQA op2 = true → CBEQA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##35: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq36 : ∀op2.eq_op CBEQX op2 = true → CBEQX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##36: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq37 : ∀op2.eq_op CLC op2 = true → CLC = op2. #op2; ncases op2; nnormalize; #H; ##[ ##37: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq38 : ∀op2.eq_op CLI op2 = true → CLI = op2. #op2; ncases op2; nnormalize; #H; ##[ ##38: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq39 : ∀op2.eq_op CLR op2 = true → CLR = op2. #op2; ncases op2; nnormalize; #H; ##[ ##39: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq40 : ∀op2.eq_op CMP op2 = true → CMP = op2. #op2; ncases op2; nnormalize; #H; ##[ ##40: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq41 : ∀op2.eq_op COM op2 = true → COM = op2. #op2; ncases op2; nnormalize; #H; ##[ ##41: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq42 : ∀op2.eq_op CPHX op2 = true → CPHX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##42: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq43 : ∀op2.eq_op CPX op2 = true → CPX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##43: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq44 : ∀op2.eq_op DAA op2 = true → DAA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##44: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq45 : ∀op2.eq_op DBNZ op2 = true → DBNZ = op2. #op2; ncases op2; nnormalize; #H; ##[ ##45: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq46 : ∀op2.eq_op DEC op2 = true → DEC = op2. #op2; ncases op2; nnormalize; #H; ##[ ##46: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq47 : ∀op2.eq_op DIV op2 = true → DIV = op2. #op2; ncases op2; nnormalize; #H; ##[ ##47: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq48 : ∀op2.eq_op EOR op2 = true → EOR = op2. #op2; ncases op2; nnormalize; #H; ##[ ##48: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq49 : ∀op2.eq_op INC op2 = true → INC = op2. #op2; ncases op2; nnormalize; #H; ##[ ##49: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq50 : ∀op2.eq_op JMP op2 = true → JMP = op2. #op2; ncases op2; nnormalize; #H; ##[ ##50: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq51 : ∀op2.eq_op JSR op2 = true → JSR = op2. #op2; ncases op2; nnormalize; #H; ##[ ##51: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq52 : ∀op2.eq_op LDA op2 = true → LDA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##52: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq53 : ∀op2.eq_op LDHX op2 = true → LDHX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##53: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq54 : ∀op2.eq_op LDX op2 = true → LDX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##54: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq55 : ∀op2.eq_op LSR op2 = true → LSR = op2. #op2; ncases op2; nnormalize; #H; ##[ ##55: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq56 : ∀op2.eq_op MOV op2 = true → MOV = op2. #op2; ncases op2; nnormalize; #H; ##[ ##56: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq57 : ∀op2.eq_op MUL op2 = true → MUL = op2. #op2; ncases op2; nnormalize; #H; ##[ ##57: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq58 : ∀op2.eq_op NEG op2 = true → NEG = op2. #op2; ncases op2; nnormalize; #H; ##[ ##58: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq59 : ∀op2.eq_op NOP op2 = true → NOP = op2. #op2; ncases op2; nnormalize; #H; ##[ ##59: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq60 : ∀op2.eq_op NSA op2 = true → NSA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##60: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq61 : ∀op2.eq_op ORA op2 = true → ORA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##61: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq62 : ∀op2.eq_op PSHA op2 = true → PSHA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##62: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq63 : ∀op2.eq_op PSHH op2 = true → PSHH = op2. #op2; ncases op2; nnormalize; #H; ##[ ##63: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq64 : ∀op2.eq_op PSHX op2 = true → PSHX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##64: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq65 : ∀op2.eq_op PULA op2 = true → PULA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##65: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq66 : ∀op2.eq_op PULH op2 = true → PULH = op2. #op2; ncases op2; nnormalize; #H; ##[ ##66: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq67 : ∀op2.eq_op PULX op2 = true → PULX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##67: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq68 : ∀op2.eq_op ROL op2 = true → ROL = op2. #op2; ncases op2; nnormalize; #H; ##[ ##68: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq69 : ∀op2.eq_op ROR op2 = true → ROR = op2. #op2; ncases op2; nnormalize; #H; ##[ ##69: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq70 : ∀op2.eq_op RSP op2 = true → RSP = op2. #op2; ncases op2; nnormalize; #H; ##[ ##70: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq71 : ∀op2.eq_op RTI op2 = true → RTI = op2. #op2; ncases op2; nnormalize; #H; ##[ ##71: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq72 : ∀op2.eq_op RTS op2 = true → RTS = op2. #op2; ncases op2; nnormalize; #H; ##[ ##72: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq73 : ∀op2.eq_op SBC op2 = true → SBC = op2. #op2; ncases op2; nnormalize; #H; ##[ ##73: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq74 : ∀op2.eq_op SEC op2 = true → SEC = op2. #op2; ncases op2; nnormalize; #H; ##[ ##74: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq75 : ∀op2.eq_op SEI op2 = true → SEI = op2. #op2; ncases op2; nnormalize; #H; ##[ ##75: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq76 : ∀op2.eq_op SHA op2 = true → SHA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##76: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq77 : ∀op2.eq_op SLA op2 = true → SLA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##77: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq78 : ∀op2.eq_op STA op2 = true → STA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##78: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq79 : ∀op2.eq_op STHX op2 = true → STHX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##79: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq80 : ∀op2.eq_op STOP op2 = true → STOP = op2. #op2; ncases op2; nnormalize; #H; ##[ ##80: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq81 : ∀op2.eq_op STX op2 = true → STX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##81: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq82 : ∀op2.eq_op SUB op2 = true → SUB = op2. #op2; ncases op2; nnormalize; #H; ##[ ##82: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq83 : ∀op2.eq_op SWI op2 = true → SWI = op2. #op2; ncases op2; nnormalize; #H; ##[ ##83: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq84 : ∀op2.eq_op TAP op2 = true → TAP = op2. #op2; ncases op2; nnormalize; #H; ##[ ##84: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq85 : ∀op2.eq_op TAX op2 = true → TAX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##85: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq86 : ∀op2.eq_op TPA op2 = true → TPA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##86: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq87 : ∀op2.eq_op TST op2 = true → TST = op2. #op2; ncases op2; nnormalize; #H; ##[ ##87: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq88 : ∀op2.eq_op TSX op2 = true → TSX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##88: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq89 : ∀op2.eq_op TXA op2 = true → TXA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##89: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq90 : ∀op2.eq_op TXS op2 = true → TXS = op2. #op2; ncases op2; nnormalize; #H; ##[ ##90: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
-nlemma eqop_to_eq91 : ∀op2.eq_op WAIT op2 = true → WAIT = op2. #op2; ncases op2; nnormalize; #H; ##[ ##91: napply (refl_eq ??) ##| ##*: napply (bool_destruct ??? H) ##]nqed.
+nlemma eqop_to_eq1 : ∀op2.eq_op ADC op2 = true → ADC = op2. #op2; ncases op2; nnormalize; #H; ##[ ##1: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq2 : ∀op2.eq_op ADD op2 = true → ADD = op2. #op2; ncases op2; nnormalize; #H; ##[ ##2: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq3 : ∀op2.eq_op AIS op2 = true → AIS = op2. #op2; ncases op2; nnormalize; #H; ##[ ##3: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq4 : ∀op2.eq_op AIX op2 = true → AIX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##4: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq5 : ∀op2.eq_op AND op2 = true → AND = op2. #op2; ncases op2; nnormalize; #H; ##[ ##5: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq6 : ∀op2.eq_op ASL op2 = true → ASL = op2. #op2; ncases op2; nnormalize; #H; ##[ ##6: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq7 : ∀op2.eq_op ASR op2 = true → ASR = op2. #op2; ncases op2; nnormalize; #H; ##[ ##7: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq8 : ∀op2.eq_op BCC op2 = true → BCC = op2. #op2; ncases op2; nnormalize; #H; ##[ ##8: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq9 : ∀op2.eq_op BCLRn op2 = true → BCLRn = op2. #op2; ncases op2; nnormalize; #H; ##[ ##9: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq10 : ∀op2.eq_op BCS op2 = true → BCS = op2. #op2; ncases op2; nnormalize; #H; ##[ ##10: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq11 : ∀op2.eq_op BEQ op2 = true → BEQ = op2. #op2; ncases op2; nnormalize; #H; ##[ ##11: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq12 : ∀op2.eq_op BGE op2 = true → BGE = op2. #op2; ncases op2; nnormalize; #H; ##[ ##12: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq13 : ∀op2.eq_op BGND op2 = true → BGND = op2. #op2; ncases op2; nnormalize; #H; ##[ ##13: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq14 : ∀op2.eq_op BGT op2 = true → BGT = op2. #op2; ncases op2; nnormalize; #H; ##[ ##14: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq15 : ∀op2.eq_op BHCC op2 = true → BHCC = op2. #op2; ncases op2; nnormalize; #H; ##[ ##15: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq16 : ∀op2.eq_op BHCS op2 = true → BHCS = op2. #op2; ncases op2; nnormalize; #H; ##[ ##16: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq17 : ∀op2.eq_op BHI op2 = true → BHI = op2. #op2; ncases op2; nnormalize; #H; ##[ ##17: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq18 : ∀op2.eq_op BIH op2 = true → BIH = op2. #op2; ncases op2; nnormalize; #H; ##[ ##18: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq19 : ∀op2.eq_op BIL op2 = true → BIL = op2. #op2; ncases op2; nnormalize; #H; ##[ ##19: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq20 : ∀op2.eq_op BIT op2 = true → BIT = op2. #op2; ncases op2; nnormalize; #H; ##[ ##20: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq21 : ∀op2.eq_op BLE op2 = true → BLE = op2. #op2; ncases op2; nnormalize; #H; ##[ ##21: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq22 : ∀op2.eq_op BLS op2 = true → BLS = op2. #op2; ncases op2; nnormalize; #H; ##[ ##22: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq23 : ∀op2.eq_op BLT op2 = true → BLT = op2. #op2; ncases op2; nnormalize; #H; ##[ ##23: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq24 : ∀op2.eq_op BMC op2 = true → BMC = op2. #op2; ncases op2; nnormalize; #H; ##[ ##24: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq25 : ∀op2.eq_op BMI op2 = true → BMI = op2. #op2; ncases op2; nnormalize; #H; ##[ ##25: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq26 : ∀op2.eq_op BMS op2 = true → BMS = op2. #op2; ncases op2; nnormalize; #H; ##[ ##26: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq27 : ∀op2.eq_op BNE op2 = true → BNE = op2. #op2; ncases op2; nnormalize; #H; ##[ ##27: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq28 : ∀op2.eq_op BPL op2 = true → BPL = op2. #op2; ncases op2; nnormalize; #H; ##[ ##28: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq29 : ∀op2.eq_op BRA op2 = true → BRA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##29: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq30 : ∀op2.eq_op BRCLRn op2 = true → BRCLRn = op2. #op2; ncases op2; nnormalize; #H; ##[ ##30: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq31 : ∀op2.eq_op BRN op2 = true → BRN = op2. #op2; ncases op2; nnormalize; #H; ##[ ##31: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq32 : ∀op2.eq_op BRSETn op2 = true → BRSETn = op2. #op2; ncases op2; nnormalize; #H; ##[ ##32: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq33 : ∀op2.eq_op BSETn op2 = true → BSETn = op2. #op2; ncases op2; nnormalize; #H; ##[ ##33: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq34 : ∀op2.eq_op BSR op2 = true → BSR = op2. #op2; ncases op2; nnormalize; #H; ##[ ##34: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq35 : ∀op2.eq_op CBEQA op2 = true → CBEQA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##35: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq36 : ∀op2.eq_op CBEQX op2 = true → CBEQX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##36: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq37 : ∀op2.eq_op CLC op2 = true → CLC = op2. #op2; ncases op2; nnormalize; #H; ##[ ##37: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq38 : ∀op2.eq_op CLI op2 = true → CLI = op2. #op2; ncases op2; nnormalize; #H; ##[ ##38: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq39 : ∀op2.eq_op CLR op2 = true → CLR = op2. #op2; ncases op2; nnormalize; #H; ##[ ##39: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq40 : ∀op2.eq_op CMP op2 = true → CMP = op2. #op2; ncases op2; nnormalize; #H; ##[ ##40: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq41 : ∀op2.eq_op COM op2 = true → COM = op2. #op2; ncases op2; nnormalize; #H; ##[ ##41: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq42 : ∀op2.eq_op CPHX op2 = true → CPHX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##42: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq43 : ∀op2.eq_op CPX op2 = true → CPX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##43: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq44 : ∀op2.eq_op DAA op2 = true → DAA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##44: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq45 : ∀op2.eq_op DBNZ op2 = true → DBNZ = op2. #op2; ncases op2; nnormalize; #H; ##[ ##45: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq46 : ∀op2.eq_op DEC op2 = true → DEC = op2. #op2; ncases op2; nnormalize; #H; ##[ ##46: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq47 : ∀op2.eq_op DIV op2 = true → DIV = op2. #op2; ncases op2; nnormalize; #H; ##[ ##47: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq48 : ∀op2.eq_op EOR op2 = true → EOR = op2. #op2; ncases op2; nnormalize; #H; ##[ ##48: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq49 : ∀op2.eq_op INC op2 = true → INC = op2. #op2; ncases op2; nnormalize; #H; ##[ ##49: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq50 : ∀op2.eq_op JMP op2 = true → JMP = op2. #op2; ncases op2; nnormalize; #H; ##[ ##50: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq51 : ∀op2.eq_op JSR op2 = true → JSR = op2. #op2; ncases op2; nnormalize; #H; ##[ ##51: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq52 : ∀op2.eq_op LDA op2 = true → LDA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##52: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq53 : ∀op2.eq_op LDHX op2 = true → LDHX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##53: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq54 : ∀op2.eq_op LDX op2 = true → LDX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##54: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq55 : ∀op2.eq_op LSR op2 = true → LSR = op2. #op2; ncases op2; nnormalize; #H; ##[ ##55: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq56 : ∀op2.eq_op MOV op2 = true → MOV = op2. #op2; ncases op2; nnormalize; #H; ##[ ##56: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq57 : ∀op2.eq_op MUL op2 = true → MUL = op2. #op2; ncases op2; nnormalize; #H; ##[ ##57: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq58 : ∀op2.eq_op NEG op2 = true → NEG = op2. #op2; ncases op2; nnormalize; #H; ##[ ##58: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq59 : ∀op2.eq_op NOP op2 = true → NOP = op2. #op2; ncases op2; nnormalize; #H; ##[ ##59: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq60 : ∀op2.eq_op NSA op2 = true → NSA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##60: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq61 : ∀op2.eq_op ORA op2 = true → ORA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##61: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq62 : ∀op2.eq_op PSHA op2 = true → PSHA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##62: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq63 : ∀op2.eq_op PSHH op2 = true → PSHH = op2. #op2; ncases op2; nnormalize; #H; ##[ ##63: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq64 : ∀op2.eq_op PSHX op2 = true → PSHX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##64: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq65 : ∀op2.eq_op PULA op2 = true → PULA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##65: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq66 : ∀op2.eq_op PULH op2 = true → PULH = op2. #op2; ncases op2; nnormalize; #H; ##[ ##66: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq67 : ∀op2.eq_op PULX op2 = true → PULX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##67: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq68 : ∀op2.eq_op ROL op2 = true → ROL = op2. #op2; ncases op2; nnormalize; #H; ##[ ##68: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq69 : ∀op2.eq_op ROR op2 = true → ROR = op2. #op2; ncases op2; nnormalize; #H; ##[ ##69: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq70 : ∀op2.eq_op RSP op2 = true → RSP = op2. #op2; ncases op2; nnormalize; #H; ##[ ##70: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq71 : ∀op2.eq_op RTI op2 = true → RTI = op2. #op2; ncases op2; nnormalize; #H; ##[ ##71: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq72 : ∀op2.eq_op RTS op2 = true → RTS = op2. #op2; ncases op2; nnormalize; #H; ##[ ##72: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq73 : ∀op2.eq_op SBC op2 = true → SBC = op2. #op2; ncases op2; nnormalize; #H; ##[ ##73: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq74 : ∀op2.eq_op SEC op2 = true → SEC = op2. #op2; ncases op2; nnormalize; #H; ##[ ##74: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq75 : ∀op2.eq_op SEI op2 = true → SEI = op2. #op2; ncases op2; nnormalize; #H; ##[ ##75: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq76 : ∀op2.eq_op SHA op2 = true → SHA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##76: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq77 : ∀op2.eq_op SLA op2 = true → SLA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##77: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq78 : ∀op2.eq_op STA op2 = true → STA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##78: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq79 : ∀op2.eq_op STHX op2 = true → STHX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##79: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq80 : ∀op2.eq_op STOP op2 = true → STOP = op2. #op2; ncases op2; nnormalize; #H; ##[ ##80: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq81 : ∀op2.eq_op STX op2 = true → STX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##81: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq82 : ∀op2.eq_op SUB op2 = true → SUB = op2. #op2; ncases op2; nnormalize; #H; ##[ ##82: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq83 : ∀op2.eq_op SWI op2 = true → SWI = op2. #op2; ncases op2; nnormalize; #H; ##[ ##83: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq84 : ∀op2.eq_op TAP op2 = true → TAP = op2. #op2; ncases op2; nnormalize; #H; ##[ ##84: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq85 : ∀op2.eq_op TAX op2 = true → TAX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##85: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq86 : ∀op2.eq_op TPA op2 = true → TPA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##86: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq87 : ∀op2.eq_op TST op2 = true → TST = op2. #op2; ncases op2; nnormalize; #H; ##[ ##87: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq88 : ∀op2.eq_op TSX op2 = true → TSX = op2. #op2; ncases op2; nnormalize; #H; ##[ ##88: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq89 : ∀op2.eq_op TXA op2 = true → TXA = op2. #op2; ncases op2; nnormalize; #H; ##[ ##89: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq90 : ∀op2.eq_op TXS op2 = true → TXS = op2. #op2; ncases op2; nnormalize; #H; ##[ ##90: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
+nlemma eqop_to_eq91 : ∀op2.eq_op WAIT op2 = true → WAIT = op2. #op2; ncases op2; nnormalize; #H; ##[ ##91: napply refl_eq ##| ##*: napply (bool_destruct … H) ##]nqed.
nlemma eqop_to_eq : ∀op1,op2.eq_op op1 op2 = true → op1 = op2.
#op1; ncases op1;
##| ##89: napply eqop_to_eq89 ##| ##90: napply eqop_to_eq90 ##| ##91: napply eqop_to_eq91 ##]
nqed.
-nlemma eq_to_eqop1 : ∀op2.ADC = op2 → eq_op ADC op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##1: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop2 : ∀op2.ADD = op2 → eq_op ADD op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##2: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop3 : ∀op2.AIS = op2 → eq_op AIS op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##3: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop4 : ∀op2.AIX = op2 → eq_op AIX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##4: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop5 : ∀op2.AND = op2 → eq_op AND op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##5: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop6 : ∀op2.ASL = op2 → eq_op ASL op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##6: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop7 : ∀op2.ASR = op2 → eq_op ASR op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##7: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop8 : ∀op2.BCC = op2 → eq_op BCC op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##8: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop9 : ∀op2.BCLRn = op2 → eq_op BCLRn op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##9: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop10 : ∀op2.BCS = op2 → eq_op BCS op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##10: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop11 : ∀op2.BEQ = op2 → eq_op BEQ op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##11: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop12 : ∀op2.BGE = op2 → eq_op BGE op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##12: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop13 : ∀op2.BGND = op2 → eq_op BGND op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##13: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop14 : ∀op2.BGT = op2 → eq_op BGT op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##14: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop15 : ∀op2.BHCC = op2 → eq_op BHCC op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##15: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop16 : ∀op2.BHCS = op2 → eq_op BHCS op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##16: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop17 : ∀op2.BHI = op2 → eq_op BHI op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##17: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop18 : ∀op2.BIH = op2 → eq_op BIH op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##18: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop19 : ∀op2.BIL = op2 → eq_op BIL op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##19: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop20 : ∀op2.BIT = op2 → eq_op BIT op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##20: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop21 : ∀op2.BLE = op2 → eq_op BLE op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##21: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop22 : ∀op2.BLS = op2 → eq_op BLS op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##22: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop23 : ∀op2.BLT = op2 → eq_op BLT op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##23: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop24 : ∀op2.BMC = op2 → eq_op BMC op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##24: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop25 : ∀op2.BMI = op2 → eq_op BMI op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##25: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop26 : ∀op2.BMS = op2 → eq_op BMS op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##26: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop27 : ∀op2.BNE = op2 → eq_op BNE op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##27: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop28 : ∀op2.BPL = op2 → eq_op BPL op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##28: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop29 : ∀op2.BRA = op2 → eq_op BRA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##29: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop30 : ∀op2.BRCLRn = op2 → eq_op BRCLRn op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##30: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop31 : ∀op2.BRN = op2 → eq_op BRN op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##31: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop32 : ∀op2.BRSETn = op2 → eq_op BRSETn op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##32: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop33 : ∀op2.BSETn = op2 → eq_op BSETn op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##33: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop34 : ∀op2.BSR = op2 → eq_op BSR op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##34: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop35 : ∀op2.CBEQA = op2 → eq_op CBEQA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##35: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop36 : ∀op2.CBEQX = op2 → eq_op CBEQX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##36: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop37 : ∀op2.CLC = op2 → eq_op CLC op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##37: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop38 : ∀op2.CLI = op2 → eq_op CLI op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##38: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop39 : ∀op2.CLR = op2 → eq_op CLR op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##39: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop40 : ∀op2.CMP = op2 → eq_op CMP op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##40: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop41 : ∀op2.COM = op2 → eq_op COM op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##41: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop42 : ∀op2.CPHX = op2 → eq_op CPHX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##42: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop43 : ∀op2.CPX = op2 → eq_op CPX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##43: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop44 : ∀op2.DAA = op2 → eq_op DAA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##44: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop45 : ∀op2.DBNZ = op2 → eq_op DBNZ op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##45: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop46 : ∀op2.DEC = op2 → eq_op DEC op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##46: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop47 : ∀op2.DIV = op2 → eq_op DIV op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##47: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop48 : ∀op2.EOR = op2 → eq_op EOR op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##48: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop49 : ∀op2.INC = op2 → eq_op INC op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##49: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop50 : ∀op2.JMP = op2 → eq_op JMP op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##50: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop51 : ∀op2.JSR = op2 → eq_op JSR op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##51: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop52 : ∀op2.LDA = op2 → eq_op LDA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##52: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop53 : ∀op2.LDHX = op2 → eq_op LDHX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##53: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop54 : ∀op2.LDX = op2 → eq_op LDX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##54: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop55 : ∀op2.LSR = op2 → eq_op LSR op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##55: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop56 : ∀op2.MOV = op2 → eq_op MOV op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##56: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop57 : ∀op2.MUL = op2 → eq_op MUL op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##57: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop58 : ∀op2.NEG = op2 → eq_op NEG op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##58: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop59 : ∀op2.NOP = op2 → eq_op NOP op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##59: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop60 : ∀op2.NSA = op2 → eq_op NSA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##60: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop61 : ∀op2.ORA = op2 → eq_op ORA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##61: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop62 : ∀op2.PSHA = op2 → eq_op PSHA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##62: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop63 : ∀op2.PSHH = op2 → eq_op PSHH op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##63: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop64 : ∀op2.PSHX = op2 → eq_op PSHX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##64: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop65 : ∀op2.PULA = op2 → eq_op PULA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##65: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop66 : ∀op2.PULH = op2 → eq_op PULH op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##66: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop67 : ∀op2.PULX = op2 → eq_op PULX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##67: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop68 : ∀op2.ROL = op2 → eq_op ROL op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##68: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop69 : ∀op2.ROR = op2 → eq_op ROR op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##69: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop70 : ∀op2.RSP = op2 → eq_op RSP op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##70: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop71 : ∀op2.RTI = op2 → eq_op RTI op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##71: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop72 : ∀op2.RTS = op2 → eq_op RTS op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##72: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop73 : ∀op2.SBC = op2 → eq_op SBC op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##73: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop74 : ∀op2.SEC = op2 → eq_op SEC op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##74: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop75 : ∀op2.SEI = op2 → eq_op SEI op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##75: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop76 : ∀op2.SHA = op2 → eq_op SHA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##76: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop77 : ∀op2.SLA = op2 → eq_op SLA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##77: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop78 : ∀op2.STA = op2 → eq_op STA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##78: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop79 : ∀op2.STHX = op2 → eq_op STHX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##79: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop80 : ∀op2.STOP = op2 → eq_op STOP op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##80: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop81 : ∀op2.STX = op2 → eq_op STX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##81: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop82 : ∀op2.SUB = op2 → eq_op SUB op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##82: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop83 : ∀op2.SWI = op2 → eq_op SWI op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##83: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop84 : ∀op2.TAP = op2 → eq_op TAP op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##84: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop85 : ∀op2.TAX = op2 → eq_op TAX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##85: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop86 : ∀op2.TPA = op2 → eq_op TPA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##86: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop87 : ∀op2.TST = op2 → eq_op TST op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##87: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop88 : ∀op2.TSX = op2 → eq_op TSX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##88: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop89 : ∀op2.TXA = op2 → eq_op TXA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##89: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop90 : ∀op2.TXS = op2 → eq_op TXS op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##90: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
-nlemma eq_to_eqop91 : ∀op2.WAIT = op2 → eq_op WAIT op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##91: napply (refl_eq ??) ##| ##*: napply (opcode_destruct ?? (false = true) H) ##]nqed.
+nlemma eq_to_eqop1 : ∀op2.ADC = op2 → eq_op ADC op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##1: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop2 : ∀op2.ADD = op2 → eq_op ADD op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##2: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop3 : ∀op2.AIS = op2 → eq_op AIS op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##3: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop4 : ∀op2.AIX = op2 → eq_op AIX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##4: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop5 : ∀op2.AND = op2 → eq_op AND op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##5: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop6 : ∀op2.ASL = op2 → eq_op ASL op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##6: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop7 : ∀op2.ASR = op2 → eq_op ASR op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##7: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop8 : ∀op2.BCC = op2 → eq_op BCC op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##8: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop9 : ∀op2.BCLRn = op2 → eq_op BCLRn op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##9: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop10 : ∀op2.BCS = op2 → eq_op BCS op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##10: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop11 : ∀op2.BEQ = op2 → eq_op BEQ op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##11: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop12 : ∀op2.BGE = op2 → eq_op BGE op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##12: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop13 : ∀op2.BGND = op2 → eq_op BGND op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##13: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop14 : ∀op2.BGT = op2 → eq_op BGT op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##14: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop15 : ∀op2.BHCC = op2 → eq_op BHCC op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##15: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop16 : ∀op2.BHCS = op2 → eq_op BHCS op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##16: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop17 : ∀op2.BHI = op2 → eq_op BHI op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##17: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop18 : ∀op2.BIH = op2 → eq_op BIH op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##18: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop19 : ∀op2.BIL = op2 → eq_op BIL op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##19: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop20 : ∀op2.BIT = op2 → eq_op BIT op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##20: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop21 : ∀op2.BLE = op2 → eq_op BLE op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##21: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop22 : ∀op2.BLS = op2 → eq_op BLS op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##22: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop23 : ∀op2.BLT = op2 → eq_op BLT op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##23: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop24 : ∀op2.BMC = op2 → eq_op BMC op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##24: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop25 : ∀op2.BMI = op2 → eq_op BMI op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##25: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop26 : ∀op2.BMS = op2 → eq_op BMS op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##26: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop27 : ∀op2.BNE = op2 → eq_op BNE op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##27: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop28 : ∀op2.BPL = op2 → eq_op BPL op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##28: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop29 : ∀op2.BRA = op2 → eq_op BRA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##29: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop30 : ∀op2.BRCLRn = op2 → eq_op BRCLRn op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##30: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop31 : ∀op2.BRN = op2 → eq_op BRN op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##31: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop32 : ∀op2.BRSETn = op2 → eq_op BRSETn op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##32: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop33 : ∀op2.BSETn = op2 → eq_op BSETn op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##33: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop34 : ∀op2.BSR = op2 → eq_op BSR op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##34: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop35 : ∀op2.CBEQA = op2 → eq_op CBEQA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##35: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop36 : ∀op2.CBEQX = op2 → eq_op CBEQX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##36: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop37 : ∀op2.CLC = op2 → eq_op CLC op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##37: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop38 : ∀op2.CLI = op2 → eq_op CLI op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##38: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop39 : ∀op2.CLR = op2 → eq_op CLR op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##39: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop40 : ∀op2.CMP = op2 → eq_op CMP op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##40: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop41 : ∀op2.COM = op2 → eq_op COM op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##41: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop42 : ∀op2.CPHX = op2 → eq_op CPHX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##42: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop43 : ∀op2.CPX = op2 → eq_op CPX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##43: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop44 : ∀op2.DAA = op2 → eq_op DAA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##44: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop45 : ∀op2.DBNZ = op2 → eq_op DBNZ op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##45: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop46 : ∀op2.DEC = op2 → eq_op DEC op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##46: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop47 : ∀op2.DIV = op2 → eq_op DIV op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##47: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop48 : ∀op2.EOR = op2 → eq_op EOR op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##48: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop49 : ∀op2.INC = op2 → eq_op INC op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##49: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop50 : ∀op2.JMP = op2 → eq_op JMP op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##50: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop51 : ∀op2.JSR = op2 → eq_op JSR op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##51: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop52 : ∀op2.LDA = op2 → eq_op LDA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##52: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop53 : ∀op2.LDHX = op2 → eq_op LDHX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##53: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop54 : ∀op2.LDX = op2 → eq_op LDX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##54: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop55 : ∀op2.LSR = op2 → eq_op LSR op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##55: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop56 : ∀op2.MOV = op2 → eq_op MOV op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##56: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop57 : ∀op2.MUL = op2 → eq_op MUL op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##57: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop58 : ∀op2.NEG = op2 → eq_op NEG op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##58: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop59 : ∀op2.NOP = op2 → eq_op NOP op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##59: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop60 : ∀op2.NSA = op2 → eq_op NSA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##60: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop61 : ∀op2.ORA = op2 → eq_op ORA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##61: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop62 : ∀op2.PSHA = op2 → eq_op PSHA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##62: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop63 : ∀op2.PSHH = op2 → eq_op PSHH op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##63: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop64 : ∀op2.PSHX = op2 → eq_op PSHX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##64: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop65 : ∀op2.PULA = op2 → eq_op PULA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##65: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop66 : ∀op2.PULH = op2 → eq_op PULH op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##66: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop67 : ∀op2.PULX = op2 → eq_op PULX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##67: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop68 : ∀op2.ROL = op2 → eq_op ROL op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##68: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop69 : ∀op2.ROR = op2 → eq_op ROR op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##69: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop70 : ∀op2.RSP = op2 → eq_op RSP op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##70: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop71 : ∀op2.RTI = op2 → eq_op RTI op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##71: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop72 : ∀op2.RTS = op2 → eq_op RTS op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##72: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop73 : ∀op2.SBC = op2 → eq_op SBC op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##73: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop74 : ∀op2.SEC = op2 → eq_op SEC op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##74: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop75 : ∀op2.SEI = op2 → eq_op SEI op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##75: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop76 : ∀op2.SHA = op2 → eq_op SHA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##76: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop77 : ∀op2.SLA = op2 → eq_op SLA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##77: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop78 : ∀op2.STA = op2 → eq_op STA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##78: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop79 : ∀op2.STHX = op2 → eq_op STHX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##79: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop80 : ∀op2.STOP = op2 → eq_op STOP op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##80: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop81 : ∀op2.STX = op2 → eq_op STX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##81: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop82 : ∀op2.SUB = op2 → eq_op SUB op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##82: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop83 : ∀op2.SWI = op2 → eq_op SWI op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##83: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop84 : ∀op2.TAP = op2 → eq_op TAP op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##84: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop85 : ∀op2.TAX = op2 → eq_op TAX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##85: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop86 : ∀op2.TPA = op2 → eq_op TPA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##86: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop87 : ∀op2.TST = op2 → eq_op TST op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##87: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop88 : ∀op2.TSX = op2 → eq_op TSX op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##88: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop89 : ∀op2.TXA = op2 → eq_op TXA op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##89: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop90 : ∀op2.TXS = op2 → eq_op TXS op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##90: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
+nlemma eq_to_eqop91 : ∀op2.WAIT = op2 → eq_op WAIT op2 = true. #op2; ncases op2; nnormalize; #H; ##[ ##91: napply refl_eq ##| ##*: napply (opcode_destruct … (false = true) H) ##]nqed.
nlemma eq_to_eqop : ∀op1,op2.op1 = op2 → eq_op op1 op2 = true.
#op1; ncases op1;
nchange with (match Some T x2 with [ None ⇒ False | Some a ⇒ x1 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma option_destruct_some_none : ∀T.∀x:T.Some T x = None T → False.
nelim op1;
nelim op2;
nnormalize;
- ##[ ##1: napply (refl_eq ??)
- ##| ##2,3: #H; napply (refl_eq ??)
+ ##[ ##1: napply refl_eq
+ ##| ##2,3: #H; napply refl_eq
##| ##4: #a; #a0;
nrewrite > (H a0 a);
- napply (refl_eq ??)
+ napply refl_eq
##]
nqed.
nelim op1;
nelim op2;
nnormalize;
- ##[ ##1: #H1; napply (refl_eq ??)
+ ##[ ##1: #H1; napply refl_eq
##| ##2: #a; #H1; nelim (option_destruct_none_some ?? H1)
##| ##3: #a; #H1; nelim (option_destruct_some_none ?? H1)
##| ##4: #a; #a0; #H1;
- nrewrite > (option_destruct_some_some ??? H1);
- nrewrite > (H a a (refl_eq ??));
- napply (refl_eq ??)
+ nrewrite > (option_destruct_some_some … H1);
+ nrewrite > (H a a (refl_eq …));
+ napply refl_eq
##]
nqed.
nelim op1;
nelim op2;
nnormalize;
- ##[ ##1: #H1; napply (refl_eq ??)
- ##| ##2,3: #a; #H1; napply (bool_destruct ??? H1)
+ ##[ ##1: #H1; napply refl_eq
+ ##| ##2,3: #a; #H1; napply (bool_destruct … H1)
##| ##4: #a; #a0; #H1;
- nrewrite > (H ?? H1);
- napply (refl_eq ??)
+ nrewrite > (H … H1);
+ napply refl_eq
##]
nqed.
nchange with (match pair T1 T2 x2 y2 with [ pair a _ ⇒ x1 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma pair_destruct_2 :
nchange with (match pair T1 T2 x2 y2 with [ pair _ b ⇒ y1 = b ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_eqpair :
nrewrite > (H x1 x2);
ncases (f1 x2 x1);
nnormalize;
- ##[ ##1: nrewrite > (H1 y1 y2); napply (refl_eq ??)
- ##| ##2: napply (refl_eq ??)
+ ##[ ##1: nrewrite > (H1 y1 y2); napply refl_eq
+ ##| ##2: napply refl_eq
##]
nqed.
nelim p2;
#x2; #y2; #H;
nnormalize;
- nrewrite > (H1 ?? (pair_destruct_1 ?????? H));
+ nrewrite > (H1 … (pair_destruct_1 … H));
nnormalize;
- nrewrite > (H2 ?? (pair_destruct_2 ?????? H));
- napply (refl_eq ??).
+ nrewrite > (H2 … (pair_destruct_2 … H));
+ napply refl_eq.
nqed.
nlemma eqpair_to_eq :
ncases (f1 x1 x2) in H:(%) K:(%);
nnormalize;
#H3;
- ##[ ##2: napply (bool_destruct ??? H3) ##]
+ ##[ ##2: napply (bool_destruct … H3) ##]
#H4;
- nrewrite > (H4 (refl_eq ??));
+ nrewrite > (H4 (refl_eq …));
nrewrite > (H2 y1 y2 H3);
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma triple_destruct_1 :
nchange with (match triple T1 T2 T3 x2 y2 z2 with [ triple a _ _ ⇒ x1 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma triple_destruct_2 :
nchange with (match triple T1 T2 T3 x2 y2 z2 with [ triple _ b _ ⇒ y1 = b ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma triple_destruct_3 :
nchange with (match triple T1 T2 T3 x2 y2 z2 with [ triple _ _ c ⇒ z1 = c ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_eqtriple :
##[ ##1: nrewrite > (H1 y1 y2);
ncases (f2 y2 y1);
nnormalize;
- ##[ ##1: nrewrite > (H2 z1 z2); napply (refl_eq ??)
- ##| ##2: napply (refl_eq ??)
+ ##[ ##1: nrewrite > (H2 z1 z2); napply refl_eq
+ ##| ##2: napply refl_eq
##]
- ##| ##2: napply (refl_eq ??)
+ ##| ##2: napply refl_eq
##]
nqed.
nelim p2;
#x2; #y2; #z2; #H;
nnormalize;
- nrewrite > (H1 ?? (triple_destruct_1 ????????? H));
+ nrewrite > (H1 … (triple_destruct_1 … H));
nnormalize;
- nrewrite > (H2 ?? (triple_destruct_2 ????????? H));
+ nrewrite > (H2 … (triple_destruct_2 … H));
nnormalize;
- nrewrite > (H3 ?? (triple_destruct_3 ????????? H));
- napply (refl_eq ??).
+ nrewrite > (H3 … (triple_destruct_3 … H));
+ napply refl_eq.
nqed.
nlemma eqtriple_to_eq :
nletin K ≝ (H1 x1 x2);
ncases (f1 x1 x2) in H:(%) K:(%);
nnormalize;
- ##[ ##2: #H4; napply (bool_destruct ??? H4) ##]
+ ##[ ##2: #H4; napply (bool_destruct … H4) ##]
nletin K1 ≝ (H2 y1 y2);
ncases (f2 y1 y2) in K1:(%) ⊢ %;
nnormalize;
- ##[ ##2: #H4; #H5; napply (bool_destruct ??? H5) ##]
+ ##[ ##2: #H4; #H5; napply (bool_destruct … H5) ##]
#H4; #H5; #H6;
- nrewrite > (H4 (refl_eq ??));
- nrewrite > (H6 (refl_eq ??));
+ nrewrite > (H4 (refl_eq …));
+ nrewrite > (H6 (refl_eq …));
nrewrite > (H3 z1 z2 H5);
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma quadruple_destruct_1 :
nchange with (match quadruple T1 T2 T3 T4 x2 y2 z2 v2 with [ quadruple a _ _ _ ⇒ x1 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma quadruple_destruct_2 :
nchange with (match quadruple T1 T2 T3 T4 x2 y2 z2 v2 with [ quadruple _ b _ _ ⇒ y1 = b ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma quadruple_destruct_3 :
nchange with (match quadruple T1 T2 T3 T4 x2 y2 z2 v2 with [ quadruple _ _ c _ ⇒ z1 = c ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma quadruple_destruct_4 :
nchange with (match quadruple T1 T2 T3 T4 x2 y2 z2 v2 with [ quadruple _ _ _ d ⇒ v1 = d ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_eqquadruple :
##[ ##1: nrewrite > (H2 z1 z2);
ncases (f3 z2 z1);
nnormalize;
- ##[ ##1: nrewrite > (H3 v1 v2); napply (refl_eq ??)
- ##| ##2: napply (refl_eq ??)
+ ##[ ##1: nrewrite > (H3 v1 v2); napply refl_eq
+ ##| ##2: napply refl_eq
##]
- ##| ##2: napply (refl_eq ??)
+ ##| ##2: napply refl_eq
##]
- ##| ##2: napply (refl_eq ??)
+ ##| ##2: napply refl_eq
##]
nqed.
nelim p2;
#x2; #y2; #z2; #v2; #H;
nnormalize;
- nrewrite > (H1 ?? (quadruple_destruct_1 ???????????? H));
+ nrewrite > (H1 … (quadruple_destruct_1 … H));
nnormalize;
- nrewrite > (H2 ?? (quadruple_destruct_2 ???????????? H));
+ nrewrite > (H2 … (quadruple_destruct_2 … H));
nnormalize;
- nrewrite > (H3 ?? (quadruple_destruct_3 ???????????? H));
+ nrewrite > (H3 … (quadruple_destruct_3 … H));
nnormalize;
- nrewrite > (H4 ?? (quadruple_destruct_4 ???????????? H));
- napply (refl_eq ??).
+ nrewrite > (H4 … (quadruple_destruct_4 … H));
+ napply refl_eq.
nqed.
nlemma eqquadruple_to_eq :
nletin K ≝ (H1 x1 x2);
ncases (f1 x1 x2) in H:(%) K:(%);
nnormalize;
- ##[ ##2: #H5; napply (bool_destruct ??? H5) ##]
+ ##[ ##2: #H5; napply (bool_destruct … H5) ##]
nletin K1 ≝ (H2 y1 y2);
ncases (f2 y1 y2) in K1:(%) ⊢ %;
nnormalize;
- ##[ ##2: #H5; #H6; napply (bool_destruct ??? H6) ##]
+ ##[ ##2: #H5; #H6; napply (bool_destruct … H6) ##]
nletin K2 ≝ (H3 z1 z2);
ncases (f3 z1 z2) in K2:(%) ⊢ %;
nnormalize;
- ##[ ##2: #H5; #H6; #H7; napply (bool_destruct ??? H7) ##]
+ ##[ ##2: #H5; #H6; #H7; napply (bool_destruct … H7) ##]
#H5; #H6; #H7; #H8;
- nrewrite > (H5 (refl_eq ??));
- nrewrite > (H6 (refl_eq ??));
- nrewrite > (H8 (refl_eq ??));
+ nrewrite > (H5 (refl_eq …));
+ nrewrite > (H6 (refl_eq …));
+ nrewrite > (H8 (refl_eq …));
nrewrite > (H4 v1 v2 H7);
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma quintuple_destruct_1 :
nchange with (match quintuple T1 T2 T3 T4 T5 x2 y2 z2 v2 w2 with [ quintuple a _ _ _ _ ⇒ x1 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma quintuple_destruct_2 :
nchange with (match quintuple T1 T2 T3 T4 T5 x2 y2 z2 v2 w2 with [ quintuple _ b _ _ _ ⇒ y1 = b ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma quintuple_destruct_3 :
nchange with (match quintuple T1 T2 T3 T4 T5 x2 y2 z2 v2 w2 with [ quintuple _ _ c _ _ ⇒ z1 = c ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma quintuple_destruct_4 :
nchange with (match quintuple T1 T2 T3 T4 T5 x2 y2 z2 v2 w2 with [ quintuple _ _ _ d _ ⇒ v1 = d ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma quintuple_destruct_5 :
nchange with (match quintuple T1 T2 T3 T4 T5 x2 y2 z2 v2 w2 with [ quintuple _ _ _ _ e ⇒ w1 = e ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_eqquintuple :
##[ ##1: nrewrite > (H3 v1 v2);
ncases (f4 v2 v1);
nnormalize;
- ##[ ##1: nrewrite > (H4 w1 w2); napply (refl_eq ??)
- ##| ##2: napply (refl_eq ??)
+ ##[ ##1: nrewrite > (H4 w1 w2); napply refl_eq
+ ##| ##2: napply refl_eq
##]
- ##| ##2: napply (refl_eq ??)
+ ##| ##2: napply refl_eq
##]
- ##| ##2: napply (refl_eq ??)
+ ##| ##2: napply refl_eq
##]
- ##| ##2: napply (refl_eq ??)
+ ##| ##2: napply refl_eq
##]
nqed.
nelim p2;
#x2; #y2; #z2; #v2; #w2; #H;
nnormalize;
- nrewrite > (H1 ?? (quintuple_destruct_1 ??????????????? H));
+ nrewrite > (H1 … (quintuple_destruct_1 … H));
nnormalize;
- nrewrite > (H2 ?? (quintuple_destruct_2 ??????????????? H));
+ nrewrite > (H2 … (quintuple_destruct_2 … H));
nnormalize;
- nrewrite > (H3 ?? (quintuple_destruct_3 ??????????????? H));
+ nrewrite > (H3 … (quintuple_destruct_3 … H));
nnormalize;
- nrewrite > (H4 ?? (quintuple_destruct_4 ??????????????? H));
+ nrewrite > (H4 … (quintuple_destruct_4 … H));
nnormalize;
- nrewrite > (H5 ?? (quintuple_destruct_5 ??????????????? H));
- napply (refl_eq ??).
+ nrewrite > (H5 … (quintuple_destruct_5 … H));
+ napply refl_eq.
nqed.
nlemma eqquintuple_to_eq :
nletin K ≝ (H1 x1 x2);
ncases (f1 x1 x2) in H:(%) K:(%);
nnormalize;
- ##[ ##2: #H6; napply (bool_destruct ??? H6) ##]
+ ##[ ##2: #H6; napply (bool_destruct … H6) ##]
nletin K1 ≝ (H2 y1 y2);
ncases (f2 y1 y2) in K1:(%) ⊢ %;
nnormalize;
- ##[ ##2: #H6; #H7; napply (bool_destruct ??? H7) ##]
+ ##[ ##2: #H6; #H7; napply (bool_destruct … H7) ##]
nletin K2 ≝ (H3 z1 z2);
ncases (f3 z1 z2) in K2:(%) ⊢ %;
nnormalize;
- ##[ ##2: #H6; #H7; #H8; napply (bool_destruct ??? H8) ##]
+ ##[ ##2: #H6; #H7; #H8; napply (bool_destruct … H8) ##]
nletin K3 ≝ (H4 v1 v2);
ncases (f4 v1 v2) in K3:(%) ⊢ %;
nnormalize;
- ##[ ##2: #H6; #H7; #H8; #H9; napply (bool_destruct ??? H9) ##]
+ ##[ ##2: #H6; #H7; #H8; #H9; napply (bool_destruct … H9) ##]
#H6; #H7; #H8; #H9; #H10;
- nrewrite > (H6 (refl_eq ??));
- nrewrite > (H7 (refl_eq ??));
- nrewrite > (H8 (refl_eq ??));
- nrewrite > (H10 (refl_eq ??));
+ nrewrite > (H6 (refl_eq …));
+ nrewrite > (H7 (refl_eq …));
+ nrewrite > (H8 (refl_eq …));
+ nrewrite > (H10 (refl_eq …));
nrewrite > (H5 w1 w2 H9);
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
(* *)
(* ********************************************************************** *)
-universe constraint Type[0] < Type[1].
\ No newline at end of file
+universe constraint Type[0] < Type[1].
(* setter forte di X *)
ndefinition set_indX_8_low_reg ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λindxlow':byte8.
- opt_map ?? (match m return aux_set_typing_opt byte8 with
+ opt_map … (match m return aux_set_typing_opt byte8 with
[ HC05 ⇒ Some ? set_indX_8_low_reg_HC05
| HC08 ⇒ Some ? set_indX_8_low_reg_HC08
| HCS08 ⇒ Some ? set_indX_8_low_reg_HC08
(* setter forte di H *)
ndefinition set_indX_8_high_reg ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λindxhigh':byte8.
- opt_map ?? (match m return aux_set_typing_opt byte8 with
+ opt_map … (match m return aux_set_typing_opt byte8 with
[ HC05 ⇒ None ?
| HC08 ⇒ Some ? set_indX_8_high_reg_HC08
| HCS08 ⇒ Some ? set_indX_8_high_reg_HC08
(* setter forte di H:X *)
ndefinition set_indX_16_reg ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λindx16:word16.
- opt_map ?? (match m return aux_set_typing_opt word16 with
+ opt_map … (match m return aux_set_typing_opt word16 with
[ HC05 ⇒ None ?
| HC08 ⇒ Some ? set_indX_16_reg_HC08
| HCS08 ⇒ Some ? set_indX_16_reg_HC08
(* setter forte di SP *)
ndefinition set_sp_reg ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λsp':word16.
- opt_map ?? (match m return aux_set_typing_opt word16 with
+ opt_map … (match m return aux_set_typing_opt word16 with
[ HC05 ⇒ Some ? set_sp_reg_HC05
| HC08 ⇒ Some ? set_sp_reg_HC08
| HCS08 ⇒ Some ? set_sp_reg_HC08
(* setter forte di SPC *)
ndefinition set_spc_reg ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λspc':word16.
- opt_map ?? (match m return aux_set_typing_opt word16 with
+ opt_map … (match m return aux_set_typing_opt word16 with
[ HC05 ⇒ None ?
| HC08 ⇒ None ?
| HCS08 ⇒ None ?
(* setter forte di memory mapped X *)
ndefinition set_x_map ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λxm':byte8.
- opt_map ?? (match m return aux_set_typing_opt byte8 with
+ opt_map … (match m return aux_set_typing_opt byte8 with
[ HC05 ⇒ None ?
| HC08 ⇒ None ?
| HCS08 ⇒ None ?
(* setter forte di memory mapped PS *)
ndefinition set_ps_map ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λpsm':byte8.
- opt_map ?? (match m return aux_set_typing_opt byte8 with
+ opt_map … (match m return aux_set_typing_opt byte8 with
[ HC05 ⇒ None ?
| HC08 ⇒ None ?
| HCS08 ⇒ None ?
(* setter forte di V *)
ndefinition set_v_flag ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λvfl':bool.
- opt_map ?? (match m return aux_set_typing_opt bool with
+ opt_map … (match m return aux_set_typing_opt bool with
[ HC05 ⇒ None ?
| HC08 ⇒ Some ? set_v_flag_HC08
| HCS08 ⇒ Some ? set_v_flag_HC08
(* setter forte di H *)
ndefinition set_h_flag ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λhfl':bool.
- opt_map ?? (match m return aux_set_typing_opt bool with
+ opt_map … (match m return aux_set_typing_opt bool with
[ HC05 ⇒ Some ? set_h_flag_HC05
| HC08 ⇒ Some ? set_h_flag_HC08
| HCS08 ⇒ Some ? set_h_flag_HC08
(* setter forte di I *)
ndefinition set_i_flag ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λifl':bool.
- opt_map ?? (match m return aux_set_typing_opt bool with
+ opt_map … (match m return aux_set_typing_opt bool with
[ HC05 ⇒ Some ? set_i_flag_HC05
| HC08 ⇒ Some ? set_i_flag_HC08
| HCS08 ⇒ Some ? set_i_flag_HC08
(* setter forte di N *)
ndefinition set_n_flag ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λnfl':bool.
- opt_map ?? (match m return aux_set_typing_opt bool with
+ opt_map … (match m return aux_set_typing_opt bool with
[ HC05 ⇒ Some ? set_n_flag_HC05
| HC08 ⇒ Some ? set_n_flag_HC08
| HCS08 ⇒ Some ? set_n_flag_HC08
(* setter forte di IRQ *)
ndefinition set_irq_flag ≝
λm:mcu_type.λt:memory_impl.λs:any_status m t.λirqfl':bool.
- opt_map ?? (match m return aux_set_typing_opt bool with
+ opt_map … (match m return aux_set_typing_opt bool with
[ HC05 ⇒ Some ? set_irq_flag_HC05
| HC08 ⇒ Some ? set_irq_flag_HC08
| HCS08 ⇒ Some ? set_irq_flag_HC08
[ None ⇒ match c2 with
[ None ⇒ true | Some _ ⇒ false ]
| Some c1' ⇒ match c2 with
- [ None ⇒ false | Some c2' ⇒ (eq_b8 (fst5T ????? c1') (fst5T ????? c2')) ⊗
- (eq_anyop m (snd5T ????? c1') (snd5T ????? c2')) ⊗
- (eq_instrmode (thd5T ????? c1') (thd5T ????? c2')) ⊗
- (eq_b8 (frth5T ????? c1') (frth5T ????? c2')) ⊗
- (eq_w16 (ffth5T ????? c1') (ffth5T ????? c2')) ]
+ [ None ⇒ false | Some c2' ⇒ (eq_b8 (fst5T … c1') (fst5T … c2')) ⊗
+ (eq_anyop m (snd5T … c1') (snd5T … c2')) ⊗
+ (eq_instrmode (thd5T … c1') (thd5T … c2')) ⊗
+ (eq_b8 (frth5T … c1') (frth5T … c2')) ⊗
+ (eq_w16 (ffth5T … c1') (ffth5T … c2')) ]
].
(* generalizzazione del confronto fra stati *)
with [ mk_alu_HC05 a _ _ _ _ _ _ _ _ _ _ _ _ ⇒ x1 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC05_destruct_2 :
with [ mk_alu_HC05 _ a _ _ _ _ _ _ _ _ _ _ _ ⇒ x2 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC05_destruct_3 :
with [ mk_alu_HC05 _ _ a _ _ _ _ _ _ _ _ _ _ ⇒ x3 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC05_destruct_4 :
with [ mk_alu_HC05 _ _ _ a _ _ _ _ _ _ _ _ _ ⇒ x4 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC05_destruct_5 :
with [ mk_alu_HC05 _ _ _ _ a _ _ _ _ _ _ _ _ ⇒ x5 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC05_destruct_6 :
with [ mk_alu_HC05 _ _ _ _ _ a _ _ _ _ _ _ _ ⇒ x6 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC05_destruct_7 :
with [ mk_alu_HC05 _ _ _ _ _ _ a _ _ _ _ _ _ ⇒ x7 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC05_destruct_8 :
with [ mk_alu_HC05 _ _ _ _ _ _ _ a _ _ _ _ _ ⇒ x8 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC05_destruct_9 :
with [ mk_alu_HC05 _ _ _ _ _ _ _ _ a _ _ _ _ ⇒ x9 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC05_destruct_10 :
with [ mk_alu_HC05 _ _ _ _ _ _ _ _ _ a _ _ _ ⇒ x10 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC05_destruct_11 :
with [ mk_alu_HC05 _ _ _ _ _ _ _ _ _ _ a _ _ ⇒ x11 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC05_destruct_12 :
with [ mk_alu_HC05 _ _ _ _ _ _ _ _ _ _ _ a _ ⇒ x12 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC05_destruct_13 :
with [ mk_alu_HC05 _ _ _ _ _ _ _ _ _ _ _ _ a ⇒ x13 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_eqaluHC05 : symmetricT alu_HC05 bool eq_alu_HC05.
nrewrite > (symmetric_eqbool x11 y11);
nrewrite > (symmetric_eqbool x12 y12);
nrewrite > (symmetric_eqbool x13 y13);
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma eqaluHC05_to_eq : ∀alu1,alu2.eq_alu_HC05 alu1 alu2 = true → alu1 = alu2.
(eq_bool x9 y9) ⊗ (eq_bool x10 y10) ⊗
(eq_bool x11 y11) ⊗ (eq_bool x12 y12) ⊗
(eq_bool x13 y13)) = true);
- nrewrite > (eqbool_to_eq ?? (andb_true_true_r ?? H));
- nletin H1 ≝ (andb_true_true_l ?? H);
- nrewrite > (eqbool_to_eq x12 y12 (andb_true_true_r ?? (andb_true_true_l ?? H)));
- nletin H2 ≝ (andb_true_true_l ?? H1);
- nrewrite > (eqbool_to_eq ?? (andb_true_true_r ?? H2));
- nletin H3 ≝ (andb_true_true_l ?? H2);
- nrewrite > (eqbool_to_eq ?? (andb_true_true_r ?? H3));
- nletin H4 ≝ (andb_true_true_l ?? H3);
- nrewrite > (eqbool_to_eq ?? (andb_true_true_r ?? H4));
- nletin H5 ≝ (andb_true_true_l ?? H4);
- nrewrite > (eqbool_to_eq ?? (andb_true_true_r ?? H5));
- nletin H6 ≝ (andb_true_true_l ?? H5);
- nrewrite > (eqw16_to_eq ?? (andb_true_true_r ?? H6));
- nletin H7 ≝ (andb_true_true_l ?? H6);
- nrewrite > (eqw16_to_eq ?? (andb_true_true_r ?? H7));
- nletin H8 ≝ (andb_true_true_l ?? H7);
- nrewrite > (eqw16_to_eq ?? (andb_true_true_r ?? H8));
- nletin H9 ≝ (andb_true_true_l ?? H8);
- nrewrite > (eqw16_to_eq ?? (andb_true_true_r ?? H9));
- nletin H10 ≝ (andb_true_true_l ?? H9);
- nrewrite > (eqw16_to_eq ?? (andb_true_true_r ?? H10));
- nletin H11 ≝ (andb_true_true_l ?? H10);
- nrewrite > (eqb8_to_eq ?? (andb_true_true_r ?? H11));
- nrewrite > (eqb8_to_eq ?? (andb_true_true_l ?? H11));
- napply (refl_eq ??).
+ nrewrite > (eqbool_to_eq … (andb_true_true_r … H));
+ nletin H1 ≝ (andb_true_true_l … H);
+ nrewrite > (eqbool_to_eq x12 y12 (andb_true_true_r … (andb_true_true_l … H)));
+ nletin H2 ≝ (andb_true_true_l … H1);
+ nrewrite > (eqbool_to_eq … (andb_true_true_r … H2));
+ nletin H3 ≝ (andb_true_true_l … H2);
+ nrewrite > (eqbool_to_eq … (andb_true_true_r … H3));
+ nletin H4 ≝ (andb_true_true_l … H3);
+ nrewrite > (eqbool_to_eq … (andb_true_true_r … H4));
+ nletin H5 ≝ (andb_true_true_l … H4);
+ nrewrite > (eqbool_to_eq … (andb_true_true_r … H5));
+ nletin H6 ≝ (andb_true_true_l … H5);
+ nrewrite > (eqw16_to_eq … (andb_true_true_r … H6));
+ nletin H7 ≝ (andb_true_true_l … H6);
+ nrewrite > (eqw16_to_eq … (andb_true_true_r … H7));
+ nletin H8 ≝ (andb_true_true_l … H7);
+ nrewrite > (eqw16_to_eq … (andb_true_true_r … H8));
+ nletin H9 ≝ (andb_true_true_l … H8);
+ nrewrite > (eqw16_to_eq … (andb_true_true_r … H9));
+ nletin H10 ≝ (andb_true_true_l … H9);
+ nrewrite > (eqw16_to_eq … (andb_true_true_r … H10));
+ nletin H11 ≝ (andb_true_true_l … H10);
+ nrewrite > (eqb8_to_eq … (andb_true_true_r … H11));
+ nrewrite > (eqb8_to_eq … (andb_true_true_l … H11));
+ napply refl_eq.
nqed.
nlemma eq_to_eqaluHC05 : ∀alu1,alu2.alu1 = alu2 → eq_alu_HC05 alu1 alu2 = true.
#x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12; #x13;
ncases alu2;
#y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12; #y13; #H;
- nrewrite > (aluHC05_destruct_1 ?????????????????????????? H);
- nrewrite > (aluHC05_destruct_2 ?????????????????????????? H);
- nrewrite > (aluHC05_destruct_3 ?????????????????????????? H);
- nrewrite > (aluHC05_destruct_4 ?????????????????????????? H);
- nrewrite > (aluHC05_destruct_5 ?????????????????????????? H);
- nrewrite > (aluHC05_destruct_6 ?????????????????????????? H);
- nrewrite > (aluHC05_destruct_7 ?????????????????????????? H);
- nrewrite > (aluHC05_destruct_8 ?????????????????????????? H);
- nrewrite > (aluHC05_destruct_9 ?????????????????????????? H);
- nrewrite > (aluHC05_destruct_10 ?????????????????????????? H);
- nrewrite > (aluHC05_destruct_11 ?????????????????????????? H);
- nrewrite > (aluHC05_destruct_12 ?????????????????????????? H);
- nrewrite > (aluHC05_destruct_13 ?????????????????????????? H);
+ nrewrite > (aluHC05_destruct_1 … H);
+ nrewrite > (aluHC05_destruct_2 … H);
+ nrewrite > (aluHC05_destruct_3 … H);
+ nrewrite > (aluHC05_destruct_4 … H);
+ nrewrite > (aluHC05_destruct_5 … H);
+ nrewrite > (aluHC05_destruct_6 … H);
+ nrewrite > (aluHC05_destruct_7 … H);
+ nrewrite > (aluHC05_destruct_8 … H);
+ nrewrite > (aluHC05_destruct_9 … H);
+ nrewrite > (aluHC05_destruct_10 … H);
+ nrewrite > (aluHC05_destruct_11 … H);
+ nrewrite > (aluHC05_destruct_12 … H);
+ nrewrite > (aluHC05_destruct_13 … H);
nchange with (
((eq_b8 y1 y1) ⊗ (eq_b8 y2 y2) ⊗
(eq_w16 y3 y3) ⊗ (eq_w16 y4 y4) ⊗
(eq_bool y9 y9) ⊗ (eq_bool y10 y10) ⊗
(eq_bool y11 y11) ⊗ (eq_bool y12 y12) ⊗
(eq_bool y13 y13)) = true);
- nrewrite > (eq_to_eqb8 y1 y1 (refl_eq ??));
- nrewrite > (eq_to_eqb8 y2 y2 (refl_eq ??));
- nrewrite > (eq_to_eqw16 y3 y3 (refl_eq ??));
- nrewrite > (eq_to_eqw16 y4 y4 (refl_eq ??));
- nrewrite > (eq_to_eqw16 y5 y5 (refl_eq ??));
- nrewrite > (eq_to_eqw16 y6 y6 (refl_eq ??));
- nrewrite > (eq_to_eqw16 y7 y7 (refl_eq ??));
- nrewrite > (eq_to_eqbool y8 y8 (refl_eq ??));
- nrewrite > (eq_to_eqbool y9 y9 (refl_eq ??));
- nrewrite > (eq_to_eqbool y10 y10 (refl_eq ??));
- nrewrite > (eq_to_eqbool y11 y11 (refl_eq ??));
- nrewrite > (eq_to_eqbool y12 y12 (refl_eq ??));
- nrewrite > (eq_to_eqbool y13 y13 (refl_eq ??));
- napply (refl_eq ??).
+ nrewrite > (eq_to_eqb8 y1 y1 (refl_eq …));
+ nrewrite > (eq_to_eqb8 y2 y2 (refl_eq …));
+ nrewrite > (eq_to_eqw16 y3 y3 (refl_eq …));
+ nrewrite > (eq_to_eqw16 y4 y4 (refl_eq …));
+ nrewrite > (eq_to_eqw16 y5 y5 (refl_eq …));
+ nrewrite > (eq_to_eqw16 y6 y6 (refl_eq …));
+ nrewrite > (eq_to_eqw16 y7 y7 (refl_eq …));
+ nrewrite > (eq_to_eqbool y8 y8 (refl_eq …));
+ nrewrite > (eq_to_eqbool y9 y9 (refl_eq …));
+ nrewrite > (eq_to_eqbool y10 y10 (refl_eq …));
+ nrewrite > (eq_to_eqbool y11 y11 (refl_eq …));
+ nrewrite > (eq_to_eqbool y12 y12 (refl_eq …));
+ nrewrite > (eq_to_eqbool y13 y13 (refl_eq …));
+ napply refl_eq.
nqed.
nlemma aluHC08_destruct_1 :
with [ mk_alu_HC08 a _ _ _ _ _ _ _ _ _ _ _ ⇒ x1 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC08_destruct_2 :
with [ mk_alu_HC08 _ a _ _ _ _ _ _ _ _ _ _ ⇒ x2 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC08_destruct_3 :
with [ mk_alu_HC08 _ _ a _ _ _ _ _ _ _ _ _ ⇒ x3 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC08_destruct_4 :
with [ mk_alu_HC08 _ _ _ a _ _ _ _ _ _ _ _ ⇒ x4 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC08_destruct_5 :
with [ mk_alu_HC08 _ _ _ _ a _ _ _ _ _ _ _ ⇒ x5 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC08_destruct_6 :
with [ mk_alu_HC08 _ _ _ _ _ a _ _ _ _ _ _ ⇒ x6 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC08_destruct_7 :
with [ mk_alu_HC08 _ _ _ _ _ _ a _ _ _ _ _ ⇒ x7 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC08_destruct_8 :
with [ mk_alu_HC08 _ _ _ _ _ _ _ a _ _ _ _ ⇒ x8 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC08_destruct_9 :
with [ mk_alu_HC08 _ _ _ _ _ _ _ _ a _ _ _ ⇒ x9 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC08_destruct_10 :
with [ mk_alu_HC08 _ _ _ _ _ _ _ _ _ a _ _ ⇒ x10 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC08_destruct_11 :
with [ mk_alu_HC08 _ _ _ _ _ _ _ _ _ _ a _ ⇒ x11 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluHC08_destruct_12 :
with [ mk_alu_HC08 _ _ _ _ _ _ _ _ _ _ _ a ⇒ x12 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_eqaluHC08 : symmetricT alu_HC08 bool eq_alu_HC08.
nrewrite > (symmetric_eqbool x10 y10);
nrewrite > (symmetric_eqbool x11 y11);
nrewrite > (symmetric_eqbool x12 y12);
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma eqaluHC08_to_eq : ∀alu1,alu2.eq_alu_HC08 alu1 alu2 = true → alu1 = alu2.
((eq_b8 x1 y1) ⊗ (eq_b8 x2 y2) ⊗ (eq_b8 x3 y3) ⊗ (eq_w16 x4 y4) ⊗
(eq_w16 x5 y5) ⊗ (eq_bool x6 y6) ⊗ (eq_bool x7 y7) ⊗ (eq_bool x8 y8) ⊗
(eq_bool x9 y9) ⊗ (eq_bool x10 y10) ⊗ (eq_bool x11 y11) ⊗ (eq_bool x12 y12)) = true);
- nrewrite > (eqbool_to_eq ?? (andb_true_true_r ?? H));
- nletin H1 ≝ (andb_true_true_l ?? H);
- nrewrite > (eqbool_to_eq ?? (andb_true_true_r ?? H1));
- nletin H2 ≝ (andb_true_true_l ?? H1);
- nrewrite > (eqbool_to_eq ?? (andb_true_true_r ?? H2));
- nletin H3 ≝ (andb_true_true_l ?? H2);
- nrewrite > (eqbool_to_eq ?? (andb_true_true_r ?? H3));
- nletin H4 ≝ (andb_true_true_l ?? H3);
- nrewrite > (eqbool_to_eq ?? (andb_true_true_r ?? H4));
- nletin H5 ≝ (andb_true_true_l ?? H4);
- nrewrite > (eqbool_to_eq ?? (andb_true_true_r ?? H5));
- nletin H6 ≝ (andb_true_true_l ?? H5);
- nrewrite > (eqbool_to_eq ?? (andb_true_true_r ?? H6));
- nletin H7 ≝ (andb_true_true_l ?? H6);
- nrewrite > (eqw16_to_eq ?? (andb_true_true_r ?? H7));
- nletin H8 ≝ (andb_true_true_l ?? H7);
- nrewrite > (eqw16_to_eq ?? (andb_true_true_r ?? H8));
- nletin H9 ≝ (andb_true_true_l ?? H8);
- nrewrite > (eqb8_to_eq ?? (andb_true_true_r ?? H9));
- nletin H10 ≝ (andb_true_true_l ?? H9);
- nrewrite > (eqb8_to_eq ?? (andb_true_true_r ?? H10));
- nrewrite > (eqb8_to_eq ?? (andb_true_true_l ?? H10));
- napply (refl_eq ??).
+ nrewrite > (eqbool_to_eq … (andb_true_true_r … H));
+ nletin H1 ≝ (andb_true_true_l … H);
+ nrewrite > (eqbool_to_eq … (andb_true_true_r … H1));
+ nletin H2 ≝ (andb_true_true_l … H1);
+ nrewrite > (eqbool_to_eq … (andb_true_true_r … H2));
+ nletin H3 ≝ (andb_true_true_l … H2);
+ nrewrite > (eqbool_to_eq … (andb_true_true_r … H3));
+ nletin H4 ≝ (andb_true_true_l … H3);
+ nrewrite > (eqbool_to_eq … (andb_true_true_r … H4));
+ nletin H5 ≝ (andb_true_true_l … H4);
+ nrewrite > (eqbool_to_eq … (andb_true_true_r … H5));
+ nletin H6 ≝ (andb_true_true_l … H5);
+ nrewrite > (eqbool_to_eq … (andb_true_true_r … H6));
+ nletin H7 ≝ (andb_true_true_l … H6);
+ nrewrite > (eqw16_to_eq … (andb_true_true_r … H7));
+ nletin H8 ≝ (andb_true_true_l … H7);
+ nrewrite > (eqw16_to_eq … (andb_true_true_r … H8));
+ nletin H9 ≝ (andb_true_true_l … H8);
+ nrewrite > (eqb8_to_eq … (andb_true_true_r … H9));
+ nletin H10 ≝ (andb_true_true_l … H9);
+ nrewrite > (eqb8_to_eq … (andb_true_true_r … H10));
+ nrewrite > (eqb8_to_eq … (andb_true_true_l … H10));
+ napply refl_eq.
nqed.
nlemma eq_to_eqaluHC08 : ∀alu1,alu2.alu1 = alu2 → eq_alu_HC08 alu1 alu2 = true.
#x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12;
ncases alu2;
#y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12; #H;
- nrewrite > (aluHC08_destruct_1 ???????????????????????? H);
- nrewrite > (aluHC08_destruct_2 ???????????????????????? H);
- nrewrite > (aluHC08_destruct_3 ???????????????????????? H);
- nrewrite > (aluHC08_destruct_4 ???????????????????????? H);
- nrewrite > (aluHC08_destruct_5 ???????????????????????? H);
- nrewrite > (aluHC08_destruct_6 ???????????????????????? H);
- nrewrite > (aluHC08_destruct_7 ???????????????????????? H);
- nrewrite > (aluHC08_destruct_8 ???????????????????????? H);
- nrewrite > (aluHC08_destruct_9 ???????????????????????? H);
- nrewrite > (aluHC08_destruct_10 ???????????????????????? H);
- nrewrite > (aluHC08_destruct_11 ???????????????????????? H);
- nrewrite > (aluHC08_destruct_12 ???????????????????????? H);
+ nrewrite > (aluHC08_destruct_1 … H);
+ nrewrite > (aluHC08_destruct_2 … H);
+ nrewrite > (aluHC08_destruct_3 … H);
+ nrewrite > (aluHC08_destruct_4 … H);
+ nrewrite > (aluHC08_destruct_5 … H);
+ nrewrite > (aluHC08_destruct_6 … H);
+ nrewrite > (aluHC08_destruct_7 … H);
+ nrewrite > (aluHC08_destruct_8 … H);
+ nrewrite > (aluHC08_destruct_9 … H);
+ nrewrite > (aluHC08_destruct_10 … H);
+ nrewrite > (aluHC08_destruct_11 … H);
+ nrewrite > (aluHC08_destruct_12 … H);
nchange with (
((eq_b8 y1 y1) ⊗ (eq_b8 y2 y2) ⊗ (eq_b8 y3 y3) ⊗ (eq_w16 y4 y4) ⊗
(eq_w16 y5 y5) ⊗ (eq_bool y6 y6) ⊗ (eq_bool y7 y7) ⊗ (eq_bool y8 y8) ⊗
(eq_bool y9 y9) ⊗ (eq_bool y10 y10) ⊗ (eq_bool y11 y11) ⊗ (eq_bool y12 y12)) = true);
- nrewrite > (eq_to_eqb8 y1 y1 (refl_eq ??));
- nrewrite > (eq_to_eqb8 y2 y2 (refl_eq ??));
- nrewrite > (eq_to_eqb8 y3 y3 (refl_eq ??));
- nrewrite > (eq_to_eqw16 y4 y4 (refl_eq ??));
- nrewrite > (eq_to_eqw16 y5 y5 (refl_eq ??));
- nrewrite > (eq_to_eqbool y6 y6 (refl_eq ??));
- nrewrite > (eq_to_eqbool y7 y7 (refl_eq ??));
- nrewrite > (eq_to_eqbool y8 y8 (refl_eq ??));
- nrewrite > (eq_to_eqbool y9 y9 (refl_eq ??));
- nrewrite > (eq_to_eqbool y10 y10 (refl_eq ??));
- nrewrite > (eq_to_eqbool y11 y11 (refl_eq ??));
- nrewrite > (eq_to_eqbool y12 y12 (refl_eq ??));
- napply (refl_eq ??).
+ nrewrite > (eq_to_eqb8 y1 y1 (refl_eq …));
+ nrewrite > (eq_to_eqb8 y2 y2 (refl_eq …));
+ nrewrite > (eq_to_eqb8 y3 y3 (refl_eq …));
+ nrewrite > (eq_to_eqw16 y4 y4 (refl_eq …));
+ nrewrite > (eq_to_eqw16 y5 y5 (refl_eq …));
+ nrewrite > (eq_to_eqbool y6 y6 (refl_eq …));
+ nrewrite > (eq_to_eqbool y7 y7 (refl_eq …));
+ nrewrite > (eq_to_eqbool y8 y8 (refl_eq …));
+ nrewrite > (eq_to_eqbool y9 y9 (refl_eq …));
+ nrewrite > (eq_to_eqbool y10 y10 (refl_eq …));
+ nrewrite > (eq_to_eqbool y11 y11 (refl_eq …));
+ nrewrite > (eq_to_eqbool y12 y12 (refl_eq …));
+ napply refl_eq.
nqed.
nlemma aluRS08_destruct_1 :
with [ mk_alu_RS08 a _ _ _ _ _ _ _ ⇒ x1 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluRS08_destruct_2 :
with [ mk_alu_RS08 _ a _ _ _ _ _ _ ⇒ x2 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluRS08_destruct_3 :
with [ mk_alu_RS08 _ _ a _ _ _ _ _ ⇒ x3 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluRS08_destruct_4 :
with [ mk_alu_RS08 _ _ _ a _ _ _ _ ⇒ x4 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluRS08_destruct_5 :
with [ mk_alu_RS08 _ _ _ _ a _ _ _ ⇒ x5 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluRS08_destruct_6 :
with [ mk_alu_RS08 _ _ _ _ _ a _ _ ⇒ x6 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluRS08_destruct_7 :
with [ mk_alu_RS08 _ _ _ _ _ _ a _ ⇒ x7 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma aluRS08_destruct_8 :
with [ mk_alu_RS08 _ _ _ _ _ _ _ a ⇒ x8 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_eqaluRS08 : symmetricT alu_RS08 bool eq_alu_RS08.
nrewrite > (symmetric_eqb8 x6 y6);
nrewrite > (symmetric_eqbool x7 y7);
nrewrite > (symmetric_eqbool x8 y8);
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma eqaluRS08_to_eq : ∀alu1,alu2.eq_alu_RS08 alu1 alu2 = true → alu1 = alu2.
(eq_w16 x3 y3) ⊗ (eq_w16 x4 y4) ⊗
(eq_b8 x5 y5) ⊗ (eq_b8 x6 y6) ⊗
(eq_bool x7 y7) ⊗ (eq_bool x8 y8)) = true);
- nrewrite > (eqbool_to_eq ?? (andb_true_true_r ?? H));
- nletin H1 ≝ (andb_true_true_l ?? H);
- nrewrite > (eqbool_to_eq ?? (andb_true_true_r ?? H1));
- nletin H2 ≝ (andb_true_true_l ?? H1);
- nrewrite > (eqb8_to_eq ?? (andb_true_true_r ?? H2));
- nletin H3 ≝ (andb_true_true_l ?? H2);
- nrewrite > (eqb8_to_eq ?? (andb_true_true_r ?? H3));
- nletin H4 ≝ (andb_true_true_l ?? H3);
- nrewrite > (eqw16_to_eq ?? (andb_true_true_r ?? H4));
- nletin H5 ≝ (andb_true_true_l ?? H4);
- nrewrite > (eqw16_to_eq ?? (andb_true_true_r ?? H5));
- nletin H6 ≝ (andb_true_true_l ?? H5);
- nrewrite > (eqw16_to_eq ?? (andb_true_true_r ?? H6));
- nrewrite > (eqb8_to_eq ?? (andb_true_true_l ?? H6));
- napply (refl_eq ??).
+ nrewrite > (eqbool_to_eq … (andb_true_true_r … H));
+ nletin H1 ≝ (andb_true_true_l … H);
+ nrewrite > (eqbool_to_eq … (andb_true_true_r … H1));
+ nletin H2 ≝ (andb_true_true_l … H1);
+ nrewrite > (eqb8_to_eq … (andb_true_true_r … H2));
+ nletin H3 ≝ (andb_true_true_l … H2);
+ nrewrite > (eqb8_to_eq … (andb_true_true_r … H3));
+ nletin H4 ≝ (andb_true_true_l … H3);
+ nrewrite > (eqw16_to_eq … (andb_true_true_r … H4));
+ nletin H5 ≝ (andb_true_true_l … H4);
+ nrewrite > (eqw16_to_eq … (andb_true_true_r … H5));
+ nletin H6 ≝ (andb_true_true_l … H5);
+ nrewrite > (eqw16_to_eq … (andb_true_true_r … H6));
+ nrewrite > (eqb8_to_eq … (andb_true_true_l … H6));
+ napply refl_eq.
nqed.
nlemma eq_to_eqaluRS08 : ∀alu1,alu2.alu1 = alu2 → eq_alu_RS08 alu1 alu2 = true.
#x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8;
ncases alu2;
#y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #H;
- nrewrite > (aluRS08_destruct_1 ???????????????? H);
- nrewrite > (aluRS08_destruct_2 ???????????????? H);
- nrewrite > (aluRS08_destruct_3 ???????????????? H);
- nrewrite > (aluRS08_destruct_4 ???????????????? H);
- nrewrite > (aluRS08_destruct_5 ???????????????? H);
- nrewrite > (aluRS08_destruct_6 ???????????????? H);
- nrewrite > (aluRS08_destruct_7 ???????????????? H);
- nrewrite > (aluRS08_destruct_8 ???????????????? H);
+ nrewrite > (aluRS08_destruct_1 … H);
+ nrewrite > (aluRS08_destruct_2 … H);
+ nrewrite > (aluRS08_destruct_3 … H);
+ nrewrite > (aluRS08_destruct_4 … H);
+ nrewrite > (aluRS08_destruct_5 … H);
+ nrewrite > (aluRS08_destruct_6 … H);
+ nrewrite > (aluRS08_destruct_7 … H);
+ nrewrite > (aluRS08_destruct_8 … H);
nchange with (
((eq_b8 y1 y1) ⊗ (eq_w16 y2 y2) ⊗
(eq_w16 y3 y3) ⊗ (eq_w16 y4 y4) ⊗
(eq_b8 y5 y5) ⊗ (eq_b8 y6 y6) ⊗
(eq_bool y7 y7) ⊗ (eq_bool y8 y8)) = true);
- nrewrite > (eq_to_eqb8 y1 y1 (refl_eq ??));
- nrewrite > (eq_to_eqw16 y2 y2 (refl_eq ??));
- nrewrite > (eq_to_eqw16 y3 y3 (refl_eq ??));
- nrewrite > (eq_to_eqw16 y4 y4 (refl_eq ??));
- nrewrite > (eq_to_eqb8 y5 y5 (refl_eq ??));
- nrewrite > (eq_to_eqb8 y6 y6 (refl_eq ??));
- nrewrite > (eq_to_eqbool y7 y7 (refl_eq ??));
- nrewrite > (eq_to_eqbool y8 y8 (refl_eq ??));
- napply (refl_eq ??).
+ nrewrite > (eq_to_eqb8 y1 y1 (refl_eq …));
+ nrewrite > (eq_to_eqw16 y2 y2 (refl_eq …));
+ nrewrite > (eq_to_eqw16 y3 y3 (refl_eq …));
+ nrewrite > (eq_to_eqw16 y4 y4 (refl_eq …));
+ nrewrite > (eq_to_eqb8 y5 y5 (refl_eq …));
+ nrewrite > (eq_to_eqb8 y6 y6 (refl_eq …));
+ nrewrite > (eq_to_eqbool y7 y7 (refl_eq …));
+ nrewrite > (eq_to_eqbool y8 y8 (refl_eq …));
+ napply refl_eq.
nqed.
nlemma symmetric_eqclk : ∀mcu,clk1,clk2.eq_clk mcu clk1 clk2 = eq_clk mcu clk2 clk1.
#mcu; #clk1; #clk2;
ncases clk1;
ncases clk2;
- ##[ ##1: napply (refl_eq ??)
- ##| ##2,3: nnormalize; #H; napply (refl_eq ??)
+ ##[ ##1: napply refl_eq
+ ##| ##2,3: nnormalize; #H; napply refl_eq
##| ##4: #p1; ncases p1; #x1; #x2; #x3; #x4; #x5;
#p2; ncases p2; #y1; #y2; #y3; #y4; #y5;
nchange with (
nrewrite > (symmetric_eqinstrmode x3 y3);
nrewrite > (symmetric_eqb8 x4 y4);
nrewrite > (symmetric_eqw16 x5 y5);
- napply (refl_eq ??)
+ napply refl_eq
##]
nqed.
#mcu; #clk1; #clk2;
ncases clk1;
ncases clk2;
- ##[ ##1: nnormalize; #H; napply (refl_eq ??)
- ##| ##2,3: nnormalize; #H; #H1; napply (bool_destruct ??? H1)
+ ##[ ##1: nnormalize; #H; napply refl_eq
+ ##| ##2,3: nnormalize; #H; #H1; napply (bool_destruct … H1)
##| ##4: #p1; ncases p1; #x1; #x2; #x3; #x4; #x5;
#p2; ncases p2; #y1; #y2; #y3; #y4; #y5; #H;
nchange in H:(%) with (
((eq_b8 y1 x1) ⊗ (eq_anyop ? y2 x2) ⊗ (eq_instrmode y3 x3) ⊗ (eq_b8 y4 x4) ⊗ (eq_w16 y5 x5)) = true);
- nrewrite > (eqw16_to_eq ?? (andb_true_true_r ?? H));
- nletin H1 ≝ (andb_true_true_l ?? H);
- nrewrite > (eqb8_to_eq ?? (andb_true_true_r ?? H1));
- nletin H2 ≝ (andb_true_true_l ?? H1);
- nrewrite > (eqinstrmode_to_eq ?? (andb_true_true_r ?? H2));
- nletin H3 ≝ (andb_true_true_l ?? H2);
- nrewrite > (eqanyop_to_eq ??? (andb_true_true_r ?? H3));
- nrewrite > (eqb8_to_eq ?? (andb_true_true_l ?? H3));
- napply (refl_eq ??)
+ nrewrite > (eqw16_to_eq … (andb_true_true_r … H));
+ nletin H1 ≝ (andb_true_true_l … H);
+ nrewrite > (eqb8_to_eq … (andb_true_true_r … H1));
+ nletin H2 ≝ (andb_true_true_l … H1);
+ nrewrite > (eqinstrmode_to_eq … (andb_true_true_r … H2));
+ nletin H3 ≝ (andb_true_true_l … H2);
+ nrewrite > (eqanyop_to_eq … (andb_true_true_r … H3));
+ nrewrite > (eqb8_to_eq … (andb_true_true_l … H3));
+ napply refl_eq
##]
nqed.
#mcu; #clk1; #clk2;
ncases clk1;
ncases clk2;
- ##[ ##1: nnormalize; #H; napply (refl_eq ??)
- ##| ##2: nnormalize; #p; #H1; nelim (option_destruct_none_some ?? H1)
- ##| ##3: nnormalize; #p; #H1; nelim (option_destruct_some_none ?? H1)
+ ##[ ##1: nnormalize; #H; napply refl_eq
+ ##| ##2: nnormalize; #p; #H1; nelim (option_destruct_none_some … H1)
+ ##| ##3: nnormalize; #p; #H1; nelim (option_destruct_some_none … H1)
##| ##4: #p1; ncases p1; #x1; #x2; #x3; #x4; #x5;
#p2; ncases p2; #y1; #y2; #y3; #y4; #y5; #H;
- nrewrite > (quintuple_destruct_1 ??????????????? (option_destruct_some_some ??? H));
- nrewrite > (quintuple_destruct_2 ??????????????? (option_destruct_some_some ??? H));
- nrewrite > (quintuple_destruct_3 ??????????????? (option_destruct_some_some ??? H));
- nrewrite > (quintuple_destruct_4 ??????????????? (option_destruct_some_some ??? H));
- nrewrite > (quintuple_destruct_5 ??????????????? (option_destruct_some_some ??? H));
+ nrewrite > (quintuple_destruct_1 … (option_destruct_some_some … H));
+ nrewrite > (quintuple_destruct_2 … (option_destruct_some_some … H));
+ nrewrite > (quintuple_destruct_3 … (option_destruct_some_some … H));
+ nrewrite > (quintuple_destruct_4 … (option_destruct_some_some … H));
+ nrewrite > (quintuple_destruct_5 … (option_destruct_some_some … H));
nchange with (
((eq_b8 x1 x1) ⊗ (eq_anyop ? x2 x2) ⊗ (eq_instrmode x3 x3) ⊗ (eq_b8 x4 x4) ⊗ (eq_w16 x5 x5)) = true);
- nrewrite > (eq_to_eqb8 x1 x1 (refl_eq ??));
+ nrewrite > (eq_to_eqb8 x1 x1 (refl_eq …));
nrewrite > (eq_to_eqanyop mcu x2 x2 (refl_eq ? x2));
- nrewrite > (eq_to_eqinstrmode x3 x3 (refl_eq ??));
- nrewrite > (eq_to_eqb8 x4 x4 (refl_eq ??));
- nrewrite > (eq_to_eqw16 x5 x5 (refl_eq ??));
+ nrewrite > (eq_to_eqinstrmode x3 x3 (refl_eq …));
+ nrewrite > (eq_to_eqb8 x4 x4 (refl_eq …));
+ nrewrite > (eq_to_eqw16 x5 x5 (refl_eq …));
nnormalize;
- napply (refl_eq ??)
+ napply refl_eq
##]
nqed.
forall_memory_ranged t chk1 chk2 mem1 mem2 addrl =
forall_memory_ranged t chk2 chk1 mem2 mem1 addrl.
#t; #chk1; #chk2; #mem1; #mem2; #addrl;
- napply (list_ind word16 ??? addrl);
- ##[ ##1: nnormalize; napply (refl_eq ??)
+ napply (list_ind word16 … addrl);
+ ##[ ##1: nnormalize; napply refl_eq
##| ##2: #a; #l; #H;
nchange with (
((eq_option byte8 (mem_read t mem1 chk1 a)
(forall_memory_ranged t chk2 chk1 mem2 mem1 l)));
nrewrite > H;
nrewrite > (symmetric_eqoption ? (mem_read t mem1 chk1 a) (mem_read t mem2 chk2 a) eq_b8 symmetric_eqb8);
- napply (refl_eq ??)
+ napply refl_eq
##]
nqed.
with [ mk_any_status a _ _ _ ⇒ x1 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma anystatus_destruct_2 :
with [ mk_any_status _ a _ _ ⇒ x2 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma anystatus_destruct_3 :
with [ mk_any_status _ _ a _ ⇒ x3 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma anystatus_destruct_4 :
with [ mk_any_status _ _ _ a ⇒ x4 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_eqstatus :
##]
nrewrite > (symmetric_forallmemoryranged t x3 y3 x2 y2 addrl);
nrewrite > (symmetric_eqclk ? x4 y4);
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma eqstatus_to_eq :
#s2; ncases s2; #y1; #y2; #y3; #y4; #H;
nchange in H:(%) with (
((eq_alu_HC05 x1 y1) ⊗ (forall_memory_ranged t x3 y3 x2 y2 addrl) ⊗ (eq_clk ? x4 y4)) = true);
- nrewrite > (eqaluHC05_to_eq ?? (andb_true_true_l ?? (andb_true_true_l ?? H)))
+ nrewrite > (eqaluHC05_to_eq … (andb_true_true_l … (andb_true_true_l … H)))
##| ##2,3: ncases s1; #x1; #x2; #x3; #x4;
#s2; ncases s2; #y1; #y2; #y3; #y4; #H;
nchange in H:(%) with (
((eq_alu_HC08 x1 y1) ⊗ (forall_memory_ranged t x3 y3 x2 y2 addrl) ⊗ (eq_clk ? x4 y4)) = true);
- nrewrite > (eqaluHC08_to_eq ?? (andb_true_true_l ?? (andb_true_true_l ?? H)))
+ nrewrite > (eqaluHC08_to_eq … (andb_true_true_l … (andb_true_true_l … H)))
##| ##4: ncases s1; #x1; #x2; #x3; #x4;
#s2; ncases s2; #y1; #y2; #y3; #y4; #H;
nchange in H:(%) with (
((eq_alu_RS08 x1 y1) ⊗ (forall_memory_ranged t x3 y3 x2 y2 addrl) ⊗ (eq_clk ? x4 y4)) = true);
- nrewrite > (eqaluRS08_to_eq ?? (andb_true_true_l ?? (andb_true_true_l ?? H)))
+ nrewrite > (eqaluRS08_to_eq … (andb_true_true_l … (andb_true_true_l … H)))
##]
nchange with ((y1 = y1) ∧ (x4 = y4) ∧ (forall_memory_ranged t x3 y3 x2 y2 addrl = true));
- nrewrite > (andb_true_true_r ?? (andb_true_true_l ?? H));
- nrewrite > (eqclk_to_eq ??? (andb_true_true_r ?? H));
- napply (conj ?? (conj ?? (refl_eq ??) (refl_eq ??)) (refl_eq ??)).
+ nrewrite > (andb_true_true_r … (andb_true_true_l … H));
+ nrewrite > (eqclk_to_eq … (andb_true_true_r … H));
+ napply (conj … (conj … (refl_eq …) (refl_eq …)) (refl_eq …)).
nqed.
nlemma eq_to_eqstatus_strong :
#s2; ncases s2; #y1; #y2; #y3; #y4; #H;
nchange with (
((eq_alu_HC05 x1 y1) ⊗ (forall_memory_ranged t x3 y3 x2 y2 addrl) ⊗ (eq_clk ? x4 y4)) = true);
- nrewrite > (anystatus_destruct_1 ?????????? H);
- nrewrite > (eq_to_eqaluHC05 y1 y1 (refl_eq ??))
+ nrewrite > (anystatus_destruct_1 … H);
+ nrewrite > (eq_to_eqaluHC05 y1 y1 (refl_eq …))
##| ##2,3: #s1; ncases s1; #x1; #x2; #x3; #x4;
#s2; ncases s2; #y1; #y2; #y3; #y4; #H;
nchange with (
((eq_alu_HC08 x1 y1) ⊗ (forall_memory_ranged t x3 y3 x2 y2 addrl) ⊗ (eq_clk ? x4 y4)) = true);
- nrewrite > (anystatus_destruct_1 ?????????? H);
- nrewrite > (eq_to_eqaluHC08 y1 y1 (refl_eq ??))
+ nrewrite > (anystatus_destruct_1 … H);
+ nrewrite > (eq_to_eqaluHC08 y1 y1 (refl_eq …))
##| ##4: #s1; ncases s1; #x1; #x2; #x3; #x4;
#s2; ncases s2; #y1; #y2; #y3; #y4; #H;
nchange with (
((eq_alu_RS08 x1 y1) ⊗ (forall_memory_ranged t x3 y3 x2 y2 addrl) ⊗ (eq_clk ? x4 y4)) = true);
- nrewrite > (anystatus_destruct_1 ?????????? H);
- nrewrite > (eq_to_eqaluRS08 y1 y1 (refl_eq ??))
+ nrewrite > (anystatus_destruct_1 … H);
+ nrewrite > (eq_to_eqaluRS08 y1 y1 (refl_eq …))
##]
- nrewrite > (anystatus_destruct_2 ?????????? H);
- nrewrite > (anystatus_destruct_3 ?????????? H);
- nrewrite > (anystatus_destruct_4 ?????????? H);
- nrewrite > (eq_to_eqclk ? y4 y4 (refl_eq ??));
- nchange with ((forall_memory_ranged ?????? ⊗ true) =true);
+ nrewrite > (anystatus_destruct_2 … H);
+ nrewrite > (anystatus_destruct_3 … H);
+ nrewrite > (anystatus_destruct_4 … H);
+ nrewrite > (eq_to_eqclk ? y4 y4 (refl_eq …));
+ nchange with ((forall_memory_ranged … ⊗ true) =true);
nrewrite > (symmetric_andbool (forall_memory_ranged t y3 y3 y2 y2 addrl) true);
- nchange with (forall_memory_ranged ?????? = true);
- napply (list_ind word16 ??? addrl);
- ##[ ##1,3,5,7: nnormalize; napply (refl_eq ??)
+ nchange with (forall_memory_ranged … = true);
+ napply (list_ind word16 … addrl);
+ ##[ ##1,3,5,7: nnormalize; napply refl_eq
##| ##2,4,6,8: #a; #l'; #H;
nchange with (
((eq_option byte8 (mem_read t y2 y3 a)
(mem_read t y2 y3 a) eq_b8) ⊗
(forall_memory_ranged t y3 y3 y2 y2 l')) = true);
nrewrite > H;
- nrewrite > (eq_to_eqoption ? (mem_read t y2 y3 a) (mem_read t y2 y3 a) eq_b8 eq_to_eqb8 (refl_eq ??));
+ nrewrite > (eq_to_eqoption ? (mem_read t y2 y3 a) (mem_read t y2 y3 a) eq_b8 eq_to_eqb8 (refl_eq …));
nnormalize;
- napply (refl_eq ??)
+ napply refl_eq
##]
nqed.
nchange with (((eq_alu_HC05 x1 y1) ⊗ (forall_memory_ranged t x3 y3 x2 y2 addrl) ⊗ (eq_clk ? x4 y4)) = true);
nchange in H:(%) with (x1 = y1);
nrewrite > H;
- nrewrite > (eq_to_eqaluHC05 y1 y1 (refl_eq ??))
+ nrewrite > (eq_to_eqaluHC05 y1 y1 (refl_eq …))
##| ##2,3: #s1; ncases s1; #x1; #x2; #x3; #x4;
#s2; ncases s2; #y1; #y2; #y3; #y4; #H; #H1; #H2;
nchange with (((eq_alu_HC08 x1 y1) ⊗ (forall_memory_ranged t x3 y3 x2 y2 addrl) ⊗ (eq_clk ? x4 y4)) = true);
nchange in H:(%) with (x1 = y1);
nrewrite > H;
- nrewrite > (eq_to_eqaluHC08 y1 y1 (refl_eq ??))
+ nrewrite > (eq_to_eqaluHC08 y1 y1 (refl_eq …))
##| ##4: #s1; ncases s1; #x1; #x2; #x3; #x4;
#s2; ncases s2; #y1; #y2; #y3; #y4; #H; #H1; #H2;
nchange with (((eq_alu_RS08 x1 y1) ⊗ (forall_memory_ranged t x3 y3 x2 y2 addrl) ⊗ (eq_clk ? x4 y4)) = true);
nchange in H:(%) with (x1 = y1);
nrewrite > H;
- nrewrite > (eq_to_eqaluRS08 y1 y1 (refl_eq ??))
+ nrewrite > (eq_to_eqaluRS08 y1 y1 (refl_eq …))
##]
nchange in H2:(%) with (forall_memory_ranged t x3 y3 x2 y2 addrl = true);
nrewrite > H2;
nchange in H1:(%) with (x4 = y4);
nrewrite > H1;
- nrewrite > (eq_to_eqclk ? y4 y4 (refl_eq ??));
+ nrewrite > (eq_to_eqclk ? y4 y4 (refl_eq …));
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
ndefinition opcode_table_HC05_1 ≝
[
- quadruple ???? (anyOP HC05 ADC) MODE_IMM1 (Byte 〈xA,x9〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC05 ADC) MODE_DIR1 (Byte 〈xB,x9〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 ADC) MODE_DIR2 (Byte 〈xC,x9〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 ADC) MODE_IX2 (Byte 〈xD,x9〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 ADC) MODE_IX1 (Byte 〈xE,x9〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 ADC) MODE_IX0 (Byte 〈xF,x9〉) 〈x0,x4〉
+ quadruple … (anyOP HC05 ADC) MODE_IMM1 (Byte 〈xA,x9〉) 〈x0,x2〉
+; quadruple … (anyOP HC05 ADC) MODE_DIR1 (Byte 〈xB,x9〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 ADC) MODE_DIR2 (Byte 〈xC,x9〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 ADC) MODE_IX2 (Byte 〈xD,x9〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 ADC) MODE_IX1 (Byte 〈xE,x9〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 ADC) MODE_IX0 (Byte 〈xF,x9〉) 〈x0,x4〉
].
ndefinition opcode_table_HC05_2 ≝
[
- quadruple ???? (anyOP HC05 ADD) MODE_IMM1 (Byte 〈xA,xB〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC05 ADD) MODE_DIR1 (Byte 〈xB,xB〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 ADD) MODE_DIR2 (Byte 〈xC,xB〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 ADD) MODE_IX2 (Byte 〈xD,xB〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 ADD) MODE_IX1 (Byte 〈xE,xB〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 ADD) MODE_IX0 (Byte 〈xF,xB〉) 〈x0,x3〉
+ quadruple … (anyOP HC05 ADD) MODE_IMM1 (Byte 〈xA,xB〉) 〈x0,x2〉
+; quadruple … (anyOP HC05 ADD) MODE_DIR1 (Byte 〈xB,xB〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 ADD) MODE_DIR2 (Byte 〈xC,xB〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 ADD) MODE_IX2 (Byte 〈xD,xB〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 ADD) MODE_IX1 (Byte 〈xE,xB〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 ADD) MODE_IX0 (Byte 〈xF,xB〉) 〈x0,x3〉
].
ndefinition opcode_table_HC05_3 ≝
[
- quadruple ???? (anyOP HC05 AND) MODE_IMM1 (Byte 〈xA,x4〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC05 AND) MODE_DIR1 (Byte 〈xB,x4〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 AND) MODE_DIR2 (Byte 〈xC,x4〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 AND) MODE_IX2 (Byte 〈xD,x4〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 AND) MODE_IX1 (Byte 〈xE,x4〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 AND) MODE_IX0 (Byte 〈xF,x4〉) 〈x0,x3〉
+ quadruple … (anyOP HC05 AND) MODE_IMM1 (Byte 〈xA,x4〉) 〈x0,x2〉
+; quadruple … (anyOP HC05 AND) MODE_DIR1 (Byte 〈xB,x4〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 AND) MODE_DIR2 (Byte 〈xC,x4〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 AND) MODE_IX2 (Byte 〈xD,x4〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 AND) MODE_IX1 (Byte 〈xE,x4〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 AND) MODE_IX0 (Byte 〈xF,x4〉) 〈x0,x3〉
].
ndefinition opcode_table_HC05_4 ≝
[
- quadruple ???? (anyOP HC05 ASL) MODE_DIR1 (Byte 〈x3,x8〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 ASL) MODE_INHA (Byte 〈x4,x8〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 ASL) MODE_INHX (Byte 〈x5,x8〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 ASL) MODE_IX1 (Byte 〈x6,x8〉) 〈x0,x6〉
-; quadruple ???? (anyOP HC05 ASL) MODE_IX0 (Byte 〈x7,x8〉) 〈x0,x5〉
+ quadruple … (anyOP HC05 ASL) MODE_DIR1 (Byte 〈x3,x8〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 ASL) MODE_INHA (Byte 〈x4,x8〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 ASL) MODE_INHX (Byte 〈x5,x8〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 ASL) MODE_IX1 (Byte 〈x6,x8〉) 〈x0,x6〉
+; quadruple … (anyOP HC05 ASL) MODE_IX0 (Byte 〈x7,x8〉) 〈x0,x5〉
].
ndefinition opcode_table_HC05_5 ≝
[
- quadruple ???? (anyOP HC05 ASR) MODE_DIR1 (Byte 〈x3,x7〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 ASR) MODE_INHA (Byte 〈x4,x7〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 ASR) MODE_INHX (Byte 〈x5,x7〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 ASR) MODE_IX1 (Byte 〈x6,x7〉) 〈x0,x6〉
-; quadruple ???? (anyOP HC05 ASR) MODE_IX0 (Byte 〈x7,x7〉) 〈x0,x5〉
+ quadruple … (anyOP HC05 ASR) MODE_DIR1 (Byte 〈x3,x7〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 ASR) MODE_INHA (Byte 〈x4,x7〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 ASR) MODE_INHX (Byte 〈x5,x7〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 ASR) MODE_IX1 (Byte 〈x6,x7〉) 〈x0,x6〉
+; quadruple … (anyOP HC05 ASR) MODE_IX0 (Byte 〈x7,x7〉) 〈x0,x5〉
].
ndefinition opcode_table_HC05_6 ≝
[
- quadruple ???? (anyOP HC05 BRA ) MODE_IMM1 (Byte 〈x2,x0〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 BRN ) MODE_IMM1 (Byte 〈x2,x1〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 BHI ) MODE_IMM1 (Byte 〈x2,x2〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 BLS ) MODE_IMM1 (Byte 〈x2,x3〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 BCC ) MODE_IMM1 (Byte 〈x2,x4〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 BCS ) MODE_IMM1 (Byte 〈x2,x5〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 BNE ) MODE_IMM1 (Byte 〈x2,x6〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 BEQ ) MODE_IMM1 (Byte 〈x2,x7〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 BHCC) MODE_IMM1 (Byte 〈x2,x8〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 BHCS) MODE_IMM1 (Byte 〈x2,x9〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 BPL ) MODE_IMM1 (Byte 〈x2,xA〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 BMI ) MODE_IMM1 (Byte 〈x2,xB〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 BMC ) MODE_IMM1 (Byte 〈x2,xC〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 BMS ) MODE_IMM1 (Byte 〈x2,xD〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 BIL ) MODE_IMM1 (Byte 〈x2,xE〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 BIH ) MODE_IMM1 (Byte 〈x2,xF〉) 〈x0,x3〉
+ quadruple … (anyOP HC05 BRA ) MODE_IMM1 (Byte 〈x2,x0〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 BRN ) MODE_IMM1 (Byte 〈x2,x1〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 BHI ) MODE_IMM1 (Byte 〈x2,x2〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 BLS ) MODE_IMM1 (Byte 〈x2,x3〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 BCC ) MODE_IMM1 (Byte 〈x2,x4〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 BCS ) MODE_IMM1 (Byte 〈x2,x5〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 BNE ) MODE_IMM1 (Byte 〈x2,x6〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 BEQ ) MODE_IMM1 (Byte 〈x2,x7〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 BHCC) MODE_IMM1 (Byte 〈x2,x8〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 BHCS) MODE_IMM1 (Byte 〈x2,x9〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 BPL ) MODE_IMM1 (Byte 〈x2,xA〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 BMI ) MODE_IMM1 (Byte 〈x2,xB〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 BMC ) MODE_IMM1 (Byte 〈x2,xC〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 BMS ) MODE_IMM1 (Byte 〈x2,xD〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 BIL ) MODE_IMM1 (Byte 〈x2,xE〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 BIH ) MODE_IMM1 (Byte 〈x2,xF〉) 〈x0,x3〉
].
ndefinition opcode_table_HC05_7 ≝
[
- quadruple ???? (anyOP HC05 BSETn) (MODE_DIRn o0) (Byte 〈x1,x0〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BCLRn) (MODE_DIRn o0) (Byte 〈x1,x1〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BSETn) (MODE_DIRn o1) (Byte 〈x1,x2〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BCLRn) (MODE_DIRn o1) (Byte 〈x1,x3〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BSETn) (MODE_DIRn o2) (Byte 〈x1,x4〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BCLRn) (MODE_DIRn o2) (Byte 〈x1,x5〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BSETn) (MODE_DIRn o3) (Byte 〈x1,x6〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BCLRn) (MODE_DIRn o3) (Byte 〈x1,x7〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BSETn) (MODE_DIRn o4) (Byte 〈x1,x8〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BCLRn) (MODE_DIRn o4) (Byte 〈x1,x9〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BSETn) (MODE_DIRn o5) (Byte 〈x1,xA〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BCLRn) (MODE_DIRn o5) (Byte 〈x1,xB〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BSETn) (MODE_DIRn o6) (Byte 〈x1,xC〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BCLRn) (MODE_DIRn o6) (Byte 〈x1,xD〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BSETn) (MODE_DIRn o7) (Byte 〈x1,xE〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BCLRn) (MODE_DIRn o7) (Byte 〈x1,xF〉) 〈x0,x5〉
+ quadruple … (anyOP HC05 BSETn) (MODE_DIRn o0) (Byte 〈x1,x0〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BCLRn) (MODE_DIRn o0) (Byte 〈x1,x1〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BSETn) (MODE_DIRn o1) (Byte 〈x1,x2〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BCLRn) (MODE_DIRn o1) (Byte 〈x1,x3〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BSETn) (MODE_DIRn o2) (Byte 〈x1,x4〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BCLRn) (MODE_DIRn o2) (Byte 〈x1,x5〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BSETn) (MODE_DIRn o3) (Byte 〈x1,x6〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BCLRn) (MODE_DIRn o3) (Byte 〈x1,x7〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BSETn) (MODE_DIRn o4) (Byte 〈x1,x8〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BCLRn) (MODE_DIRn o4) (Byte 〈x1,x9〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BSETn) (MODE_DIRn o5) (Byte 〈x1,xA〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BCLRn) (MODE_DIRn o5) (Byte 〈x1,xB〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BSETn) (MODE_DIRn o6) (Byte 〈x1,xC〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BCLRn) (MODE_DIRn o6) (Byte 〈x1,xD〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BSETn) (MODE_DIRn o7) (Byte 〈x1,xE〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BCLRn) (MODE_DIRn o7) (Byte 〈x1,xF〉) 〈x0,x5〉
].
ndefinition opcode_table_HC05_8 ≝
[
- quadruple ???? (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x0〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x1〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x2〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x3〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x4〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x5〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x6〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x7〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x8〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x9〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xA〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xB〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xC〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xD〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xE〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xF〉) 〈x0,x5〉
+ quadruple … (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x0〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x1〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x2〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x3〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x4〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x5〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x6〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x7〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x8〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x9〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xA〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xB〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xC〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xD〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xE〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xF〉) 〈x0,x5〉
].
ndefinition opcode_table_HC05_9 ≝
[
- quadruple ???? (anyOP HC05 BIT) MODE_IMM1 (Byte 〈xA,x5〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC05 BIT) MODE_DIR1 (Byte 〈xB,x5〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 BIT) MODE_DIR2 (Byte 〈xC,x5〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 BIT) MODE_IX2 (Byte 〈xD,x5〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 BIT) MODE_IX1 (Byte 〈xE,x5〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 BIT) MODE_IX0 (Byte 〈xF,x5〉) 〈x0,x3〉
+ quadruple … (anyOP HC05 BIT) MODE_IMM1 (Byte 〈xA,x5〉) 〈x0,x2〉
+; quadruple … (anyOP HC05 BIT) MODE_DIR1 (Byte 〈xB,x5〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 BIT) MODE_DIR2 (Byte 〈xC,x5〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 BIT) MODE_IX2 (Byte 〈xD,x5〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 BIT) MODE_IX1 (Byte 〈xE,x5〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 BIT) MODE_IX0 (Byte 〈xF,x5〉) 〈x0,x3〉
].
ndefinition opcode_table_HC05_10 ≝
[
- quadruple ???? (anyOP HC05 MUL ) MODE_INH (Byte 〈x4,x2〉) 〈x0,xB〉
-; quadruple ???? (anyOP HC05 RTI ) MODE_INH (Byte 〈x8,x0〉) 〈x0,x9〉
-; quadruple ???? (anyOP HC05 RTS ) MODE_INH (Byte 〈x8,x1〉) 〈x0,x6〉
-; quadruple ???? (anyOP HC05 SWI ) MODE_INH (Byte 〈x8,x3〉) 〈x0,xA〉
-; quadruple ???? (anyOP HC05 STOP) MODE_INH (Byte 〈x8,xE〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC05 WAIT) MODE_INH (Byte 〈x8,xF〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC05 TAX ) MODE_INH (Byte 〈x9,x7〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC05 CLC ) MODE_INH (Byte 〈x9,x8〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC05 SEC ) MODE_INH (Byte 〈x9,x9〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC05 CLI ) MODE_INH (Byte 〈x9,xA〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC05 SEI ) MODE_INH (Byte 〈x9,xB〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC05 RSP ) MODE_INH (Byte 〈x9,xC〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC05 NOP ) MODE_INH (Byte 〈x9,xD〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC05 TXA ) MODE_INH (Byte 〈x9,xF〉) 〈x0,x2〉
+ quadruple … (anyOP HC05 MUL ) MODE_INH (Byte 〈x4,x2〉) 〈x0,xB〉
+; quadruple … (anyOP HC05 RTI ) MODE_INH (Byte 〈x8,x0〉) 〈x0,x9〉
+; quadruple … (anyOP HC05 RTS ) MODE_INH (Byte 〈x8,x1〉) 〈x0,x6〉
+; quadruple … (anyOP HC05 SWI ) MODE_INH (Byte 〈x8,x3〉) 〈x0,xA〉
+; quadruple … (anyOP HC05 STOP) MODE_INH (Byte 〈x8,xE〉) 〈x0,x2〉
+; quadruple … (anyOP HC05 WAIT) MODE_INH (Byte 〈x8,xF〉) 〈x0,x2〉
+; quadruple … (anyOP HC05 TAX ) MODE_INH (Byte 〈x9,x7〉) 〈x0,x2〉
+; quadruple … (anyOP HC05 CLC ) MODE_INH (Byte 〈x9,x8〉) 〈x0,x2〉
+; quadruple … (anyOP HC05 SEC ) MODE_INH (Byte 〈x9,x9〉) 〈x0,x2〉
+; quadruple … (anyOP HC05 CLI ) MODE_INH (Byte 〈x9,xA〉) 〈x0,x2〉
+; quadruple … (anyOP HC05 SEI ) MODE_INH (Byte 〈x9,xB〉) 〈x0,x2〉
+; quadruple … (anyOP HC05 RSP ) MODE_INH (Byte 〈x9,xC〉) 〈x0,x2〉
+; quadruple … (anyOP HC05 NOP ) MODE_INH (Byte 〈x9,xD〉) 〈x0,x2〉
+; quadruple … (anyOP HC05 TXA ) MODE_INH (Byte 〈x9,xF〉) 〈x0,x2〉
].
ndefinition opcode_table_HC05_11 ≝
[
- quadruple ???? (anyOP HC05 CLR) MODE_DIR1 (Byte 〈x3,xF〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 CLR) MODE_INHA (Byte 〈x4,xF〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 CLR) MODE_INHX (Byte 〈x5,xF〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 CLR) MODE_IX1 (Byte 〈x6,xF〉) 〈x0,x6〉
-; quadruple ???? (anyOP HC05 CLR) MODE_IX0 (Byte 〈x7,xF〉) 〈x0,x5〉
+ quadruple … (anyOP HC05 CLR) MODE_DIR1 (Byte 〈x3,xF〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 CLR) MODE_INHA (Byte 〈x4,xF〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 CLR) MODE_INHX (Byte 〈x5,xF〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 CLR) MODE_IX1 (Byte 〈x6,xF〉) 〈x0,x6〉
+; quadruple … (anyOP HC05 CLR) MODE_IX0 (Byte 〈x7,xF〉) 〈x0,x5〉
].
ndefinition opcode_table_HC05_12 ≝
[
- quadruple ???? (anyOP HC05 CMP) MODE_IMM1 (Byte 〈xA,x1〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC05 CMP) MODE_DIR1 (Byte 〈xB,x1〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 CMP) MODE_DIR2 (Byte 〈xC,x1〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 CMP) MODE_IX2 (Byte 〈xD,x1〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 CMP) MODE_IX1 (Byte 〈xE,x1〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 CMP) MODE_IX0 (Byte 〈xF,x1〉) 〈x0,x3〉
+ quadruple … (anyOP HC05 CMP) MODE_IMM1 (Byte 〈xA,x1〉) 〈x0,x2〉
+; quadruple … (anyOP HC05 CMP) MODE_DIR1 (Byte 〈xB,x1〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 CMP) MODE_DIR2 (Byte 〈xC,x1〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 CMP) MODE_IX2 (Byte 〈xD,x1〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 CMP) MODE_IX1 (Byte 〈xE,x1〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 CMP) MODE_IX0 (Byte 〈xF,x1〉) 〈x0,x3〉
].
ndefinition opcode_table_HC05_13 ≝
[
- quadruple ???? (anyOP HC05 COM) MODE_DIR1 (Byte 〈x3,x3〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 COM) MODE_INHA (Byte 〈x4,x3〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 COM) MODE_INHX (Byte 〈x5,x3〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 COM) MODE_IX1 (Byte 〈x6,x3〉) 〈x0,x6〉
-; quadruple ???? (anyOP HC05 COM) MODE_IX0 (Byte 〈x7,x3〉) 〈x0,x5〉
+ quadruple … (anyOP HC05 COM) MODE_DIR1 (Byte 〈x3,x3〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 COM) MODE_INHA (Byte 〈x4,x3〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 COM) MODE_INHX (Byte 〈x5,x3〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 COM) MODE_IX1 (Byte 〈x6,x3〉) 〈x0,x6〉
+; quadruple … (anyOP HC05 COM) MODE_IX0 (Byte 〈x7,x3〉) 〈x0,x5〉
].
ndefinition opcode_table_HC05_14 ≝
[
- quadruple ???? (anyOP HC05 CPX) MODE_IMM1 (Byte 〈xA,x3〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC05 CPX) MODE_DIR1 (Byte 〈xB,x3〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 CPX) MODE_DIR2 (Byte 〈xC,x3〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 CPX) MODE_IX2 (Byte 〈xD,x3〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 CPX) MODE_IX1 (Byte 〈xE,x3〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 CPX) MODE_IX0 (Byte 〈xF,x3〉) 〈x0,x3〉
+ quadruple … (anyOP HC05 CPX) MODE_IMM1 (Byte 〈xA,x3〉) 〈x0,x2〉
+; quadruple … (anyOP HC05 CPX) MODE_DIR1 (Byte 〈xB,x3〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 CPX) MODE_DIR2 (Byte 〈xC,x3〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 CPX) MODE_IX2 (Byte 〈xD,x3〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 CPX) MODE_IX1 (Byte 〈xE,x3〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 CPX) MODE_IX0 (Byte 〈xF,x3〉) 〈x0,x3〉
].
ndefinition opcode_table_HC05_15 ≝
[
- quadruple ???? (anyOP HC05 DEC) MODE_DIR1 (Byte 〈x3,xA〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 DEC) MODE_INHA (Byte 〈x4,xA〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 DEC) MODE_INHX (Byte 〈x5,xA〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 DEC) MODE_IX1 (Byte 〈x6,xA〉) 〈x0,x6〉
-; quadruple ???? (anyOP HC05 DEC) MODE_IX0 (Byte 〈x7,xA〉) 〈x0,x5〉
+ quadruple … (anyOP HC05 DEC) MODE_DIR1 (Byte 〈x3,xA〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 DEC) MODE_INHA (Byte 〈x4,xA〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 DEC) MODE_INHX (Byte 〈x5,xA〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 DEC) MODE_IX1 (Byte 〈x6,xA〉) 〈x0,x6〉
+; quadruple … (anyOP HC05 DEC) MODE_IX0 (Byte 〈x7,xA〉) 〈x0,x5〉
].
ndefinition opcode_table_HC05_16 ≝
[
- quadruple ???? (anyOP HC05 EOR) MODE_IMM1 (Byte 〈xA,x8〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC05 EOR) MODE_DIR1 (Byte 〈xB,x8〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 EOR) MODE_DIR2 (Byte 〈xC,x8〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 EOR) MODE_IX2 (Byte 〈xD,x8〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 EOR) MODE_IX1 (Byte 〈xE,x8〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 EOR) MODE_IX0 (Byte 〈xF,x8〉) 〈x0,x3〉
+ quadruple … (anyOP HC05 EOR) MODE_IMM1 (Byte 〈xA,x8〉) 〈x0,x2〉
+; quadruple … (anyOP HC05 EOR) MODE_DIR1 (Byte 〈xB,x8〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 EOR) MODE_DIR2 (Byte 〈xC,x8〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 EOR) MODE_IX2 (Byte 〈xD,x8〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 EOR) MODE_IX1 (Byte 〈xE,x8〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 EOR) MODE_IX0 (Byte 〈xF,x8〉) 〈x0,x3〉
].
ndefinition opcode_table_HC05_17 ≝
[
- quadruple ???? (anyOP HC05 INC) MODE_DIR1 (Byte 〈x3,xC〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 INC) MODE_INHA (Byte 〈x4,xC〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 INC) MODE_INHX (Byte 〈x5,xC〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 INC) MODE_IX1 (Byte 〈x6,xC〉) 〈x0,x6〉
-; quadruple ???? (anyOP HC05 INC) MODE_IX0 (Byte 〈x7,xC〉) 〈x0,x5〉
+ quadruple … (anyOP HC05 INC) MODE_DIR1 (Byte 〈x3,xC〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 INC) MODE_INHA (Byte 〈x4,xC〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 INC) MODE_INHX (Byte 〈x5,xC〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 INC) MODE_IX1 (Byte 〈x6,xC〉) 〈x0,x6〉
+; quadruple … (anyOP HC05 INC) MODE_IX0 (Byte 〈x7,xC〉) 〈x0,x5〉
].
ndefinition opcode_table_HC05_18 ≝
[
- quadruple ???? (anyOP HC05 JMP) MODE_IMM1EXT (Byte 〈xB,xC〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC05 JMP) MODE_IMM2 (Byte 〈xC,xC〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 JMP) MODE_INHX2ADD (Byte 〈xD,xC〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 JMP) MODE_INHX1ADD (Byte 〈xE,xC〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 JMP) MODE_INHX0ADD (Byte 〈xF,xC〉) 〈x0,x2〉
+ quadruple … (anyOP HC05 JMP) MODE_IMM1EXT (Byte 〈xB,xC〉) 〈x0,x2〉
+; quadruple … (anyOP HC05 JMP) MODE_IMM2 (Byte 〈xC,xC〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 JMP) MODE_INHX2ADD (Byte 〈xD,xC〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 JMP) MODE_INHX1ADD (Byte 〈xE,xC〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 JMP) MODE_INHX0ADD (Byte 〈xF,xC〉) 〈x0,x2〉
].
ndefinition opcode_table_HC05_19 ≝
[
- quadruple ???? (anyOP HC05 BSR) MODE_IMM1 (Byte 〈xA,xD〉) 〈x0,x6〉
-; quadruple ???? (anyOP HC05 JSR) MODE_IMM1EXT (Byte 〈xB,xD〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 JSR) MODE_IMM2 (Byte 〈xC,xD〉) 〈x0,x6〉
-; quadruple ???? (anyOP HC05 JSR) MODE_INHX2ADD (Byte 〈xD,xD〉) 〈x0,x7〉
-; quadruple ???? (anyOP HC05 JSR) MODE_INHX1ADD (Byte 〈xE,xD〉) 〈x0,x6〉
-; quadruple ???? (anyOP HC05 JSR) MODE_INHX0ADD (Byte 〈xF,xD〉) 〈x0,x5〉
+ quadruple … (anyOP HC05 BSR) MODE_IMM1 (Byte 〈xA,xD〉) 〈x0,x6〉
+; quadruple … (anyOP HC05 JSR) MODE_IMM1EXT (Byte 〈xB,xD〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 JSR) MODE_IMM2 (Byte 〈xC,xD〉) 〈x0,x6〉
+; quadruple … (anyOP HC05 JSR) MODE_INHX2ADD (Byte 〈xD,xD〉) 〈x0,x7〉
+; quadruple … (anyOP HC05 JSR) MODE_INHX1ADD (Byte 〈xE,xD〉) 〈x0,x6〉
+; quadruple … (anyOP HC05 JSR) MODE_INHX0ADD (Byte 〈xF,xD〉) 〈x0,x5〉
].
ndefinition opcode_table_HC05_20 ≝
[
- quadruple ???? (anyOP HC05 LDA) MODE_IMM1 (Byte 〈xA,x6〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC05 LDA) MODE_DIR1 (Byte 〈xB,x6〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 LDA) MODE_DIR2 (Byte 〈xC,x6〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 LDA) MODE_IX2 (Byte 〈xD,x6〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 LDA) MODE_IX1 (Byte 〈xE,x6〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 LDA) MODE_IX0 (Byte 〈xF,x6〉) 〈x0,x3〉
+ quadruple … (anyOP HC05 LDA) MODE_IMM1 (Byte 〈xA,x6〉) 〈x0,x2〉
+; quadruple … (anyOP HC05 LDA) MODE_DIR1 (Byte 〈xB,x6〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 LDA) MODE_DIR2 (Byte 〈xC,x6〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 LDA) MODE_IX2 (Byte 〈xD,x6〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 LDA) MODE_IX1 (Byte 〈xE,x6〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 LDA) MODE_IX0 (Byte 〈xF,x6〉) 〈x0,x3〉
].
ndefinition opcode_table_HC05_21 ≝
[
- quadruple ???? (anyOP HC05 LDX) MODE_IMM1 (Byte 〈xA,xE〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC05 LDX) MODE_DIR1 (Byte 〈xB,xE〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 LDX) MODE_DIR2 (Byte 〈xC,xE〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 LDX) MODE_IX2 (Byte 〈xD,xE〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 LDX) MODE_IX1 (Byte 〈xE,xE〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 LDX) MODE_IX0 (Byte 〈xF,xE〉) 〈x0,x3〉
+ quadruple … (anyOP HC05 LDX) MODE_IMM1 (Byte 〈xA,xE〉) 〈x0,x2〉
+; quadruple … (anyOP HC05 LDX) MODE_DIR1 (Byte 〈xB,xE〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 LDX) MODE_DIR2 (Byte 〈xC,xE〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 LDX) MODE_IX2 (Byte 〈xD,xE〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 LDX) MODE_IX1 (Byte 〈xE,xE〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 LDX) MODE_IX0 (Byte 〈xF,xE〉) 〈x0,x3〉
].
ndefinition opcode_table_HC05_22 ≝
[
- quadruple ???? (anyOP HC05 LSR) MODE_DIR1 (Byte 〈x3,x4〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 LSR) MODE_INHA (Byte 〈x4,x4〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 LSR) MODE_INHX (Byte 〈x5,x4〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 LSR) MODE_IX1 (Byte 〈x6,x4〉) 〈x0,x6〉
-; quadruple ???? (anyOP HC05 LSR) MODE_IX0 (Byte 〈x7,x4〉) 〈x0,x5〉
+ quadruple … (anyOP HC05 LSR) MODE_DIR1 (Byte 〈x3,x4〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 LSR) MODE_INHA (Byte 〈x4,x4〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 LSR) MODE_INHX (Byte 〈x5,x4〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 LSR) MODE_IX1 (Byte 〈x6,x4〉) 〈x0,x6〉
+; quadruple … (anyOP HC05 LSR) MODE_IX0 (Byte 〈x7,x4〉) 〈x0,x5〉
].
ndefinition opcode_table_HC05_23 ≝
[
- quadruple ???? (anyOP HC05 NEG) MODE_DIR1 (Byte 〈x3,x0〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 NEG) MODE_INHA (Byte 〈x4,x0〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 NEG) MODE_INHX (Byte 〈x5,x0〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 NEG) MODE_IX1 (Byte 〈x6,x0〉) 〈x0,x6〉
-; quadruple ???? (anyOP HC05 NEG) MODE_IX0 (Byte 〈x7,x0〉) 〈x0,x5〉
+ quadruple … (anyOP HC05 NEG) MODE_DIR1 (Byte 〈x3,x0〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 NEG) MODE_INHA (Byte 〈x4,x0〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 NEG) MODE_INHX (Byte 〈x5,x0〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 NEG) MODE_IX1 (Byte 〈x6,x0〉) 〈x0,x6〉
+; quadruple … (anyOP HC05 NEG) MODE_IX0 (Byte 〈x7,x0〉) 〈x0,x5〉
].
ndefinition opcode_table_HC05_24 ≝
[
- quadruple ???? (anyOP HC05 ORA) MODE_IMM1 (Byte 〈xA,xA〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC05 ORA) MODE_DIR1 (Byte 〈xB,xA〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 ORA) MODE_DIR2 (Byte 〈xC,xA〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 ORA) MODE_IX2 (Byte 〈xD,xA〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 ORA) MODE_IX1 (Byte 〈xE,xA〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 ORA) MODE_IX0 (Byte 〈xF,xA〉) 〈x0,x3〉
+ quadruple … (anyOP HC05 ORA) MODE_IMM1 (Byte 〈xA,xA〉) 〈x0,x2〉
+; quadruple … (anyOP HC05 ORA) MODE_DIR1 (Byte 〈xB,xA〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 ORA) MODE_DIR2 (Byte 〈xC,xA〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 ORA) MODE_IX2 (Byte 〈xD,xA〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 ORA) MODE_IX1 (Byte 〈xE,xA〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 ORA) MODE_IX0 (Byte 〈xF,xA〉) 〈x0,x3〉
].
ndefinition opcode_table_HC05_25 ≝
[
- quadruple ???? (anyOP HC05 ROL) MODE_DIR1 (Byte 〈x3,x9〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 ROL) MODE_INHA (Byte 〈x4,x9〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 ROL) MODE_INHX (Byte 〈x5,x9〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 ROL) MODE_IX1 (Byte 〈x6,x9〉) 〈x0,x6〉
-; quadruple ???? (anyOP HC05 ROL) MODE_IX0 (Byte 〈x7,x9〉) 〈x0,x5〉
+ quadruple … (anyOP HC05 ROL) MODE_DIR1 (Byte 〈x3,x9〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 ROL) MODE_INHA (Byte 〈x4,x9〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 ROL) MODE_INHX (Byte 〈x5,x9〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 ROL) MODE_IX1 (Byte 〈x6,x9〉) 〈x0,x6〉
+; quadruple … (anyOP HC05 ROL) MODE_IX0 (Byte 〈x7,x9〉) 〈x0,x5〉
].
ndefinition opcode_table_HC05_26 ≝
[
- quadruple ???? (anyOP HC05 ROR) MODE_DIR1 (Byte 〈x3,x6〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 ROR) MODE_INHA (Byte 〈x4,x6〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 ROR) MODE_INHX (Byte 〈x5,x6〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 ROR) MODE_IX1 (Byte 〈x6,x6〉) 〈x0,x6〉
-; quadruple ???? (anyOP HC05 ROR) MODE_IX0 (Byte 〈x7,x6〉) 〈x0,x5〉
+ quadruple … (anyOP HC05 ROR) MODE_DIR1 (Byte 〈x3,x6〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 ROR) MODE_INHA (Byte 〈x4,x6〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 ROR) MODE_INHX (Byte 〈x5,x6〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 ROR) MODE_IX1 (Byte 〈x6,x6〉) 〈x0,x6〉
+; quadruple … (anyOP HC05 ROR) MODE_IX0 (Byte 〈x7,x6〉) 〈x0,x5〉
].
ndefinition opcode_table_HC05_27 ≝
[
- quadruple ???? (anyOP HC05 SBC) MODE_IMM1 (Byte 〈xA,x2〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC05 SBC) MODE_DIR1 (Byte 〈xB,x2〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 SBC) MODE_DIR2 (Byte 〈xC,x2〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 SBC) MODE_IX2 (Byte 〈xD,x2〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 SBC) MODE_IX1 (Byte 〈xE,x2〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 SBC) MODE_IX0 (Byte 〈xF,x2〉) 〈x0,x3〉
+ quadruple … (anyOP HC05 SBC) MODE_IMM1 (Byte 〈xA,x2〉) 〈x0,x2〉
+; quadruple … (anyOP HC05 SBC) MODE_DIR1 (Byte 〈xB,x2〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 SBC) MODE_DIR2 (Byte 〈xC,x2〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 SBC) MODE_IX2 (Byte 〈xD,x2〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 SBC) MODE_IX1 (Byte 〈xE,x2〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 SBC) MODE_IX0 (Byte 〈xF,x2〉) 〈x0,x3〉
].
ndefinition opcode_table_HC05_28 ≝
[
- quadruple ???? (anyOP HC05 STA) MODE_DIR1 (Byte 〈xB,x7〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 STA) MODE_DIR2 (Byte 〈xC,x7〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 STA) MODE_IX2 (Byte 〈xD,x7〉) 〈x0,x6〉
-; quadruple ???? (anyOP HC05 STA) MODE_IX1 (Byte 〈xE,x7〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 STA) MODE_IX0 (Byte 〈xF,x7〉) 〈x0,x4〉
+ quadruple … (anyOP HC05 STA) MODE_DIR1 (Byte 〈xB,x7〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 STA) MODE_DIR2 (Byte 〈xC,x7〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 STA) MODE_IX2 (Byte 〈xD,x7〉) 〈x0,x6〉
+; quadruple … (anyOP HC05 STA) MODE_IX1 (Byte 〈xE,x7〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 STA) MODE_IX0 (Byte 〈xF,x7〉) 〈x0,x4〉
].
ndefinition opcode_table_HC05_29 ≝
[
- quadruple ???? (anyOP HC05 STX) MODE_DIR1 (Byte 〈xB,xF〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 STX) MODE_DIR2 (Byte 〈xC,xF〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 STX) MODE_IX2 (Byte 〈xD,xF〉) 〈x0,x6〉
-; quadruple ???? (anyOP HC05 STX) MODE_IX1 (Byte 〈xE,xF〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 STX) MODE_IX0 (Byte 〈xF,xF〉) 〈x0,x4〉
+ quadruple … (anyOP HC05 STX) MODE_DIR1 (Byte 〈xB,xF〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 STX) MODE_DIR2 (Byte 〈xC,xF〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 STX) MODE_IX2 (Byte 〈xD,xF〉) 〈x0,x6〉
+; quadruple … (anyOP HC05 STX) MODE_IX1 (Byte 〈xE,xF〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 STX) MODE_IX0 (Byte 〈xF,xF〉) 〈x0,x4〉
].
ndefinition opcode_table_HC05_30 ≝
[
- quadruple ???? (anyOP HC05 SUB) MODE_IMM1 (Byte 〈xA,x0〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC05 SUB) MODE_DIR1 (Byte 〈xB,x0〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 SUB) MODE_DIR2 (Byte 〈xC,x0〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 SUB) MODE_IX2 (Byte 〈xD,x0〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 SUB) MODE_IX1 (Byte 〈xE,x0〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 SUB) MODE_IX0 (Byte 〈xF,x0〉) 〈x0,x3〉
+ quadruple … (anyOP HC05 SUB) MODE_IMM1 (Byte 〈xA,x0〉) 〈x0,x2〉
+; quadruple … (anyOP HC05 SUB) MODE_DIR1 (Byte 〈xB,x0〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 SUB) MODE_DIR2 (Byte 〈xC,x0〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 SUB) MODE_IX2 (Byte 〈xD,x0〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 SUB) MODE_IX1 (Byte 〈xE,x0〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 SUB) MODE_IX0 (Byte 〈xF,x0〉) 〈x0,x3〉
].
ndefinition opcode_table_HC05_31 ≝
[
- quadruple ???? (anyOP HC05 TST) MODE_DIR1 (Byte 〈x3,xD〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC05 TST) MODE_INHA (Byte 〈x4,xD〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 TST) MODE_INHX (Byte 〈x5,xD〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC05 TST) MODE_IX1 (Byte 〈x6,xD〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC05 TST) MODE_IX0 (Byte 〈x7,xD〉) 〈x0,x4〉
+ quadruple … (anyOP HC05 TST) MODE_DIR1 (Byte 〈x3,xD〉) 〈x0,x4〉
+; quadruple … (anyOP HC05 TST) MODE_INHA (Byte 〈x4,xD〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 TST) MODE_INHX (Byte 〈x5,xD〉) 〈x0,x3〉
+; quadruple … (anyOP HC05 TST) MODE_IX1 (Byte 〈x6,xD〉) 〈x0,x5〉
+; quadruple … (anyOP HC05 TST) MODE_IX0 (Byte 〈x7,xD〉) 〈x0,x4〉
].
ndefinition opcode_table_HC05 ≝
(test_not_impl_byte b HC05_not_impl_byte ⊙ eq_nat (get_byte_count HC05 b 0 opcode_table_HC05) 1) ⊗
(⊖ (test_not_impl_byte b HC05_not_impl_byte) ⊙ eq_nat (get_byte_count HC05 b 0 opcode_table_HC05) 0))
= true.
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
(* HC05: pseudocodici non implementati come da manuale *)
(test_not_impl_pseudo o HC05_not_impl_pseudo ⊙ le_nat 1 (get_pseudo_count HC05 o 0 opcode_table_HC05)) ⊗
(⊖ (test_not_impl_pseudo o HC05_not_impl_pseudo) ⊙ eq_nat (get_pseudo_count HC05 o 0 opcode_table_HC05) 0))
= true.
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
(* HC05: modalita' non implementate come da manuale *)
(test_not_impl_mode i HC05_not_impl_mode ⊙ le_nat 1 (get_mode_count HC05 i 0 opcode_table_HC05)) ⊗
(⊖ (test_not_impl_mode i HC05_not_impl_mode) ⊙ eq_nat (get_mode_count HC05 i 0 opcode_table_HC05) 0))
= true.
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma ok_OpIm_table_HC05 :
forall_instr_mode (λi:instr_mode.
forall_opcode (λop:opcode.
le_nat (get_OpIm_count HC05 (anyOP HC05 op) i 0 opcode_table_HC05) 1)) = true.
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
ndefinition opcode_table_HC08_1 ≝
[
- quadruple ???? (anyOP HC08 ADC) MODE_IMM1 (Byte 〈xA,x9〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 ADC) MODE_DIR1 (Byte 〈xB,x9〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 ADC) MODE_DIR2 (Byte 〈xC,x9〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 ADC) MODE_IX2 (Byte 〈xD,x9〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 ADC) MODE_IX1 (Byte 〈xE,x9〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 ADC) MODE_IX0 (Byte 〈xF,x9〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 ADC) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x9〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 ADC) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x9〉〉) 〈x0,x4〉
+ quadruple … (anyOP HC08 ADC) MODE_IMM1 (Byte 〈xA,x9〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 ADC) MODE_DIR1 (Byte 〈xB,x9〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 ADC) MODE_DIR2 (Byte 〈xC,x9〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 ADC) MODE_IX2 (Byte 〈xD,x9〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 ADC) MODE_IX1 (Byte 〈xE,x9〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 ADC) MODE_IX0 (Byte 〈xF,x9〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 ADC) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x9〉〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 ADC) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x9〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HC08_2 ≝
[
- quadruple ???? (anyOP HC08 ADD) MODE_IMM1 (Byte 〈xA,xB〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 ADD) MODE_DIR1 (Byte 〈xB,xB〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 ADD) MODE_DIR2 (Byte 〈xC,xB〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 ADD) MODE_IX2 (Byte 〈xD,xB〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 ADD) MODE_IX1 (Byte 〈xE,xB〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 ADD) MODE_IX0 (Byte 〈xF,xB〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 ADD) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,xB〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 ADD) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,xB〉〉) 〈x0,x4〉
+ quadruple … (anyOP HC08 ADD) MODE_IMM1 (Byte 〈xA,xB〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 ADD) MODE_DIR1 (Byte 〈xB,xB〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 ADD) MODE_DIR2 (Byte 〈xC,xB〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 ADD) MODE_IX2 (Byte 〈xD,xB〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 ADD) MODE_IX1 (Byte 〈xE,xB〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 ADD) MODE_IX0 (Byte 〈xF,xB〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 ADD) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,xB〉〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 ADD) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,xB〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HC08_3 ≝
[
- quadruple ???? (anyOP HC08 AND) MODE_IMM1 (Byte 〈xA,x4〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 AND) MODE_DIR1 (Byte 〈xB,x4〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 AND) MODE_DIR2 (Byte 〈xC,x4〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 AND) MODE_IX2 (Byte 〈xD,x4〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 AND) MODE_IX1 (Byte 〈xE,x4〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 AND) MODE_IX0 (Byte 〈xF,x4〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 AND) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x4〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 AND) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x4〉〉) 〈x0,x4〉
+ quadruple … (anyOP HC08 AND) MODE_IMM1 (Byte 〈xA,x4〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 AND) MODE_DIR1 (Byte 〈xB,x4〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 AND) MODE_DIR2 (Byte 〈xC,x4〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 AND) MODE_IX2 (Byte 〈xD,x4〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 AND) MODE_IX1 (Byte 〈xE,x4〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 AND) MODE_IX0 (Byte 〈xF,x4〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 AND) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x4〉〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 AND) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x4〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HC08_4 ≝
[
- quadruple ???? (anyOP HC08 ASL) MODE_DIR1 (Byte 〈x3,x8〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 ASL) MODE_INHA (Byte 〈x4,x8〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 ASL) MODE_INHX (Byte 〈x5,x8〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 ASL) MODE_IX1 (Byte 〈x6,x8〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 ASL) MODE_IX0 (Byte 〈x7,x8〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 ASL) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x8〉〉) 〈x0,x5〉
+ quadruple … (anyOP HC08 ASL) MODE_DIR1 (Byte 〈x3,x8〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 ASL) MODE_INHA (Byte 〈x4,x8〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 ASL) MODE_INHX (Byte 〈x5,x8〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 ASL) MODE_IX1 (Byte 〈x6,x8〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 ASL) MODE_IX0 (Byte 〈x7,x8〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 ASL) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x8〉〉) 〈x0,x5〉
].
ndefinition opcode_table_HC08_5 ≝
[
- quadruple ???? (anyOP HC08 ASR) MODE_DIR1 (Byte 〈x3,x7〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 ASR) MODE_INHA (Byte 〈x4,x7〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 ASR) MODE_INHX (Byte 〈x5,x7〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 ASR) MODE_IX1 (Byte 〈x6,x7〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 ASR) MODE_IX0 (Byte 〈x7,x7〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 ASR) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x7〉〉) 〈x0,x5〉
+ quadruple … (anyOP HC08 ASR) MODE_DIR1 (Byte 〈x3,x7〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 ASR) MODE_INHA (Byte 〈x4,x7〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 ASR) MODE_INHX (Byte 〈x5,x7〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 ASR) MODE_IX1 (Byte 〈x6,x7〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 ASR) MODE_IX0 (Byte 〈x7,x7〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 ASR) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x7〉〉) 〈x0,x5〉
].
ndefinition opcode_table_HC08_6 ≝
[
- quadruple ???? (anyOP HC08 BRA ) MODE_IMM1 (Byte 〈x2,x0〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 BRN ) MODE_IMM1 (Byte 〈x2,x1〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 BHI ) MODE_IMM1 (Byte 〈x2,x2〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 BLS ) MODE_IMM1 (Byte 〈x2,x3〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 BCC ) MODE_IMM1 (Byte 〈x2,x4〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 BCS ) MODE_IMM1 (Byte 〈x2,x5〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 BNE ) MODE_IMM1 (Byte 〈x2,x6〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 BEQ ) MODE_IMM1 (Byte 〈x2,x7〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 BHCC) MODE_IMM1 (Byte 〈x2,x8〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 BHCS) MODE_IMM1 (Byte 〈x2,x9〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 BPL ) MODE_IMM1 (Byte 〈x2,xA〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 BMI ) MODE_IMM1 (Byte 〈x2,xB〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 BMC ) MODE_IMM1 (Byte 〈x2,xC〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 BMS ) MODE_IMM1 (Byte 〈x2,xD〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 BIL ) MODE_IMM1 (Byte 〈x2,xE〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 BIH ) MODE_IMM1 (Byte 〈x2,xF〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 BGE ) MODE_IMM1 (Byte 〈x9,x0〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 BLT ) MODE_IMM1 (Byte 〈x9,x1〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 BGT ) MODE_IMM1 (Byte 〈x9,x2〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 BLE ) MODE_IMM1 (Byte 〈x9,x3〉) 〈x0,x3〉
+ quadruple … (anyOP HC08 BRA ) MODE_IMM1 (Byte 〈x2,x0〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 BRN ) MODE_IMM1 (Byte 〈x2,x1〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 BHI ) MODE_IMM1 (Byte 〈x2,x2〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 BLS ) MODE_IMM1 (Byte 〈x2,x3〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 BCC ) MODE_IMM1 (Byte 〈x2,x4〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 BCS ) MODE_IMM1 (Byte 〈x2,x5〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 BNE ) MODE_IMM1 (Byte 〈x2,x6〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 BEQ ) MODE_IMM1 (Byte 〈x2,x7〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 BHCC) MODE_IMM1 (Byte 〈x2,x8〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 BHCS) MODE_IMM1 (Byte 〈x2,x9〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 BPL ) MODE_IMM1 (Byte 〈x2,xA〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 BMI ) MODE_IMM1 (Byte 〈x2,xB〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 BMC ) MODE_IMM1 (Byte 〈x2,xC〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 BMS ) MODE_IMM1 (Byte 〈x2,xD〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 BIL ) MODE_IMM1 (Byte 〈x2,xE〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 BIH ) MODE_IMM1 (Byte 〈x2,xF〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 BGE ) MODE_IMM1 (Byte 〈x9,x0〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 BLT ) MODE_IMM1 (Byte 〈x9,x1〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 BGT ) MODE_IMM1 (Byte 〈x9,x2〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 BLE ) MODE_IMM1 (Byte 〈x9,x3〉) 〈x0,x3〉
].
ndefinition opcode_table_HC08_7 ≝
[
- quadruple ???? (anyOP HC08 BSETn) (MODE_DIRn o0) (Byte 〈x1,x0〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 BCLRn) (MODE_DIRn o0) (Byte 〈x1,x1〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 BSETn) (MODE_DIRn o1) (Byte 〈x1,x2〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 BCLRn) (MODE_DIRn o1) (Byte 〈x1,x3〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 BSETn) (MODE_DIRn o2) (Byte 〈x1,x4〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 BCLRn) (MODE_DIRn o2) (Byte 〈x1,x5〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 BSETn) (MODE_DIRn o3) (Byte 〈x1,x6〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 BCLRn) (MODE_DIRn o3) (Byte 〈x1,x7〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 BSETn) (MODE_DIRn o4) (Byte 〈x1,x8〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 BCLRn) (MODE_DIRn o4) (Byte 〈x1,x9〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 BSETn) (MODE_DIRn o5) (Byte 〈x1,xA〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 BCLRn) (MODE_DIRn o5) (Byte 〈x1,xB〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 BSETn) (MODE_DIRn o6) (Byte 〈x1,xC〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 BCLRn) (MODE_DIRn o6) (Byte 〈x1,xD〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 BSETn) (MODE_DIRn o7) (Byte 〈x1,xE〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 BCLRn) (MODE_DIRn o7) (Byte 〈x1,xF〉) 〈x0,x4〉
+ quadruple … (anyOP HC08 BSETn) (MODE_DIRn o0) (Byte 〈x1,x0〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 BCLRn) (MODE_DIRn o0) (Byte 〈x1,x1〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 BSETn) (MODE_DIRn o1) (Byte 〈x1,x2〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 BCLRn) (MODE_DIRn o1) (Byte 〈x1,x3〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 BSETn) (MODE_DIRn o2) (Byte 〈x1,x4〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 BCLRn) (MODE_DIRn o2) (Byte 〈x1,x5〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 BSETn) (MODE_DIRn o3) (Byte 〈x1,x6〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 BCLRn) (MODE_DIRn o3) (Byte 〈x1,x7〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 BSETn) (MODE_DIRn o4) (Byte 〈x1,x8〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 BCLRn) (MODE_DIRn o4) (Byte 〈x1,x9〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 BSETn) (MODE_DIRn o5) (Byte 〈x1,xA〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 BCLRn) (MODE_DIRn o5) (Byte 〈x1,xB〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 BSETn) (MODE_DIRn o6) (Byte 〈x1,xC〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 BCLRn) (MODE_DIRn o6) (Byte 〈x1,xD〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 BSETn) (MODE_DIRn o7) (Byte 〈x1,xE〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 BCLRn) (MODE_DIRn o7) (Byte 〈x1,xF〉) 〈x0,x4〉
].
ndefinition opcode_table_HC08_8 ≝
[
- quadruple ???? (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x0〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x1〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x2〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x3〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x4〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x5〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x6〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x7〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x8〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x9〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xA〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xB〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xC〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xD〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xE〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xF〉) 〈x0,x5〉
+ quadruple … (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x0〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x1〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x2〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x3〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x4〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x5〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x6〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x7〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x8〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x9〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xA〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xB〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xC〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xD〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xE〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xF〉) 〈x0,x5〉
].
ndefinition opcode_table_HC08_9 ≝
[
- quadruple ???? (anyOP HC08 BIT) MODE_IMM1 (Byte 〈xA,x5〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 BIT) MODE_DIR1 (Byte 〈xB,x5〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 BIT) MODE_DIR2 (Byte 〈xC,x5〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 BIT) MODE_IX2 (Byte 〈xD,x5〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 BIT) MODE_IX1 (Byte 〈xE,x5〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 BIT) MODE_IX0 (Byte 〈xF,x5〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 BIT) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x5〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 BIT) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x5〉〉) 〈x0,x4〉
+ quadruple … (anyOP HC08 BIT) MODE_IMM1 (Byte 〈xA,x5〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 BIT) MODE_DIR1 (Byte 〈xB,x5〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 BIT) MODE_DIR2 (Byte 〈xC,x5〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 BIT) MODE_IX2 (Byte 〈xD,x5〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 BIT) MODE_IX1 (Byte 〈xE,x5〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 BIT) MODE_IX0 (Byte 〈xF,x5〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 BIT) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x5〉〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 BIT) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x5〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HC08_10 ≝
[
- quadruple ???? (anyOP HC08 MUL ) MODE_INH (Byte 〈x4,x2〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 DIV ) MODE_INH (Byte 〈x5,x2〉) 〈x0,x7〉
-; quadruple ???? (anyOP HC08 NSA ) MODE_INH (Byte 〈x6,x2〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 DAA ) MODE_INH (Byte 〈x7,x2〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 RTI ) MODE_INH (Byte 〈x8,x0〉) 〈x0,x7〉
-; quadruple ???? (anyOP HC08 RTS ) MODE_INH (Byte 〈x8,x1〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 SWI ) MODE_INH (Byte 〈x8,x3〉) 〈x0,x9〉
-; quadruple ???? (anyOP HC08 TAP ) MODE_INH (Byte 〈x8,x4〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 TPA ) MODE_INH (Byte 〈x8,x5〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 PULA) MODE_INH (Byte 〈x8,x6〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 PSHA) MODE_INH (Byte 〈x8,x7〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 PULX) MODE_INH (Byte 〈x8,x8〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 PSHX) MODE_INH (Byte 〈x8,x9〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 PULH) MODE_INH (Byte 〈x8,xA〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 PSHH) MODE_INH (Byte 〈x8,xB〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 STOP) MODE_INH (Byte 〈x8,xE〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 WAIT) MODE_INH (Byte 〈x8,xF〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 TXS ) MODE_INH (Byte 〈x9,x4〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 TSX ) MODE_INH (Byte 〈x9,x5〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 TAX ) MODE_INH (Byte 〈x9,x7〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 CLC ) MODE_INH (Byte 〈x9,x8〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 SEC ) MODE_INH (Byte 〈x9,x9〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 CLI ) MODE_INH (Byte 〈x9,xA〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 SEI ) MODE_INH (Byte 〈x9,xB〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 RSP ) MODE_INH (Byte 〈x9,xC〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 NOP ) MODE_INH (Byte 〈x9,xD〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 TXA ) MODE_INH (Byte 〈x9,xF〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 AIS ) MODE_IMM1 (Byte 〈xA,x7〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 AIX ) MODE_IMM1 (Byte 〈xA,xF〉) 〈x0,x2〉
+ quadruple … (anyOP HC08 MUL ) MODE_INH (Byte 〈x4,x2〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 DIV ) MODE_INH (Byte 〈x5,x2〉) 〈x0,x7〉
+; quadruple … (anyOP HC08 NSA ) MODE_INH (Byte 〈x6,x2〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 DAA ) MODE_INH (Byte 〈x7,x2〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 RTI ) MODE_INH (Byte 〈x8,x0〉) 〈x0,x7〉
+; quadruple … (anyOP HC08 RTS ) MODE_INH (Byte 〈x8,x1〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 SWI ) MODE_INH (Byte 〈x8,x3〉) 〈x0,x9〉
+; quadruple … (anyOP HC08 TAP ) MODE_INH (Byte 〈x8,x4〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 TPA ) MODE_INH (Byte 〈x8,x5〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 PULA) MODE_INH (Byte 〈x8,x6〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 PSHA) MODE_INH (Byte 〈x8,x7〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 PULX) MODE_INH (Byte 〈x8,x8〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 PSHX) MODE_INH (Byte 〈x8,x9〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 PULH) MODE_INH (Byte 〈x8,xA〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 PSHH) MODE_INH (Byte 〈x8,xB〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 STOP) MODE_INH (Byte 〈x8,xE〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 WAIT) MODE_INH (Byte 〈x8,xF〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 TXS ) MODE_INH (Byte 〈x9,x4〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 TSX ) MODE_INH (Byte 〈x9,x5〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 TAX ) MODE_INH (Byte 〈x9,x7〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 CLC ) MODE_INH (Byte 〈x9,x8〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 SEC ) MODE_INH (Byte 〈x9,x9〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 CLI ) MODE_INH (Byte 〈x9,xA〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 SEI ) MODE_INH (Byte 〈x9,xB〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 RSP ) MODE_INH (Byte 〈x9,xC〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 NOP ) MODE_INH (Byte 〈x9,xD〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 TXA ) MODE_INH (Byte 〈x9,xF〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 AIS ) MODE_IMM1 (Byte 〈xA,x7〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 AIX ) MODE_IMM1 (Byte 〈xA,xF〉) 〈x0,x2〉
].
ndefinition opcode_table_HC08_11 ≝
[
- quadruple ???? (anyOP HC08 CBEQA) MODE_DIR1_and_IMM1 (Byte 〈x3,x1〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 CBEQA) MODE_IMM1_and_IMM1 (Byte 〈x4,x1〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 CBEQX) MODE_IMM1_and_IMM1 (Byte 〈x5,x1〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 CBEQA) MODE_IX1p_and_IMM1 (Byte 〈x6,x1〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 CBEQA) MODE_IX0p_and_IMM1 (Byte 〈x7,x1〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 CBEQA) MODE_SP1_and_IMM1 (Word 〈〈x9,xE〉:〈x6,x1〉〉) 〈x0,x6〉
+ quadruple … (anyOP HC08 CBEQA) MODE_DIR1_and_IMM1 (Byte 〈x3,x1〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 CBEQA) MODE_IMM1_and_IMM1 (Byte 〈x4,x1〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 CBEQX) MODE_IMM1_and_IMM1 (Byte 〈x5,x1〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 CBEQA) MODE_IX1p_and_IMM1 (Byte 〈x6,x1〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 CBEQA) MODE_IX0p_and_IMM1 (Byte 〈x7,x1〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 CBEQA) MODE_SP1_and_IMM1 (Word 〈〈x9,xE〉:〈x6,x1〉〉) 〈x0,x6〉
].
ndefinition opcode_table_HC08_12 ≝
[
- quadruple ???? (anyOP HC08 CLR) MODE_DIR1 (Byte 〈x3,xF〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 CLR) MODE_INHA (Byte 〈x4,xF〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 CLR) MODE_INHX (Byte 〈x5,xF〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 CLR) MODE_IX1 (Byte 〈x6,xF〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 CLR) MODE_IX0 (Byte 〈x7,xF〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 CLR) MODE_INHH (Byte 〈x8,xC〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 CLR) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,xF〉〉) 〈x0,x4〉
+ quadruple … (anyOP HC08 CLR) MODE_DIR1 (Byte 〈x3,xF〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 CLR) MODE_INHA (Byte 〈x4,xF〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 CLR) MODE_INHX (Byte 〈x5,xF〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 CLR) MODE_IX1 (Byte 〈x6,xF〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 CLR) MODE_IX0 (Byte 〈x7,xF〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 CLR) MODE_INHH (Byte 〈x8,xC〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 CLR) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,xF〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HC08_13 ≝
[
- quadruple ???? (anyOP HC08 CMP) MODE_IMM1 (Byte 〈xA,x1〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 CMP) MODE_DIR1 (Byte 〈xB,x1〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 CMP) MODE_DIR2 (Byte 〈xC,x1〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 CMP) MODE_IX2 (Byte 〈xD,x1〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 CMP) MODE_IX1 (Byte 〈xE,x1〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 CMP) MODE_IX0 (Byte 〈xF,x1〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 CMP) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x1〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 CMP) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x1〉〉) 〈x0,x4〉
+ quadruple … (anyOP HC08 CMP) MODE_IMM1 (Byte 〈xA,x1〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 CMP) MODE_DIR1 (Byte 〈xB,x1〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 CMP) MODE_DIR2 (Byte 〈xC,x1〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 CMP) MODE_IX2 (Byte 〈xD,x1〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 CMP) MODE_IX1 (Byte 〈xE,x1〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 CMP) MODE_IX0 (Byte 〈xF,x1〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 CMP) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x1〉〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 CMP) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x1〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HC08_14 ≝
[
- quadruple ???? (anyOP HC08 COM) MODE_DIR1 (Byte 〈x3,x3〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 COM) MODE_INHA (Byte 〈x4,x3〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 COM) MODE_INHX (Byte 〈x5,x3〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 COM) MODE_IX1 (Byte 〈x6,x3〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 COM) MODE_IX0 (Byte 〈x7,x3〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 COM) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x3〉〉) 〈x0,x5〉
+ quadruple … (anyOP HC08 COM) MODE_DIR1 (Byte 〈x3,x3〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 COM) MODE_INHA (Byte 〈x4,x3〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 COM) MODE_INHX (Byte 〈x5,x3〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 COM) MODE_IX1 (Byte 〈x6,x3〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 COM) MODE_IX0 (Byte 〈x7,x3〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 COM) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x3〉〉) 〈x0,x5〉
].
ndefinition opcode_table_HC08_15 ≝
[
- quadruple ???? (anyOP HC08 STHX) MODE_DIR1 (Byte 〈x3,x5〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 LDHX) MODE_IMM2 (Byte 〈x4,x5〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 LDHX) MODE_DIR1 (Byte 〈x5,x5〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 CPHX) MODE_IMM2 (Byte 〈x6,x5〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 CPHX) MODE_DIR1 (Byte 〈x7,x5〉) 〈x0,x4〉
+ quadruple … (anyOP HC08 STHX) MODE_DIR1 (Byte 〈x3,x5〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 LDHX) MODE_IMM2 (Byte 〈x4,x5〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 LDHX) MODE_DIR1 (Byte 〈x5,x5〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 CPHX) MODE_IMM2 (Byte 〈x6,x5〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 CPHX) MODE_DIR1 (Byte 〈x7,x5〉) 〈x0,x4〉
].
ndefinition opcode_table_HC08_16 ≝
[
- quadruple ???? (anyOP HC08 CPX) MODE_IMM1 (Byte 〈xA,x3〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 CPX) MODE_DIR1 (Byte 〈xB,x3〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 CPX) MODE_DIR2 (Byte 〈xC,x3〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 CPX) MODE_IX2 (Byte 〈xD,x3〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 CPX) MODE_IX1 (Byte 〈xE,x3〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 CPX) MODE_IX0 (Byte 〈xF,x3〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 CPX) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x3〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 CPX) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x3〉〉) 〈x0,x4〉
+ quadruple … (anyOP HC08 CPX) MODE_IMM1 (Byte 〈xA,x3〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 CPX) MODE_DIR1 (Byte 〈xB,x3〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 CPX) MODE_DIR2 (Byte 〈xC,x3〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 CPX) MODE_IX2 (Byte 〈xD,x3〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 CPX) MODE_IX1 (Byte 〈xE,x3〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 CPX) MODE_IX0 (Byte 〈xF,x3〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 CPX) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x3〉〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 CPX) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x3〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HC08_17 ≝
[
- quadruple ???? (anyOP HC08 DBNZ) MODE_DIR1_and_IMM1 (Byte 〈x3,xB〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 DBNZ) MODE_INHA_and_IMM1 (Byte 〈x4,xB〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 DBNZ) MODE_INHX_and_IMM1 (Byte 〈x5,xB〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 DBNZ) MODE_IX1_and_IMM1 (Byte 〈x6,xB〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 DBNZ) MODE_IX0_and_IMM1 (Byte 〈x7,xB〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 DBNZ) MODE_SP1_and_IMM1 (Word 〈〈x9,xE〉:〈x6,xB〉〉) 〈x0,x6〉
+ quadruple … (anyOP HC08 DBNZ) MODE_DIR1_and_IMM1 (Byte 〈x3,xB〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 DBNZ) MODE_INHA_and_IMM1 (Byte 〈x4,xB〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 DBNZ) MODE_INHX_and_IMM1 (Byte 〈x5,xB〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 DBNZ) MODE_IX1_and_IMM1 (Byte 〈x6,xB〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 DBNZ) MODE_IX0_and_IMM1 (Byte 〈x7,xB〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 DBNZ) MODE_SP1_and_IMM1 (Word 〈〈x9,xE〉:〈x6,xB〉〉) 〈x0,x6〉
].
ndefinition opcode_table_HC08_18 ≝
[
- quadruple ???? (anyOP HC08 DEC) MODE_DIR1 (Byte 〈x3,xA〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 DEC) MODE_INHA (Byte 〈x4,xA〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 DEC) MODE_INHX (Byte 〈x5,xA〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 DEC) MODE_IX1 (Byte 〈x6,xA〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 DEC) MODE_IX0 (Byte 〈x7,xA〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 DEC) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,xA〉〉) 〈x0,x5〉
+ quadruple … (anyOP HC08 DEC) MODE_DIR1 (Byte 〈x3,xA〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 DEC) MODE_INHA (Byte 〈x4,xA〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 DEC) MODE_INHX (Byte 〈x5,xA〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 DEC) MODE_IX1 (Byte 〈x6,xA〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 DEC) MODE_IX0 (Byte 〈x7,xA〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 DEC) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,xA〉〉) 〈x0,x5〉
].
ndefinition opcode_table_HC08_19 ≝
[
- quadruple ???? (anyOP HC08 EOR) MODE_IMM1 (Byte 〈xA,x8〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 EOR) MODE_DIR1 (Byte 〈xB,x8〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 EOR) MODE_DIR2 (Byte 〈xC,x8〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 EOR) MODE_IX2 (Byte 〈xD,x8〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 EOR) MODE_IX1 (Byte 〈xE,x8〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 EOR) MODE_IX0 (Byte 〈xF,x8〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 EOR) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x8〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 EOR) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x8〉〉) 〈x0,x4〉
+ quadruple … (anyOP HC08 EOR) MODE_IMM1 (Byte 〈xA,x8〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 EOR) MODE_DIR1 (Byte 〈xB,x8〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 EOR) MODE_DIR2 (Byte 〈xC,x8〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 EOR) MODE_IX2 (Byte 〈xD,x8〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 EOR) MODE_IX1 (Byte 〈xE,x8〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 EOR) MODE_IX0 (Byte 〈xF,x8〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 EOR) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x8〉〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 EOR) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x8〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HC08_20 ≝
[
- quadruple ???? (anyOP HC08 INC) MODE_DIR1 (Byte 〈x3,xC〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 INC) MODE_INHA (Byte 〈x4,xC〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 INC) MODE_INHX (Byte 〈x5,xC〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 INC) MODE_IX1 (Byte 〈x6,xC〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 INC) MODE_IX0 (Byte 〈x7,xC〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 INC) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,xC〉〉) 〈x0,x5〉
+ quadruple … (anyOP HC08 INC) MODE_DIR1 (Byte 〈x3,xC〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 INC) MODE_INHA (Byte 〈x4,xC〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 INC) MODE_INHX (Byte 〈x5,xC〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 INC) MODE_IX1 (Byte 〈x6,xC〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 INC) MODE_IX0 (Byte 〈x7,xC〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 INC) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,xC〉〉) 〈x0,x5〉
].
ndefinition opcode_table_HC08_21 ≝
[
- quadruple ???? (anyOP HC08 JMP) MODE_IMM1EXT (Byte 〈xB,xC〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 JMP) MODE_IMM2 (Byte 〈xC,xC〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 JMP) MODE_INHX2ADD (Byte 〈xD,xC〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 JMP) MODE_INHX1ADD (Byte 〈xE,xC〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 JMP) MODE_INHX0ADD (Byte 〈xF,xC〉) 〈x0,x3〉
+ quadruple … (anyOP HC08 JMP) MODE_IMM1EXT (Byte 〈xB,xC〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 JMP) MODE_IMM2 (Byte 〈xC,xC〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 JMP) MODE_INHX2ADD (Byte 〈xD,xC〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 JMP) MODE_INHX1ADD (Byte 〈xE,xC〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 JMP) MODE_INHX0ADD (Byte 〈xF,xC〉) 〈x0,x3〉
].
ndefinition opcode_table_HC08_22 ≝
[
- quadruple ???? (anyOP HC08 BSR) MODE_IMM1 (Byte 〈xA,xD〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 JSR) MODE_IMM1EXT (Byte 〈xB,xD〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 JSR) MODE_IMM2 (Byte 〈xC,xD〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 JSR) MODE_INHX2ADD (Byte 〈xD,xD〉) 〈x0,x6〉
-; quadruple ???? (anyOP HC08 JSR) MODE_INHX1ADD (Byte 〈xE,xD〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 JSR) MODE_INHX0ADD (Byte 〈xF,xD〉) 〈x0,x4〉
+ quadruple … (anyOP HC08 BSR) MODE_IMM1 (Byte 〈xA,xD〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 JSR) MODE_IMM1EXT (Byte 〈xB,xD〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 JSR) MODE_IMM2 (Byte 〈xC,xD〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 JSR) MODE_INHX2ADD (Byte 〈xD,xD〉) 〈x0,x6〉
+; quadruple … (anyOP HC08 JSR) MODE_INHX1ADD (Byte 〈xE,xD〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 JSR) MODE_INHX0ADD (Byte 〈xF,xD〉) 〈x0,x4〉
].
ndefinition opcode_table_HC08_23 ≝
[
- quadruple ???? (anyOP HC08 LDA) MODE_IMM1 (Byte 〈xA,x6〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 LDA) MODE_DIR1 (Byte 〈xB,x6〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 LDA) MODE_DIR2 (Byte 〈xC,x6〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 LDA) MODE_IX2 (Byte 〈xD,x6〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 LDA) MODE_IX1 (Byte 〈xE,x6〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 LDA) MODE_IX0 (Byte 〈xF,x6〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 LDA) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x6〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 LDA) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x6〉〉) 〈x0,x4〉
+ quadruple … (anyOP HC08 LDA) MODE_IMM1 (Byte 〈xA,x6〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 LDA) MODE_DIR1 (Byte 〈xB,x6〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 LDA) MODE_DIR2 (Byte 〈xC,x6〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 LDA) MODE_IX2 (Byte 〈xD,x6〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 LDA) MODE_IX1 (Byte 〈xE,x6〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 LDA) MODE_IX0 (Byte 〈xF,x6〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 LDA) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x6〉〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 LDA) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x6〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HC08_24 ≝
[
- quadruple ???? (anyOP HC08 LDX) MODE_IMM1 (Byte 〈xA,xE〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 LDX) MODE_DIR1 (Byte 〈xB,xE〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 LDX) MODE_DIR2 (Byte 〈xC,xE〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 LDX) MODE_IX2 (Byte 〈xD,xE〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 LDX) MODE_IX1 (Byte 〈xE,xE〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 LDX) MODE_IX0 (Byte 〈xF,xE〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 LDX) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,xE〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 LDX) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,xE〉〉) 〈x0,x4〉
+ quadruple … (anyOP HC08 LDX) MODE_IMM1 (Byte 〈xA,xE〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 LDX) MODE_DIR1 (Byte 〈xB,xE〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 LDX) MODE_DIR2 (Byte 〈xC,xE〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 LDX) MODE_IX2 (Byte 〈xD,xE〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 LDX) MODE_IX1 (Byte 〈xE,xE〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 LDX) MODE_IX0 (Byte 〈xF,xE〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 LDX) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,xE〉〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 LDX) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,xE〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HC08_25 ≝
[
- quadruple ???? (anyOP HC08 LSR) MODE_DIR1 (Byte 〈x3,x4〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 LSR) MODE_INHA (Byte 〈x4,x4〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 LSR) MODE_INHX (Byte 〈x5,x4〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 LSR) MODE_IX1 (Byte 〈x6,x4〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 LSR) MODE_IX0 (Byte 〈x7,x4〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 LSR) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x4〉〉) 〈x0,x5〉
+ quadruple … (anyOP HC08 LSR) MODE_DIR1 (Byte 〈x3,x4〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 LSR) MODE_INHA (Byte 〈x4,x4〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 LSR) MODE_INHX (Byte 〈x5,x4〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 LSR) MODE_IX1 (Byte 〈x6,x4〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 LSR) MODE_IX0 (Byte 〈x7,x4〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 LSR) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x4〉〉) 〈x0,x5〉
].
ndefinition opcode_table_HC08_26 ≝
[
- quadruple ???? (anyOP HC08 MOV) MODE_DIR1_to_DIR1 (Byte 〈x4,xE〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 MOV) MODE_DIR1_to_IX0p (Byte 〈x5,xE〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 MOV) MODE_IMM1_to_DIR1 (Byte 〈x6,xE〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 MOV) MODE_IX0p_to_DIR1 (Byte 〈x7,xE〉) 〈x0,x4〉
+ quadruple … (anyOP HC08 MOV) MODE_DIR1_to_DIR1 (Byte 〈x4,xE〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 MOV) MODE_DIR1_to_IX0p (Byte 〈x5,xE〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 MOV) MODE_IMM1_to_DIR1 (Byte 〈x6,xE〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 MOV) MODE_IX0p_to_DIR1 (Byte 〈x7,xE〉) 〈x0,x4〉
].
ndefinition opcode_table_HC08_27 ≝
[
- quadruple ???? (anyOP HC08 NEG) MODE_DIR1 (Byte 〈x3,x0〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 NEG) MODE_INHA (Byte 〈x4,x0〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 NEG) MODE_INHX (Byte 〈x5,x0〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 NEG) MODE_IX1 (Byte 〈x6,x0〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 NEG) MODE_IX0 (Byte 〈x7,x0〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 NEG) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x0〉〉) 〈x0,x5〉
+ quadruple … (anyOP HC08 NEG) MODE_DIR1 (Byte 〈x3,x0〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 NEG) MODE_INHA (Byte 〈x4,x0〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 NEG) MODE_INHX (Byte 〈x5,x0〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 NEG) MODE_IX1 (Byte 〈x6,x0〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 NEG) MODE_IX0 (Byte 〈x7,x0〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 NEG) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x0〉〉) 〈x0,x5〉
].
ndefinition opcode_table_HC08_28 ≝
[
- quadruple ???? (anyOP HC08 ORA) MODE_IMM1 (Byte 〈xA,xA〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 ORA) MODE_DIR1 (Byte 〈xB,xA〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 ORA) MODE_DIR2 (Byte 〈xC,xA〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 ORA) MODE_IX2 (Byte 〈xD,xA〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 ORA) MODE_IX1 (Byte 〈xE,xA〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 ORA) MODE_IX0 (Byte 〈xF,xA〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 ORA) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,xA〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 ORA) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,xA〉〉) 〈x0,x4〉
+ quadruple … (anyOP HC08 ORA) MODE_IMM1 (Byte 〈xA,xA〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 ORA) MODE_DIR1 (Byte 〈xB,xA〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 ORA) MODE_DIR2 (Byte 〈xC,xA〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 ORA) MODE_IX2 (Byte 〈xD,xA〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 ORA) MODE_IX1 (Byte 〈xE,xA〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 ORA) MODE_IX0 (Byte 〈xF,xA〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 ORA) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,xA〉〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 ORA) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,xA〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HC08_29 ≝
[
- quadruple ???? (anyOP HC08 ROL) MODE_DIR1 (Byte 〈x3,x9〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 ROL) MODE_INHA (Byte 〈x4,x9〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 ROL) MODE_INHX (Byte 〈x5,x9〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 ROL) MODE_IX1 (Byte 〈x6,x9〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 ROL) MODE_IX0 (Byte 〈x7,x9〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 ROL) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x9〉〉) 〈x0,x5〉
+ quadruple … (anyOP HC08 ROL) MODE_DIR1 (Byte 〈x3,x9〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 ROL) MODE_INHA (Byte 〈x4,x9〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 ROL) MODE_INHX (Byte 〈x5,x9〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 ROL) MODE_IX1 (Byte 〈x6,x9〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 ROL) MODE_IX0 (Byte 〈x7,x9〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 ROL) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x9〉〉) 〈x0,x5〉
].
ndefinition opcode_table_HC08_30 ≝
[
- quadruple ???? (anyOP HC08 ROR) MODE_DIR1 (Byte 〈x3,x6〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 ROR) MODE_INHA (Byte 〈x4,x6〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 ROR) MODE_INHX (Byte 〈x5,x6〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 ROR) MODE_IX1 (Byte 〈x6,x6〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 ROR) MODE_IX0 (Byte 〈x7,x6〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 ROR) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x6〉〉) 〈x0,x5〉
+ quadruple … (anyOP HC08 ROR) MODE_DIR1 (Byte 〈x3,x6〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 ROR) MODE_INHA (Byte 〈x4,x6〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 ROR) MODE_INHX (Byte 〈x5,x6〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 ROR) MODE_IX1 (Byte 〈x6,x6〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 ROR) MODE_IX0 (Byte 〈x7,x6〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 ROR) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x6〉〉) 〈x0,x5〉
].
ndefinition opcode_table_HC08_31 ≝
[
- quadruple ???? (anyOP HC08 SBC) MODE_IMM1 (Byte 〈xA,x2〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 SBC) MODE_DIR1 (Byte 〈xB,x2〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 SBC) MODE_DIR2 (Byte 〈xC,x2〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 SBC) MODE_IX2 (Byte 〈xD,x2〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 SBC) MODE_IX1 (Byte 〈xE,x2〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 SBC) MODE_IX0 (Byte 〈xF,x2〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 SBC) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x2〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 SBC) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x2〉〉) 〈x0,x4〉
+ quadruple … (anyOP HC08 SBC) MODE_IMM1 (Byte 〈xA,x2〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 SBC) MODE_DIR1 (Byte 〈xB,x2〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 SBC) MODE_DIR2 (Byte 〈xC,x2〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 SBC) MODE_IX2 (Byte 〈xD,x2〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 SBC) MODE_IX1 (Byte 〈xE,x2〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 SBC) MODE_IX0 (Byte 〈xF,x2〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 SBC) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x2〉〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 SBC) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x2〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HC08_32 ≝
[
- quadruple ???? (anyOP HC08 STA) MODE_DIR1 (Byte 〈xB,x7〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 STA) MODE_DIR2 (Byte 〈xC,x7〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 STA) MODE_IX2 (Byte 〈xD,x7〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 STA) MODE_IX1 (Byte 〈xE,x7〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 STA) MODE_IX0 (Byte 〈xF,x7〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 STA) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x7〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 STA) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x7〉〉) 〈x0,x4〉
+ quadruple … (anyOP HC08 STA) MODE_DIR1 (Byte 〈xB,x7〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 STA) MODE_DIR2 (Byte 〈xC,x7〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 STA) MODE_IX2 (Byte 〈xD,x7〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 STA) MODE_IX1 (Byte 〈xE,x7〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 STA) MODE_IX0 (Byte 〈xF,x7〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 STA) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x7〉〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 STA) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x7〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HC08_33 ≝
[
- quadruple ???? (anyOP HC08 STX) MODE_DIR1 (Byte 〈xB,xF〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 STX) MODE_DIR2 (Byte 〈xC,xF〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 STX) MODE_IX2 (Byte 〈xD,xF〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 STX) MODE_IX1 (Byte 〈xE,xF〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 STX) MODE_IX0 (Byte 〈xF,xF〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 STX) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,xF〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 STX) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,xF〉〉) 〈x0,x4〉
+ quadruple … (anyOP HC08 STX) MODE_DIR1 (Byte 〈xB,xF〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 STX) MODE_DIR2 (Byte 〈xC,xF〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 STX) MODE_IX2 (Byte 〈xD,xF〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 STX) MODE_IX1 (Byte 〈xE,xF〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 STX) MODE_IX0 (Byte 〈xF,xF〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 STX) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,xF〉〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 STX) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,xF〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HC08_34 ≝
[
- quadruple ???? (anyOP HC08 SUB) MODE_IMM1 (Byte 〈xA,x0〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 SUB) MODE_DIR1 (Byte 〈xB,x0〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 SUB) MODE_DIR2 (Byte 〈xC,x0〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 SUB) MODE_IX2 (Byte 〈xD,x0〉) 〈x0,x4〉
-; quadruple ???? (anyOP HC08 SUB) MODE_IX1 (Byte 〈xE,x0〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 SUB) MODE_IX0 (Byte 〈xF,x0〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 SUB) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x0〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HC08 SUB) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x0〉〉) 〈x0,x4〉
+ quadruple … (anyOP HC08 SUB) MODE_IMM1 (Byte 〈xA,x0〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 SUB) MODE_DIR1 (Byte 〈xB,x0〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 SUB) MODE_DIR2 (Byte 〈xC,x0〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 SUB) MODE_IX2 (Byte 〈xD,x0〉) 〈x0,x4〉
+; quadruple … (anyOP HC08 SUB) MODE_IX1 (Byte 〈xE,x0〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 SUB) MODE_IX0 (Byte 〈xF,x0〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 SUB) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x0〉〉) 〈x0,x5〉
+; quadruple … (anyOP HC08 SUB) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x0〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HC08_35 ≝
[
- quadruple ???? (anyOP HC08 TST) MODE_DIR1 (Byte 〈x3,xD〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 TST) MODE_INHA (Byte 〈x4,xD〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 TST) MODE_INHX (Byte 〈x5,xD〉) 〈x0,x1〉
-; quadruple ???? (anyOP HC08 TST) MODE_IX1 (Byte 〈x6,xD〉) 〈x0,x3〉
-; quadruple ???? (anyOP HC08 TST) MODE_IX0 (Byte 〈x7,xD〉) 〈x0,x2〉
-; quadruple ???? (anyOP HC08 TST) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,xD〉〉) 〈x0,x4〉
+ quadruple … (anyOP HC08 TST) MODE_DIR1 (Byte 〈x3,xD〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 TST) MODE_INHA (Byte 〈x4,xD〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 TST) MODE_INHX (Byte 〈x5,xD〉) 〈x0,x1〉
+; quadruple … (anyOP HC08 TST) MODE_IX1 (Byte 〈x6,xD〉) 〈x0,x3〉
+; quadruple … (anyOP HC08 TST) MODE_IX0 (Byte 〈x7,xD〉) 〈x0,x2〉
+; quadruple … (anyOP HC08 TST) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,xD〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HC08 ≝
(test_not_impl_byte b HC08_not_impl_byte ⊙ eq_nat (get_byte_count HC08 b 0 opcode_table_HC08) 1) ⊗
(⊖ (test_not_impl_byte b HC08_not_impl_byte) ⊙ eq_nat (get_byte_count HC08 b 0 opcode_table_HC08) 0))
= true.
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
(* HC08: opcode non implementati come da manuale (0x9E+byte) *)
(test_not_impl_byte b HC08_not_impl_word ⊙ eq_nat (get_word_count HC08 b 0 opcode_table_HC08) 1) ⊗
(⊖ (test_not_impl_byte b HC08_not_impl_word) ⊙ eq_nat (get_word_count HC08 b 0 opcode_table_HC08) 0))
= true.
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
(* HC08: pseudocodici non implementati come da manuale *)
(test_not_impl_pseudo o HC08_not_impl_pseudo ⊙ le_nat 1 (get_pseudo_count HC08 o 0 opcode_table_HC08)) ⊗
(⊖ (test_not_impl_pseudo o HC08_not_impl_pseudo) ⊙ eq_nat (get_pseudo_count HC08 o 0 opcode_table_HC08) 0))
= true.
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
(* HC08: modalita' non implementate come da manuale *)
(test_not_impl_mode i HC08_not_impl_mode ⊙ le_nat 1 (get_mode_count HC08 i 0 opcode_table_HC08)) ⊗
(⊖ (test_not_impl_mode i HC08_not_impl_mode) ⊙ eq_nat (get_mode_count HC08 i 0 opcode_table_HC08) 0))
= true.
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma ok_OpIm_table_HC08 :
forall_instr_mode (λi:instr_mode.
forall_opcode (λop:opcode.
le_nat (get_OpIm_count HC08 (anyOP HC08 op) i 0 opcode_table_HC08) 1)) = true.
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
ndefinition opcode_table_HCS08_1 ≝
[
- quadruple ???? (anyOP HCS08 ADC) MODE_IMM1 (Byte 〈xA,x9〉) 〈x0,x2〉
-; quadruple ???? (anyOP HCS08 ADC) MODE_DIR1 (Byte 〈xB,x9〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 ADC) MODE_DIR2 (Byte 〈xC,x9〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 ADC) MODE_IX2 (Byte 〈xD,x9〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 ADC) MODE_IX1 (Byte 〈xE,x9〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 ADC) MODE_IX0 (Byte 〈xF,x9〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 ADC) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x9〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 ADC) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x9〉〉) 〈x0,x4〉
+ quadruple … (anyOP HCS08 ADC) MODE_IMM1 (Byte 〈xA,x9〉) 〈x0,x2〉
+; quadruple … (anyOP HCS08 ADC) MODE_DIR1 (Byte 〈xB,x9〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 ADC) MODE_DIR2 (Byte 〈xC,x9〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 ADC) MODE_IX2 (Byte 〈xD,x9〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 ADC) MODE_IX1 (Byte 〈xE,x9〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 ADC) MODE_IX0 (Byte 〈xF,x9〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 ADC) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x9〉〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 ADC) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x9〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HCS08_2 ≝
[
- quadruple ???? (anyOP HCS08 ADD) MODE_IMM1 (Byte 〈xA,xB〉) 〈x0,x2〉
-; quadruple ???? (anyOP HCS08 ADD) MODE_DIR1 (Byte 〈xB,xB〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 ADD) MODE_DIR2 (Byte 〈xC,xB〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 ADD) MODE_IX2 (Byte 〈xD,xB〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 ADD) MODE_IX1 (Byte 〈xE,xB〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 ADD) MODE_IX0 (Byte 〈xF,xB〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 ADD) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,xB〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 ADD) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,xB〉〉) 〈x0,x4〉
+ quadruple … (anyOP HCS08 ADD) MODE_IMM1 (Byte 〈xA,xB〉) 〈x0,x2〉
+; quadruple … (anyOP HCS08 ADD) MODE_DIR1 (Byte 〈xB,xB〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 ADD) MODE_DIR2 (Byte 〈xC,xB〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 ADD) MODE_IX2 (Byte 〈xD,xB〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 ADD) MODE_IX1 (Byte 〈xE,xB〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 ADD) MODE_IX0 (Byte 〈xF,xB〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 ADD) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,xB〉〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 ADD) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,xB〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HCS08_3 ≝
[
- quadruple ???? (anyOP HCS08 AND) MODE_IMM1 (Byte 〈xA,x4〉) 〈x0,x2〉
-; quadruple ???? (anyOP HCS08 AND) MODE_DIR1 (Byte 〈xB,x4〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 AND) MODE_DIR2 (Byte 〈xC,x4〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 AND) MODE_IX2 (Byte 〈xD,x4〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 AND) MODE_IX1 (Byte 〈xE,x4〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 AND) MODE_IX0 (Byte 〈xF,x4〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 AND) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x4〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 AND) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x4〉〉) 〈x0,x4〉
+ quadruple … (anyOP HCS08 AND) MODE_IMM1 (Byte 〈xA,x4〉) 〈x0,x2〉
+; quadruple … (anyOP HCS08 AND) MODE_DIR1 (Byte 〈xB,x4〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 AND) MODE_DIR2 (Byte 〈xC,x4〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 AND) MODE_IX2 (Byte 〈xD,x4〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 AND) MODE_IX1 (Byte 〈xE,x4〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 AND) MODE_IX0 (Byte 〈xF,x4〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 AND) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x4〉〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 AND) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x4〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HCS08_4 ≝
[
- quadruple ???? (anyOP HCS08 ASL) MODE_DIR1 (Byte 〈x3,x8〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 ASL) MODE_INHA (Byte 〈x4,x8〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 ASL) MODE_INHX (Byte 〈x5,x8〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 ASL) MODE_IX1 (Byte 〈x6,x8〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 ASL) MODE_IX0 (Byte 〈x7,x8〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 ASL) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x8〉〉) 〈x0,x6〉
+ quadruple … (anyOP HCS08 ASL) MODE_DIR1 (Byte 〈x3,x8〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 ASL) MODE_INHA (Byte 〈x4,x8〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 ASL) MODE_INHX (Byte 〈x5,x8〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 ASL) MODE_IX1 (Byte 〈x6,x8〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 ASL) MODE_IX0 (Byte 〈x7,x8〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 ASL) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x8〉〉) 〈x0,x6〉
].
ndefinition opcode_table_HCS08_5 ≝
[
- quadruple ???? (anyOP HCS08 ASR) MODE_DIR1 (Byte 〈x3,x7〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 ASR) MODE_INHA (Byte 〈x4,x7〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 ASR) MODE_INHX (Byte 〈x5,x7〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 ASR) MODE_IX1 (Byte 〈x6,x7〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 ASR) MODE_IX0 (Byte 〈x7,x7〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 ASR) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x7〉〉) 〈x0,x6〉
+ quadruple … (anyOP HCS08 ASR) MODE_DIR1 (Byte 〈x3,x7〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 ASR) MODE_INHA (Byte 〈x4,x7〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 ASR) MODE_INHX (Byte 〈x5,x7〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 ASR) MODE_IX1 (Byte 〈x6,x7〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 ASR) MODE_IX0 (Byte 〈x7,x7〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 ASR) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x7〉〉) 〈x0,x6〉
].
ndefinition opcode_table_HCS08_6 ≝
[
- quadruple ???? (anyOP HCS08 BRA ) MODE_IMM1 (Byte 〈x2,x0〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 BRN ) MODE_IMM1 (Byte 〈x2,x1〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 BHI ) MODE_IMM1 (Byte 〈x2,x2〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 BLS ) MODE_IMM1 (Byte 〈x2,x3〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 BCC ) MODE_IMM1 (Byte 〈x2,x4〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 BCS ) MODE_IMM1 (Byte 〈x2,x5〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 BNE ) MODE_IMM1 (Byte 〈x2,x6〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 BEQ ) MODE_IMM1 (Byte 〈x2,x7〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 BHCC) MODE_IMM1 (Byte 〈x2,x8〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 BHCS) MODE_IMM1 (Byte 〈x2,x9〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 BPL ) MODE_IMM1 (Byte 〈x2,xA〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 BMI ) MODE_IMM1 (Byte 〈x2,xB〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 BMC ) MODE_IMM1 (Byte 〈x2,xC〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 BMS ) MODE_IMM1 (Byte 〈x2,xD〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 BIL ) MODE_IMM1 (Byte 〈x2,xE〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 BIH ) MODE_IMM1 (Byte 〈x2,xF〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 BGE ) MODE_IMM1 (Byte 〈x9,x0〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 BLT ) MODE_IMM1 (Byte 〈x9,x1〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 BGT ) MODE_IMM1 (Byte 〈x9,x2〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 BLE ) MODE_IMM1 (Byte 〈x9,x3〉) 〈x0,x3〉
+ quadruple … (anyOP HCS08 BRA ) MODE_IMM1 (Byte 〈x2,x0〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 BRN ) MODE_IMM1 (Byte 〈x2,x1〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 BHI ) MODE_IMM1 (Byte 〈x2,x2〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 BLS ) MODE_IMM1 (Byte 〈x2,x3〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 BCC ) MODE_IMM1 (Byte 〈x2,x4〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 BCS ) MODE_IMM1 (Byte 〈x2,x5〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 BNE ) MODE_IMM1 (Byte 〈x2,x6〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 BEQ ) MODE_IMM1 (Byte 〈x2,x7〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 BHCC) MODE_IMM1 (Byte 〈x2,x8〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 BHCS) MODE_IMM1 (Byte 〈x2,x9〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 BPL ) MODE_IMM1 (Byte 〈x2,xA〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 BMI ) MODE_IMM1 (Byte 〈x2,xB〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 BMC ) MODE_IMM1 (Byte 〈x2,xC〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 BMS ) MODE_IMM1 (Byte 〈x2,xD〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 BIL ) MODE_IMM1 (Byte 〈x2,xE〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 BIH ) MODE_IMM1 (Byte 〈x2,xF〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 BGE ) MODE_IMM1 (Byte 〈x9,x0〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 BLT ) MODE_IMM1 (Byte 〈x9,x1〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 BGT ) MODE_IMM1 (Byte 〈x9,x2〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 BLE ) MODE_IMM1 (Byte 〈x9,x3〉) 〈x0,x3〉
].
ndefinition opcode_table_HCS08_7 ≝
[
- quadruple ???? (anyOP HCS08 BSETn) (MODE_DIRn o0) (Byte 〈x1,x0〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BCLRn) (MODE_DIRn o0) (Byte 〈x1,x1〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BSETn) (MODE_DIRn o1) (Byte 〈x1,x2〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BCLRn) (MODE_DIRn o1) (Byte 〈x1,x3〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BSETn) (MODE_DIRn o2) (Byte 〈x1,x4〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BCLRn) (MODE_DIRn o2) (Byte 〈x1,x5〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BSETn) (MODE_DIRn o3) (Byte 〈x1,x6〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BCLRn) (MODE_DIRn o3) (Byte 〈x1,x7〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BSETn) (MODE_DIRn o4) (Byte 〈x1,x8〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BCLRn) (MODE_DIRn o4) (Byte 〈x1,x9〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BSETn) (MODE_DIRn o5) (Byte 〈x1,xA〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BCLRn) (MODE_DIRn o5) (Byte 〈x1,xB〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BSETn) (MODE_DIRn o6) (Byte 〈x1,xC〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BCLRn) (MODE_DIRn o6) (Byte 〈x1,xD〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BSETn) (MODE_DIRn o7) (Byte 〈x1,xE〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BCLRn) (MODE_DIRn o7) (Byte 〈x1,xF〉) 〈x0,x5〉
+ quadruple … (anyOP HCS08 BSETn) (MODE_DIRn o0) (Byte 〈x1,x0〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BCLRn) (MODE_DIRn o0) (Byte 〈x1,x1〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BSETn) (MODE_DIRn o1) (Byte 〈x1,x2〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BCLRn) (MODE_DIRn o1) (Byte 〈x1,x3〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BSETn) (MODE_DIRn o2) (Byte 〈x1,x4〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BCLRn) (MODE_DIRn o2) (Byte 〈x1,x5〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BSETn) (MODE_DIRn o3) (Byte 〈x1,x6〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BCLRn) (MODE_DIRn o3) (Byte 〈x1,x7〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BSETn) (MODE_DIRn o4) (Byte 〈x1,x8〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BCLRn) (MODE_DIRn o4) (Byte 〈x1,x9〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BSETn) (MODE_DIRn o5) (Byte 〈x1,xA〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BCLRn) (MODE_DIRn o5) (Byte 〈x1,xB〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BSETn) (MODE_DIRn o6) (Byte 〈x1,xC〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BCLRn) (MODE_DIRn o6) (Byte 〈x1,xD〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BSETn) (MODE_DIRn o7) (Byte 〈x1,xE〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BCLRn) (MODE_DIRn o7) (Byte 〈x1,xF〉) 〈x0,x5〉
].
ndefinition opcode_table_HCS08_8 ≝
[
- quadruple ???? (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x0〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x1〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x2〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x3〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x4〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x5〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x6〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x7〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x8〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x9〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xA〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xB〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xC〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xD〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xE〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xF〉) 〈x0,x5〉
+ quadruple … (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x0〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x1〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x2〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x3〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x4〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x5〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x6〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x7〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x8〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x9〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xA〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xB〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xC〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xD〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xE〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xF〉) 〈x0,x5〉
].
ndefinition opcode_table_HCS08_9 ≝
[
- quadruple ???? (anyOP HCS08 BIT) MODE_IMM1 (Byte 〈xA,x5〉) 〈x0,x2〉
-; quadruple ???? (anyOP HCS08 BIT) MODE_DIR1 (Byte 〈xB,x5〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 BIT) MODE_DIR2 (Byte 〈xC,x5〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 BIT) MODE_IX2 (Byte 〈xD,x5〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 BIT) MODE_IX1 (Byte 〈xE,x5〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 BIT) MODE_IX0 (Byte 〈xF,x5〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 BIT) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x5〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 BIT) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x5〉〉) 〈x0,x4〉
+ quadruple … (anyOP HCS08 BIT) MODE_IMM1 (Byte 〈xA,x5〉) 〈x0,x2〉
+; quadruple … (anyOP HCS08 BIT) MODE_DIR1 (Byte 〈xB,x5〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 BIT) MODE_DIR2 (Byte 〈xC,x5〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 BIT) MODE_IX2 (Byte 〈xD,x5〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 BIT) MODE_IX1 (Byte 〈xE,x5〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 BIT) MODE_IX0 (Byte 〈xF,x5〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 BIT) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x5〉〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 BIT) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x5〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HCS08_10 ≝
[
- quadruple ???? (anyOP HCS08 MUL ) MODE_INH (Byte 〈x4,x2〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 DIV ) MODE_INH (Byte 〈x5,x2〉) 〈x0,x6〉
-; quadruple ???? (anyOP HCS08 NSA ) MODE_INH (Byte 〈x6,x2〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 DAA ) MODE_INH (Byte 〈x7,x2〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 RTI ) MODE_INH (Byte 〈x8,x0〉) 〈x0,x9〉
-; quadruple ???? (anyOP HCS08 RTS ) MODE_INH (Byte 〈x8,x1〉) 〈x0,x6〉
-; quadruple ???? (anyOP HCS08 SWI ) MODE_INH (Byte 〈x8,x3〉) 〈x0,xB〉
-; quadruple ???? (anyOP HCS08 BGND) MODE_INH (Byte 〈x8,x2〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 TAP ) MODE_INH (Byte 〈x8,x4〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 TPA ) MODE_INH (Byte 〈x8,x5〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 PULA) MODE_INH (Byte 〈x8,x6〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 PSHA) MODE_INH (Byte 〈x8,x7〉) 〈x0,x2〉
-; quadruple ???? (anyOP HCS08 PULX) MODE_INH (Byte 〈x8,x8〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 PSHX) MODE_INH (Byte 〈x8,x9〉) 〈x0,x2〉
-; quadruple ???? (anyOP HCS08 PULH) MODE_INH (Byte 〈x8,xA〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 PSHH) MODE_INH (Byte 〈x8,xB〉) 〈x0,x2〉
-; quadruple ???? (anyOP HCS08 STOP) MODE_INH (Byte 〈x8,xE〉) 〈x0,x2〉
-; quadruple ???? (anyOP HCS08 WAIT) MODE_INH (Byte 〈x8,xF〉) 〈x0,x2〉
-; quadruple ???? (anyOP HCS08 TXS ) MODE_INH (Byte 〈x9,x4〉) 〈x0,x2〉
-; quadruple ???? (anyOP HCS08 TSX ) MODE_INH (Byte 〈x9,x5〉) 〈x0,x2〉
-; quadruple ???? (anyOP HCS08 TAX ) MODE_INH (Byte 〈x9,x7〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 CLC ) MODE_INH (Byte 〈x9,x8〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 SEC ) MODE_INH (Byte 〈x9,x9〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 CLI ) MODE_INH (Byte 〈x9,xA〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 SEI ) MODE_INH (Byte 〈x9,xB〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 RSP ) MODE_INH (Byte 〈x9,xC〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 NOP ) MODE_INH (Byte 〈x9,xD〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 TXA ) MODE_INH (Byte 〈x9,xF〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 AIS ) MODE_IMM1 (Byte 〈xA,x7〉) 〈x0,x2〉
-; quadruple ???? (anyOP HCS08 AIX ) MODE_IMM1 (Byte 〈xA,xF〉) 〈x0,x2〉
+ quadruple … (anyOP HCS08 MUL ) MODE_INH (Byte 〈x4,x2〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 DIV ) MODE_INH (Byte 〈x5,x2〉) 〈x0,x6〉
+; quadruple … (anyOP HCS08 NSA ) MODE_INH (Byte 〈x6,x2〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 DAA ) MODE_INH (Byte 〈x7,x2〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 RTI ) MODE_INH (Byte 〈x8,x0〉) 〈x0,x9〉
+; quadruple … (anyOP HCS08 RTS ) MODE_INH (Byte 〈x8,x1〉) 〈x0,x6〉
+; quadruple … (anyOP HCS08 SWI ) MODE_INH (Byte 〈x8,x3〉) 〈x0,xB〉
+; quadruple … (anyOP HCS08 BGND) MODE_INH (Byte 〈x8,x2〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 TAP ) MODE_INH (Byte 〈x8,x4〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 TPA ) MODE_INH (Byte 〈x8,x5〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 PULA) MODE_INH (Byte 〈x8,x6〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 PSHA) MODE_INH (Byte 〈x8,x7〉) 〈x0,x2〉
+; quadruple … (anyOP HCS08 PULX) MODE_INH (Byte 〈x8,x8〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 PSHX) MODE_INH (Byte 〈x8,x9〉) 〈x0,x2〉
+; quadruple … (anyOP HCS08 PULH) MODE_INH (Byte 〈x8,xA〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 PSHH) MODE_INH (Byte 〈x8,xB〉) 〈x0,x2〉
+; quadruple … (anyOP HCS08 STOP) MODE_INH (Byte 〈x8,xE〉) 〈x0,x2〉
+; quadruple … (anyOP HCS08 WAIT) MODE_INH (Byte 〈x8,xF〉) 〈x0,x2〉
+; quadruple … (anyOP HCS08 TXS ) MODE_INH (Byte 〈x9,x4〉) 〈x0,x2〉
+; quadruple … (anyOP HCS08 TSX ) MODE_INH (Byte 〈x9,x5〉) 〈x0,x2〉
+; quadruple … (anyOP HCS08 TAX ) MODE_INH (Byte 〈x9,x7〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 CLC ) MODE_INH (Byte 〈x9,x8〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 SEC ) MODE_INH (Byte 〈x9,x9〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 CLI ) MODE_INH (Byte 〈x9,xA〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 SEI ) MODE_INH (Byte 〈x9,xB〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 RSP ) MODE_INH (Byte 〈x9,xC〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 NOP ) MODE_INH (Byte 〈x9,xD〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 TXA ) MODE_INH (Byte 〈x9,xF〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 AIS ) MODE_IMM1 (Byte 〈xA,x7〉) 〈x0,x2〉
+; quadruple … (anyOP HCS08 AIX ) MODE_IMM1 (Byte 〈xA,xF〉) 〈x0,x2〉
].
ndefinition opcode_table_HCS08_11 ≝
[
- quadruple ???? (anyOP HCS08 CBEQA) MODE_DIR1_and_IMM1 (Byte 〈x3,x1〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 CBEQA) MODE_IMM1_and_IMM1 (Byte 〈x4,x1〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 CBEQX) MODE_IMM1_and_IMM1 (Byte 〈x5,x1〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 CBEQA) MODE_IX1p_and_IMM1 (Byte 〈x6,x1〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 CBEQA) MODE_IX0p_and_IMM1 (Byte 〈x7,x1〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 CBEQA) MODE_SP1_and_IMM1 (Word 〈〈x9,xE〉:〈x6,x1〉〉) 〈x0,x6〉
+ quadruple … (anyOP HCS08 CBEQA) MODE_DIR1_and_IMM1 (Byte 〈x3,x1〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 CBEQA) MODE_IMM1_and_IMM1 (Byte 〈x4,x1〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 CBEQX) MODE_IMM1_and_IMM1 (Byte 〈x5,x1〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 CBEQA) MODE_IX1p_and_IMM1 (Byte 〈x6,x1〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 CBEQA) MODE_IX0p_and_IMM1 (Byte 〈x7,x1〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 CBEQA) MODE_SP1_and_IMM1 (Word 〈〈x9,xE〉:〈x6,x1〉〉) 〈x0,x6〉
].
ndefinition opcode_table_HCS08_12 ≝
[
- quadruple ???? (anyOP HCS08 CLR) MODE_DIR1 (Byte 〈x3,xF〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 CLR) MODE_INHA (Byte 〈x4,xF〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 CLR) MODE_INHX (Byte 〈x5,xF〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 CLR) MODE_IX1 (Byte 〈x6,xF〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 CLR) MODE_IX0 (Byte 〈x7,xF〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 CLR) MODE_INHH (Byte 〈x8,xC〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 CLR) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,xF〉〉) 〈x0,x6〉
+ quadruple … (anyOP HCS08 CLR) MODE_DIR1 (Byte 〈x3,xF〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 CLR) MODE_INHA (Byte 〈x4,xF〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 CLR) MODE_INHX (Byte 〈x5,xF〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 CLR) MODE_IX1 (Byte 〈x6,xF〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 CLR) MODE_IX0 (Byte 〈x7,xF〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 CLR) MODE_INHH (Byte 〈x8,xC〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 CLR) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,xF〉〉) 〈x0,x6〉
].
ndefinition opcode_table_HCS08_13 ≝
[
- quadruple ???? (anyOP HCS08 CMP) MODE_IMM1 (Byte 〈xA,x1〉) 〈x0,x2〉
-; quadruple ???? (anyOP HCS08 CMP) MODE_DIR1 (Byte 〈xB,x1〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 CMP) MODE_DIR2 (Byte 〈xC,x1〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 CMP) MODE_IX2 (Byte 〈xD,x1〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 CMP) MODE_IX1 (Byte 〈xE,x1〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 CMP) MODE_IX0 (Byte 〈xF,x1〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 CMP) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x1〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 CMP) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x1〉〉) 〈x0,x4〉
+ quadruple … (anyOP HCS08 CMP) MODE_IMM1 (Byte 〈xA,x1〉) 〈x0,x2〉
+; quadruple … (anyOP HCS08 CMP) MODE_DIR1 (Byte 〈xB,x1〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 CMP) MODE_DIR2 (Byte 〈xC,x1〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 CMP) MODE_IX2 (Byte 〈xD,x1〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 CMP) MODE_IX1 (Byte 〈xE,x1〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 CMP) MODE_IX0 (Byte 〈xF,x1〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 CMP) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x1〉〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 CMP) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x1〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HCS08_14 ≝
[
- quadruple ???? (anyOP HCS08 COM) MODE_DIR1 (Byte 〈x3,x3〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 COM) MODE_INHA (Byte 〈x4,x3〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 COM) MODE_INHX (Byte 〈x5,x3〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 COM) MODE_IX1 (Byte 〈x6,x3〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 COM) MODE_IX0 (Byte 〈x7,x3〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 COM) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x3〉〉) 〈x0,x6〉
+ quadruple … (anyOP HCS08 COM) MODE_DIR1 (Byte 〈x3,x3〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 COM) MODE_INHA (Byte 〈x4,x3〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 COM) MODE_INHX (Byte 〈x5,x3〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 COM) MODE_IX1 (Byte 〈x6,x3〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 COM) MODE_IX0 (Byte 〈x7,x3〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 COM) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x3〉〉) 〈x0,x6〉
].
ndefinition opcode_table_HCS08_15 ≝
[
- quadruple ???? (anyOP HCS08 CPHX) MODE_DIR2 (Byte 〈x3,xE〉) 〈x0,x6〉
-; quadruple ???? (anyOP HCS08 CPHX) MODE_IMM2 (Byte 〈x6,x5〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 CPHX) MODE_DIR1 (Byte 〈x7,x5〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 CPHX) MODE_SP1 (Word 〈〈x9,xE〉:〈xF,x3〉〉) 〈x0,x6〉
+ quadruple … (anyOP HCS08 CPHX) MODE_DIR2 (Byte 〈x3,xE〉) 〈x0,x6〉
+; quadruple … (anyOP HCS08 CPHX) MODE_IMM2 (Byte 〈x6,x5〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 CPHX) MODE_DIR1 (Byte 〈x7,x5〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 CPHX) MODE_SP1 (Word 〈〈x9,xE〉:〈xF,x3〉〉) 〈x0,x6〉
-; quadruple ???? (anyOP HCS08 LDHX) MODE_DIR2 (Byte 〈x3,x2〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 LDHX) MODE_IMM2 (Byte 〈x4,x5〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 LDHX) MODE_DIR1 (Byte 〈x5,x5〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 LDHX) MODE_IX0 (Word 〈〈x9,xE〉:〈xA,xE〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 LDHX) MODE_IX2 (Word 〈〈x9,xE〉:〈xB,xE〉〉) 〈x0,x6〉
-; quadruple ???? (anyOP HCS08 LDHX) MODE_IX1 (Word 〈〈x9,xE〉:〈xC,xE〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 LDHX) MODE_SP1 (Word 〈〈x9,xE〉:〈xF,xE〉〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 LDHX) MODE_DIR2 (Byte 〈x3,x2〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 LDHX) MODE_IMM2 (Byte 〈x4,x5〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 LDHX) MODE_DIR1 (Byte 〈x5,x5〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 LDHX) MODE_IX0 (Word 〈〈x9,xE〉:〈xA,xE〉〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 LDHX) MODE_IX2 (Word 〈〈x9,xE〉:〈xB,xE〉〉) 〈x0,x6〉
+; quadruple … (anyOP HCS08 LDHX) MODE_IX1 (Word 〈〈x9,xE〉:〈xC,xE〉〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 LDHX) MODE_SP1 (Word 〈〈x9,xE〉:〈xF,xE〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 STHX) MODE_DIR1 (Byte 〈x3,x5〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 STHX) MODE_DIR2 (Byte 〈x9,x6〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 STHX) MODE_SP1 (Word 〈〈x9,xE〉:〈xF,xF〉〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 STHX) MODE_DIR1 (Byte 〈x3,x5〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 STHX) MODE_DIR2 (Byte 〈x9,x6〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 STHX) MODE_SP1 (Word 〈〈x9,xE〉:〈xF,xF〉〉) 〈x0,x5〉
].
ndefinition opcode_table_HCS08_16 ≝
[
- quadruple ???? (anyOP HCS08 CPX) MODE_IMM1 (Byte 〈xA,x3〉) 〈x0,x2〉
-; quadruple ???? (anyOP HCS08 CPX) MODE_DIR1 (Byte 〈xB,x3〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 CPX) MODE_DIR2 (Byte 〈xC,x3〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 CPX) MODE_IX2 (Byte 〈xD,x3〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 CPX) MODE_IX1 (Byte 〈xE,x3〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 CPX) MODE_IX0 (Byte 〈xF,x3〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 CPX) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x3〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 CPX) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x3〉〉) 〈x0,x4〉
+ quadruple … (anyOP HCS08 CPX) MODE_IMM1 (Byte 〈xA,x3〉) 〈x0,x2〉
+; quadruple … (anyOP HCS08 CPX) MODE_DIR1 (Byte 〈xB,x3〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 CPX) MODE_DIR2 (Byte 〈xC,x3〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 CPX) MODE_IX2 (Byte 〈xD,x3〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 CPX) MODE_IX1 (Byte 〈xE,x3〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 CPX) MODE_IX0 (Byte 〈xF,x3〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 CPX) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x3〉〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 CPX) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x3〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HCS08_17 ≝
[
- quadruple ???? (anyOP HCS08 DBNZ) MODE_DIR1_and_IMM1 (Byte 〈x3,xB〉) 〈x0,x7〉
-; quadruple ???? (anyOP HCS08 DBNZ) MODE_INHA_and_IMM1 (Byte 〈x4,xB〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 DBNZ) MODE_INHX_and_IMM1 (Byte 〈x5,xB〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 DBNZ) MODE_IX1_and_IMM1 (Byte 〈x6,xB〉) 〈x0,x7〉
-; quadruple ???? (anyOP HCS08 DBNZ) MODE_IX0_and_IMM1 (Byte 〈x7,xB〉) 〈x0,x6〉
-; quadruple ???? (anyOP HCS08 DBNZ) MODE_SP1_and_IMM1 (Word 〈〈x9,xE〉:〈x6,xB〉〉) 〈x0,x8〉
+ quadruple … (anyOP HCS08 DBNZ) MODE_DIR1_and_IMM1 (Byte 〈x3,xB〉) 〈x0,x7〉
+; quadruple … (anyOP HCS08 DBNZ) MODE_INHA_and_IMM1 (Byte 〈x4,xB〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 DBNZ) MODE_INHX_and_IMM1 (Byte 〈x5,xB〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 DBNZ) MODE_IX1_and_IMM1 (Byte 〈x6,xB〉) 〈x0,x7〉
+; quadruple … (anyOP HCS08 DBNZ) MODE_IX0_and_IMM1 (Byte 〈x7,xB〉) 〈x0,x6〉
+; quadruple … (anyOP HCS08 DBNZ) MODE_SP1_and_IMM1 (Word 〈〈x9,xE〉:〈x6,xB〉〉) 〈x0,x8〉
].
ndefinition opcode_table_HCS08_18 ≝
[
- quadruple ???? (anyOP HCS08 DEC) MODE_DIR1 (Byte 〈x3,xA〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 DEC) MODE_INHA (Byte 〈x4,xA〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 DEC) MODE_INHX (Byte 〈x5,xA〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 DEC) MODE_IX1 (Byte 〈x6,xA〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 DEC) MODE_IX0 (Byte 〈x7,xA〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 DEC) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,xA〉〉) 〈x0,x6〉
+ quadruple … (anyOP HCS08 DEC) MODE_DIR1 (Byte 〈x3,xA〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 DEC) MODE_INHA (Byte 〈x4,xA〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 DEC) MODE_INHX (Byte 〈x5,xA〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 DEC) MODE_IX1 (Byte 〈x6,xA〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 DEC) MODE_IX0 (Byte 〈x7,xA〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 DEC) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,xA〉〉) 〈x0,x6〉
].
ndefinition opcode_table_HCS08_19 ≝
[
- quadruple ???? (anyOP HCS08 EOR) MODE_IMM1 (Byte 〈xA,x8〉) 〈x0,x2〉
-; quadruple ???? (anyOP HCS08 EOR) MODE_DIR1 (Byte 〈xB,x8〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 EOR) MODE_DIR2 (Byte 〈xC,x8〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 EOR) MODE_IX2 (Byte 〈xD,x8〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 EOR) MODE_IX1 (Byte 〈xE,x8〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 EOR) MODE_IX0 (Byte 〈xF,x8〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 EOR) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x8〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 EOR) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x8〉〉) 〈x0,x4〉
+ quadruple … (anyOP HCS08 EOR) MODE_IMM1 (Byte 〈xA,x8〉) 〈x0,x2〉
+; quadruple … (anyOP HCS08 EOR) MODE_DIR1 (Byte 〈xB,x8〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 EOR) MODE_DIR2 (Byte 〈xC,x8〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 EOR) MODE_IX2 (Byte 〈xD,x8〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 EOR) MODE_IX1 (Byte 〈xE,x8〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 EOR) MODE_IX0 (Byte 〈xF,x8〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 EOR) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x8〉〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 EOR) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x8〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HCS08_20 ≝
[
- quadruple ???? (anyOP HCS08 INC) MODE_DIR1 (Byte 〈x3,xC〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 INC) MODE_INHA (Byte 〈x4,xC〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 INC) MODE_INHX (Byte 〈x5,xC〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 INC) MODE_IX1 (Byte 〈x6,xC〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 INC) MODE_IX0 (Byte 〈x7,xC〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 INC) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,xC〉〉) 〈x0,x6〉
+ quadruple … (anyOP HCS08 INC) MODE_DIR1 (Byte 〈x3,xC〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 INC) MODE_INHA (Byte 〈x4,xC〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 INC) MODE_INHX (Byte 〈x5,xC〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 INC) MODE_IX1 (Byte 〈x6,xC〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 INC) MODE_IX0 (Byte 〈x7,xC〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 INC) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,xC〉〉) 〈x0,x6〉
].
ndefinition opcode_table_HCS08_21 ≝
[
- quadruple ???? (anyOP HCS08 JMP) MODE_IMM1EXT (Byte 〈xB,xC〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 JMP) MODE_IMM2 (Byte 〈xC,xC〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 JMP) MODE_INHX2ADD (Byte 〈xD,xC〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 JMP) MODE_INHX1ADD (Byte 〈xE,xC〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 JMP) MODE_INHX0ADD (Byte 〈xF,xC〉) 〈x0,x3〉
+ quadruple … (anyOP HCS08 JMP) MODE_IMM1EXT (Byte 〈xB,xC〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 JMP) MODE_IMM2 (Byte 〈xC,xC〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 JMP) MODE_INHX2ADD (Byte 〈xD,xC〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 JMP) MODE_INHX1ADD (Byte 〈xE,xC〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 JMP) MODE_INHX0ADD (Byte 〈xF,xC〉) 〈x0,x3〉
].
ndefinition opcode_table_HCS08_22 ≝
[
- quadruple ???? (anyOP HCS08 BSR) MODE_IMM1 (Byte 〈xA,xD〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 JSR) MODE_IMM1EXT (Byte 〈xB,xD〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 JSR) MODE_IMM2 (Byte 〈xC,xD〉) 〈x0,x6〉
-; quadruple ???? (anyOP HCS08 JSR) MODE_INHX2ADD (Byte 〈xD,xD〉) 〈x0,x6〉
-; quadruple ???? (anyOP HCS08 JSR) MODE_INHX1ADD (Byte 〈xE,xD〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 JSR) MODE_INHX0ADD (Byte 〈xF,xD〉) 〈x0,x5〉
+ quadruple … (anyOP HCS08 BSR) MODE_IMM1 (Byte 〈xA,xD〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 JSR) MODE_IMM1EXT (Byte 〈xB,xD〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 JSR) MODE_IMM2 (Byte 〈xC,xD〉) 〈x0,x6〉
+; quadruple … (anyOP HCS08 JSR) MODE_INHX2ADD (Byte 〈xD,xD〉) 〈x0,x6〉
+; quadruple … (anyOP HCS08 JSR) MODE_INHX1ADD (Byte 〈xE,xD〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 JSR) MODE_INHX0ADD (Byte 〈xF,xD〉) 〈x0,x5〉
].
ndefinition opcode_table_HCS08_23 ≝
[
- quadruple ???? (anyOP HCS08 LDA) MODE_IMM1 (Byte 〈xA,x6〉) 〈x0,x2〉
-; quadruple ???? (anyOP HCS08 LDA) MODE_DIR1 (Byte 〈xB,x6〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 LDA) MODE_DIR2 (Byte 〈xC,x6〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 LDA) MODE_IX2 (Byte 〈xD,x6〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 LDA) MODE_IX1 (Byte 〈xE,x6〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 LDA) MODE_IX0 (Byte 〈xF,x6〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 LDA) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x6〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 LDA) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x6〉〉) 〈x0,x4〉
+ quadruple … (anyOP HCS08 LDA) MODE_IMM1 (Byte 〈xA,x6〉) 〈x0,x2〉
+; quadruple … (anyOP HCS08 LDA) MODE_DIR1 (Byte 〈xB,x6〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 LDA) MODE_DIR2 (Byte 〈xC,x6〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 LDA) MODE_IX2 (Byte 〈xD,x6〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 LDA) MODE_IX1 (Byte 〈xE,x6〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 LDA) MODE_IX0 (Byte 〈xF,x6〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 LDA) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x6〉〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 LDA) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x6〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HCS08_24 ≝
[
- quadruple ???? (anyOP HCS08 LDX) MODE_IMM1 (Byte 〈xA,xE〉) 〈x0,x2〉
-; quadruple ???? (anyOP HCS08 LDX) MODE_DIR1 (Byte 〈xB,xE〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 LDX) MODE_DIR2 (Byte 〈xC,xE〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 LDX) MODE_IX2 (Byte 〈xD,xE〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 LDX) MODE_IX1 (Byte 〈xE,xE〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 LDX) MODE_IX0 (Byte 〈xF,xE〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 LDX) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,xE〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 LDX) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,xE〉〉) 〈x0,x4〉
+ quadruple … (anyOP HCS08 LDX) MODE_IMM1 (Byte 〈xA,xE〉) 〈x0,x2〉
+; quadruple … (anyOP HCS08 LDX) MODE_DIR1 (Byte 〈xB,xE〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 LDX) MODE_DIR2 (Byte 〈xC,xE〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 LDX) MODE_IX2 (Byte 〈xD,xE〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 LDX) MODE_IX1 (Byte 〈xE,xE〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 LDX) MODE_IX0 (Byte 〈xF,xE〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 LDX) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,xE〉〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 LDX) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,xE〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HCS08_25 ≝
[
- quadruple ???? (anyOP HCS08 LSR) MODE_DIR1 (Byte 〈x3,x4〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 LSR) MODE_INHA (Byte 〈x4,x4〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 LSR) MODE_INHX (Byte 〈x5,x4〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 LSR) MODE_IX1 (Byte 〈x6,x4〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 LSR) MODE_IX0 (Byte 〈x7,x4〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 LSR) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x4〉〉) 〈x0,x6〉
+ quadruple … (anyOP HCS08 LSR) MODE_DIR1 (Byte 〈x3,x4〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 LSR) MODE_INHA (Byte 〈x4,x4〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 LSR) MODE_INHX (Byte 〈x5,x4〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 LSR) MODE_IX1 (Byte 〈x6,x4〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 LSR) MODE_IX0 (Byte 〈x7,x4〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 LSR) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x4〉〉) 〈x0,x6〉
].
ndefinition opcode_table_HCS08_26 ≝
[
- quadruple ???? (anyOP HCS08 MOV) MODE_DIR1_to_DIR1 (Byte 〈x4,xE〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 MOV) MODE_DIR1_to_IX0p (Byte 〈x5,xE〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 MOV) MODE_IMM1_to_DIR1 (Byte 〈x6,xE〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 MOV) MODE_IX0p_to_DIR1 (Byte 〈x7,xE〉) 〈x0,x5〉
+ quadruple … (anyOP HCS08 MOV) MODE_DIR1_to_DIR1 (Byte 〈x4,xE〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 MOV) MODE_DIR1_to_IX0p (Byte 〈x5,xE〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 MOV) MODE_IMM1_to_DIR1 (Byte 〈x6,xE〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 MOV) MODE_IX0p_to_DIR1 (Byte 〈x7,xE〉) 〈x0,x5〉
].
ndefinition opcode_table_HCS08_27 ≝
[
- quadruple ???? (anyOP HCS08 NEG) MODE_DIR1 (Byte 〈x3,x0〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 NEG) MODE_INHA (Byte 〈x4,x0〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 NEG) MODE_INHX (Byte 〈x5,x0〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 NEG) MODE_IX1 (Byte 〈x6,x0〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 NEG) MODE_IX0 (Byte 〈x7,x0〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 NEG) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x0〉〉) 〈x0,x6〉
+ quadruple … (anyOP HCS08 NEG) MODE_DIR1 (Byte 〈x3,x0〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 NEG) MODE_INHA (Byte 〈x4,x0〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 NEG) MODE_INHX (Byte 〈x5,x0〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 NEG) MODE_IX1 (Byte 〈x6,x0〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 NEG) MODE_IX0 (Byte 〈x7,x0〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 NEG) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x0〉〉) 〈x0,x6〉
].
ndefinition opcode_table_HCS08_28 ≝
[
- quadruple ???? (anyOP HCS08 ORA) MODE_IMM1 (Byte 〈xA,xA〉) 〈x0,x2〉
-; quadruple ???? (anyOP HCS08 ORA) MODE_DIR1 (Byte 〈xB,xA〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 ORA) MODE_DIR2 (Byte 〈xC,xA〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 ORA) MODE_IX2 (Byte 〈xD,xA〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 ORA) MODE_IX1 (Byte 〈xE,xA〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 ORA) MODE_IX0 (Byte 〈xF,xA〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 ORA) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,xA〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 ORA) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,xA〉〉) 〈x0,x4〉
+ quadruple … (anyOP HCS08 ORA) MODE_IMM1 (Byte 〈xA,xA〉) 〈x0,x2〉
+; quadruple … (anyOP HCS08 ORA) MODE_DIR1 (Byte 〈xB,xA〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 ORA) MODE_DIR2 (Byte 〈xC,xA〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 ORA) MODE_IX2 (Byte 〈xD,xA〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 ORA) MODE_IX1 (Byte 〈xE,xA〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 ORA) MODE_IX0 (Byte 〈xF,xA〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 ORA) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,xA〉〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 ORA) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,xA〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HCS08_29 ≝
[
- quadruple ???? (anyOP HCS08 ROL) MODE_DIR1 (Byte 〈x3,x9〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 ROL) MODE_INHA (Byte 〈x4,x9〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 ROL) MODE_INHX (Byte 〈x5,x9〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 ROL) MODE_IX1 (Byte 〈x6,x9〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 ROL) MODE_IX0 (Byte 〈x7,x9〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 ROL) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x9〉〉) 〈x0,x6〉
+ quadruple … (anyOP HCS08 ROL) MODE_DIR1 (Byte 〈x3,x9〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 ROL) MODE_INHA (Byte 〈x4,x9〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 ROL) MODE_INHX (Byte 〈x5,x9〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 ROL) MODE_IX1 (Byte 〈x6,x9〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 ROL) MODE_IX0 (Byte 〈x7,x9〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 ROL) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x9〉〉) 〈x0,x6〉
].
ndefinition opcode_table_HCS08_30 ≝
[
- quadruple ???? (anyOP HCS08 ROR) MODE_DIR1 (Byte 〈x3,x6〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 ROR) MODE_INHA (Byte 〈x4,x6〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 ROR) MODE_INHX (Byte 〈x5,x6〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 ROR) MODE_IX1 (Byte 〈x6,x6〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 ROR) MODE_IX0 (Byte 〈x7,x6〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 ROR) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x6〉〉) 〈x0,x6〉
+ quadruple … (anyOP HCS08 ROR) MODE_DIR1 (Byte 〈x3,x6〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 ROR) MODE_INHA (Byte 〈x4,x6〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 ROR) MODE_INHX (Byte 〈x5,x6〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 ROR) MODE_IX1 (Byte 〈x6,x6〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 ROR) MODE_IX0 (Byte 〈x7,x6〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 ROR) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,x6〉〉) 〈x0,x6〉
].
ndefinition opcode_table_HCS08_31 ≝
[
- quadruple ???? (anyOP HCS08 SBC) MODE_IMM1 (Byte 〈xA,x2〉) 〈x0,x2〉
-; quadruple ???? (anyOP HCS08 SBC) MODE_DIR1 (Byte 〈xB,x2〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 SBC) MODE_DIR2 (Byte 〈xC,x2〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 SBC) MODE_IX2 (Byte 〈xD,x2〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 SBC) MODE_IX1 (Byte 〈xE,x2〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 SBC) MODE_IX0 (Byte 〈xF,x2〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 SBC) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x2〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 SBC) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x2〉〉) 〈x0,x4〉
+ quadruple … (anyOP HCS08 SBC) MODE_IMM1 (Byte 〈xA,x2〉) 〈x0,x2〉
+; quadruple … (anyOP HCS08 SBC) MODE_DIR1 (Byte 〈xB,x2〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 SBC) MODE_DIR2 (Byte 〈xC,x2〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 SBC) MODE_IX2 (Byte 〈xD,x2〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 SBC) MODE_IX1 (Byte 〈xE,x2〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 SBC) MODE_IX0 (Byte 〈xF,x2〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 SBC) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x2〉〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 SBC) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x2〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HCS08_32 ≝
[
- quadruple ???? (anyOP HCS08 STA) MODE_DIR1 (Byte 〈xB,x7〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 STA) MODE_DIR2 (Byte 〈xC,x7〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 STA) MODE_IX2 (Byte 〈xD,x7〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 STA) MODE_IX1 (Byte 〈xE,x7〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 STA) MODE_IX0 (Byte 〈xF,x7〉) 〈x0,x2〉
-; quadruple ???? (anyOP HCS08 STA) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x7〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 STA) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x7〉〉) 〈x0,x4〉
+ quadruple … (anyOP HCS08 STA) MODE_DIR1 (Byte 〈xB,x7〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 STA) MODE_DIR2 (Byte 〈xC,x7〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 STA) MODE_IX2 (Byte 〈xD,x7〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 STA) MODE_IX1 (Byte 〈xE,x7〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 STA) MODE_IX0 (Byte 〈xF,x7〉) 〈x0,x2〉
+; quadruple … (anyOP HCS08 STA) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x7〉〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 STA) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x7〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HCS08_33 ≝
[
- quadruple ???? (anyOP HCS08 STX) MODE_DIR1 (Byte 〈xB,xF〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 STX) MODE_DIR2 (Byte 〈xC,xF〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 STX) MODE_IX2 (Byte 〈xD,xF〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 STX) MODE_IX1 (Byte 〈xE,xF〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 STX) MODE_IX0 (Byte 〈xF,xF〉) 〈x0,x2〉
-; quadruple ???? (anyOP HCS08 STX) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,xF〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 STX) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,xF〉〉) 〈x0,x4〉
+ quadruple … (anyOP HCS08 STX) MODE_DIR1 (Byte 〈xB,xF〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 STX) MODE_DIR2 (Byte 〈xC,xF〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 STX) MODE_IX2 (Byte 〈xD,xF〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 STX) MODE_IX1 (Byte 〈xE,xF〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 STX) MODE_IX0 (Byte 〈xF,xF〉) 〈x0,x2〉
+; quadruple … (anyOP HCS08 STX) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,xF〉〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 STX) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,xF〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HCS08_34 ≝
[
- quadruple ???? (anyOP HCS08 SUB) MODE_IMM1 (Byte 〈xA,x0〉) 〈x0,x2〉
-; quadruple ???? (anyOP HCS08 SUB) MODE_DIR1 (Byte 〈xB,x0〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 SUB) MODE_DIR2 (Byte 〈xC,x0〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 SUB) MODE_IX2 (Byte 〈xD,x0〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 SUB) MODE_IX1 (Byte 〈xE,x0〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 SUB) MODE_IX0 (Byte 〈xF,x0〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 SUB) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x0〉〉) 〈x0,x5〉
-; quadruple ???? (anyOP HCS08 SUB) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x0〉〉) 〈x0,x4〉
+ quadruple … (anyOP HCS08 SUB) MODE_IMM1 (Byte 〈xA,x0〉) 〈x0,x2〉
+; quadruple … (anyOP HCS08 SUB) MODE_DIR1 (Byte 〈xB,x0〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 SUB) MODE_DIR2 (Byte 〈xC,x0〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 SUB) MODE_IX2 (Byte 〈xD,x0〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 SUB) MODE_IX1 (Byte 〈xE,x0〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 SUB) MODE_IX0 (Byte 〈xF,x0〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 SUB) MODE_SP2 (Word 〈〈x9,xE〉:〈xD,x0〉〉) 〈x0,x5〉
+; quadruple … (anyOP HCS08 SUB) MODE_SP1 (Word 〈〈x9,xE〉:〈xE,x0〉〉) 〈x0,x4〉
].
ndefinition opcode_table_HCS08_35 ≝
[
- quadruple ???? (anyOP HCS08 TST) MODE_DIR1 (Byte 〈x3,xD〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 TST) MODE_INHA (Byte 〈x4,xD〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 TST) MODE_INHX (Byte 〈x5,xD〉) 〈x0,x1〉
-; quadruple ???? (anyOP HCS08 TST) MODE_IX1 (Byte 〈x6,xD〉) 〈x0,x4〉
-; quadruple ???? (anyOP HCS08 TST) MODE_IX0 (Byte 〈x7,xD〉) 〈x0,x3〉
-; quadruple ???? (anyOP HCS08 TST) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,xD〉〉) 〈x0,x5〉
+ quadruple … (anyOP HCS08 TST) MODE_DIR1 (Byte 〈x3,xD〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 TST) MODE_INHA (Byte 〈x4,xD〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 TST) MODE_INHX (Byte 〈x5,xD〉) 〈x0,x1〉
+; quadruple … (anyOP HCS08 TST) MODE_IX1 (Byte 〈x6,xD〉) 〈x0,x4〉
+; quadruple … (anyOP HCS08 TST) MODE_IX0 (Byte 〈x7,xD〉) 〈x0,x3〉
+; quadruple … (anyOP HCS08 TST) MODE_SP1 (Word 〈〈x9,xE〉:〈x6,xD〉〉) 〈x0,x5〉
].
ndefinition opcode_table_HCS08 ≝
(test_not_impl_byte b HCS08_not_impl_byte ⊙ eq_nat (get_byte_count HCS08 b 0 opcode_table_HCS08) 1) ⊗
(⊖ (test_not_impl_byte b HCS08_not_impl_byte) ⊙ eq_nat (get_byte_count HCS08 b 0 opcode_table_HCS08) 0))
= true.
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
(* HCS08: opcode non implementati come da manuale (0x9E+byte) *)
(test_not_impl_byte b HCS08_not_impl_word ⊙ eq_nat (get_word_count HCS08 b 0 opcode_table_HCS08) 1) ⊗
(⊖ (test_not_impl_byte b HCS08_not_impl_word) ⊙ eq_nat (get_word_count HCS08 b 0 opcode_table_HCS08) 0))
= true.
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
(* HCS08: pseudocodici non implementati come da manuale *)
(test_not_impl_pseudo o HCS08_not_impl_pseudo ⊙ le_nat 1 (get_pseudo_count HCS08 o 0 opcode_table_HCS08)) ⊗
(⊖ (test_not_impl_pseudo o HCS08_not_impl_pseudo) ⊙ eq_nat (get_pseudo_count HCS08 o 0 opcode_table_HCS08) 0))
= true.
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
(* HCS08: modalita' non implementate come da manuale *)
(test_not_impl_mode i HCS08_not_impl_mode ⊙ le_nat 1 (get_mode_count HCS08 i 0 opcode_table_HCS08)) ⊗
(⊖ (test_not_impl_mode i HCS08_not_impl_mode) ⊙ eq_nat (get_mode_count HCS08 i 0 opcode_table_HCS08) 0))
= true.
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma ok_OpIm_table_HCS08 :
forall_instr_mode (λi:instr_mode.
forall_opcode (λop:opcode.
le_nat (get_OpIm_count HCS08 (anyOP HCS08 op) i 0 opcode_table_HCS08) 1)) = true.
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
ndefinition opcode_table_RS08_1 ≝
[
- quadruple ???? (anyOP RS08 ADC) MODE_IMM1 (Byte 〈xA,x9〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 ADC) MODE_DIR1 (Byte 〈xB,x9〉) 〈x0,x3〉
+ quadruple … (anyOP RS08 ADC) MODE_IMM1 (Byte 〈xA,x9〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 ADC) MODE_DIR1 (Byte 〈xB,x9〉) 〈x0,x3〉
].
ndefinition opcode_table_RS08_2 ≝
[
- quadruple ???? (anyOP RS08 ADD) MODE_IMM1 (Byte 〈xA,xB〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 ADD) MODE_DIR1 (Byte 〈xB,xB〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 ADD) (MODE_TNY x0) (Byte 〈x6,x0〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 ADD) (MODE_TNY x1) (Byte 〈x6,x1〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 ADD) (MODE_TNY x2) (Byte 〈x6,x2〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 ADD) (MODE_TNY x3) (Byte 〈x6,x3〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 ADD) (MODE_TNY x4) (Byte 〈x6,x4〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 ADD) (MODE_TNY x5) (Byte 〈x6,x5〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 ADD) (MODE_TNY x6) (Byte 〈x6,x6〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 ADD) (MODE_TNY x7) (Byte 〈x6,x7〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 ADD) (MODE_TNY x8) (Byte 〈x6,x8〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 ADD) (MODE_TNY x9) (Byte 〈x6,x9〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 ADD) (MODE_TNY xA) (Byte 〈x6,xA〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 ADD) (MODE_TNY xB) (Byte 〈x6,xB〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 ADD) (MODE_TNY xC) (Byte 〈x6,xC〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 ADD) (MODE_TNY xD) (Byte 〈x6,xD〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 ADD) (MODE_TNY xE) (Byte 〈x6,xE〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 ADD) (MODE_TNY xF) (Byte 〈x6,xF〉) 〈x0,x3〉
+ quadruple … (anyOP RS08 ADD) MODE_IMM1 (Byte 〈xA,xB〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 ADD) MODE_DIR1 (Byte 〈xB,xB〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 ADD) (MODE_TNY x0) (Byte 〈x6,x0〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 ADD) (MODE_TNY x1) (Byte 〈x6,x1〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 ADD) (MODE_TNY x2) (Byte 〈x6,x2〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 ADD) (MODE_TNY x3) (Byte 〈x6,x3〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 ADD) (MODE_TNY x4) (Byte 〈x6,x4〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 ADD) (MODE_TNY x5) (Byte 〈x6,x5〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 ADD) (MODE_TNY x6) (Byte 〈x6,x6〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 ADD) (MODE_TNY x7) (Byte 〈x6,x7〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 ADD) (MODE_TNY x8) (Byte 〈x6,x8〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 ADD) (MODE_TNY x9) (Byte 〈x6,x9〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 ADD) (MODE_TNY xA) (Byte 〈x6,xA〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 ADD) (MODE_TNY xB) (Byte 〈x6,xB〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 ADD) (MODE_TNY xC) (Byte 〈x6,xC〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 ADD) (MODE_TNY xD) (Byte 〈x6,xD〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 ADD) (MODE_TNY xE) (Byte 〈x6,xE〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 ADD) (MODE_TNY xF) (Byte 〈x6,xF〉) 〈x0,x3〉
].
ndefinition opcode_table_RS08_3 ≝
[
- quadruple ???? (anyOP RS08 AND) MODE_IMM1 (Byte 〈xA,x4〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 AND) MODE_DIR1 (Byte 〈xB,x4〉) 〈x0,x3〉
+ quadruple … (anyOP RS08 AND) MODE_IMM1 (Byte 〈xA,x4〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 AND) MODE_DIR1 (Byte 〈xB,x4〉) 〈x0,x3〉
].
ndefinition opcode_table_RS08_4 ≝
[
- quadruple ???? (anyOP RS08 ASL) MODE_INHA (Byte 〈x4,x8〉) 〈x0,x1〉
+ quadruple … (anyOP RS08 ASL) MODE_INHA (Byte 〈x4,x8〉) 〈x0,x1〉
].
ndefinition opcode_table_RS08_5 ≝
[
- quadruple ???? (anyOP RS08 BRA) MODE_IMM1 (Byte 〈x3,x0〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 BCC) MODE_IMM1 (Byte 〈x3,x4〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 BCS) MODE_IMM1 (Byte 〈x3,x5〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 BNE) MODE_IMM1 (Byte 〈x3,x6〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 BEQ) MODE_IMM1 (Byte 〈x3,x7〉) 〈x0,x3〉
+ quadruple … (anyOP RS08 BRA) MODE_IMM1 (Byte 〈x3,x0〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 BCC) MODE_IMM1 (Byte 〈x3,x4〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 BCS) MODE_IMM1 (Byte 〈x3,x5〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 BNE) MODE_IMM1 (Byte 〈x3,x6〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 BEQ) MODE_IMM1 (Byte 〈x3,x7〉) 〈x0,x3〉
].
ndefinition opcode_table_RS08_6 ≝
[
- quadruple ???? (anyOP RS08 BSETn) (MODE_DIRn o0) (Byte 〈x1,x0〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BCLRn) (MODE_DIRn o0) (Byte 〈x1,x1〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BSETn) (MODE_DIRn o1) (Byte 〈x1,x2〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BCLRn) (MODE_DIRn o1) (Byte 〈x1,x3〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BSETn) (MODE_DIRn o2) (Byte 〈x1,x4〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BCLRn) (MODE_DIRn o2) (Byte 〈x1,x5〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BSETn) (MODE_DIRn o3) (Byte 〈x1,x6〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BCLRn) (MODE_DIRn o3) (Byte 〈x1,x7〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BSETn) (MODE_DIRn o4) (Byte 〈x1,x8〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BCLRn) (MODE_DIRn o4) (Byte 〈x1,x9〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BSETn) (MODE_DIRn o5) (Byte 〈x1,xA〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BCLRn) (MODE_DIRn o5) (Byte 〈x1,xB〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BSETn) (MODE_DIRn o6) (Byte 〈x1,xC〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BCLRn) (MODE_DIRn o6) (Byte 〈x1,xD〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BSETn) (MODE_DIRn o7) (Byte 〈x1,xE〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BCLRn) (MODE_DIRn o7) (Byte 〈x1,xF〉) 〈x0,x5〉
+ quadruple … (anyOP RS08 BSETn) (MODE_DIRn o0) (Byte 〈x1,x0〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BCLRn) (MODE_DIRn o0) (Byte 〈x1,x1〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BSETn) (MODE_DIRn o1) (Byte 〈x1,x2〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BCLRn) (MODE_DIRn o1) (Byte 〈x1,x3〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BSETn) (MODE_DIRn o2) (Byte 〈x1,x4〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BCLRn) (MODE_DIRn o2) (Byte 〈x1,x5〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BSETn) (MODE_DIRn o3) (Byte 〈x1,x6〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BCLRn) (MODE_DIRn o3) (Byte 〈x1,x7〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BSETn) (MODE_DIRn o4) (Byte 〈x1,x8〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BCLRn) (MODE_DIRn o4) (Byte 〈x1,x9〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BSETn) (MODE_DIRn o5) (Byte 〈x1,xA〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BCLRn) (MODE_DIRn o5) (Byte 〈x1,xB〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BSETn) (MODE_DIRn o6) (Byte 〈x1,xC〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BCLRn) (MODE_DIRn o6) (Byte 〈x1,xD〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BSETn) (MODE_DIRn o7) (Byte 〈x1,xE〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BCLRn) (MODE_DIRn o7) (Byte 〈x1,xF〉) 〈x0,x5〉
].
ndefinition opcode_table_RS08_7 ≝
[
- quadruple ???? (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x0〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x1〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x2〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x3〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x4〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x5〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x6〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x7〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x8〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x9〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xA〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xB〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xC〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xD〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xE〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xF〉) 〈x0,x5〉
+ quadruple … (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x0〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x1〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x2〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x3〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x4〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x5〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x6〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x7〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x8〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x9〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xA〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xB〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xC〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xD〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xE〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xF〉) 〈x0,x5〉
].
ndefinition opcode_table_RS08_8 ≝
[
- quadruple ???? (anyOP RS08 CLC ) MODE_INH (Byte 〈x3,x8〉) 〈x0,x1〉
-; quadruple ???? (anyOP RS08 SEC ) MODE_INH (Byte 〈x3,x9〉) 〈x0,x1〉
-; quadruple ???? (anyOP RS08 SLA ) MODE_INH (Byte 〈x4,x2〉) 〈x0,x1〉
-; quadruple ???? (anyOP RS08 SHA ) MODE_INH (Byte 〈x4,x5〉) 〈x0,x1〉
-; quadruple ???? (anyOP RS08 NOP ) MODE_INH (Byte 〈xA,xC〉) 〈x0,x1〉
-; quadruple ???? (anyOP RS08 STOP) MODE_INH (Byte 〈xA,xE〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 WAIT) MODE_INH (Byte 〈xA,xF〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 RTS ) MODE_INH (Byte 〈xB,xE〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 BGND) MODE_INH (Byte 〈xB,xF〉) 〈x0,x5〉
+ quadruple … (anyOP RS08 CLC ) MODE_INH (Byte 〈x3,x8〉) 〈x0,x1〉
+; quadruple … (anyOP RS08 SEC ) MODE_INH (Byte 〈x3,x9〉) 〈x0,x1〉
+; quadruple … (anyOP RS08 SLA ) MODE_INH (Byte 〈x4,x2〉) 〈x0,x1〉
+; quadruple … (anyOP RS08 SHA ) MODE_INH (Byte 〈x4,x5〉) 〈x0,x1〉
+; quadruple … (anyOP RS08 NOP ) MODE_INH (Byte 〈xA,xC〉) 〈x0,x1〉
+; quadruple … (anyOP RS08 STOP) MODE_INH (Byte 〈xA,xE〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 WAIT) MODE_INH (Byte 〈xA,xF〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 RTS ) MODE_INH (Byte 〈xB,xE〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 BGND) MODE_INH (Byte 〈xB,xF〉) 〈x0,x5〉
].
ndefinition opcode_table_RS08_9 ≝
[
- quadruple ???? (anyOP RS08 CBEQA) MODE_DIR1_and_IMM1 (Byte 〈x3,x1〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 CBEQA) MODE_IMM1_and_IMM1 (Byte 〈x4,x1〉) 〈x0,x4〉
+ quadruple … (anyOP RS08 CBEQA) MODE_DIR1_and_IMM1 (Byte 〈x3,x1〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 CBEQA) MODE_IMM1_and_IMM1 (Byte 〈x4,x1〉) 〈x0,x4〉
].
ndefinition opcode_table_RS08_10 ≝
[
- quadruple ???? (anyOP RS08 CLR) MODE_DIR1 (Byte 〈x3,xF〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 CLR) MODE_INHA (Byte 〈x4,xF〉) 〈x0,x1〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t00) (Byte 〈x8,x0〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t01) (Byte 〈x8,x1〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t02) (Byte 〈x8,x2〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t03) (Byte 〈x8,x3〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t04) (Byte 〈x8,x4〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t05) (Byte 〈x8,x5〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t06) (Byte 〈x8,x6〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t07) (Byte 〈x8,x7〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t08) (Byte 〈x8,x8〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t09) (Byte 〈x8,x9〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t0A) (Byte 〈x8,xA〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t0B) (Byte 〈x8,xB〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t0C) (Byte 〈x8,xC〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t0D) (Byte 〈x8,xD〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t0E) (Byte 〈x8,xE〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t0F) (Byte 〈x8,xF〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t10) (Byte 〈x9,x0〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t11) (Byte 〈x9,x1〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t12) (Byte 〈x9,x2〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t13) (Byte 〈x9,x3〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t14) (Byte 〈x9,x4〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t15) (Byte 〈x9,x5〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t16) (Byte 〈x9,x6〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t17) (Byte 〈x9,x7〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t18) (Byte 〈x9,x8〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t19) (Byte 〈x9,x9〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t1A) (Byte 〈x9,xA〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t1B) (Byte 〈x9,xB〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t1C) (Byte 〈x9,xC〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t1D) (Byte 〈x9,xD〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t1E) (Byte 〈x9,xE〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CLR) (MODE_SRT t1F) (Byte 〈x9,xF〉) 〈x0,x2〉
+ quadruple … (anyOP RS08 CLR) MODE_DIR1 (Byte 〈x3,xF〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 CLR) MODE_INHA (Byte 〈x4,xF〉) 〈x0,x1〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t00) (Byte 〈x8,x0〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t01) (Byte 〈x8,x1〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t02) (Byte 〈x8,x2〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t03) (Byte 〈x8,x3〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t04) (Byte 〈x8,x4〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t05) (Byte 〈x8,x5〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t06) (Byte 〈x8,x6〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t07) (Byte 〈x8,x7〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t08) (Byte 〈x8,x8〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t09) (Byte 〈x8,x9〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t0A) (Byte 〈x8,xA〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t0B) (Byte 〈x8,xB〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t0C) (Byte 〈x8,xC〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t0D) (Byte 〈x8,xD〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t0E) (Byte 〈x8,xE〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t0F) (Byte 〈x8,xF〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t10) (Byte 〈x9,x0〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t11) (Byte 〈x9,x1〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t12) (Byte 〈x9,x2〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t13) (Byte 〈x9,x3〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t14) (Byte 〈x9,x4〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t15) (Byte 〈x9,x5〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t16) (Byte 〈x9,x6〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t17) (Byte 〈x9,x7〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t18) (Byte 〈x9,x8〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t19) (Byte 〈x9,x9〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t1A) (Byte 〈x9,xA〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t1B) (Byte 〈x9,xB〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t1C) (Byte 〈x9,xC〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t1D) (Byte 〈x9,xD〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t1E) (Byte 〈x9,xE〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CLR) (MODE_SRT t1F) (Byte 〈x9,xF〉) 〈x0,x2〉
].
ndefinition opcode_table_RS08_11 ≝
[
- quadruple ???? (anyOP RS08 CMP) MODE_IMM1 (Byte 〈xA,x1〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 CMP) MODE_DIR1 (Byte 〈xB,x1〉) 〈x0,x3〉
+ quadruple … (anyOP RS08 CMP) MODE_IMM1 (Byte 〈xA,x1〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 CMP) MODE_DIR1 (Byte 〈xB,x1〉) 〈x0,x3〉
].
ndefinition opcode_table_RS08_12 ≝
[
- quadruple ???? (anyOP RS08 COM) MODE_INHA (Byte 〈x4,x3〉) 〈x0,x1〉
+ quadruple … (anyOP RS08 COM) MODE_INHA (Byte 〈x4,x3〉) 〈x0,x1〉
].
ndefinition opcode_table_RS08_13 ≝
[
- quadruple ???? (anyOP RS08 DBNZ) MODE_DIR1_and_IMM1 (Byte 〈x3,xB〉) 〈x0,x7〉
-; quadruple ???? (anyOP RS08 DBNZ) MODE_INHA_and_IMM1 (Byte 〈x4,xB〉) 〈x0,x4〉
+ quadruple … (anyOP RS08 DBNZ) MODE_DIR1_and_IMM1 (Byte 〈x3,xB〉) 〈x0,x7〉
+; quadruple … (anyOP RS08 DBNZ) MODE_INHA_and_IMM1 (Byte 〈x4,xB〉) 〈x0,x4〉
].
ndefinition opcode_table_RS08_14 ≝
[
- quadruple ???? (anyOP RS08 DEC) MODE_DIR1 (Byte 〈x3,xA〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 DEC) MODE_INHA (Byte 〈x4,xA〉) 〈x0,x1〉
-; quadruple ???? (anyOP RS08 DEC) (MODE_TNY x0) (Byte 〈x5,x0〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 DEC) (MODE_TNY x1) (Byte 〈x5,x1〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 DEC) (MODE_TNY x2) (Byte 〈x5,x2〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 DEC) (MODE_TNY x3) (Byte 〈x5,x3〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 DEC) (MODE_TNY x4) (Byte 〈x5,x4〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 DEC) (MODE_TNY x5) (Byte 〈x5,x5〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 DEC) (MODE_TNY x6) (Byte 〈x5,x6〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 DEC) (MODE_TNY x7) (Byte 〈x5,x7〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 DEC) (MODE_TNY x8) (Byte 〈x5,x8〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 DEC) (MODE_TNY x9) (Byte 〈x5,x9〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 DEC) (MODE_TNY xA) (Byte 〈x5,xA〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 DEC) (MODE_TNY xB) (Byte 〈x5,xB〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 DEC) (MODE_TNY xC) (Byte 〈x5,xC〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 DEC) (MODE_TNY xD) (Byte 〈x5,xD〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 DEC) (MODE_TNY xE) (Byte 〈x5,xE〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 DEC) (MODE_TNY xF) (Byte 〈x5,xF〉) 〈x0,x4〉
+ quadruple … (anyOP RS08 DEC) MODE_DIR1 (Byte 〈x3,xA〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 DEC) MODE_INHA (Byte 〈x4,xA〉) 〈x0,x1〉
+; quadruple … (anyOP RS08 DEC) (MODE_TNY x0) (Byte 〈x5,x0〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 DEC) (MODE_TNY x1) (Byte 〈x5,x1〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 DEC) (MODE_TNY x2) (Byte 〈x5,x2〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 DEC) (MODE_TNY x3) (Byte 〈x5,x3〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 DEC) (MODE_TNY x4) (Byte 〈x5,x4〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 DEC) (MODE_TNY x5) (Byte 〈x5,x5〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 DEC) (MODE_TNY x6) (Byte 〈x5,x6〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 DEC) (MODE_TNY x7) (Byte 〈x5,x7〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 DEC) (MODE_TNY x8) (Byte 〈x5,x8〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 DEC) (MODE_TNY x9) (Byte 〈x5,x9〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 DEC) (MODE_TNY xA) (Byte 〈x5,xA〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 DEC) (MODE_TNY xB) (Byte 〈x5,xB〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 DEC) (MODE_TNY xC) (Byte 〈x5,xC〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 DEC) (MODE_TNY xD) (Byte 〈x5,xD〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 DEC) (MODE_TNY xE) (Byte 〈x5,xE〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 DEC) (MODE_TNY xF) (Byte 〈x5,xF〉) 〈x0,x4〉
].
ndefinition opcode_table_RS08_15 ≝
[
- quadruple ???? (anyOP RS08 EOR) MODE_IMM1 (Byte 〈xA,x8〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 EOR) MODE_DIR1 (Byte 〈xB,x8〉) 〈x0,x3〉
+ quadruple … (anyOP RS08 EOR) MODE_IMM1 (Byte 〈xA,x8〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 EOR) MODE_DIR1 (Byte 〈xB,x8〉) 〈x0,x3〉
].
ndefinition opcode_table_RS08_16 ≝
[
- quadruple ???? (anyOP RS08 INC) MODE_DIR1 (Byte 〈x3,xC〉) 〈x0,x5〉
-; quadruple ???? (anyOP RS08 INC) MODE_INHA (Byte 〈x4,xC〉) 〈x0,x1〉
-; quadruple ???? (anyOP RS08 INC) (MODE_TNY x0) (Byte 〈x2,x0〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 INC) (MODE_TNY x1) (Byte 〈x2,x1〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 INC) (MODE_TNY x2) (Byte 〈x2,x2〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 INC) (MODE_TNY x3) (Byte 〈x2,x3〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 INC) (MODE_TNY x4) (Byte 〈x2,x4〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 INC) (MODE_TNY x5) (Byte 〈x2,x5〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 INC) (MODE_TNY x6) (Byte 〈x2,x6〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 INC) (MODE_TNY x7) (Byte 〈x2,x7〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 INC) (MODE_TNY x8) (Byte 〈x2,x8〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 INC) (MODE_TNY x9) (Byte 〈x2,x9〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 INC) (MODE_TNY xA) (Byte 〈x2,xA〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 INC) (MODE_TNY xB) (Byte 〈x2,xB〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 INC) (MODE_TNY xC) (Byte 〈x2,xC〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 INC) (MODE_TNY xD) (Byte 〈x2,xD〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 INC) (MODE_TNY xE) (Byte 〈x2,xE〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 INC) (MODE_TNY xF) (Byte 〈x2,xF〉) 〈x0,x4〉
+ quadruple … (anyOP RS08 INC) MODE_DIR1 (Byte 〈x3,xC〉) 〈x0,x5〉
+; quadruple … (anyOP RS08 INC) MODE_INHA (Byte 〈x4,xC〉) 〈x0,x1〉
+; quadruple … (anyOP RS08 INC) (MODE_TNY x0) (Byte 〈x2,x0〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 INC) (MODE_TNY x1) (Byte 〈x2,x1〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 INC) (MODE_TNY x2) (Byte 〈x2,x2〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 INC) (MODE_TNY x3) (Byte 〈x2,x3〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 INC) (MODE_TNY x4) (Byte 〈x2,x4〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 INC) (MODE_TNY x5) (Byte 〈x2,x5〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 INC) (MODE_TNY x6) (Byte 〈x2,x6〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 INC) (MODE_TNY x7) (Byte 〈x2,x7〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 INC) (MODE_TNY x8) (Byte 〈x2,x8〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 INC) (MODE_TNY x9) (Byte 〈x2,x9〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 INC) (MODE_TNY xA) (Byte 〈x2,xA〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 INC) (MODE_TNY xB) (Byte 〈x2,xB〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 INC) (MODE_TNY xC) (Byte 〈x2,xC〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 INC) (MODE_TNY xD) (Byte 〈x2,xD〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 INC) (MODE_TNY xE) (Byte 〈x2,xE〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 INC) (MODE_TNY xF) (Byte 〈x2,xF〉) 〈x0,x4〉
].
ndefinition opcode_table_RS08_17 ≝
[
- quadruple ???? (anyOP RS08 JMP) MODE_IMM2 (Byte 〈xB,xC〉) 〈x0,x4〉
+ quadruple … (anyOP RS08 JMP) MODE_IMM2 (Byte 〈xB,xC〉) 〈x0,x4〉
].
ndefinition opcode_table_RS08_18 ≝
[
- quadruple ???? (anyOP RS08 BSR) MODE_IMM1 (Byte 〈xA,xD〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 JSR) MODE_IMM2 (Byte 〈xB,xD〉) 〈x0,x4〉
+ quadruple … (anyOP RS08 BSR) MODE_IMM1 (Byte 〈xA,xD〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 JSR) MODE_IMM2 (Byte 〈xB,xD〉) 〈x0,x4〉
].
ndefinition opcode_table_RS08_19 ≝
[
- quadruple ???? (anyOP RS08 LDA) MODE_IMM1 (Byte 〈xA,x6〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 LDA) MODE_DIR1 (Byte 〈xB,x6〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t00) (Byte 〈xC,x0〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t01) (Byte 〈xC,x1〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t02) (Byte 〈xC,x2〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t03) (Byte 〈xC,x3〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t04) (Byte 〈xC,x4〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t05) (Byte 〈xC,x5〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t06) (Byte 〈xC,x6〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t07) (Byte 〈xC,x7〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t08) (Byte 〈xC,x8〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t09) (Byte 〈xC,x9〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t0A) (Byte 〈xC,xA〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t0B) (Byte 〈xC,xB〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t0C) (Byte 〈xC,xC〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t0D) (Byte 〈xC,xD〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t0E) (Byte 〈xC,xE〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t0F) (Byte 〈xC,xF〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t10) (Byte 〈xD,x0〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t11) (Byte 〈xD,x1〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t12) (Byte 〈xD,x2〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t13) (Byte 〈xD,x3〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t14) (Byte 〈xD,x4〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t15) (Byte 〈xD,x5〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t16) (Byte 〈xD,x6〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t17) (Byte 〈xD,x7〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t18) (Byte 〈xD,x8〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t19) (Byte 〈xD,x9〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t1A) (Byte 〈xD,xA〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t1B) (Byte 〈xD,xB〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t1C) (Byte 〈xD,xC〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t1D) (Byte 〈xD,xD〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t1E) (Byte 〈xD,xE〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 LDA) (MODE_SRT t1F) (Byte 〈xD,xF〉) 〈x0,x3〉
+ quadruple … (anyOP RS08 LDA) MODE_IMM1 (Byte 〈xA,x6〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 LDA) MODE_DIR1 (Byte 〈xB,x6〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t00) (Byte 〈xC,x0〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t01) (Byte 〈xC,x1〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t02) (Byte 〈xC,x2〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t03) (Byte 〈xC,x3〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t04) (Byte 〈xC,x4〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t05) (Byte 〈xC,x5〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t06) (Byte 〈xC,x6〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t07) (Byte 〈xC,x7〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t08) (Byte 〈xC,x8〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t09) (Byte 〈xC,x9〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t0A) (Byte 〈xC,xA〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t0B) (Byte 〈xC,xB〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t0C) (Byte 〈xC,xC〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t0D) (Byte 〈xC,xD〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t0E) (Byte 〈xC,xE〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t0F) (Byte 〈xC,xF〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t10) (Byte 〈xD,x0〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t11) (Byte 〈xD,x1〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t12) (Byte 〈xD,x2〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t13) (Byte 〈xD,x3〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t14) (Byte 〈xD,x4〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t15) (Byte 〈xD,x5〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t16) (Byte 〈xD,x6〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t17) (Byte 〈xD,x7〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t18) (Byte 〈xD,x8〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t19) (Byte 〈xD,x9〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t1A) (Byte 〈xD,xA〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t1B) (Byte 〈xD,xB〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t1C) (Byte 〈xD,xC〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t1D) (Byte 〈xD,xD〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t1E) (Byte 〈xD,xE〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 LDA) (MODE_SRT t1F) (Byte 〈xD,xF〉) 〈x0,x3〉
].
ndefinition opcode_table_RS08_20 ≝
[
- quadruple ???? (anyOP RS08 LSR) MODE_INHA (Byte 〈x4,x4〉) 〈x0,x1〉
+ quadruple … (anyOP RS08 LSR) MODE_INHA (Byte 〈x4,x4〉) 〈x0,x1〉
].
ndefinition opcode_table_RS08_21 ≝
[
- quadruple ???? (anyOP RS08 MOV) MODE_IMM1_to_DIR1 (Byte 〈x3,xE〉) 〈x0,x4〉
-; quadruple ???? (anyOP RS08 MOV) MODE_DIR1_to_DIR1 (Byte 〈x4,xE〉) 〈x0,x5〉
+ quadruple … (anyOP RS08 MOV) MODE_IMM1_to_DIR1 (Byte 〈x3,xE〉) 〈x0,x4〉
+; quadruple … (anyOP RS08 MOV) MODE_DIR1_to_DIR1 (Byte 〈x4,xE〉) 〈x0,x5〉
].
ndefinition opcode_table_RS08_22 ≝
[
- quadruple ???? (anyOP RS08 ORA) MODE_IMM1 (Byte 〈xA,xA〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 ORA) MODE_DIR1 (Byte 〈xB,xA〉) 〈x0,x3〉
+ quadruple … (anyOP RS08 ORA) MODE_IMM1 (Byte 〈xA,xA〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 ORA) MODE_DIR1 (Byte 〈xB,xA〉) 〈x0,x3〉
].
ndefinition opcode_table_RS08_23 ≝
[
- quadruple ???? (anyOP RS08 ROL) MODE_INHA (Byte 〈x4,x9〉) 〈x0,x1〉
+ quadruple … (anyOP RS08 ROL) MODE_INHA (Byte 〈x4,x9〉) 〈x0,x1〉
].
ndefinition opcode_table_RS08_24 ≝
[
- quadruple ???? (anyOP RS08 ROR) MODE_INHA (Byte 〈x4,x6〉) 〈x0,x1〉
+ quadruple … (anyOP RS08 ROR) MODE_INHA (Byte 〈x4,x6〉) 〈x0,x1〉
].
ndefinition opcode_table_RS08_25 ≝
[
- quadruple ???? (anyOP RS08 SBC) MODE_IMM1 (Byte 〈xA,x2〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 SBC) MODE_DIR1 (Byte 〈xB,x2〉) 〈x0,x3〉
+ quadruple … (anyOP RS08 SBC) MODE_IMM1 (Byte 〈xA,x2〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 SBC) MODE_DIR1 (Byte 〈xB,x2〉) 〈x0,x3〉
].
ndefinition opcode_table_RS08_26 ≝
[
- quadruple ???? (anyOP RS08 STA) MODE_DIR1 (Byte 〈xB,x7〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t00) (Byte 〈xE,x0〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t01) (Byte 〈xE,x1〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t02) (Byte 〈xE,x2〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t03) (Byte 〈xE,x3〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t04) (Byte 〈xE,x4〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t05) (Byte 〈xE,x5〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t06) (Byte 〈xE,x6〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t07) (Byte 〈xE,x7〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t08) (Byte 〈xE,x8〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t09) (Byte 〈xE,x9〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t0A) (Byte 〈xE,xA〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t0B) (Byte 〈xE,xB〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t0C) (Byte 〈xE,xC〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t0D) (Byte 〈xE,xD〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t0E) (Byte 〈xE,xE〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t0F) (Byte 〈xE,xF〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t10) (Byte 〈xF,x0〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t11) (Byte 〈xF,x1〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t12) (Byte 〈xF,x2〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t13) (Byte 〈xF,x3〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t14) (Byte 〈xF,x4〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t15) (Byte 〈xF,x5〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t16) (Byte 〈xF,x6〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t17) (Byte 〈xF,x7〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t18) (Byte 〈xF,x8〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t19) (Byte 〈xF,x9〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t1A) (Byte 〈xF,xA〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t1B) (Byte 〈xF,xB〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t1C) (Byte 〈xF,xC〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t1D) (Byte 〈xF,xD〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t1E) (Byte 〈xF,xE〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 STA) (MODE_SRT t1F) (Byte 〈xF,xF〉) 〈x0,x2〉
+ quadruple … (anyOP RS08 STA) MODE_DIR1 (Byte 〈xB,x7〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t00) (Byte 〈xE,x0〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t01) (Byte 〈xE,x1〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t02) (Byte 〈xE,x2〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t03) (Byte 〈xE,x3〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t04) (Byte 〈xE,x4〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t05) (Byte 〈xE,x5〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t06) (Byte 〈xE,x6〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t07) (Byte 〈xE,x7〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t08) (Byte 〈xE,x8〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t09) (Byte 〈xE,x9〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t0A) (Byte 〈xE,xA〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t0B) (Byte 〈xE,xB〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t0C) (Byte 〈xE,xC〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t0D) (Byte 〈xE,xD〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t0E) (Byte 〈xE,xE〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t0F) (Byte 〈xE,xF〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t10) (Byte 〈xF,x0〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t11) (Byte 〈xF,x1〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t12) (Byte 〈xF,x2〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t13) (Byte 〈xF,x3〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t14) (Byte 〈xF,x4〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t15) (Byte 〈xF,x5〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t16) (Byte 〈xF,x6〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t17) (Byte 〈xF,x7〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t18) (Byte 〈xF,x8〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t19) (Byte 〈xF,x9〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t1A) (Byte 〈xF,xA〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t1B) (Byte 〈xF,xB〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t1C) (Byte 〈xF,xC〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t1D) (Byte 〈xF,xD〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t1E) (Byte 〈xF,xE〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 STA) (MODE_SRT t1F) (Byte 〈xF,xF〉) 〈x0,x2〉
].
ndefinition opcode_table_RS08_27 ≝
[
- quadruple ???? (anyOP RS08 SUB) MODE_IMM1 (Byte 〈xA,x0〉) 〈x0,x2〉
-; quadruple ???? (anyOP RS08 SUB) MODE_DIR1 (Byte 〈xB,x0〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 SUB) (MODE_TNY x0) (Byte 〈x7,x0〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 SUB) (MODE_TNY x1) (Byte 〈x7,x1〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 SUB) (MODE_TNY x2) (Byte 〈x7,x2〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 SUB) (MODE_TNY x3) (Byte 〈x7,x3〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 SUB) (MODE_TNY x4) (Byte 〈x7,x4〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 SUB) (MODE_TNY x5) (Byte 〈x7,x5〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 SUB) (MODE_TNY x6) (Byte 〈x7,x6〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 SUB) (MODE_TNY x7) (Byte 〈x7,x7〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 SUB) (MODE_TNY x8) (Byte 〈x7,x8〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 SUB) (MODE_TNY x9) (Byte 〈x7,x9〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 SUB) (MODE_TNY xA) (Byte 〈x7,xA〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 SUB) (MODE_TNY xB) (Byte 〈x7,xB〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 SUB) (MODE_TNY xC) (Byte 〈x7,xC〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 SUB) (MODE_TNY xD) (Byte 〈x7,xD〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 SUB) (MODE_TNY xE) (Byte 〈x7,xE〉) 〈x0,x3〉
-; quadruple ???? (anyOP RS08 SUB) (MODE_TNY xF) (Byte 〈x7,xF〉) 〈x0,x3〉
+ quadruple … (anyOP RS08 SUB) MODE_IMM1 (Byte 〈xA,x0〉) 〈x0,x2〉
+; quadruple … (anyOP RS08 SUB) MODE_DIR1 (Byte 〈xB,x0〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 SUB) (MODE_TNY x0) (Byte 〈x7,x0〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 SUB) (MODE_TNY x1) (Byte 〈x7,x1〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 SUB) (MODE_TNY x2) (Byte 〈x7,x2〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 SUB) (MODE_TNY x3) (Byte 〈x7,x3〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 SUB) (MODE_TNY x4) (Byte 〈x7,x4〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 SUB) (MODE_TNY x5) (Byte 〈x7,x5〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 SUB) (MODE_TNY x6) (Byte 〈x7,x6〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 SUB) (MODE_TNY x7) (Byte 〈x7,x7〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 SUB) (MODE_TNY x8) (Byte 〈x7,x8〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 SUB) (MODE_TNY x9) (Byte 〈x7,x9〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 SUB) (MODE_TNY xA) (Byte 〈x7,xA〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 SUB) (MODE_TNY xB) (Byte 〈x7,xB〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 SUB) (MODE_TNY xC) (Byte 〈x7,xC〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 SUB) (MODE_TNY xD) (Byte 〈x7,xD〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 SUB) (MODE_TNY xE) (Byte 〈x7,xE〉) 〈x0,x3〉
+; quadruple … (anyOP RS08 SUB) (MODE_TNY xF) (Byte 〈x7,xF〉) 〈x0,x3〉
].
ndefinition opcode_table_RS08 ≝
(test_not_impl_byte b RS08_not_impl_byte ⊙ eq_nat (get_byte_count RS08 b 0 opcode_table_RS08) 1) ⊗
(⊖ (test_not_impl_byte b RS08_not_impl_byte) ⊙ eq_nat (get_byte_count RS08 b 0 opcode_table_RS08) 0))
= true.
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
(* RS08: pseudocodici non implementati come da manuale *)
(test_not_impl_pseudo o RS08_not_impl_pseudo ⊙ le_nat 1 (get_pseudo_count RS08 o 0 opcode_table_RS08)) ⊗
(⊖ (test_not_impl_pseudo o RS08_not_impl_pseudo) ⊙ eq_nat (get_pseudo_count RS08 o 0 opcode_table_RS08) 0))
= true.
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
(* RS08: modalita' non implementate come da manuale *)
(test_not_impl_mode i RS08_not_impl_mode ⊙ le_nat 1 (get_mode_count RS08 i 0 opcode_table_RS08)) ⊗
(⊖ (test_not_impl_mode i RS08_not_impl_mode) ⊙ eq_nat (get_mode_count RS08 i 0 opcode_table_RS08) 0))
= true.
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma ok_OpIm_table_RS08 :
forall_instr_mode (λi:instr_mode.
forall_opcode (λop:opcode.
le_nat (get_OpIm_count RS08 (anyOP RS08 op) i 0 opcode_table_RS08) 1)) = true.
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma proj1: ∀A,B:Prop.A ∧ B → A.
#A; #B; #H;
- napply (And_ind A B ?? H);
+ napply (And_ind A B … H);
#H1; #H2;
napply H1.
nqed.
nlemma proj2: ∀A,B:Prop.A ∧ B → B.
#A; #B; #H;
- napply (And_ind A B ?? H);
+ napply (And_ind A B … H);
#H1; #H2;
napply H2.
nqed.
nnormalize;
#x; #y; #H;
nrewrite < H;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma eq_elim_r: ∀A:Type.∀x:A.∀P:A → Prop.P x → ∀y:A.y=x → P y.
#A; #x; #P; #H; #y; #H1;
- nrewrite < (symmetric_eq ??? H1);
+ nrewrite < (symmetric_eq … H1);
napply H.
nqed.
nchange with (match cons T x2 y2 with [ nil ⇒ False | cons a _ ⇒ x1 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma list_destruct_2 : ∀T.∀x1,x2:T.∀y1,y2:list T.cons T x1 y1 = cons T x2 y2 → y1 = y2.
nchange with (match cons T x2 y2 with [ nil ⇒ False | cons _ b ⇒ y1 = b ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma list_destruct_cons_nil : ∀T.∀x:T.∀y:list T.cons T x y = nil T → False.
#T; #l;
nelim l;
nnormalize;
- ##[ ##1: napply (refl_eq ??)
+ ##[ ##1: napply refl_eq
##| ##2: #x; #y; #H;
nrewrite > H;
- napply (refl_eq ??)
+ napply refl_eq
##]
nqed.
#T; #x; #y; #z;
nelim x;
nnormalize;
- ##[ ##1: napply (refl_eq ??)
+ ##[ ##1: napply refl_eq
##| ##2: #a; #b; #H;
nrewrite > H;
- napply (refl_eq ??)
+ napply refl_eq
##]
nqed.
nlemma cons_append_commute : ∀T:Type.∀l1,l2:list T.∀a:T.a :: (l1 @ l2) = (a :: l1) @ l2.
#T; #l1; #l2; #a;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma append_cons_commute : ∀T:Type.∀a:T.∀l,l1:list T.l @ (a::l1) = (l@[a]) @ l1.
#T; #a; #l; #l1;
nrewrite > (associative_list T l [a] l1);
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
match param with
[ nil ⇒ None ?
| cons hd tl ⇒
- match thd4T ???? hd with
+ match thd4T … hd with
[ Byte b ⇒ match borw with
[ Byte borw' ⇒ match eq_b8 borw' b with
[ true ⇒ Some ? hd
nchange with (match TByte m b2 with [ TByte a ⇒ b1 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
(* introduzione di un tipo dipendente (dalla modalita') per gli argomenti *)
(param:list (Prod4T (any_opcode m) instr_mode byte8_or_word16 byte8)) on param ≝
match param with
[ nil ⇒ None ? | cons hd tl ⇒
- match (eq_anyop m o (fst4T ???? hd)) ⊗ (eq_instrmode i (snd4T ???? hd)) with
- [ true ⇒ match thd4T ???? hd with
+ match (eq_anyop m o (fst4T … hd)) ⊗ (eq_instrmode i (snd4T … hd)) with
+ [ true ⇒ match thd4T … hd with
[ Byte isab ⇒
Some ? ([ (TByte m isab) ]@(args_picker m i p))
| Word isaw ⇒
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'' ]].
+ [ pair wl' c'' ⇒ pair … (mk_word16 wh' wl') c'' ]].
(* operatore shift destro *)
ndefinition shr_w16 ≝
λw:word16.match rcr_b8 (w16h w) false with
[ pair wh' c' ⇒ match rcr_b8 (w16l w) c' with
- [ pair wl' c'' ⇒ pair ?? (mk_word16 wh' wl') c'' ]].
+ [ pair wl' c'' ⇒ pair … (mk_word16 wh' wl') c'' ]].
(* operatore rotazione destra *)
ndefinition ror_w16 ≝
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'' ]].
+ [ pair wh' c'' ⇒ pair … (mk_word16 wh' wl') c'' ]].
(* operatore shift sinistro *)
ndefinition shl_w16 ≝
λw:word16.match rcl_b8 (w16l w) false with
[ pair wl' c' ⇒ match rcl_b8 (w16h w) c' with
- [ pair wh' c'' ⇒ pair ?? (mk_word16 wh' wl') c'' ]].
+ [ pair wh' c'' ⇒ pair … (mk_word16 wh' wl') c'' ]].
(* operatore rotazione sinistra *)
ndefinition rol_w16 ≝
λ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' ]].
+ [ pair h c' ⇒ pair … 〈h:l〉 c' ]].
(* operatore somma con data+carry → data *)
ndefinition plus_w16_dc_d ≝
λ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' ]].
+ [ pair h c' ⇒ pair … 〈h:l〉 c' ]].
(* operatore somma con data → data *)
ndefinition plus_w16_d_d ≝
let w' ≝ plus_w16_d_d divd (compl_w16 divs) in
match c with
[ 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〉)) ]
+ [ 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〉)) ]
| S c' ⇒ match le_w16 divs divd with
[ true ⇒ div_b8_aux w' (ror_w16 divs) (ror_b8 molt) (or_b8 molt q) c'
| false ⇒ div_b8_aux divd (ror_w16 divs) (ror_b8 molt) q c' ]].
(*
la combinazione n/0 e' illegale, segnala solo overflow senza dare risultato
*)
- [ true ⇒ triple ??? 〈xF,xF〉 (w16l w) true
+ [ 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
+ [ 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 *)
nchange with (match mk_word16 x2 y2 with [ mk_word16 a _ ⇒ x1 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma word16_destruct_2 :
nchange with (match mk_word16 x2 y2 with [ mk_word16 _ b ⇒ y1 = b ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_eqw16 : symmetricT word16 bool eq_w16.
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 ??).
+ napply refl_eq.
nqed.
nlemma symmetric_andw16 : symmetricT word16 word16 and_w16.
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 ??).
+ napply refl_eq.
nqed.
nlemma associative_andw16 : ∀w1,w2,w3.(and_w16 (and_w16 w1 w2) w3) = (and_w16 w1 (and_w16 w2 w3)).
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 ??).
+ napply refl_eq.
nqed.
nlemma symmetric_orw16 : symmetricT word16 word16 or_w16.
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 ??).
+ napply refl_eq.
nqed.
nlemma associative_orw16 : ∀w1,w2,w3.(or_w16 (or_w16 w1 w2) w3) = (or_w16 w1 (or_w16 w2 w3)).
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 ??).
+ napply refl_eq.
nqed.
nlemma symmetric_xorw16 : symmetricT word16 word16 xor_w16.
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 ??).
+ napply refl_eq.
nqed.
nlemma associative_xorw16 : ∀w1,w2,w3.(xor_w16 (xor_w16 w1 w2) w3) = (xor_w16 w1 (xor_w16 w2 w3)).
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 ??).
+ 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.
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' ]] =
+ [ 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' ]]);
+ [ 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' ] =
+ [ pair h c' ⇒ pair … 〈h:b5〉 c' ] =
match plus_b8_dc_dc b3 b1 c1 with
- [ pair h c' ⇒ pair ?? 〈h:b5〉 c' ]);
+ [ pair h c' ⇒ pair … 〈h:b5〉 c' ]);
nrewrite > (symmetric_plusb8_dc_dc b1 b3 c1);
- napply (refl_eq ??).
+ 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.
#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 ??).
+ 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.
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 ??).
+ napply refl_eq.
nqed.
nlemma symmetric_plusw16_d_dc : ∀w1,w2.plus_w16_d_dc w1 w2 = plus_w16_d_dc w2 w1.
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' ]] =
+ [ 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' ]]);
+ [ 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' ] =
+ [ pair h c' ⇒ pair … 〈h:b5〉 c' ] =
match plus_b8_dc_dc b3 b1 c with
- [ pair h c' ⇒ pair ?? 〈h:b5〉 c' ]);
+ [ pair h c' ⇒ pair … 〈h:b5〉 c' ]);
nrewrite > (symmetric_plusb8_dc_dc b1 b3 c);
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_plusw16_d_d : ∀w1,w2.plus_w16_d_d w1 w2 = plus_w16_d_d w2 w1.
#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 ??).
+ napply refl_eq.
nqed.
nlemma symmetric_plusw16_d_c : ∀w1,w2.plus_w16_d_c w1 w2 = plus_w16_d_c w2 w1.
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 ??).
+ napply refl_eq.
nqed.
nlemma symmetric_mulb8 : symmetricT byte8 word16 mul_b8.
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 ??).
+ napply refl_eq.
nqed.
nlemma eqw16_to_eq : ∀w1,w2:word16.(eq_w16 w1 w2 = true) → (w1 = 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 ??).
+ 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.
nelim w2;
#b1; #b2; #b3; #b4;
#H;
- nrewrite < (word16_destruct_1 ???? H);
- nrewrite < (word16_destruct_2 ???? 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 ??));
+ nrewrite > (eq_to_eqb8 b3 b3 (refl_eq …));
+ nrewrite > (eq_to_eqb8 b4 b4 (refl_eq …));
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
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'' ]].
+ [ pair wl' c'' ⇒ pair … (mk_word32 wh' wl') c'' ]].
(* 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'' ]].
+ [ pair wl' c'' ⇒ pair … (mk_word32 wh' wl') c'' ]].
(* operatore rotazione destra *)
ndefinition ror_w32 ≝
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'' ]].
+ [ pair wh' c'' ⇒ pair … (mk_word32 wh' wl') c'' ]].
(* 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'' ]].
+ [ pair wh' c'' ⇒ pair … (mk_word32 wh' wl') c'' ]].
(* operatore rotazione sinistra *)
ndefinition rol_w32 ≝
λ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' ]].
+ [ pair h c' ⇒ pair … 〈h.l〉 c' ]].
(* operatore somma con data+carry → data *)
ndefinition plus_w32_dc_d ≝
λ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' ]].
+ [ pair h c' ⇒ pair … 〈h.l〉 c' ]].
(* operatore somma con data → data *)
ndefinition plus_w32_d_d ≝
let w' ≝ plus_w32_d_d divd (compl_w32 divs) in
match c with
[ 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〉〉)) ]
+ [ 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〉〉)) ]
| S c' ⇒ match le_w32 divs divd with
[ true ⇒ div_w16_aux w' (ror_w32 divs) (ror_w16 molt) (or_w16 molt q) c'
| false ⇒ div_w16_aux divd (ror_w32 divs) (ror_w16 molt) q c' ]].
(*
la combinazione n/0 e' illegale, segnala solo overflow senza dare risultato
*)
- [ true ⇒ triple ??? 〈〈xF,xF〉:〈xF,xF〉〉 (w32l w) true
+ [ 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
+ [ 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 *)
nchange with (match mk_word32 x2 y2 with [ mk_word32 a _ ⇒ x1 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma word32_destruct_2 :
nchange with (match mk_word32 x2 y2 with [ mk_word32 _ b ⇒ y1 = b ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_eqw32 : symmetricT word32 bool eq_w32.
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 ??).
+ napply refl_eq.
nqed.
nlemma symmetric_andw32 : symmetricT word32 word32 and_w32.
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 ??).
+ napply refl_eq.
nqed.
nlemma associative_andw32 : ∀dw1,dw2,dw3.(and_w32 (and_w32 dw1 dw2) dw3) = (and_w32 dw1 (and_w32 dw2 dw3)).
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 ??).
+ napply refl_eq.
nqed.
nlemma symmetric_orw32 : symmetricT word32 word32 or_w32.
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 ??).
+ napply refl_eq.
nqed.
nlemma associative_orw32 : ∀dw1,dw2,dw3.(or_w32 (or_w32 dw1 dw2) dw3) = (or_w32 dw1 (or_w32 dw2 dw3)).
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 ??).
+ napply refl_eq.
nqed.
nlemma symmetric_xorw32 : symmetricT word32 word32 xor_w32.
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 ??).
+ napply refl_eq.
nqed.
nlemma associative_xorw32 : ∀dw1,dw2,dw3.(xor_w32 (xor_w32 dw1 dw2) dw3) = (xor_w32 dw1 (xor_w32 dw2 dw3)).
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 ??).
+ 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.
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' ]] =
+ [ 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' ]]);
+ [ 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' ] =
+ [ pair h c' ⇒ pair … 〈h.w5〉 c' ] =
match plus_w16_dc_dc w3 w1 c1 with
- [ pair h c' ⇒ pair ?? 〈h.w5〉 c' ]);
+ [ pair h c' ⇒ pair … 〈h.w5〉 c' ]);
nrewrite > (symmetric_plusw16_dc_dc w1 w3 c1);
- napply (refl_eq ??).
+ 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.
#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 ??).
+ 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.
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 ??).
+ napply refl_eq.
nqed.
nlemma symmetric_plusw32_d_dc : ∀dw1,dw2.plus_w32_d_dc dw1 dw2 = plus_w32_d_dc dw2 dw1.
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' ]] =
+ [ 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' ]]);
+ [ 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' ] =
+ [ pair h c' ⇒ pair … 〈h.w5〉 c' ] =
match plus_w16_dc_dc w3 w1 c with
- [ pair h c' ⇒ pair ?? 〈h.w5〉 c' ]);
+ [ pair h c' ⇒ pair … 〈h.w5〉 c' ]);
nrewrite > (symmetric_plusw16_dc_dc w1 w3 c);
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_plusw32_d_d : ∀dw1,dw2.plus_w32_d_d dw1 dw2 = plus_w32_d_d dw2 dw1.
#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 ??).
+ napply refl_eq.
nqed.
nlemma symmetric_plusw32_d_c : ∀dw1,dw2.plus_w32_d_c dw1 dw2 = plus_w32_d_c dw2 dw1.
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 ??).
+ napply refl_eq.
nqed.
nlemma symmetric_mulw16 : symmetricT word16 word32 mul_w16.
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 ??).
+ napply refl_eq.
nqed.
nlemma eqw32_to_eq : ∀dw1,dw2:word32.(eq_w32 dw1 dw2 = true) → (dw1 = 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 ??).
+ 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.
nelim dw2;
#w1; #w2; #w3; #w4;
#H;
- nrewrite < (word32_destruct_1 ???? H);
- nrewrite < (word32_destruct_2 ???? 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 ??));
+ nrewrite > (eq_to_eqw16 w3 w3 (refl_eq …));
+ nrewrite > (eq_to_eqw16 w4 w4 (refl_eq …));
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
(* DEFINIZIONE ASCII MINIMALE *)
(* ************************** *)
-nlemma symmetric_eqascii1 : ∀a2.eq_ascii ch_0 a2 = eq_ascii a2 ch_0. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii2 : ∀a2.eq_ascii ch_1 a2 = eq_ascii a2 ch_1. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii3 : ∀a2.eq_ascii ch_2 a2 = eq_ascii a2 ch_2. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii4 : ∀a2.eq_ascii ch_3 a2 = eq_ascii a2 ch_3. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii5 : ∀a2.eq_ascii ch_4 a2 = eq_ascii a2 ch_4. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii6 : ∀a2.eq_ascii ch_5 a2 = eq_ascii a2 ch_5. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii7 : ∀a2.eq_ascii ch_6 a2 = eq_ascii a2 ch_6. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii8 : ∀a2.eq_ascii ch_7 a2 = eq_ascii a2 ch_7. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii9 : ∀a2.eq_ascii ch_8 a2 = eq_ascii a2 ch_8. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii10 : ∀a2.eq_ascii ch_9 a2 = eq_ascii a2 ch_9. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii11 : ∀a2.eq_ascii ch__ a2 = eq_ascii a2 ch__. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii12 : ∀a2.eq_ascii ch_A a2 = eq_ascii a2 ch_A. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii13 : ∀a2.eq_ascii ch_B a2 = eq_ascii a2 ch_B. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii14 : ∀a2.eq_ascii ch_C a2 = eq_ascii a2 ch_C. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii15 : ∀a2.eq_ascii ch_D a2 = eq_ascii a2 ch_D. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii16 : ∀a2.eq_ascii ch_E a2 = eq_ascii a2 ch_E. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii17 : ∀a2.eq_ascii ch_F a2 = eq_ascii a2 ch_F. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii18 : ∀a2.eq_ascii ch_G a2 = eq_ascii a2 ch_G. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii19 : ∀a2.eq_ascii ch_H a2 = eq_ascii a2 ch_H. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii20 : ∀a2.eq_ascii ch_I a2 = eq_ascii a2 ch_I. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii21 : ∀a2.eq_ascii ch_J a2 = eq_ascii a2 ch_J. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii22 : ∀a2.eq_ascii ch_K a2 = eq_ascii a2 ch_K. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii23 : ∀a2.eq_ascii ch_L a2 = eq_ascii a2 ch_L. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii24 : ∀a2.eq_ascii ch_M a2 = eq_ascii a2 ch_M. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii25 : ∀a2.eq_ascii ch_N a2 = eq_ascii a2 ch_N. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii26 : ∀a2.eq_ascii ch_O a2 = eq_ascii a2 ch_O. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii27 : ∀a2.eq_ascii ch_P a2 = eq_ascii a2 ch_P. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii28 : ∀a2.eq_ascii ch_Q a2 = eq_ascii a2 ch_Q. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii29 : ∀a2.eq_ascii ch_R a2 = eq_ascii a2 ch_R. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii30 : ∀a2.eq_ascii ch_S a2 = eq_ascii a2 ch_S. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii31 : ∀a2.eq_ascii ch_T a2 = eq_ascii a2 ch_T. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii32 : ∀a2.eq_ascii ch_U a2 = eq_ascii a2 ch_U. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii33 : ∀a2.eq_ascii ch_V a2 = eq_ascii a2 ch_V. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii34 : ∀a2.eq_ascii ch_W a2 = eq_ascii a2 ch_W. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii35 : ∀a2.eq_ascii ch_X a2 = eq_ascii a2 ch_X. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii36 : ∀a2.eq_ascii ch_Y a2 = eq_ascii a2 ch_Y. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii37 : ∀a2.eq_ascii ch_Z a2 = eq_ascii a2 ch_Z. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii38 : ∀a2.eq_ascii ch_a a2 = eq_ascii a2 ch_a. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii39 : ∀a2.eq_ascii ch_b a2 = eq_ascii a2 ch_b. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii40 : ∀a2.eq_ascii ch_c a2 = eq_ascii a2 ch_c. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii41 : ∀a2.eq_ascii ch_d a2 = eq_ascii a2 ch_d. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii42 : ∀a2.eq_ascii ch_e a2 = eq_ascii a2 ch_e. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii43 : ∀a2.eq_ascii ch_f a2 = eq_ascii a2 ch_f. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii44 : ∀a2.eq_ascii ch_g a2 = eq_ascii a2 ch_g. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii45 : ∀a2.eq_ascii ch_h a2 = eq_ascii a2 ch_h. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii46 : ∀a2.eq_ascii ch_i a2 = eq_ascii a2 ch_i. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii47 : ∀a2.eq_ascii ch_j a2 = eq_ascii a2 ch_j. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii48 : ∀a2.eq_ascii ch_k a2 = eq_ascii a2 ch_k. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii49 : ∀a2.eq_ascii ch_l a2 = eq_ascii a2 ch_l. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii50 : ∀a2.eq_ascii ch_m a2 = eq_ascii a2 ch_m. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii51 : ∀a2.eq_ascii ch_n a2 = eq_ascii a2 ch_n. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii52 : ∀a2.eq_ascii ch_o a2 = eq_ascii a2 ch_o. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii53 : ∀a2.eq_ascii ch_p a2 = eq_ascii a2 ch_p. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii54 : ∀a2.eq_ascii ch_q a2 = eq_ascii a2 ch_q. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii55 : ∀a2.eq_ascii ch_r a2 = eq_ascii a2 ch_r. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii56 : ∀a2.eq_ascii ch_s a2 = eq_ascii a2 ch_s. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii57 : ∀a2.eq_ascii ch_t a2 = eq_ascii a2 ch_t. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii58 : ∀a2.eq_ascii ch_u a2 = eq_ascii a2 ch_u. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii59 : ∀a2.eq_ascii ch_v a2 = eq_ascii a2 ch_v. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii60 : ∀a2.eq_ascii ch_w a2 = eq_ascii a2 ch_w. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii61 : ∀a2.eq_ascii ch_x a2 = eq_ascii a2 ch_x. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii62 : ∀a2.eq_ascii ch_y a2 = eq_ascii a2 ch_y. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
-nlemma symmetric_eqascii63 : ∀a2.eq_ascii ch_z a2 = eq_ascii a2 ch_z. #a2; ncases a2; nnormalize; napply (refl_eq ??).nqed.
+nlemma symmetric_eqascii1 : ∀a2.eq_ascii ch_0 a2 = eq_ascii a2 ch_0. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii2 : ∀a2.eq_ascii ch_1 a2 = eq_ascii a2 ch_1. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii3 : ∀a2.eq_ascii ch_2 a2 = eq_ascii a2 ch_2. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii4 : ∀a2.eq_ascii ch_3 a2 = eq_ascii a2 ch_3. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii5 : ∀a2.eq_ascii ch_4 a2 = eq_ascii a2 ch_4. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii6 : ∀a2.eq_ascii ch_5 a2 = eq_ascii a2 ch_5. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii7 : ∀a2.eq_ascii ch_6 a2 = eq_ascii a2 ch_6. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii8 : ∀a2.eq_ascii ch_7 a2 = eq_ascii a2 ch_7. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii9 : ∀a2.eq_ascii ch_8 a2 = eq_ascii a2 ch_8. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii10 : ∀a2.eq_ascii ch_9 a2 = eq_ascii a2 ch_9. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii11 : ∀a2.eq_ascii ch__ a2 = eq_ascii a2 ch__. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii12 : ∀a2.eq_ascii ch_A a2 = eq_ascii a2 ch_A. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii13 : ∀a2.eq_ascii ch_B a2 = eq_ascii a2 ch_B. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii14 : ∀a2.eq_ascii ch_C a2 = eq_ascii a2 ch_C. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii15 : ∀a2.eq_ascii ch_D a2 = eq_ascii a2 ch_D. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii16 : ∀a2.eq_ascii ch_E a2 = eq_ascii a2 ch_E. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii17 : ∀a2.eq_ascii ch_F a2 = eq_ascii a2 ch_F. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii18 : ∀a2.eq_ascii ch_G a2 = eq_ascii a2 ch_G. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii19 : ∀a2.eq_ascii ch_H a2 = eq_ascii a2 ch_H. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii20 : ∀a2.eq_ascii ch_I a2 = eq_ascii a2 ch_I. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii21 : ∀a2.eq_ascii ch_J a2 = eq_ascii a2 ch_J. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii22 : ∀a2.eq_ascii ch_K a2 = eq_ascii a2 ch_K. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii23 : ∀a2.eq_ascii ch_L a2 = eq_ascii a2 ch_L. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii24 : ∀a2.eq_ascii ch_M a2 = eq_ascii a2 ch_M. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii25 : ∀a2.eq_ascii ch_N a2 = eq_ascii a2 ch_N. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii26 : ∀a2.eq_ascii ch_O a2 = eq_ascii a2 ch_O. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii27 : ∀a2.eq_ascii ch_P a2 = eq_ascii a2 ch_P. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii28 : ∀a2.eq_ascii ch_Q a2 = eq_ascii a2 ch_Q. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii29 : ∀a2.eq_ascii ch_R a2 = eq_ascii a2 ch_R. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii30 : ∀a2.eq_ascii ch_S a2 = eq_ascii a2 ch_S. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii31 : ∀a2.eq_ascii ch_T a2 = eq_ascii a2 ch_T. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii32 : ∀a2.eq_ascii ch_U a2 = eq_ascii a2 ch_U. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii33 : ∀a2.eq_ascii ch_V a2 = eq_ascii a2 ch_V. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii34 : ∀a2.eq_ascii ch_W a2 = eq_ascii a2 ch_W. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii35 : ∀a2.eq_ascii ch_X a2 = eq_ascii a2 ch_X. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii36 : ∀a2.eq_ascii ch_Y a2 = eq_ascii a2 ch_Y. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii37 : ∀a2.eq_ascii ch_Z a2 = eq_ascii a2 ch_Z. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii38 : ∀a2.eq_ascii ch_a a2 = eq_ascii a2 ch_a. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii39 : ∀a2.eq_ascii ch_b a2 = eq_ascii a2 ch_b. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii40 : ∀a2.eq_ascii ch_c a2 = eq_ascii a2 ch_c. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii41 : ∀a2.eq_ascii ch_d a2 = eq_ascii a2 ch_d. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii42 : ∀a2.eq_ascii ch_e a2 = eq_ascii a2 ch_e. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii43 : ∀a2.eq_ascii ch_f a2 = eq_ascii a2 ch_f. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii44 : ∀a2.eq_ascii ch_g a2 = eq_ascii a2 ch_g. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii45 : ∀a2.eq_ascii ch_h a2 = eq_ascii a2 ch_h. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii46 : ∀a2.eq_ascii ch_i a2 = eq_ascii a2 ch_i. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii47 : ∀a2.eq_ascii ch_j a2 = eq_ascii a2 ch_j. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii48 : ∀a2.eq_ascii ch_k a2 = eq_ascii a2 ch_k. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii49 : ∀a2.eq_ascii ch_l a2 = eq_ascii a2 ch_l. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii50 : ∀a2.eq_ascii ch_m a2 = eq_ascii a2 ch_m. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii51 : ∀a2.eq_ascii ch_n a2 = eq_ascii a2 ch_n. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii52 : ∀a2.eq_ascii ch_o a2 = eq_ascii a2 ch_o. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii53 : ∀a2.eq_ascii ch_p a2 = eq_ascii a2 ch_p. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii54 : ∀a2.eq_ascii ch_q a2 = eq_ascii a2 ch_q. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii55 : ∀a2.eq_ascii ch_r a2 = eq_ascii a2 ch_r. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii56 : ∀a2.eq_ascii ch_s a2 = eq_ascii a2 ch_s. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii57 : ∀a2.eq_ascii ch_t a2 = eq_ascii a2 ch_t. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii58 : ∀a2.eq_ascii ch_u a2 = eq_ascii a2 ch_u. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii59 : ∀a2.eq_ascii ch_v a2 = eq_ascii a2 ch_v. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii60 : ∀a2.eq_ascii ch_w a2 = eq_ascii a2 ch_w. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii61 : ∀a2.eq_ascii ch_x a2 = eq_ascii a2 ch_x. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii62 : ∀a2.eq_ascii ch_y a2 = eq_ascii a2 ch_y. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
+nlemma symmetric_eqascii63 : ∀a2.eq_ascii ch_z a2 = eq_ascii a2 ch_z. #a2; ncases a2; nnormalize; napply refl_eq.nqed.
nlemma symmetric_eqascii : symmetricT ascii bool eq_ascii.
#a1; ncases a1;
##| ##63: napply symmetric_eqascii63 ##]
nqed.
-nlemma eqascii_to_eq1 : ∀a2.eq_ascii ch_0 a2 = true → ch_0 = a2. #a2; ncases a2; nnormalize; #H; ##[ ##1: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq2 : ∀a2.eq_ascii ch_1 a2 = true → ch_1 = a2. #a2; ncases a2; nnormalize; #H; ##[ ##2: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq3 : ∀a2.eq_ascii ch_2 a2 = true → ch_2 = a2. #a2; ncases a2; nnormalize; #H; ##[ ##3: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq4 : ∀a2.eq_ascii ch_3 a2 = true → ch_3 = a2. #a2; ncases a2; nnormalize; #H; ##[ ##4: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq5 : ∀a2.eq_ascii ch_4 a2 = true → ch_4 = a2. #a2; ncases a2; nnormalize; #H; ##[ ##5: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq6 : ∀a2.eq_ascii ch_5 a2 = true → ch_5 = a2. #a2; ncases a2; nnormalize; #H; ##[ ##6: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq7 : ∀a2.eq_ascii ch_6 a2 = true → ch_6 = a2. #a2; ncases a2; nnormalize; #H; ##[ ##7: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq8 : ∀a2.eq_ascii ch_7 a2 = true → ch_7 = a2. #a2; ncases a2; nnormalize; #H; ##[ ##8: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq9 : ∀a2.eq_ascii ch_8 a2 = true → ch_8 = a2. #a2; ncases a2; nnormalize; #H; ##[ ##9: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq10 : ∀a2.eq_ascii ch_9 a2 = true → ch_9 = a2. #a2; ncases a2; nnormalize; #H; ##[ ##10: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq11 : ∀a2.eq_ascii ch__ a2 = true → ch__ = a2. #a2; ncases a2; nnormalize; #H; ##[ ##11: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq12 : ∀a2.eq_ascii ch_A a2 = true → ch_A = a2. #a2; ncases a2; nnormalize; #H; ##[ ##12: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq13 : ∀a2.eq_ascii ch_B a2 = true → ch_B = a2. #a2; ncases a2; nnormalize; #H; ##[ ##13: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq14 : ∀a2.eq_ascii ch_C a2 = true → ch_C = a2. #a2; ncases a2; nnormalize; #H; ##[ ##14: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq15 : ∀a2.eq_ascii ch_D a2 = true → ch_D = a2. #a2; ncases a2; nnormalize; #H; ##[ ##15: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq16 : ∀a2.eq_ascii ch_E a2 = true → ch_E = a2. #a2; ncases a2; nnormalize; #H; ##[ ##16: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq17 : ∀a2.eq_ascii ch_F a2 = true → ch_F = a2. #a2; ncases a2; nnormalize; #H; ##[ ##17: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq18 : ∀a2.eq_ascii ch_G a2 = true → ch_G = a2. #a2; ncases a2; nnormalize; #H; ##[ ##18: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq19 : ∀a2.eq_ascii ch_H a2 = true → ch_H = a2. #a2; ncases a2; nnormalize; #H; ##[ ##19: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq20 : ∀a2.eq_ascii ch_I a2 = true → ch_I = a2. #a2; ncases a2; nnormalize; #H; ##[ ##20: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq21 : ∀a2.eq_ascii ch_J a2 = true → ch_J = a2. #a2; ncases a2; nnormalize; #H; ##[ ##21: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq22 : ∀a2.eq_ascii ch_K a2 = true → ch_K = a2. #a2; ncases a2; nnormalize; #H; ##[ ##22: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq23 : ∀a2.eq_ascii ch_L a2 = true → ch_L = a2. #a2; ncases a2; nnormalize; #H; ##[ ##23: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq24 : ∀a2.eq_ascii ch_M a2 = true → ch_M = a2. #a2; ncases a2; nnormalize; #H; ##[ ##24: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq25 : ∀a2.eq_ascii ch_N a2 = true → ch_N = a2. #a2; ncases a2; nnormalize; #H; ##[ ##25: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq26 : ∀a2.eq_ascii ch_O a2 = true → ch_O = a2. #a2; ncases a2; nnormalize; #H; ##[ ##26: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq27 : ∀a2.eq_ascii ch_P a2 = true → ch_P = a2. #a2; ncases a2; nnormalize; #H; ##[ ##27: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq28 : ∀a2.eq_ascii ch_Q a2 = true → ch_Q = a2. #a2; ncases a2; nnormalize; #H; ##[ ##28: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq29 : ∀a2.eq_ascii ch_R a2 = true → ch_R = a2. #a2; ncases a2; nnormalize; #H; ##[ ##29: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq30 : ∀a2.eq_ascii ch_S a2 = true → ch_S = a2. #a2; ncases a2; nnormalize; #H; ##[ ##30: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq31 : ∀a2.eq_ascii ch_T a2 = true → ch_T = a2. #a2; ncases a2; nnormalize; #H; ##[ ##31: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq32 : ∀a2.eq_ascii ch_U a2 = true → ch_U = a2. #a2; ncases a2; nnormalize; #H; ##[ ##32: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq33 : ∀a2.eq_ascii ch_V a2 = true → ch_V = a2. #a2; ncases a2; nnormalize; #H; ##[ ##33: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq34 : ∀a2.eq_ascii ch_W a2 = true → ch_W = a2. #a2; ncases a2; nnormalize; #H; ##[ ##34: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq35 : ∀a2.eq_ascii ch_X a2 = true → ch_X = a2. #a2; ncases a2; nnormalize; #H; ##[ ##35: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq36 : ∀a2.eq_ascii ch_Y a2 = true → ch_Y = a2. #a2; ncases a2; nnormalize; #H; ##[ ##36: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq37 : ∀a2.eq_ascii ch_Z a2 = true → ch_Z = a2. #a2; ncases a2; nnormalize; #H; ##[ ##37: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq38 : ∀a2.eq_ascii ch_a a2 = true → ch_a = a2. #a2; ncases a2; nnormalize; #H; ##[ ##38: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq39 : ∀a2.eq_ascii ch_b a2 = true → ch_b = a2. #a2; ncases a2; nnormalize; #H; ##[ ##39: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq40 : ∀a2.eq_ascii ch_c a2 = true → ch_c = a2. #a2; ncases a2; nnormalize; #H; ##[ ##40: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq41 : ∀a2.eq_ascii ch_d a2 = true → ch_d = a2. #a2; ncases a2; nnormalize; #H; ##[ ##41: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq42 : ∀a2.eq_ascii ch_e a2 = true → ch_e = a2. #a2; ncases a2; nnormalize; #H; ##[ ##42: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq43 : ∀a2.eq_ascii ch_f a2 = true → ch_f = a2. #a2; ncases a2; nnormalize; #H; ##[ ##43: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq44 : ∀a2.eq_ascii ch_g a2 = true → ch_g = a2. #a2; ncases a2; nnormalize; #H; ##[ ##44: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq45 : ∀a2.eq_ascii ch_h a2 = true → ch_h = a2. #a2; ncases a2; nnormalize; #H; ##[ ##45: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq46 : ∀a2.eq_ascii ch_i a2 = true → ch_i = a2. #a2; ncases a2; nnormalize; #H; ##[ ##46: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq47 : ∀a2.eq_ascii ch_j a2 = true → ch_j = a2. #a2; ncases a2; nnormalize; #H; ##[ ##47: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq48 : ∀a2.eq_ascii ch_k a2 = true → ch_k = a2. #a2; ncases a2; nnormalize; #H; ##[ ##48: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq49 : ∀a2.eq_ascii ch_l a2 = true → ch_l = a2. #a2; ncases a2; nnormalize; #H; ##[ ##49: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq50 : ∀a2.eq_ascii ch_m a2 = true → ch_m = a2. #a2; ncases a2; nnormalize; #H; ##[ ##50: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq51 : ∀a2.eq_ascii ch_n a2 = true → ch_n = a2. #a2; ncases a2; nnormalize; #H; ##[ ##51: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq52 : ∀a2.eq_ascii ch_o a2 = true → ch_o = a2. #a2; ncases a2; nnormalize; #H; ##[ ##52: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq53 : ∀a2.eq_ascii ch_p a2 = true → ch_p = a2. #a2; ncases a2; nnormalize; #H; ##[ ##53: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq54 : ∀a2.eq_ascii ch_q a2 = true → ch_q = a2. #a2; ncases a2; nnormalize; #H; ##[ ##54: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq55 : ∀a2.eq_ascii ch_r a2 = true → ch_r = a2. #a2; ncases a2; nnormalize; #H; ##[ ##55: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq56 : ∀a2.eq_ascii ch_s a2 = true → ch_s = a2. #a2; ncases a2; nnormalize; #H; ##[ ##56: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq57 : ∀a2.eq_ascii ch_t a2 = true → ch_t = a2. #a2; ncases a2; nnormalize; #H; ##[ ##57: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq58 : ∀a2.eq_ascii ch_u a2 = true → ch_u = a2. #a2; ncases a2; nnormalize; #H; ##[ ##58: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq59 : ∀a2.eq_ascii ch_v a2 = true → ch_v = a2. #a2; ncases a2; nnormalize; #H; ##[ ##59: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq60 : ∀a2.eq_ascii ch_w a2 = true → ch_w = a2. #a2; ncases a2; nnormalize; #H; ##[ ##60: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq61 : ∀a2.eq_ascii ch_x a2 = true → ch_x = a2. #a2; ncases a2; nnormalize; #H; ##[ ##61: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq62 : ∀a2.eq_ascii ch_y a2 = true → ch_y = a2. #a2; ncases a2; nnormalize; #H; ##[ ##62: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
-nlemma eqascii_to_eq63 : ∀a2.eq_ascii ch_z a2 = true → ch_z = a2. #a2; ncases a2; nnormalize; #H; ##[ ##63: napply (refl_eq ??) | ##*: napply (bool_destruct ??? H) ##] nqed.
+nlemma eqascii_to_eq1 : ∀a2.eq_ascii ch_0 a2 = true → ch_0 = a2. #a2; ncases a2; nnormalize; #H; ##[ ##1: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq2 : ∀a2.eq_ascii ch_1 a2 = true → ch_1 = a2. #a2; ncases a2; nnormalize; #H; ##[ ##2: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq3 : ∀a2.eq_ascii ch_2 a2 = true → ch_2 = a2. #a2; ncases a2; nnormalize; #H; ##[ ##3: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq4 : ∀a2.eq_ascii ch_3 a2 = true → ch_3 = a2. #a2; ncases a2; nnormalize; #H; ##[ ##4: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq5 : ∀a2.eq_ascii ch_4 a2 = true → ch_4 = a2. #a2; ncases a2; nnormalize; #H; ##[ ##5: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq6 : ∀a2.eq_ascii ch_5 a2 = true → ch_5 = a2. #a2; ncases a2; nnormalize; #H; ##[ ##6: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq7 : ∀a2.eq_ascii ch_6 a2 = true → ch_6 = a2. #a2; ncases a2; nnormalize; #H; ##[ ##7: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq8 : ∀a2.eq_ascii ch_7 a2 = true → ch_7 = a2. #a2; ncases a2; nnormalize; #H; ##[ ##8: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq9 : ∀a2.eq_ascii ch_8 a2 = true → ch_8 = a2. #a2; ncases a2; nnormalize; #H; ##[ ##9: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq10 : ∀a2.eq_ascii ch_9 a2 = true → ch_9 = a2. #a2; ncases a2; nnormalize; #H; ##[ ##10: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq11 : ∀a2.eq_ascii ch__ a2 = true → ch__ = a2. #a2; ncases a2; nnormalize; #H; ##[ ##11: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq12 : ∀a2.eq_ascii ch_A a2 = true → ch_A = a2. #a2; ncases a2; nnormalize; #H; ##[ ##12: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq13 : ∀a2.eq_ascii ch_B a2 = true → ch_B = a2. #a2; ncases a2; nnormalize; #H; ##[ ##13: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq14 : ∀a2.eq_ascii ch_C a2 = true → ch_C = a2. #a2; ncases a2; nnormalize; #H; ##[ ##14: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq15 : ∀a2.eq_ascii ch_D a2 = true → ch_D = a2. #a2; ncases a2; nnormalize; #H; ##[ ##15: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq16 : ∀a2.eq_ascii ch_E a2 = true → ch_E = a2. #a2; ncases a2; nnormalize; #H; ##[ ##16: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq17 : ∀a2.eq_ascii ch_F a2 = true → ch_F = a2. #a2; ncases a2; nnormalize; #H; ##[ ##17: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq18 : ∀a2.eq_ascii ch_G a2 = true → ch_G = a2. #a2; ncases a2; nnormalize; #H; ##[ ##18: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq19 : ∀a2.eq_ascii ch_H a2 = true → ch_H = a2. #a2; ncases a2; nnormalize; #H; ##[ ##19: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq20 : ∀a2.eq_ascii ch_I a2 = true → ch_I = a2. #a2; ncases a2; nnormalize; #H; ##[ ##20: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq21 : ∀a2.eq_ascii ch_J a2 = true → ch_J = a2. #a2; ncases a2; nnormalize; #H; ##[ ##21: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq22 : ∀a2.eq_ascii ch_K a2 = true → ch_K = a2. #a2; ncases a2; nnormalize; #H; ##[ ##22: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq23 : ∀a2.eq_ascii ch_L a2 = true → ch_L = a2. #a2; ncases a2; nnormalize; #H; ##[ ##23: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq24 : ∀a2.eq_ascii ch_M a2 = true → ch_M = a2. #a2; ncases a2; nnormalize; #H; ##[ ##24: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq25 : ∀a2.eq_ascii ch_N a2 = true → ch_N = a2. #a2; ncases a2; nnormalize; #H; ##[ ##25: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq26 : ∀a2.eq_ascii ch_O a2 = true → ch_O = a2. #a2; ncases a2; nnormalize; #H; ##[ ##26: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq27 : ∀a2.eq_ascii ch_P a2 = true → ch_P = a2. #a2; ncases a2; nnormalize; #H; ##[ ##27: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq28 : ∀a2.eq_ascii ch_Q a2 = true → ch_Q = a2. #a2; ncases a2; nnormalize; #H; ##[ ##28: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq29 : ∀a2.eq_ascii ch_R a2 = true → ch_R = a2. #a2; ncases a2; nnormalize; #H; ##[ ##29: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq30 : ∀a2.eq_ascii ch_S a2 = true → ch_S = a2. #a2; ncases a2; nnormalize; #H; ##[ ##30: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq31 : ∀a2.eq_ascii ch_T a2 = true → ch_T = a2. #a2; ncases a2; nnormalize; #H; ##[ ##31: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq32 : ∀a2.eq_ascii ch_U a2 = true → ch_U = a2. #a2; ncases a2; nnormalize; #H; ##[ ##32: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq33 : ∀a2.eq_ascii ch_V a2 = true → ch_V = a2. #a2; ncases a2; nnormalize; #H; ##[ ##33: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq34 : ∀a2.eq_ascii ch_W a2 = true → ch_W = a2. #a2; ncases a2; nnormalize; #H; ##[ ##34: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq35 : ∀a2.eq_ascii ch_X a2 = true → ch_X = a2. #a2; ncases a2; nnormalize; #H; ##[ ##35: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq36 : ∀a2.eq_ascii ch_Y a2 = true → ch_Y = a2. #a2; ncases a2; nnormalize; #H; ##[ ##36: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq37 : ∀a2.eq_ascii ch_Z a2 = true → ch_Z = a2. #a2; ncases a2; nnormalize; #H; ##[ ##37: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq38 : ∀a2.eq_ascii ch_a a2 = true → ch_a = a2. #a2; ncases a2; nnormalize; #H; ##[ ##38: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq39 : ∀a2.eq_ascii ch_b a2 = true → ch_b = a2. #a2; ncases a2; nnormalize; #H; ##[ ##39: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq40 : ∀a2.eq_ascii ch_c a2 = true → ch_c = a2. #a2; ncases a2; nnormalize; #H; ##[ ##40: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq41 : ∀a2.eq_ascii ch_d a2 = true → ch_d = a2. #a2; ncases a2; nnormalize; #H; ##[ ##41: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq42 : ∀a2.eq_ascii ch_e a2 = true → ch_e = a2. #a2; ncases a2; nnormalize; #H; ##[ ##42: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq43 : ∀a2.eq_ascii ch_f a2 = true → ch_f = a2. #a2; ncases a2; nnormalize; #H; ##[ ##43: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq44 : ∀a2.eq_ascii ch_g a2 = true → ch_g = a2. #a2; ncases a2; nnormalize; #H; ##[ ##44: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq45 : ∀a2.eq_ascii ch_h a2 = true → ch_h = a2. #a2; ncases a2; nnormalize; #H; ##[ ##45: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq46 : ∀a2.eq_ascii ch_i a2 = true → ch_i = a2. #a2; ncases a2; nnormalize; #H; ##[ ##46: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq47 : ∀a2.eq_ascii ch_j a2 = true → ch_j = a2. #a2; ncases a2; nnormalize; #H; ##[ ##47: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq48 : ∀a2.eq_ascii ch_k a2 = true → ch_k = a2. #a2; ncases a2; nnormalize; #H; ##[ ##48: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq49 : ∀a2.eq_ascii ch_l a2 = true → ch_l = a2. #a2; ncases a2; nnormalize; #H; ##[ ##49: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq50 : ∀a2.eq_ascii ch_m a2 = true → ch_m = a2. #a2; ncases a2; nnormalize; #H; ##[ ##50: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq51 : ∀a2.eq_ascii ch_n a2 = true → ch_n = a2. #a2; ncases a2; nnormalize; #H; ##[ ##51: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq52 : ∀a2.eq_ascii ch_o a2 = true → ch_o = a2. #a2; ncases a2; nnormalize; #H; ##[ ##52: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq53 : ∀a2.eq_ascii ch_p a2 = true → ch_p = a2. #a2; ncases a2; nnormalize; #H; ##[ ##53: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq54 : ∀a2.eq_ascii ch_q a2 = true → ch_q = a2. #a2; ncases a2; nnormalize; #H; ##[ ##54: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq55 : ∀a2.eq_ascii ch_r a2 = true → ch_r = a2. #a2; ncases a2; nnormalize; #H; ##[ ##55: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq56 : ∀a2.eq_ascii ch_s a2 = true → ch_s = a2. #a2; ncases a2; nnormalize; #H; ##[ ##56: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq57 : ∀a2.eq_ascii ch_t a2 = true → ch_t = a2. #a2; ncases a2; nnormalize; #H; ##[ ##57: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq58 : ∀a2.eq_ascii ch_u a2 = true → ch_u = a2. #a2; ncases a2; nnormalize; #H; ##[ ##58: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq59 : ∀a2.eq_ascii ch_v a2 = true → ch_v = a2. #a2; ncases a2; nnormalize; #H; ##[ ##59: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq60 : ∀a2.eq_ascii ch_w a2 = true → ch_w = a2. #a2; ncases a2; nnormalize; #H; ##[ ##60: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq61 : ∀a2.eq_ascii ch_x a2 = true → ch_x = a2. #a2; ncases a2; nnormalize; #H; ##[ ##61: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq62 : ∀a2.eq_ascii ch_y a2 = true → ch_y = a2. #a2; ncases a2; nnormalize; #H; ##[ ##62: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
+nlemma eqascii_to_eq63 : ∀a2.eq_ascii ch_z a2 = true → ch_z = a2. #a2; ncases a2; nnormalize; #H; ##[ ##63: napply refl_eq | ##*: napply (bool_destruct … H) ##] nqed.
nlemma eqascii_to_eq : ∀c1,c2.eq_ascii c1 c2 = true → c1 = c2.
#c1; ncases c1;
##| ##63: napply eqascii_to_eq63 ##]
nqed.
-nlemma eq_to_eqascii1 : ∀a2.ch_0 = a2 → eq_ascii ch_0 a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##1: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii2 : ∀a2.ch_1 = a2 → eq_ascii ch_1 a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##2: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii3 : ∀a2.ch_2 = a2 → eq_ascii ch_2 a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##3: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii4 : ∀a2.ch_3 = a2 → eq_ascii ch_3 a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##4: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii5 : ∀a2.ch_4 = a2 → eq_ascii ch_4 a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##5: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii6 : ∀a2.ch_5 = a2 → eq_ascii ch_5 a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##6: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii7 : ∀a2.ch_6 = a2 → eq_ascii ch_6 a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##7: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii8 : ∀a2.ch_7 = a2 → eq_ascii ch_7 a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##8: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii9 : ∀a2.ch_8 = a2 → eq_ascii ch_8 a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##9: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii10 : ∀a2.ch_9 = a2 → eq_ascii ch_9 a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##10: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii11 : ∀a2.ch__ = a2 → eq_ascii ch__ a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##11: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii12 : ∀a2.ch_A = a2 → eq_ascii ch_A a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##12: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii13 : ∀a2.ch_B = a2 → eq_ascii ch_B a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##13: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii14 : ∀a2.ch_C = a2 → eq_ascii ch_C a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##14: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii15 : ∀a2.ch_D = a2 → eq_ascii ch_D a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##15: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii16 : ∀a2.ch_E = a2 → eq_ascii ch_E a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##16: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii17 : ∀a2.ch_F = a2 → eq_ascii ch_F a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##17: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii18 : ∀a2.ch_G = a2 → eq_ascii ch_G a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##18: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii19 : ∀a2.ch_H = a2 → eq_ascii ch_H a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##19: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii20 : ∀a2.ch_I = a2 → eq_ascii ch_I a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##20: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii21 : ∀a2.ch_J = a2 → eq_ascii ch_J a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##21: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii22 : ∀a2.ch_K = a2 → eq_ascii ch_K a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##22: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii23 : ∀a2.ch_L = a2 → eq_ascii ch_L a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##23: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii24 : ∀a2.ch_M = a2 → eq_ascii ch_M a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##24: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii25 : ∀a2.ch_N = a2 → eq_ascii ch_N a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##25: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii26 : ∀a2.ch_O = a2 → eq_ascii ch_O a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##26: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii27 : ∀a2.ch_P = a2 → eq_ascii ch_P a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##27: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii28 : ∀a2.ch_Q = a2 → eq_ascii ch_Q a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##28: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii29 : ∀a2.ch_R = a2 → eq_ascii ch_R a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##29: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii30 : ∀a2.ch_S = a2 → eq_ascii ch_S a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##30: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii31 : ∀a2.ch_T = a2 → eq_ascii ch_T a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##31: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii32 : ∀a2.ch_U = a2 → eq_ascii ch_U a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##32: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii33 : ∀a2.ch_V = a2 → eq_ascii ch_V a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##33: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii34 : ∀a2.ch_W = a2 → eq_ascii ch_W a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##34: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii35 : ∀a2.ch_X = a2 → eq_ascii ch_X a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##35: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii36 : ∀a2.ch_Y = a2 → eq_ascii ch_Y a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##36: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii37 : ∀a2.ch_Z = a2 → eq_ascii ch_Z a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##37: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii38 : ∀a2.ch_a = a2 → eq_ascii ch_a a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##38: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii39 : ∀a2.ch_b = a2 → eq_ascii ch_b a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##39: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii40 : ∀a2.ch_c = a2 → eq_ascii ch_c a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##40: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii41 : ∀a2.ch_d = a2 → eq_ascii ch_d a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##41: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii42 : ∀a2.ch_e = a2 → eq_ascii ch_e a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##42: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii43 : ∀a2.ch_f = a2 → eq_ascii ch_f a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##43: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii44 : ∀a2.ch_g = a2 → eq_ascii ch_g a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##44: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii45 : ∀a2.ch_h = a2 → eq_ascii ch_h a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##45: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii46 : ∀a2.ch_i = a2 → eq_ascii ch_i a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##46: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii47 : ∀a2.ch_j = a2 → eq_ascii ch_j a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##47: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii48 : ∀a2.ch_k = a2 → eq_ascii ch_k a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##48: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii49 : ∀a2.ch_l = a2 → eq_ascii ch_l a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##49: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii50 : ∀a2.ch_m = a2 → eq_ascii ch_m a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##50: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii51 : ∀a2.ch_n = a2 → eq_ascii ch_n a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##51: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii52 : ∀a2.ch_o = a2 → eq_ascii ch_o a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##52: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii53 : ∀a2.ch_p = a2 → eq_ascii ch_p a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##53: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii54 : ∀a2.ch_q = a2 → eq_ascii ch_q a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##54: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii55 : ∀a2.ch_r = a2 → eq_ascii ch_r a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##55: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii56 : ∀a2.ch_s = a2 → eq_ascii ch_s a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##56: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii57 : ∀a2.ch_t = a2 → eq_ascii ch_t a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##57: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii58 : ∀a2.ch_u = a2 → eq_ascii ch_u a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##58: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii59 : ∀a2.ch_v = a2 → eq_ascii ch_v a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##59: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii60 : ∀a2.ch_w = a2 → eq_ascii ch_w a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##60: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii61 : ∀a2.ch_x = a2 → eq_ascii ch_x a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##61: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii62 : ∀a2.ch_y = a2 → eq_ascii ch_y a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##62: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
-nlemma eq_to_eqascii63 : ∀a2.ch_z = a2 → eq_ascii ch_z a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##63: napply (refl_eq ??) | ##*: napply (ascii_destruct ??? H) ##] nqed.
+nlemma eq_to_eqascii1 : ∀a2.ch_0 = a2 → eq_ascii ch_0 a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##1: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii2 : ∀a2.ch_1 = a2 → eq_ascii ch_1 a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##2: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii3 : ∀a2.ch_2 = a2 → eq_ascii ch_2 a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##3: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii4 : ∀a2.ch_3 = a2 → eq_ascii ch_3 a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##4: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii5 : ∀a2.ch_4 = a2 → eq_ascii ch_4 a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##5: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii6 : ∀a2.ch_5 = a2 → eq_ascii ch_5 a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##6: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii7 : ∀a2.ch_6 = a2 → eq_ascii ch_6 a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##7: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii8 : ∀a2.ch_7 = a2 → eq_ascii ch_7 a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##8: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii9 : ∀a2.ch_8 = a2 → eq_ascii ch_8 a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##9: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii10 : ∀a2.ch_9 = a2 → eq_ascii ch_9 a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##10: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii11 : ∀a2.ch__ = a2 → eq_ascii ch__ a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##11: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii12 : ∀a2.ch_A = a2 → eq_ascii ch_A a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##12: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii13 : ∀a2.ch_B = a2 → eq_ascii ch_B a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##13: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii14 : ∀a2.ch_C = a2 → eq_ascii ch_C a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##14: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii15 : ∀a2.ch_D = a2 → eq_ascii ch_D a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##15: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii16 : ∀a2.ch_E = a2 → eq_ascii ch_E a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##16: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii17 : ∀a2.ch_F = a2 → eq_ascii ch_F a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##17: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii18 : ∀a2.ch_G = a2 → eq_ascii ch_G a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##18: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii19 : ∀a2.ch_H = a2 → eq_ascii ch_H a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##19: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii20 : ∀a2.ch_I = a2 → eq_ascii ch_I a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##20: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii21 : ∀a2.ch_J = a2 → eq_ascii ch_J a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##21: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii22 : ∀a2.ch_K = a2 → eq_ascii ch_K a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##22: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii23 : ∀a2.ch_L = a2 → eq_ascii ch_L a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##23: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii24 : ∀a2.ch_M = a2 → eq_ascii ch_M a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##24: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii25 : ∀a2.ch_N = a2 → eq_ascii ch_N a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##25: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii26 : ∀a2.ch_O = a2 → eq_ascii ch_O a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##26: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii27 : ∀a2.ch_P = a2 → eq_ascii ch_P a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##27: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii28 : ∀a2.ch_Q = a2 → eq_ascii ch_Q a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##28: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii29 : ∀a2.ch_R = a2 → eq_ascii ch_R a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##29: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii30 : ∀a2.ch_S = a2 → eq_ascii ch_S a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##30: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii31 : ∀a2.ch_T = a2 → eq_ascii ch_T a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##31: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii32 : ∀a2.ch_U = a2 → eq_ascii ch_U a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##32: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii33 : ∀a2.ch_V = a2 → eq_ascii ch_V a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##33: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii34 : ∀a2.ch_W = a2 → eq_ascii ch_W a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##34: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii35 : ∀a2.ch_X = a2 → eq_ascii ch_X a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##35: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii36 : ∀a2.ch_Y = a2 → eq_ascii ch_Y a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##36: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii37 : ∀a2.ch_Z = a2 → eq_ascii ch_Z a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##37: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii38 : ∀a2.ch_a = a2 → eq_ascii ch_a a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##38: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii39 : ∀a2.ch_b = a2 → eq_ascii ch_b a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##39: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii40 : ∀a2.ch_c = a2 → eq_ascii ch_c a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##40: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii41 : ∀a2.ch_d = a2 → eq_ascii ch_d a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##41: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii42 : ∀a2.ch_e = a2 → eq_ascii ch_e a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##42: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii43 : ∀a2.ch_f = a2 → eq_ascii ch_f a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##43: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii44 : ∀a2.ch_g = a2 → eq_ascii ch_g a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##44: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii45 : ∀a2.ch_h = a2 → eq_ascii ch_h a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##45: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii46 : ∀a2.ch_i = a2 → eq_ascii ch_i a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##46: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii47 : ∀a2.ch_j = a2 → eq_ascii ch_j a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##47: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii48 : ∀a2.ch_k = a2 → eq_ascii ch_k a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##48: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii49 : ∀a2.ch_l = a2 → eq_ascii ch_l a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##49: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii50 : ∀a2.ch_m = a2 → eq_ascii ch_m a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##50: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii51 : ∀a2.ch_n = a2 → eq_ascii ch_n a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##51: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii52 : ∀a2.ch_o = a2 → eq_ascii ch_o a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##52: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii53 : ∀a2.ch_p = a2 → eq_ascii ch_p a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##53: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii54 : ∀a2.ch_q = a2 → eq_ascii ch_q a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##54: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii55 : ∀a2.ch_r = a2 → eq_ascii ch_r a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##55: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii56 : ∀a2.ch_s = a2 → eq_ascii ch_s a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##56: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii57 : ∀a2.ch_t = a2 → eq_ascii ch_t a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##57: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii58 : ∀a2.ch_u = a2 → eq_ascii ch_u a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##58: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii59 : ∀a2.ch_v = a2 → eq_ascii ch_v a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##59: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii60 : ∀a2.ch_w = a2 → eq_ascii ch_w a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##60: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii61 : ∀a2.ch_x = a2 → eq_ascii ch_x a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##61: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii62 : ∀a2.ch_y = a2 → eq_ascii ch_y a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##62: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
+nlemma eq_to_eqascii63 : ∀a2.ch_z = a2 → eq_ascii ch_z a2 = true. #a2; ncases a2; nnormalize; #H; ##[ ##63: napply refl_eq | ##*: napply (ascii_destruct … H) ##] nqed.
nlemma eq_to_eqascii : ∀c1,c2.c1 = c2 → eq_ascii c1 c2 = true.
#c1; ncases c1;
nchange with (match mk_strId x2 y2 with [ mk_strId a _ ⇒ x1 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma strid_destruct_2 : ∀x1,x2,y1,y2.mk_strId x1 y1 = mk_strId x2 y2 → y1 = y2.
nchange with (match mk_strId x2 y2 with [ mk_strId _ b ⇒ y1 = b ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_eqstrid : symmetricT strId bool eq_strId.
((eq_str (str_elem si2) (str_elem si1))⊗(eq_nat (id_elem si2) (id_elem si1))));
nrewrite > (symmetric_eqstr (str_elem si1) (str_elem si2));
nrewrite > (symmetric_eqnat (id_elem si1) (id_elem si2));
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma eqstrid_to_eq : ∀s,s'.eq_strId s s' = true → s = s'.
nelim si2;
#l2; #n2; #H;
nchange in H:(%) with (((eq_str l1 l2)⊗(eq_nat n1 n2)) = true);
- nrewrite > (eqstr_to_eq l1 l2 (andb_true_true_l ?? H));
- nrewrite > (eqnat_to_eq n1 n2 (andb_true_true_r ?? H));
- napply (refl_eq ??).
+ nrewrite > (eqstr_to_eq l1 l2 (andb_true_true_l … H));
+ nrewrite > (eqnat_to_eq n1 n2 (andb_true_true_r … H));
+ napply refl_eq.
nqed.
nlemma eq_to_eqstrid : ∀s,s'.s = s' → eq_strId s s' = true.
nelim si2;
#l2; #n2; #H;
nchange with (((eq_str l1 l2)⊗(eq_nat n1 n2)) = true);
- nrewrite > (strid_destruct_1 ???? H);
- nrewrite > (strid_destruct_2 ???? H);
- nrewrite > (eq_to_eqstr l2 l2 (refl_eq ??));
- nrewrite > (eq_to_eqnat n2 n2 (refl_eq ??));
+ nrewrite > (strid_destruct_1 … H);
+ nrewrite > (strid_destruct_2 … H);
+ nrewrite > (eq_to_eqstr l2 l2 (refl_eq …));
+ nrewrite > (eq_to_eqnat n2 n2 (refl_eq …));
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
#T; #h; #h'; #t; #t';
nelim t;
nelim t';
- ##[ ##1: nnormalize; #H; napply (refl_eq ??)
+ ##[ ##1: nnormalize; #H; napply refl_eq
##| ##2: #a; #l'; #H; #H1;
nchange in H1:(%) with ((S O) = (S (S (len_list T l'))));
- nelim (nat_destruct_0_S ? (nat_destruct_S_S ?? H1))
+ nelim (nat_destruct_0_S ? (nat_destruct_S_S … H1))
##| ##3: #a; #l'; #H; #H1;
nchange in H1:(%) with ((S (S (len_list T l'))) = (S O));
- nelim (nat_destruct_S_0 ? (nat_destruct_S_S ?? H1))
+ nelim (nat_destruct_S_0 ? (nat_destruct_S_S … H1))
##| ##4: #a; #l; #H; #a1; #l1; #H1; #H2;
nchange in H2:(%) with ((S (S (len_list T l1))) = (S (S (len_list T l))));
nchange with ((S (len_list T l1)) = (S (len_list T l)));
- nrewrite > (nat_destruct_S_S ?? H2);
- napply (refl_eq ??)
+ nrewrite > (nat_destruct_S_S … H2);
+ napply refl_eq
##]
nqed.
nnormalize;
ncases t';
nnormalize;
- ##[ ##1: #x; #H; nelim (nat_destruct_0_S ? (nat_destruct_S_S ?? H))
- ##| ##2: #x; #l; #H; nelim (nat_destruct_0_S ? (nat_destruct_S_S ?? H))
+ ##[ ##1: #x; #H; nelim (nat_destruct_0_S ? (nat_destruct_S_S … H))
+ ##| ##2: #x; #l; #H; nelim (nat_destruct_0_S ? (nat_destruct_S_S … H))
##]
nqed.
nnormalize;
ncases t;
nnormalize;
- ##[ ##1: #x; #H; nelim (nat_destruct_S_0 ? (nat_destruct_S_S ?? H))
- ##| ##2: #x; #l; #H; nelim (nat_destruct_S_0 ? (nat_destruct_S_S ?? H))
+ ##[ ##1: #x; #H; nelim (nat_destruct_S_0 ? (nat_destruct_S_S … H))
+ ##| ##2: #x; #l; #H; nelim (nat_destruct_S_0 ? (nat_destruct_S_S … H))
##]
nqed.
#T; #h; #h'; #t; #t';
nelim t;
nelim t';
- ##[ ##1: nnormalize; #x; #y; #H; napply (refl_eq ??)
+ ##[ ##1: nnormalize; #x; #y; #H; napply refl_eq
##| ##2: #a; #l'; #H; #x; #H1;
nchange in H1:(%) with ((S (len_neList T «£x»)) = (S (len_neList T (a§§l'))));
- nrewrite > (nat_destruct_S_S ?? H1);
- napply (refl_eq ??)
+ nrewrite > (nat_destruct_S_S … H1);
+ napply refl_eq
##| ##3: #x; #a; #l'; #H; #H1;
nchange in H1:(%) with ((S (len_neList T (a§§l')))= (S (len_neList T «£x»)));
- nrewrite > (nat_destruct_S_S ?? H1);
- napply (refl_eq ??)
+ nrewrite > (nat_destruct_S_S … H1);
+ napply refl_eq
##| ##4: #a; #l; #H; #a1; #l1; #H1; #H2;
nchange in H2:(%) with ((S (len_neList T (a1§§l1))) = (S (len_neList T (a§§l))));
- nrewrite > (nat_destruct_S_S ?? H2);
- napply (refl_eq ??)
+ nrewrite > (nat_destruct_S_S … H2);
+ napply refl_eq
##]
nqed.
nchange with (match ne_nil T x2 with [ ne_cons _ _ ⇒ False | ne_nil a ⇒ x1 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma nelist_destruct_cons_cons_1 : ∀T.∀x1,x2:T.∀y1,y2:ne_list T.ne_cons T x1 y1 = ne_cons T x2 y2 → x1 = x2.
nchange with (match ne_cons T x2 y2 with [ ne_nil _ ⇒ False | ne_cons a _ ⇒ x1 = a ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma nelist_destruct_cons_cons_2 : ∀T.∀x1,x2:T.∀y1,y2:ne_list T.ne_cons T x1 y1 = ne_cons T x2 y2 → y1 = y2.
nchange with (match ne_cons T x2 y2 with [ ne_nil _ ⇒ False | ne_cons _ b ⇒ y1 = b ]);
nrewrite < H;
nnormalize;
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma nelist_destruct_cons_nil : ∀T.∀x1,x2:T.∀y1:ne_list T.ne_cons T x1 y1 = ne_nil T x2 → False.
#T; #l1;
nelim l1;
##[ ##1: #l2; ncases l2; nnormalize;
- ##[ ##1: #H; napply (refl_eq ??)
+ ##[ ##1: #H; napply refl_eq
##| ##2: #h; #t; #H; nelim (nat_destruct_0_S ? H)
##]
##| ##2: #h; #l2; ncases l2; nnormalize;
- ##[ ##1: #H; #l; #H1; nrewrite < H1; napply (refl_eq ??)
- ##| ##2: #h; #l; #H; #l3; #H1; nrewrite < H1; napply (refl_eq ??)
+ ##[ ##1: #H; #l; #H1; nrewrite < H1; napply refl_eq
+ ##| ##2: #h; #l; #H; #l3; #H1; nrewrite < H1; napply refl_eq
##]
##]
nqed.
#T1; #T2; #f; #acc; #l1;
nelim l1;
##[ ##1: #l2; ncases l2;
- ##[ ##1: nnormalize; #H1; #H2; #H3; napply (refl_eq ??)
+ ##[ ##1: nnormalize; #H1; #H2; #H3; napply refl_eq
##| ##2: #h; #l; #H1; #H2; #H3;
nchange in H1:(%) with (O = (S (len_list ? l)));
nelim (nat_destruct_0_S ? H1)
(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 ?));
- napply (refl_eq ??)
+ napply refl_eq
##]
##]
nqed.
fold_right_list2 T1 T2 f acc l1 l2 H = fold_right_list2 T1 T2 f acc l2 l1 (symmetric_eqlenlist T1 l1 l2 H).
#T1; #T2; #f; #acc; #l1; #l2; #H; #H1;
nrewrite > (symmetric_foldrightlist2_aux T1 T2 f acc l1 l2 H (symmetric_eqlenlist T1 l1 l2 H) H1);
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_bfoldrightlist2
#T; #f; #l1;
nelim l1;
##[ ##1: #l2; ncases l2;
- ##[ ##1: #H; nnormalize; napply (refl_eq ??)
- ##| ##2: #hh2; #ll2; #H; nnormalize; napply (refl_eq ??)
+ ##[ ##1: #H; nnormalize; napply refl_eq
+ ##| ##2: #hh2; #ll2; #H; nnormalize; napply refl_eq
##]
##| ##2: #hh1; #ll1; #H; #l2; ncases l2;
- ##[ ##1: #H1; nnormalize; napply (refl_eq ??)
+ ##[ ##1: #H1; nnormalize; napply refl_eq
##| ##2: #hh2; #ll2; #H1; nnormalize;
nrewrite > (H ll2 H1);
nrewrite > (H1 hh1 hh2);
- napply (refl_eq ??)
+ napply refl_eq
##]
##]
nqed.
#T; #f; #l1;
nelim l1;
##[ ##1: #l2; ncases l2;
- ##[ ##1: #H; #H1; napply (refl_eq ??)
- ##| ##2: #hh2; #ll2; #H; nnormalize; #H1; napply (bool_destruct ??? H1)
+ ##[ ##1: #H; #H1; napply refl_eq
+ ##| ##2: #hh2; #ll2; #H; nnormalize; #H1; napply (bool_destruct … H1)
##]
##| ##2: #hh1; #ll1; #H; #l2; ncases l2;
- ##[ ##1: #H1; nnormalize; #H2; napply (bool_destruct ??? H2)
+ ##[ ##1: #H1; nnormalize; #H2; napply (bool_destruct … H2)
##| ##2: #hh2; #ll2; #H1; #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));
- napply (refl_eq ??)
+ nrewrite > (H1 hh1 hh2 (andb_true_true_l … H2));
+ nrewrite > (H ll2 H1 (andb_true_true_r … H2));
+ napply refl_eq
##]
##]
nqed.
#T; #f; #l1;
nelim l1;
##[ ##1: #l2; ncases l2;
- ##[ ##1: #H; #H1; nnormalize; napply (refl_eq ??)
+ ##[ ##1: #H; #H1; nnormalize; napply refl_eq
##| ##2: #hh2; #ll2; #H; #H1; nelim (list_destruct_nil_cons ??? H1)
##]
##| ##2: #hh1; #ll1; #H; #l2; ncases l2;
##[ ##1: #H1; #H2; nelim (list_destruct_cons_nil ??? H2)
##| ##2: #hh2; #ll2; #H1; #H2; nnormalize;
- nrewrite > (list_destruct_1 ????? H2);
- nrewrite > (H1 hh2 hh2 (refl_eq ??));
+ nrewrite > (list_destruct_1 … H2);
+ nrewrite > (H1 hh2 hh2 (refl_eq …));
nnormalize;
- nrewrite > (H ll2 H1 (list_destruct_2 ????? H2));
- napply (refl_eq ??)
+ nrewrite > (H ll2 H1 (list_destruct_2 … H2));
+ napply refl_eq
##]
##]
nqed.
#T; #l1;
nelim l1;
##[ ##1: #h; #l2; ncases l2; nnormalize;
- ##[ ##1: #H; #H1; napply (refl_eq ??)
- ##| ##2: #h; #t; #H; nrewrite > H; napply (refl_eq ??)
+ ##[ ##1: #H; #H1; napply refl_eq
+ ##| ##2: #h; #t; #H; nrewrite > H; napply refl_eq
##]
##| ##2: #h; #l2; ncases l2; nnormalize;
- ##[ ##1: #h1; #H; #l; #H1; nrewrite < H1; napply (refl_eq ??)
- ##| ##2: #h; #l; #H; #l3; #H1; nrewrite < H1; napply (refl_eq ??)
+ ##[ ##1: #h1; #H; #l; #H1; nrewrite < H1; napply refl_eq
+ ##| ##2: #h; #l; #H; #l3; #H1; nrewrite < H1; napply refl_eq
##]
##]
nqed.
#T1; #T2; #f; #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; #H3; nrewrite > (H3 h h1 acc); napply refl_eq
##| ##2: #h1; #l; ncases l;
##[ ##1: #h3; #H1; #H2; #H3;
nchange in H1:(%) with ((S O) = (S (S O)));
- nelim (nat_destruct_0_S ? (nat_destruct_S_S ?? H1))
+ nelim (nat_destruct_0_S ? (nat_destruct_S_S … H1))
##| ##2: #h3; #l3; #H1; #H2; #H3;
nchange in H1:(%) with ((S O) = (S (S (len_neList ? l3))));
- nelim (nat_destruct_0_S ? (nat_destruct_S_S ?? H1))
+ nelim (nat_destruct_0_S ? (nat_destruct_S_S … H1))
##]
##]
##| ##2: #h3; #l3; #H; #l2; ncases l2;
##[ ##1: #h4; ncases l3;
##[ ##1: #h5; #H1; #H2; #H3;
nchange in H1:(%) with ((S (S O)) = (S O));
- nelim (nat_destruct_S_0 ? (nat_destruct_S_S ?? 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 ?? H1))
+ nelim (nat_destruct_S_0 ? (nat_destruct_S_S … H1))
##]
##| ##2: #h4; #l4; #H1; #H2; #H3;
nchange in H1:(%) with ((S (len_neList ? l3)) = (S (len_neList ? 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 ?));
- napply (refl_eq ??)
+ napply refl_eq
##]
##]
nqed.
fold_right_neList2 T1 T2 f acc l1 l2 H = fold_right_neList2 T1 T2 f acc l2 l1 (symmetric_eqlennelist T1 l1 l2 H).
#T1; #T2; #f; #acc; #l1; #l2; #H; #H1;
nrewrite > (symmetric_foldrightnelist2_aux T1 T2 f acc l1 l2 H (symmetric_eqlennelist T1 l1 l2 H) H1);
- napply (refl_eq ??).
+ napply refl_eq.
nqed.
nlemma symmetric_bfoldrightnelist2
#T; #f; #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 ??)
+ ##[ ##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 ??)
+ ##[ ##1: #hh2; #H1; nnormalize; napply refl_eq
##| ##2: #hh2; #ll2; #H1; nnormalize;
nrewrite > (H ll2 H1);
nrewrite > (H1 hh1 hh2);
- napply (refl_eq ??)
+ napply refl_eq
##]
##]
nqed.
#T; #f; #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; napply (bool_destruct ??? H1)
+ ##[ ##1: #hh2; #H; #H1; nnormalize in H1:(%); nrewrite > (H hh1 hh2 H1); napply refl_eq
+ ##| ##2: #hh2; #ll2; #H; nnormalize; #H1; napply (bool_destruct … H1)
##]
##| ##2: #hh1; #ll1; #H; #l2; ncases l2;
- ##[ ##1: #hh2; #H1; nnormalize; #H2; napply (bool_destruct ??? H2)
+ ##[ ##1: #hh2; #H1; nnormalize; #H2; napply (bool_destruct … H2)
##| ##2: #hh2; #ll2; #H1; #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));
- napply (refl_eq ??)
+ nrewrite > (H1 hh1 hh2 (andb_true_true_l … H2));
+ nrewrite > (H ll2 H1 (andb_true_true_r … H2));
+ napply refl_eq
##]
##]
nqed.
nelim l1;
##[ ##1: #hh1; #l2; ncases l2;
##[ ##1: #hh2; #H; #H1; nnormalize;
- nrewrite > (H hh1 hh2 (nelist_destruct_nil_nil ??? H1));
- napply (refl_eq ??)
+ nrewrite > (H hh1 hh2 (nelist_destruct_nil_nil … H1));
+ napply refl_eq
##| ##2: #hh2; #ll2; #H; #H1; nelim (nelist_destruct_nil_cons ???? H1)
##]
##| ##2: #hh1; #ll1; #H; #l2; ncases l2;
##[ ##1: #hh2; #H1; #H2; nelim (nelist_destruct_cons_nil ???? H2)
##| ##2: #hh2; #ll2; #H1; #H2; nnormalize;
- nrewrite > (nelist_destruct_cons_cons_1 ????? H2);
- nrewrite > (H1 hh2 hh2 (refl_eq ??));
+ nrewrite > (nelist_destruct_cons_cons_1 … H2);
+ nrewrite > (H1 hh2 hh2 (refl_eq …));
nnormalize;
- nrewrite > (H ll2 H1 (nelist_destruct_cons_cons_2 ????? H2));
- napply (refl_eq ??)
+ nrewrite > (H ll2 H1 (nelist_destruct_cons_cons_2 … H2));
+ napply refl_eq
##]
##]
nqed.
ncases l;
nnormalize;
##[ ##1: #H; napply I
- ##| ##2: #x; #l; #H; napply (bool_destruct ??? H)
+ ##| ##2: #x; #l; #H; napply (bool_destruct … H)
##]
nqed.
#T; #l;
ncases l;
nnormalize;
- ##[ ##1: #H; napply (bool_destruct ??? H)
+ ##[ ##1: #H; napply (bool_destruct … H)
##| ##2: #x; #l; #H; napply I
##]
nqed.
#n;
ncases n;
##[ ##1: nnormalize; #H; napply I
- ##| ##2: #n1; nnormalize; #H; napply (bool_destruct ??? H)
+ ##| ##2: #n1; nnormalize; #H; napply (bool_destruct … H)
##]
nqed.