# HG changeset patch # User haftmann # Date 1313516870 -7200 # Node ID a5cb6aa77ffc0a1b283b779ae544dc77ac5131cc # Parent 5f974bead4365a583087acc052b4b03904f133e8 avoid Collect_def in proof diff -r 5f974bead436 -r a5cb6aa77ffc src/HOL/Quotient.thy --- a/src/HOL/Quotient.thy Tue Aug 16 07:56:17 2011 -0700 +++ b/src/HOL/Quotient.thy Tue Aug 16 19:47:50 2011 +0200 @@ -643,10 +643,18 @@ have "Collect (R (SOME x. x \ Rep a)) = (Rep a)" by (metis some_collect rep_prop) then show "Abs (Collect (R (SOME x. x \ Rep a))) = a" using rep_inverse by auto have "R r r \ R s s \ Abs (Collect (R r)) = Abs (Collect (R s)) \ R r = R s" - by (metis Collect_def abs_inverse) + proof - + assume "R r r" and "R s s" + then have "Abs (Collect (R r)) = Abs (Collect (R s)) \ Collect (R r) = Collect (R s)" + by (metis abs_inverse) + also have "Collect (R r) = Collect (R s) \ (\A x. x \ A) (Collect (R r)) = (\A x. x \ A) (Collect (R s))" + by rule simp_all + finally show "Abs (Collect (R r)) = Abs (Collect (R s)) \ R r = R s" by simp + qed then show "R r s \ R r r \ R s s \ (Abs (Collect (R r)) = Abs (Collect (R s)))" using equivp[simplified part_equivp_def] by metis qed + end subsection {* ML setup *}