--- a/QuotMain.thy Thu Dec 03 12:22:19 2009 +0100
+++ b/QuotMain.thy Thu Dec 03 12:31:05 2009 +0100
@@ -1088,8 +1088,10 @@
*)
ML {*
-fun clean_tac lthy quot defs =
+fun clean_tac lthy quot =
let
+ val thy = ProofContext.theory_of lthy;
+ val defs = map (#def) (qconsts_dest thy);
val absrep = map (fn x => @{thm QUOTIENT_ABS_REP} OF [x]) quot
val meta_absrep = map (fn x => @{thm eq_reflection} OF [x]) absrep;
val reps_same = map (fn x => @{thm QUOTIENT_REL_REP} OF [x]) quot
@@ -1213,7 +1215,7 @@
ML {*
(* FIXME/TODO should only get as arguments the rthm like procedure_tac *)
-fun lift_tac lthy rthm rel_eqv quot defs =
+fun lift_tac lthy rthm rel_eqv quot =
ObjectLogic.full_atomize_tac
THEN' gen_frees_tac lthy
THEN' Subgoal.FOCUS (fn {context, concl, ...} =>
@@ -1229,9 +1231,12 @@
RANGE [rtac rthm',
regularize_tac lthy rel_eqv,
rtac thm THEN' all_inj_repabs_tac lthy quot rel_refl trans2,
- clean_tac lthy quot defs]]
+ clean_tac lthy quot]]
end) lthy
*}
+print_quotients
+
+
end