diff -r bb23a7393de3 -r 6cdba30c6d66 LamEx.thy --- a/LamEx.thy Thu Dec 03 14:00:43 2009 +0100 +++ b/LamEx.thy Thu Dec 03 14:02:05 2009 +0100 @@ -170,13 +170,11 @@ done ML {* val qty = @{typ "lam"} *} -ML {* val defs = @{thms Var_def App_def Lam_def perm_lam_def fv_def} *} ML {* val rsp_thms = @{thms perm_rsp fresh_rsp rVar_rsp rApp_rsp rLam_rsp rfv_rsp} *} 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] defs *} +ML {* fun lift_tac_lam lthy t = lift_tac lthy t [rel_eqv] [quot] *} lemma pi_var: "(pi\('x \ 'x) list) \ Var a = Var (pi \ a)" apply (tactic {* lift_tac_lam @{context} @{thm pi_var_com} 1 *}) @@ -337,7 +335,7 @@ prefer 2 apply (tactic {* all_inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *}) prefer 3 -apply (tactic {* clean_tac @{context} [quot] defs 1 *}) +apply (tactic {* clean_tac @{context} [quot] 1 *}) thm all_prs thm REP_ABS_RSP