]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_assembly/compiler/ast_type_lemmas.ma
freescale porting, work in progress
[helm.git] / helm / software / matita / contribs / ng_assembly / compiler / ast_type_lemmas.ma
index d7948c38690ddb735c2e402644ccacbc15fc3321..b6042164cf354027cfb0f281b22b2877de6611dd 100755 (executable)
@@ -21,6 +21,7 @@
 (* ********************************************************************** *)
 
 include "compiler/ast_type.ma".
+include "utility/utility_lemmas.ma".
 
 (* ************************* *)
 (* dimensioni degli elementi *)
@@ -100,7 +101,7 @@ nlemma asttype_destruct_array_array_2 : ∀x1,x2,y1,y2.AST_TYPE_ARRAY x1 y1 = AS
  napply (refl_eq ??).
 nqed.
 
-nlemma asttype_destruct_struc_struct : ∀b1,b2.AST_TYPE_STRUCT b1 = AST_TYPE_STRUCT b2 → b1 = b2.
+nlemma asttype_destruct_struct_struct : ∀b1,b2.AST_TYPE_STRUCT b1 = AST_TYPE_STRUCT b2 → b1 = b2.
  #b1; #b2; #H;
  nchange with (match AST_TYPE_STRUCT b2 with [ AST_TYPE_STRUCT a ⇒ b1 = a | _ ⇒ False ]);
  nrewrite < H;
@@ -203,8 +204,102 @@ nlemma symmetric_eqasttype : symmetricT ast_type bool eq_ast_type.
  ##]
 nqed.
 
-...
+nlemma eqasttype_to_eq : ∀t1,t2.eq_ast_type 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);
+                   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);
+                   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)
+                   ##]
+          ##| ##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)
+                   ##| ##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 ??)
+                   ##]
+          ##| ##1: #b2; nnormalize; #H2; napply (bool_destruct ??? H2)
+          ##| ##2: #st2; #n2; nnormalize; #H2; 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.
+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)
+          ##]
+ ##]
+nqed.
 
 nlemma isbastbasetype_to_isastbasetype : ∀ast.isb_ast_base_type ast = true → is_ast_base_type ast.
  #ast;