# HG changeset patch # User Cezary Kaliszyk # Date 1269872990 -7200 # Node ID e42ee5947b5c3dc06eba7a81834266a921ab7c03 # Parent 15d2d46bf89e91f2895d4657b9a388dea37fccc3 Changed to Lists. diff -r 15d2d46bf89e -r e42ee5947b5c Nominal/ExCoreHaskell.thy --- a/Nominal/ExCoreHaskell.thy Mon Mar 29 12:06:22 2010 +0200 +++ b/Nominal/ExCoreHaskell.thy Mon Mar 29 16:29:50 2010 +0200 @@ -11,7 +11,7 @@ (* there are types, coercion types and regular types list-data-structure *) -ML {* val _ = alpha_type := AlphaGen *} +ML {* val _ = alpha_type := AlphaLst *} nominal_datatype tkind = KStar | KFun "tkind" "tkind" @@ -71,18 +71,18 @@ TVCKNil | TVCKCons "tvar" "ckind" "tvck_lst" binder - bv :: "pat \ atom set" -and bv_vt :: "vt_lst \ atom set" -and bv_tvtk :: "tvtk_lst \ atom set" -and bv_tvck :: "tvck_lst \ atom set" + bv :: "pat \ atom list" +and bv_vt :: "vt_lst \ atom list" +and bv_tvtk :: "tvtk_lst \ atom list" +and bv_tvck :: "tvck_lst \ atom list" where - "bv (K s tvts tvcs vs) = (bv_tvtk tvts) \ (bv_tvck tvcs) \ (bv_vt vs)" -| "bv_vt VTNil = {}" -| "bv_vt (VTCons v k t) = {atom v} \ bv_vt t" -| "bv_tvtk TVTKNil = {}" -| "bv_tvtk (TVTKCons v k t) = {atom v} \ bv_tvtk t" -| "bv_tvck TVCKNil = {}" -| "bv_tvck (TVCKCons v k t) = {atom v} \ bv_tvck t" + "bv (K s tvts tvcs vs) = append (bv_tvtk tvts) (append (bv_tvck tvcs) (bv_vt vs))" +| "bv_vt VTNil = []" +| "bv_vt (VTCons v k t) = (atom v) # bv_vt t" +| "bv_tvtk TVTKNil = []" +| "bv_tvtk (TVTKCons v k t) = (atom v) # bv_tvtk t" +| "bv_tvck TVCKNil = []" +| "bv_tvck (TVCKCons v k t) = (atom v) # bv_tvck t" lemmas fv_supp=tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vt_lst_tvtk_lst_tvck_lst.supp(1-9,11,13,15) lemmas supp=tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vt_lst_tvtk_lst_tvck_lst.fv[simplified fv_supp] @@ -201,7 +201,7 @@ done lemma ACons_subst: - "supp (Abs (bv pat) trm) \* q \ (ACons pat trm al) = ACons (permute_bv q pat) (q \ trm) al" + "supp (Abs_lst (bv pat) trm) \* q \ (ACons pat trm al) = ACons (permute_bv q pat) (q \ trm) al" apply (simp only: eq_iff) apply (rule_tac x="q" in exI) apply (simp add: alphas) @@ -295,36 +295,30 @@ apply (simp add: fin1 fin2 fin3) done -lemma finb1: "finite (bv_tvtk x)" +lemma finb1: "finite (set (bv_tvtk x))" apply (induct x rule: inducts(11)) apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *}) apply (simp_all add: fv_supp finite_supp) done -lemma finb2: "finite (bv_tvck x)" +lemma finb2: "finite (set (bv_tvck x))" apply (induct x rule: inducts(12)) apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *}) apply (simp_all add: fv_supp finite_supp) done -lemma finb3: "finite (bv_vt x)" +lemma finb3: "finite (set (bv_vt x))" apply (induct x rule: inducts(10)) apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *}) apply (simp_all add: fv_supp finite_supp) done -lemma fin_bv: "finite (bv x)" +lemma fin_bv: "finite (set (bv x))" apply (induct x rule: inducts(9)) apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *}) apply (simp add: finb1 finb2 finb3) done -lemma "bv x \* fv_bv x" -apply (induct x rule: inducts(9)) -apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *}) -apply (simp) -oops - lemma strong_inudction_principle: assumes a01: "\b. P1 b KStar" and a02: "\tkind1 tkind2 b. \\c. P1 c tkind1; \c. P1 c tkind2\ \ P1 b (KFun tkind1 tkind2)" @@ -368,7 +362,7 @@ and a35: "\trm assoc_lst b. \\c. P7 c trm; \c. P8 c assoc_lst\ \ P7 b (Case trm assoc_lst)" and a36: "\trm ty b. \\c. P7 c trm; \c. P3 c ty\ \ P7 b (Cast trm ty)" and a37: "\b. P8 b ANil" - and a38: "\pat trm assoc_lst b. \\c. P9 c pat; \c. P7 c trm; \c. P8 c assoc_lst; bv(pat) \* b\ + and a38: "\pat trm assoc_lst b. \\c. P9 c pat; \c. P7 c trm; \c. P8 c assoc_lst; set (bv (pat)) \* b\ \ P8 b (ACons pat trm assoc_lst)" and a39: "\string tvtk_lst tvck_lst vt_lst b. \\c. P11 c tvtk_lst; \c. P12 c tvck_lst; \c. P10 c vt_lst\ \ P9 b (K string tvtk_lst tvck_lst vt_lst)" @@ -643,8 +637,8 @@ apply blast (* MAIN ACons Goal *) - apply(subgoal_tac "\pa. ((pa \ (bv (p \ pat))) \* h \ - supp (Abs (p \ (bv pat)) (p \ trm)) \* pa)") + apply(subgoal_tac "\pa. ((pa \ (set (bv (p \ pat)))) \* h \ + supp (Abs_lst (p \ (bv pat)) (p \ trm)) \* pa)") apply clarify apply (simp only: perm eqvts) apply (subst ACons_subst) @@ -661,7 +655,6 @@ apply (simp add: fin_bv) apply (simp add: finite_supp) apply (simp add: supp_abs) - apply (rule finite_Diff) apply (simp add: finite_supp) apply (simp add: fresh_star_def fresh_def supp_abs eqvts) done