| author | Christian Urban <urbanc@in.tum.de> | 
| Sat, 22 Jan 2011 18:59:48 -0600 | |
| changeset 2701 | 7b2691911fbc | 
| parent 2693 | 2abc8cb46a5c | 
| child 2706 | 8ae1c2e6369e | 
| permissions | -rw-r--r-- | 
| 2691 
abb6c3ac2df2
separated type preservation and progress into a separate file
 Christian Urban <urbanc@in.tum.de> parents: 
2689diff
changeset | 1 | |
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 2 | theory Tutorial4 | 
| 2701 | 3 | imports Tutorial1 Tutorial2 | 
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 4 | begin | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 5 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 6 | section {* The CBV Reduction Relation (Small-Step Semantics) *}
 | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 7 | |
| 2693 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 8 | |
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 9 | inductive | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 10 |   cbv :: "lam \<Rightarrow> lam \<Rightarrow> bool" ("_ \<longrightarrow>cbv _" [60, 60] 60) 
 | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 11 | where | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 12 | cbv1: "\<lbrakk>val v; atom x \<sharp> v\<rbrakk> \<Longrightarrow> App (Lam [x].t) v \<longrightarrow>cbv t[x ::= v]" | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 13 | | cbv2: "t \<longrightarrow>cbv t' \<Longrightarrow> App t t2 \<longrightarrow>cbv App t' t2" | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 14 | | cbv3: "t \<longrightarrow>cbv t' \<Longrightarrow> App t2 t \<longrightarrow>cbv App t2 t'" | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 15 | |
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 16 | inductive | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 17 |   "cbv_star" :: "lam \<Rightarrow> lam \<Rightarrow> bool" (" _ \<longrightarrow>cbv* _" [60, 60] 60)
 | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 18 | where | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 19 | cbvs1: "e \<longrightarrow>cbv* e" | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 20 | | cbvs2: "\<lbrakk>e1\<longrightarrow>cbv e2; e2 \<longrightarrow>cbv* e3\<rbrakk> \<Longrightarrow> e1 \<longrightarrow>cbv* e3" | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 21 | |
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 22 | declare cbv.intros[intro] cbv_star.intros[intro] | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 23 | |
| 2701 | 24 | subsection {* EXERCISE 11 *}
 | 
| 2693 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 25 | |
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 26 | text {*
 | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 27 | Show that cbv* is transitive, by filling the gaps in the | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 28 | proof below. | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 29 | *} | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 30 | |
