]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/assembly/string/string.ma
mod change (-x)
[helm.git] / helm / software / matita / contribs / assembly / string / string.ma
old mode 100755 (executable)
new mode 100644 (file)
index 7681c11..3226eac
@@ -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;