diff -r b9292bbcffb6 -r c39d4fe31100 Nominal/Ex/CoreHaskell.thy --- a/Nominal/Ex/CoreHaskell.thy Sun May 16 11:00:44 2010 +0100 +++ b/Nominal/Ex/CoreHaskell.thy Sun May 16 12:41:27 2010 +0100 @@ -84,7 +84,7 @@ | "bv_tvs TvsNil = []" | "bv_tvs (TvsCons v k t) = (atom v) # bv_tvs t" | "bv_cvs CvsNil = []" -| "bv_cvs (CvsCons v k t) = (atom v) # bv_cvs t" +| "bv_cvs (CvsCons v k t) = (atom v) # bv_cvs t" lemmas fv_supp=tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vars_tvars_cvars.supp(1-9,11,13,15) lemmas supp=tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vars_tvars_cvars.fv[simplified fv_supp] @@ -191,13 +191,13 @@ done lemma alpha_perm_bn: - "alpha_bv pat (permute_bv q pat)" - apply(induct pat rule: inducts(9)) + "alpha_bv pt (permute_bv q pt)" + apply(induct pt rule: inducts(9)) apply(simp_all add:permute_bv eqvts eq_iff alpha_perm_bn1) done lemma ACons_subst: - "supp (Abs_lst (bv pat) trm) \* q \ (ACons pat trm al) = ACons (permute_bv q pat) (q \ trm) al" + "supp (Abs_lst (bv pt) trm) \* q \ (ACons pt trm al) = ACons (permute_bv q pt) (q \ trm) al" apply (simp only: eq_iff) apply (simp add: alpha_perm_bn) apply (rule_tac x="q" in exI) @@ -352,8 +352,8 @@ 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; set (bv (pat)) \* b\ - \ P8 b (ACons pat trm assoc_lst)" + and a38: "\pt trm assoc_lst b. \\c. P9 c pt; \c. P7 c trm; \c. P8 c assoc_lst; set (bv (pt)) \* b\ + \ P8 b (ACons pt trm assoc_lst)" and a39: "\string tvars cvars vars b. \\c. P11 c tvars; \c. P12 c cvars; \c. P10 c vars\ \ P9 b (Kpat string tvars cvars vars)" and a40: "\b. P10 b VsNil" @@ -372,12 +372,12 @@ P6 (f :: 'f :: pt) co_lst \ P7 (g :: 'g :: {pt,fs}) trm \ P8 (h :: 'h :: {pt,fs}) assoc_lst \ - P9 (i :: 'i :: pt) pat \ + P9 (i :: 'i :: pt) pt \ P10 (j :: 'j :: pt) vars \ P11 (k :: 'k :: pt) tvars \ P12 (l :: 'l :: pt) cvars" proof - - have a1: "(\p a. P1 a (p \ tkind))" and "(\p b. P2 b (p \ ckind))" and "(\p c. P3 c (p \ ty))" and "(\p d. P4 d (p \ ty_lst))" and "(\p e. P5 e (p \ co))" and " (\p f. P6 f (p \ co_lst))" and "(\p g. P7 g (p \ trm))" and "(\p h. P8 h (p \ assoc_lst))" and a1:"(\p q i. P9 i (permute_bv p (q \ pat)))" and a2:"(\p q j. P10 j (permute_bv_vs q (p \ vars)))" and a3:"(\p q k. P11 k ( permute_bv_tvs q (p \ tvars)))" and a4:"(\p q l. P12 l (permute_bv_cvs q (p \ cvars)))" + have a1: "(\p a. P1 a (p \ tkind))" and "(\p b. P2 b (p \ ckind))" and "(\p c. P3 c (p \ ty))" and "(\p d. P4 d (p \ ty_lst))" and "(\p e. P5 e (p \ co))" and " (\p f. P6 f (p \ co_lst))" and "(\p g. P7 g (p \ trm))" and "(\p h. P8 h (p \ assoc_lst))" and a1:"(\p q i. P9 i (permute_bv p (q \ pt)))" and a2:"(\p q j. P10 j (permute_bv_vs q (p \ vars)))" and a3:"(\p q k. P11 k ( permute_bv_tvs q (p \ tvars)))" and a4:"(\p q l. P12 l (permute_bv_cvs q (p \ cvars)))" apply (induct rule: tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vars_tvars_cvars.inducts) apply (tactic {* ALLGOALS (REPEAT o rtac allI) *}) apply (tactic {* ALLGOALS (TRY o SOLVED' (simp_tac @{simpset} THEN_ALL_NEW resolve_tac @{thms assms} THEN_ALL_NEW asm_full_simp_tac @{simpset})) *}) @@ -637,7 +637,7 @@ apply (simp add: fresh_star_def fresh_def supp_abs eqvts) done then have b: "P1 a (0 \ tkind)" and "P2 b (0 \ ckind)" "P3 c (0 \ ty)" and "P4 d (0 \ ty_lst)" and "P5 e (0 \ co)" and "P6 f (0 \ co_lst)" and "P7 g (0 \ trm)" and "P8 h (0 \ assoc_lst)" by (blast+) - moreover have "P9 i (permute_bv 0 (0 \ pat))" and "P10 j (permute_bv_vs 0 (0 \ vars))" and "P11 k (permute_bv_tvs 0 (0 \ tvars))" and "P12 l (permute_bv_cvs 0 (0 \ cvars))" using a1 a2 a3 a4 by (blast+) + moreover have "P9 i (permute_bv 0 (0 \ pt))" and "P10 j (permute_bv_vs 0 (0 \ vars))" and "P11 k (permute_bv_tvs 0 (0 \ tvars))" and "P12 l (permute_bv_cvs 0 (0 \ cvars))" using a1 a2 a3 a4 by (blast+) ultimately show ?thesis by (simp_all add: permute_bv_zero1 permute_bv_zero2) qed