diff -r 09cd71fac4ec -r 287ea842a7d4 QuotMain.thy --- a/QuotMain.thy Sat Dec 05 16:26:18 2009 +0100 +++ b/QuotMain.thy Sat Dec 05 21:44:01 2009 +0100 @@ -1137,25 +1137,31 @@ The rest are a simp_tac and are fast. *) +ML {* fun OF1 thm1 thm2 = thm2 RS thm1 *} +ML {* +fun quotient_solver_tac ss = quotient_tac (Simplifier.the_context ss) +val quotient_solver = Simplifier.mk_solver' "Quotient goal solver" quotient_solver_tac +*} + ML {* fun clean_tac lthy = let val thy = ProofContext.theory_of lthy; - val quotients = quotient_rules_get lthy - val defs = map (Thm.varifyT o #def) (qconsts_dest thy); - val absrep = map (fn x => @{thm Quotient_abs_rep} OF [x]) quotients - val reps_same = map (fn x => @{thm Quotient_rel_rep} OF [x]) quotients - val simp_ctxt = (Simplifier.context lthy empty_ss) addsimps - (@{thm eq_reflection[OF fun_map.simps]} :: @{thms id_simps} @ absrep @ reps_same @ defs) + val defs = map (Thm.varifyT o #def) (qconsts_dest thy) + val thms = @{thms eq_reflection[OF fun_map.simps]} + @ @{thms id_simps Quotient_abs_rep Quotient_rel_rep} + @ defs + val simp_ctxt = HOL_basic_ss addsimps thms addSolver quotient_solver in EVERY' [ - (* (Rep1 ---> Abs2) (\x. Rep2 (f (Abs1 x))) ----> f *) - (* \x\Respects R. (abs ---> id) f ----> \x. f *) + (* (rep1 ---> abs2) (\x. rep2 (f (abs1 x))) ----> f *) + (* Ball (Respects R) ((abs ---> id) f) ----> All f *) NDT lthy "a" (TRY o lambda_allex_prs_tac lthy), - (* NewConst ----> (rep ---> abs) oldConst *) - (* Abs (Rep x) ----> x *) - (* id_simps; fun_map.simps *) + (* NewConst -----> (rep ---> abs) oldConst *) + (* abs (rep x) -----> x *) + (* R (Rep x) (Rep y) -----> x = y *) + (* id_simps; fun_map.simps *) NDT lthy "c" (TRY o simp_tac simp_ctxt), (* final step *)