diff -r faeac3ae43e5 -r 25d813c5042d Nominal/Ex/CR.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Nominal/Ex/CR.thy Wed Dec 21 13:06:09 2011 +0900 @@ -0,0 +1,377 @@ +(* CR_Takahashi from Nominal1 ported to Nominal2 *) +theory CR imports Nominal2 begin + +atom_decl name + +nominal_datatype lam = + Var "name" +| App "lam" "lam" +| Lam x::"name" l::"lam" binds x in l ("Lam [_]. _" [100, 100] 100) + +nominal_primrec + subst :: "lam \ name \ lam \ lam" ("_ [_ ::= _]" [90, 90, 90] 90) +where + "(Var x)[y ::= s] = (if x = y then s else (Var x))" +| "(App t1 t2)[y ::= s] = App (t1[y ::= s]) (t2[y ::= s])" +| "atom x \ (y, s) \ (Lam [x]. t)[y ::= s] = Lam [x].(t[y ::= s])" + unfolding eqvt_def subst_graph_def + apply (rule, perm_simp, rule) + apply(rule TrueI) + apply(auto simp add: lam.distinct lam.eq_iff) + apply(rule_tac y="a" and c="(aa, b)" in lam.strong_exhaust) + apply(blast)+ + apply(simp_all add: fresh_star_def fresh_Pair_elim) + apply (erule_tac c="(ya,sa)" in Abs_lst1_fcb2) + apply(simp_all add: Abs_fresh_iff) + apply(simp add: fresh_star_def fresh_Pair) + apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq) + apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq) + done + +termination (eqvt) + by lexicographic_order + +lemma forget: + shows "atom x \ t \ t[x ::= s] = t" + by (nominal_induct t avoiding: x s rule: lam.strong_induct) + (auto simp add: lam.fresh fresh_at_base) + +lemma fresh_fact: + fixes z::"name" + assumes a: "atom z \ s" + and b: "z = y \ atom z \ t" + shows "atom z \ t[y ::= s]" + using a b + by (nominal_induct t avoiding: z y s rule: lam.strong_induct) + (auto simp add: lam.fresh fresh_at_base) + +lemma substitution_lemma: + assumes a: "x \ y" "atom x \ u" + shows "t[x ::= s][y ::= u] = t[y ::= u][x ::= s[y ::= u]]" + using a + by (nominal_induct t avoiding: x y s u rule: lam.strong_induct) + (auto simp add: fresh_fact forget) + +lemma subst_rename: + assumes a: "atom y \ t" + shows "t[x ::= s] = ((y \ x) \t)[y ::= s]" + using a + by (nominal_induct t avoiding: x y s rule: lam.strong_induct) + (auto simp add: lam.fresh fresh_at_base) + +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) + +inductive + beta :: "lam \ lam \ bool" (" _ \b _" [80,80] 80) +where + b1[intro]: "t1 \b t2 \ App t1 s \b App t2 s" +| b2[intro]: "s1 \b s2 \ App t s1 \b App t s2" +| b3[intro]: "t1 \b t2 \ Lam [x]. t1 \b Lam [x]. t2" +| b4[intro]: "atom x \ s \ App (Lam [x]. t) s \b t[x ::= s]" + +equivariance beta + +nominal_inductive beta + avoids b3: x + | b4: x + by (simp_all add: fresh_star_def fresh_Pair lam.fresh fresh_fact) + +section {* Transitive Closure of Beta *} + +inductive + beta_star :: "lam \ lam \ bool" (" _ \b* _" [80,80] 80) +where + bs1[intro, simp]: "M \b* M" +| bs2[intro]: "\M1\b* M2; M2 \b M3\ \ M1 \b* M3" + +lemma bs3[intro, trans]: + assumes "A \b* B" + and "B \b* C" + shows "A \b* C" + using assms(2) assms(1) + by induct auto + +section {* One-Reduction *} + +inductive + One :: "lam \ lam \ bool" (" _ \1 _" [80,80] 80) +where + o1[intro]: "Var x \1 Var x" +| o2[intro]: "\t1 \1 t2; s1 \1 s2\ \ App t1 s1 \1 App t2 s2" +| o3[intro]: "t1 \1 t2 \ Lam [x].t1 \1 Lam [x].t2" +| o4[intro]: "\atom x \ (s1, s2); t1 \1 t2; s1 \1 s2\ \ App (Lam [x].t1) s1 \1 t2[x ::= s2]" + +equivariance One + +nominal_inductive One + avoids o3: "x" + | o4: "x" + by (simp_all add: fresh_star_def fresh_Pair lam.fresh fresh_fact) + +lemma One_refl: + shows "t \1 t" + by (nominal_induct t rule: lam.strong_induct) (auto) + +lemma One_subst: + assumes a: "t1 \1 t2" "s1 \1 s2" + shows "t1[x ::= s1] \1 t2[x ::= s2]" + using a + by (nominal_induct t1 t2 avoiding: s1 s2 x rule: One.strong_induct) + (auto simp add: substitution_lemma fresh_at_base fresh_fact fresh_Pair) + +lemma better_o4_intro: + assumes a: "t1 \1 t2" "s1 \1 s2" + shows "App (Lam [x]. t1) s1 \1 t2[ x ::= s2]" +proof - + obtain y::"name" where fs: "atom y \ (x, t1, s1, t2, s2)" by (rule obtain_fresh) + have "App (Lam [x]. t1) s1 = App (Lam [y]. ((y \ x) \ t1)) s1" using fs + by (auto simp add: lam.eq_iff Abs1_eq_iff' flip_def fresh_Pair fresh_at_base) + also have "\ \1 ((y \ x) \ t2)[y ::= s2]" using fs a by (auto simp add: One.eqvt) + also have "\ = t2[x ::= s2]" using fs by (simp add: subst_rename[symmetric]) + finally show "App (Lam [x].t1) s1 \1 t2[x ::= s2]" by simp +qed + +lemma One_Var: + assumes a: "Var x \1 M" + shows "M = Var x" + using a by (cases rule: One.cases) (simp_all) + +lemma One_Lam: + assumes a: "Lam [x].t \1 s" "atom x \ s" + shows "\t'. s = Lam [x].t' \ t \1 t'" + using a + apply (cases rule: One.cases) + apply (auto simp add: Abs1_eq_iff) + apply (rule_tac x="(atom xa \ atom x) \ t2" in exI) + apply (auto simp add: fresh_permute_left lam.fresh) + by (metis swap_commute One.eqvt) + +lemma One_App: + assumes a: "App t s \1 r" + shows "(\t' s'. r = App t' s' \ t \1 t' \ s \1 s') \ + (\x p p' s'. r = p'[x ::= s'] \ t = Lam [x].p \ p \1 p' \ s \1 s' \ atom x \ (s,s'))" + using a by (cases rule: One.cases) auto + +lemma One_preserves_fresh: + fixes x::"name" + assumes a: "M \1 N" + shows "atom x \ M \ atom x \ N" + using a + by (induct, auto simp add: lam.fresh) + (metis fresh_fact)+ + +(* TODO *) +lemma One_strong_cases[consumes 1]: + "\ a1 \1 a2; \x. \a1 = Var x; a2 = Var x\ \ P; + \t1 t2 s1 s2. \a1 = App t1 s1; a2 = App t2 s2; t1 \1 t2; s1 \1 s2\ \ P; + \t1 t2. (\atom xa \ a1; atom xa \ a2\ \ a1 = Lam [xa].t1 \ a2 = Lam [xa].t2 \ t1 \1 t2) \ P; + \s1 s2 t1 t2. + (\atom xaa \ a1; atom xaa \ a2\ + \ a1 = App (Lam [xaa].t1) s1 \ a2 = t2[xaa::=s2] \ atom xaa \ (s1, s2) \ t1 \1 t2 \ s1 \1 s2) \ + P\ + \ P" + apply (nominal_induct avoiding: a1 a2 rule: One.strong_induct) + apply blast + apply blast + apply (simp add: fresh_Pair_elim Abs1_eq_iff lam.fresh) + apply (case_tac "xa = x") + apply (simp_all)[2] + apply blast + apply (rotate_tac 6) + apply (drule_tac x="(atom x \ atom xa) \ t1" in meta_spec) + apply (rotate_tac -1) + apply (drule_tac x="(atom x \ atom xa) \ t2" in meta_spec) + apply (simp add: One.eqvt fresh_permute_left) + apply (simp add: fresh_Pair_elim Abs1_eq_iff lam.fresh) + apply (case_tac "xaa = x") + apply (simp_all add: fresh_Pair)[2] + apply blast + apply (rotate_tac -2) + apply (drule_tac x="s1" in meta_spec) + apply (rotate_tac -1) + apply (drule_tac x="s2" in meta_spec) + apply (rotate_tac -1) + apply (drule_tac x="(atom x \ atom xaa) \ t1" in meta_spec) + apply (rotate_tac -1) + apply (drule_tac x="(atom x \ atom xaa) \ t2" in meta_spec) + apply (rotate_tac -1) + apply (simp add: One_preserves_fresh fresh_permute_left One.eqvt) + by (metis Nominal2_Base.swap_commute One_preserves_fresh flip_def subst_rename) + +lemma One_Redex: + assumes a: "App (Lam [x].t) s \1 r" "atom x \ (s,r)" + shows "(\t' s'. r = App (Lam [x].t') s' \ t \1 t' \ s \1 s') \ + (\t' s'. r = t'[x ::= s'] \ t \1 t' \ s \1 s')" + using a + by (cases rule: One_strong_cases) + (auto dest!: One_Lam simp add: fresh_Pair lam.fresh Abs1_eq_iff) + +inductive + One_star :: "lam\lam\bool" (" _ \1* _" [80,80] 80) +where + os1[intro]: "t \1* t" +| os2[intro]: "t \1 s \ t \1* s" +| os3[intro, trans]: "t1 \1* t2 \ t2 \1* t3 \ t1 \1* t3" + +section {* Complete Development Reduction *} + +inductive + Dev :: "lam \ lam \ bool" (" _ \d _" [80,80] 80) +where + d1[intro]: "Var x \d Var x" +| d2[intro]: "t \d s \ Lam [x].t \d Lam[x].s" +| d3[intro]: "\\(\y t'. t1 = Lam [y].t'); t1 \d t2; s1 \d s2\ \ App t1 s1 \d App t2 s2" +| d4[intro]: "\atom x \ (s1,s2); t1 \d t2; s1 \d s2\ \ App (Lam [x].t1) s1 \d t2[x::=s2]" + +equivariance Dev +nominal_inductive Dev + avoids d2: "x" + | d4: "x" + by (simp_all add: fresh_star_def fresh_Pair lam.fresh fresh_fact) + +lemma better_d4_intro: + assumes a: "t1 \d t2" "s1 \d s2" + shows "App (Lam [x].t1) s1 \d t2[x::=s2]" +proof - + obtain y::"name" where fs: "atom y\(x,t1,s1,t2,s2)" by (rule obtain_fresh) + have "App (Lam [x].t1) s1 = App (Lam [y].((y \ x)\t1)) s1" using fs + by (auto simp add: Abs1_eq_iff' flip_def fresh_Pair fresh_at_base) + also have "\ \d ((y \ x) \ t2)[y ::= s2]" using fs a by (auto simp add: Dev.eqvt) + also have "\ = t2[x::=s2]" using fs by (simp add: subst_rename[symmetric]) + finally show "App (Lam [x].t1) s1 \d t2[x::=s2]" by simp +qed + +lemma Dev_preserves_fresh: + fixes x::"name" + assumes a: "M\d N" + shows "atom x\M \ atom x\N" + using a + by (induct, auto simp add: lam.fresh) + (metis fresh_fact)+ + +lemma Dev_Lam: + assumes a: "Lam [x].M \d N" + shows "\N'. N = Lam [x].N' \ M \d N'" +proof - + from a have "atom x \ Lam [x].M" by (simp add: lam.fresh) + with a have "atom x \ N" by (simp add: Dev_preserves_fresh) + with a show "\N'. N = Lam [x].N' \ M \d N'" + apply (cases rule: Dev.cases) + apply (auto simp add: Abs1_eq_iff lam.fresh) + apply (rule_tac x="(atom xa \ atom x) \ s" in exI) + apply (auto simp add: fresh_permute_left lam.fresh) + by (metis swap_commute Dev.eqvt) +qed + +lemma Development_existence: + shows "\M'. M \d M'" +by (nominal_induct M rule: lam.strong_induct) + (auto dest!: Dev_Lam intro: better_d4_intro) + +lemma Triangle: + assumes a: "t \d t1" "t \1 t2" + shows "t2 \1 t1" +using a +proof(nominal_induct avoiding: t2 rule: Dev.strong_induct) + case (d4 x s1 s2 t1 t1' t2) + have fc: "atom x\t2" "atom x\s1" by fact+ + have "App (Lam [x].t1) s1 \1 t2" by fact + then obtain t' s' where reds: + "(t2 = App (Lam [x].t') s' \ t1 \1 t' \ s1 \1 s') \ + (t2 = t'[x::=s'] \ t1 \1 t' \ s1 \1 s')" + using fc by (auto dest!: One_Redex) + have ih1: "t1 \1 t' \ t' \1 t1'" by fact + have ih2: "s1 \1 s' \ s' \1 s2" by fact + { assume "t1 \1 t'" "s1 \1 s'" + then have "App (Lam [x].t') s' \1 t1'[x::=s2]" + using ih1 ih2 by (auto intro: better_o4_intro) + } + moreover + { assume "t1 \1 t'" "s1 \1 s'" + then have "t'[x::=s'] \1 t1'[x::=s2]" + using ih1 ih2 by (auto intro: One_subst) + } + ultimately show "t2 \1 t1'[x::=s2]" using reds by auto +qed (auto dest!: One_Lam One_Var One_App) + +lemma Diamond_for_One: + assumes a: "t \1 t1" "t \1 t2" + shows "\t3. t2 \1 t3 \ t1 \1 t3" +proof - + obtain tc where "t \d tc" using Development_existence by blast + with a have "t2 \1 tc" and "t1 \1 tc" by (simp_all add: Triangle) + then show "\t3. t2 \1 t3 \ t1 \1 t3" by blast +qed + +lemma Rectangle_for_One: + assumes a: "t \1* t1" "t \1 t2" + shows "\t3. t1 \1 t3 \ t2 \1* t3" +using a Diamond_for_One by (induct arbitrary: t2) (blast)+ + +lemma CR_for_One_star: + assumes a: "t \1* t1" "t \1* t2" + shows "\t3. t2 \1* t3 \ t1 \1* t3" +using a Rectangle_for_One by (induct arbitrary: t2) (blast)+ + +section {* Establishing the Equivalence of Beta-star and One-star *} + +lemma Beta_Lam_cong: + assumes a: "t1 \b* t2" + shows "Lam [x].t1 \b* Lam [x].t2" +using a by (induct) (blast)+ + +lemma Beta_App_cong_aux: + assumes a: "t1 \b* t2" + shows "App t1 s\b* App t2 s" + and "App s t1 \b* App s t2" +using a by (induct) (blast)+ + +lemma Beta_App_cong: + assumes a: "t1 \b* t2" "s1 \b* s2" + shows "App t1 s1 \b* App t2 s2" +using a by (blast intro: Beta_App_cong_aux) + +lemmas Beta_congs = Beta_Lam_cong Beta_App_cong + +lemma One_implies_Beta_star: + assumes a: "t \1 s" + shows "t \b* s" +using a by (induct, auto intro!: Beta_congs) + (metis (hide_lams, no_types) Beta_App_cong_aux(1) Beta_App_cong_aux(2) Beta_Lam_cong b4 bs2 bs3 fresh_PairD(2)) + +lemma One_congs: + assumes a: "t1 \1* t2" + shows "Lam [x].t1 \1* Lam [x].t2" + and "App t1 s \1* App t2 s" + and "App s t1 \1* App s t2" +using a by (induct) (auto intro: One_refl) + +lemma Beta_implies_One_star: + assumes a: "t1 \b t2" + shows "t1 \1* t2" +using a by (induct) (auto intro: One_refl One_congs better_o4_intro) + +lemma Beta_star_equals_One_star: + shows "t1 \1* t2 = t1 \b* t2" +proof + assume "t1 \1* t2" + then show "t1 \b* t2" by (induct) (auto intro: One_implies_Beta_star) +next + assume "t1 \b* t2" + then show "t1 \1* t2" by (induct) (auto intro: Beta_implies_One_star) +qed + +section {* The Church-Rosser Theorem *} + +theorem CR_for_Beta_star: + assumes a: "t \b* t1" "t\b* t2" + shows "\t3. t1 \b* t3 \ t2 \b* t3" +proof - + from a have "t \1* t1" and "t\1* t2" by (simp_all add: Beta_star_equals_One_star) + then have "\t3. t1 \1* t3 \ t2 \1* t3" by (simp add: CR_for_One_star) + then show "\t3. t1 \b* t3 \ t2 \b* t3" by (simp add: Beta_star_equals_One_star) +qed + +end