]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/ng_assembly/test_errori.ma
freescale porting, work in progress
[helm.git] / helm / software / matita / contribs / ng_assembly / test_errori.ma
1
2 include "utility/utility.ma".
3
4 (*
5 nlemma fold_right_neList2_aux3 :
6 \forall T. \forall h,h',t,t'.len_neList T (h§§t) = len_neList T (h'§§t') → len_neList T t = len_neList T t'.
7  #T; #h; #h'; #t; #t';
8  napply (ne_list_ind T ??? t);
9  napply (ne_list_ind T ??? t');
10  ##[ ##1: nnormalize; #x; #y; #H; napply (refl_eq ??)
11  ##| ##2: #a; #l'; #H; #x; #H1;
12           nchange in H1:(%) with ((S (len_neList T «£x»)) = (S (len_neList T (a§§l'))));
13           nrewrite > (nat_destruct_S_S ?? H1);
14           napply (refl_eq ??)
15  ##| ##3: #x; #a; #l'; #H; #H1;
16           nchange in H1:(%) with ((S (len_neList T (a§§l')))= (S (len_neList T «£x»)));
17           nrewrite > (nat_destruct_S_S ?? H1);
18           napply (refl_eq ??)
19  ##| ##4: #a; #l; #H; #a1; #l1; #H1; #H2;
20           (* sarebbe nchange in H2:(%) with ((S (len_neList T (a1§§l1))) = (S (len_neList T (a§§l)))); *)
21           (* ma fa passare il seguente errato ... *) 
22           nchange in H2:(%) with ((S (len_neList T (a§§l1))) = (S (len_neList T (a§§l))));
23 *)
24
25 include "freescale/byte8_lemmas.ma".
26
27 (*
28 nlemma associative_plusb8_aux
29  : \forall e1,e2,e3,e4.
30   match plus_ex_d_dc e2 e4 with
31    [ pair l c ⇒ 〈plus_ex_dc_d e1 e3 c,l〉 ] =
32   〈plus_ex_dc_d e1 e3 (snd ?? (plus_ex_d_dc e2 e4)),(fst ?? (plus_ex_d_dc e2 e4))〉.
33  #e1; #e2; #e3; #e4;
34  (* anche qui appare un T1 *)
35  ncases (plus_ex_d_dc e2 e4);
36  #e5; #c;
37  napply (refl_eq ??).
38 nqed.
39
40 nlemma associative_plusb8
41  : \forall b1,b2,b3.(plus_b8_d_d (plus_b8_d_d b1 b2) b3) = (plus_b8_d_d b1 (plus_b8_d_d b2 b3)).
42  #b1; ncases b1; #e1; #e2;
43  #b2; ncases b2; #e3; #e4;
44  #b3; ncases b3; #e5; #e6;
45
46 (* perche' volendo posso introdurre anche 2 premesse diverse con lo stesso nome? tipo #e2; #e2 *)
47
48  nchange with ( 
49  match plus_ex_d_dc (b8l (match plus_ex_d_dc e2 e4 with
50                            [ pair l1 c1 ⇒ 〈plus_ex_dc_d e1 e3 c1,l1〉 ])) e6 with
51   [ pair l2 c2 ⇒ 〈plus_ex_dc_d (b8h (match plus_ex_d_dc e2 e4 with
52                                     [ pair l3 c3 ⇒ 〈plus_ex_dc_d e1 e3 c3,l3〉 ])) e5 c2,l2〉 ] =
53  match plus_ex_d_dc e2 (b8l (match plus_ex_d_dc e4 e6 with
54                                     [ pair l4 c4 ⇒ 〈plus_ex_dc_d e3 e5 c4,l4〉 ])) with
55   [ pair l5 c5 ⇒ 〈plus_ex_dc_d e1 (b8h (match plus_ex_d_dc e4 e6 with
56                                              [ pair l6 c6 ⇒ 〈plus_ex_dc_d e3 e5 c6,l6〉 ])) c5,l5〉 ]);
57
58 (* gia' qua ci sono T1, T2 che appaiono dal nulla al posto delle variabili *)
59
60  nrewrite > (associative_plusb8_aux e1 e2 e3 e4);
61  nrewrite > (associative_plusb8_aux e3 e4 e5 e6);
62  nrewrite > (plusex_d_dc_to_d_c e2 e4);
63  nrewrite > (plusex_d_dc_to_d_d e2 e4);
64  nrewrite > (plusex_d_dc_to_d_c e4 e6);
65
66  (* nel visualizzatore era (snd ?? (plus_ex_d_dc e5 T2)) ma ha accettato la versione corretta *)
67
68  nrewrite > (plusex_d_dc_to_d_d e4 e6);
69  
70  (*...*)
71 *)
72
73 include "compiler/ast_type_lemmas.ma".
74
75 (*
76 nlemma symmetric_eqasttype_aux1
77  : ∀x1,x2,y2.eq_nat (len_neList ast_type («£x1»)) (len_neList ast_type (x2§§y2)) = false.
78  #x1; #x2; #y2; ncases y2; nnormalize;
79  ##[ ##1: #t; napply (refl_eq ??)
80  ##| ##2: #t; #l; napply (refl_eq ??)
81  ##]
82 nqed.
83
84 nlemma symmetric_eqasttype_aux2
85  : ∀x1,y1,x2.eq_nat (len_neList ast_type (x1§§y1)) (len_neList ast_type («£x2»)) = false.
86  #x1; #y1; #x2; ncases y1; nnormalize;
87  ##[ ##1: #t; napply (refl_eq ??)
88  ##| ##2: #t; #l; napply (refl_eq ??)
89  ##]
90 nqed.
91
92 ndefinition symmetric_eqasttype_aux3 ≝
93 λnelSubType1,nelSubType2:ne_list ast_type.
94  match eq_nat (len_neList ast_type nelSubType1) (len_neList ast_type nelSubType2)
95   return λx.eq_nat (len_neList ast_type nelSubType1) (len_neList ast_type nelSubType2) = x → bool
96  with
97   [ true ⇒ λp:(eq_nat (len_neList ast_type nelSubType1) (len_neList ast_type nelSubType2) = true).
98    fold_right_neList2 ?? (λx1,x2,acc.(eq_ast_type x1 x2)⊗acc)
99                       true nelSubType1 nelSubType2
100                       (eqnat_to_eq (len_neList ? nelSubType1) (len_neList ? nelSubType2) p)
101   | false ⇒ λp:(eq_nat (len_neList ast_type nelSubType1) (len_neList ast_type nelSubType2) = false).false
102   ] (refl_eq ? (eq_nat (len_neList ast_type nelSubType1) (len_neList ast_type nelSubType2))).
103
104 nlemma symmetric_eqasttype : symmetricT ast_type bool eq_ast_type.
105  #t1;
106  napply (ast_type_ind ????? t1);
107  ##[ ##1: #b1; #t2; ncases t2;
108           ##[ ##1: #b2; nchange with ((eq_ast_base_type b1 b2) = (eq_ast_base_type b2 b1));
109                    nrewrite > (symmetric_eqastbasetype b1 b2);
110                    napply (refl_eq ??)
111           ##| ##2: #tt; #n; nnormalize; napply (refl_eq ??)
112           ##| ##3: #l; nnormalize; napply (refl_eq ??)
113           ##]
114  ##| ##2: #tt1; #n1; #H; #t2; ncases t2;
115           ##[ ##2: #tt2; #n2; nchange with (((eq_ast_type tt1 tt2)⊗(eq_nat n1 n2)) = ((eq_ast_type tt2 tt1)⊗(eq_nat n2 n1)));
116                    nrewrite > (H tt2);
117                    nrewrite > (symmetric_eqnat n1 n2);
118                    napply (refl_eq ??)
119           ##| ##1: #b2; nnormalize; napply (refl_eq ??)
120           ##| ##3: #l; nnormalize; napply (refl_eq ??)
121           ##]
122  ##| ##3: #tt1; #H; #t2; ncases t2;
123           ##[ ##3: #l; ncases l;
124                    ##[ ##1: #tt2; nnormalize; nrewrite > (H tt2); napply (refl_eq ??)
125                    ##| ##2: #tt2; #l1;
126                             nchange with (
127                              (match eq_nat (len_neList ast_type «£tt1») (len_neList ast_type (tt2§§l1))
128                                return λx.eq_nat (len_neList ast_type «£tt1») (len_neList ast_type (tt2§§l1)) = x → bool
129                               with
130                                [ true ⇒ λp:(eq_nat (len_neList ast_type «£tt1») (len_neList ast_type (tt2§§l1)) = true).
131                                 fold_right_neList2 ?? (λx1,x2,acc.(eq_ast_type x1 x2)⊗acc)
132                                                    true «£tt1» (tt2§§l1)
133                                                    (eqnat_to_eq (len_neList ? «£tt1») (len_neList ? (tt2§§l1)) p)
134                                | false ⇒ λp:(eq_nat (len_neList ast_type «£tt1») (len_neList ast_type (tt2§§l1)) = false).false
135                                ] (refl_eq ? (eq_nat (len_neList ast_type «£tt1») (len_neList ast_type (tt2§§l1))))) = ?);
136
137                             (* eseguendo questa sequenza e' ok *)
138                             nrewrite > (symmetric_eqasttype_aux1 tt1 tt2 l1);
139                             nchange with (
140                              false = (match eq_nat (len_neList ast_type (tt2§§l1)) (len_neList ast_type «£tt1»)
141                                return λx.eq_nat (len_neList ast_type (tt2§§l1)) (len_neList ast_type «£tt1») = x → bool
142                               with
143                                [ true ⇒ λp:(eq_nat (len_neList ast_type (tt2§§l1)) (len_neList ast_type «£tt1») = true).
144                                 fold_right_neList2 ?? (λx1,x2,acc.(eq_ast_type x1 x2)⊗acc)
145                                                    true (tt2§§l1) «£tt1»
146                                                    (eqnat_to_eq (len_neList ? (tt2§§l1)) (len_neList ? «£tt1») p)
147                                | false ⇒ λp:(eq_nat (len_neList ast_type (tt2§§l1)) (len_neList ast_type «£tt1») = false).false
148                                ] (refl_eq ? (eq_nat (len_neList ast_type (tt2§§l1)) (len_neList ast_type «£tt1»)))));
149                             nrewrite > (symmetric_eqasttype_aux2 tt2 l1 tt1);
150                             nnormalize;
151
152                             (* se commentiamo invece la prima sequenza ed eseguiamo questa *)
153                             (* come e' possibile che rigetti la seconda rewrite? *)
154                             nchange with (
155                              ? = (match eq_nat (len_neList ast_type (tt2§§l1)) (len_neList ast_type «£tt1»)
156                                return λx.eq_nat (len_neList ast_type (tt2§§l1)) (len_neList ast_type «£tt1») = x → bool
157                               with
158                                [ true ⇒ λp:(eq_nat (len_neList ast_type (tt2§§l1)) (len_neList ast_type «£tt1») = true).
159                                 fold_right_neList2 ?? (λx1,x2,acc.(eq_ast_type x1 x2)⊗acc)
160                                                    true (tt2§§l1) «£tt1»
161                                                    (eqnat_to_eq (len_neList ? (tt2§§l1)) (len_neList ? «£tt1») p)
162                                | false ⇒ λp:(eq_nat (len_neList ast_type (tt2§§l1)) (len_neList ast_type «£tt1») = false).false
163                                ] (refl_eq ? (eq_nat (len_neList ast_type (tt2§§l1)) (len_neList ast_type «£tt1»)))));
164                             nrewrite > (symmetric_eqasttype_aux1 tt1 tt2 l1);                            
165                             nrewrite > (symmetric_eqasttype_aux2 tt2 l1 tt1);
166 *)