# HG changeset patch # User Christian Urban # Date 1316593386 -7200 # Node ID 6fd3fc3254ee5c3597ef80a415afeb92ee045345 # Parent aa5059a00f41215c37d7188776aeed256688874e deleted PNil diff -r aa5059a00f41 -r 6fd3fc3254ee LMCS-Paper/Paper.thy --- a/LMCS-Paper/Paper.thy Wed Sep 21 17:16:11 2011 +0900 +++ b/LMCS-Paper/Paper.thy Wed Sep 21 10:23:06 2011 +0200 @@ -1121,12 +1121,10 @@ \hspace{5mm}$\mid$~@{text "Let_pat p::pat trm t::trm"} \;\;\isacommand{binds} @{text "bn(p)"} \isacommand{in} @{text t}\\ \isacommand{and} @{text pat} $=$\\ - \hspace{5mm}\phantom{$\mid$}~@{text PNil}\\ - \hspace{5mm}$\mid$~@{text "PVar name"}\\ + \hspace{5mm}\phantom{$\mid$}~@{text "PVar name"}\\ \hspace{5mm}$\mid$~@{text "PTup pat pat"}\\ \isacommand{binder}~@{text "bn::pat \ atom list"}\\ - \isacommand{where}~@{text "bn(PNil) = []"}\\ - \hspace{5mm}$\mid$~@{text "bn(PVar x) = [atom x]"}\\ + \isacommand{where}~@{text "bn(PVar x) = [atom x]"}\\ \hspace{5mm}$\mid$~@{text "bn(PTup p\<^isub>1 p\<^isub>2) = bn(p\<^isub>1) @ bn(p\<^isub>2)"}\smallskip\\ \end{tabular}} \end{equation}\smallskip @@ -1972,7 +1970,6 @@ \infer{@{text "P\<^bsub>pat\<^esub>"}} {\begin{array}{@ {}l@ {}} - @{text "y = PNil\<^sup>\ \ P\<^bsub>pat\<^esub>"}\\ @{text "\x. y = PVar\<^sup>\ x \ P\<^bsub>pat\<^esub>"}\\ @{text "\x\<^isub>1 x\<^isub>2. y = PTup\<^sup>\ x\<^isub>1 x\<^isub>2 \ P\<^bsub>pat\<^esub>"} \end{array}}\medskip\\ @@ -1985,7 +1982,6 @@ @{text "\x\<^isub>1 x\<^isub>2. P\<^bsub>trm\<^esub> x\<^isub>1 \ P\<^bsub>trm\<^esub> x\<^isub>2 \ P\<^bsub>trm\<^esub> (App\<^sup>\ x\<^isub>1 x\<^isub>2)"}\\ @{text "\x\<^isub>1 x\<^isub>2. P\<^bsub>trm\<^esub> x\<^isub>2 \ P\<^bsub>trm\<^esub> (Lam\<^sup>\ x\<^isub>1 x\<^isub>2)"}\\ @{text "\x\<^isub>1 x\<^isub>2 x\<^isub>3. P\<^bsub>pat\<^esub> x\<^isub>1 \ P\<^bsub>trm\<^esub> x\<^isub>2 \ P\<^bsub>trm\<^esub> x\<^isub>3 \ P\<^bsub>trm\<^esub> (Let_pat\<^sup>\ x\<^isub>1 x\<^isub>2 x\<^isub>3)"}\\ - @{text "P\<^bsub>pat\<^esub> (PNil\<^sup>\)"}\\ @{text "\x. P\<^bsub>pat\<^esub> (PVar\<^sup>\ x)"}\\ @{text "\x\<^isub>1 x\<^isub>2. P\<^bsub>pat\<^esub> x\<^isub>1 \ P\<^bsub>pat\<^esub> x\<^isub>2 \ P\<^bsub>pat\<^esub> (PTup\<^sup>\ x\<^isub>1 x\<^isub>2)"} \end{array}} @@ -2137,7 +2133,6 @@ @{text "\x\<^isub>1 x\<^isub>2 c. atom x\<^isub>1 # c \ (\d. P\<^bsub>trm\<^esub> d x\<^isub>2) \ P\<^bsub>trm\<^esub> c (Lam\<^sup>\ x\<^isub>1 x\<^isub>2)"}\\ @{text "\x\<^isub>1 x\<^isub>2 x\<^isub>3 c. (set (bn\<^sup>\ x\<^isub>1)) #\<^sup>* c \"}\\ \hspace{10mm}@{text "(\d. P\<^bsub>pat\<^esub> d x\<^isub>1) \ (\d. P\<^bsub>trm\<^esub> d x\<^isub>2) \ (\d. P\<^bsub>trm\<^esub> d x\<^isub>3) \ P\<^bsub>trm\<^esub> c (Let_pat\<^sup>\ x\<^isub>1 x\<^isub>2 x\<^isub>3)"}\\ - @{text "\c. P\<^bsub>pat\<^esub> c (PNil\<^sup>\)"}\\ @{text "\x c. P\<^bsub>pat\<^esub> c (PVar\<^sup>\ x)"}\\ @{text "\x\<^isub>1 x\<^isub>2 c. (\d. P\<^bsub>pat\<^esub> d x\<^isub>1) \ (\d. P\<^bsub>pat\<^esub> d x\<^isub>2) \ P\<^bsub>pat\<^esub> c (PTup\<^sup>\ x\<^isub>1 x\<^isub>2)"} \end{array}} @@ -2200,7 +2195,6 @@ \infer{@{text "P\<^bsub>pat\<^esub>"}} {\begin{array}{@ {}l@ {}} - @{text "y = PNil\<^sup>\ \ P\<^bsub>pat\<^esub>"}\\ @{text "\x. y = PVar\<^sup>\ x \ P\<^bsub>pat\<^esub>"}\\ @{text "\x\<^isub>1 x\<^isub>2. y = PTup\<^sup>\ x\<^isub>1 x\<^isub>2 \ P\<^bsub>pat\<^esub>"} \end{array}} diff -r aa5059a00f41 -r 6fd3fc3254ee Nominal/Ex/LetFun.thy --- a/Nominal/Ex/LetFun.thy Wed Sep 21 17:16:11 2011 +0900 +++ b/Nominal/Ex/LetFun.thy Wed Sep 21 10:23:06 2011 +0200 @@ -34,6 +34,7 @@ thm exp_pat.eq_iff thm exp_pat.fv_bn_eqvt thm exp_pat.size_eqvt +thm exp_pat.size thm exp_pat.supports thm exp_pat.fsupp thm exp_pat.supp diff -r aa5059a00f41 -r 6fd3fc3254ee Nominal/Ex/LetPat.thy --- a/Nominal/Ex/LetPat.thy Wed Sep 21 17:16:11 2011 +0900 +++ b/Nominal/Ex/LetPat.thy Wed Sep 21 10:23:06 2011 +0200 @@ -7,28 +7,96 @@ nominal_datatype trm = Var "name" | App "trm" "trm" -| Lam x::"name" t::"trm" binds (set) x in t -| Let p::"pat" "trm" t::"trm" binds (set) "f p" in t +| Lam x::"name" t::"trm" binds x in t +| Let p::"pat" "trm" t::"trm" binds "bn p" in t and pat = - PN -| PS "name" -| PD "name" "name" + PNil +| PVar "name" +| PTup "pat" "pat" binder - f::"pat \ atom set" + bn::"pat \ atom list" where - "f PN = {}" -| "f (PD x y) = {atom x, atom y}" -| "f (PS x) = {atom x}" + "bn PNil = []" +| "bn (PVar x) = [atom x]" +| "bn (PTup p1 p2) = bn p1 @ bn p2" +thm trm_pat.eq_iff thm trm_pat.distinct thm trm_pat.induct +thm trm_pat.strong_induct[no_vars] thm trm_pat.exhaust +thm trm_pat.strong_exhaust[no_vars] thm trm_pat.fv_defs thm trm_pat.bn_defs thm trm_pat.perm_simps thm trm_pat.eq_iff thm trm_pat.fv_bn_eqvt -thm trm_pat.size_eqvt +thm trm_pat.size + +(* Nominal_Abs test *) + +lemma alpha_res_alpha_set: + "(bs, x) \res op = supp p (cs, y) \ + (bs \ supp x, x) \set op = supp p (cs \ supp y, y)" + using alpha_abs_set_abs_res alpha_abs_res_abs_set by blast + + +lemma + fixes x::"'a::fs" + assumes "(supp x - as) \* p" + and "p \ x = y" + and "p \ (as \ supp x) = bs \ supp y" + shows "(as, x) \res (op =) supp p (bs, y)" + using assms + unfolding alpha_res_alpha_set + unfolding alphas + apply simp + apply rule + apply (rule trans) + apply (rule supp_perm_eq[symmetric, of _ p]) + apply(subst supp_finite_atom_set) + apply (metis finite_Diff finite_supp) + apply (simp add: fresh_star_def) + apply blast + apply(perm_simp) + apply(simp) + done + +lemma + fixes x::"'a::fs" + assumes "(supp x - as) \* p" + and "p \ x = y" + and "p \ as = bs" + shows "(as, x) \set (op =) supp p (bs, y)" + using assms + unfolding alphas + apply simp + apply (rule trans) + apply (rule supp_perm_eq[symmetric, of _ p]) + apply(subst supp_finite_atom_set) + apply (metis finite_Diff finite_supp) + apply(simp) + apply(perm_simp) + apply(simp) + done + +lemma + fixes x::"'a::fs" + assumes "(supp x - set as) \* p" + and "p \ x = y" + and "p \ as = bs" + shows "(as, x) \lst (op =) supp p (bs, y)" + using assms + unfolding alphas + apply simp + apply (rule trans) + apply (rule supp_perm_eq[symmetric, of _ p]) + apply(subst supp_finite_atom_set) + apply (metis finite_Diff finite_supp) + apply(simp) + apply(perm_simp) + apply(simp) + done end diff -r aa5059a00f41 -r 6fd3fc3254ee Nominal/Ex/LetRec2.thy --- a/Nominal/Ex/LetRec2.thy Wed Sep 21 17:16:11 2011 +0900 +++ b/Nominal/Ex/LetRec2.thy Wed Sep 21 10:23:06 2011 +0200 @@ -5,26 +5,29 @@ atom_decl name nominal_datatype trm = - Vr "name" -| Ap "trm" "trm" -| Lm x::"name" t::"trm" binds (set) x in t -| Lt a::"lts" t::"trm" binds "bn a" in a t -and lts = - Lnil -| Lcons "name" "trm" "lts" + Var "name" +| App "trm" "trm" +| Lam x::"name" t::"trm" binds x in t +| Let as::"assn" t::"trm" binds "bn as" in t +| Let_rec as::"assn" t::"trm" binds "bn as" in as t +and assn = + ANil +| ACons "name" "trm" "assn" binder bn where - "bn Lnil = []" -| "bn (Lcons x t l) = (atom x) # (bn l)" + "bn (ANil) = []" +| "bn (ACons x t as) = (atom x) # (bn as)" +thm trm_assn.eq_iff +thm trm_assn.supp -thm trm_lts.fv_defs -thm trm_lts.eq_iff -thm trm_lts.bn_defs -thm trm_lts.perm_simps -thm trm_lts.induct -thm trm_lts.distinct +thm trm_assn.fv_defs +thm trm_assn.eq_iff +thm trm_assn.bn_defs +thm trm_assn.perm_simps +thm trm_assn.induct +thm trm_assn.distinct diff -r aa5059a00f41 -r 6fd3fc3254ee Nominal/Ex/TypeVarsTest.thy --- a/Nominal/Ex/TypeVarsTest.thy Wed Sep 21 17:16:11 2011 +0900 +++ b/Nominal/Ex/TypeVarsTest.thy Wed Sep 21 10:23:06 2011 +0200 @@ -44,6 +44,15 @@ | Output "'alpha" "'alpha" "('alpha, 'beta, 'gamma) psi" +nominal_datatype 'a foo = + Node x::"name" f::"'a foo" binds x in f +| Leaf "'a" + +term "Leaf" + + + + end diff -r aa5059a00f41 -r 6fd3fc3254ee Nominal/nominal_dt_alpha.ML --- a/Nominal/nominal_dt_alpha.ML Wed Sep 21 17:16:11 2011 +0900 +++ b/Nominal/nominal_dt_alpha.ML Wed Sep 21 10:23:06 2011 +0200 @@ -564,7 +564,7 @@ [ etac @{thm exE}, etac @{thm exE}, perm_inst_tac ctxt, - resolve_tac @{thms alpha_trans_eqvt}, + resolve_tac @{thms alpha_trans_eqvt}, atac, aatac pred_names ctxt, eqvt_tac ctxt (eqvt_relaxed_config addpres alpha_eqvt) THEN' rtac @{thm refl},