diff -r 19efc2af3e6c -r f967a16dfcdd src/HOL/Quotient_Examples/Quotient_Int.thy --- a/src/HOL/Quotient_Examples/Quotient_Int.thy Mon Sep 06 22:58:06 2010 +0200 +++ b/src/HOL/Quotient_Examples/Quotient_Int.thy Tue Sep 07 10:05:19 2010 +0200 @@ -14,7 +14,7 @@ "intrel (x, y) (u, v) = (x + v = u + y)" quotient_type int = "nat \ nat" / intrel - by (auto simp add: equivp_def expand_fun_eq) + by (auto simp add: equivp_def ext_iff) instantiation int :: "{zero, one, plus, uminus, minus, times, ord, abs, sgn}" begin