--- a/FSet.thy Fri Nov 27 07:16:16 2009 +0100
+++ b/FSet.thy Fri Nov 27 08:15:23 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 rel_refl rty quot trans2 rsp_thms reps_same absrep defs *}
+ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] rty [quot] trans2 rsp_thms defs *}
lemma "IN x EMPTY = False"
by (tactic {* lift_tac_fset @{context} @{thm m1} 1 *})
@@ -345,10 +345,10 @@
lemma "\<lbrakk>P EMPTY; \<And>a x. P x \<Longrightarrow> P (INSERT a x)\<rbrakk> \<Longrightarrow> P l"
apply(tactic {* procedure_tac @{context} @{thm list.induct} 1 *})
-apply(tactic {* regularize_tac @{context} [rel_eqv] rel_refl 1 *})
+apply(tactic {* regularize_tac @{context} [rel_eqv] [rel_refl] 1 *})
prefer 2
-apply (tactic {* clean_tac @{context} quot defs reps_same absrep [(@{typ "('a list \<Rightarrow> bool)"},@{typ "('a fset \<Rightarrow> bool)"})] 1 *})
-apply(tactic {* r_mk_comb_tac' @{context} rty quot rel_refl trans2 rsp_thms 1*})
+apply (tactic {* clean_tac @{context} [quot] defs [(@{typ "('a list \<Rightarrow> bool)"},@{typ "('a fset \<Rightarrow> bool)"})] 1 *})
+apply(tactic {* r_mk_comb_tac' @{context} rty [quot] rel_refl trans2 rsp_thms 1*})
done
quotient_def
@@ -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 rel_refl rty quot trans2 rsp_thms reps_same absrep defs *}
+ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] rty [quot] trans2 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 *}
@@ -405,7 +405,7 @@
(* Construction site starts here *)
lemma "P (x :: 'a list) (EMPTY :: 'c fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"
apply (tactic {* procedure_tac @{context} @{thm list_induct_part} 1 *})
-apply (tactic {* regularize_tac @{context} [rel_eqv] rel_refl 1 *})
+apply (tactic {* regularize_tac @{context} [rel_eqv] [rel_refl] 1 *})
apply (tactic {* (APPLY_RSP_TAC rty @{context}) 1 *})
apply (rule FUN_QUOTIENT)
apply (rule FUN_QUOTIENT)
@@ -454,25 +454,25 @@
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
apply (tactic {* instantiate_tac @{thm APPLY_RSP2} @{context} 1 *})
-apply (tactic {* (instantiate_tac @{thm REP_ABS_RSP(1)} @{context} THEN' (RANGE [quotient_tac quot])) 1 *})
+apply (tactic {* (instantiate_tac @{thm REP_ABS_RSP(1)} @{context} THEN' (RANGE [quotient_tac [quot]])) 1 *})
apply assumption
apply (rule refl)
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
apply (tactic {* instantiate_tac @{thm APPLY_RSP2} @{context} 1 *})
apply (tactic {* instantiate_tac @{thm APPLY_RSP2} @{context} 1 *})
-apply (tactic {* (instantiate_tac @{thm REP_ABS_RSP(1)} @{context} THEN' (RANGE [quotient_tac quot])) 1 *})
+apply (tactic {* (instantiate_tac @{thm REP_ABS_RSP(1)} @{context} THEN' (RANGE [quotient_tac [quot]])) 1 *})
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
apply (tactic {* REPEAT_ALL_NEW (r_mk_comb_tac_fset @{context}) 1 *})
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
apply (tactic {* instantiate_tac @{thm APPLY_RSP2} @{context} 1 *})
-apply (tactic {* (instantiate_tac @{thm REP_ABS_RSP(1)} @{context} THEN' (RANGE [quotient_tac quot])) 1 *})
+apply (tactic {* (instantiate_tac @{thm REP_ABS_RSP(1)} @{context} THEN' (RANGE [quotient_tac [quot]])) 1 *})
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
-apply (tactic {* clean_tac @{context} quot defs reps_same absrep [(@{typ "('a list \<Rightarrow> 'c list \<Rightarrow> bool)"},@{typ "('a list \<Rightarrow> 'c fset \<Rightarrow> bool)"})] 1 *})
+apply (tactic {* clean_tac @{context} [quot] defs [(@{typ "('a list \<Rightarrow> 'c list \<Rightarrow> bool)"},@{typ "('a list \<Rightarrow> 'c fset \<Rightarrow> bool)"})] 1 *})
done
end