--- a/Quot/QuotProd.thy Thu Dec 10 12:25:12 2009 +0100
+++ b/Quot/QuotProd.thy Thu Dec 10 16:56:03 2009 +0100
@@ -26,7 +26,8 @@
assumes q2: "Quotient R2 Abs2 Rep2"
shows "Quotient (prod_rel R1 R2) (prod_fun Abs1 Abs2) (prod_fun Rep1 Rep2)"
unfolding Quotient_def
-apply (simp add: Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] Quotient_rel_rep[OF q1] Quotient_rel_rep[OF q2])
+using q1 q2
+apply (simp add: Quotient_abs_rep Quotient_abs_rep Quotient_rel_rep Quotient_rel_rep)
using Quotient_rel[OF q1] Quotient_rel[OF q2] by blast
lemma pair_rsp:
@@ -41,8 +42,6 @@
shows "(prod_fun Abs1 Abs2) (Rep1 l, Rep2 r) \<equiv> (l, r)"
by (simp add: Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2])
-term Pair
-
lemma pair_prs[quot_preserve]:
assumes q1: "Quotient R1 Abs1 Rep1"
assumes q2: "Quotient R2 Abs2 Rep2"