# HG changeset patch # User Christian Urban # Date 1323966011 0 # Node ID ade2f8fcf8e84c2be044a9df624f986b5934e98a # Parent 32abaea424bda7335f0557fc852b6a6a17298fd0 a bit more on alpha-beta-equated terms diff -r 32abaea424bd -r ade2f8fcf8e8 Nominal/Ex/Beta.thy --- a/Nominal/Ex/Beta.thy Wed Dec 14 01:32:42 2011 +0000 +++ b/Nominal/Ex/Beta.thy Thu Dec 15 16:20:11 2011 +0000 @@ -6,6 +6,7 @@ atom_decl name + nominal_datatype lam = Var "name" | App "lam" "lam" @@ -58,6 +59,21 @@ (auto simp add: fresh_fact forget) inductive + beta :: "lam \ lam \ bool" ("_ \ _") +where + beta_beta: "atom x \ s \ App (Lam [x].t) s \ t[x ::= s]" +| beta_Lam: "t \ s \ Lam [x].t \ Lam [x].s" +| beta_App1: "t \ s \ App t u \ App s u" +| beta_App2: "t \ s \ App u t \ App u s" + +equivariance beta + +nominal_inductive beta + avoids beta_beta: "x" + | beta_Lam: "x" +by (simp_all add: fresh_star_def fresh_Pair lam.fresh fresh_fact) + +inductive equ :: "lam \ lam \ bool" ("_ \ _") where equ_beta: "atom x \ s \ App (Lam [x].t) s \ t[x ::= s]" @@ -75,6 +91,64 @@ | equ_Lam: "x" by (simp_all add: fresh_star_def fresh_Pair lam.fresh fresh_fact) +inductive + equ2 :: "lam \ lam \ bool" ("_ \2 _") +where + equ2_beta1: "\atom x \ (t2, s2); s1 \2 t1; t2 \2 s2\ \ App (Lam [x].t1) t2 \2 s1[x ::= s2]" +| equ2_beta2: "\atom x \ (t2, s2); s1 \2 t1; t2 \2 s2\ \ s1[x ::= s2] \2 App (Lam [x].t1) t2" +| equ2_Var: "Var x \2 Var x" +| equ2_Lam: "t \2 s \ Lam [x].t \2 Lam [x].s" +| equ2_App: "\t1 \2 s1; t2 \2 s2\ \ App t1 t2 \2 App s1 s2" + +equivariance equ2 + +nominal_inductive equ2 + avoids equ2_beta1: "x" + | equ2_beta2: "x" + | equ2_Lam: "x" +by (simp_all add: fresh_star_def fresh_Pair lam.fresh fresh_fact) + +lemma equ2_refl: + fixes t::"lam" + shows "t \2 t" +apply(induct t rule: lam.induct) +apply(auto intro: equ2.intros) +done + +lemma equ2_symm: + fixes t s::"lam" + assumes "t \2 s" + shows "s \2 t" +using assms +apply(induct) +apply(auto intro: equ2.intros) +done + +lemma equ2_trans: + fixes t1 t2 t3::"lam" + assumes "t1 \2 t2" "t2 \2 t3" + shows "t1 \2 t3" +using assms +apply(nominal_induct t1 t2 avoiding: t3 rule: equ2.strong_induct) +defer +defer +apply(erule equ2.cases) +apply(auto) +apply(rule equ2_beta2) +apply(simp) +apply(simp) +apply(simp) +apply(rule equ2_refl) +defer +apply(rotate_tac 4) +apply(erule equ2.cases) +apply(auto) +oops + + + + + lemma [quot_respect]: shows "(op = ===> equ) Var Var" and "(equ ===> equ ===> equ) App App" @@ -123,7 +197,7 @@ lemma [quot_respect]: shows "(op = ===> equ ===> equ) permute permute" unfolding fun_rel_def -by (auto intro: eqvt) +by (auto intro: eqvts) quotient_type qlam = lam / equ apply(rule equivpI) @@ -161,8 +235,52 @@ end +lemma qlam_perm[simp]: + shows "p \ (QVar x) = QVar (p \ x)" + and "p \ (QApp t s) = QApp (p \ t) (p \ s)" + and "p \ (QLam [x].t) = QLam [p \ x].(p \ t)" +apply(descending) +apply(simp add: equ_refl) +apply(descending) +apply(simp add: equ_refl) +apply(descending) +apply(simp add: equ_refl) +done + +lemma qlam_supports: + shows "{atom x} supports (QVar x)" + and "supp (t, s) supports (QApp t s)" + and "supp (x, t) supports (QLam [x].t)" +unfolding supports_def fresh_def[symmetric] +by (auto simp add: swap_fresh_fresh) + +lemma qlam_fs: + fixes t::"qlam" + shows "finite (supp t)" +apply(induct t rule: qlam_induct) +apply(rule supports_finite) +apply(rule qlam_supports) +apply(simp) +apply(rule supports_finite) +apply(rule qlam_supports) +apply(simp add: supp_Pair) +apply(rule supports_finite) +apply(rule qlam_supports) +apply(simp add: supp_Pair finite_supp) +done + +instantiation qlam :: fs +begin + +instance +apply(default) +apply(rule qlam_fs) +done + +end + quotient_definition subst_qlam ("_[_ ::q= _]") - where "subst_qlam::qlam \ name \ qlam \ qlam" is subst + where "subst_qlam::qlam \ name \ qlam \ qlam" is subst lemma "(QVar x)[y ::q= s] = (if x = y then s else (QVar x))" @@ -178,12 +296,136 @@ apply(rule equ_refl) done +definition + "Supp t = \{supp s | s. s \ t}" + +lemma [quot_respect]: + shows "(equ ===> op=) Supp Supp" +unfolding fun_rel_def +unfolding Supp_def +apply(rule allI)+ +apply(rule impI) +apply(rule_tac f="Inter" in arg_cong) +apply(auto) +apply (metis equ_trans) +by (metis equivp_def qlam_equivp) + +quotient_definition "supp_qlam::qlam \ atom set" + is "Supp::lam \ atom set" + +lemma Supp_supp: + "Supp t \ supp t" +unfolding Supp_def +apply(auto) +apply(drule_tac x="supp t" in spec) +apply(auto simp add: equ_refl) +done + +lemma Supp_Lam: + "atom a \ Supp (Lam [a].t)" +proof - + have "atom a \ supp (Lam [a].t)" by (simp add: lam.supp) + then show ?thesis using Supp_supp + by blast +qed + +lemmas s = Supp_Lam[quot_lifted] + +definition + ssupp :: "atom set \ 'a::pt \ bool" ("_ ssupp _") +where + "A ssupp x \ \p. (\a \ A. (p \ a) = a) \ (p \ x) = x" + +lemma ssupp_supports: + "A ssupp x \ A supports x" +unfolding ssupp_def supports_def +apply(rule allI)+ +apply(drule_tac x="(a \ b)" in spec) +apply(auto) +by (metis swap_atom_simps(3)) + +lemma ssupp_supp: + assumes a: "finite A" + and b: "A ssupp x" + shows "supp x = A" +proof - + { assume 0: "A - supp x \ {}" + then obtain a where 1: "a \ A - supp x" by auto + obtain a' where *: "a' \ UNIV - A" and **: "sort_of a' = sort_of a" + by (rule obtain_atom[OF a]) (auto) + have "(a \ a') \ a = a'" using 1 ** * by (auto) + then have w0: "(a \ a') \ x \ x" + using b unfolding ssupp_def + apply - + apply(drule_tac x="(a \ a')" in spec) + apply(auto) + using 1 * + apply(auto) + done + have w1: "a \ x" unfolding fresh_def using 1 by auto + have w2: "a' \ x" using * + apply(rule_tac supports_fresh) + apply(rule ssupp_supports) + apply(rule b) + apply(rule a) + apply(auto) + done + have "False" using w0 w1 w2 by (simp add: swap_fresh_fresh) + then have "supp x = A" by simp } + moreover + { assume "A - supp x = {}" + moreover + have "supp x \ A" + apply(rule supp_is_subset) + apply(rule ssupp_supports) + apply(rule b) + apply(rule a) + done + ultimately have "supp x = A" + by blast + } + ultimately show "supp x = A" by blast +qed + + lemma - "(QLam [x].t)[y ::q= s] = QLam [x].(t[y ::q= s])" + "(supp x) ssupp x" +unfolding ssupp_def +apply(auto) +apply(rule supp_perm_eq) +apply(simp add: fresh_star_def) +apply(auto) +apply(simp add: fresh_perm) +apply(simp add: fresh_perm[symmetric]) +(*Tzevelekos 2008, section 2.1.2 property 2.5*) +oops + + +function + qsubst :: "qlam \ name \ qlam \ qlam" ("_ [_ ::qq= _]" [90, 90, 90] 90) +where + "(QVar x)[y ::qq= s] = (if x = y then s else (QVar x))" +| "(QApp t1 t2)[y ::qq= s] = QApp (t1[y ::qq= s]) (t2[y ::qq= s])" +| "\atom x \ y; atom x \ s\ \ (QLam [x]. t)[y ::qq= s] = QLam [x].(t[y ::qq= s])" +apply(simp_all add:) +oops + + +lemma + assumes a: "Lam [x].t \ s" + shows "\x' t'. s = Lam [x']. t' \ t \ t'" +using a +apply(induct s rule:lam.induct) +apply(erule equ.cases) +apply(auto) +apply(erule equ.cases) +apply(auto) + + + +lemma + "atom x \ y \ atom x \ supp_qlam s \ (QLam [x].t)[y ::q= s] = QLam [x].(t[y ::q= s])" apply(descending) -apply(subst subst.simps) -prefer 3 -apply(rule equ_refl) oops @@ -222,23 +464,6 @@ section {* Supp *} -definition - "Supp t = \{supp s | s. s \ t}" - -lemma [quot_respect]: - shows "(equ ===> op=) Supp Supp" -unfolding fun_rel_def -unfolding Supp_def -apply(rule allI)+ -apply(rule impI) -apply(rule_tac f="Inter" in arg_cong) -apply(auto) -apply (metis equ_trans) -by (metis equivp_def qlam_equivp) - -quotient_definition "supp_qlam::qlam \ atom set" - is "Supp::lam \ atom set" - lemma s: assumes "s \ t" shows "a \ (abs_qlam s) \ a \ (abs_qlam t)" @@ -277,7 +502,6 @@ apply(rule_tac x="QVar x" in Abs_qlam_cases) apply(auto) thm QVar_def -apply(descending) oops @@ -326,11 +550,16 @@ "size (QLam [x].t) = Suc (size t)" apply(descending) apply(simp add: Size_def) -apply(rule_tac n="Suc (size t)" and m="size t" in Least_Suc2) +thm Least_Suc +apply(rule trans) +apply(rule_tac n="Suc (size t)" in Least_Suc) apply(simp add: Collect_def) apply(rule_tac x="Lam [x].t" in exI) apply(auto intro: equ_refl)[1] apply(simp add: Collect_def) +apply(auto) +defer +apply(simp add: Collect_def) apply(rule_tac x="t" in exI) apply(auto intro: equ_refl)[1] apply(simp add: Collect_def)