diff -r da9bed7baf23 -r 2abc8cb46a5c Tutorial/Tutorial4.thy --- a/Tutorial/Tutorial4.thy Fri Jan 21 22:58:03 2011 +0100 +++ b/Tutorial/Tutorial4.thy Sat Jan 22 12:46:01 2011 -0600 @@ -5,18 +5,62 @@ section {* The CBV Reduction Relation (Small-Step Semantics) *} + +inductive + cbv :: "lam \ lam \ bool" ("_ \cbv _" [60, 60] 60) +where + cbv1: "\val v; atom x \ v\ \ App (Lam [x].t) v \cbv t[x ::= v]" +| cbv2: "t \cbv t' \ App t t2 \cbv App t' t2" +| cbv3: "t \cbv t' \ App t2 t \cbv App t2 t'" + +inductive + "cbv_star" :: "lam \ lam \ bool" (" _ \cbv* _" [60, 60] 60) +where + cbvs1: "e \cbv* e" +| cbvs2: "\e1\cbv e2; e2 \cbv* e3\ \ e1 \cbv* e3" + +declare cbv.intros[intro] cbv_star.intros[intro] + +subsection {* EXERCISE 3 *} + +text {* + Show that cbv* is transitive, by filling the gaps in the + proof below. +*} + +lemma + assumes a: "e1 \cbv* e2" "e2 \cbv* e3" + shows "e1 \cbv* e3" +using a +proof (induct) + case (cbvs1 e1) + have asm: "e1 \cbv* e3" by fact + show "e1 \cbv* e3" sorry +next + case (cbvs2 e1 e2 e3') + have "e1 \cbv e2" by fact + have "e2 \cbv* e3'" by fact + have "e3' \cbv* e3 \ e2 \cbv* e3" by fact + have "e3' \cbv* e3" by fact + + show "e1 \cbv* e3" sorry +qed + +lemma cbvs3 [intro]: + assumes a: "e1 \cbv* e2" "e2 \cbv* e3" + shows "e1 \cbv* e3" +using a by (induct) (auto) + + + + + text {* In order to help establishing the property that the CK Machine calculates a nomrmalform that corresponds to the evaluation relation, we introduce the call-by-value small-step semantics. *} -inductive - cbv :: "lam \ lam \ bool" ("_ \cbv _" [60, 60] 60) -where - cbv1: "\val v; atom x \ v\ \ App (Lam [x].t) v \cbv t[x ::= v]" -| cbv2[intro]: "t \cbv t' \ App t t2 \cbv App t' t2" -| cbv3[intro]: "t \cbv t' \ App t2 t \cbv App t2 t'" equivariance val equivariance cbv @@ -44,7 +88,7 @@ lemma better_cbv1 [intro]: assumes a: "val v" - shows "App (Lam [x].t) v \cbv t[x::=v]" + shows "App (Lam [x].t) v \cbv t[x ::= v]" proof - obtain y::"name" where fs: "atom y \ (x, t, v)" by (rule obtain_fresh) have "App (Lam [x].t) v = App (Lam [y].((y \ x) \ t)) v" using fs @@ -58,16 +102,8 @@ The transitive closure of the cbv-reduction relation: *} -inductive - "cbvs" :: "lam \ lam \ bool" (" _ \cbv* _" [60, 60] 60) -where - cbvs1[intro]: "e \cbv* e" -| cbvs2[intro]: "\e1\cbv e2; e2 \cbv* e3\ \ e1 \cbv* e3" -lemma cbvs3 [intro]: - assumes a: "e1 \cbv* e2" "e2 \cbv* e3" - shows "e1 \cbv* e3" -using a by (induct) (auto) + subsection {* EXERCISE 8 *}