]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_assembly/common/string_lemmas.ma
(no commit message)
[helm.git] / helm / software / matita / contribs / ng_assembly / common / string_lemmas.ma
index 170581a672a50ba1cc72f83e07bc734464094800..53dcc50d100daf8b10c1eb6df59656d4350b22c2 100755 (executable)
 (*                          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".
 
 (* ************************ *)
@@ -153,7 +150,7 @@ nlemma neqstrid_to_neq : ∀sid1,sid2:strId.(eq_strId sid1 sid2 = false) → (si
  #sid2; nelim sid2; #s2; #n2;
  nchange with ((((eq_str s1 s2) ⊗ (eq_nat n1 n2)) = false) → ?);
  #H;
- napply (or2_elim ((eq_str s1 s2) = false) ((eq_nat n1 n2) = false) ? (andb_false … H) …);
+ napply (or2_elim ((eq_str s1 s2) = false) ((eq_nat n1 n2) = false) ? (andb_false2 … H) …);
  ##[ ##1: #H1; napply (decidable_strid_aux1 … (neqstr_to_neq … H1))
  ##| ##2: #H1; napply (decidable_strid_aux2 … (neqnat_to_neq … H1))
  ##]