1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 (* ********************************************************************** *)
16 (* Progetto FreeScale *)
18 (* Sviluppato da: Cosimo Oliboni, oliboni@cs.unibo.it *)
19 (* Cosimo Oliboni, oliboni@cs.unibo.it *)
21 (* ********************************************************************** *)
23 include "utility/string.ma".
24 include "utility/ascii_lemmas2.ma".
25 include "freescale/nat_lemmas.ma".
26 include "utility/utility_lemmas.ma".
28 (* ************************ *)
29 (* MANIPOLAZIONE DI STRINGA *)
30 (* ************************ *)
32 nlemma symmetric_eqstr : symmetricT (list ascii) bool eq_str.
34 napply (symmetric_bfoldrightlist2 ascii eq_ascii s1 s2 symmetric_eqascii).
37 nlemma eqstr_to_eq : ∀s,s'.eq_str s s' = true → s = s'.
39 napply (bfoldrightlist2_to_eq ascii eq_ascii s1 s2 eqascii_to_eq).
42 nlemma eq_to_eqstr : ∀s,s'.s = s' → eq_str s s' = true.
44 napply (eq_to_bfoldrightlist2 ascii eq_ascii s1 s2 eq_to_eqascii).
51 nlemma strid_destruct_1 : ∀x1,x2,y1,y2.mk_strId x1 y1 = mk_strId x2 y2 → x1 = x2.
52 #x1; #x2; #y1; #y2; #H;
53 nchange with (match mk_strId x2 y2 with [ mk_strId a _ ⇒ x1 = a ]);
59 nlemma strid_destruct_2 : ∀x1,x2,y1,y2.mk_strId x1 y1 = mk_strId x2 y2 → y1 = y2.
60 #x1; #x2; #y1; #y2; #H;
61 nchange with (match mk_strId x2 y2 with [ mk_strId _ b ⇒ y1 = b ]);
67 nlemma symmetric_eqstrid : symmetricT strId bool eq_strId.
70 ((eq_str (str_elem si1) (str_elem si2))⊗(eq_nat (id_elem si1) (id_elem si2))) =
71 ((eq_str (str_elem si2) (str_elem si1))⊗(eq_nat (id_elem si2) (id_elem si1))));
72 nrewrite > (symmetric_eqstr (str_elem si1) (str_elem si2));
73 nrewrite > (symmetric_eqnat (id_elem si1) (id_elem si2));
77 nlemma eqstrid_to_eq : ∀s,s'.eq_strId s s' = true → s = s'.
83 nchange in H:(%) with (((eq_str l1 l2)⊗(eq_nat n1 n2)) = true);
84 nrewrite > (eqstr_to_eq l1 l2 (andb_true_true_l ?? H));
85 nrewrite > (eqnat_to_eq n1 n2 (andb_true_true_r ?? H));
89 nlemma eq_to_eqstrid : ∀s,s'.s = s' → eq_strId s s' = true.
95 nchange with (((eq_str l1 l2)⊗(eq_nat n1 n2)) = true);
96 nrewrite > (strid_destruct_1 ???? H);
97 nrewrite > (strid_destruct_2 ???? H);
98 nrewrite > (eq_to_eqstr l2 l2 (refl_eq ??));
99 nrewrite > (eq_to_eqnat n2 n2 (refl_eq ??));