# HG changeset patch # User Cezary Kaliszyk # Date 1259845019 -3600 # Node ID 375e28eedee7fead2abb6036a23652186952d14f # Parent 184d74813679254e248e8ea3af50dc60069e9f09 Reintroduced varifyT; we still need it for permutation definition. diff -r 184d74813679 -r 375e28eedee7 LFex.thy --- a/LFex.thy Thu Dec 03 13:45:52 2009 +0100 +++ b/LFex.thy Thu Dec 03 13:56:59 2009 +0100 @@ -216,11 +216,6 @@ thm akind_aty_atrm.induct thm kind_ty_trm.induct -ML {* val defs = - @{thms TYP_def KPI_def TCONST_def TAPP_def TPI_def VAR_def CONS_def APP_def LAM_def - FV_kind_def FV_ty_def FV_trm_def perm_kind_def perm_ty_def perm_trm_def} -*} - ML {* val quot = @{thms QUOTIENT_KIND QUOTIENT_TY QUOTIENT_TRM} val rel_refl = map (fn x => @{thm EQUIV_REFL} OF [x]) @{thms alpha_EQUIVs} @@ -266,7 +261,6 @@ apply(tactic {* procedure_tac @{context} @{thm akind_aty_atrm.induct} 1 *}) apply(tactic {* regularize_tac @{context} @{thms alpha_EQUIVs} 1 *}) prefer 2 -apply(tactic {* simp_tac ((Simplifier.context @{context} empty_ss) addsimps @{thms perm_kind_def perm_ty_def perm_trm_def}) 1 *}) apply(tactic {* clean_tac @{context} quot 1 *}) (* Profiling: diff -r 184d74813679 -r 375e28eedee7 LamEx.thy --- a/LamEx.thy Thu Dec 03 13:45:52 2009 +0100 +++ b/LamEx.thy Thu Dec 03 13:56:59 2009 +0100 @@ -210,7 +210,6 @@ lemma a3: "\(x\lam) = [(a\name, b\name)] \ (xa\lam); a \ fv (Lam b x)\ \ Lam a x = Lam b xa" apply (tactic {* lift_tac_lam @{context} @{thm a3} 1 *}) -apply (simp add:perm_lam_def) done lemma alpha_cases: "\a1 = a2; \a b. \a1 = Var a; a2 = Var b; a = b\ \ P; @@ -218,7 +217,6 @@ \x a b xa. \a1 = Lam a x; a2 = Lam b xa; x = [(a, b)] \ xa; a \ fv (Lam b x)\ \ P\ \ P" apply (tactic {* lift_tac_lam @{context} @{thm alpha.cases} 1 *}) -apply (simp add:perm_lam_def) done lemma alpha_induct: "\(qx\lam) = (qxa\lam); \(a\name) b\name. a = b \ (qxb\lam \ lam \ bool) (Var a) (Var b); @@ -227,7 +225,6 @@ \x = [(a, b)] \ xa; qxb x ([(a, b)] \ xa); a \ fv (Lam b x)\ \ qxb (Lam a x) (Lam b xa)\ \ qxb qx qxa" apply (tactic {* lift_tac_lam @{context} @{thm alpha.induct} 1 *}) -apply (simp add:perm_lam_def) done lemma var_inject: "(Var a = Var b) = (a = b)" diff -r 184d74813679 -r 375e28eedee7 QuotMain.thy --- a/QuotMain.thy Thu Dec 03 13:45:52 2009 +0100 +++ b/QuotMain.thy Thu Dec 03 13:56:59 2009 +0100 @@ -1091,7 +1091,7 @@ fun clean_tac lthy quot = let val thy = ProofContext.theory_of lthy; - val defs = map (#def) (qconsts_dest thy); + val defs = map (Thm.varifyT o #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