| 27162 |      1 | (* $Id$ *)
 | 
|  |      2 | 
 | 
|  |      3 | theory CK_Machine 
 | 
|  |      4 |   imports "../Nominal" 
 | 
|  |      5 | begin
 | 
|  |      6 | 
 | 
|  |      7 | text {* 
 | 
|  |      8 |   
 | 
|  |      9 |   This theory establishes soundness and completeness for a CK-machine
 | 
|  |     10 |   with respect to a cbv-big-step semantics. The language includes 
 | 
|  |     11 |   functions, recursion, booleans and numbers. In the soundness proof 
 | 
| 27247 |     12 |   the small-step cbv-reduction relation is used in order to get the 
 | 
|  |     13 |   induction through. The type-preservation property is proved for the 
 | 
|  |     14 |   machine and  also for the small- and big-step semantics. Finally, 
 | 
|  |     15 |   the progress property is proved for the small-step semantics.
 | 
| 27162 |     16 | 
 | 
| 27247 |     17 |   The development is inspired by notes about context machines written
 | 
|  |     18 |   by Roshan James (Indiana University) and also by the lecture notes 
 | 
|  |     19 |   written by Andy Pitts for his semantics course. See
 | 
| 27162 |     20 |   
 | 
| 27247 |     21 |      http://www.cs.indiana.edu/~rpjames/lm.pdf
 | 
|  |     22 |      http://www.cl.cam.ac.uk/teaching/2001/Semantics/
 | 
| 27162 |     23 | 
 | 
|  |     24 | *}
 | 
|  |     25 | 
 | 
|  |     26 | atom_decl name
 | 
|  |     27 | 
 | 
|  |     28 | nominal_datatype lam =
 | 
|  |     29 |   VAR "name"
 | 
|  |     30 | | APP "lam" "lam" 
 | 
|  |     31 | | LAM "\<guillemotleft>name\<guillemotright>lam" ("LAM [_]._")
 | 
|  |     32 | | NUM "nat"
 | 
| 27247 |     33 | | DIFF "lam" "lam" ("_ -- _")    (* subtraction *) 
 | 
|  |     34 | | PLUS "lam" "lam" ("_ ++ _")    (* addition *)
 | 
| 27162 |     35 | | TRUE
 | 
|  |     36 | | FALSE
 | 
|  |     37 | | IF "lam" "lam" "lam"
 | 
| 27247 |     38 | | FIX "\<guillemotleft>name\<guillemotright>lam" ("FIX [_]._")  (* recursion *)
 | 
| 27162 |     39 | | ZET "lam"                      (* zero test *)
 | 
|  |     40 | | EQI "lam" "lam"                (* equality test on numbers *)
 | 
|  |     41 | 
 | 
|  |     42 | section {* Capture-Avoiding Substitution *}
 | 
|  |     43 | 
 | 
|  |     44 | consts subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam"  ("_[_::=_]" [100,100,100] 100)
 | 
|  |     45 | 
 | 
|  |     46 | nominal_primrec
 | 
|  |     47 |   "(VAR x)[y::=s] = (if x=y then s else (VAR x))"
 | 
|  |     48 |   "(APP t\<^isub>1 t\<^isub>2)[y::=s] = APP (t\<^isub>1[y::=s]) (t\<^isub>2[y::=s])"
 | 
|  |     49 |   "x\<sharp>(y,s) \<Longrightarrow> (LAM [x].t)[y::=s] = LAM [x].(t[y::=s])"
 | 
|  |     50 |   "(NUM n)[y::=s] = NUM n"
 | 
|  |     51 |   "(t\<^isub>1 -- t\<^isub>2)[y::=s] = (t\<^isub>1[y::=s]) -- (t\<^isub>2[y::=s])"
 | 
|  |     52 |   "(t\<^isub>1 ++ t\<^isub>2)[y::=s] = (t\<^isub>1[y::=s]) ++ (t\<^isub>2[y::=s])"
 | 
|  |     53 |   "x\<sharp>(y,s) \<Longrightarrow> (FIX [x].t)[y::=s] = FIX [x].(t[y::=s])"
 | 
|  |     54 |   "TRUE[y::=s] = TRUE"
 | 
|  |     55 |   "FALSE[y::=s] = FALSE"
 | 
|  |     56 |   "(IF t1 t2 t3)[y::=s] = IF (t1[y::=s]) (t2[y::=s]) (t3[y::=s])"
 | 
|  |     57 |   "(ZET t)[y::=s] = ZET (t[y::=s])"
 | 
|  |     58 |   "(EQI t1 t2)[y::=s] = EQI (t1[y::=s]) (t2[y::=s])"
 | 
|  |     59 | apply(finite_guess)+
 | 
|  |     60 | apply(rule TrueI)+
 | 
|  |     61 | apply(simp add: abs_fresh)+
 | 
|  |     62 | apply(fresh_guess)+
 | 
|  |     63 | done
 | 
|  |     64 | 
 | 
|  |     65 | lemma  subst_eqvt[eqvt]:
 | 
|  |     66 |   fixes pi::"name prm"
 | 
|  |     67 |   shows "pi\<bullet>(t1[x::=t2]) = (pi\<bullet>t1)[(pi\<bullet>x)::=(pi\<bullet>t2)]"
 | 
|  |     68 | by (nominal_induct t1 avoiding: x t2 rule: lam.strong_induct)
 | 
|  |     69 |    (auto simp add: perm_bij fresh_atm fresh_bij)
 | 
|  |     70 | 
 | 
|  |     71 | lemma fresh_fact:
 | 
|  |     72 |   fixes z::"name"
 | 
|  |     73 |   shows "\<lbrakk>z\<sharp>s; (z=y \<or> z\<sharp>t)\<rbrakk> \<Longrightarrow> z\<sharp>t[y::=s]"
 | 
|  |     74 | by (nominal_induct t avoiding: z y s rule: lam.strong_induct)
 | 
|  |     75 |    (auto simp add: abs_fresh fresh_prod fresh_atm fresh_nat)
 | 
|  |     76 | 
 | 
|  |     77 | lemma subst_rename: 
 | 
|  |     78 |   assumes a: "y\<sharp>t"
 | 
|  |     79 |   shows "t[x::=s] = ([(y,x)]\<bullet>t)[y::=s]"
 | 
|  |     80 | using a 
 | 
|  |     81 | by (nominal_induct t avoiding: x y s rule: lam.strong_induct)
 | 
|  |     82 |    (auto simp add: calc_atm fresh_atm abs_fresh perm_nat_def)
 | 
|  |     83 | 
 | 
|  |     84 | section {* Evaluation Contexts *}
 | 
