# HG changeset patch # User Cezary Kaliszyk # Date 1269876737 -7200 # Node ID b5141d1ab24f36f6ff4a75d97d62442375f3e07a # Parent 2dca07aca28745801079dc53d9727cc100b16215 Initial renaming diff -r 2dca07aca287 -r b5141d1ab24f Nominal/ExCoreHaskell.thy --- a/Nominal/ExCoreHaskell.thy Mon Mar 29 17:14:02 2010 +0200 +++ b/Nominal/ExCoreHaskell.thy Mon Mar 29 17:32:17 2010 +0200 @@ -64,69 +64,69 @@ ANil | ACons p::"pat" t::"trm" "assoc_lst" bind "bv p" in t and pat = - Kpat "string" "tvtk_lst" "tvck_lst" "vt_lst" -and vt_lst = - VTNil -| VTCons "var" "ty" "vt_lst" -and tvtk_lst = - TVTKNil -| TVTKCons "tvar" "tkind" "tvtk_lst" -and tvck_lst = - TVCKNil -| TVCKCons "cvar" "ckind" "tvck_lst" + Kpat "string" "tvars" "cvars" "vars" +and vars = + VsNil +| VsCons "var" "ty" "vars" +and tvars = + TvsNil +| TvsCons "tvar" "tkind" "tvars" +and cvars = + CvsNil +| CvsCons "cvar" "ckind" "cvars" binder 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" +and bv_vs :: "vars \ atom list" +and bv_tvs :: "tvars \ atom list" +and bv_cvs :: "cvars \ atom list" where - "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" + "bv (Kpat s tvts tvcs vs) = append (bv_tvs tvts) (append (bv_cvs tvcs) (bv_vs vs))" +| "bv_vs VsNil = []" +| "bv_vs (VsCons v k t) = (atom v) # bv_vs t" +| "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" -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] -lemmas perm=tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vt_lst_tvtk_lst_tvck_lst.perm -lemmas eq_iff=tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vt_lst_tvtk_lst_tvck_lst.eq_iff -lemmas inducts=tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vt_lst_tvtk_lst_tvck_lst.inducts +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] +lemmas perm=tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vars_tvars_cvars.perm +lemmas eq_iff=tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vars_tvars_cvars.eq_iff +lemmas inducts=tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vars_tvars_cvars.inducts -lemmas alpha_inducts=alpha_tkind_raw_alpha_ckind_raw_alpha_ty_raw_alpha_ty_lst_raw_alpha_co_raw_alpha_co_lst_raw_alpha_trm_raw_alpha_assoc_lst_raw_alpha_pat_raw_alpha_vt_lst_raw_alpha_tvtk_lst_raw_alpha_tvck_lst_raw_alpha_bv_raw_alpha_bv_vt_raw_alpha_bv_tvtk_raw_alpha_bv_tvck_raw.inducts -lemmas alpha_intros=alpha_tkind_raw_alpha_ckind_raw_alpha_ty_raw_alpha_ty_lst_raw_alpha_co_raw_alpha_co_lst_raw_alpha_trm_raw_alpha_assoc_lst_raw_alpha_pat_raw_alpha_vt_lst_raw_alpha_tvtk_lst_raw_alpha_tvck_lst_raw_alpha_bv_raw_alpha_bv_vt_raw_alpha_bv_tvtk_raw_alpha_bv_tvck_raw.intros +lemmas alpha_inducts=alpha_tkind_raw_alpha_ckind_raw_alpha_ty_raw_alpha_ty_lst_raw_alpha_co_raw_alpha_co_lst_raw_alpha_trm_raw_alpha_assoc_lst_raw_alpha_pat_raw_alpha_vars_raw_alpha_tvars_raw_alpha_cvars_raw_alpha_bv_raw_alpha_bv_vs_raw_alpha_bv_tvs_raw_alpha_bv_cvs_raw.inducts +lemmas alpha_intros=alpha_tkind_raw_alpha_ckind_raw_alpha_ty_raw_alpha_ty_lst_raw_alpha_co_raw_alpha_co_lst_raw_alpha_trm_raw_alpha_assoc_lst_raw_alpha_pat_raw_alpha_vars_raw_alpha_tvars_raw_alpha_cvars_raw_alpha_bv_raw_alpha_bv_vs_raw_alpha_bv_tvs_raw_alpha_bv_cvs_raw.intros lemma fresh_star_minus_perm: "as \* - p = as \* (p :: perm)" unfolding fresh_star_def Ball_def by auto (simp_all add: fresh_minus_perm) -primrec permute_bv_vt_raw -where "permute_bv_vt_raw p VTNil_raw = VTNil_raw" -| "permute_bv_vt_raw p (VTCons_raw v t l) = VTCons_raw (p \ v) t (permute_bv_vt_raw p l)" -primrec permute_bv_tvck_raw -where "permute_bv_tvck_raw p TVCKNil_raw = TVCKNil_raw" -| "permute_bv_tvck_raw p (TVCKCons_raw v t l) = TVCKCons_raw (p \ v) t (permute_bv_tvck_raw p l)" -primrec permute_bv_tvtk_raw -where "permute_bv_tvtk_raw p TVTKNil_raw = TVTKNil_raw" -| "permute_bv_tvtk_raw p (TVTKCons_raw v t l) = TVTKCons_raw (p \ v) t (permute_bv_tvtk_raw p l)" +primrec permute_bv_vs_raw +where "permute_bv_vs_raw p VsNil_raw = VsNil_raw" +| "permute_bv_vs_raw p (VsCons_raw v t l) = VsCons_raw (p \ v) t (permute_bv_vs_raw p l)" +primrec permute_bv_cvs_raw +where "permute_bv_cvs_raw p CvsNil_raw = CvsNil_raw" +| "permute_bv_cvs_raw p (CvsCons_raw v t l) = CvsCons_raw (p \ v) t (permute_bv_cvs_raw p l)" +primrec permute_bv_tvs_raw +where "permute_bv_tvs_raw p TvsNil_raw = TvsNil_raw" +| "permute_bv_tvs_raw p (TvsCons_raw v t l) = TvsCons_raw (p \ v) t (permute_bv_tvs_raw p l)" primrec permute_bv_raw -where "permute_bv_raw p (K_raw c l1 l2 l3) = - K_raw c (permute_bv_tvtk_raw p l1) (permute_bv_tvck_raw p l2) (permute_bv_vt_raw p l3)" +where "permute_bv_raw p (Kpat_raw c l1 l2 l3) = + Kpat_raw c (permute_bv_tvs_raw p l1) (permute_bv_cvs_raw p l2) (permute_bv_vs_raw p l3)" -quotient_definition "permute_bv_vt :: perm \ vt_lst \ vt_lst" -is "permute_bv_vt_raw" -quotient_definition "permute_bv_tvck :: perm \ tvck_lst \ tvck_lst" -is "permute_bv_tvck_raw" -quotient_definition "permute_bv_tvtk :: perm \ tvtk_lst \ tvtk_lst" -is "permute_bv_tvtk_raw" +quotient_definition "permute_bv_vs :: perm \ vars \ vars" +is "permute_bv_vs_raw" +quotient_definition "permute_bv_cvs :: perm \ cvars \ cvars" +is "permute_bv_cvs_raw" +quotient_definition "permute_bv_tvs :: perm \ tvars \ tvars" +is "permute_bv_tvs_raw" quotient_definition "permute_bv :: perm \ pat \ pat" is "permute_bv_raw" lemma rsp_pre: - "alpha_tvtk_lst_raw d a \ alpha_tvtk_lst_raw (permute_bv_tvtk_raw x d) (permute_bv_tvtk_raw x a)" - "alpha_tvck_lst_raw e b \ alpha_tvck_lst_raw (permute_bv_tvck_raw x e) (permute_bv_tvck_raw x b)" - "alpha_vt_lst_raw f c \ alpha_vt_lst_raw (permute_bv_vt_raw x f) (permute_bv_vt_raw x c)" + "alpha_tvars_raw d a \ alpha_tvars_raw (permute_bv_tvs_raw x d) (permute_bv_tvs_raw x a)" + "alpha_cvars_raw e b \ alpha_cvars_raw (permute_bv_cvs_raw x e) (permute_bv_cvs_raw x b)" + "alpha_vars_raw f c \ alpha_vars_raw (permute_bv_vs_raw x f) (permute_bv_vs_raw x c)" apply (erule_tac [!] alpha_inducts) apply simp_all apply (rule_tac [!] alpha_intros) @@ -135,9 +135,9 @@ lemma [quot_respect]: "(op = ===> alpha_pat_raw ===> alpha_pat_raw) permute_bv_raw permute_bv_raw" - "(op = ===> alpha_tvtk_lst_raw ===> alpha_tvtk_lst_raw) permute_bv_tvtk_raw permute_bv_tvtk_raw" - "(op = ===> alpha_tvck_lst_raw ===> alpha_tvck_lst_raw) permute_bv_tvck_raw permute_bv_tvck_raw" - "(op = ===> alpha_vt_lst_raw ===> alpha_vt_lst_raw) permute_bv_vt_raw permute_bv_vt_raw" + "(op = ===> alpha_tvars_raw ===> alpha_tvars_raw) permute_bv_tvs_raw permute_bv_tvs_raw" + "(op = ===> alpha_cvars_raw ===> alpha_cvars_raw) permute_bv_cvs_raw permute_bv_cvs_raw" + "(op = ===> alpha_vars_raw ===> alpha_vars_raw) permute_bv_vs_raw permute_bv_vs_raw" apply (simp_all add: rsp_pre) apply clarify apply (erule_tac alpha_inducts) @@ -147,25 +147,25 @@ done thm permute_bv_raw.simps[no_vars] - permute_bv_vt_raw.simps[quot_lifted] - permute_bv_tvck_raw.simps[quot_lifted] - permute_bv_tvtk_raw.simps[quot_lifted] + permute_bv_vs_raw.simps[quot_lifted] + permute_bv_cvs_raw.simps[quot_lifted] + permute_bv_tvs_raw.simps[quot_lifted] lemma permute_bv_pre: - "permute_bv p (K c l1 l2 l3) = - K c (permute_bv_tvtk p l1) (permute_bv_tvck p l2) (permute_bv_vt p l3)" + "permute_bv p (Kpat c l1 l2 l3) = + Kpat c (permute_bv_tvs p l1) (permute_bv_cvs p l2) (permute_bv_vs p l3)" by (lifting permute_bv_raw.simps) lemmas permute_bv[simp] = permute_bv_pre - permute_bv_vt_raw.simps[quot_lifted] - permute_bv_tvck_raw.simps[quot_lifted] - permute_bv_tvtk_raw.simps[quot_lifted] + permute_bv_vs_raw.simps[quot_lifted] + permute_bv_cvs_raw.simps[quot_lifted] + permute_bv_tvs_raw.simps[quot_lifted] lemma perm_bv1: - "p \ bv_tvck b = bv_tvck (permute_bv_tvck p b)" - "p \ bv_tvtk c = bv_tvtk (permute_bv_tvtk p c)" - "p \ bv_vt d = bv_vt (permute_bv_vt p d)" + "p \ bv_cvs b = bv_cvs (permute_bv_cvs p b)" + "p \ bv_tvs c = bv_tvs (permute_bv_tvs p c)" + "p \ bv_vs d = bv_vs (permute_bv_vs p d)" apply(induct b rule: inducts(12)) apply(rule TrueI) apply(simp_all add:permute_bv eqvts) @@ -187,14 +187,14 @@ done lemma alpha_perm_bn1: - " alpha_bv_tvtk tvtk_lst (permute_bv_tvtk q tvtk_lst)" - "alpha_bv_tvck tvck_lst (permute_bv_tvck q tvck_lst)" - "alpha_bv_vt vt_lst (permute_bv_vt q vt_lst)" - apply(induct tvtk_lst rule: inducts(11)) + " alpha_bv_tvs tvars (permute_bv_tvs q tvars)" + "alpha_bv_cvs cvars (permute_bv_cvs q cvars)" + "alpha_bv_vs vars (permute_bv_vs q vars)" + apply(induct tvars rule: inducts(11)) apply(simp_all add:permute_bv eqvts eq_iff) - apply(induct tvck_lst rule: inducts(12)) + apply(induct cvars rule: inducts(12)) apply(simp_all add:permute_bv eqvts eq_iff) - apply(induct vt_lst rule: inducts(10)) + apply(induct vars rule: inducts(10)) apply(simp_all add:permute_bv eqvts eq_iff) done @@ -222,9 +222,9 @@ done lemma permute_bv_zero1: - "permute_bv_tvck 0 b = b" - "permute_bv_tvtk 0 c = c" - "permute_bv_vt 0 d = d" + "permute_bv_cvs 0 b = b" + "permute_bv_tvs 0 c = c" + "permute_bv_vs 0 d = d" apply(induct b rule: inducts(12)) apply(rule TrueI) apply(simp_all add:permute_bv eqvts) @@ -243,7 +243,7 @@ apply(simp_all add:permute_bv eqvts permute_bv_zero1) done -lemma fv_alpha1: "fv_bv_tvtk x \* pa \ alpha_bv_tvtk (pa \ x) x" +lemma fv_alpha1: "fv_bv_tvs x \* pa \ alpha_bv_tvs (pa \ x) x" apply (induct x rule: inducts(11)) apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *}) apply (simp_all add: eq_iff fresh_star_union) @@ -251,7 +251,7 @@ apply (simp_all add: fv_supp) done -lemma fv_alpha2: "fv_bv_tvck x \* pa \ alpha_bv_tvck (pa \ x) x" +lemma fv_alpha2: "fv_bv_cvs x \* pa \ alpha_bv_cvs (pa \ x) x" apply (induct x rule: inducts(12)) apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *}) apply (simp_all add: eq_iff fresh_star_union) @@ -259,7 +259,7 @@ apply (simp_all add: fv_supp) done -lemma fv_alpha3: "fv_bv_vt x \* pa \ alpha_bv_vt (pa \ x) x" +lemma fv_alpha3: "fv_bv_vs x \* pa \ alpha_bv_vs (pa \ x) x" apply (induct x rule: inducts(10)) apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *}) apply (simp_all add: eq_iff fresh_star_union) @@ -275,19 +275,19 @@ apply (simp add: eqvts) done -lemma fin1: "finite (fv_bv_tvtk x)" +lemma fin1: "finite (fv_bv_tvs 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 fin2: "finite (fv_bv_tvck x)" +lemma fin2: "finite (fv_bv_cvs 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 fin3: "finite (fv_bv_vt x)" +lemma fin3: "finite (fv_bv_vs x)" apply (induct x rule: inducts(10)) apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *}) apply (simp_all add: fv_supp finite_supp) @@ -299,19 +299,19 @@ apply (simp add: fin1 fin2 fin3) done -lemma finb1: "finite (set (bv_tvtk x))" +lemma finb1: "finite (set (bv_tvs 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 (set (bv_tvck x))" +lemma finb2: "finite (set (bv_cvs 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 (set (bv_vt x))" +lemma finb3: "finite (set (bv_vs x))" apply (induct x rule: inducts(10)) apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *}) apply (simp_all add: fv_supp finite_supp) @@ -334,8 +334,8 @@ and a08: "\tvar tkind ty b. \\c. P1 c tkind; \c. P3 c ty; atom tvar \ b\ \ P3 b (TAll tvar tkind ty)" and a09: "\ty1 ty2 ty3 b. \\c. P3 c ty1; \c. P3 c ty2; \c. P3 c ty3\ \ P3 b (TEq ty1 ty2 ty3)" - and a10: "\b. P4 b TsNil" - and a11: "\ty ty_lst b. \\c. P3 c ty; \c. P4 c ty_lst\ \ P4 b (TsCons ty ty_lst)" + and a10: "\b. P4 b TvsNil" + and a11: "\ty ty_lst b. \\c. P3 c ty; \c. P4 c ty_lst\ \ P4 b (TvsCons ty ty_lst)" and a12: "\string b. P5 b (CC string)" and a13: "\co1 co2 b. \\c. P5 c co1; \c. P5 c co2\ \ P5 b (CApp co1 co2)" and a14: "\string co_lst b. \\c. P6 c co_lst\ \ P5 b (CFun string co_lst)" @@ -368,16 +368,16 @@ 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 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)" - and a40: "\b. P10 b VTNil" - and a41: "\var ty vt_lst b. \\c. P3 c ty; \c. P10 c vt_lst\ \ P10 b (VTCons var ty vt_lst)" - and a42: "\b. P11 b TVTKNil" - and a43: "\tvar tkind tvtk_lst b. \\c. P1 c tkind; \c. P11 c tvtk_lst\ - \ P11 b (TVTKCons tvar tkind tvtk_lst)" - and a44: "\b. P12 b TVCKNil" - and a45: "\tvar ckind tvck_lst b. \\c. P2 c ckind; \c. P12 c tvck_lst\ - \ P12 b (TVCKCons tvar ckind tvck_lst)" + and a39: "\string tvars cvars vars b. \\c. P11 c tvars; \c. P12 c cvars; \c. P10 c vars\ + \ P9 b (K string tvars cvars vars)" + and a40: "\b. P10 b VsNil" + and a41: "\var ty vars b. \\c. P3 c ty; \c. P10 c vars\ \ P10 b (VsCons var ty vars)" + and a42: "\b. P11 b TvsNil" + and a43: "\tvar tkind tvars b. \\c. P1 c tkind; \c. P11 c tvars\ + \ P11 b (TvsCons tvar tkind tvars)" + and a44: "\b. P12 b CvsNil" + and a45: "\tvar ckind cvars b. \\c. P2 c ckind; \c. P12 c cvars\ + \ P12 b (CvsCons tvar ckind cvars)" shows "P1 (a :: 'a :: pt) tkind \ P2 (b :: 'b :: pt) ckind \ P3 (c :: 'c :: {pt,fs}) ty \ @@ -387,12 +387,12 @@ P7 (g :: 'g :: {pt,fs}) trm \ P8 (h :: 'h :: {pt,fs}) assoc_lst \ P9 (i :: 'i :: pt) pat \ - P10 (j :: 'j :: pt) vt_lst \ - P11 (k :: 'k :: pt) tvtk_lst \ - P12 (l :: 'l :: pt) tvck_lst" + 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_vt q (p \ vt_lst)))" and a3:"(\p q k. P11 k ( permute_bv_tvtk q (p \ tvtk_lst)))" and a4:"(\p q l. P12 l (permute_bv_tvck q (p \ tvck_lst)))" - apply (induct rule: tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vt_lst_tvtk_lst_tvck_lst.inducts) + 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)))" + 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})) *}) @@ -663,7 +663,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_vt 0 (0 \ vt_lst))" and "P11 k (permute_bv_tvtk 0 (0 \ tvtk_lst))" and "P12 l (permute_bv_tvck 0 (0 \ tvck_lst))" using a1 a2 a3 a4 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+) ultimately show ?thesis by (simp_all add: permute_bv_zero1 permute_bv_zero2) qed