# HG changeset patch # User Cezary Kaliszyk # Date 1260078112 -3600 # Node ID 06e54c862a390afd36e481cf803b33e526792a68 # Parent 14682786c3564f20104c3a7b6fd4c2d933f2321d# Parent a68c51dd85b3de96f429cd3935b4ac1bc58c1287 merge diff -r a68c51dd85b3 -r 06e54c862a39 LFex.thy --- a/LFex.thy Sun Dec 06 04:03:08 2009 +0100 +++ b/LFex.thy Sun Dec 06 06:41:52 2009 +0100 @@ -260,7 +260,7 @@ apply(tactic {* procedure_tac @{context} @{thm akind_aty_atrm.induct} 1 *}) apply(tactic {* regularize_tac @{context} @{thms alpha_equivps} 1 *}) apply(tactic {* all_inj_repabs_tac @{context} rel_refl trans2 1 *}) -apply(fold perm_kind_def perm_ty_def perm_trm_def FV_ty_def[simplified id_simps] FV_kind_def[simplified id_simps] FV_trm_def[simplified id_simps]) +apply(fold perm_kind_def perm_ty_def perm_trm_def) apply(tactic {* clean_tac @{context} 1 *}) (* Profiling: diff -r a68c51dd85b3 -r 06e54c862a39 LamEx.thy --- a/LamEx.thy Sun Dec 06 04:03:08 2009 +0100 +++ b/LamEx.thy Sun Dec 06 06:41:52 2009 +0100 @@ -219,20 +219,14 @@ lemma fv_var: "fv (Var (a\name)) = {a}" apply (tactic {* lift_tac_lam @{context} @{thm rfv_var} 1 *}) -apply (unfold fv_def[simplified id_simps]) -apply (tactic {* clean_tac @{context} 1 *}) done lemma fv_app: "fv (App (x\lam) (xa\lam)) = fv x \ fv xa" apply (tactic {* lift_tac_lam @{context} @{thm rfv_app} 1 *}) -apply (unfold fv_def[simplified id_simps]) -apply (tactic {* clean_tac @{context} 1 *}) done lemma fv_lam: "fv (Lam (a\name) (x\lam)) = fv x - {a}" apply (tactic {* lift_tac_lam @{context} @{thm rfv_lam} 1 *}) -apply (unfold fv_def[simplified id_simps]) -apply (tactic {* clean_tac @{context} 1 *}) done lemma a1: "(a\name) = (b\name) \ Var a = Var b" @@ -245,8 +239,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 (unfold fv_def[simplified id_simps]) -apply (tactic {* clean_tac @{context} 1 *}) done lemma alpha_cases: "\a1 = a2; \a b. \a1 = Var a; a2 = Var b; a = b\ \ P; @@ -254,8 +246,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 (unfold fv_def[simplified id_simps]) -apply (tactic {* clean_tac @{context} 1 *}) done lemma alpha_induct: "\(qx\lam) = (qxa\lam); \(a\name) b\name. a = b \ (qxb\lam \ lam \ bool) (Var a) (Var b); @@ -264,8 +254,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 (unfold fv_def[simplified id_simps]) -apply (tactic {* clean_tac @{context} 1 *}) done lemma var_inject: "(Var a = Var b) = (a = b)" diff -r a68c51dd85b3 -r 06e54c862a39 quotient_def.ML --- a/quotient_def.ML Sun Dec 06 04:03:08 2009 +0100 +++ b/quotient_def.ML Sun Dec 06 06:41:52 2009 +0100 @@ -81,17 +81,7 @@ fun make_def qconst_bname qty mx attr rhs lthy = let - val rty = fastype_of rhs - val (arg_rtys, res_rty) = strip_type rty - val (arg_qtys, res_qty) = strip_type qty - - val rep_fns = map (get_fun repF lthy) (arg_rtys ~~ arg_qtys) - val abs_fn = get_fun absF lthy (res_rty, res_qty) - - fun mk_fun_map t s = - Const (@{const_name "fun_map"}, dummyT) $ t $ s - - val absrep_trm = (fold_rev mk_fun_map rep_fns abs_fn $ rhs) + val absrep_trm = get_fun absF lthy (fastype_of rhs, qty) $ rhs |> Syntax.check_term lthy val ((trm, thm), lthy') = define qconst_bname mx attr absrep_trm lthy