|  |     85 | 
 | 
|  |     86 | datatype ctx = 
 | 
|  |     87 |     Hole ("\<box>")  
 | 
|  |     88 |   | CAPPL "ctx" "lam"
 | 
|  |     89 |   | CAPPR "lam" "ctx"
 | 
|  |     90 |   | CDIFFL "ctx" "lam"
 | 
|  |     91 |   | CDIFFR "lam" "ctx"
 | 
|  |     92 |   | CPLUSL "ctx" "lam"
 | 
|  |     93 |   | CPLUSR "lam" "ctx"
 | 
|  |     94 |   | CIF "ctx" "lam" "lam"
 | 
|  |     95 |   | CZET "ctx"
 | 
|  |     96 |   | CEQIL "ctx" "lam"
 | 
|  |     97 |   | CEQIR "lam" "ctx"
 | 
|  |     98 | 
 | 
| 27247 |     99 | text {* The operation of filling a term into a context: *} 
 | 
|  |    100 | 
 | 
| 27162 |    101 | fun
 | 
|  |    102 |   filling :: "ctx \<Rightarrow> lam \<Rightarrow> lam" ("_\<lbrakk>_\<rbrakk>")
 | 
|  |    103 | where
 | 
|  |    104 |   "\<box>\<lbrakk>t\<rbrakk> = t"
 | 
|  |    105 | | "(CAPPL E t')\<lbrakk>t\<rbrakk> = APP (E\<lbrakk>t\<rbrakk>) t'"
 | 
|  |    106 | | "(CAPPR t' E)\<lbrakk>t\<rbrakk> = APP t' (E\<lbrakk>t\<rbrakk>)"
 | 
|  |    107 | | "(CDIFFL E t')\<lbrakk>t\<rbrakk> = (E\<lbrakk>t\<rbrakk>) -- t'"
 | 
|  |    108 | | "(CDIFFR t' E)\<lbrakk>t\<rbrakk> = t' -- (E\<lbrakk>t\<rbrakk>)"
 | 
|  |    109 | | "(CPLUSL E t')\<lbrakk>t\<rbrakk> = (E\<lbrakk>t\<rbrakk>) ++ t'"
 | 
|  |    110 | | "(CPLUSR t' E)\<lbrakk>t\<rbrakk> = t' ++ (E\<lbrakk>t\<rbrakk>)"
 | 
|  |    111 | | "(CIF E t1 t2)\<lbrakk>t\<rbrakk> = IF (E\<lbrakk>t\<rbrakk>) t1 t2"
 | 
|  |    112 | | "(CZET E)\<lbrakk>t\<rbrakk> = ZET (E\<lbrakk>t\<rbrakk>)"
 | 
|  |    113 | | "(CEQIL E t')\<lbrakk>t\<rbrakk> = EQI (E\<lbrakk>t\<rbrakk>) t'"
 | 
|  |    114 | | "(CEQIR t' E)\<lbrakk>t\<rbrakk> = EQI t' (E\<lbrakk>t\<rbrakk>)"
 | 
|  |    115 | 
 | 
| 27247 |    116 | text {* The operation of composing two contexts: *}
 | 
|  |    117 | 
 | 
| 27162 |    118 | fun 
 | 
|  |    119 |  ctx_compose :: "ctx \<Rightarrow> ctx \<Rightarrow> ctx" ("_ \<circ> _")
 | 
|  |    120 | where
 | 
|  |    121 |   "\<box> \<circ> E' = E'"
 | 
|  |    122 | | "(CAPPL E t') \<circ> E' = CAPPL (E \<circ> E') t'"
 | 
|  |    123 | | "(CAPPR t' E) \<circ> E' = CAPPR t' (E \<circ> E')"
 | 
|  |    124 | | "(CDIFFL E t') \<circ> E' = CDIFFL (E \<circ> E') t'"
 | 
|  |    125 | | "(CDIFFR t' E) \<circ> E' = CDIFFR t' (E \<circ> E')"
 | 
|  |    126 | | "(CPLUSL E t') \<circ> E' = CPLUSL (E \<circ> E') t'"
 | 
|  |    127 | | "(CPLUSR t' E) \<circ> E' = CPLUSR t' (E \<circ> E')"
 | 
|  |    128 | | "(CIF E t1 t2) \<circ> E' = CIF (E \<circ> E') t1 t2"
 | 
|  |    129 | | "(CZET E) \<circ> E' = CZET (E \<circ> E')"
 | 
|  |    130 | | "(CEQIL E t') \<circ> E' = CEQIL (E \<circ> E') t'"
 | 
|  |    131 | | "(CEQIR t' E) \<circ> E' = CEQIR t' (E \<circ> E')"
 | 
|  |    132 | 
 | 
|  |    133 | lemma ctx_compose:
 | 
|  |    134 |   shows "(E1 \<circ> E2)\<lbrakk>t\<rbrakk> = E1\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>"
 | 
|  |    135 | by (induct E1 rule: ctx.induct) (auto)
 | 
|  |    136 | 
 | 
| 27247 |    137 | text {* Composing a list (stack) of contexts. *}
 | 
|  |    138 | 
 | 
| 27162 |    139 | fun
 | 
|  |    140 |   ctx_composes :: "ctx list \<Rightarrow> ctx" ("_\<down>")
 | 
|  |    141 | where
 | 
|  |    142 |     "[]\<down> = \<box>"
 | 
|  |    143 |   | "(E#Es)\<down> = (Es\<down>) \<circ> E"
 | 
|  |    144 | 
 | 
| 27247 |    145 | section {* The CK-Machine *}
 | 
| 27162 |    146 | 
 | 
|  |    147 | inductive
 | 
|  |    148 |   val :: "lam\<Rightarrow>bool" 
 | 
|  |    149 | where
 | 
|  |    150 |   v_LAM[intro]:   "val (LAM [x].e)"
 | 
|  |    151 | | v_NUM[intro]:   "val (NUM n)"  
 | 
|  |    152 | | v_FALSE[intro]: "val FALSE"
 | 
|  |    153 | | v_TRUE[intro]:  "val TRUE"
 | 
|  |    154 | 
 | 
|  |    155 | equivariance val 
 | 
|  |    156 | 
 | 
|  |    157 | inductive
 | 
|  |    158 |   machine :: "lam\<Rightarrow>ctx list\<Rightarrow>lam\<Rightarrow>ctx list\<Rightarrow>bool" ("<_,_> \<mapsto> <_,_>")
 | 
|  |    159 | where
 | 
|  |    160 |   m1[intro]: "<APP e1 e2,Es> \<mapsto> <e1,(CAPPL \<box> e2)#Es>"
 | 
