Simplifying arguments; got rid of trans2_thm.
authorCezary Kaliszyk <kaliszyk@in.tum.de>
Fri, 27 Nov 2009 10:04:49 +0100
changeset 419 b1cd040ff5f7
parent 418 f24fd4689d00
child 420 dcfe009c98aa
child 421 2b64936f8fab
Simplifying arguments; got rid of trans2_thm.
FSet.thy
IntEx.thy
LFex.thy
LamEx.thy
QuotMain.thy
--- a/FSet.thy	Fri Nov 27 09:16:32 2009 +0100
+++ b/FSet.thy	Fri Nov 27 10:04:49 2009 +0100
@@ -304,7 +304,7 @@
 ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
 ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "fset"; *}
 ML {* val consts = lookup_quot_consts defs *}
-ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] rty [quot] trans2 rsp_thms defs *}
+ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] rty [quot] rsp_thms defs *}
 
 lemma "IN x EMPTY = False"
 by (tactic {* lift_tac_fset @{context} @{thm m1} 1 *})
@@ -376,7 +376,7 @@
 
 ML {* val rsp_thms = @{thms list_rec_rsp list_case_rsp} @ rsp_thms *}
 ML {* val defs = @{thms fset_rec_def fset_case_def} @ defs *}
-ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] rty [quot] trans2 rsp_thms defs *}
+ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] rty [quot] rsp_thms defs *}
 
 lemma "fset_rec (f1::'t) x (INSERT a xa) = x a xa (fset_rec f1 x xa)"
 apply (tactic {* lift_tac_fset @{context} @{thm list.recs(2)} 1 *})
@@ -397,7 +397,7 @@
   done
 
 
-ML {* fun r_mk_comb_tac_fset lthy = r_mk_comb_tac lthy rty [quot] [rel_refl] trans2 rsp_thms *}
+ML {* fun r_mk_comb_tac_fset lthy = r_mk_comb_tac lthy rty [quot] [rel_refl] [trans2] rsp_thms *}
 
 
 
--- a/IntEx.thy	Fri Nov 27 09:16:32 2009 +0100
+++ b/IntEx.thy	Fri Nov 27 10:04:49 2009 +0100
@@ -142,7 +142,7 @@
 
 ML {*
 fun lift_tac_fset lthy t =
-  lift_tac lthy t [rel_eqv] rty [quot] trans2 rsp_thms defs
+  lift_tac lthy t [rel_eqv] rty [quot] rsp_thms defs
 *}
 
 lemma "PLUS a b = PLUS b a"
@@ -191,7 +191,7 @@
 lemma "PLUS (PLUS i j) k = PLUS i (PLUS j k)"
 apply(tactic {* procedure_tac @{context} @{thm plus_assoc_pre} 1 *})
 apply(tactic {* regularize_tac @{context} [rel_eqv] [rel_refl] 1 *})
-apply(tactic {* REPEAT_ALL_NEW (r_mk_comb_tac @{context} rty [quot] [rel_refl] trans2 rsp_thms) 1 *})
+apply(tactic {* REPEAT_ALL_NEW (r_mk_comb_tac @{context} rty [quot] [rel_refl] [trans2] rsp_thms) 1 *})
 oops
 
 
--- a/LFex.thy	Fri Nov 27 09:16:32 2009 +0100
+++ b/LFex.thy	Fri Nov 27 10:04:49 2009 +0100
@@ -318,7 +318,7 @@
 apply (tactic {* clean_tac @{context}  defs aps 1 *})
 ML_prf {*  *}
 print_quotients
-apply(tactic {* r_mk_comb_tac' @{context} rty [quot] rel_refl trans2 [] 1*})
+apply(tactic {* r_mk_comb_tac' @{context} rty [quot] rel_refl [trans2] [] 1*})
 
 
 ML {* val consts = lookup_quot_consts defs *}
--- a/LamEx.thy	Fri Nov 27 09:16:32 2009 +0100
+++ b/LamEx.thy	Fri Nov 27 10:04:49 2009 +0100
@@ -179,7 +179,7 @@
 ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
 ML {* val consts = lookup_quot_consts defs *}
 ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "lam" *}
-ML {* fun lift_tac_lam lthy t = lift_tac lthy t [rel_eqv] rty [quot] trans2 rsp_thms defs *}
+ML {* fun lift_tac_lam lthy t = lift_tac lthy t [rel_eqv] rty [quot] rsp_thms defs *}
 
 lemma pi_var: "(pi\<Colon>('x \<times> 'x) list) \<bullet> Var a = Var (pi \<bullet> a)"
 apply (tactic {* lift_tac_lam @{context} @{thm pi_var_com} 1 *})
--- 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,