diff -r a9f3600c9ae6 -r 589b1a0c75e6 Nominal/Ex/SFT/Consts.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Nominal/Ex/SFT/Consts.thy Fri Jun 24 10:30:06 2011 +0900 @@ -0,0 +1,316 @@ +header {* Constant definitions *} +theory Consts imports Utils Lambda begin + +fun Umn :: "nat \ nat \ lam" +where + [simp del]: "Umn 0 n = \(cn 0). V (cn n)" +| [simp del]: "Umn (Suc m) n = \(cn (Suc m)). Umn m n" + +lemma [simp]: "2 = Suc 1" + by auto + +lemma Lam_U: + "x \ y \ y \ z \ x \ z \ Umn 2 0 = \x. \y. \z. V z" + "x \ y \ y \ z \ x \ z \ Umn 2 1 = \x. \y. \z. V y" + "x \ y \ y \ z \ x \ z \ Umn 2 2 = \x. \y. \z. V x" + apply (simp_all add: Umn.simps Abs1_eq_iff lam.fresh fresh_at_base flip_def[symmetric] Umn.simps) + apply (smt Zero_not_Suc cnd flip_at_base_simps flip_at_simps)+ + done + +lemma a: "n \ m \ atom (cn n) \ supp (Umn m n)" + apply (induct m) + apply (auto simp add: lam.supp supp_at_base Umn.simps) + by smt + +lemma b: "supp (Umn m n) \ {atom (cn n)}" + by (induct m) (auto simp add: lam.supp supp_at_base Umn.simps) + +lemma supp_U[simp]: "n \ m \ supp (Umn m n) = {}" + using a b + by blast + +lemma U_eqvt: + "n \ m \ p \ (Umn m n) = Umn m n" + by (rule_tac [!] perm_supp_eq) (simp_all add: fresh_star_def fresh_def) + +definition Var where "Var \ \cx. \cy. (V cy \ (Umn 2 2) \ V cx \ V cy)" +definition "App \ \cx. \cy. \cz. (V cz \ Umn 2 1 \ V cx \ V cy \ V cz)" +definition "Abs \ \cx. \cy. (V cy \ Umn 2 0 \ V cx \ V cy)" + +lemma Var_App_Abs: + "x \ e \ Var = \x. \e. (V e \ Umn 2 2 \ V x \ V e)" + "e \ x \ e \ y \ x \ y \ App = \x. \y. \e. (V e \ Umn 2 1 \ V x \ V y \ V e)" + "x \ e \ Abs = \x. \e. (V e \ Umn 2 0 \ V x \ V e)" + unfolding Var_def App_def Abs_def + by (simp_all add: Abs1_eq_iff lam.fresh flip_def[symmetric] U_eqvt fresh_def lam.supp supp_at_base) + (smt cx_cy_cz permute_flip_at Zero_not_Suc cnd flip_at_base_simps flip_at_simps)+ + +lemma Var_app: + "Var \ x \ e \ e \ Umn 2 2 \ x \ e" + by (rule lam2_fast_app) (simp_all add: Var_App_Abs) + +lemma App_app: + "App \ x \ y \ e \ e \ Umn 2 1 \ x \ y \ e" + by (rule lam3_fast_app[OF Var_App_Abs(2)]) (simp_all) + +lemma Abs_app: + "Abs \ x \ e \ e \ Umn 2 0 \ x \ e" + by (rule lam2_fast_app) (simp_all add: Var_App_Abs) + +lemma supp_Var_App_Abs[simp]: + "supp Var = {}" "supp App = {}" "supp Abs = {}" + by (simp_all add: Var_def App_def Abs_def lam.supp supp_at_base) blast+ + +lemma Var_App_Abs_eqvt[eqvt]: + "p \ Var = Var" "p \ App = App" "p \ Abs = Abs" + by (rule_tac [!] perm_supp_eq) (simp_all add: fresh_star_def fresh_def) + +nominal_primrec + Numeral :: "lam \ lam" ("\_\" 1000) +where + "\V x\ = Var \ (V x)" +| Ap: "\M \ N\ = App \ \M\ \ \N\" +| "\\x. M\ = Abs \ (\x. \M\)" +proof auto + fix x :: lam and P + assume "\xa. x = V xa \ P" "\M N. x = M \ N \ P" "\xa M. x = \ xa. M \ P" + then show "P" + by (rule_tac y="x" and c="0 :: perm" in lam.strong_exhaust) + (auto simp add: Abs1_eq_iff fresh_star_def)[3] +next + fix x :: var and M and xa :: var and Ma + assume "[[atom x]]lst. M = [[atom xa]]lst. Ma" + "eqvt_at Numeral_sumC M" + then show "[[atom x]]lst. Numeral_sumC M = [[atom xa]]lst. Numeral_sumC Ma" + apply - + apply (erule Abs_lst1_fcb) + apply (simp_all add: Abs_fresh_iff) + apply (erule fresh_eqvt_at) + apply (simp_all add: finite_supp Abs1_eq_iff eqvt_at_def) + done +next + show "eqvt Numeral_graph" unfolding eqvt_def Numeral_graph_def + by (rule, perm_simp, rule) +qed + +termination + by (relation "measure (\(t). size t)") + (simp_all add: lam.size) + +lemma numeral_eqvt[eqvt]: "p \ \x\ = \p \ x\" + by (induct x rule: lam.induct) + (simp_all add: Var_App_Abs_eqvt) + +lemma supp_numeral[simp]: + "supp \x\ = supp x" + by (induct x rule: lam.induct) + (simp_all add: lam.supp) + +lemma fresh_numeral[simp]: + "x \ \y\ = x \ y" + unfolding fresh_def by simp + +fun app_lst :: "var \ lam list \ lam" where + "app_lst n [] = V n" +| "app_lst n (h # t) = Ap (app_lst n t) h" + +lemma app_lst_eqvt[eqvt]: "p \ (app_lst t ts) = app_lst (p \ t) (p \ ts)" + by (induct ts arbitrary: t p) (simp_all add: eqvts) + +lemma supp_app_lst: "supp (app_lst x l) = {atom x} \ supp l" + apply (induct l) + apply (simp_all add: supp_Nil lam.supp supp_at_base supp_Cons) + by blast + +lemma app_lst_eq_iff: "app_lst n M = app_lst n N \ M = N" + by (induct M N rule: list_induct2') simp_all + +lemma app_lst_rev_eq_iff: "app_lst n (rev M) = app_lst n (rev N) \ M = N" + by (drule app_lst_eq_iff) simp + +nominal_primrec + Ltgt :: "lam list \ lam" ("\_\" 1000) +where + [simp del]: "atom x \ l \ \l\ = \x. (app_lst x (rev l))" + unfolding eqvt_def Ltgt_graph_def + apply (rule, perm_simp, rule, rule) + apply (rule_tac x="x" and ?'a="var" in obtain_fresh) + apply (simp_all add: Abs1_eq_iff lam.fresh swap_fresh_fresh fresh_at_base) + apply (simp add: eqvts swap_fresh_fresh) + apply (case_tac "x = xa") + apply simp_all + apply (subgoal_tac "eqvt app_lst") + apply (erule fresh_fun_eqvt_app2) + apply (simp_all add: fresh_at_base lam.fresh eqvt_def eqvts_raw fresh_rev) + done + +termination + by (relation "measure (\t. size t)") + (simp_all add: lam.size) + +lemma ltgt_eqvt[eqvt]: + "p \ \t\ = \p \ t\" +proof - + obtain x :: var where "atom x \ (t, p \ t)" using obtain_fresh by auto + then have *: "atom x \ t" "atom x \ (p \ t)" using fresh_Pair by simp_all + then show ?thesis using *[unfolded fresh_def] + apply (simp add: Abs1_eq_iff lam.fresh app_lst_eqvt Ltgt.simps) + apply (case_tac "p \ x = x") + apply (simp_all add: eqvts) + apply rule + apply (subst swap_fresh_fresh) + apply (simp_all add: fresh_at_base_permute_iff fresh_def[symmetric] fresh_at_base) + apply (subgoal_tac "eqvt app_lst") + apply (erule fresh_fun_eqvt_app2) + apply (simp_all add: fresh_at_base lam.fresh eqvt_def eqvts_raw fresh_rev) + done +qed + +lemma ltgt_eq_iff[simp]: + "\M\ = \N\ \ M = N" +proof auto + obtain x :: var where "atom x \ (M, N)" using obtain_fresh by auto + then have *: "atom x \ M" "atom x \ N" using fresh_Pair by simp_all + then show "(\M\ = \N\) \ (M = N)" by (simp add: Abs1_eq_iff app_lst_rev_eq_iff Ltgt.simps) +qed + +lemma Ltgt1_app: "\[M]\ \ N \ N \ M" +proof - + obtain x :: var where "atom x \ (M, N)" using obtain_fresh by auto + then have "atom x \ M" "atom x \ N" using fresh_Pair by simp_all + then show ?thesis + apply (subst Ltgt.simps) + apply (simp add: fresh_Cons fresh_Nil) + apply (rule b3, rule bI, simp add: b1) + done +qed + +lemma Ltgt3_app: "\[M,N,P]\ \ R \ R \ M \ N \ P" +proof - + obtain x :: var where "atom x \ (M, N, P, R)" using obtain_fresh by auto + then have *: "atom x \ (M,N,P)" "atom x \ R" using fresh_Pair by simp_all + then have s: "V x \ M \ N \ P [x ::= R] \ R \ M \ N \ P" using b1 by simp + show ?thesis using * + apply (subst Ltgt.simps) + apply (simp add: fresh_Cons fresh_Nil fresh_Pair_elim) + apply auto[1] + apply (rule b3, rule bI, simp add: b1) + done +qed + +lemma supp_ltgt[simp]: + "supp \t\ = supp t" +proof - + obtain x :: var where *:"atom x \ t" using obtain_fresh by auto + show ?thesis using * + by (simp_all add: Ltgt.simps lam.supp supp_at_base supp_Nil supp_app_lst supp_rev fresh_def) +qed + +lemma fresh_ltgt[simp]: + "x \ \[y]\ = x \ y" + "x \ \[t,r,s]\ = x \ (t,r,s)" + by (simp_all add: fresh_def supp_Cons supp_Nil supp_Pair) + +lemma Ltgt1_subst[simp]: + "\[M]\ [y ::= A] = \[M [y ::= A]]\" +proof - + obtain x :: var where a: "atom x \ (M, A, y, M [y ::= A])" using obtain_fresh by blast + have "x \ y" using a[simplified fresh_Pair fresh_at_base] by simp + then show ?thesis + apply (subst Ltgt.simps) + using a apply (simp add: fresh_Nil fresh_Cons fresh_Pair_elim) + apply (subst Ltgt.simps) + using a apply (simp add: fresh_Pair_elim fresh_Nil fresh_Cons) + apply (simp add: a) + done +qed + +lemma U_app: + "\[A,B,C]\ \ Umn 2 2 \ A" "\[A,B,C]\ \ Umn 2 1 \ B" "\[A,B,C]\ \ Umn 2 0 \ C" + by (rule b3, rule Ltgt3_app, rule lam3_fast_app, rule Lam_U, simp_all) + (rule b3, rule Ltgt3_app, rule lam3_fast_app, rule Lam_U[simplified], simp_all)+ + +definition "F1 \ \cx. (App \ \Var\ \ (Var \ V cx))" +definition "F2 \ \cx. \cy. \cz. ((App \ (App \ \App\ \ (V cz \ V cx))) \ (V cz \ V cy))" +definition "F3 \ \cx. \cy. (App \ \Abs\ \ (Abs \ (\cz. (V cy \ (V cx \ V cz)))))" + + +lemma Lam_F: + "F1 = \x. (App \ \Var\ \ (Var \ V x))" + "a \ b \ a \ c \ c \ b \ F2 = \a. \b. \c. ((App \ (App \ \App\ \ (V c \ V a))) \ (V c \ V b))" + "a \ b \ a \ x \ x \ b \ F3 = \a. \b. (App \ \Abs\ \ (Abs \ (\x. (V b \ (V a \ V x)))))" + apply (simp_all add: F1_def F2_def F3_def Abs1_eq_iff lam.fresh supp_at_base Var_App_Abs_eqvt numeral_eqvt flip_def[symmetric] fresh_at_base) + apply (smt cx_cy_cz permute_flip_at)+ + done + +lemma supp_F[simp]: + "supp F1 = {}" "supp F2 = {}" "supp F3 = {}" + by (simp_all add: F1_def F2_def F3_def lam.supp supp_at_base) + blast+ + +lemma F_eqvt[eqvt]: + "p \ F1 = F1" "p \ F2 = F2" "p \ F3 = F3" + by (rule_tac [!] perm_supp_eq) + (simp_all add: fresh_star_def fresh_def) + +lemma F_app: + "F1 \ A \ App \ \Var\ \ (Var \ A)" + "F2 \ A \ B \ C \ (App \ (App \ \App\ \ (C \ A))) \ (C \ B)" + by (rule lam1_fast_app, rule Lam_F, simp_all) + (rule lam3_fast_app, rule Lam_F, simp_all) + +lemma F3_app: + assumes f: "atom x \ A" "atom x \ B" (* or A and B have empty support *) + shows "F3 \ A \ B \ App \ \Abs\ \ (Abs \ (\x. (B \ (A \ V x))))" +proof - + obtain y :: var where b: "atom y \ (x, A, B)" using obtain_fresh by blast + obtain z :: var where c: "atom z \ (x, y, A, B)" using obtain_fresh by blast + have *: "x \ z" "x \ y" "y \ z" + using b c by (simp_all add: fresh_Pair fresh_at_base) blast+ + have **: + "atom y \ z" "atom x \ z" "atom y \ x" + "atom z \ y" "atom z \ x" "atom x \ y" + "atom x \ A" "atom y \ A" "atom z \ A" + "atom x \ B" "atom y \ B" "atom z \ B" + using b c f by (simp_all add: fresh_Pair fresh_at_base) blast+ + show ?thesis + apply (simp add: Lam_F(3)[of y z x] * *[symmetric]) + apply (rule b3) apply (rule b5) apply (rule bI) + apply (simp add: ** fresh_Pair * *[symmetric]) + apply (rule b3) apply (rule bI) + apply (simp add: ** fresh_Pair * *[symmetric]) + apply (rule b1) + done +qed + +definition Lam_A1_pre : "A1 \ \cx. \cy. (F1 \ V cx)" +definition Lam_A2_pre : "A2 \ \cx. \cy. \cz. (F2 \ V cx \ V cy \ \[V cz]\)" +definition Lam_A3_pre : "A3 \ \cx. \cy. (F3 \ V cx \ \[V cy]\)" +lemma Lam_A: + "x \ y \ A1 = \x. \y. (F1 \ V x)" + "a \ b \ a \ c \ c \ b \ A2 = \a. \b. \c. (F2 \ V a \ V b \ \[V c]\)" + "a \ b \ A3 = \a. \b. (F3 \ V a \ \[V b]\)" + apply (simp_all add: Lam_A1_pre Lam_A2_pre Lam_A3_pre Abs1_eq_iff lam.fresh supp_at_base Var_App_Abs_eqvt numeral_eqvt flip_def[symmetric] fresh_at_base F_eqvt ltgt_eqvt) + apply (smt cx_cy_cz permute_flip_at)+ + done + +lemma supp_A[simp]: + "supp A1 = {}" "supp A2 = {}" "supp A3 = {}" + by (auto simp add: Lam_A1_pre Lam_A2_pre Lam_A3_pre lam.supp supp_at_base supp_Cons supp_Nil) + +lemma A_app: + "A1 \ A \ B \ F1 \ A" + "A2 \ A \ B \ C \ F2 \ A \ B \ \[C]\" + "A3 \ A \ B \ F3 \ A \ \[B]\" + apply (rule lam2_fast_app, rule Lam_A, simp_all) + apply (rule lam3_fast_app, rule Lam_A, simp_all) + apply (rule lam2_fast_app, rule Lam_A, simp_all) + done + +definition "Num \ \[\[A1,A2,A3]\]\" + +lemma supp_Num[simp]: + "supp Num = {}" + by (auto simp only: Num_def supp_ltgt supp_Pair supp_A supp_Cons supp_Nil) + +end