Simpler definition code that works with any type maps.
authorCezary Kaliszyk <kaliszyk@in.tum.de>
Sun, 06 Dec 2009 06:39:32 +0100
changeset 573 14682786c356
parent 567 5dffcd087e30
child 574 06e54c862a39
Simpler definition code that works with any type maps.
LFex.thy
LamEx.thy
quotient_def.ML
--- a/LFex.thy	Sun Dec 06 00:00:47 2009 +0100
+++ b/LFex.thy	Sun Dec 06 06:39:32 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:
--- a/LamEx.thy	Sun Dec 06 00:00:47 2009 +0100
+++ b/LamEx.thy	Sun Dec 06 06:39:32 2009 +0100
@@ -190,20 +190,14 @@
 
 lemma fv_var: "fv (Var (a\<Colon>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\<Colon>lam) (xa\<Colon>lam)) = fv x \<union> 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\<Colon>name) (x\<Colon>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\<Colon>name) = (b\<Colon>name) \<Longrightarrow> Var a = Var b"
@@ -216,8 +210,6 @@
 
 lemma a3: "\<lbrakk>(x\<Colon>lam) = [(a\<Colon>name, b\<Colon>name)] \<bullet> (xa\<Colon>lam); a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> 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: "\<lbrakk>a1 = a2; \<And>a b. \<lbrakk>a1 = Var a; a2 = Var b; a = b\<rbrakk> \<Longrightarrow> P;
@@ -225,8 +217,6 @@
      \<And>x a b xa. \<lbrakk>a1 = Lam a x; a2 = Lam b xa; x = [(a, b)] \<bullet> xa; a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> P\<rbrakk>
     \<Longrightarrow> 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: "\<lbrakk>(qx\<Colon>lam) = (qxa\<Colon>lam); \<And>(a\<Colon>name) b\<Colon>name. a = b \<Longrightarrow> (qxb\<Colon>lam \<Rightarrow> lam \<Rightarrow> bool) (Var a) (Var b);
@@ -235,8 +225,6 @@
         \<lbrakk>x = [(a, b)] \<bullet> xa; qxb x ([(a, b)] \<bullet> xa); a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> qxb (Lam a x) (Lam b xa)\<rbrakk>
     \<Longrightarrow> 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)"
--- a/quotient_def.ML	Sun Dec 06 00:00:47 2009 +0100
+++ b/quotient_def.ML	Sun Dec 06 06:39:32 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