avoid accidental dependency of automated proof on sort equiv;
authorwenzelm
Thu, 16 Oct 2008 22:44:25 +0200
changeset 28616 ac1da69fbc5a
parent 28615 4c8fa015ec7f
child 28617 c9c1c8b28a62
avoid accidental dependency of automated proof on sort equiv;
src/HOL/ex/PER.thy
--- a/src/HOL/ex/PER.thy	Thu Oct 16 22:44:24 2008 +0200
+++ b/src/HOL/ex/PER.thy	Thu Oct 16 22:44:25 2008 +0200
@@ -223,7 +223,7 @@
 qed
 
 lemma eqv_class_eq' [simp]: "a \<in> domain ==> (\<lfloor>a\<rfloor> = \<lfloor>b\<rfloor>) = (a \<sim> b)"
-  using eqv_class_eqI eqv_class_eqD' by blast
+  using eqv_class_eqI eqv_class_eqD' by (blast del: eqv_refl)
 
 lemma eqv_class_eq [simp]: "(\<lfloor>a\<rfloor> = \<lfloor>b\<rfloor>) = (a \<sim> (b::'a::equiv))"
   using eqv_class_eqI eqv_class_eqD by blast