X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fassembly%2Fstring%2Fstring.ma;h=3226eac436c3f01716b938ac75cc30bc2718a768;hb=ce3886fa05ff0a2fbd4d7b6cf68225d90686eafe;hp=7681c112e5a1ce07219b293f3cc76f105c2b25b5;hpb=55e5bef77f163b29feeb9ad4e83376c5aa301297;p=helm.git diff --git a/helm/software/matita/contribs/assembly/string/string.ma b/helm/software/matita/contribs/assembly/string/string.ma index 7681c112e..3226eac43 100755 --- a/helm/software/matita/contribs/assembly/string/string.ma +++ b/helm/software/matita/contribs/assembly/string/string.ma @@ -49,6 +49,30 @@ let rec eqStr_str (str,str':aux_str_type) ≝ ] ]. +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;