diff -r 25a7f421a3ba -r 5635a968fd3f Nominal/CPS/CPS1_Plotkin.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Nominal/CPS/CPS1_Plotkin.thy Thu Jun 16 20:56:30 2011 +0900 @@ -0,0 +1,366 @@ +header {* CPS conversion *} +theory Plotkin +imports Lt +begin + +nominal_primrec + CPS :: "lt \ lt" ("_*" [250] 250) +where + "atom k \ x \ (x~)* = (Abs k ((k~) $ (x~)))" +| "atom k \ (x, M) \ (Abs x M)* = Abs k (k~ $ Abs x (M*))" +| "atom k \ (M, N) \ atom m \ (N, k) \ atom n \ (k, m) \ + (M $ N)* = Abs k (M* $ Abs m (N* $ Abs n (m~ $ n~ $ k~)))" +unfolding eqvt_def CPS_graph_def +apply (rule, perm_simp, rule, rule) +apply (simp_all add: fresh_Pair_elim) +apply (rule_tac y="x" in lt.exhaust) +apply (auto) +apply (rule_tac x="name" and ?'a="name" in obtain_fresh) +apply (simp_all add: fresh_at_base)[3] +apply (rule_tac x="(name, lt)" and ?'a="name" in obtain_fresh) +apply (simp add: fresh_Pair_elim fresh_at_base)[2] +apply (rule_tac x="(lt1, lt2)" and ?'a="name" in obtain_fresh) +apply (rule_tac x="(lt2, a)" and ?'a="name" in obtain_fresh) +apply (rule_tac x="(a, aa)" and ?'a="name" in obtain_fresh) +apply (simp add: fresh_Pair_elim fresh_at_base) +apply (simp add: Abs1_eq_iff lt.fresh fresh_at_base) +--"-" +apply(rule_tac s="[[atom ka]]lst. ka~ $ Abs x (CPS_sumC M)" in trans) +apply (case_tac "k = ka") +apply simp +apply(simp (no_asm) add: Abs1_eq_iff del:eqvts) +apply (simp del: eqvts add: lt.fresh fresh_at_base) +apply (simp only: lt.perm_simps(1) lt.perm_simps(3) flip_def[symmetric] lt.eq_iff(3)) +apply (subst flip_at_base_simps(2)) +apply simp +apply (intro conjI refl) +apply (rule flip_fresh_fresh[symmetric]) +apply (simp_all add: lt.fresh) +apply (metis fresh_eqvt_at lt.fsupp) +apply (case_tac "ka = x") +apply simp_all[2] +apply (metis Abs_fresh_iff(3) atom_eq_iff finite_set fresh_Cons fresh_Nil fresh_atom fresh_eqvt_at fresh_finite_atom_set fresh_set lt.fsupp) +apply (metis Abs_fresh_iff(3) atom_eq_iff finite_set fresh_Cons fresh_Nil fresh_atom fresh_eqvt_at fresh_finite_atom_set fresh_set lt.fsupp) +--"-" +apply (simp add: Abs1_eq(3)) +apply (erule Abs_lst1_fcb) +apply (simp add: fresh_def supp_Abs) +apply (drule_tac a="atom xa" in fresh_eqvt_at) +apply (simp add: finite_supp) +apply assumption +apply (simp add: fresh_def supp_Abs) +apply (simp add: eqvts eqvt_at_def) +apply simp +--"-" +apply (rename_tac k' M N m' n') +apply (subgoal_tac "atom k \ CPS_sumC M \ atom k' \ CPS_sumC M \ atom k \ CPS_sumC N \ atom k' \ CPS_sumC N \ + atom m \ CPS_sumC N \ atom m' \ CPS_sumC N") +prefer 2 +apply (intro conjI) +apply (erule fresh_eqvt_at, simp add: finite_supp, assumption)+ +apply clarify +apply (case_tac "k = k'", case_tac [!] "m' = k",case_tac [!]"m = k'",case_tac[!] "m = m'") +apply (simp_all add: Abs1_eq_iff lt.fresh flip_def[symmetric] fresh_at_base flip_fresh_fresh permute_eq_iff) +by (metis flip_at_base_simps(3) flip_at_simps(2) flip_commute permute_flip_at)+ + +termination + by (relation "measure size") (simp_all add: lt.size) + +lemmas [simp] = fresh_Pair_elim CPS.simps(2,3)[simplified fresh_Pair_elim] + +lemma [simp]: "supp (M*) = supp M" + by (induct rule: CPS.induct, simp_all add: lt.supp supp_at_base fresh_at_base fresh_def supp_Pair) + (simp_all only: atom_eq_iff[symmetric], blast+) + +lemma [simp]: "x \ M* = x \ M" + unfolding fresh_def by simp + +(* Will be provided automatically by nominal_primrec *) +lemma CPS_eqvt[eqvt]: + shows "p \ M* = (p \ M)*" + apply (induct M rule: lt.induct) + apply (rule_tac x="(name, p \ name, p)" and ?'a="name" in obtain_fresh) + apply simp + apply (simp add: Abs1_eq_iff lt.fresh flip_def[symmetric]) + apply (metis atom_eqvt flip_fresh_fresh fresh_perm atom_eq_iff fresh_at_base) + apply (rule_tac x="(name, lt, p \ name, p \ lt, p)" and ?'a="name" in obtain_fresh) + apply simp + apply (metis atom_eqvt fresh_perm atom_eq_iff) + apply (rule_tac x="(lt1, p \ lt1, lt2, p \ lt2, p)" and ?'a="name" in obtain_fresh) + apply (rule_tac x="(a, lt2, p \ lt2, p)" and ?'a="name" in obtain_fresh) + apply (rule_tac x="(a, aa, p)" and ?'a="name" in obtain_fresh) + apply (simp) + apply (simp add: Abs1_eq_iff lt.fresh flip_def[symmetric]) + apply (metis atom_eqvt fresh_perm atom_eq_iff) + done + +nominal_primrec + convert:: "lt => lt" ("_+" [250] 250) +where + "(Var x)+ = Var x" +| "(Abs x M)+ = Abs x (M*)" +| "(M $ N)+ = M $ N" + unfolding convert_graph_def eqvt_def + apply (rule, perm_simp, rule, rule) + apply (erule lt.exhaust) + apply (simp_all) + apply blast + apply (simp add: Abs1_eq_iff CPS_eqvt) + by blast + +termination + by (relation "measure size") (simp_all add: lt.size) + +lemma convert_supp[simp]: + shows "supp (M+) = supp M" + by (induct M rule: lt.induct, simp_all add: lt.supp) + +lemma convert_fresh[simp]: + shows "x \ (M+) = x \ M" + unfolding fresh_def by simp + +lemma convert_eqvt[eqvt]: + shows "p \ (M+) = (p \ M)+" + by (nominal_induct M rule: lt.strong_induct, auto simp add: CPS_eqvt) + +lemma [simp]: + shows "isValue (p \ (M::lt)) = isValue M" + by (nominal_induct M rule: lt.strong_induct) auto + +lemma [eqvt]: + shows "p \ isValue M = isValue (p \ M)" + by (induct M rule: lt.induct) (perm_simp, rule refl)+ + +nominal_primrec + Kapply :: "lt \ lt \ lt" (infixl ";" 100) +where + "Kapply (Abs x M) K = K $ (Abs x M)+" +| "Kapply (Var x) K = K $ Var x" +| "isValue M \ isValue N \ Kapply (M $ N) K = M+ $ N+ $ K" +| "isValue M \ \isValue N \ atom n \ M \ atom n \ K \ + Kapply (M $ N) K = N; (Abs n (M+ $ Var n $ K))" +| "\isValue M \ atom m \ N \ atom m \ K \ atom n \ m \ atom n \ K \ + Kapply (M $ N) K = M; (Abs m (N* $ (Abs n (Var m $ Var n $ K))))" + unfolding Kapply_graph_def eqvt_def + apply (rule, perm_simp, rule, rule) +apply (simp_all) +apply (case_tac x) +apply (rule_tac y="a" in lt.exhaust) +apply (auto) +apply (case_tac "isValue lt1") +apply (case_tac "isValue lt2") +apply (auto)[1] +apply (rule_tac x="(lt1, ba)" and ?'a="name" in obtain_fresh) +apply (simp add: fresh_Pair_elim fresh_at_base) +apply (rule_tac x="(lt2, ba)" and ?'a="name" in obtain_fresh) +apply (rule_tac x="(a, ba)" and ?'a="name" in obtain_fresh) +apply (simp add: fresh_Pair_elim fresh_at_base) +apply (auto simp add: Abs1_eq_iff eqvts)[1] +apply (rename_tac M N u K) +apply (subgoal_tac "Abs n (M+ $ n~ $ K) = Abs u (M+ $ u~ $ K)") +apply (simp only:) +apply (auto simp add: Abs1_eq_iff flip_def[symmetric] lt.fresh fresh_at_base flip_fresh_fresh[symmetric])[1] +apply (subgoal_tac "Abs m (Na* $ Abs n (m~ $ n~ $ Ka)) = Abs ma (Na* $ Abs na (ma~ $ na~ $ Ka))") +apply (simp only:) +apply (simp add: Abs1_eq_iff flip_def[symmetric] lt.fresh fresh_at_base) +apply (subgoal_tac "Ka = (n \ na) \ Ka") +apply (subgoal_tac "Ka = (m \ ma) \ Ka") +apply (subgoal_tac "Ka = (n \ (m \ ma) \ na) \ Ka") +apply (case_tac "m = ma") +apply simp_all +apply rule +apply (auto simp add: flip_fresh_fresh[symmetric]) +apply (metis flip_at_base_simps(3) flip_fresh_fresh permute_flip_at)+ +done + +termination + by (relation "measure (\(t, _). size t)") (simp_all add: lt.size) + +section{* lemma related to Kapply *} + +lemma [simp]: "isValue V \ V; K = K $ (V+)" + by (nominal_induct V rule: lt.strong_induct) auto + +section{* lemma related to CPS conversion *} + +lemma value_CPS: + assumes "isValue V" + and "atom a \ V" + shows "V* = Abs a (a~ $ V+)" + using assms +proof (nominal_induct V avoiding: a rule: lt.strong_induct, simp_all add: lt.fresh) + fix name :: name and lt aa + assume a: "atom name \ aa" "\b. \isValue lt; atom b \ lt\ \ lt* = Abs b (b~ $ lt+)" + "atom aa \ lt \ aa = name" + obtain ab :: name where b: "atom ab \ (name, lt, a)" using obtain_fresh by blast + show "Abs name lt* = Abs aa (aa~ $ Abs name (lt*))" using a b + by (simp add: Abs1_eq_iff fresh_at_base lt.fresh) +qed + +section{* first lemma CPS subst *} + +lemma CPS_subst_fv: + assumes *:"isValue V" + shows "((M[V/x])* = (M*)[V+/x])" +using * proof (nominal_induct M avoiding: V x rule: lt.strong_induct) + case (Var name) + assume *: "isValue V" + obtain a :: name where a: "atom a \ (x, name, V)" using obtain_fresh by blast + show "((name~)[V/x])* = (name~)*[V+/x]" using a + by (simp add: fresh_at_base * value_CPS) +next + case (Abs name lt V x) + assume *: "atom name \ V" "atom name \ x" "\b ba. isValue b \ (lt[b/ba])* = lt*[b+/ba]" + "isValue V" + obtain a :: name where a: "atom a \ (name, lt, lt[V/x], x, V)" using obtain_fresh by blast + show "(Abs name lt[V/x])* = Abs name lt*[V+/x]" using * a + by (simp add: fresh_at_base) +next + case (App lt1 lt2 V x) + assume *: "\b ba. isValue b \ (lt1[b/ba])* = lt1*[b+/ba]" "\b ba. isValue b \ (lt2[b/ba])* = lt2*[b+/ba]" + "isValue V" + obtain a :: name where a: "atom a \ (lt1[V/x], lt1, lt2[V/x], lt2, V, x)" using obtain_fresh by blast + obtain b :: name where b: "atom b \ (lt2[V/x], lt2, a, V, x)" using obtain_fresh by blast + obtain c :: name where c: "atom c \ (a, b, V, x)" using obtain_fresh by blast + show "((lt1 $ lt2)[V/x])* = (lt1 $ lt2)*[V+/x]" using * a b c + by (simp add: fresh_at_base) +qed + +lemma [simp]: "isValue V \ isValue (V+)" + by (nominal_induct V rule: lt.strong_induct, auto) + +lemma CPS_eval_Kapply: + assumes a: "isValue K" + shows "(M* $ K) \\<^isub>\\<^sup>* (M ; K)" + using a +proof (nominal_induct M avoiding: K rule: lt.strong_induct, simp_all) + case (Var name K) + assume *: "isValue K" + obtain a :: name where a: "atom a \ (name, K)" using obtain_fresh by blast + show "(name~)* $ K \\<^isub>\\<^sup>* K $ name~" using * a + by simp (rule evbeta', simp_all add: fresh_at_base) +next + fix name :: name and lt K + assume *: "atom name \ K" "\b. isValue b \ lt* $ b \\<^isub>\\<^sup>* lt ; b" "isValue K" + obtain a :: name where a: "atom a \ (name, K, lt)" using obtain_fresh by blast + then have b: "atom name \ a" using fresh_PairD(1) fresh_at_base atom_eq_iff by metis + show "Abs name lt* $ K \\<^isub>\\<^sup>* K $ Abs name (lt*)" using * a b + by simp (rule evbeta', simp_all) +next + fix lt1 lt2 K + assume *: "\b. isValue b \ lt1* $ b \\<^isub>\\<^sup>* lt1 ; b" "\b. isValue b \ lt2* $ b \\<^isub>\\<^sup>* lt2 ; b" "isValue K" + obtain a :: name where a: "atom a \ (lt1, lt2, K)" using obtain_fresh by blast + obtain b :: name where b: "atom b \ (lt1, lt2, K, a)" using obtain_fresh by blast + obtain c :: name where c: "atom c \ (lt1, lt2, K, a, b)" using obtain_fresh by blast + have d: "atom a \ lt1" "atom a \ lt2" "atom a \ K" "atom b \ lt1" "atom b \ lt2" "atom b \ K" "atom b \ a" + "atom c \ lt1" "atom c \ lt2" "atom c \ K" "atom c \ a" "atom c \ b" using fresh_Pair a b c by simp_all + have "(lt1 $ lt2)* $ K \\<^isub>\\<^sup>* lt1* $ Abs b (lt2* $ Abs c (b~ $ c~ $ K))" using * d + by (simp add: fresh_at_base) (rule evbeta', simp_all add: fresh_at_base) + also have "... \\<^isub>\\<^sup>* lt1 $ lt2 ; K" proof (cases "isValue lt1") + assume e: "isValue lt1" + have "lt1* $ Abs b (lt2* $ Abs c (b~ $ c~ $ K)) \\<^isub>\\<^sup>* Abs b (lt2* $ Abs c (b~ $ c~ $ K)) $ lt1+" + using * d e by simp + also have "... \\<^isub>\\<^sup>* lt2* $ Abs c (lt1+ $ c~ $ K)" + by (rule evbeta', simp_all add: * d e, metis d(12) fresh_at_base) + also have "... \\<^isub>\\<^sup>* lt1 $ lt2 ; K" proof (cases "isValue lt2") + assume f: "isValue lt2" + have "lt2* $ Abs c (lt1+ $ c~ $ K) \\<^isub>\\<^sup>* Abs c (lt1+ $ c~ $ K) $ lt2+" using * d e f by simp + also have "... \\<^isub>\\<^sup>* lt1+ $ lt2+ $ K" + by (rule evbeta', simp_all add: d e f) + finally show ?thesis using * d e f by simp + next + assume f: "\ isValue lt2" + have "lt2* $ Abs c (lt1+ $ c~ $ K) \\<^isub>\\<^sup>* lt2 ; Abs c (lt1+ $ c~ $ K)" using * d e f by simp + also have "... \\<^isub>\\<^sup>* lt2 ; Abs a (lt1+ $ a~ $ K)" using Kapply.simps(4) d e evs1 f by metis + finally show ?thesis using * d e f by simp + qed + finally show ?thesis . + qed (metis Kapply.simps(5) isValue.simps(2) * d) + finally show "(lt1 $ lt2)* $ K \\<^isub>\\<^sup>* lt1 $ lt2 ; K" . +qed + +lemma Kapply_eval: + assumes a: "M \\<^isub>\ N" "isValue K" + shows "(M; K) \\<^isub>\\<^sup>* (N; K)" + using assms +proof (induct arbitrary: K rule: eval.induct) + case (evbeta x V M) + fix K + assume a: "isValue K" "isValue V" "atom x \ V" + have "Abs x (M*) $ V+ $ K \\<^isub>\\<^sup>* ((M*)[V+/x] $ K)" + by (rule evs2,rule ev2,rule Lt.evbeta) (simp_all add: fresh_def a[simplified fresh_def] evs1) + also have "... = ((M[V/x])* $ K)" by (simp add: CPS_subst_fv a) + also have "... \\<^isub>\\<^sup>* (M[V/x] ; K)" by (rule CPS_eval_Kapply, simp_all add: a) + finally show "(Abs x M $ V ; K) \\<^isub>\\<^sup>* (M[V/x] ; K)" using a by simp +next + case (ev1 V M N) + fix V M N K + assume a: "isValue V" "M \\<^isub>\ N" "\K. isValue K \ M ; K \\<^isub>\\<^sup>* N ; K" "isValue K" + obtain a :: name where b: "atom a \ (V, K, M, N)" using obtain_fresh by blast + show "V $ M ; K \\<^isub>\\<^sup>* V $ N ; K" proof (cases "isValue N") + assume "\ isValue N" + then show "V $ M ; K \\<^isub>\\<^sup>* V $ N ; K" using a b by simp + next + assume n: "isValue N" + have c: "M; Abs a (V+ $ a~ $ K) \\<^isub>\\<^sup>* Abs a (V+ $ a~ $ K) $ N+" using a b by (simp add: n) + also have d: "... \\<^isub>\\<^sup>* V+ $ N+ $ K" by (rule evbeta') (simp_all add: a b n) + finally show "V $ M ; K \\<^isub>\\<^sup>* V $ N ; K" using a b by (simp add: n) + qed +next + case (ev2 M M' N) + assume *: "M \\<^isub>\ M'" "\K. isValue K \ M ; K \\<^isub>\\<^sup>* M' ; K" "isValue K" + obtain a :: name where a: "atom a \ (K, M, N, M')" using obtain_fresh by blast + obtain b :: name where b: "atom b \ (a, K, M, N, M', N+)" using obtain_fresh by blast + have d: "atom a \ K" "atom a \ M" "atom a \ N" "atom a \ M'" "atom b \ a" "atom b \ K" + "atom b \ M" "atom b \ N" "atom b \ M'" using a b fresh_Pair by simp_all + have "M $ N ; K \\<^isub>\\<^sup>* M' ; Abs a (N* $ Abs b (a~ $ b~ $ K))" using * d by simp + also have "... \\<^isub>\\<^sup>* M' $ N ; K" proof (cases "isValue M'") + assume "\ isValue M'" + then show ?thesis using * d by (simp_all add: evs1) + next + assume e: "isValue M'" + then have "M' ; Abs a (N* $ Abs b (a~ $ b~ $ K)) = Abs a (N* $ Abs b (a~ $ b~ $ K)) $ M'+" by simp + also have "... \\<^isub>\\<^sup>* (N* $ Abs b (a~ $ b~ $ K))[M'+/a]" + by (rule evbeta') (simp_all add: fresh_at_base e d) + also have "... = N* $ Abs b (M'+ $ b~ $ K)" using * d by (simp add: fresh_at_base) + also have "... \\<^isub>\\<^sup>* M' $ N ; K" proof (cases "isValue N") + assume f: "isValue N" + have "N* $ Abs b (M'+ $ b~ $ K) \\<^isub>\\<^sup>* Abs b (M'+ $ b~ $ K) $ N+" + by (rule eval_trans, rule CPS_eval_Kapply) (simp_all add: d e f * evs1) + also have "... \\<^isub>\\<^sup>* M' $ N ; K" by (rule evbeta') (simp_all add: d e f evs1) + finally show ?thesis . + next + assume "\ isValue N" + then show ?thesis using d e + by (metis CPS_eval_Kapply Kapply.simps(4) isValue.simps(2)) + qed + finally show ?thesis . + qed + finally show ?case . +qed + +lemma Kapply_eval_rtrancl: + assumes H: "M \\<^isub>\\<^sup>* N" and "isValue K" + shows "(M;K) \\<^isub>\\<^sup>* (N;K)" + using H + by (induct) (metis Kapply_eval assms(2) eval_trans evs1)+ + +lemma + assumes "isValue V" "M \\<^isub>\\<^sup>* V" + shows "M* $ (Abs x (x~)) \\<^isub>\\<^sup>* V+" +proof- + obtain y::name where *: "atom y \ V" using obtain_fresh by blast + have e: "Abs x (x~) = Abs y (y~)" + by (simp add: Abs1_eq_iff lt.fresh fresh_at_base) + have "M* $ Abs x (x~) \\<^isub>\\<^sup>* M ; Abs x (x~)" + by(rule CPS_eval_Kapply,simp_all add: assms) + also have "... \\<^isub>\\<^sup>* (V ; Abs x (x~))" by (rule Kapply_eval_rtrancl, simp_all add: assms) + also have "... = V ; Abs y (y~)" using e by (simp only:) + also have "... \\<^isub>\\<^sup>* (V+)" by (simp add: assms, rule evbeta') (simp_all add: assms *) + finally show "M* $ (Abs x (x~)) \\<^isub>\\<^sup>* (V+)" . +qed + +end + + +