All eq_reflections apart from the one of 'id_apply' can be removed.
--- a/Quot/QuotMain.thy Tue Jan 26 08:55:55 2010 +0100
+++ b/Quot/QuotMain.thy Tue Jan 26 09:28:32 2010 +0100
@@ -47,8 +47,7 @@
qed
theorem thm10:
- shows "abs (rep a) \<equiv> a"
- apply (rule eq_reflection)
+ shows "abs (rep a) = a"
unfolding abs_def rep_def
proof -
from rep_prop
@@ -112,10 +111,10 @@
text {* Lemmas about simplifying id's. *}
lemmas [id_simps] =
id_def[symmetric]
- fun_map_id[THEN eq_reflection]
+ fun_map_id
id_apply[THEN eq_reflection]
- id_o[THEN eq_reflection]
- o_id[THEN eq_reflection]
+ id_o
+ o_id
eq_comp_r
text {* Translation functions for the lifting process. *}
--- a/Quot/quotient_tacs.ML Tue Jan 26 08:55:55 2010 +0100
+++ b/Quot/quotient_tacs.ML Tue Jan 26 09:28:32 2010 +0100
@@ -526,8 +526,7 @@
fun mk_simps thms = (mk_minimal_ss lthy) addsimps thms addSolver quotient_solver
val ss1 = mk_simps (defs @ prs @ @{thms babs_prs all_prs ex_prs ex1_prs})
- val ss2 = mk_simps (@{thms Quotient_abs_rep[THEN eq_reflection]
- Quotient_rel_rep[THEN eq_reflection] ex1_prs} @ ids)
+ val ss2 = mk_simps (@{thms Quotient_abs_rep Quotient_rel_rep ex1_prs} @ ids)
in
EVERY' [simp_tac ss1,
fun_map_tac lthy,