|  |    161 | | m2[intro]: "val v \<Longrightarrow> <v,(CAPPL \<box> e2)#Es> \<mapsto> <e2,(CAPPR v \<box>)#Es>"
 | 
|  |    162 | | m3[intro]: "val v \<Longrightarrow> <v,(CAPPR (LAM [y].e) \<box>)#Es> \<mapsto> <e[y::=v],Es>"
 | 
|  |    163 | | m4[intro]: "<e1 -- e2, Es> \<mapsto> <e1,(CDIFFL \<box> e2)#Es>"
 | 
|  |    164 | | m5[intro]: "<NUM n1,(CDIFFL \<box> e2)#Es> \<mapsto> <e2,(CDIFFR (NUM n1) \<box>)#Es>"
 | 
|  |    165 | | m6[intro]: "<NUM n2,(CDIFFR (NUM n1) \<box>)#Es> \<mapsto> <NUM (n1 - n2),Es>"
 | 
|  |    166 | | m4'[intro]:"<e1 ++ e2, Es> \<mapsto> <e1,(CPLUSL \<box> e2)#Es>"
 | 
|  |    167 | | m5'[intro]:"<NUM n1,(CPLUSL \<box> e2)#Es> \<mapsto> <e2,(CPLUSR (NUM n1) \<box>)#Es>"
 | 
|  |    168 | | m6'[intro]:"<NUM n2,(CPLUSR (NUM n1) \<box>)#Es> \<mapsto> <NUM (n1+n2),Es>"
 | 
|  |    169 | | m7[intro]: "<IF e1 e2 e3,Es> \<mapsto> <e1,(CIF \<box> e2 e3)#Es>"
 | 
|  |    170 | | m8[intro]: "<TRUE,(CIF \<box> e1 e2)#Es> \<mapsto> <e1,Es>"
 | 
|  |    171 | | m9[intro]: "<FALSE,(CIF \<box> e1 e2)#Es> \<mapsto> <e2,Es>"
 | 
|  |    172 | | mA[intro]: "<FIX [x].t,Es> \<mapsto> <t[x::=FIX [x].t],Es>"
 | 
|  |    173 | | mB[intro]: "<ZET e,Es> \<mapsto> <e,(CZET \<box>)#Es>"
 | 
|  |    174 | | mC[intro]: "<NUM 0,(CZET \<box>)#Es> \<mapsto> <TRUE,Es>"
 | 
|  |    175 | | mD[intro]: "0 < n \<Longrightarrow> <NUM n,(CZET \<box>)#Es> \<mapsto> <FALSE,Es>"
 | 
|  |    176 | | mE[intro]: "<EQI e1 e2,Es> \<mapsto> <e1,(CEQIL \<box> e2)#Es>"
 | 
|  |    177 | | mF[intro]: "<NUM n1,(CEQIL \<box> e2)#Es> \<mapsto> <e2,(CEQIR (NUM n1) \<box>)#Es>"
 | 
|  |    178 | | mG[intro]: "<NUM n,(CEQIR (NUM n) \<box>)#Es> \<mapsto> <TRUE,Es>"
 | 
|  |    179 | | mH[intro]: "n1\<noteq>n2 \<Longrightarrow> <NUM n1,(CEQIR (NUM n2) \<box>)#Es> \<mapsto> <FALSE,Es>"
 | 
|  |    180 | 
 | 
|  |    181 | inductive 
 | 
|  |    182 |   "machine_star" :: "lam\<Rightarrow>ctx list\<Rightarrow>lam\<Rightarrow>ctx list\<Rightarrow>bool" ("<_,_> \<mapsto>* <_,_>")
 | 
|  |    183 | where
 | 
|  |    184 |   ms1[intro]: "<e,Es> \<mapsto>* <e,Es>"
 | 
|  |    185 | | ms2[intro]: "\<lbrakk><e1,Es1> \<mapsto> <e2,Es2>; <e2,Es2> \<mapsto>* <e3,Es3>\<rbrakk> \<Longrightarrow> <e1,Es1> \<mapsto>* <e3,Es3>"
 | 
|  |    186 | 
 | 
|  |    187 | lemma ms3[intro,trans]:
 | 
|  |    188 |   assumes a: "<e1,Es1> \<mapsto>* <e2,Es2>" "<e2,Es2> \<mapsto>* <e3,Es3>"
 | 
|  |    189 |   shows "<e1,Es1> \<mapsto>* <e3,Es3>"
 | 
|  |    190 | using a by (induct) (auto) 
 | 
|  |    191 | 
 | 
|  |    192 | lemma ms4[intro]:
 | 
|  |    193 |   assumes a: "<e1,Es1> \<mapsto> <e2,Es2>" 
 | 
|  |    194 |   shows "<e1,Es1> \<mapsto>* <e2,Es2>"
 | 
|  |    195 | using a by (rule ms2) (rule ms1)
 | 
|  |    196 | 
 | 
| 27247 |    197 | section {* The Evaluation Relation (Big-Step Semantics) *}
 | 
| 27162 |    198 | 
 | 
|  |    199 | inductive
 | 
|  |    200 |   eval :: "lam\<Rightarrow>lam\<Rightarrow>bool" ("_ \<Down> _") 
 | 
|  |    201 | where
 | 
|  |    202 |   eval_NUM[intro]:  "NUM n \<Down> NUM n" 
 | 
|  |    203 | | eval_DIFF[intro]: "\<lbrakk>t1 \<Down> (NUM n1); t2 \<Down> (NUM n2)\<rbrakk> \<Longrightarrow> t1 -- t2 \<Down> NUM (n1 - n2)"
 | 
|  |    204 | | eval_PLUS[intro]: "\<lbrakk>t1 \<Down> (NUM n1); t2 \<Down> (NUM n2)\<rbrakk> \<Longrightarrow> t1 ++ t2 \<Down> NUM (n1 + n2)"
 | 
|  |    205 | | eval_LAM[intro]:  "LAM [x].t \<Down> LAM [x].t"
 | 
|  |    206 | | eval_APP[intro]:  "\<lbrakk>t1\<Down> LAM [x].t; t2\<Down> t2'; t[x::=t2']\<Down> t'\<rbrakk> \<Longrightarrow> APP t1 t2 \<Down> t'"
 | 
|  |    207 | | eval_FIX[intro]:  "t[x::= FIX [x].t] \<Down> t' \<Longrightarrow> FIX [x].t \<Down> t'"
 | 
|  |    208 | | eval_IF1[intro]:  "\<lbrakk>t1 \<Down> TRUE; t2 \<Down> t'\<rbrakk> \<Longrightarrow> IF t1 t2 t3 \<Down> t'"
 | 