| 2701 | 31 | lemma cbvs3 [intro]: | 
| 2693 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 32 | assumes a: "e1 \<longrightarrow>cbv* e2" "e2 \<longrightarrow>cbv* e3" | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 33 | shows "e1 \<longrightarrow>cbv* e3" | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 34 | using a | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 35 | proof (induct) | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 36 | case (cbvs1 e1) | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 37 | have asm: "e1 \<longrightarrow>cbv* e3" by fact | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 38 | show "e1 \<longrightarrow>cbv* e3" sorry | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 39 | next | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 40 | case (cbvs2 e1 e2 e3') | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 41 | have "e1 \<longrightarrow>cbv e2" by fact | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 42 | have "e2 \<longrightarrow>cbv* e3'" by fact | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 43 | have "e3' \<longrightarrow>cbv* e3 \<Longrightarrow> e2 \<longrightarrow>cbv* e3" by fact | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 44 | have "e3' \<longrightarrow>cbv* e3" by fact | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 45 | |
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 46 | show "e1 \<longrightarrow>cbv* e3" sorry | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 47 | qed | 
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 48 | |
| 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 49 | |
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 50 | text {*
 | 
| 2701 | 51 | In order to help establishing the property that the machine | 
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 52 | calculates a nomrmalform that corresponds to the evaluation | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 53 | relation, we introduce the call-by-value small-step semantics. | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 54 | *} | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 55 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 56 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 57 | equivariance val | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 58 | equivariance cbv | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 59 | nominal_inductive cbv | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 60 | avoids cbv1: "x" | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 61 | unfolding fresh_star_def | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 62 | by (simp_all add: lam.fresh Abs_fresh_iff fresh_Pair fresh_fact) | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 63 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 64 | text {*
 | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 65 | In order to satisfy the vc-condition we have to formulate | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 66 | this relation with the additional freshness constraint | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 67 | atom x \<sharp> v. Although this makes the definition vc-ompatible, it | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 68 | makes the definition less useful. We can with a little bit of | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 69 | pain show that the more restricted rule is equivalent to the | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 70 | usual rule. | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 71 | *} | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 72 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 73 | lemma subst_rename: | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 74 | assumes a: "atom y \<sharp> t" | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 75 | shows "t[x ::= s] = ((y \<leftrightarrow> x) \<bullet> t)[y ::= s]" | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 76 | using a | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 77 | by (nominal_induct t avoiding: x y s rule: lam.strong_induct) | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 78 | (auto simp add: lam.fresh fresh_at_base) | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 79 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 80 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 81 | lemma better_cbv1 [intro]: | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 82 | assumes a: "val v" | 
| 2693 
2abc8cb46a5c
better version of Tutorial 1
 Christian Urban <urbanc@in.tum.de> parents: 
2691diff
changeset | 83 | shows "App (Lam [x].t) v \<longrightarrow>cbv t[x ::= v]" | 
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 84 | proof - | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 85 | obtain y::"name" where fs: "atom y \<sharp> (x, t, v)" by (rule obtain_fresh) | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 86 | have "App (Lam [x].t) v = App (Lam [y].((y \<leftrightarrow> x) \<bullet> t)) v" using fs | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 87 | by (auto simp add: lam.eq_iff Abs1_eq_iff' flip_def fresh_Pair fresh_at_base) | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 88 | also have "\<dots> \<longrightarrow>cbv ((y \<leftrightarrow> x) \<bullet> t)[y ::= v]" using fs a cbv1 by auto | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 89 | also have "\<dots> = t[x ::= v]" using fs subst_rename[symmetric] by simp | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 90 | finally show "App (Lam [x].t) v \<longrightarrow>cbv t[x ::= v]" by simp | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 91 | qed | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 92 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 93 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 94 | |
| 2701 | 95 | subsection {* EXERCISE 12 *}
 | 
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 96 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 97 | text {*  
 | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 98 | If more simple exercises are needed, then complete the following proof. | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 99 | *} | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 100 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 101 | lemma cbv_in_ctx: | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 102 | assumes a: "t \<longrightarrow>cbv t'" | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 103 | shows "E\<lbrakk>t\<rbrakk> \<longrightarrow>cbv E\<lbrakk>t'\<rbrakk>" | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 104 | using a | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 105 | proof (induct E) | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 106 | case Hole | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 107 | have "t \<longrightarrow>cbv t'" by fact | 
| 2701 | 108 | show "\<box>\<lbrakk>t\<rbrakk> \<longrightarrow>cbv \<box>\<lbrakk>t'\<rbrakk>" sorry | 
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 109 | next | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 110 | case (CAppL E s) | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 111 | have ih: "t \<longrightarrow>cbv t' \<Longrightarrow> E\<lbrakk>t\<rbrakk> \<longrightarrow>cbv E\<lbrakk>t'\<rbrakk>" by fact | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 112 | have "t \<longrightarrow>cbv t'" by fact | 
| 2701 | 113 | |
| 114 | show "(CAppL E s)\<lbrakk>t\<rbrakk> \<longrightarrow>cbv (CAppL E s)\<lbrakk>t'\<rbrakk>" sorry | |
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 115 | next | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 116 | case (CAppR s E) | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 117 | have ih: "t \<longrightarrow>cbv t' \<Longrightarrow> E\<lbrakk>t\<rbrakk> \<longrightarrow>cbv E\<lbrakk>t'\<rbrakk>" by fact | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 118 | have a: "t \<longrightarrow>cbv t'" by fact | 
| 2701 | 119 | |
| 120 | show "(CAppR s E)\<lbrakk>t\<rbrakk> \<longrightarrow>cbv (CAppR s E)\<lbrakk>t'\<rbrakk>" sorry | |
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 121 | qed | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 122 | |
| 2701 | 123 | section {* EXERCISE 13 *} 
 | 
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 124 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 125 | text {*
 | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 126 | The point of the cbv-reduction was that we can easily relatively | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 127 | establish the follwoing property: | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 128 | *} | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 129 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 130 | lemma machine_implies_cbvs_ctx: | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 131 | assumes a: "<e, Es> \<mapsto> <e', Es'>" | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 132 | shows "(Es\<down>)\<lbrakk>e\<rbrakk> \<longrightarrow>cbv* (Es'\<down>)\<lbrakk>e'\<rbrakk>" | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 133 | using a | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 134 | proof (induct) | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 135 | case (m1 t1 t2 Es) | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 136 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 137 | show "Es\<down>\<lbrakk>App t1 t2\<rbrakk> \<longrightarrow>cbv* ((CAppL \<box> t2) # Es)\<down>\<lbrakk>t1\<rbrakk>" sorry | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 138 | next | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 139 | case (m2 v t2 Es) | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 140 | have "val v" by fact | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 141 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 142 | show "((CAppL \<box> t2) # Es)\<down>\<lbrakk>v\<rbrakk> \<longrightarrow>cbv* (CAppR v \<box> # Es)\<down>\<lbrakk>t2\<rbrakk>" sorry | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 143 | next | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 144 | case (m3 v x t Es) | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 145 | have "val v" by fact | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 146 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 147 | show "(((CAppR (Lam [x].t) \<box>) # Es)\<down>)\<lbrakk>v\<rbrakk> \<longrightarrow>cbv* (Es\<down>)\<lbrakk>(t[x ::= v])\<rbrakk>" sorry | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 148 | qed | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 149 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 150 | text {* 
 | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 151 | It is not difficult to extend the lemma above to | 
| 2701 | 152 | arbitrary reductions sequences of the machine. | 
| 153 | *} | |
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 154 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 155 | lemma machines_implies_cbvs_ctx: | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 156 | assumes a: "<e, Es> \<mapsto>* <e', Es'>" | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 157 | shows "(Es\<down>)\<lbrakk>e\<rbrakk> \<longrightarrow>cbv* (Es'\<down>)\<lbrakk>e'\<rbrakk>" | 
| 2689 | 158 | using a machine_implies_cbvs_ctx | 
| 159 | by (induct) (blast)+ | |
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 160 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 161 | text {* 
 | 
| 2701 | 162 | So whenever we let the machine start in an initial | 
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 163 | state and it arrives at a final state, then there exists | 
| 2689 | 164 | a corresponding cbv-reduction sequence. | 
| 165 | *} | |
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 166 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 167 | corollary machines_implies_cbvs: | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 168 | assumes a: "<e, []> \<mapsto>* <e', []>" | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 169 | shows "e \<longrightarrow>cbv* e'" | 
| 2689 | 170 | proof - | 
| 171 | have "[]\<down>\<lbrakk>e\<rbrakk> \<longrightarrow>cbv* []\<down>\<lbrakk>e'\<rbrakk>" | |
| 172 | using a machines_implies_cbvs_ctx by blast | |
| 173 | then show "e \<longrightarrow>cbv* e'" by simp | |
| 174 | qed | |
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 175 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 176 | text {*
 | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 177 | We now want to relate the cbv-reduction to the evaluation | 
| 2701 | 178 | relation. For this we need one auxiliary lemma about | 
| 179 | inverting the e_App rule. | |
| 2689 | 180 | *} | 
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 181 | |
| 2691 
abb6c3ac2df2
separated type preservation and progress into a separate file
 Christian Urban <urbanc@in.tum.de> parents: 
2689diff
changeset | 182 | |
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 183 | lemma e_App_elim: | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 184 | assumes a: "App t1 t2 \<Down> v" | 
| 2689 | 185 | obtains x t v' where "t1 \<Down> Lam [x].t" "t2 \<Down> v'" "t[x::=v'] \<Down> v" | 
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 186 | using a by (cases) (auto simp add: lam.eq_iff lam.distinct) | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 187 | |
| 2689 | 188 | |
| 2701 | 189 | subsection {* EXERCISE 13 *}
 | 
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 190 | |
| 2689 | 191 | text {*
 | 
| 192 | Complete the first and second case in the | |
| 193 | proof below. | |
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 194 | *} | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 195 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 196 | lemma cbv_eval: | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 197 | assumes a: "t1 \<longrightarrow>cbv t2" "t2 \<Down> t3" | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 198 | shows "t1 \<Down> t3" | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 199 | using a | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 200 | proof(induct arbitrary: t3) | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 201 | case (cbv1 v x t t3) | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 202 | have a1: "val v" by fact | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 203 | have a2: "t[x ::= v] \<Down> t3" by fact | 
| 2701 | 204 | |
| 205 | show "App (Lam [x].t) v \<Down> t3" sorry | |
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 206 | next | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 207 | case (cbv2 t t' t2 t3) | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 208 | have ih: "\<And>t3. t' \<Down> t3 \<Longrightarrow> t \<Down> t3" by fact | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 209 | have "App t' t2 \<Down> t3" by fact | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 210 | then obtain x t'' v' | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 211 | where a1: "t' \<Down> Lam [x].t''" | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 212 | and a2: "t2 \<Down> v'" | 
| 2689 | 213 | and a3: "t''[x ::= v'] \<Down> t3" by (rule e_App_elim) | 
| 2701 | 214 | |
| 215 | show "App t t2 \<Down> t3" sorry | |
| 2689 | 216 | qed (auto elim!: e_App_elim) | 
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 217 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 218 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 219 | text {* 
 | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 220 | Next we extend the lemma above to arbitray initial | 
| 2701 | 221 | sequences of cbv-reductions. | 
| 222 | *} | |
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 223 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 224 | lemma cbvs_eval: | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 225 | assumes a: "t1 \<longrightarrow>cbv* t2" "t2 \<Down> t3" | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 226 | shows "t1 \<Down> t3" | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 227 | using a by (induct) (auto intro: cbv_eval) | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 228 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 229 | text {* 
 | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 230 | Finally, we can show that if from a term t we reach a value | 
| 2689 | 231 | by a cbv-reduction sequence, then t evaluates to this value. | 
| 2701 | 232 | |
| 233 | This proof is not by induction. So we have to set up the proof | |
| 234 | with | |
| 235 | ||
| 236 | proof - | |
| 237 | ||
| 238 | in order to prevent Isabelle from applying a default introduction | |
| 239 | rule. | |
| 2689 | 240 | *} | 
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 241 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 242 | lemma cbvs_implies_eval: | 
| 2701 | 243 | assumes a: "t \<longrightarrow>cbv* v" | 
| 244 | and b: "val v" | |
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 245 | shows "t \<Down> v" | 
| 2701 | 246 | proof - | 
| 247 | have "v \<Down> v" using b eval_val by simp | |
| 248 | then show "t \<Down> v" using a cbvs_eval by auto | |
| 249 | qed | |
| 250 | ||
| 251 | section {* EXERCISE 15 *}
 | |
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 252 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 253 | text {* 
 | 
| 2701 | 254 | All facts tied together give us the desired property | 
| 255 | about machines: we know that a machines transitions | |
| 256 | correspond to cbvs transitions, and with the lemma | |
| 257 | above they correspond to an eval judgement. | |
| 2689 | 258 | *} | 
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 259 | |
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 260 | theorem machines_implies_eval: | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 261 | assumes a: "<t1, []> \<mapsto>* <t2, []>" | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 262 | and b: "val t2" | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 263 | shows "t1 \<Down> t2" | 
| 2701 | 264 | proof - | 
| 265 | ||
| 266 | show "t1 \<Down> t2" sorry | |
| 2687 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 267 | qed | 
| 
d0fb94035969
first split of tutorrial theory
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 268 | |
| 2689 | 269 | end | 
| 270 |