ndefinition astbasetype_destruct_aux ≝
Πb1,b2:ast_base_type.ΠP:Prop.b1 = b2 →
- match b1 with
- [ AST_BASE_TYPE_BYTE8 ⇒ match b2 with [ AST_BASE_TYPE_BYTE8 ⇒ P → P | _ ⇒ P ]
- | AST_BASE_TYPE_WORD16 ⇒ match b2 with [ AST_BASE_TYPE_WORD16 ⇒ P → P | _ ⇒ P ]
- | AST_BASE_TYPE_WORD32 ⇒ match b2 with [ AST_BASE_TYPE_WORD32 ⇒ P → P | _ ⇒ P ]
- ].
+ match eq_astbasetype b1 b2 with [ true ⇒ P → P | false ⇒ P ].
ndefinition astbasetype_destruct : astbasetype_destruct_aux.
- #b1; #b2; #P;
+ #b1; #b2; #P; #H;
+ nrewrite < H;
nelim b1;
- nelim b2;
nnormalize;
- #H;
- ##[ ##1,5,9: napply (λx:P.x)
- ##| ##2,3: napply False_ind;
- nchange with (match AST_BASE_TYPE_BYTE8 with [ AST_BASE_TYPE_BYTE8 ⇒ False | _ ⇒ True]);
- nrewrite > H;
- nnormalize;
- napply I
- ##| ##4,6: napply False_ind;
- nchange with (match AST_BASE_TYPE_WORD16 with [ AST_BASE_TYPE_WORD16 ⇒ False | _ ⇒ True]);
- nrewrite > H;
- nnormalize;
- napply I
- ##| ##7,8: napply False_ind;
- nchange with (match AST_BASE_TYPE_WORD32 with [ AST_BASE_TYPE_WORD32 ⇒ False | _ ⇒ True]);
- nrewrite > H;
- nnormalize;
- napply I
- ##]
+ napply (λx.x).
nqed.
-nlemma symmetric_eqastbasetype : symmetricT ast_base_type bool eq_ast_base_type.
+nlemma symmetric_eqastbasetype : symmetricT ast_base_type bool eq_astbasetype.
#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.
+nlemma eqastbasetype_to_eq : ∀b1,b2.eq_astbasetype b1 b2 = true → b1 = b2.
#b1; #b2; ncases b1; ncases b2; nnormalize;
##[ ##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.
+nlemma eq_to_eqastbasetype : ∀b1,b2.b1 = b2 → eq_astbasetype b1 b2 = true.
#b1; #b2; ncases b1; ncases b2; nnormalize;
##[ ##1,5,9: #H; napply refl_eq
##| ##*: #H; napply (astbasetype_destruct … H)
##]
nqed.
+nlemma decidable_astbasetype : ∀x,y:ast_base_type.decidable (x = y).
+ #x; #y;
+ nnormalize;
+ nelim x;
+ nelim y;
+ ##[ ##1,5,9: napply (or2_intro1 (? = ?) (? ≠ ?) …); napply refl_eq
+ ##| ##*: napply (or2_intro2 (? = ?) (? ≠ ?) …);
+ nnormalize; #H;
+ napply False_ind;
+ napply (astbasetype_destruct … H)
+ ##]
+nqed.
+
+nlemma neqastbasetype_to_neq : ∀b1,b2.(eq_astbasetype b1 b2 = false) → (b1 ≠ b2).
+ #b1; #b2;
+ ncases b1;
+ ncases b2;
+ nnormalize;
+ ##[ ##1,5,9: #H; napply (bool_destruct … H)
+ ##| ##*: #H; #H1; napply (astbasetype_destruct … H1)
+ ##]
+nqed.
+
+nlemma neq_to_neqastbasetype : ∀b1,b2.b1 ≠ b2 → eq_astbasetype b1 b2 = false.
+ #b1; #b2;
+ ncases b1;
+ ncases b2;
+ nnormalize;
+ ##[ ##1,5,9: #H; nelim (H (refl_eq …))
+ ##| ##*: #H; napply refl_eq
+ ##]
+nqed.
+
nlemma asttype_destruct_base_base : ∀b1,b2.AST_TYPE_BASE b1 = AST_TYPE_BASE b2 → b1 = b2.
#b1; #b2; #H;
nchange with (match AST_TYPE_BASE b2 with [ AST_TYPE_BASE a ⇒ b1 = a | _ ⇒ False ]);
ndefinition asttype_destruct_aux ≝
Πb1,b2:ast_type.ΠP:Prop.b1 = b2 →
- match b1 with
- [ AST_TYPE_BASE s1 ⇒ match b2 with
- [ AST_TYPE_BASE s2 ⇒ match s1 with
- [ AST_BASE_TYPE_BYTE8 ⇒ match s2 with [ AST_BASE_TYPE_BYTE8 ⇒ P → P | _ ⇒ P ]
- | AST_BASE_TYPE_WORD16 ⇒ match s2 with [ AST_BASE_TYPE_WORD16 ⇒ P → P | _ ⇒ P ]
- | AST_BASE_TYPE_WORD32 ⇒ match s2 with [ AST_BASE_TYPE_WORD32 ⇒ P → P | _ ⇒ P ]
- ] | _ ⇒ P ]
- | AST_TYPE_ARRAY _ _ ⇒ match b2 with [ AST_TYPE_ARRAY _ _ ⇒ P → P | _ ⇒ P ]
- | AST_TYPE_STRUCT _ ⇒ match b2 with [ AST_TYPE_STRUCT _ ⇒ P → P | _ ⇒ P ]
- ].
+ match eq_asttype b1 b2 with [ true ⇒ P → P | false ⇒ P ].
ndefinition asttype_destruct : asttype_destruct_aux.
- #b1; #b2; #P;
- ncases b1;
- ##[ ##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))
- ##]
- ##| ##2: #t; #n; #b; nnormalize; #H
- ##| ##3: #l; #b; nnormalize; #H
- ##]
- napply False_ind;
- nchange with (match AST_TYPE_BASE b with [ AST_TYPE_BASE _ ⇒ False | _ ⇒ True ]);
- nrewrite > H; nnormalize; napply I
- ##| ##2: ncases b2;
- ##[ ##2: #t1; #n1; #t2; #n2; nnormalize; #H; napply (λx:P.x)
- ##| ##1: #b; #t; #n; nnormalize; #H
- ##| ##3: #l; #t; #n; nnormalize; #H
- ##]
- napply False_ind;
- nchange with (match AST_TYPE_ARRAY t n with [ AST_TYPE_ARRAY _ _ ⇒ False | _ ⇒ True ]);
- nrewrite > H; nnormalize; napply I
- ##| ##3: ncases b2;
- ##[ ##3: #l1; #l2; nnormalize; #H; napply (λx:P.x)
- ##| ##1: #b; #l; nnormalize; #H
- ##| ##2: #t; #n; #l; nnormalize; #H
+ #b1; #b2; #P; #H;
+ nrewrite > H;
+ napply (ast_type_index … b2);
+ ##[ ##1: #e; nchange with (match eq_astbasetype e e with [ true ⇒ P → P | false ⇒ P ]);
+ nrewrite > (eq_to_eqastbasetype e e (refl_eq …));
+ nnormalize; napply (λx.x);
+ ##| ##2: #e; #n; #H; nchange with (match (eq_asttype e e)⊗(eq_nat n n) with [ true ⇒ P → P | false ⇒ P]);
+ nrewrite > (eq_to_eqnat n n (refl_eq …));
+ nrewrite > (symmetric_andbool (eq_asttype e e) true);
+ nchange with (match eq_asttype e e with [ true ⇒ P → P | false ⇒ P]);
+ napply H;
+ ##| ##3: #e; #H; nchange with (match eq_asttype e e with [ true ⇒ P → P | false ⇒ P]);
+ napply H;
+ ##| ##4: #hh; #tt; #H;
+ nchange with (match bfold_right_neList2 ?? tt tt with [ true ⇒ P → P | false ⇒ P ] →
+ match (eq_asttype hh hh)⊗(bfold_right_neList2 ?? tt tt) with [ true ⇒ P → P | false ⇒ P ]);
+ #H1;
+ ncases (eq_asttype hh hh) in H:(%) ⊢ %; #H;
+ ncases (bfold_right_neList2 ? (λx1,x2.eq_asttype x1 x2) tt tt) in H1:(%) ⊢ %; #H1;
+ ##[ ##1: nnormalize; napply (λx.x);
+ ##| ##3: nnormalize in H:(%) ⊢ %; napply H
+ ##| ##*: nnormalize in H1:(%) ⊢ %; napply H1
##]
- napply False_ind;
- nchange with (match AST_TYPE_STRUCT l with [ AST_TYPE_STRUCT _ ⇒ False | _ ⇒ True ]);
- nrewrite > H; nnormalize; napply I
##]
nqed.
nlemma symmetric_eqasttype_aux1
: ∀nl1,nl2.
- (eq_ast_type (AST_TYPE_STRUCT nl1) (AST_TYPE_STRUCT nl2)) = (eq_ast_type (AST_TYPE_STRUCT nl2) (AST_TYPE_STRUCT nl1)) →
- (bfold_right_neList2 ? (λx,y.eq_ast_type x y) nl1 nl2) = (bfold_right_neList2 ? (λx,y.eq_ast_type x y) nl2 nl1).
+ (eq_asttype (AST_TYPE_STRUCT nl1) (AST_TYPE_STRUCT nl2)) = (eq_asttype (AST_TYPE_STRUCT nl2) (AST_TYPE_STRUCT nl1)) →
+ (bfold_right_neList2 ? (λx,y.eq_asttype x y) nl1 nl2) = (bfold_right_neList2 ? (λx,y.eq_asttype x y) nl2 nl1).
#nl1; #nl2; #H;
napply H.
nqed.
-nlemma symmetric_eqasttype : symmetricT ast_type bool eq_ast_type.
+nlemma symmetric_eqasttype : symmetricT ast_type bool eq_asttype.
#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));
+ ##[ ##1: #b2; nchange with ((eq_astbasetype b1 b2) = (eq_astbasetype b2 b1));
nrewrite > (symmetric_eqastbasetype b1 b2);
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)));
+ ##[ ##2: #st2; #n2; nchange with (((eq_asttype st1 st2)⊗(eq_nat n1 n2)) = ((eq_asttype st2 st1)⊗(eq_nat n2 n1)));
nrewrite > (symmetric_eqnat n1 n2);
nrewrite > (H st2);
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));
+ ##[ ##1: #hh2; nchange with ((eq_asttype hh1 hh2) = (eq_asttype hh2 hh1));
nrewrite > (H hh2);
napply refl_eq
##| ##2: #hh2; #ll2; nnormalize; napply refl_eq
##]
nqed.
-nlemma eqasttype_to_eq : ∀t1,t2.eq_ast_type t1 t2 = true → t1 = t2.
+nlemma eqasttype_to_eq : ∀t1,t2.eq_asttype t1 t2 = true → t1 = t2.
#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);
+ ##[ ##1: #b2; #H; nchange in H:(%) with ((eq_astbasetype 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)
##]
##| ##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);
+ ##[ ##2: #st2; #n2; #H1; nchange in H1:(%) with (((eq_asttype 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
##]
##| ##3: #hh1; #H; #t2; ncases t2;
##[ ##3: #nl2; ncases nl2;
- ##[ ##1: #hh2; #H1; nchange in H1:(%) with ((eq_ast_type hh1 hh2) = true);
+ ##[ ##1: #hh2; #H1; nchange in H1:(%) with ((eq_asttype hh1 hh2) = true);
nrewrite > (H hh2 H1);
napply refl_eq
##| ##2: #hh2; #ll2; 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)
- ##| ##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);
+ ##| ##2: #hh2; #ll2; #H2; nchange in H2:(%) with (((eq_asttype hh1 hh2)⊗(bfold_right_neList2 ? (λx,y.eq_asttype 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
nlemma eq_to_eqasttype_aux1
: ∀nl1,nl2.
- ((eq_ast_type (AST_TYPE_STRUCT nl1) (AST_TYPE_STRUCT nl2)) = true) →
- ((bfold_right_neList2 ? (λx,y.eq_ast_type x y) nl1 nl2) = true).
+ ((eq_asttype (AST_TYPE_STRUCT nl1) (AST_TYPE_STRUCT nl2)) = true) →
+ ((bfold_right_neList2 ? (λx,y.eq_asttype x y) nl1 nl2) = true).
#nl1; #nl2; #H;
napply H.
nqed.
-nlemma eq_to_eqasttype : ∀t1,t2.t1 = t2 → eq_ast_type t1 t2 = true.
+nlemma eq_to_eqasttype : ∀t1,t2.t1 = t2 → eq_asttype t1 t2 = true.
#t1;
napply (ast_type_index … t1);
##[ ##1: #b1; #t2; ncases t2;
##[ ##1: #b2; #H; nrewrite > (asttype_destruct_base_base … H);
- nchange with ((eq_ast_base_type b2 b2) = true);
+ nchange with ((eq_astbasetype 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)
##]
##| ##2: #st1; #n1; #H; #t2; ncases t2;
- ##[ ##2: #st2; #n2; #H1; nchange with (((eq_ast_type st1 st2)⊗(eq_nat n1 n2)) = true);
+ ##[ ##2: #st2; #n2; #H1; nchange with (((eq_asttype 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));
nnormalize;
##]
##| ##3: #hh1; #H; #t2; ncases t2;
##[ ##3: #nl2; ncases nl2;
- ##[ ##1: #hh2; #H1; nchange with ((eq_ast_type hh1 hh2) = true);
+ ##[ ##1: #hh2; #H1; nchange with ((eq_asttype 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))
##| ##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))
- ##| ##2: #hh2; #ll2; #H2; nchange with (((eq_ast_type hh1 hh2)⊗(bfold_right_neList2 ? (λx,y.eq_ast_type x y) ll1 ll2)) = true);
+ ##| ##2: #hh2; #ll2; #H2; nchange with (((eq_asttype hh1 hh2)⊗(bfold_right_neList2 ? (λx,y.eq_asttype x y) ll1 ll2)) = true);
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
##]
nqed.
+nlemma decidable_asttype : ∀x,y:ast_type.decidable (x = y).
+ #x;
+ napply (ast_type_index … x);
+ ##[ ##1: #b1; #y; ncases y;
+ ##[ ##1: #b2; nnormalize; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_astbasetype b1 b2));
+ ##[ ##1: #H; nrewrite > H; napply (or2_intro1 (? = ?) (? ≠ ?) (refl_eq …))
+ ##| ##2: #H; napply (or2_intro2 (? = ?) (? ≠ ?) …);
+ nnormalize; #H1; napply (H (asttype_destruct_base_base … H1))
+ ##]
+ ##| ##2: #b2; #n2; nnormalize; napply (or2_intro2 (? = ?) (? ≠ ?) …);
+ nnormalize; #H; napply (asttype_destruct … H)
+ ##| ##3: #n2; nnormalize; napply (or2_intro2 (? = ?) (? ≠ ?) …);
+ nnormalize; #H; napply (asttype_destruct … H)
+ ##]
+ ##| ##2: #b1; #n1; #H; #y; ncases y;
+ ##[ ##2: #b2; #n2; nnormalize; napply (or2_elim (? = ?) (? ≠ ?) ? (H b2));
+ ##[ ##2: #H1; napply (or2_intro2 (? = ?) (? ≠ ?) …); nnormalize;
+ #H2; napply (H1 (asttype_destruct_array_array_1 … H2))
+ ##| ##1: #H1; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_nat n1 n2));
+ ##[ ##2: #H2; napply (or2_intro2 (? = ?) (? ≠ ?) …); nnormalize;
+ #H3; napply (H2 (asttype_destruct_array_array_2 … H3))
+ ##| ##1: #H2; nrewrite > H1; nrewrite > H2;
+ napply (or2_intro1 (? = ?) (? ≠ ?) (refl_eq …))
+ ##]
+ ##]
+ ##| ##1: #b2; nnormalize; napply (or2_intro2 (? = ?) (? ≠ ?) …);
+ nnormalize; #H1; napply (asttype_destruct … H1)
+ ##| ##3: #n2; nnormalize; napply (or2_intro2 (? = ?) (? ≠ ?) …);
+ nnormalize; #H1; napply (asttype_destruct … H1)
+ ##]
+ ##| ##3: #hh1; #H; #y; ncases y;
+ ##[ ##3: #n2; ncases n2;
+ ##[ ##1: #hh2; nnormalize; napply (or2_elim (? = ?) (? ≠ ?) ? (H hh2));
+ ##[ ##2: #H1; napply (or2_intro2 (? = ?) (? ≠ ?) …);
+ nnormalize; #H2;
+ napply (H1 (nelist_destruct_nil_nil … (asttype_destruct_struct_struct … H2)))
+ ##| ##1: #H1; nrewrite > H1; napply (or2_intro1 (? = ?) (? ≠ ?) (refl_eq …))
+ ##]
+ ##| ##2: #hh2; #tt2; nnormalize; napply (or2_intro2 (? = ?) (? ≠ ?) …);
+ nnormalize; #H1; napply (nelist_destruct_nil_cons ast_type … (asttype_destruct_struct_struct … H1))
+ ##]
+ ##| ##1: #b2; nnormalize; napply (or2_intro2 (? = ?) (? ≠ ?) …);
+ nnormalize; #H1; napply (asttype_destruct … H1)
+ ##| ##2: #b2; #n2; nnormalize; napply (or2_intro2 (? = ?) (? ≠ ?) …);
+ nnormalize; #H1; napply (asttype_destruct … H1)
+ ##]
+ ##| ##4: #hh1; #tt1; #H; #H1; #y; ncases y;
+ ##[ ##3: #n2; ncases n2;
+ ##[ ##2: #hh2; #tt2; nnormalize; napply (or2_elim (? = ?) (? ≠ ?) ? (H hh2));
+ ##[ ##2: #H2; napply (or2_intro2 (? = ?) (? ≠ ?) …);
+ nnormalize; #H3; napply (H2 (nelist_destruct_cons_cons_1 ast_type … (asttype_destruct_struct_struct … H3)))
+ ##| ##1: #H2; napply (or2_elim (? = ?) (? ≠ ?) ? (H1 (AST_TYPE_STRUCT tt2)));
+ ##[ ##2: #H3; napply (or2_intro2 (? = ?) (? ≠ ?) …);
+ nnormalize; #H4; napply H3;
+ nrewrite > (nelist_destruct_cons_cons_2 ast_type … (asttype_destruct_struct_struct … H4));
+ napply refl_eq
+ ##| ##1: #H3; nrewrite > H2; nrewrite > (asttype_destruct_struct_struct … H3);
+ napply (or2_intro1 (? = ?) (? ≠ ?) (refl_eq …))
+ ##]
+ ##]
+ ##| ##1: #hh2; nnormalize; napply (or2_intro2 (? = ?) (? ≠ ?) …);
+ nnormalize; #H2; napply (nelist_destruct_cons_nil ast_type … (asttype_destruct_struct_struct … H2))
+ ##]
+ ##| ##1: #b2; nnormalize; napply (or2_intro2 (? = ?) (? ≠ ?) …);
+ nnormalize; #H2; napply (asttype_destruct … H2)
+ ##| ##2: #b2; #n2; nnormalize; napply (or2_intro2 (? = ?) (? ≠ ?) …);
+ nnormalize; #H2; napply (asttype_destruct … H2)
+ ##]
+ ##]
+nqed.
+
+nlemma neqasttype_to_neq : ∀t1,t2.eq_asttype t1 t2 = false → t1 ≠ t2.
+ #t1;
+ napply (ast_type_index … t1);
+ ##[ ##1: #b1; #t2; ncases t2;
+ ##[ ##1: #b2; nchange with ((eq_astbasetype b1 b2 = false) → ?); #H;
+ nnormalize; #H1; napply (neqastbasetype_to_neq b1 b2 H);
+ napply (asttype_destruct_base_base … H1)
+ ##| ##2: #b2; #n2; nnormalize; #H; #H1; napply (asttype_destruct … H1)
+ ##| ##3: #n2; nnormalize; #H; #H1; napply (asttype_destruct … H1)
+ ##]
+ ##| ##2: #b1; #n1; #H; #t2; ncases t2;
+ ##[ ##2: #b2; #n2; nchange with ((((eq_asttype b1 b2)⊗(eq_nat n1 n2)) = false) → ?); #H1;
+ napply (or2_elim ??? (andb_false2 … H1));
+ ##[ ##1: #H2; nnormalize; #H3; napply (H b2 H2); napply (asttype_destruct_array_array_1 … H3)
+ ##| ##2: #H2; nnormalize; #H3; napply (neqnat_to_neq n1 n2 H2); napply (asttype_destruct_array_array_2 … H3)
+ ##]
+ ##| ##1: #b2; nnormalize; #H1; #H2; napply (asttype_destruct … H2)
+ ##| ##3: #n2; nnormalize; #H1; #H2; napply (asttype_destruct … H2)
+ ##]
+ ##| ##3: #hh1; #H; #t2; ncases t2;
+ ##[ ##3: #n2; ncases n2;
+ ##[ ##1: #hh2; nchange with ((eq_asttype hh1 hh2 = false) → ?); #H1;
+ nnormalize; #H2; napply (H hh2 H1);
+ napply (nelist_destruct_nil_nil ast_type … (asttype_destruct_struct_struct … H2))
+ ##| ##2: #hh2; #tt2; nnormalize; #H1; #H2;
+ napply (nelist_destruct_nil_cons ast_type … (asttype_destruct_struct_struct … H2))
+ ##]
+ ##| ##1: #b2; nnormalize; #H1; #H2; napply (asttype_destruct … H2)
+ ##| ##2: #b2; #n2; nnormalize; #H1; #H2; napply (asttype_destruct … H2)
+ ##]
+ ##| ##4: #hh1; #tt1; #H; #H1; #t2; ncases t2;
+ ##[ ##3: #n2; ncases n2;
+ ##[ ##2: #hh2; #tt2; nchange with ((((eq_asttype hh1 hh2)⊗(bfold_right_neList2 ?? tt1 tt2)) = false) → ?);
+ #H2; napply (or2_elim ??? (andb_false2 … H2));
+ ##[ ##1: #H3; nnormalize; #H4; napply (H hh2 H3);
+ napply (nelist_destruct_cons_cons_1 ast_type … (asttype_destruct_struct_struct … H4))
+ ##| ##2: #H3; nnormalize; #H4; napply (H1 (AST_TYPE_STRUCT tt2) H3);
+ nrewrite > (nelist_destruct_cons_cons_2 ast_type … (asttype_destruct_struct_struct … H4));
+ napply refl_eq
+ ##]
+ ##| ##1: #hh2; nnormalize; #H2; #H3;
+ napply (nelist_destruct_cons_nil ast_type … (asttype_destruct_struct_struct … H3))
+ ##]
+ ##| ##1: #b2; nnormalize; #H2; #H3; napply (asttype_destruct … H3)
+ ##| ##2: #b2; #n2; nnormalize; #H2; #H3; napply (asttype_destruct … H3)
+ ##]
+ ##]
+nqed.
+
+nlemma neq_to_neqasttype : ∀t1,t2.t1 ≠ t2 → eq_asttype t1 t2 = false.
+ #t1;
+ napply (ast_type_index … t1);
+ ##[ ##1: #b1; #t2; ncases t2;
+ ##[ ##1: #b2; nchange with (? → (eq_astbasetype b1 b2 = false)); #H;
+ napply (neq_to_neqastbasetype b1 b2 ?); nnormalize; #H1;
+ napply H; nrewrite > H1; napply refl_eq
+ ##| ##2: #b2; #n2; nnormalize; #H; napply refl_eq
+ ##| ##3: #n2; nnormalize; #H; napply refl_eq
+ ##]
+ ##| ##2: #b1; #n1; #H; #t2; ncases t2;
+ ##[ ##2: #b2; #n2; nchange with (? → ((eq_asttype b1 b2)⊗(eq_nat n1 n2)) = false); #H1;
+ napply (or2_elim ??? (decidable_asttype b1 b2));
+ ##[ ##2: #H2; nrewrite > (H b2 H2); nnormalize; napply refl_eq
+ ##| ##1: #H2; napply (or2_elim ??? (decidable_nat n1 n2))
+ ##[ ##2: #H3; nrewrite > (neq_to_neqnat n1 n2 H3);
+ nrewrite > (andb_false2_2 (eq_asttype b1 b2));
+ napply refl_eq
+ ##| ##1: #H3; nrewrite > H2 in H1:(%); nrewrite > H3; #H1;
+ nelim (H1 (refl_eq …))
+ ##]
+ ##]
+ ##| ##1: #b2; nnormalize; #H1; napply refl_eq
+ ##| ##3: #n2; nnormalize; #H1; napply refl_eq
+ ##]
+ ##| ##3: #hh1; #H; #t2; ncases t2;
+ ##[ ##3: #n2; ncases n2;
+ ##[ ##1: #hh2; nchange with (? → (eq_asttype hh1 hh2 = false)); #H1;
+ napply (H hh2); nnormalize; #H2; nrewrite > H2 in H1:(%);
+ nnormalize; #H1; napply (H1 (refl_eq …))
+ ##| ##2: #hh2; #tt2; nnormalize; #H1; napply refl_eq
+ ##]
+ ##| ##1: #b2; nnormalize; #H1; napply refl_eq
+ ##| ##2: #b2; #n2; nnormalize; #H1; napply refl_eq
+ ##]
+ ##| ##4: #hh1; #tt1; #H; #H1; #t2; ncases t2;
+ ##[ ##3: #n2; ncases n2;
+ ##[ ##2: #hh2; #tt2; nchange with (? → ((eq_asttype hh1 hh2)⊗(bfold_right_neList2 ?? tt1 tt2)) = false);
+ #H2; napply (or2_elim (hh1 ≠ hh2) (tt1 ≠ tt2) …);
+ ##[ ##2: #H3; nrewrite > (H hh2 H3); nnormalize; napply refl_eq
+ ##| ##3: #H3; nchange with (((eq_asttype hh1 hh2)⊗(eq_asttype (AST_TYPE_STRUCT tt1) (AST_TYPE_STRUCT tt2))) = false);
+ nrewrite > (H1 (AST_TYPE_STRUCT tt2) ?);
+ ##[ ##1: nrewrite > (andb_false2_2 (eq_asttype hh1 hh2)); napply refl_eq
+ ##| ##2: nnormalize; #H4; napply (H3 (asttype_destruct_struct_struct … H4))
+ ##]
+ ##| ##1: napply (or2_elim ??? (decidable_asttype hh1 hh2));
+ ##[ ##2: #H3; napply (or2_intro1 … H3)
+ ##| ##1: #H3; napply (or2_elim ??? (decidable_nelist ast_type decidable_asttype tt1 tt2));
+ ##[ ##2: #H4; napply (or2_intro2 … H4)
+ ##| ##1: #H4; nrewrite > H3 in H2:(%); nrewrite > H4; #H2;
+ nelim (H2 (refl_eq …))
+ ##]
+ ##]
+ ##]
+ ##| ##1: #hh2; nnormalize; #H2; napply refl_eq
+ ##]
+ ##| ##1: #b2; nnormalize; #H2; napply refl_eq
+ ##| ##2: #b2; #n2; nnormalize; #H2; napply refl_eq
+ ##]
+ ##]
+nqed.
+
nlemma isbastbasetype_to_isastbasetype : ∀ast.isb_ast_base_type ast = true → is_ast_base_type ast.
#ast;
ncases ast;
napply refl_eq.
nqed.
-nlemma symmetric_eqaluHC05 : symmetricT alu_HC05 bool eq_alu_HC05.
+nlemma symmetric_eqaluHC05 : symmetricT alu_HC05 bool eq_aluHC05.
#alu1; #alu2;
ncases alu1;
#x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12; #x13;
napply refl_eq.
nqed.
-nlemma eqaluHC05_to_eq : ∀alu1,alu2.eq_alu_HC05 alu1 alu2 = true → alu1 = alu2.
+nlemma eqaluHC05_to_eq : ∀alu1,alu2.eq_aluHC05 alu1 alu2 = true → alu1 = alu2.
#alu1; #alu2;
ncases alu1;
#x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12; #x13;
napply refl_eq.
nqed.
-nlemma eq_to_eqaluHC05 : ∀alu1,alu2.alu1 = alu2 → eq_alu_HC05 alu1 alu2 = true.
+nlemma eq_to_eqaluHC05 : ∀alu1,alu2.alu1 = alu2 → eq_aluHC05 alu1 alu2 = true.
#alu1; #alu2;
ncases alu1;
#x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12; #x13;
napply refl_eq.
nqed.
+nlemma decidable_aluHC05_aux1
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,x13,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12,y13.
+ (x1 ≠ y1) →
+ (mk_alu_HC05 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13) ≠
+ (mk_alu_HC05 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12 y13).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12; #x13;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12; #y13;
+ nnormalize; #H; #H1;
+ napply (H (aluHC05_destruct_1 … H1)).
+nqed.
+
+nlemma decidable_aluHC05_aux2
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,x13,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12,y13.
+ (x2 ≠ y2) →
+ (mk_alu_HC05 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13) ≠
+ (mk_alu_HC05 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12 y13).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12; #x13;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12; #y13;
+ nnormalize; #H; #H1;
+ napply (H (aluHC05_destruct_2 … H1)).
+nqed.
+
+nlemma decidable_aluHC05_aux3
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,x13,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12,y13.
+ (x3 ≠ y3) →
+ (mk_alu_HC05 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13) ≠
+ (mk_alu_HC05 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12 y13).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12; #x13;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12; #y13;
+ nnormalize; #H; #H1;
+ napply (H (aluHC05_destruct_3 … H1)).
+nqed.
+
+nlemma decidable_aluHC05_aux4
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,x13,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12,y13.
+ (x4 ≠ y4) →
+ (mk_alu_HC05 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13) ≠
+ (mk_alu_HC05 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12 y13).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12; #x13;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12; #y13;
+ nnormalize; #H; #H1;
+ napply (H (aluHC05_destruct_4 … H1)).
+nqed.
+
+nlemma decidable_aluHC05_aux5
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,x13,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12,y13.
+ (x5 ≠ y5) →
+ (mk_alu_HC05 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13) ≠
+ (mk_alu_HC05 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12 y13).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12; #x13;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12; #y13;
+ nnormalize; #H; #H1;
+ napply (H (aluHC05_destruct_5 … H1)).
+nqed.
+
+nlemma decidable_aluHC05_aux6
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,x13,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12,y13.
+ (x6 ≠ y6) →
+ (mk_alu_HC05 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13) ≠
+ (mk_alu_HC05 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12 y13).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12; #x13;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12; #y13;
+ nnormalize; #H; #H1;
+ napply (H (aluHC05_destruct_6 … H1)).
+nqed.
+
+nlemma decidable_aluHC05_aux7
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,x13,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12,y13.
+ (x7 ≠ y7) →
+ (mk_alu_HC05 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13) ≠
+ (mk_alu_HC05 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12 y13).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12; #x13;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12; #y13;
+ nnormalize; #H; #H1;
+ napply (H (aluHC05_destruct_7 … H1)).
+nqed.
+
+nlemma decidable_aluHC05_aux8
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,x13,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12,y13.
+ (x8 ≠ y8) →
+ (mk_alu_HC05 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13) ≠
+ (mk_alu_HC05 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12 y13).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12; #x13;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12; #y13;
+ nnormalize; #H; #H1;
+ napply (H (aluHC05_destruct_8 … H1)).
+nqed.
+
+nlemma decidable_aluHC05_aux9
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,x13,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12,y13.
+ (x9 ≠ y9) →
+ (mk_alu_HC05 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13) ≠
+ (mk_alu_HC05 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12 y13).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12; #x13;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12; #y13;
+ nnormalize; #H; #H1;
+ napply (H (aluHC05_destruct_9 … H1)).
+nqed.
+
+nlemma decidable_aluHC05_aux10
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,x13,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12,y13.
+ (x10 ≠ y10) →
+ (mk_alu_HC05 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13) ≠
+ (mk_alu_HC05 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12 y13).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12; #x13;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12; #y13;
+ nnormalize; #H; #H1;
+ napply (H (aluHC05_destruct_10 … H1)).
+nqed.
+
+nlemma decidable_aluHC05_aux11
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,x13,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12,y13.
+ (x11 ≠ y11) →
+ (mk_alu_HC05 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13) ≠
+ (mk_alu_HC05 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12 y13).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12; #x13;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12; #y13;
+ nnormalize; #H; #H1;
+ napply (H (aluHC05_destruct_11 … H1)).
+nqed.
+
+nlemma decidable_aluHC05_aux12
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,x13,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12,y13.
+ (x12 ≠ y12) →
+ (mk_alu_HC05 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13) ≠
+ (mk_alu_HC05 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12 y13).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12; #x13;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12; #y13;
+ nnormalize; #H; #H1;
+ napply (H (aluHC05_destruct_12 … H1)).
+nqed.
+
+nlemma decidable_aluHC05_aux13
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,x13,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12,y13.
+ (x13 ≠ y13) →
+ (mk_alu_HC05 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13) ≠
+ (mk_alu_HC05 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12 y13).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12; #x13;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12; #y13;
+ nnormalize; #H; #H1;
+ napply (H (aluHC05_destruct_13 … H1)).
+nqed.
+
+nlemma decidable_aluHC05 : ∀x,y:alu_HC05.decidable (x = y).
+ #x; nelim x; #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12; #x13;
+ #y; nelim y; #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12; #y13;
+ nnormalize;
+ napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_b8 x1 y1) …);
+ ##[ ##2: #H; napply (or2_intro2 … (decidable_aluHC05_aux1 … H))
+ ##| ##1: #H; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_b8 x2 y2) …);
+ ##[ ##2: #H1; napply (or2_intro2 … (decidable_aluHC05_aux2 … H1))
+ ##| ##1: #H1; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_w16 x3 y3) …);
+ ##[ ##2: #H2; napply (or2_intro2 … (decidable_aluHC05_aux3 … H2))
+ ##| ##1: #H2; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_w16 x4 y4) …);
+ ##[ ##2: #H3; napply (or2_intro2 … (decidable_aluHC05_aux4 … H3))
+ ##| ##1: #H3; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_w16 x5 y5) …);
+ ##[ ##2: #H4; napply (or2_intro2 … (decidable_aluHC05_aux5 … H4))
+ ##| ##1: #H4; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_w16 x6 y6) …);
+ ##[ ##2: #H5; napply (or2_intro2 … (decidable_aluHC05_aux6 … H5))
+ ##| ##1: #H5; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_w16 x7 y7) …);
+ ##[ ##2: #H6; napply (or2_intro2 … (decidable_aluHC05_aux7 … H6))
+ ##| ##1: #H6; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_bool x8 y8) …);
+ ##[ ##2: #H7; napply (or2_intro2 … (decidable_aluHC05_aux8 … H7))
+ ##| ##1: #H7; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_bool x9 y9) …);
+ ##[ ##2: #H8; napply (or2_intro2 … (decidable_aluHC05_aux9 … H8))
+ ##| ##1: #H8; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_bool x10 y10) …);
+ ##[ ##2: #H9; napply (or2_intro2 … (decidable_aluHC05_aux10 … H9))
+ ##| ##1: #H9; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_bool x11 y11) …);
+ ##[ ##2: #H10; napply (or2_intro2 … (decidable_aluHC05_aux11 … H10))
+ ##| ##1: #H10; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_bool x12 y12) …);
+ ##[ ##2: #H11; napply (or2_intro2 … (decidable_aluHC05_aux12 … H11))
+ ##| ##1: #H11; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_bool x13 y13) …);
+ ##[ ##2: #H12; napply (or2_intro2 … (decidable_aluHC05_aux13 … H12))
+ ##| ##1: #H12; nrewrite > H; nrewrite > H1; nrewrite > H2; nrewrite > H3;
+ nrewrite > H4; nrewrite > H5; nrewrite > H6; nrewrite > H7;
+ nrewrite > H8; nrewrite > H9; nrewrite > H10; nrewrite > H11;
+ nrewrite > H12; napply (or2_intro1 (? = ?) (? ≠ ?) (refl_eq …))
+ ##]
+ ##]
+ ##]
+ ##]
+ ##]
+ ##]
+ ##]
+ ##]
+ ##]
+ ##]
+ ##]
+ ##]
+ ##]
+nqed.
+
nlemma aluHC08_destruct_1 :
∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12.
mk_alu_HC08 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 = mk_alu_HC08 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12 →
napply refl_eq.
nqed.
-nlemma symmetric_eqaluHC08 : symmetricT alu_HC08 bool eq_alu_HC08.
+nlemma symmetric_eqaluHC08 : symmetricT alu_HC08 bool eq_aluHC08.
#alu1; #alu2;
ncases alu1;
#x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12;
napply refl_eq.
nqed.
-nlemma eqaluHC08_to_eq : ∀alu1,alu2.eq_alu_HC08 alu1 alu2 = true → alu1 = alu2.
+nlemma eqaluHC08_to_eq : ∀alu1,alu2.eq_aluHC08 alu1 alu2 = true → alu1 = alu2.
#alu1; #alu2;
ncases alu1;
#x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12;
napply refl_eq.
nqed.
-nlemma eq_to_eqaluHC08 : ∀alu1,alu2.alu1 = alu2 → eq_alu_HC08 alu1 alu2 = true.
+nlemma eq_to_eqaluHC08 : ∀alu1,alu2.alu1 = alu2 → eq_aluHC08 alu1 alu2 = true.
#alu1; #alu2;
ncases alu1;
#x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12;
napply refl_eq.
nqed.
+nlemma decidable_aluHC08_aux1
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12.
+ (x1 ≠ y1) →
+ (mk_alu_HC08 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12) ≠
+ (mk_alu_HC08 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12;
+ nnormalize; #H; #H1;
+ napply (H (aluHC08_destruct_1 … H1)).
+nqed.
+
+nlemma decidable_aluHC08_aux2
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12.
+ (x2 ≠ y2) →
+ (mk_alu_HC08 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12) ≠
+ (mk_alu_HC08 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12;
+ nnormalize; #H; #H1;
+ napply (H (aluHC08_destruct_2 … H1)).
+nqed.
+
+nlemma decidable_aluHC08_aux3
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12.
+ (x3 ≠ y3) →
+ (mk_alu_HC08 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12) ≠
+ (mk_alu_HC08 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12;
+ nnormalize; #H; #H1;
+ napply (H (aluHC08_destruct_3 … H1)).
+nqed.
+
+nlemma decidable_aluHC08_aux4
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12.
+ (x4 ≠ y4) →
+ (mk_alu_HC08 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12) ≠
+ (mk_alu_HC08 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12;
+ nnormalize; #H; #H1;
+ napply (H (aluHC08_destruct_4 … H1)).
+nqed.
+
+nlemma decidable_aluHC08_aux5
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12.
+ (x5 ≠ y5) →
+ (mk_alu_HC08 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12) ≠
+ (mk_alu_HC08 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12;
+ nnormalize; #H; #H1;
+ napply (H (aluHC08_destruct_5 … H1)).
+nqed.
+
+nlemma decidable_aluHC08_aux6
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12.
+ (x6 ≠ y6) →
+ (mk_alu_HC08 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12) ≠
+ (mk_alu_HC08 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12;
+ nnormalize; #H; #H1;
+ napply (H (aluHC08_destruct_6 … H1)).
+nqed.
+
+nlemma decidable_aluHC08_aux7
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12.
+ (x7 ≠ y7) →
+ (mk_alu_HC08 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12) ≠
+ (mk_alu_HC08 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12;
+ nnormalize; #H; #H1;
+ napply (H (aluHC08_destruct_7 … H1)).
+nqed.
+
+nlemma decidable_aluHC08_aux8
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12.
+ (x8 ≠ y8) →
+ (mk_alu_HC08 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12) ≠
+ (mk_alu_HC08 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12;
+ nnormalize; #H; #H1;
+ napply (H (aluHC08_destruct_8 … H1)).
+nqed.
+
+nlemma decidable_aluHC08_aux9
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12.
+ (x9 ≠ y9) →
+ (mk_alu_HC08 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12) ≠
+ (mk_alu_HC08 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12;
+ nnormalize; #H; #H1;
+ napply (H (aluHC08_destruct_9 … H1)).
+nqed.
+
+nlemma decidable_aluHC08_aux10
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12.
+ (x10 ≠ y10) →
+ (mk_alu_HC08 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12) ≠
+ (mk_alu_HC08 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12;
+ nnormalize; #H; #H1;
+ napply (H (aluHC08_destruct_10 … H1)).
+nqed.
+
+nlemma decidable_aluHC08_aux11
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12.
+ (x11 ≠ y11) →
+ (mk_alu_HC08 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12) ≠
+ (mk_alu_HC08 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12;
+ nnormalize; #H; #H1;
+ napply (H (aluHC08_destruct_11 … H1)).
+nqed.
+
+nlemma decidable_aluHC08_aux12
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,x9,x10,x11,x12,y1,y2,y3,y4,y5,y6,y7,y8,y9,y10,y11,y12.
+ (x12 ≠ y12) →
+ (mk_alu_HC08 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12) ≠
+ (mk_alu_HC08 y1 y2 y3 y4 y5 y6 y7 y8 y9 y10 y11 y12).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12;
+ nnormalize; #H; #H1;
+ napply (H (aluHC08_destruct_12 … H1)).
+nqed.
+
+nlemma decidable_aluHC08 : ∀x,y:alu_HC08.decidable (x = y).
+ #x; nelim x; #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8; #x9; #x10; #x11; #x12;
+ #y; nelim y; #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8; #y9; #y10; #y11; #y12;
+ nnormalize;
+ napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_b8 x1 y1) …);
+ ##[ ##2: #H; napply (or2_intro2 … (decidable_aluHC08_aux1 … H))
+ ##| ##1: #H; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_b8 x2 y2) …);
+ ##[ ##2: #H1; napply (or2_intro2 … (decidable_aluHC08_aux2 … H1))
+ ##| ##1: #H1; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_b8 x3 y3) …);
+ ##[ ##2: #H2; napply (or2_intro2 … (decidable_aluHC08_aux3 … H2))
+ ##| ##1: #H2; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_w16 x4 y4) …);
+ ##[ ##2: #H3; napply (or2_intro2 … (decidable_aluHC08_aux4 … H3))
+ ##| ##1: #H3; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_w16 x5 y5) …);
+ ##[ ##2: #H4; napply (or2_intro2 … (decidable_aluHC08_aux5 … H4))
+ ##| ##1: #H4; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_bool x6 y6) …);
+ ##[ ##2: #H5; napply (or2_intro2 … (decidable_aluHC08_aux6 … H5))
+ ##| ##1: #H5; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_bool x7 y7) …);
+ ##[ ##2: #H6; napply (or2_intro2 … (decidable_aluHC08_aux7 … H6))
+ ##| ##1: #H6; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_bool x8 y8) …);
+ ##[ ##2: #H7; napply (or2_intro2 … (decidable_aluHC08_aux8 … H7))
+ ##| ##1: #H7; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_bool x9 y9) …);
+ ##[ ##2: #H8; napply (or2_intro2 … (decidable_aluHC08_aux9 … H8))
+ ##| ##1: #H8; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_bool x10 y10) …);
+ ##[ ##2: #H9; napply (or2_intro2 … (decidable_aluHC08_aux10 … H9))
+ ##| ##1: #H9; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_bool x11 y11) …);
+ ##[ ##2: #H10; napply (or2_intro2 … (decidable_aluHC08_aux11 … H10))
+ ##| ##1: #H10; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_bool x12 y12) …);
+ ##[ ##2: #H11; napply (or2_intro2 … (decidable_aluHC08_aux12 … H11))
+ ##| ##1: #H11; nrewrite > H; nrewrite > H1; nrewrite > H2; nrewrite > H3;
+ nrewrite > H4; nrewrite > H5; nrewrite > H6; nrewrite > H7;
+ nrewrite > H8; nrewrite > H9; nrewrite > H10; nrewrite > H11;
+ napply (or2_intro1 (? = ?) (? ≠ ?) (refl_eq …))
+ ##]
+ ##]
+ ##]
+ ##]
+ ##]
+ ##]
+ ##]
+ ##]
+ ##]
+ ##]
+ ##]
+ ##]
+nqed.
+
nlemma aluRS08_destruct_1 :
∀x1,x2,x3,x4,x5,x6,x7,x8,y1,y2,y3,y4,y5,y6,y7,y8.
mk_alu_RS08 x1 x2 x3 x4 x5 x6 x7 x8 = mk_alu_RS08 y1 y2 y3 y4 y5 y6 y7 y8 →
napply refl_eq.
nqed.
-nlemma symmetric_eqaluRS08 : symmetricT alu_RS08 bool eq_alu_RS08.
+nlemma symmetric_eqaluRS08 : symmetricT alu_RS08 bool eq_aluRS08.
#alu1; #alu2;
ncases alu1;
#x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8;
napply refl_eq.
nqed.
-nlemma eqaluRS08_to_eq : ∀alu1,alu2.eq_alu_RS08 alu1 alu2 = true → alu1 = alu2.
+nlemma eqaluRS08_to_eq : ∀alu1,alu2.eq_aluRS08 alu1 alu2 = true → alu1 = alu2.
#alu1; #alu2;
ncases alu1;
#x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8;
napply refl_eq.
nqed.
-nlemma eq_to_eqaluRS08 : ∀alu1,alu2.alu1 = alu2 → eq_alu_RS08 alu1 alu2 = true.
+nlemma eq_to_eqaluRS08 : ∀alu1,alu2.alu1 = alu2 → eq_aluRS08 alu1 alu2 = true.
#alu1; #alu2;
ncases alu1;
#x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8;
napply refl_eq.
nqed.
+nlemma decidable_aluRS08_aux1
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,y1,y2,y3,y4,y5,y6,y7,y8.
+ (x1 ≠ y1) →
+ (mk_alu_RS08 x1 x2 x3 x4 x5 x6 x7 x8) ≠
+ (mk_alu_RS08 y1 y2 y3 y4 y5 y6 y7 y8).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8;
+ nnormalize; #H; #H1;
+ napply (H (aluRS08_destruct_1 … H1)).
+nqed.
+
+nlemma decidable_aluRS08_aux2
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,y1,y2,y3,y4,y5,y6,y7,y8.
+ (x2 ≠ y2) →
+ (mk_alu_RS08 x1 x2 x3 x4 x5 x6 x7 x8) ≠
+ (mk_alu_RS08 y1 y2 y3 y4 y5 y6 y7 y8).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8;
+ nnormalize; #H; #H1;
+ napply (H (aluRS08_destruct_2 … H1)).
+nqed.
+
+nlemma decidable_aluRS08_aux3
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,y1,y2,y3,y4,y5,y6,y7,y8.
+ (x3 ≠ y3) →
+ (mk_alu_RS08 x1 x2 x3 x4 x5 x6 x7 x8) ≠
+ (mk_alu_RS08 y1 y2 y3 y4 y5 y6 y7 y8).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8;
+ nnormalize; #H; #H1;
+ napply (H (aluRS08_destruct_3 … H1)).
+nqed.
+
+nlemma decidable_aluRS08_aux4
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,y1,y2,y3,y4,y5,y6,y7,y8.
+ (x4 ≠ y4) →
+ (mk_alu_RS08 x1 x2 x3 x4 x5 x6 x7 x8) ≠
+ (mk_alu_RS08 y1 y2 y3 y4 y5 y6 y7 y8).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8;
+ nnormalize; #H; #H1;
+ napply (H (aluRS08_destruct_4 … H1)).
+nqed.
+
+nlemma decidable_aluRS08_aux5
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,y1,y2,y3,y4,y5,y6,y7,y8.
+ (x5 ≠ y5) →
+ (mk_alu_RS08 x1 x2 x3 x4 x5 x6 x7 x8) ≠
+ (mk_alu_RS08 y1 y2 y3 y4 y5 y6 y7 y8).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8;
+ nnormalize; #H; #H1;
+ napply (H (aluRS08_destruct_5 … H1)).
+nqed.
+
+nlemma decidable_aluRS08_aux6
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,y1,y2,y3,y4,y5,y6,y7,y8.
+ (x6 ≠ y6) →
+ (mk_alu_RS08 x1 x2 x3 x4 x5 x6 x7 x8) ≠
+ (mk_alu_RS08 y1 y2 y3 y4 y5 y6 y7 y8).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8;
+ nnormalize; #H; #H1;
+ napply (H (aluRS08_destruct_6 … H1)).
+nqed.
+
+nlemma decidable_aluRS08_aux7
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,y1,y2,y3,y4,y5,y6,y7,y8.
+ (x7 ≠ y7) →
+ (mk_alu_RS08 x1 x2 x3 x4 x5 x6 x7 x8) ≠
+ (mk_alu_RS08 y1 y2 y3 y4 y5 y6 y7 y8).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8;
+ nnormalize; #H; #H1;
+ napply (H (aluRS08_destruct_7 … H1)).
+nqed.
+
+nlemma decidable_aluRS08_aux8
+ : ∀x1,x2,x3,x4,x5,x6,x7,x8,y1,y2,y3,y4,y5,y6,y7,y8.
+ (x8 ≠ y8) →
+ (mk_alu_RS08 x1 x2 x3 x4 x5 x6 x7 x8) ≠
+ (mk_alu_RS08 y1 y2 y3 y4 y5 y6 y7 y8).
+ #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8;
+ #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8;
+ nnormalize; #H; #H1;
+ napply (H (aluRS08_destruct_8 … H1)).
+nqed.
+
+nlemma decidable_aluRS08 : ∀x,y:alu_RS08.decidable (x = y).
+ #x; nelim x; #x1; #x2; #x3; #x4; #x5; #x6; #x7; #x8;
+ #y; nelim y; #y1; #y2; #y3; #y4; #y5; #y6; #y7; #y8;
+ nnormalize;
+ napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_b8 x1 y1) …);
+ ##[ ##2: #H; napply (or2_intro2 … (decidable_aluRS08_aux1 … H))
+ ##| ##1: #H; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_w16 x2 y2) …);
+ ##[ ##2: #H1; napply (or2_intro2 … (decidable_aluRS08_aux2 … H1))
+ ##| ##1: #H1; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_w16 x3 y3) …);
+ ##[ ##2: #H2; napply (or2_intro2 … (decidable_aluRS08_aux3 … H2))
+ ##| ##1: #H2; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_w16 x4 y4) …);
+ ##[ ##2: #H3; napply (or2_intro2 … (decidable_aluRS08_aux4 … H3))
+ ##| ##1: #H3; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_b8 x5 y5) …);
+ ##[ ##2: #H4; napply (or2_intro2 … (decidable_aluRS08_aux5 … H4))
+ ##| ##1: #H4; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_b8 x6 y6) …);
+ ##[ ##2: #H5; napply (or2_intro2 … (decidable_aluRS08_aux6 … H5))
+ ##| ##1: #H5; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_bool x7 y7) …);
+ ##[ ##2: #H6; napply (or2_intro2 … (decidable_aluRS08_aux7 … H6))
+ ##| ##1: #H6; napply (or2_elim (? = ?) (? ≠ ?) ? (decidable_bool x8 y8) …);
+ ##[ ##2: #H7; napply (or2_intro2 … (decidable_aluRS08_aux8 … H7))
+ ##| ##1: #H7; nrewrite > H; nrewrite > H1; nrewrite > H2; nrewrite > H3;
+ nrewrite > H4; nrewrite > H5; nrewrite > H6; nrewrite > H7;
+ napply (or2_intro1 (? = ?) (? ≠ ?) (refl_eq …))
+ ##]
+ ##]
+ ##]
+ ##]
+ ##]
+ ##]
+ ##]
+ ##]
+nqed.
+
nlemma symmetric_eqclk : ∀mcu,clk1,clk2.eq_clk mcu clk1 clk2 = eq_clk mcu clk2 clk1.
#mcu; #clk1; #clk2;
ncases clk1;
napply refl_eq.
nqed.
-nlemma symmetric_eqstatus :
+nlemma symmetric_eqanystatus :
∀addrl:list word16.∀m:mcu_type.∀t:memory_impl.∀s1,s2:any_status m t.
- eq_status m t s1 s2 addrl = eq_status m t s2 s1 addrl.
+ eq_anystatus m t s1 s2 addrl = eq_anystatus m t s2 s1 addrl.
#addrl; #m;
ncases m; #t; #s1;
##[ ##1: ncases s1; #x1; #x2; #x3; #x4;
#s2; ncases s2; #y1; #y2; #y3; #y4;
nchange with (
- ((eq_alu_HC05 x1 y1) ⊗ (forall_memory_ranged t x3 y3 x2 y2 addrl) ⊗ (eq_clk HC05 x4 y4)) =
- ((eq_alu_HC05 y1 x1) ⊗ (forall_memory_ranged t y3 x3 y2 x2 addrl) ⊗ (eq_clk HC05 y4 x4)));
+ ((eq_aluHC05 x1 y1) ⊗ (forall_memory_ranged t x3 y3 x2 y2 addrl) ⊗ (eq_clk HC05 x4 y4)) =
+ ((eq_aluHC05 y1 x1) ⊗ (forall_memory_ranged t y3 x3 y2 x2 addrl) ⊗ (eq_clk HC05 y4 x4)));
nrewrite > (symmetric_eqaluHC05 x1 y1)
##| ##2,3: ncases s1; #x1; #x2; #x3; #x4;
#s2; ncases s2; #y1; #y2; #y3; #y4;
nchange with (
- ((eq_alu_HC08 x1 y1) ⊗ (forall_memory_ranged t x3 y3 x2 y2 addrl) ⊗ (eq_clk ? x4 y4)) =
- ((eq_alu_HC08 y1 x1) ⊗ (forall_memory_ranged t y3 x3 y2 x2 addrl) ⊗ (eq_clk ? y4 x4)));
+ ((eq_aluHC08 x1 y1) ⊗ (forall_memory_ranged t x3 y3 x2 y2 addrl) ⊗ (eq_clk ? x4 y4)) =
+ ((eq_aluHC08 y1 x1) ⊗ (forall_memory_ranged t y3 x3 y2 x2 addrl) ⊗ (eq_clk ? y4 x4)));
nrewrite > (symmetric_eqaluHC08 x1 y1)
##| ##4: ncases s1; #x1; #x2; #x3; #x4;
#s2; ncases s2; #y1; #y2; #y3; #y4;
nchange with (
- ((eq_alu_RS08 x1 y1) ⊗ (forall_memory_ranged t x3 y3 x2 y2 addrl) ⊗ (eq_clk RS08 x4 y4)) =
- ((eq_alu_RS08 y1 x1) ⊗ (forall_memory_ranged t y3 x3 y2 x2 addrl) ⊗ (eq_clk RS08 y4 x4)));
+ ((eq_aluRS08 x1 y1) ⊗ (forall_memory_ranged t x3 y3 x2 y2 addrl) ⊗ (eq_clk RS08 x4 y4)) =
+ ((eq_aluRS08 y1 x1) ⊗ (forall_memory_ranged t y3 x3 y2 x2 addrl) ⊗ (eq_clk RS08 y4 x4)));
nrewrite > (symmetric_eqaluRS08 x1 y1)
##]
nrewrite > (symmetric_forallmemoryranged t x3 y3 x2 y2 addrl);
napply refl_eq.
nqed.
-nlemma eqstatus_to_eq :
+nlemma eqanystatus_to_eq :
∀addrl:list word16.∀m:mcu_type.∀t:memory_impl.∀s1,s2:any_status m t.
- (eq_status m t s1 s2 addrl = true) →
- ((alu m t s1 = alu m t s2) ∧
- (clk_desc m t s1 = clk_desc m t s2) ∧
- ((forall_memory_ranged t (chk_desc m t s1) (chk_desc m t s2)
- (mem_desc m t s1) (mem_desc m t s2) addrl) = true)).
+ (eq_anystatus m t s1 s2 addrl = true) →
+ And3 (alu m t s1 = alu m t s2)
+ (clk_desc m t s1 = clk_desc m t s2)
+ ((forall_memory_ranged t (chk_desc m t s1) (chk_desc m t s2)
+ (mem_desc m t s1) (mem_desc m t s2) addrl) = true).
#addrl; #m; #t;
ncases m; #s1;
##[ ##1: ncases s1; #x1; #x2; #x3; #x4;
#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);
+ ((eq_aluHC05 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)))
##| ##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);
+ ((eq_aluHC08 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)))
##| ##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);
+ ((eq_aluRS08 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)))
##]
- nchange with ((y1 = y1) ∧ (x4 = y4) ∧ (forall_memory_ranged t x3 y3 x2 y2 addrl = true));
+ nchange with (And3 (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 ? y1) (refl_eq ? y4)) (refl_eq ? true)).
+ napply (conj3 … (refl_eq ? y1) (refl_eq ? y4) (refl_eq ? true)).
nqed.
-nlemma eq_to_eqstatus_strong :
+nlemma eq_to_eqanystatus_strong :
∀addrl:list word16.∀m:mcu_type.∀t:memory_impl.∀s1,s2:any_status m t.
- s1 = s2 → (eq_status m t s1 s2 addrl = true).
+ s1 = s2 → (eq_anystatus m t s1 s2 addrl = true).
#addrl; #m; #t;
ncases m;
##[ ##1: #s1; ncases s1; #x1; #x2; #x3; #x4;
#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);
+ ((eq_aluHC05 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 …))
##| ##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);
+ ((eq_aluHC08 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 …))
##| ##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);
+ ((eq_aluRS08 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 …))
##]
##]
nqed.
-nlemma eq_to_eqstatus_weak :
+nlemma eq_to_eqanystatus_weak :
∀addrl:list word16.∀m:mcu_type.∀t:memory_impl.∀s1,s2:any_status m t.
(alu m t s1 = alu m t s2) →
(clk_desc m t s1 = clk_desc m t s2) →
((forall_memory_ranged t (chk_desc m t s1) (chk_desc m t s2)
(mem_desc m t s1) (mem_desc m t s2) addrl) = true) →
- (eq_status m t s1 s2 addrl = true).
+ (eq_anystatus m t s1 s2 addrl = true).
#addrl; #m; #t;
ncases m;
##[ ##1: #s1; ncases s1; #x1; #x2; #x3; #x4;
#s2; ncases s2; #y1; #y2; #y3; #y4; #H; #H1; #H2;
- nchange with (((eq_alu_HC05 x1 y1) ⊗ (forall_memory_ranged t x3 y3 x2 y2 addrl) ⊗ (eq_clk ? x4 y4)) = true);
+ nchange with (((eq_aluHC05 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 …))
##| ##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 with (((eq_aluHC08 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 …))
##| ##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 with (((eq_aluRS08 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 …))