theory FSet
imports QuotMain
begin
inductive
list_eq (infix "\<approx>" 50)
where
"a#b#xs \<approx> b#a#xs"
| "[] \<approx> []"
| "xs \<approx> ys \<Longrightarrow> ys \<approx> xs"
| "a#a#xs \<approx> a#xs"
| "xs \<approx> ys \<Longrightarrow> a#xs \<approx> a#ys"
| "\<lbrakk>xs1 \<approx> xs2; xs2 \<approx> xs3\<rbrakk> \<Longrightarrow> xs1 \<approx> xs3"
lemma list_eq_refl:
shows "xs \<approx> xs"
apply (induct xs)
apply (auto intro: list_eq.intros)
done
lemma equiv_list_eq:
shows "EQUIV list_eq"
unfolding EQUIV_REFL_SYM_TRANS REFL_def SYM_def TRANS_def
apply(auto intro: list_eq.intros list_eq_refl)
done
quotient fset = "'a list" / "list_eq"
apply(rule equiv_list_eq)
done
print_theorems
typ "'a fset"
thm "Rep_fset"
thm "ABS_fset_def"
quotient_def
EMPTY :: "'a fset"
where
"EMPTY \<equiv> ([]::'a list)"
term Nil
term EMPTY
thm EMPTY_def
quotient_def
INSERT :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
where
"INSERT \<equiv> op #"
term Cons
term INSERT
thm INSERT_def
quotient_def
FUNION :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
where
"FUNION \<equiv> (op @)"
term append
term FUNION
thm FUNION_def
thm QUOTIENT_fset
thm QUOT_TYPE_I_fset.thm11
fun
membship :: "'a \<Rightarrow> 'a list \<Rightarrow> bool" (infix "memb" 100)
where
m1: "(x memb []) = False"
| m2: "(x memb (y#xs)) = ((x=y) \<or> (x memb xs))"
fun
card1 :: "'a list \<Rightarrow> nat"
where
card1_nil: "(card1 []) = 0"
| card1_cons: "(card1 (x # xs)) = (if (x memb xs) then (card1 xs) else (Suc (card1 xs)))"
quotient_def
CARD :: "'a fset \<Rightarrow> nat"
where
"CARD \<equiv> card1"
term card1
term CARD
thm CARD_def
(* text {*
Maybe make_const_def should require a theorem that says that the particular lifted function
respects the relation. With it such a definition would be impossible:
make_const_def @{binding CARD} @{term "length"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
*}*)
lemma card1_0:
fixes a :: "'a list"
shows "(card1 a = 0) = (a = [])"
by (induct a) auto
lemma not_mem_card1:
fixes x :: "'a"
fixes xs :: "'a list"
shows "(~(x memb xs)) = (card1 (x # xs) = Suc (card1 xs))"
by auto
lemma mem_cons:
fixes x :: "'a"
fixes xs :: "'a list"
assumes a : "x memb xs"
shows "x # xs \<approx> xs"
using a by (induct xs) (auto intro: list_eq.intros )
lemma card1_suc:
fixes xs :: "'a list"
fixes n :: "nat"
assumes c: "card1 xs = Suc n"
shows "\<exists>a ys. ~(a memb ys) \<and> xs \<approx> (a # ys)"
using c
apply(induct xs)
apply (metis Suc_neq_Zero card1_0)
apply (metis QUOT_TYPE_I_fset.R_trans card1_cons list_eq_refl mem_cons)
done
definition
rsp_fold
where
"rsp_fold f = ((!u v. (f u v = f v u)) \<and> (!u v w. ((f u (f v w) = f (f u v) w))))"
primrec
fold1
where
"fold1 f (g :: 'a \<Rightarrow> 'b) (z :: 'b) [] = z"
| "fold1 f g z (a # A) =
(if rsp_fold f
then (
if (a memb A) then (fold1 f g z A) else (f (g a) (fold1 f g z A))
) else z)"
(* fold1_def is not usable, but: *)
thm fold1.simps
lemma fs1_strong_cases:
fixes X :: "'a list"
shows "(X = []) \<or> (\<exists>a. \<exists> Y. (~(a memb Y) \<and> (X \<approx> a # Y)))"
apply (induct X)
apply (simp)
apply (metis QUOT_TYPE_I_fset.thm11 list_eq_refl mem_cons m1)
done
quotient_def
IN :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool"
where
"IN \<equiv> membship"
term membship
term IN
thm IN_def
term fold1
quotient_def
FOLD :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b fset \<Rightarrow> 'a"
where
"FOLD \<equiv> fold1"
term fold1
term fold
thm fold_def
quotient_def
fmap::"('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset"
where
"fmap \<equiv> map"
term map
term fmap
thm fmap_def
ML {* prop_of @{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: "list_eq x y"
shows "(z memb x) = (z memb y)"
using a by induct auto
lemma ho_memb_rsp:
"(op = ===> (op \<approx> ===> op =)) (op memb) (op memb)"
by (simp add: memb_rsp)
lemma card1_rsp:
fixes a b :: "'a list"
assumes e: "a \<approx> b"
shows "card1 a = card1 b"
using e by induct (simp_all add:memb_rsp)
lemma ho_card1_rsp: "(op \<approx> ===> op =) card1 card1"
by (simp add: card1_rsp)
lemma cons_rsp:
fixes z
assumes a: "xs \<approx> ys"
shows "(z # xs) \<approx> (z # ys)"
using a by (rule list_eq.intros(5))
lemma ho_cons_rsp:
"(op = ===> op \<approx> ===> op \<approx>) op # op #"
by (simp add: cons_rsp)
lemma append_rsp_fst:
assumes a : "list_eq l1 l2"
shows "(l1 @ s) \<approx> (l2 @ s)"
using a
by (induct) (auto intro: list_eq.intros list_eq_refl)
lemma append_end:
shows "(e # l) \<approx> (l @ [e])"
apply (induct l)
apply (auto intro: list_eq.intros list_eq_refl)
done
lemma rev_rsp:
shows "a \<approx> rev a"
apply (induct a)
apply simp
apply (rule list_eq_refl)
apply simp_all
apply (rule list_eq.intros(6))
prefer 2
apply (rule append_rsp_fst)
apply assumption
apply (rule append_end)
done
lemma append_sym_rsp:
shows "(a @ b) \<approx> (b @ a)"
apply (rule list_eq.intros(6))
apply (rule append_rsp_fst)
apply (rule rev_rsp)
apply (rule list_eq.intros(6))
apply (rule rev_rsp)
apply (simp)
apply (rule append_rsp_fst)
apply (rule list_eq.intros(3))
apply (rule rev_rsp)
done
lemma append_rsp:
assumes a : "list_eq l1 r1"
assumes b : "list_eq l2 r2 "
shows "(l1 @ l2) \<approx> (r1 @ r2)"
apply (rule list_eq.intros(6))
apply (rule append_rsp_fst)
using a apply (assumption)
apply (rule list_eq.intros(6))
apply (rule append_sym_rsp)
apply (rule list_eq.intros(6))
apply (rule append_rsp_fst)
using b apply (assumption)
apply (rule append_sym_rsp)
done
lemma ho_append_rsp:
"(op \<approx> ===> op \<approx> ===> op \<approx>) op @ op @"
by (simp add: append_rsp)
lemma map_rsp:
assumes a: "a \<approx> b"
shows "map f a \<approx> map f b"
using a
apply (induct)
apply(auto intro: list_eq.intros)
done
lemma ho_map_rsp:
"(op = ===> op \<approx> ===> op \<approx>) map map"
by (simp add: map_rsp)
lemma map_append:
"(map f (a @ b)) \<approx>
(map f a) @ (map f b)"
by simp (rule list_eq_refl)
lemma ho_fold_rsp:
"(op = ===> op = ===> op = ===> op \<approx> ===> op =) fold1 fold1"
apply (auto simp add: FUN_REL_EQ)
apply (case_tac "rsp_fold x")
prefer 2
apply (erule_tac list_eq.induct)
apply (simp_all)
apply (erule_tac list_eq.induct)
apply (simp_all)
apply (auto simp add: memb_rsp rsp_fold_def)
done
print_quotients
ML {* val qty = @{typ "'a fset"} *}
ML {* val rsp_thms =
@{thms ho_memb_rsp ho_cons_rsp ho_card1_rsp ho_map_rsp ho_append_rsp ho_fold_rsp}
@ @{thms ho_all_prs ho_ex_prs} *}
ML {* fun lift_thm_fset lthy t = lift_thm lthy qty "fset" rsp_thms defs t *}
ML {* fun lift_thm_g_fset lthy t g = lift_thm_goal lthy qty "fset" rsp_thms defs t g *}
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 defs *}
lemma "IN x EMPTY = False"
by (tactic {* lift_tac_fset @{context} @{thm m1} 1 *})
lemma "IN x (INSERT y xa) = (x = y \<or> IN x xa)"
by (tactic {* lift_tac_fset @{context} @{thm m2} 1 *})
lemma "INSERT a (INSERT a x) = INSERT a x"
apply (tactic {* lift_tac_fset @{context} @{thm list_eq.intros(4)} 1 *})
done
lemma "x = xa \<Longrightarrow> INSERT a x = INSERT a xa"
apply (tactic {* lift_tac_fset @{context} @{thm list_eq.intros(5)} 1 *})
done
lemma "CARD x = Suc n \<Longrightarrow> (\<exists>a b. \<not> IN a b & x = INSERT a b)"
apply (tactic {* lift_tac_fset @{context} @{thm card1_suc} 1 *})
done
lemma "(\<not> IN x xa) = (CARD (INSERT x xa) = Suc (CARD xa))"
apply (tactic {* lift_tac_fset @{context} @{thm not_mem_card1} 1 *})
done
lemma "\<forall>f g z a x. FOLD f g (z::'b) (INSERT a x) =
(if rsp_fold f then if IN a x then FOLD f g z x else f (g a) (FOLD f g z x) else z)"
apply (tactic {* lift_tac_fset @{context} @{thm fold1.simps(2)} 1 *})
done
lemma "fmap f (FUNION (x::'b fset) (xa::'b fset)) = FUNION (fmap f x) (fmap f xa)"
apply (tactic {* lift_tac_fset @{context} @{thm map_append} 1 *})
done
lemma "FUNION (FUNION x xa) xb = FUNION x (FUNION xa xb)"
apply (tactic {* lift_tac_fset @{context} @{thm append_assoc} 1 *})
done
ML {* val aps = findaps rty (prop_of (atomize_thm @{thm list.induct})) *}
ML {* val app_prs_thms = map (applic_prs @{context} rty qty absrep) aps *}
lemma "\<lbrakk>P EMPTY; \<And>a x. P x \<Longrightarrow> P (INSERT a x)\<rbrakk> \<Longrightarrow> P l"
apply (tactic {* lift_tac_fset @{context} @{thm list.induct} 1 *})
apply (tactic {* REPEAT_ALL_NEW (EqSubst.eqsubst_tac @{context} [0] app_prs_thms) 1 *})
apply (tactic {* simp_tac (HOL_ss addsimps [reps_same]) 1 *})
done
quotient_def
fset_rec::"'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"
where
"fset_rec \<equiv> list_rec"
quotient_def
fset_case::"'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"
where
"fset_case \<equiv> list_case"
(* Probably not true without additional assumptions about the function *)
lemma list_rec_rsp:
"(op = ===> (op = ===> op \<approx> ===> op =) ===> op \<approx> ===> op =) list_rec list_rec"
apply (auto simp add: FUN_REL_EQ)
apply (erule_tac list_eq.induct)
apply (simp_all)
sorry
lemma list_case_rsp:
"(op = ===> (op = ===> op \<approx> ===> op =) ===> op \<approx> ===> op =) list_case list_case"
apply (auto simp add: FUN_REL_EQ)
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_thm_fset lthy t = lift_thm lthy qty "fset" rsp_thms defs t *}
ML {* fun lift_thm_g_fset lthy t g = lift_thm_goal lthy qty "fset" rsp_thms defs t g *}
ML {* fun lift_tac_fset lthy t = lift_tac lthy t rel_eqv rel_refl rty quot trans2 rsp_thms reps_same defs *}
ML {* val aps = findaps rty (prop_of (atomize_thm @{thm list.recs(2)})) *}
ML {* val app_prs_thms = map (applic_prs @{context} rty qty absrep) aps *}
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 *})
apply (tactic {* REPEAT_ALL_NEW (EqSubst.eqsubst_tac @{context} [0] app_prs_thms) 1 *})
apply (tactic {* (simp_tac (HOL_ss addsimps
@{thms eq_reflection[OF FUN_MAP_I] eq_reflection[OF id_apply] id_def_sym prod_fun_id map_id})) 1 *})
ML_prf {* val lower = flat (map (add_lower_defs @{context}) defs) *}
apply (tactic {* REPEAT_ALL_NEW (EqSubst.eqsubst_tac @{context} [0] lower) 1 *})
apply (tactic {* simp_tac (HOL_ss addsimps [reps_same]) 1 *})
done
ML {* val aps = findaps rty (prop_of (atomize_thm @{thm list.cases(2)})) *}
ML {* val app_prs_thms = map (applic_prs @{context} rty qty absrep) aps *}
ML {* map (lift_thm_fset @{context}) @{thms list.cases(2)} *}
lemma "fset_case (f1::'t) f2 (INSERT a xa) = f2 a xa"
apply (tactic {* lift_tac_fset @{context} @{thm list.cases(2)} 1 *})
apply (tactic {* REPEAT_ALL_NEW (EqSubst.eqsubst_tac @{context} [0] app_prs_thms) 1 *})
apply (tactic {* simp_tac (HOL_ss addsimps [reps_same]) 1 *})
done
lemma list_induct_part:
assumes a: "P (x :: 'a list) ([] :: 'a list)"
assumes b: "\<And>e t. P x t \<Longrightarrow> P x (e # t)"
shows "P x l"
apply (rule_tac P="P x" in list.induct)
apply (rule a)
apply (rule b)
apply (assumption)
done
(* Construction site starts here *)
ML {* val consts = lookup_quot_consts defs *}
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 t_a = atomize_thm @{thm list_induct_part} *}
(* prove {* build_regularize_goal t_a rty rel @{context} *}
ML_prf {* fun tac ctxt = FIRST' [
rtac rel_refl,
atac,
rtac @{thm universal_twice},
(rtac @{thm impI} THEN' atac),
rtac @{thm implication_twice},
//comented out rtac @{thm equality_twice}, //
EqSubst.eqsubst_tac ctxt [0]
[(@{thm equiv_res_forall} OF [rel_eqv]),
(@{thm equiv_res_exists} OF [rel_eqv])],
(rtac @{thm impI} THEN' (asm_full_simp_tac (Simplifier.context ctxt HOL_ss)) THEN' rtac rel_refl),
(rtac @{thm RIGHT_RES_FORALL_REGULAR})
]; *}
apply (atomize(full))
apply (tactic {* REPEAT_ALL_NEW (tac @{context}) 1 *})
done *)
ML {* val t_r = regularize t_a rty rel rel_eqv rel_refl @{context} *}
ML {*
val rt = build_repabs_term @{context} t_r consts rty qty
val rg = Logic.mk_equals ((Thm.prop_of t_r), rt);
*}
prove {* Syntax.check_term @{context} rg *}
ML_prf {* fun r_mk_comb_tac_fset lthy = r_mk_comb_tac lthy rty quot rel_refl trans2 rsp_thms *}
apply(atomize(full))
apply (tactic {* REPEAT_ALL_NEW (r_mk_comb_tac_fset @{context}) 1 *})
done
ML {*
val t_t = repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms
*}
ML {* val abs = findabs rty (prop_of (t_a)) *}
ML {* val aps = findaps rty (prop_of (t_a)) *}
ML {* val lam_prs_thms = map (make_simp_prs_thm @{context} quot @{thm LAMBDA_PRS}) abs *}
ML {* val app_prs_thms = map (applic_prs @{context} rty qty absrep) aps *}
ML {* val lam_prs_thms = map Thm.varifyT lam_prs_thms *}
ML {* t_t *}
ML {* val (alls, exs) = findallex @{context} rty qty (prop_of t_a); *}
ML {* val allthms = map (make_allex_prs_thm @{context} quot @{thm FORALL_PRS}) alls *}
ML {* val t_l0 = repeat_eqsubst_thm @{context} (app_prs_thms) t_t *}
ML app_prs_thms
ML {* val t_l = repeat_eqsubst_thm @{context} (lam_prs_thms) t_l0 *}
ML lam_prs_thms
ML {* val t_id = simp_ids @{context} t_l *}
thm INSERT_def
ML {* val defs_sym = flat (map (add_lower_defs @{context}) defs) *}
ML {* val t_d = repeat_eqsubst_thm @{context} defs_sym t_id *}
ML allthms
thm FORALL_PRS
ML {* val t_al = MetaSimplifier.rewrite_rule (allthms) t_d *}
ML {* val t_s = MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_fset.REPS_same} t_al *}
ML {* ObjectLogic.rulify t_s *}
ML {* val gl = @{term "P (x :: 'a list) (EMPTY :: 'a fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"} *}
ML {* val gla = atomize_goal @{theory} gl *}
prove t_r: {* mk_REGULARIZE_goal @{context} (prop_of t_a) gla *}
ML_prf {* fun tac ctxt = FIRST' [
rtac rel_refl,
atac,
rtac @{thm universal_twice},
(rtac @{thm impI} THEN' atac),
rtac @{thm implication_twice},
(*rtac @{thm equality_twice},*)
EqSubst.eqsubst_tac ctxt [0]
[(@{thm equiv_res_forall} OF [rel_eqv]),
(@{thm equiv_res_exists} OF [rel_eqv])],
(rtac @{thm impI} THEN' (asm_full_simp_tac (Simplifier.context ctxt HOL_ss)) THEN' rtac rel_refl),
(rtac @{thm RIGHT_RES_FORALL_REGULAR})
]; *}
apply (atomize(full))
apply (tactic {* REPEAT_ALL_NEW (tac @{context}) 1 *})
done
ML {* val t_r = @{thm t_r} OF [t_a] *}
ML {* val ttt = mk_inj_REPABS_goal @{context} (prop_of t_r, gla) *}
ML {* val si = simp_ids_trm (cterm_of @{theory} ttt) *}
prove t_t: {* term_of si *}
ML_prf {* fun r_mk_comb_tac_fset lthy = r_mk_comb_tac lthy rty quot rel_refl trans2 rsp_thms *}
apply(atomize(full))
apply (tactic {* (APPLY_RSP_TAC rty @{context}) 1 *})
apply (rule FUN_QUOTIENT)
apply (rule FUN_QUOTIENT)
apply (rule IDENTITY_QUOTIENT)
apply (rule FUN_QUOTIENT)
apply (rule QUOTIENT_fset)
apply (rule IDENTITY_QUOTIENT)
apply (rule IDENTITY_QUOTIENT)
apply (rule IDENTITY_QUOTIENT)
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 {* (APPLY_RSP_TAC rty @{context}) 1 *})
apply (rule IDENTITY_QUOTIENT)
apply (rule IDENTITY_QUOTIENT)
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 {* (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 {* (r_mk_comb_tac_fset @{context}) 1 *})
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
apply (tactic {* (APPLY_RSP_TAC rty @{context}) 1 *})
apply (rule IDENTITY_QUOTIENT)
apply (rule FUN_QUOTIENT)
apply (rule QUOTIENT_fset)
apply (rule IDENTITY_QUOTIENT)
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 {* (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 {* (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 {* (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 {* (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 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 {* (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 {* (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 *})
done
thm t_t
ML {* val t_t = @{thm Pure.equal_elim_rule1} OF [@{thm t_t}, t_r] *}
ML {* val t_l = repeat_eqsubst_thm @{context} (lam_prs_thms) t_t *}
ML {* val t_d = repeat_eqsubst_thm @{context} defs_sym t_l *}
ML {* val t_al = MetaSimplifier.rewrite_rule (allthms) t_d *}
ML {* val t_s = MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_fset.REPS_same} t_al *}
end