--- a/Quot/QuotMain.thy Mon Jan 25 20:35:42 2010 +0100
+++ b/Quot/QuotMain.thy Mon Jan 25 20:47:20 2010 +0100
@@ -111,7 +111,7 @@
text {* Lemmas about simplifying id's. *}
lemmas [id_simps] =
- id_def[symmetric, THEN eq_reflection]
+ id_def[symmetric]
fun_map_id[THEN eq_reflection]
id_apply[THEN eq_reflection]
id_o[THEN eq_reflection]
--- a/Quot/QuotOption.thy Mon Jan 25 20:35:42 2010 +0100
+++ b/Quot/QuotOption.thy Mon Jan 25 20:47:20 2010 +0100
@@ -82,7 +82,7 @@
done
lemma option_rel_eq[id_simps]:
- shows "option_rel op = \<equiv> op ="
+ shows "option_rel (op =) \<equiv> (op =)"
apply(rule eq_reflection)
apply(auto simp add: expand_fun_eq)
apply(case_tac x)
--- a/Quot/QuotProd.thy Mon Jan 25 20:35:42 2010 +0100
+++ b/Quot/QuotProd.thy Mon Jan 25 20:47:20 2010 +0100
@@ -80,7 +80,7 @@
(simp add: prod_fun_def)
lemma prod_rel_eq[id_simps]:
- shows "prod_rel op = op = \<equiv> op ="
+ shows "prod_rel (op =) (op =) \<equiv> (op =)"
apply (rule eq_reflection)
apply (rule ext)+
apply auto