Up to index of Isabelle/HOL/HOL-Nominal/Examples
theory Contexts(* $Id: Contexts.thy,v 1.8 2008/05/22 14:34:46 urbanc Exp $ *) theory Contexts imports "../Nominal" begin text {* We show here that the Plotkin-style of defining beta-reduction (based on congruence rules) is equivalent to the Felleisen-Hieb-style representation (based on contexts). We also define cbv-evaluation via a CK-machine described by Roshan James. The interesting point in this theory is that contexts do not contain any binders. On the other hand the operation of filling a term into a context produces an alpha-equivalent term. *} atom_decl name text {* Lambda-Terms - the Lam-term constructor binds a name *} nominal_datatype lam = Var "name" | App "lam" "lam" | Lam "«name»lam" ("Lam [_]._" [100,100] 100) text {* Contexts - the lambda case in contexts does not bind a name, even if we introduce the notation [_]._ for CLam. *} nominal_datatype ctx = Hole ("\<box>" 1000) | CAppL "ctx" "lam" | CAppR "lam" "ctx" | CLam "name" "ctx" ("CLam [_]._" [100,100] 100) text {* Capture-avoiding substitution *} consts subst :: "lam => name => lam => lam" ("_[_::=_]" [100,100,100] 100) nominal_primrec "(Var x)[y::=s] = (if x=y then s else (Var x))" "(App t1 t2)[y::=s] = App (t1[y::=s]) (t2[y::=s])" "x\<sharp>(y,s) ==> (Lam [x].t)[y::=s] = Lam [x].(t[y::=s])" apply(finite_guess)+ apply(rule TrueI)+ apply(simp add: abs_fresh) apply(fresh_guess)+ done text {* Filling is the operation that fills a term into a hole. This operation is possibly capturing. *} consts filling :: "ctx => lam => lam" ("_[|_|]" [100,100] 100) nominal_primrec "\<box>[|t|] = t" "(CAppL E t')[|t|] = App (E[|t|]) t'" "(CAppR t' E)[|t|] = App t' (E[|t|])" "(CLam [x].E)[|t|] = Lam [x].(E[|t|])" by (rule TrueI)+ text {* While contexts themselves are not alpha-equivalence classes, the filling operation produces an alpha-equivalent lambda-term. *} lemma alpha_test: shows "x≠y ==> (CLam [x].\<box>) ≠ (CLam [y].\<box>)" and "(CLam [x].\<box>)[|Var x|] = (CLam [y].\<box>)[|Var y|]" by (auto simp add: alpha ctx.inject lam.inject calc_atm fresh_atm) text {* The composition of two contexts. *} consts ctx_compose :: "ctx => ctx => ctx" ("_ o _" [100,100] 100) nominal_primrec "\<box> o E' = E'" "(CAppL E t') o E' = CAppL (E o E') t'" "(CAppR t' E) o E' = CAppR t' (E o E')" "(CLam [x].E) o E' = CLam [x].(E o E')" by (rule TrueI)+ lemma ctx_compose: shows "(E1 o E2)[|t|] = E1[|E2[|t|]|]" by (induct E1 rule: ctx.induct) (auto) text {* Beta-reduction via contexts in the Felleisen-Hieb style. *} inductive ctx_red :: "lam=>lam=>bool" ("_ -->x _" [80,80] 80) where xbeta[intro]: "E[|App (Lam [x].t) t'|] -->x E[|t[x::=t']|]" text {* Beta-reduction via congruence rules in the Plotkin style. *} inductive cong_red :: "lam=>lam=>bool" ("_ -->o _" [80,80] 80) where obeta[intro]: "App (Lam [x].t) t' -->o t[x::=t']" | oapp1[intro]: "t -->o t' ==> App t t2 -->o App t' t2" | oapp2[intro]: "t -->o t' ==> App t2 t -->o App t2 t'" | olam[intro]: "t -->o t' ==> Lam [x].t -->o Lam [x].t'" text {* The proof that shows both relations are equal. *} lemma cong_red_in_ctx: assumes a: "t -->o t'" shows "E[|t|] -->o E[|t'|]" using a by (induct E rule: ctx.induct) (auto) lemma ctx_red_in_ctx: assumes a: "t -->x t'" shows "E[|t|] -->x E[|t'|]" using a by (induct) (auto simp add: ctx_compose[symmetric]) theorem ctx_red_implies_cong_red: assumes a: "t -->x t'" shows "t -->o t'" using a by (induct) (auto intro: cong_red_in_ctx) theorem cong_red_implies_ctx_red: assumes a: "t -->o t'" shows "t -->x t'" using a proof (induct) case (obeta x t' t) have "\<box>[|App (Lam [x].t) t'|] -->x \<box>[|t[x::=t']|]" by (rule xbeta) then show "App (Lam [x].t) t' -->x t[x::=t']" by simp qed (metis ctx_red_in_ctx filling.simps)+ (* found by SledgeHammer *) lemma ctx_existence: assumes a: "t -->o t'" shows "∃C s s'. t = C[|s|] ∧ t' = C[|s'|] ∧ s -->o s'" using a by (induct) (metis cong_red.intros filling.simps)+ section {* We now use context in a CBV list machine *} text {* First, a function that composes a list of contexts. *} primrec ctx_composes :: "ctx list => ctx" ("_\<down>" [80] 80) where "[]\<down> = \<box>" | "(E#Es)\<down> = (Es\<down>) o E" text {* Values *} inductive val :: "lam=>bool" where v_Lam[intro]: "val (Lam [x].e)" | v_Var[intro]: "val (Var x)" text {* CBV-reduction using contexts *} inductive cbv :: "lam=>lam=>bool" ("_ -->cbv _" [80,80] 80) where cbv_beta[intro]: "val v ==> E[|App (Lam [x].e) v|] -->cbv E[|e[x::=v]|]" text {* reflexive, transitive closure of CBV reduction *} inductive "cbv_star" :: "lam=>lam=>bool" (" _ -->cbv* _" [80,80] 80) where cbvs1[intro]: "e -->cbv* e" | cbvs2[intro]: "e -->cbv e' ==> e -->cbv* e'" | cbvs3[intro,trans]: "[|e1-->cbv* e2; e2 -->cbv* e3|] ==> e1 -->cbv* e3" text {* A little CK-machine, which builds up a list of evaluation contexts. *} inductive machine :: "lam=>ctx list=>lam=>ctx list=>bool" ("<_,_> \<mapsto> <_,_>" [60,60,60,60] 60) where m1[intro]: "<App e1 e2, Es> \<mapsto> <e1,(CAppL \<box> e2)#Es>" | m2[intro]: "val v ==> <v,(CAppL \<box> e2)#Es> \<mapsto> <e2,(CAppR v \<box>)#Es>" | m3[intro]: "val v ==> <v,(CAppR (Lam [x].e) \<box>)#Es> \<mapsto> <e[x::=v],Es>" lemma machine_red_implies_cbv_reds_aux: assumes a: "<e,Es> \<mapsto> <e',Es'>" shows "(Es\<down>)[|e|] -->cbv* (Es'\<down>)[|e'|]" using a by (induct) (auto simp add: ctx_compose) lemma machine_execution_implies_cbv_reds: assumes a: "<e,[]> \<mapsto> <e',[]>" shows "e -->cbv* e'" using a by (auto dest: machine_red_implies_cbv_reds_aux) text {* One would now like to show something about the opposite direction, by according to Amr Sabry this amounts to showing a standardisation lemma, which is hard. *} end
lemma alpha_test(1):
x ≠ y ==> CLam [x].\<box> ≠ CLam [y].\<box>
and alpha_test(2):
CLam [x].\<box>[|Var x|] = CLam [y].\<box>[|Var y|]
lemma ctx_compose:
E1.0 o E2.0[|t|] = E1.0[|E2.0[|t|]|]
lemma cong_red_in_ctx:
t -->o t' ==> E[|t|] -->o E[|t'|]
lemma ctx_red_in_ctx:
t -->x t' ==> E[|t|] -->x E[|t'|]
theorem ctx_red_implies_cong_red:
t -->x t' ==> t -->o t'
theorem cong_red_implies_ctx_red:
t -->o t' ==> t -->x t'
lemma ctx_existence:
t -->o t' ==> ∃C s s'. t = C[|s|] ∧ t' = C[|s'|] ∧ s -->o s'
lemma machine_red_implies_cbv_reds_aux:
<e,Es> \<mapsto> <e',Es'> ==> (Es\<down>)[|e|] -->cbv* (Es'\<down>)[|e'|]
lemma machine_execution_implies_cbv_reds:
<e,[]> \<mapsto> <e',[]> ==> e -->cbv* e'