|  |    209 | | eval_IF2[intro]:  "\<lbrakk>t1 \<Down> FALSE; t3 \<Down> t'\<rbrakk> \<Longrightarrow> IF t1 t2 t3 \<Down> t'"
 | 
|  |    210 | | eval_TRUE[intro]: "TRUE \<Down> TRUE"
 | 
|  |    211 | | eval_FALSE[intro]:"FALSE \<Down> FALSE"
 | 
|  |    212 | | eval_ZET1[intro]: "t \<Down> NUM 0 \<Longrightarrow> ZET t \<Down> TRUE"
 | 
|  |    213 | | eval_ZET2[intro]: "\<lbrakk>t \<Down> NUM n; 0 < n\<rbrakk> \<Longrightarrow> ZET t \<Down> FALSE"
 | 
|  |    214 | | eval_EQ1[intro]:  "\<lbrakk>t1 \<Down> NUM n; t2 \<Down> NUM n\<rbrakk> \<Longrightarrow> EQI t1 t2 \<Down> TRUE"
 | 
|  |    215 | | eval_EQ2[intro]:  "\<lbrakk>t1 \<Down> NUM n1; t2 \<Down> NUM n2; n1\<noteq>n2\<rbrakk> \<Longrightarrow> EQI t1 t2 \<Down> FALSE"
 | 
|  |    216 | 
 | 
|  |    217 | declare lam.inject[simp]
 | 
|  |    218 | inductive_cases eval_elim:
 | 
|  |    219 |   "APP t1 t2 \<Down> t'"
 | 
|  |    220 |   "IF t1 t2 t3 \<Down> t'"
 | 
|  |    221 |   "ZET t \<Down> t'"
 | 
|  |    222 |   "EQI t1 t2 \<Down> t'"
 | 
|  |    223 |   "t1 ++ t2 \<Down> t'"
 | 
|  |    224 |   "t1 -- t2 \<Down> t'"
 | 
|  |    225 |   "(NUM n) \<Down> t"
 | 
|  |    226 |   "TRUE \<Down> t"
 | 
|  |    227 |   "FALSE \<Down> t"
 | 
|  |    228 | declare lam.inject[simp del]
 | 
|  |    229 | 
 | 
|  |    230 | lemma eval_to:
 | 
|  |    231 |   assumes a: "t \<Down> t'"
 | 
|  |    232 |   shows "val t'"
 | 
|  |    233 | using a by (induct) (auto)
 | 
|  |    234 | 
 | 
|  |    235 | lemma eval_val:
 | 
|  |    236 |   assumes a: "val t"
 | 
|  |    237 |   shows "t \<Down> t"
 | 
|  |    238 | using a by (induct) (auto)
 | 
|  |    239 | 
 | 
| 27247 |    240 | text {* The Completeness Property: *}
 | 
|  |    241 | 
 | 
| 27162 |    242 | theorem eval_implies_machine_star_ctx:
 | 
|  |    243 |   assumes a: "t \<Down> t'"
 | 
|  |    244 |   shows "<t,Es> \<mapsto>* <t',Es>"
 | 
|  |    245 | using a
 | 
|  |    246 | by (induct arbitrary: Es)
 | 
|  |    247 |    (metis eval_to machine.intros ms1 ms2 ms3 ms4 v_LAM)+
 | 
|  |    248 | 
 | 
|  |    249 | corollary eval_implies_machine_star:
 | 
|  |    250 |   assumes a: "t \<Down> t'"
 | 
|  |    251 |   shows "<t,[]> \<mapsto>* <t',[]>"
 | 
|  |    252 | using a by (auto dest: eval_implies_machine_star_ctx)
 | 
|  |    253 | 
 | 
| 27247 |    254 | section {* The CBV Reduction Relation (Small-Step Semantics) *}
 | 
| 27162 |    255 | 
 | 
|  |    256 | lemma less_eqvt[eqvt]:
 | 
|  |    257 |   fixes pi::"name prm"
 | 
|  |    258 |   and   n1 n2::"nat"
 | 
|  |    259 |   shows "(pi\<bullet>(n1 < n2)) = ((pi\<bullet>n1) < (pi\<bullet>n2))"
 | 
|  |    260 | by (simp add: perm_nat_def perm_bool)
 | 
|  |    261 | 
 | 
|  |    262 | inductive
 | 
|  |    263 |   cbv :: "lam\<Rightarrow>lam\<Rightarrow>bool" ("_ \<longrightarrow>cbv _") 
 | 
|  |    264 | where
 | 
| 27247 |    265 |   cbv1: "\<lbrakk>val v; x\<sharp>v\<rbrakk> \<Longrightarrow> APP (LAM [x].t) v \<longrightarrow>cbv t[x::=v]"
 | 
| 27162 |    266 | | cbv2[intro]: "t \<longrightarrow>cbv t' \<Longrightarrow> APP t t2 \<longrightarrow>cbv APP t' t2"
 | 
|  |    267 | | cbv3[intro]: "t \<longrightarrow>cbv t' \<Longrightarrow> APP t2 t \<longrightarrow>cbv APP t2 t'"
 | 
|  |    268 | | cbv4[intro]: "t \<longrightarrow>cbv t' \<Longrightarrow> t -- t2 \<longrightarrow>cbv t' -- t2"
 | 
|  |    269 | | cbv5[intro]: "t \<longrightarrow>cbv t' \<Longrightarrow> t2 -- t \<longrightarrow>cbv t2 -- t'"
 | 
|  |    270 | | cbv6[intro]: "(NUM n1) -- (NUM n2) \<longrightarrow>cbv NUM (n1 - n2)"
 | 
|  |    271 | | cbv4'[intro]: "t \<longrightarrow>cbv t' \<Longrightarrow> t ++ t2 \<longrightarrow>cbv t' ++ t2"
 | 
|  |    272 | | cbv5'[intro]: "t \<longrightarrow>cbv t' \<Longrightarrow> t2 ++ t \<longrightarrow>cbv t2 ++ t'"
 | 
|  |    273 | | cbv6'[intro]:"(NUM n1) ++ (NUM n2) \<longrightarrow>cbv NUM (n1 + n2)"
 | 
|  |    274 | | cbv7[intro]: "t \<longrightarrow>cbv t' \<Longrightarrow> IF t t1 t2 \<longrightarrow>cbv IF t' t1 t2"
 | 
|  |    275 | | cbv8[intro]: "IF TRUE t1 t2 \<longrightarrow>cbv t1"
 | 
