src/HOL/Induct/QuoNestedDataType.thy
changeset 30240 5b25fee0362c
parent 23746 a455e69c31cc
child 32960 69916a850301
--- a/src/HOL/Induct/QuoNestedDataType.thy	Wed Mar 04 10:43:39 2009 +0100
+++ b/src/HOL/Induct/QuoNestedDataType.thy	Wed Mar 04 10:45:52 2009 +0100
@@ -44,7 +44,7 @@
 
 theorem equiv_exprel: "equiv UNIV exprel"
 proof -
-  have "reflexive exprel" by (simp add: refl_def exprel_refl)
+  have "refl exprel" by (simp add: refl_on_def exprel_refl)
   moreover have "sym exprel" by (simp add: sym_def, blast intro: exprel.SYM)
   moreover have "trans exprel" by (simp add: trans_def, blast intro: exprel.TRANS)
   ultimately show ?thesis by (simp add: equiv_def)