]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_assembly/compiler/ast_type_lemmas.ma
mod change (-x)
[helm.git] / helm / software / matita / contribs / ng_assembly / compiler / ast_type_lemmas.ma
old mode 100755 (executable)
new mode 100644 (file)
index 8b45497..4bc55e2
@@ -16,7 +16,7 @@
 (*                          Progetto FreeScale                            *)
 (*                                                                        *)
 (*   Sviluppato da: Ing. Cosimo Oliboni, oliboni@cs.unibo.it              *)
-(*   Ultima modifica: 05/08/2009                                          *)
+(*   Sviluppo: 2008-2010                                                  *)
 (*                                                                        *)
 (* ********************************************************************** *)
 
@@ -27,53 +27,65 @@ include "common/list_utility_lemmas.ma".
 (* dimensioni degli elementi *)
 (* ************************* *)
 
+(*
 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.
- #b1; #b2; ncases b1; ncases b2; nnormalize; napply refl_eq. nqed.
+nlemma eq_to_eqastbasetype : ∀n1,n2.n1 = n2 → eq_astbasetype n1 n2 = true.
+ #n1; #n2; #H;
+ nrewrite > H;
+ nelim n2;
+ 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)
+nlemma neqastbasetype_to_neq : ∀n1,n2.eq_astbasetype n1 n2 = false → n1 ≠ n2.
+ #n1; #n2; #H;
+ napply (not_to_not (n1 = n2) (eq_astbasetype n1 n2 = true) …);
+ ##[ ##1: napply (eq_to_eqastbasetype n1 n2)
+ ##| ##2: napply (eqfalse_to_neqtrue … H)
  ##]
 nqed.
 
-nlemma eq_to_eqastbasetype : ∀b1,b2.b1 = b2 → eq_ast_base_type b1 b2 = true.
+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 (astbasetype_destruct … H)
+ ##| ##*: #H; ndestruct (*napply (bool_destruct … H)*)
+ ##]
+nqed.
+
+nlemma neq_to_neqastbasetype : ∀n1,n2.n1 ≠ n2 → eq_astbasetype n1 n2 = false.
+ #n1; #n2; #H;
+ napply (neqtrue_to_eqfalse (eq_astbasetype n1 n2));
+ napply (not_to_not (eq_astbasetype n1 n2 = true) (n1 = n2) ? H);
+ napply (eqastbasetype_to_eq n1 n2).
+nqed.
+
+nlemma decidable_astbasetype : ∀x,y:ast_base_type.decidable (x = y).
+ #x; #y; nnormalize;
+ napply (or2_elim (eq_astbasetype x y = true) (eq_astbasetype x y = false) ? (decidable_bexpr ?));
+ ##[ ##1: #H; napply (or2_intro1 (x = y) (x ≠ y) (eqastbasetype_to_eq … H))
+ ##| ##2: #H; napply (or2_intro2 (x = y) (x ≠ y) (neqastbasetype_to_neq … H))
+ ##]
+nqed.
+
+nlemma symmetric_eqastbasetype : symmetricT ast_base_type bool eq_astbasetype.
+ #n1; #n2;
+ napply (or2_elim (n1 = n2) (n1 ≠ n2) ? (decidable_astbasetype n1 n2));
+ ##[ ##1: #H; nrewrite > H; napply refl_eq
+ ##| ##2: #H; nrewrite > (neq_to_neqastbasetype n1 n2 H);
+          napply (symmetric_eq ? (eq_astbasetype n2 n1) false);
+          napply (neq_to_neqastbasetype n2 n1 (symmetric_neq ? n1 n2 H))
  ##]
 nqed.
 
@@ -109,195 +121,170 @@ nlemma asttype_destruct_struct_struct : ∀b1,b2.AST_TYPE_STRUCT b1 = AST_TYPE_S
  napply refl_eq.
 nqed.
 
+(*
 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
+ #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_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
-          ##]
-          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
+nlemma eq_to_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)) = true) →
+  ((bfold_right_neList2 ? (λx,y.eq_asttype x y) nl1 nl2) = true).
  #nl1; #nl2; #H;
  napply H.
 nqed.
 