|  |    276 | | cbv9[intro]: "IF FALSE t1 t2 \<longrightarrow>cbv t2"
 | 
|  |    277 | | cbvA[intro]: "FIX [x].t \<longrightarrow>cbv t[x::=FIX [x].t]"
 | 
|  |    278 | | cbvB[intro]: "t \<longrightarrow>cbv t' \<Longrightarrow> ZET t \<longrightarrow>cbv ZET t'"
 | 
|  |    279 | | cbvC[intro]: "ZET (NUM 0) \<longrightarrow>cbv TRUE"
 | 
|  |    280 | | cbvD[intro]: "0 < n \<Longrightarrow> ZET (NUM n) \<longrightarrow>cbv FALSE"
 | 
|  |    281 | | cbvE[intro]: "t \<longrightarrow>cbv t' \<Longrightarrow> EQI t t2 \<longrightarrow>cbv EQI t' t2"
 | 
|  |    282 | | cbvF[intro]: "t \<longrightarrow>cbv t' \<Longrightarrow> EQI t2 t \<longrightarrow>cbv EQI t2 t'"
 | 
|  |    283 | | cbvG[intro]: "EQI (NUM n) (NUM n) \<longrightarrow>cbv TRUE"
 | 
|  |    284 | | cbvH[intro]: "n1\<noteq>n2 \<Longrightarrow> EQI (NUM n1) (NUM n2) \<longrightarrow>cbv FALSE"
 | 
|  |    285 | 
 | 
|  |    286 | equivariance cbv
 | 
|  |    287 | nominal_inductive cbv
 | 
|  |    288 |  by (simp_all add: abs_fresh fresh_fact)
 | 
|  |    289 | 
 | 
| 27247 |    290 | lemma better_cbv1[intro]: 
 | 
| 27162 |    291 |   assumes a: "val v" 
 | 
|  |    292 |   shows "APP (LAM [x].t) v \<longrightarrow>cbv t[x::=v]"
 | 
|  |    293 | proof -
 | 
|  |    294 |   obtain y::"name" where fs: "y\<sharp>(x,t,v)" by (rule exists_fresh, rule fin_supp, blast)
 | 
|  |    295 |   have "APP (LAM [x].t) v = APP (LAM [y].([(y,x)]\<bullet>t)) v" using fs
 | 
|  |    296 |     by (auto simp add: lam.inject alpha' fresh_prod fresh_atm)
 | 
| 27247 |    297 |   also have "\<dots> \<longrightarrow>cbv  ([(y,x)]\<bullet>t)[y::=v]" using fs a by (auto simp add: cbv.eqvt cbv1)
 | 
| 27162 |    298 |   also have "\<dots> = t[x::=v]" using fs by (simp add: subst_rename[symmetric])
 | 
|  |    299 |   finally show "APP (LAM [x].t) v \<longrightarrow>cbv t[x::=v]" by simp
 | 
|  |    300 | qed
 | 
|  |    301 | 
 | 
|  |    302 | inductive 
 | 
|  |    303 |   "cbv_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>cbv* _")
 | 
|  |    304 | where
 | 
|  |    305 |   cbvs1[intro]: "e \<longrightarrow>cbv* e"
 | 
|  |    306 | | cbvs2[intro]: "\<lbrakk>e1\<longrightarrow>cbv e2; e2 \<longrightarrow>cbv* e3\<rbrakk> \<Longrightarrow> e1 \<longrightarrow>cbv* e3"
 | 
|  |    307 | 
 | 
|  |    308 | lemma cbvs3[intro,trans]:
 | 
|  |    309 |   assumes a: "e1 \<longrightarrow>cbv* e2" "e2 \<longrightarrow>cbv* e3"
 | 
|  |    310 |   shows "e1 \<longrightarrow>cbv* e3"
 | 
|  |    311 | using a by (induct) (auto) 
 | 
|  |    312 | 
 | 
|  |    313 | lemma cbv_in_ctx:
 | 
|  |    314 |   assumes a: "t \<longrightarrow>cbv t'"
 | 
|  |    315 |   shows "E\<lbrakk>t\<rbrakk> \<longrightarrow>cbv E\<lbrakk>t'\<rbrakk>"
 | 
|  |    316 | using a by (induct E) (auto)
 | 
|  |    317 | 
 | 
|  |    318 | lemma machine_implies_cbv_star_ctx:
 | 
|  |    319 |   assumes a: "<e,Es> \<mapsto> <e',Es'>"
 | 
|  |    320 |   shows "(Es\<down>)\<lbrakk>e\<rbrakk> \<longrightarrow>cbv* (Es'\<down>)\<lbrakk>e'\<rbrakk>"
 | 
|  |    321 | using a by (induct) (auto simp add: ctx_compose intro: cbv_in_ctx)
 | 
|  |    322 | 
 | 
|  |    323 | lemma machine_star_implies_cbv_star_ctx:
 | 
|  |    324 |   assumes a: "<e,Es> \<mapsto>* <e',Es'>"
 | 
|  |    325 |   shows "(Es\<down>)\<lbrakk>e\<rbrakk> \<longrightarrow>cbv* (Es'\<down>)\<lbrakk>e'\<rbrakk>"
 | 
|  |    326 | using a 
 | 
|  |    327 | by (induct) (auto dest: machine_implies_cbv_star_ctx)
 | 
|  |    328 | 
 | 
|  |    329 | lemma machine_star_implies_cbv_star:
 | 
|  |    330 |   assumes a: "<e,[]> \<mapsto>* <e',[]>"
 | 
|  |    331 |   shows "e \<longrightarrow>cbv* e'"
 | 
|  |    332 | using a by (auto dest: machine_star_implies_cbv_star_ctx)
 | 
|  |    333 | 
 | 
|  |    334 | lemma cbv_eval:
 | 
|  |    335 |   assumes a: "t1 \<longrightarrow>cbv t2" "t2 \<Down> t3"
 | 
|  |    336 |   shows "t1 \<Down> t3"
 | 
|  |    337 | using a
 | 
|  |    338 | by (induct arbitrary: t3)
 | 
|  |    339 |    (auto elim!: eval_elim intro: eval_val)
 | 
|  |    340 | 
 | 
|  |    341 | lemma cbv_star_eval:
 | 
|  |    342 |   assumes a: "t1 \<longrightarrow>cbv* t2" "t2 \<Down> t3"
 | 
|  |    343 |   shows "t1 \<Down> t3"
 | 
|  |    344 | using a by (induct) (auto simp add: cbv_eval)
 | 
|  |    345 | 
 | 
|  |    346 | lemma cbv_star_implies_eval:
 | 
|  |    347 |   assumes a: "t \<longrightarrow>cbv* v" "val v"
 | 
