X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fng_assembly%2Fcommon%2Fstring_lemmas.ma;h=1b596170bbf0afbf5609db4d28ea8410d2c58d6d;hb=eb4144a401147a44a9620169eb6dafeb8f5a2c17;hp=b30befbf9a41ced577d514b1b3bd8b99a4c88532;hpb=601baed778a190b580982b588ebe49ba3f762b30;p=helm.git diff --git a/helm/software/matita/contribs/ng_assembly/common/string_lemmas.ma b/helm/software/matita/contribs/ng_assembly/common/string_lemmas.ma index b30befbf9..1b596170b 100755 --- a/helm/software/matita/contribs/ng_assembly/common/string_lemmas.ma +++ b/helm/software/matita/contribs/ng_assembly/common/string_lemmas.ma @@ -16,15 +16,12 @@ (* Progetto FreeScale *) (* *) (* Sviluppato da: Ing. Cosimo Oliboni, oliboni@cs.unibo.it *) -(* Ultima modifica: 05/08/2009 *) +(* Sviluppo: 2008-2010 *) (* *) (* ********************************************************************** *) include "common/string.ma". -include "common/ascii_lemmas2.ma". -include "common/ascii_lemmas3.ma". -include "common/ascii_lemmas4.ma". -include "common/ascii_lemmas5.ma". +include "common/ascii_lemmas.ma". include "common/list_utility_lemmas.ma". (* ************************ *) @@ -33,17 +30,17 @@ include "common/list_utility_lemmas.ma". nlemma symmetric_eqstr : symmetricT (list ascii) bool eq_str. #s1; #s2; - napply (symmetric_bfoldrightlist2 ascii eq_ascii s1 s2 symmetric_eqascii). + napply (symmetric_bfoldrightlist2 ascii eq_ascii symmetric_eqascii s1 s2). nqed. nlemma eqstr_to_eq : ∀s,s'.eq_str s s' = true → s = s'. #s1; #s2; - napply (bfoldrightlist2_to_eq ascii eq_ascii s1 s2 eqascii_to_eq). + napply (bfoldrightlist2_to_eq ascii eq_ascii eqascii_to_eq s1 s2). nqed. nlemma eq_to_eqstr : ∀s,s'.s = s' → eq_str s s' = true. #s1; #s2; - napply (eq_to_bfoldrightlist2 ascii eq_ascii s1 s2 eq_to_eqascii). + napply (eq_to_bfoldrightlist2 ascii eq_ascii eq_to_eqascii s1 s2). nqed. nlemma decidable_str : ∀x,y:list ascii.decidable (x = y). @@ -53,13 +50,13 @@ nqed. nlemma neqstr_to_neq : ∀s,s'.eq_str s s' = false → s ≠ s'. #s1; #s2; - napply (nbfoldrightlist2_to_neq ascii eq_ascii s1 s2 …); + napply (nbfoldrightlist2_to_neq ascii eq_ascii ? s1 s2 …); napply neqascii_to_neq. nqed. nlemma neq_to_neqstr : ∀s,s'.s ≠ s' → eq_str s s' = false. #s1; #s2; - napply (neq_to_nbfoldrightlist2 ascii eq_ascii s1 s2 …); + napply (neq_to_nbfoldrightlist2 ascii eq_ascii ?? s1 s2 …); ##[ ##1: napply decidable_ascii ##| ##2: napply neq_to_neqascii ##] @@ -93,7 +90,7 @@ nlemma symmetric_eqstrid : symmetricT strId bool eq_strId. nrewrite > (symmetric_eqstr (str_elem si1) (str_elem si2)); nrewrite > (symmetric_eqnat (id_elem si1) (id_elem si2)); napply refl_eq. -nqed. +nqed. nlemma eqstrid_to_eq : ∀s,s'.eq_strId s s' = true → s = s'. #si1; #si2;