-nlemma symmetric_eqasttype : symmetricT ast_type bool eq_ast_type.
- #t1; napply (ast_type_index … t1);
+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; nchange with ((eq_ast_base_type b1 b2) = (eq_ast_base_type b2 b1));
-                   nrewrite > (symmetric_eqastbasetype b1 b2);
+          ##[ ##1: #b2; #H; nrewrite > (asttype_destruct_base_base … H);
+                   nchange with ((eq_astbasetype b2 b2) = true);
+                   nrewrite > (eq_to_eqastbasetype b2 b2 (refl_eq …));
                    napply refl_eq
-          ##| ##2: #st2; #n2; nnormalize; napply refl_eq
-          ##| ##3: #nl2; nnormalize; napply refl_eq
+          ##| ##2: #st2; #n2; #H; ndestruct (*napply (asttype_destruct … H)*)
+          ##| ##3: #nl2; #H; ndestruct (*napply (asttype_destruct … H)*)
           ##]
  ##| ##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);
+          ##[ ##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;
                    napply refl_eq
-          ##| ##1: #b2; nnormalize; napply refl_eq
-          ##| ##3: #nl2; nnormalize; napply refl_eq
+          ##| ##1: #b2; #H1; ndestruct (*napply (asttype_destruct … H1)*)
+          ##| ##3: #nl2; #H1; ndestruct (*napply (asttype_destruct … H1)*)
           ##]
  ##| ##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);
+                   ##[ ##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; nnormalize; napply refl_eq
+                   ##| ##2: #hh2; #ll2; #H1;
+                            (* !!! ndestruct non va *)
+                            nelim (nelist_destruct_nil_cons ? hh1 hh2 ll2 (asttype_destruct_struct_struct … H1))
                    ##]
-          ##| ##1: #b2; nnormalize; napply refl_eq
-          ##| ##2: #st2; #n2; nnormalize; napply refl_eq
+          ##| ##1: #b2; #H1; ndestruct (*napply (asttype_destruct … H1)*)
+          ##| ##2: #st2; #n2; #H1; ndestruct (*napply (asttype_destruct … H1)*)
           ##]
  ##| ##4: #hh1; #ll1; #H; #H1; #t2; ncases t2;
           ##[ ##3: #nl2; ncases nl2;
-                   ##[ ##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
+                   ##[ ##1: #hh2; #H2;
+                            (* !!! ndestruct non va *)
+                            nelim (nelist_destruct_cons_nil ? hh1 hh2 ll1 (asttype_destruct_struct_struct … H2))
+                   ##| ##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
+                            ##| ##2: nrewrite > (nelist_destruct_cons_cons_2 … (asttype_destruct_struct_struct … H2));
+                                     napply refl_eq
+                            ##]
                    ##]
-          ##| ##1: #b2; nnormalize; napply refl_eq
-          ##| ##2: #st2; #n2; nnormalize; napply refl_eq
+          ##| ##1: #b2; #H2; ndestruct (*napply (asttype_destruct … H2)*)
+          ##| ##2: #st2; #n2; #H2; ndestruct (*napply (asttype_destruct … H2)*)
           ##]
  ##]
 nqed.
 
-nlemma eqasttype_to_eq : ∀t1,t2.eq_ast_type t1 t2 = true → t1 = t2.
+nlemma neqasttype_to_neq : ∀n1,n2.eq_asttype n1 n2 = false → n1 ≠ n2.
+ #n1; #n2; #H;
+ napply (not_to_not (n1 = n2) (eq_asttype n1 n2 = true) …);
+ ##[ ##1: napply (eq_to_eqasttype n1 n2)
+ ##| ##2: napply (eqfalse_to_neqtrue … H)
+ ##]
+nqed.
+
+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: #st2; #n2; nnormalize; #H; ndestruct (*napply (bool_destruct … H)*)
+          ##| ##3: #nl2; nnormalize; #H; ndestruct (*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
-          ##| ##1: #b2; nnormalize; #H1; napply (bool_destruct … H1)
-          ##| ##3: #nl2; nnormalize; #H1; napply (bool_destruct … H1)
+          ##| ##1: #b2; nnormalize; #H1; ndestruct (*napply (bool_destruct … H1)*)
+          ##| ##3: #nl2; nnormalize; #H1; ndestruct (*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);
+                   ##[ ##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)
+                   ##| ##2: #hh2; #ll2; nnormalize; #H1; ndestruct (*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; ndestruct (*napply (bool_destruct … H1)*)
+          ##| ##2: #st2; #n2; nnormalize; #H1; ndestruct (*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);
+                   ##[ ##1: #hh2; nnormalize; #H2; ndestruct (*napply (bool_destruct … H2)*)
+                   ##| ##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
                    ##]
