diff -r c0ab7451b20d -r 751d1349329b Nominal/Ex/ExCoreHaskell.thy --- a/Nominal/Ex/ExCoreHaskell.thy Mon May 10 15:44:49 2010 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,675 +0,0 @@ -theory ExCoreHaskell -imports "../NewParser" -begin - -(* core haskell *) - -atom_decl var -atom_decl cvar -atom_decl tvar - -(* there are types, coercion types and regular types list-data-structure *) - -nominal_datatype tkind = - KStar -| KFun "tkind" "tkind" -and ckind = - CKEq "ty" "ty" -and ty = - TVar "tvar" -| TC "string" -| TApp "ty" "ty" -| TFun "string" "ty_lst" -| TAll tv::"tvar" "tkind" T::"ty" bind_set tv in T -| TEq "ckind" "ty" -and ty_lst = - TsNil -| TsCons "ty" "ty_lst" -and co = - CVar "cvar" -| CConst "string" -| CApp "co" "co" -| CFun "string" "co_lst" -| CAll cv::"cvar" "ckind" C::"co" bind_set cv in C -| CEq "ckind" "co" -| CRefl "ty" -| CSym "co" -| CCir "co" "co" -| CAt "co" "ty" -| CLeft "co" -| CRight "co" -| CSim "co" "co" -| CRightc "co" -| CLeftc "co" -| CCoe "co" "co" -and co_lst = - CsNil -| CsCons "co" "co_lst" -and trm = - Var "var" -| K "string" -| LAMT tv::"tvar" "tkind" t::"trm" bind_set tv in t -| LAMC cv::"cvar" "ckind" t::"trm" bind_set cv in t -| AppT "trm" "ty" -| AppC "trm" "co" -| Lam v::"var" "ty" t::"trm" bind_set v in t -| App "trm" "trm" -| Let x::"var" "ty" "trm" t::"trm" bind_set x in t -| Case "trm" "assoc_lst" -| Cast "trm" "ty" --"ty is supposed to be a coercion type only" -and assoc_lst = - ANil -| ACons p::"pat" t::"trm" "assoc_lst" bind "bv p" in t -and pat = - 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_vs :: "vars \ atom list" -and bv_tvs :: "tvars \ atom list" -and bv_cvs :: "cvars \ atom list" -where - "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_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_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_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 (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_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_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 only: alpha_intros perm permute_bv_tvs_raw.simps permute_bv_cvs_raw.simps permute_bv_vs_raw.simps) - done - -lemma [quot_respect]: - "(op = ===> alpha_pat_raw ===> alpha_pat_raw) permute_bv_raw permute_bv_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) - apply (simp_all) - apply (rule alpha_intros) - apply (simp_all add: rsp_pre) - done - -thm permute_bv_raw.simps[no_vars] - 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 (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_vs_raw.simps[quot_lifted] - permute_bv_cvs_raw.simps[quot_lifted] - permute_bv_tvs_raw.simps[quot_lifted] - -lemma perm_bv1: - "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(simp_all add:permute_bv eqvts) - apply(induct c rule: inducts(11)) - apply(simp_all add:permute_bv eqvts) - apply(induct d rule: inducts(10)) - apply(simp_all add:permute_bv eqvts) - done - -lemma perm_bv2: - "p \ bv l = bv (permute_bv p l)" - apply(induct l rule: inducts(9)) - apply(simp_all add:permute_bv) - apply(simp add: perm_bv1[symmetric]) - apply(simp add: eqvts) - done - -lemma alpha_perm_bn1: - "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 cvars rule: inducts(12)) - apply(simp_all add:permute_bv eqvts eq_iff) - apply(induct vars rule: inducts(10)) - apply(simp_all add:permute_bv eqvts eq_iff) - done - -lemma alpha_perm_bn: - "alpha_bv pat (permute_bv q pat)" - apply(induct pat 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" - apply (simp only: eq_iff) - apply (simp add: alpha_perm_bn) - apply (rule_tac x="q" in exI) - apply (simp add: alphas) - apply (simp add: perm_bv2[symmetric]) - apply (simp add: eqvts[symmetric]) - apply (simp add: supp_abs) - apply (simp add: fv_supp) - apply (rule supp_perm_eq[symmetric]) - apply (subst supp_finite_atom_set) - apply (rule finite_Diff) - apply (simp add: finite_supp) - apply (assumption) - done - -lemma permute_bv_zero1: - "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(simp_all add:permute_bv eqvts) - apply(induct c rule: inducts(11)) - apply(simp_all add:permute_bv eqvts) - apply(induct d rule: inducts(10)) - apply(simp_all add:permute_bv eqvts) - done - -lemma permute_bv_zero2: - "permute_bv 0 a = a" - apply(induct a rule: inducts(9)) - apply(simp_all add:permute_bv eqvts permute_bv_zero1) - done - -lemma fv_alpha1: "fv_bv_tvs x \* pa \ alpha_bv_tvs (pa \ x) x" - apply (induct x rule: inducts(11)) - apply (simp_all add: eq_iff fresh_star_union) - done - -lemma fv_alpha2: "fv_bv_cvs x \* pa \ alpha_bv_cvs (pa \ x) x" -apply (induct x rule: inducts(12)) -apply (rule TrueI)+ -apply (simp_all add: eq_iff fresh_star_union) -apply (subst supp_perm_eq) -apply (simp_all add: fv_supp) -done - -lemma fv_alpha3: "fv_bv_vs x \* pa \ alpha_bv_vs (pa \ x) x" -apply (induct x rule: inducts(10)) -apply (rule TrueI)+ -apply (simp_all add: fresh_star_union eq_iff) -apply (subst supp_perm_eq) -apply (simp_all add: fv_supp) -done - -lemma fv_alpha: "fv_bv x \* pa \ alpha_bv (pa \ x) x" -apply (induct x rule: inducts(9)) -apply (rule TrueI)+ -apply (simp_all add: eq_iff fresh_star_union) -apply (simp add: fv_alpha1 fv_alpha2 fv_alpha3) -apply (subst supp_perm_eq) -apply (simp_all add: fv_supp) -done - -lemma fin1: "finite (fv_bv_tvs x)" -apply (induct x rule: inducts(11)) -apply (simp_all add: fv_supp finite_supp) -done - -lemma fin2: "finite (fv_bv_cvs x)" -apply (induct x rule: inducts(12)) -apply (simp_all add: fv_supp finite_supp) -done - -lemma fin3: "finite (fv_bv_vs x)" -apply (induct x rule: inducts(10)) -apply (simp_all add: fv_supp finite_supp) -done - -lemma fin_fv_bv: "finite (fv_bv x)" -apply (induct x rule: inducts(9)) -apply (rule TrueI)+ -defer -apply (rule TrueI)+ -apply (simp add: fin1 fin2 fin3) -apply (rule finite_supp) -done - -lemma finb1: "finite (set (bv_tvs x))" -apply (induct x rule: inducts(11)) -apply (simp_all add: fv_supp finite_supp) -done - -lemma finb2: "finite (set (bv_cvs x))" -apply (induct x rule: inducts(12)) -apply (simp_all add: fv_supp finite_supp) -done - -lemma finb3: "finite (set (bv_vs x))" -apply (induct x rule: inducts(10)) -apply (simp_all add: fv_supp finite_supp) -done - -lemma fin_bv: "finite (set (bv x))" -apply (induct x rule: inducts(9)) -apply (simp_all add: finb1 finb2 finb3) -done - -lemma strong_induction_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)" - and a03: "\ty1 ty2 b. \\c. P3 c ty1; \c. P3 c ty2\ \ P2 b (CKEq ty1 ty2)" - and a04: "\tvar b. P3 b (TVar tvar)" - and a05: "\string b. P3 b (TC string)" - and a06: "\ty1 ty2 b. \\c. P3 c ty1; \c. P3 c ty2\ \ P3 b (TApp ty1 ty2)" - and a07: "\string ty_lst b. \\c. P4 c ty_lst\ \ P3 b (TFun string ty_lst)" - 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: "\ck ty b. \\c. P2 c ck; \c. P3 c ty\ \ P3 b (TEq ck ty)" - 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 a12: "\string b. P5 b (CVar string)" - and a12a:"\str b. P5 b (CConst str)" - 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)" - and a15: "\tvar ckind co b. \\c. P2 c ckind; \c. P5 c co; atom tvar \ b\ - \ P5 b (CAll tvar ckind co)" - and a16: "\ck co b. \\c. P2 c ck; \c. P5 c co\ \ P5 b (CEq ck co)" - and a17: "\ty b. \\c. P3 c ty\ \ P5 b (CRefl ty)" - and a17a: "\co b. \\c. P5 c co\ \ P5 b (CSym co)" - and a18: "\co1 co2 b. \\c. P5 c co1; \c. P5 c co2\ \ P5 b (CCir co1 co2)" - and a18a:"\co ty b. \\c. P5 c co; \c. P3 c ty\ \ P5 b (CAt co ty)" - and a19: "\co b. \\c. P5 c co\ \ P5 b (CLeft co)" - and a20: "\co b. \\c. P5 c co\ \ P5 b (CRight co)" - and a21: "\co1 co2 b. \\c. P5 c co1; \c. P5 c co2\ \ P5 b (CSim co1 co2)" - and a22: "\co b. \\c. P5 c co\ \ P5 b (CRightc co)" - and a23: "\co b. \\c. P5 c co\ \ P5 b (CLeftc co)" - and a24: "\co1 co2 b. \\c. P5 c co1; \c. P5 c co2\ \ P5 b (CCoe co1 co2)" - and a25: "\b. P6 b CsNil" - and a26: "\co co_lst b. \\c. P5 c co; \c. P6 c co_lst\ \ P6 b (CsCons co co_lst)" - and a27: "\var b. P7 b (Var var)" - and a28: "\string b. P7 b (K string)" - and a29: "\tvar tkind trm b. \\c. P1 c tkind; \c. P7 c trm; atom tvar \ b\ - \ P7 b (LAMT tvar tkind trm)" - and a30: "\tvar ckind trm b. \\c. P2 c ckind; \c. P7 c trm; atom tvar \ b\ - \ P7 b (LAMC tvar ckind trm)" - and a31: "\trm ty b. \\c. P7 c trm; \c. P3 c ty\ \ P7 b (AppT trm ty)" - and a31a:"\trm co b. \\c. P7 c trm; \c. P5 c co\ \ P7 b (AppC trm co)" - and a32: "\var ty trm b. \\c. P3 c ty; \c. P7 c trm; atom var \ b\ \ P7 b (Lam var ty trm)" - and a33: "\trm1 trm2 b. \\c. P7 c trm1; \c. P7 c trm2\ \ P7 b (App trm1 trm2)" - and a34: "\var ty trm1 trm2 b. \\c. P3 c ty; \c. P7 c trm1; \c. P7 c trm2; atom var \ b\ - \ P7 b (Let var ty trm1 trm2)" - 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 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" - 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 \ - P4 (d :: 'd :: pt) ty_lst \ - P5 (e :: 'e :: {pt,fs}) co \ - P6 (f :: 'f :: pt) co_lst \ - P7 (g :: 'g :: {pt,fs}) trm \ - P8 (h :: 'h :: {pt,fs}) assoc_lst \ - P9 (i :: 'i :: pt) pat \ - 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)))" - 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})) *}) - -(* GOAL1 *) - apply(subgoal_tac "\pa. ((pa \ (atom (p \ tvar))) \ c \ - supp (Abs (p \ {atom tvar}) (p \ ty)) \* pa)") - apply clarify - apply (simp only: perm) - apply(rule_tac t="TAll (p \ tvar) (p \ tkind) (p \ ty)" - and s="TAll (pa \ p \ tvar) (p \ tkind) (pa \ p \ ty)" in subst) - apply (simp add: eq_iff) - apply (rule_tac x="-pa" in exI) - apply (simp add: alphas eqvts eqvts_raw supp_abs fv_supp) - apply (rule_tac t="supp (pa \ p \ ty) - {atom (pa \ p \ tvar)}" - and s="pa \ (p \ supp ty - {p \ atom tvar})" in subst) - apply (simp add: eqvts) - apply (simp add: eqvts[symmetric]) - apply (rule conjI) - apply (rule supp_perm_eq) - apply (simp add: eqvts) - apply (subst supp_finite_atom_set) - apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp) - apply (simp add: eqvts) - apply (simp add: eqvts) - apply (subst supp_perm_eq) - apply (subst supp_finite_atom_set) - apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp) - apply (simp add: eqvts) - apply assumption - apply (simp add: fresh_star_minus_perm) - apply (rule a08) - apply simp - apply(rotate_tac 1) - apply(erule_tac x="(pa + p)" in allE) - apply simp - apply (simp add: eqvts eqvts_raw) - apply (rule at_set_avoiding2_atom) - apply (simp add: finite_supp) - apply (simp add: finite_supp) - apply (simp add: fresh_def) - apply (simp only: supp_abs eqvts) - apply blast - -(* GOAL2 *) - apply(subgoal_tac "\pa. ((pa \ (atom (p \ cvar))) \ e \ - supp (Abs (p \ {atom cvar}) (p \ co)) \* pa)") - apply clarify - apply (simp only: perm) - apply(rule_tac t="CAll (p \ cvar) (p \ ckind) (p \ co)" - and s="CAll (pa \ p \ cvar) (p \ ckind) (pa \ p \ co)" in subst) - apply (simp add: eq_iff) - apply (rule_tac x="-pa" in exI) - apply (simp add: alphas eqvts eqvts_raw supp_abs fv_supp) - apply (rule_tac t="supp (pa \ p \ co) - {atom (pa \ p \ cvar)}" - and s="pa \ (p \ supp co - {p \ atom cvar})" in subst) - apply (simp add: eqvts) - apply (simp add: eqvts[symmetric]) - apply (rule conjI) - apply (rule supp_perm_eq) - apply (simp add: eqvts) - apply (subst supp_finite_atom_set) - apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp) - apply (simp add: eqvts) - apply (simp add: eqvts) - apply (subst supp_perm_eq) - apply (subst supp_finite_atom_set) - apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp) - apply (simp add: eqvts) - apply assumption - apply (simp add: fresh_star_minus_perm) - apply (rule a15) - apply simp - apply(rotate_tac 1) - apply(erule_tac x="(pa + p)" in allE) - apply simp - apply (simp add: eqvts eqvts_raw) - apply (rule at_set_avoiding2_atom) - apply (simp add: finite_supp) - apply (simp add: finite_supp) - apply (simp add: fresh_def) - apply (simp only: supp_abs eqvts) - apply blast - - -(* GOAL3 a copy-and-paste of Goal2 with consts and variable names changed *) - apply(subgoal_tac "\pa. ((pa \ (atom (p \ tvar))) \ g \ - supp (Abs (p \ {atom tvar}) (p \ trm)) \* pa)") - apply clarify - apply (simp only: perm) - apply(rule_tac t="LAMT (p \ tvar) (p \ tkind) (p \ trm)" - and s="LAMT (pa \ p \ tvar) (p \ tkind) (pa \ p \ trm)" in subst) - apply (simp add: eq_iff) - apply (rule_tac x="-pa" in exI) - apply (simp add: alphas eqvts eqvts_raw supp_abs fv_supp) - apply (rule_tac t="supp (pa \ p \ trm) - {atom (pa \ p \ tvar)}" - and s="pa \ (p \ supp trm - {p \ atom tvar})" in subst) - apply (simp add: eqvts) - apply (simp add: eqvts[symmetric]) - apply (rule conjI) - apply (rule supp_perm_eq) - apply (simp add: eqvts) - apply (subst supp_finite_atom_set) - apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp) - apply (simp add: eqvts) - apply (simp add: eqvts) - apply (subst supp_perm_eq) - apply (subst supp_finite_atom_set) - apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp) - apply (simp add: eqvts) - apply assumption - apply (simp add: fresh_star_minus_perm) - apply (rule a29) - apply simp - apply(rotate_tac 1) - apply(erule_tac x="(pa + p)" in allE) - apply simp - apply (simp add: eqvts eqvts_raw) - apply (rule at_set_avoiding2_atom) - apply (simp add: finite_supp) - apply (simp add: finite_supp) - apply (simp add: fresh_def) - apply (simp only: supp_abs eqvts) - apply blast - -(* GOAL4 a copy-and-paste *) - apply(subgoal_tac "\pa. ((pa \ (atom (p \ cvar))) \ g \ - supp (Abs (p \ {atom cvar}) (p \ trm)) \* pa)") - apply clarify - apply (simp only: perm) - apply(rule_tac t="LAMC (p \ cvar) (p \ ckind) (p \ trm)" - and s="LAMC (pa \ p \ cvar) (p \ ckind) (pa \ p \ trm)" in subst) - apply (simp add: eq_iff) - apply (rule_tac x="-pa" in exI) - apply (simp add: alphas eqvts eqvts_raw supp_abs fv_supp) - apply (rule_tac t="supp (pa \ p \ trm) - {atom (pa \ p \ cvar)}" - and s="pa \ (p \ supp trm - {p \ atom cvar})" in subst) - apply (simp add: eqvts) - apply (simp add: eqvts[symmetric]) - apply (rule conjI) - apply (rule supp_perm_eq) - apply (simp add: eqvts) - apply (subst supp_finite_atom_set) - apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp) - apply (simp add: eqvts) - apply (simp add: eqvts) - apply (subst supp_perm_eq) - apply (subst supp_finite_atom_set) - apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp) - apply (simp add: eqvts) - apply assumption - apply (simp add: fresh_star_minus_perm) - apply (rule a30) - apply simp - apply(rotate_tac 1) - apply(erule_tac x="(pa + p)" in allE) - apply simp - apply (simp add: eqvts eqvts_raw) - apply (rule at_set_avoiding2_atom) - apply (simp add: finite_supp) - apply (simp add: finite_supp) - apply (simp add: fresh_def) - apply (simp only: supp_abs eqvts) - apply blast - - -(* GOAL5 a copy-and-paste *) - apply(subgoal_tac "\pa. ((pa \ (atom (p \ var))) \ g \ - supp (Abs (p \ {atom var}) (p \ trm)) \* pa)") - apply clarify - apply (simp only: perm) - apply(rule_tac t="Lam (p \ var) (p \ ty) (p \ trm)" - and s="Lam (pa \ p \ var) (p \ ty) (pa \ p \ trm)" in subst) - apply (simp add: eq_iff) - apply (rule_tac x="-pa" in exI) - apply (simp add: alphas eqvts eqvts_raw supp_abs fv_supp) - apply (rule_tac t="supp (pa \ p \ trm) - {atom (pa \ p \ var)}" - and s="pa \ (p \ supp trm - {p \ atom var})" in subst) - apply (simp add: eqvts) - apply (simp add: eqvts[symmetric]) - apply (rule conjI) - apply (rule supp_perm_eq) - apply (simp add: eqvts) - apply (subst supp_finite_atom_set) - apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp) - apply (simp add: eqvts) - apply (simp add: eqvts) - apply (subst supp_perm_eq) - apply (subst supp_finite_atom_set) - apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp) - apply (simp add: eqvts) - apply assumption - apply (simp add: fresh_star_minus_perm) - apply (rule a32) - apply simp - apply(rotate_tac 1) - apply(erule_tac x="(pa + p)" in allE) - apply simp - apply (simp add: eqvts eqvts_raw) - apply (rule at_set_avoiding2_atom) - apply (simp add: finite_supp) - apply (simp add: finite_supp) - apply (simp add: fresh_def) - apply (simp only: supp_abs eqvts) - apply blast - - -(* GOAL6 a copy-and-paste *) - apply(subgoal_tac "\pa. ((pa \ (atom (p \ var))) \ g \ - supp (Abs (p \ {atom var}) (p \ trm2)) \* pa)") - apply clarify - apply (simp only: perm) - apply(rule_tac t="Let (p \ var) (p \ ty) (p \ trm1) (p \ trm2)" - and s="Let (pa \ p \ var) (p \ ty) (p \ trm1) (pa \ p \ trm2)" in subst) - apply (simp add: eq_iff) - apply (rule_tac x="-pa" in exI) - apply (simp add: alphas eqvts eqvts_raw supp_abs fv_supp) - apply (rule_tac t="supp (pa \ p \ trm2) - {atom (pa \ p \ var)}" - and s="pa \ (p \ supp trm2 - {p \ atom var})" in subst) - apply (simp add: eqvts) - apply (simp add: eqvts[symmetric]) - apply (rule conjI) - apply (rule supp_perm_eq) - apply (simp add: eqvts) - apply (subst supp_finite_atom_set) - apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp) - apply (simp add: eqvts) - apply (simp add: eqvts) - apply (subst supp_perm_eq) - apply (subst supp_finite_atom_set) - apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp) - apply (simp add: eqvts) - apply assumption - apply (simp add: fresh_star_minus_perm) - apply (rule a34) - apply simp - apply simp - apply(rotate_tac 2) - apply(erule_tac x="(pa + p)" in allE) - apply simp - apply (simp add: eqvts eqvts_raw) - apply (rule at_set_avoiding2_atom) - apply (simp add: finite_supp) - apply (simp add: finite_supp) - apply (simp add: fresh_def) - apply (simp only: supp_abs eqvts) - apply blast - -(* MAIN ACons Goal *) - 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) - apply assumption - apply (rule a38) - apply simp - apply(rotate_tac 1) - apply(erule_tac x="(pa + p)" in allE) - apply simp - apply simp - apply (simp add: perm_bv2[symmetric]) - apply (simp add: eqvts eqvts_raw) - apply (rule at_set_avoiding2) - apply (simp add: fin_bv) - apply (simp add: finite_supp) - apply (simp add: supp_abs) - apply (simp add: finite_supp) - 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+) - ultimately show ?thesis by (simp_all add: permute_bv_zero1 permute_bv_zero2) -qed - -section {* test about equivariance for alpha *} - -(* this should not be an equivariance rule *) -(* for the moment, we force it to be *) - -(*declare permute_pure[eqvt]*) -(*setup {* Context.theory_map (Nominal_ThmDecls.add_thm @{thm "permute_pure"}) *) - -thm eqvts -thm eqvts_raw - -declare permute_tkind_raw_permute_ckind_raw_permute_ty_raw_permute_ty_lst_raw_permute_co_raw_permute_co_lst_raw_permute_trm_raw_permute_assoc_lst_raw_permute_pat_raw_permute_vars_raw_permute_tvars_raw_permute_cvars_raw.simps[eqvt] -declare alpha_gen_eqvt[eqvt] - -equivariance alpha_tkind_raw - -thm eqvts -thm eqvts_raw - -end -