diff -r a303ef51cd97 -r 9253984db291 Nominal/Ex/BetaCR.thy --- a/Nominal/Ex/BetaCR.thy Tue Dec 20 11:40:04 2011 +0900 +++ b/Nominal/Ex/BetaCR.thy Tue Dec 20 16:12:32 2011 +0900 @@ -153,6 +153,13 @@ lemma Lam_beta: "A \b* B \ Lam [x]. A \b* Lam [x]. B" by (erule beta_star.induct) auto +lemma Lam_rsp: "A \ B \ Lam [x]. A \ Lam [x]. B" + unfolding equ_def + apply auto + apply (rule_tac x="Lam [x]. r" in exI) + apply (auto simp add: Lam_beta) + done + lemma [quot_respect]: shows "(op = ===> equ) Var Var" and "(equ ===> equ ===> equ) App App" @@ -221,13 +228,17 @@ lemma beta_subst: "A \b* B \ C \b* D \ A [x ::= C] \b* B [x ::= D]" using beta_subst1 beta_subst2 beta_star_trans by metis +lemma subst_rsp_pre: + "x \ y \ xb \ ya \ x [xa ::= xb] \ y [xa ::= ya]" + unfolding equ_def + apply auto + apply (rule_tac x="r [xa ::= ra]" in exI) + apply (simp add: beta_subst) + done + lemma [quot_respect]: shows "(equ ===> op = ===> equ ===> equ) subst subst" -unfolding fun_rel_def equ_def -apply auto -apply (rule_tac x="r [xa ::= ra]" in exI) -apply (simp add: beta_subst) -done +unfolding fun_rel_def by (auto simp add: subst_rsp_pre) lemma [quot_respect]: shows "(op = ===> equ ===> equ) permute permute" @@ -307,20 +318,21 @@ quotient_definition subst_qlam ("_[_ ::q= _]") where "subst_qlam::qlam \ name \ qlam \ qlam" is subst -lemmas qsubst = subst.simps(1-2)[quot_lifted] - definition "Supp t = \{supp s | s. s \ t}" -lemma [quot_respect]: - shows "(equ ===> op=) Supp Supp" - unfolding fun_rel_def Supp_def - apply(intro allI impI) +lemma Supp_rsp: + "a \ b \ Supp a = Supp b" + unfolding Supp_def apply(rule_tac f="Inter" in arg_cong) apply(auto) apply (metis equ_trans) by (metis equivp_def qlam_equivp) +lemma [quot_respect]: + shows "(equ ===> op=) Supp Supp" + unfolding fun_rel_def by (auto simp add: Supp_rsp) + quotient_definition "supp_qlam::qlam \ atom set" is "Supp::lam \ atom set" @@ -332,6 +344,99 @@ apply(auto simp add: equ_refl) done +lemma supp_subst: + shows "supp (t[x ::= s]) \ (supp t - {atom x}) \ supp s" + by (induct t x s rule: subst.induct) (auto simp add: lam.supp supp_at_base) + +lemma supp_beta: "x \b r \ supp r \ supp x" + apply (induct rule: beta.induct) + apply (simp_all add: lam.supp) + apply blast+ + using supp_subst by blast + +lemma supp_beta_star: "x \b* r \ supp r \ supp x" + apply (erule beta_star.induct) + apply auto + using supp_beta by blast + +lemma supp_equ: "x \ y \ \z. z \ x \ supp z \ supp x \ supp y" + unfolding equ_def + apply (simp (no_asm) only: equ_def[symmetric]) + apply (elim exE) + apply (rule_tac x=r in exI) + apply rule + apply (metis bs1 equ_def) + using supp_beta_star by blast + +lemma supp_psubset: "Supp x \ supp x \ \t. t \ x \ supp t \ supp x" +proof - + assume "Supp x \ supp x" + then obtain a where "a \ supp x" "a \ Supp x" by blast + then obtain y where nin: "y \ x" "a \ supp y" unfolding Supp_def by blast + then obtain t where eq: "t \ x" "supp t \ supp x \ supp y" + using supp_equ equ_sym by blast + then have "supp t \ supp x" using nin + by (metis `(a\atom) \ supp (x\lam)` le_infE psubset_eq set_rev_mp) + then show "\t. t \ x \ supp t \ supp x" using eq by blast +qed + +lemma Supp_exists: "\t. supp t = Supp t \ t \ x" +proof - + have "\fs x. supp x = fs \ \t. supp t = Supp t \ t \ x" + proof - + fix fs + show "\x. supp x = fs \ \t\lam. supp t = Supp t \ t \ x" + proof (cases "finite fs") + case True + assume fs: "finite fs" + then show "\x. supp x = fs \ \t\lam. supp t = Supp t \ t \ x" + proof (induct fs rule: finite_psubset_induct, clarify) + fix A :: "atom set" fix x :: lam + assume IH: "\B xa. \B \ supp x; supp xa = B\ \ \t\lam. supp t = Supp t \ t \ xa" + show "\t\lam. supp t = Supp t \ t \ x" + proof (cases "supp x = Supp x") + assume "supp x = Supp x" + then show "\t\lam. supp t = Supp t \ t \ x" + by (rule_tac x="x" in exI) (simp add: equ_refl) + next + assume "supp x \ Supp x" + then have "Supp x \ supp x" using Supp_supp by blast + then obtain y where a1: "supp y \ supp x" "y \ x" + using supp_psubset by blast + then obtain t where "supp t = Supp t \ t \ y" + using IH[of "supp y" y] by blast + then show "\t\lam. supp t = Supp t \ t \ x" using a1 equ_trans + by blast + qed + qed + next + case False + fix x :: lam + assume "supp x = fs" "infinite fs" + then show "\t\lam. supp t = Supp t \ t \ x" + by (auto simp add: finite_supp) + qed simp + qed + then show "\t\lam. supp t = Supp t \ t \ x" by blast +qed + +lemma subst3: "x \ y \ atom x \ Supp s \ Lam [x]. t [y ::= s] \ Lam [x]. (t [y ::= s])" +proof - + assume as: "x \ y \ atom x \ Supp s" + obtain s' where s: "supp s' = Supp s'" "s' \ s" using Supp_exists[of s] by blast + then have lhs: "Lam [x]. t [y ::= s] \ Lam [x]. t [y ::= s']" using subst_rsp_pre equ_refl equ_sym by blast + have supp: "Supp s' = Supp s" using Supp_rsp s(2) by blast + have lhs_rhs: "Lam [x]. t [y ::= s'] = Lam [x]. (t [y ::= s'])" + by (simp add: fresh_def supp_Pair s supp as supp_at_base) + have "t [y ::= s'] \ t [y ::= s]" + using subst_rsp_pre[OF equ_refl s(2)] . + with Lam_rsp have rhs: "Lam [x]. (t [y ::= s']) \ Lam [x]. (t [y ::= s])" . + show ?thesis + using lhs[unfolded lhs_rhs] rhs equ_trans by blast +qed + +thm subst3[quot_lifted] + lemma Supp_Lam: "atom a \ Supp (Lam [a].t)" proof -