|  |    348 |   shows "t \<Down> v"
 | 
|  |    349 | using a
 | 
|  |    350 | by (induct)
 | 
|  |    351 |    (auto simp add: eval_val cbv_star_eval dest: cbvs2)
 | 
|  |    352 | 
 | 
| 27247 |    353 | text {* The Soundness Property *}
 | 
| 27162 |    354 | 
 | 
|  |    355 | theorem machine_star_implies_eval:
 | 
|  |    356 |   assumes a: "<t1,[]> \<mapsto>* <t2,[]>" 
 | 
|  |    357 |   and     b: "val t2" 
 | 
|  |    358 |   shows "t1 \<Down> t2"
 | 
|  |    359 | proof -
 | 
|  |    360 |   from a have "t1 \<longrightarrow>cbv* t2" by (simp add: machine_star_implies_cbv_star)
 | 
|  |    361 |   then show "t1 \<Down> t2" using b by (simp add: cbv_star_implies_eval)
 | 
|  |    362 | qed
 | 
|  |    363 | 
 | 
|  |    364 | section {* Typing *}
 | 
|  |    365 | 
 | 
| 27247 |    366 | text {* Types *}
 | 
|  |    367 | 
 | 
| 27162 |    368 | nominal_datatype ty =
 | 
|  |    369 |   tVAR "string"
 | 
|  |    370 | | tBOOL 
 | 
|  |    371 | | tINT
 | 
|  |    372 | | tARR "ty" "ty" ("_ \<rightarrow> _")
 | 
|  |    373 | 
 | 
|  |    374 | declare ty.inject[simp]
 | 
|  |    375 | 
 | 
|  |    376 | lemma ty_fresh:
 | 
|  |    377 |   fixes x::"name"
 | 
|  |    378 |   and   T::"ty"
 | 
|  |    379 |   shows "x\<sharp>T"
 | 
|  |    380 | by (induct T rule: ty.induct)
 | 
|  |    381 |    (auto simp add: fresh_string)
 | 
|  |    382 | 
 | 
| 27247 |    383 | text {* Typing Contexts *}
 | 
|  |    384 | 
 | 
| 27162 |    385 | types tctx = "(name\<times>ty) list"
 | 
|  |    386 | 
 | 
| 27247 |    387 | text {* Sub-Typing Contexts *}
 | 
|  |    388 | 
 | 
| 27162 |    389 | abbreviation
 | 
|  |    390 |   "sub_tctx" :: "tctx \<Rightarrow> tctx \<Rightarrow> bool" ("_ \<subseteq> _") 
 | 
|  |    391 | where
 | 
|  |    392 |   "\<Gamma>\<^isub>1 \<subseteq> \<Gamma>\<^isub>2 \<equiv> \<forall>x. x \<in> set \<Gamma>\<^isub>1 \<longrightarrow> x \<in> set \<Gamma>\<^isub>2"
 | 
|  |    393 | 
 | 
| 27247 |    394 | text {* Valid Typing Contexts *}
 | 
|  |    395 | 
 | 
| 27162 |    396 | inductive
 | 
|  |    397 |   valid :: "tctx \<Rightarrow> bool"
 | 
|  |    398 | where
 | 
|  |    399 |   v1[intro]: "valid []"
 | 
|  |    400 | | v2[intro]: "\<lbrakk>valid \<Gamma>; x\<sharp>\<Gamma>\<rbrakk>\<Longrightarrow> valid ((x,T)#\<Gamma>)"
 | 
|  |    401 | 
 | 
|  |    402 | equivariance valid
 | 
|  |    403 | 
 | 
|  |    404 | lemma valid_elim[dest]:
 | 
|  |    405 |   assumes a: "valid ((x,T)#\<Gamma>)"
 | 
|  |    406 |   shows "x\<sharp>\<Gamma> \<and> valid \<Gamma>"
 | 
|  |    407 | using a by (cases) (auto)
 | 
|  |    408 | 
 | 
|  |    409 | lemma valid_insert:
 | 
|  |    410 |   assumes a: "valid (\<Delta>@[(x,T)]@\<Gamma>)"
 | 
|  |    411 |   shows "valid (\<Delta> @ \<Gamma>)" 
 | 
|  |    412 | using a
 | 
|  |    413 | by (induct \<Delta>)
 | 
|  |    414 |    (auto simp add:  fresh_list_append fresh_list_cons dest!: valid_elim)
 | 
|  |    415 | 
 | 
|  |    416 | lemma fresh_set: 
 | 
|  |    417 |   shows "y\<sharp>xs = (\<forall>x\<in>set xs. y\<sharp>x)"
 | 
|  |    418 | by (induct xs) (simp_all add: fresh_list_nil fresh_list_cons)
 | 
|  |    419 | 
 | 
|  |    420 | lemma context_unique:
 | 
|  |    421 |   assumes a1: "valid \<Gamma>"
 | 
|  |    422 |   and     a2: "(x,T) \<in> set \<Gamma>"
 | 
|  |    423 |   and     a3: "(x,U) \<in> set \<Gamma>"
 | 
|  |    424 |   shows "T = U" 
 | 
|  |    425 | using a1 a2 a3
 | 
|  |    426 | by (induct) (auto simp add: fresh_set fresh_prod fresh_atm)
 | 
|  |    427 | 
 | 
| 27247 |    428 | section {* The Typing Relation *}
 | 
| 27162 |    429 | 
 | 
|  |    430 | inductive
 | 
|  |    431 |   typing :: "tctx \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _") 
 | 
|  |    432 | where
 | 
|  |    433 |   t_VAR[intro]:  "\<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> VAR x : T"
 | 
|  |    434 | | t_APP[intro]:  "\<lbrakk>\<Gamma> \<turnstile> t\<^isub>1 : T\<^isub>1\<rightarrow>T\<^isub>2; \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> APP t\<^isub>1 t\<^isub>2 : T\<^isub>2"
 | 
|  |    435 | | t_LAM[intro]:  "\<lbrakk>x\<sharp>\<Gamma>; (x,T\<^isub>1)#\<Gamma> \<turnstile> t : T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> LAM [x].t : T\<^isub>1 \<rightarrow> T\<^isub>2"
 | 
|  |    436 | | t_NUM[intro]:  "\<Gamma> \<turnstile> (NUM n) : tINT"
 | 
|  |    437 | | t_DIFF[intro]: "\<lbrakk>\<Gamma> \<turnstile> t\<^isub>1 : tINT; \<Gamma> \<turnstile> t\<^isub>2 : tINT\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> t\<^isub>1 -- t\<^isub>2 : tINT"
 | 