-          ##| ##1: #b2; nnormalize; #H2; napply (bool_destruct … H2)
-          ##| ##2: #st2; #n2; nnormalize; #H2; napply (bool_destruct … H2)
+          ##| ##1: #b2; nnormalize; #H2; ndestruct (*napply (bool_destruct … H2)*)
+          ##| ##2: #st2; #n2; nnormalize; #H2; ndestruct (*napply (bool_destruct … H2)*)
           ##]
  ##]
 nqed.
 
-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).
- #nl1; #nl2; #H;
- napply H.
+nlemma neq_to_neqasttype : ∀n1,n2.n1 ≠ n2 → eq_asttype n1 n2 = false.
+ #n1; #n2; #H;
+ napply (neqtrue_to_eqfalse (eq_asttype n1 n2));
+ napply (not_to_not (eq_asttype n1 n2 = true) (n1 = n2) ? H);
+ napply (eqasttype_to_eq n1 n2).
 nqed.
 
-nlemma eq_to_eqasttype : ∀t1,t2.t1 = t2 → eq_ast_type 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);
-                   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));
-                   nnormalize;
-                   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))
-                   ##]
-          ##| ##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))
-                   ##| ##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 > (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: #b2; #H2; napply (asttype_destruct … H2)
-          ##| ##2: #st2; #n2; #H2; napply (asttype_destruct … H2)
-          ##]
+nlemma decidable_asttype : ∀x,y:ast_type.decidable (x = y).
+ #x; #y; nnormalize;
+ napply (or2_elim (eq_asttype x y = true) (eq_asttype x y = false) ? (decidable_bexpr ?));
+ ##[ ##1: #H; napply (or2_intro1 (x = y) (x ≠ y) (eqasttype_to_eq … H))
+ ##| ##2: #H; napply (or2_intro2 (x = y) (x ≠ y) (neqasttype_to_neq … H))
+ ##]
+nqed.
+
+nlemma symmetric_eqasttype : symmetricT ast_type bool eq_asttype.
+ #n1; #n2;
+ napply (or2_elim (n1 = n2) (n1 ≠ n2) ? (decidable_asttype n1 n2));
+ ##[ ##1: #H; nrewrite > H; napply refl_eq
+ ##| ##2: #H; nrewrite > (neq_to_neqasttype n1 n2 H);
+          napply (symmetric_eq ? (eq_asttype n2 n1) false);
+          napply (neq_to_neqasttype n2 n1 (symmetric_neq ? n1 n2 H))
  ##]
 nqed.
 
@@ -306,8 +293,8 @@ nlemma isbastbasetype_to_isastbasetype : ∀ast.isb_ast_base_type ast = true →
  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; ndestruct (*napply (bool_destruct … H)*)
+ ##| ##3: #t; #H; ndestruct (*napply (bool_destruct … H)*)
  ##]
 nqed.
 
@@ -315,7 +302,7 @@ nlemma isntbastbasetype_to_isntastbasetype : ∀ast.isntb_ast_base_type ast = tr
  #ast;
  ncases ast;
  nnormalize;
- ##[ ##1: #t; #H; napply (bool_destruct … H)
+ ##[ ##1: #t; #H; ndestruct (*napply (bool_destruct … H)*)
  ##| ##2: #t; #n; #H; napply I
  ##| ##3: #l; #H; napply I
  ##]