--- a/QuotMain.thy Fri Nov 27 09:16:32 2009 +0100
+++ b/QuotMain.thy Fri Nov 27 10:04:49 2009 +0100
@@ -136,6 +136,18 @@
end
+lemma tst: "EQUIV bla"
+sorry
+
+lemma equiv_trans2:
+ assumes e: "EQUIV R"
+ and ac: "R a c"
+ and bd: "R b d"
+ shows "R a b = R c d"
+using ac bd e
+unfolding EQUIV_REFL_SYM_TRANS TRANS_def SYM_def
+by (blast)
+
section {* type definition for the quotient type *}
(* the auxiliary data for the quotient types *)
@@ -803,9 +815,9 @@
*}
ML {*
-fun r_mk_comb_tac ctxt rty quot_thms rel_refl trans_thm rsp_thms =
+fun r_mk_comb_tac ctxt rty quot_thms rel_refl trans2 rsp_thms =
(FIRST' [
- rtac trans_thm,
+ FIRST' (map rtac trans2),
LAMBDA_RES_TAC ctxt,
rtac @{thm RES_FORALL_RSP},
ball_rsp_tac ctxt,
@@ -1116,7 +1128,7 @@
ML {*
(* FIXME/TODO should only get as arguments the rthm like procedure_tac *)
-fun lift_tac lthy rthm rel_eqv rty quot trans2 rsp_thms defs =
+fun lift_tac lthy rthm rel_eqv rty quot rsp_thms defs =
ObjectLogic.full_atomize_tac
THEN' gen_frees_tac lthy
THEN' Subgoal.FOCUS (fn {context, concl, ...} =>
@@ -1125,6 +1137,7 @@
val rule = procedure_inst context (prop_of rthm') (term_of concl)
val aps = find_aps (prop_of rthm') (term_of concl)
val rel_refl = map (fn x => @{thm EQUIV_REFL} OF [x]) rel_eqv
+ val trans2 = map (fn x => @{thm equiv_trans2} OF [x]) rel_eqv
in
EVERY1
[rtac rule,