|  |    438 | | t_PLUS[intro]: "\<lbrakk>\<Gamma> \<turnstile> t\<^isub>1 : tINT; \<Gamma> \<turnstile> t\<^isub>2 : tINT\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> t\<^isub>1 ++ t\<^isub>2 : tINT"
 | 
|  |    439 | | t_TRUE[intro]:  "\<Gamma> \<turnstile> TRUE : tBOOL"
 | 
|  |    440 | | t_FALSE[intro]: "\<Gamma> \<turnstile> FALSE : tBOOL"
 | 
|  |    441 | | t_IF[intro]:    "\<lbrakk>\<Gamma> \<turnstile> t1 : tBOOL; \<Gamma> \<turnstile> t2 : T; \<Gamma> \<turnstile> t3 : T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> IF t1 t2 t3 : T"
 | 
|  |    442 | | t_ZET[intro]:   "\<Gamma> \<turnstile> t : tINT \<Longrightarrow> \<Gamma> \<turnstile> ZET t : tBOOL"
 | 
|  |    443 | | t_EQI[intro]:   "\<lbrakk>\<Gamma> \<turnstile> t1 : tINT; \<Gamma> \<turnstile> t2 : tINT\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> EQI t1 t2 : tBOOL"
 | 
|  |    444 | | t_FIX[intro]:   "\<lbrakk>x\<sharp>\<Gamma>; (x,T)#\<Gamma> \<turnstile> t : T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> FIX [x].t : T"  
 | 
|  |    445 | 
 | 
|  |    446 | declare lam.inject[simp]
 | 
|  |    447 | inductive_cases typing_inversion[elim]:
 | 
|  |    448 |   "\<Gamma> \<turnstile> t\<^isub>1 -- t\<^isub>2 : T"
 | 
|  |    449 |   "\<Gamma> \<turnstile> t\<^isub>1 ++ t\<^isub>2 : T"
 | 
|  |    450 |   "\<Gamma> \<turnstile> IF t1 t2 t3 : T"
 | 
|  |    451 |   "\<Gamma> \<turnstile> ZET t : T"
 | 
|  |    452 |   "\<Gamma> \<turnstile> EQI t1 t2 : T"
 | 
|  |    453 |   "\<Gamma> \<turnstile> APP t1 t2 : T"
 | 
| 27247 |    454 |   "\<Gamma> \<turnstile> TRUE : T"
 | 
|  |    455 |   "\<Gamma> \<turnstile> FALSE : T"
 | 
|  |    456 |   "\<Gamma> \<turnstile> NUM n : T"
 | 
| 27162 |    457 | declare lam.inject[simp del]
 | 
|  |    458 | 
 | 
|  |    459 | equivariance typing
 | 
|  |    460 | nominal_inductive typing
 | 
|  |    461 |   by (simp_all add: abs_fresh ty_fresh)
 | 
|  |    462 | 
 | 
|  |    463 | lemma t_LAM_inversion[dest]:
 | 
|  |    464 |   assumes ty: "\<Gamma> \<turnstile> LAM [x].t : T" 
 | 
|  |    465 |   and     fc: "x\<sharp>\<Gamma>" 
 | 
|  |    466 |   shows "\<exists>T\<^isub>1 T\<^isub>2. T = T\<^isub>1 \<rightarrow> T\<^isub>2 \<and> (x,T\<^isub>1)#\<Gamma> \<turnstile> t : T\<^isub>2"
 | 
|  |    467 | using ty fc 
 | 
|  |    468 | by (cases rule: typing.strong_cases) 
 | 
|  |    469 |    (auto simp add: alpha lam.inject abs_fresh ty_fresh)
 | 
|  |    470 | 
 | 
|  |    471 | lemma t_FIX_inversion[dest]:
 | 
|  |    472 |   assumes ty: "\<Gamma> \<turnstile> FIX [x].t : T" 
 | 
|  |    473 |   and     fc: "x\<sharp>\<Gamma>" 
 | 
|  |    474 |   shows "(x,T)#\<Gamma> \<turnstile> t : T"
 | 
|  |    475 | using ty fc 
 | 
|  |    476 | by (cases rule: typing.strong_cases) 
 | 
|  |    477 |    (auto simp add: alpha lam.inject abs_fresh ty_fresh)
 | 
|  |    478 | 
 | 
| 27247 |    479 | section {* The Type-Preservation Property for the CBV Reduction Relation *}
 | 
| 27162 |    480 | 
 | 
|  |    481 | lemma weakening: 
 | 
|  |    482 |   fixes \<Gamma>1 \<Gamma>2::"tctx"
 | 
|  |    483 |   assumes a: "\<Gamma>1 \<turnstile> t : T" 
 | 
|  |    484 |   and     b: "valid \<Gamma>2" 
 | 
|  |    485 |   and     c: "\<Gamma>1 \<subseteq> \<Gamma>2"
 | 
|  |    486 |   shows "\<Gamma>2 \<turnstile> t : T"
 | 
|  |    487 | using a b c
 | 
|  |    488 | by (nominal_induct \<Gamma>1 t T avoiding: \<Gamma>2 rule: typing.strong_induct)
 | 
|  |    489 |    (auto | atomize)+
 | 
|  |    490 | 
 | 
|  |    491 | lemma type_substitution_aux:
 | 
|  |    492 |   assumes a: "(\<Delta>@[(x,T')]@\<Gamma>) \<turnstile> e : T"
 | 
|  |    493 |   and     b: "\<Gamma> \<turnstile> e' : T'"
 | 
|  |    494 |   shows "(\<Delta>@\<Gamma>) \<turnstile> e[x::=e'] : T" 
 | 
|  |    495 | using a b 
 | 
|  |    496 | proof (nominal_induct \<Gamma>'\<equiv>"\<Delta>@[(x,T')]@\<Gamma>" e T avoiding: x e' \<Delta> rule: typing.strong_induct)
 | 
|  |    497 |   case (t_VAR \<Gamma>' y T x e' \<Delta>)
 | 
|  |    498 |   then have a1: "valid (\<Delta>@[(x,T')]@\<Gamma>)" 
 | 
|  |    499 |        and  a2: "(y,T) \<in> set (\<Delta>@[(x,T')]@\<Gamma>)" 
 | 
|  |    500 |        and  a3: "\<Gamma> \<turnstile> e' : T'" by simp_all
 | 
|  |    501 |   from a1 have a4: "valid (\<Delta>@\<Gamma>)" by (rule valid_insert)
 | 
