]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_assembly/common/string_lemmas.ma
mod change (-x)
[helm.git] / helm / software / matita / contribs / ng_assembly / common / string_lemmas.ma
old mode 100755 (executable)
new mode 100644 (file)
index b30befb..1b59617
 (*                          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;