defs used automatically by clean_tac
authorCezary Kaliszyk <kaliszyk@in.tum.de>
Thu, 03 Dec 2009 12:31:05 +0100
changeset 498 e7bb6bbe7576
parent 497 b663bc007d00
child 499 f122816d7729
defs used automatically by clean_tac
FSet.thy
QuotMain.thy
--- a/FSet.thy	Thu Dec 03 12:22:19 2009 +0100
+++ b/FSet.thy	Thu Dec 03 12:31:05 2009 +0100
@@ -174,8 +174,6 @@
 term fmap
 thm fmap_def
 
-ML {* val defs = @{thms EMPTY_def IN_def FUNION_def CARD_def INSERT_def fmap_def FOLD_def} *}
-
 lemma memb_rsp:
   fixes z
   assumes a: "x \<approx> y"
@@ -299,14 +297,13 @@
 
 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] [quot] defs *}
+ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] [quot] *}
 
 lemma "IN x EMPTY = False"
 apply(tactic {* procedure_tac @{context} @{thm m1} 1 *})
 apply(tactic {* regularize_tac @{context} [rel_eqv] 1 *})
 apply(tactic {* all_inj_repabs_tac @{context} [quot] [rel_refl] [trans2] 1 *})
-apply(tactic {* clean_tac @{context} [quot] defs 1*})
+apply(tactic {* clean_tac @{context} [quot] 1*})
 done
 
 lemma "IN x (INSERT y xa) = (x = y \<or> IN x xa)"
@@ -351,7 +348,7 @@
 apply(tactic {* procedure_tac @{context} @{thm list.induct} 1 *})
 apply(tactic {* regularize_tac @{context} [rel_eqv] 1 *})
 prefer 2
-apply(tactic {* clean_tac @{context} [quot] defs 1 *})
+apply(tactic {* clean_tac @{context} [quot] 1 *})
 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 3 *) (* Ball-Ball *)
 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *)
 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 2 *) (* lam-lam-elim for R = (===>) *)
@@ -433,10 +430,9 @@
 where
   "INSERT2 \<equiv> op #"
 
-ML {* val defs = @{thms EMPTY2_def INSERT2_def} @ defs *}
 ML {* val quot = @{thms QUOTIENT_fset QUOTIENT_fset2} *}
 ML {* fun inj_repabs_tac_fset lthy = inj_repabs_tac lthy quot [rel_refl] [trans2] *}
-ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] quot defs *}
+ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] quot *}
 
 lemma "P (x :: 'a fset2) (EMPTY :: 'c fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"
 apply (tactic {* lift_tac_fset @{context} @{thm list_induct_part} 1 *})
@@ -470,8 +466,7 @@
   sorry
 
 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] quot defs *}
+ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] quot *}
 
 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 *})
@@ -481,8 +476,5 @@
 apply (tactic {* lift_tac_fset @{context} @{thm list.cases(2)} 1 *})
 done
 
-ML {* qconsts_lookup @{theory} "EMPTY" *}
-thm EMPTY_def
-
 
 end
--- 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