]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/assembly/string/string.ma
- setters for data structures now support "commuting conversion"-like
[helm.git] / helm / software / matita / contribs / assembly / string / string.ma
index f60b8a804c46f3a7bb5d42de6de1ffd1db4355ac..7681c112e5a1ce07219b293f3cc76f105c2b25b5 100755 (executable)
@@ -38,17 +38,53 @@ definition isNull_str ≝
  [ nil ⇒ true | cons _ _ ⇒ false ].
 
 (* strcmp *)
-let rec strCmp_str (str,str':aux_str_type) ≝
+let rec eqStr_str (str,str':aux_str_type) ≝
  match str with
   [ nil ⇒ match str' with
-   [ nil ⇒ true | cons _ _ ⇒ false ]
+   [ nil => true
+   | cons _ _ => false ]
   | cons h t ⇒ match str' with
    [ nil ⇒ false
-   | cons h' t' ⇒ match eqAsciiMin h h' with
-    [ true ⇒ strCmp_str t t' | false ⇒ false ]
-  ]].
+   | cons h' t' ⇒ (eqAsciiMin h h')⊗(eqStr_str t t')
+   ]
+  ].
 
-axiom eqbstrcmpstr_to_eq : ∀s,s'.strCmp_str s s' = true → s = s'.
+lemma eq_to_eqstr : ∀s,s'.s = s' → eqStr_str s s' = true.
+ do 2 intro;
+ intro;
+ rewrite < H;
+ elim s;
+ [ 1: reflexivity
+ | 2: simplify;
+      rewrite > (eq_to_eqasciimin a a (refl_eq ??));
+      rewrite > H1;
+      reflexivity
+ ].
+qed.
+
+lemma eqstr_to_eq : ∀s,s'.eqStr_str s s' = true → s = s'.
+ intros 1;
+ elim s 0;
+ intros;
+ [ simplify in H:(%);
+   cases s' in H;
+   simplify;
+   intros;
+   [ reflexivity
+   | destruct H
+   ]
+ | elim s' in H1;
+   [ simplify in H1;
+     destruct H1
+   | simplify in H2;
+     lapply (andb_true_true ?? H2);
+     lapply (andb_true_true_r ?? H2);
+     rewrite > (H ? Hletin1);
+     rewrite > (eqasciimin_to_eq ?? Hletin);
+     reflexivity
+   ]
+ ].
+qed.
 
 (* strcat *)
 definition strCat_str ≝
@@ -70,5 +106,27 @@ definition get_name_strId ≝ λsid:aux_strId_type.match sid with [ STR_ID n _ 
 definition get_id_strId ≝ λsid:aux_strId_type.match sid with [ STR_ID _ d ⇒ d ].
 
 (* confronto *)
-definition strCmp_strId ≝
-λsid,sid':aux_strId_type.(strCmp_str (get_name_strId sid) (get_name_strId sid'))⊗(eqb (get_id_strId sid) (get_id_strId sid')).
+definition eqStrId_str ≝
+λsid,sid':aux_strId_type.(eqStr_str (get_name_strId sid) (get_name_strId sid'))⊗(eqb (get_id_strId sid) (get_id_strId sid')).
+
+lemma eq_to_eqstrid : ∀s,s'.s = s' → eqStrId_str s s' = true.
+ intros 3;
+ rewrite < H;
+ elim s;
+ simplify;
+ rewrite > (eq_to_eqstr a a (refl_eq ??));
+ rewrite > (eq_to_eqb_true n n (refl_eq ??));
+ reflexivity.
+qed.
+
+lemma eqstrid_to_eq : ∀s,s'.eqStrId_str s s' = true → s = s'.
+ intros 2;
+ elim s 0;
+ elim s' 0;
+ intros 4;
+ simplify;
+ intro;
+ rewrite > (eqstr_to_eq a1 a (andb_true_true ?? H));
+ rewrite > (eqb_true_to_eq n1 n (andb_true_true_r ?? H));
+ reflexivity.
+qed.