X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fassembly%2Fstring%2Fstring.ma;h=3226eac436c3f01716b938ac75cc30bc2718a768;hb=34e2c8f59dd7924e15a7746644182d12ad09fed3;hp=f60b8a804c46f3a7bb5d42de6de1ffd1db4355ac;hpb=05e6e4771934d95be8b4cffcc87eeb7b27250536;p=helm.git diff --git a/helm/software/matita/contribs/assembly/string/string.ma b/helm/software/matita/contribs/assembly/string/string.ma index f60b8a804..3226eac43 100755 --- a/helm/software/matita/contribs/assembly/string/string.ma +++ b/helm/software/matita/contribs/assembly/string/string.ma @@ -38,17 +38,77 @@ 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 eqex_switch : ∀e1,e2.eq_ex e1 e2 = eq_ex e2 e1. + intros; + elim e1; + elim e2; + reflexivity. +qed. + +lemma eqb8_switch : ∀b1,b2.eq_b8 b1 b2 = eq_b8 b2 b1. + intros; + elim b1; + elim b2; + unfold eq_b8; + rewrite > eqex_switch in ⊢ (? ? % ?); + rewrite > eqex_switch in ⊢ (? ? (? ? %) ?); + reflexivity. +qed. + +lemma eqasciimin_switch : ∀a1,a2.eqAsciiMin a1 a2 = eqAsciiMin a2 a1. + intros; + unfold eqAsciiMin; + rewrite > eqb8_switch in ⊢ (? ? % ?); + reflexivity. +qed. + +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 +130,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.