|  |    502 |   { assume eq: "x=y"
 | 
|  |    503 |     from a1 a2 have "T=T'" using eq by (auto intro: context_unique)
 | 
|  |    504 |     with a3 have "\<Delta>@\<Gamma> \<turnstile> VAR y[x::=e'] : T" using eq a4 by (auto intro: weakening)
 | 
|  |    505 |   }
 | 
|  |    506 |   moreover
 | 
|  |    507 |   { assume ineq: "x\<noteq>y"
 | 
|  |    508 |     from a2 have "(y,T) \<in> set (\<Delta>@\<Gamma>)" using ineq by simp
 | 
|  |    509 |     then have "\<Delta>@\<Gamma> \<turnstile> VAR y[x::=e'] : T" using ineq a4 by auto
 | 
|  |    510 |   }
 | 
|  |    511 |   ultimately show "\<Delta>@\<Gamma> \<turnstile> VAR y[x::=e'] : T" by blast
 | 
|  |    512 | qed (auto | force simp add: fresh_list_append fresh_list_cons)+
 | 
|  |    513 | 
 | 
|  |    514 | corollary type_substitution:
 | 
|  |    515 |   assumes a: "(x,T')#\<Gamma> \<turnstile> e : T"
 | 
|  |    516 |   and     b: "\<Gamma> \<turnstile> e' : T'"
 | 
|  |    517 |   shows "\<Gamma> \<turnstile> e[x::=e'] : T"
 | 
|  |    518 | using a b
 | 
|  |    519 | by (auto intro: type_substitution_aux[where \<Delta>="[]",simplified])
 | 
|  |    520 | 
 | 
|  |    521 | theorem cbv_type_preservation:
 | 
|  |    522 |   assumes a: "t \<longrightarrow>cbv t'"
 | 
|  |    523 |   and     b: "\<Gamma> \<turnstile> t : T" 
 | 
|  |    524 |   shows "\<Gamma> \<turnstile> t' : T"
 | 
|  |    525 | using a b
 | 
|  |    526 | apply(nominal_induct avoiding: \<Gamma> T rule: cbv.strong_induct)
 | 
|  |    527 | apply(auto elim!: typing_inversion dest: t_LAM_inversion simp add: type_substitution)
 | 
|  |    528 | apply(frule t_FIX_inversion)
 | 
|  |    529 | apply(auto simp add: type_substitution)
 | 
|  |    530 | done
 | 
|  |    531 | 
 | 
|  |    532 | corollary cbv_star_type_preservation:
 | 
|  |    533 |   assumes a: "t \<longrightarrow>cbv* t'"
 | 
|  |    534 |   and     b: "\<Gamma> \<turnstile> t : T" 
 | 
|  |    535 |   shows "\<Gamma> \<turnstile> t' : T"
 | 
|  |    536 | using a b
 | 
|  |    537 | by (induct) (auto intro: cbv_type_preservation)
 | 
|  |    538 | 
 | 
| 27247 |    539 | section {* The Type-Preservation Property for the Machine and Evaluation Relation *}
 | 
| 27162 |    540 | 
 | 
|  |    541 | theorem machine_type_preservation:
 | 
|  |    542 |   assumes a: "<t,[]> \<mapsto>* <t',[]>"
 | 
|  |    543 |   and     b: "\<Gamma> \<turnstile> t : T" 
 | 
|  |    544 |   shows "\<Gamma> \<turnstile> t' : T"
 | 
|  |    545 | proof -
 | 
|  |    546 |   from a have "t \<longrightarrow>cbv* t'" by (simp add: machine_star_implies_cbv_star)
 | 
|  |    547 |   then show "\<Gamma> \<turnstile> t' : T" using b by (simp add: cbv_star_type_preservation)
 | 
|  |    548 | qed
 | 
|  |    549 | 
 | 
|  |    550 | theorem eval_type_preservation:
 | 
|  |    551 |   assumes a: "t \<Down> t'"
 | 
|  |    552 |   and     b: "\<Gamma> \<turnstile> t : T" 
 | 
|  |    553 |   shows "\<Gamma> \<turnstile> t' : T"
 | 
|  |    554 | proof -
 | 
|  |    555 |   from a have "<t,[]> \<mapsto>* <t',[]>" by (simp add: eval_implies_machine_star)
 | 
|  |    556 |   then show "\<Gamma> \<turnstile> t' : T" using b by (simp add: machine_type_preservation)
 | 
|  |    557 | qed
 | 
|  |    558 | 
 | 
| 27247 |    559 | text {* The Progress Property *}
 | 
|  |    560 | 
 | 
|  |    561 | lemma canonical_tARR[dest]:
 | 
|  |    562 |   assumes a: "[] \<turnstile> t : T1 \<rightarrow> T2"
 | 
|  |    563 |   and     b: "val t"
 | 
|  |    564 |   shows "\<exists>x t'. t = LAM [x].t'"
 | 
|  |    565 | using b a by (induct) (auto) 
 | 
|  |    566 | 
 | 
|  |    567 | lemma canonical_tINT[dest]:
 | 
|  |    568 |   assumes a: "[] \<turnstile> t : tINT"
 | 
|  |    569 |   and     b: "val t"
 | 
|  |    570 |   shows "\<exists>n. t = NUM n"
 | 
|  |    571 | using b a 
 | 
|  |    572 | by (induct) (auto simp add: fresh_list_nil)
 | 
|  |    573 | 
 | 
|  |    574 | lemma canonical_tBOOL[dest]:
 | 
|  |    575 |   assumes a: "[] \<turnstile> t : tBOOL"
 | 
|  |    576 |   and     b: "val t"
 | 
|  |    577 |   shows "t = TRUE \<or> t = FALSE"
 | 
|  |    578 | using b a 
 | 
|  |    579 | by (induct) (auto simp add: fresh_list_nil)
 | 
|  |    580 | 
 | 
|  |    581 | theorem progress:
 | 
|  |    582 |   assumes a: "[] \<turnstile> t : T"
 | 
|  |    583 |   shows "(\<exists>t'. t \<longrightarrow>cbv t') \<or> (val t)"
 | 
|  |    584 | using a
 | 
|  |    585 | by (induct \<Gamma>\<equiv>"[]::tctx" t T)
 | 
|  |    586 |    (auto dest!: canonical_tINT intro!: cbv.intros gr0I)
 | 
|  |    587 | 
 | 
| 27162 |    588 | end    
 | 
|  |    589 |    
 | 
|  |    590 | 
 | 
|  |    591 | 
 | 
|  |    592 | 
 | 
|  |    593 |   
 | 
|  |    594 | 
 | 
|  |    595 |   
 | 
|  |    596 |   
 | 
|  |    597 | 
 | 
|  |    598 | 
 |