| author | urbanc | 
| Thu, 12 Jun 2008 09:56:28 +0200 | |
| changeset 27162 | 8d747de5c73e | 
| parent 26966 | 071f40487734 | 
| child 28568 | e1659c30f48d | 
| permissions | -rw-r--r-- | 
| 22447 | 1 | (* "$Id$" *) | 
| 25867 | 2 | (* *) | 
| 3 | (* Formalisation of some typical SOS-proofs. *) | |
| 4 | (* *) | |
| 5 | (* This work was inspired by challenge suggested by Adam *) | |
| 6 | (* Chlipala on the POPLmark mailing list. *) | |
| 7 | (* *) | |
| 8 | (* We thank Nick Benton for helping us with the *) | |
| 9 | (* termination-proof for evaluation. *) | |
| 10 | (* *) | |
| 11 | (* The formalisation was done by Julien Narboux and *) | |
| 12 | (* Christian Urban. *) | |
| 22447 | 13 | |
| 14 | theory SOS | |
| 22534 | 15 | imports "../Nominal" | 
| 22447 | 16 | begin | 
| 17 | ||
| 23158 
749b6870b1a1
introduced symmetric variants of the lemmas for alpha-equivalence
 urbanc parents: 
22730diff
changeset | 18 | atom_decl name | 
| 22447 | 19 | |
| 25832 | 20 | text {* types and terms *}
 | 
| 22447 | 21 | nominal_datatype ty = | 
| 25832 | 22 | TVar "nat" | 
| 22447 | 23 |   | Arrow "ty" "ty" ("_\<rightarrow>_" [100,100] 100)
 | 
| 24 | ||
| 25 | nominal_datatype trm = | |
| 26 | Var "name" | |
| 27 |   | Lam "\<guillemotleft>name\<guillemotright>trm" ("Lam [_]._" [100,100] 100)
 | |
| 28 | | App "trm" "trm" | |
| 29 | ||
| 25832 | 30 | lemma fresh_ty: | 
| 22447 | 31 | fixes x::"name" | 
| 32 | and T::"ty" | |
| 33 | shows "x\<sharp>T" | |
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26806diff
changeset | 34 | by (induct T rule: ty.induct) | 
| 25832 | 35 | (auto simp add: fresh_nat) | 
| 22447 | 36 | |
| 25832 | 37 | text {* Parallel and single substitution. *}
 | 
| 22447 | 38 | fun | 
| 39 | lookup :: "(name\<times>trm) list \<Rightarrow> name \<Rightarrow> trm" | |
| 40 | where | |
| 41 | "lookup [] x = Var x" | |
| 22502 | 42 | | "lookup ((y,e)#\<theta>) x = (if x=y then e else lookup \<theta> x)" | 
| 22447 | 43 | |
| 25832 | 44 | lemma lookup_eqvt[eqvt]: | 
| 22447 | 45 | fixes pi::"name prm" | 
| 46 | and \<theta>::"(name\<times>trm) list" | |
| 47 | and X::"name" | |
| 48 | shows "pi\<bullet>(lookup \<theta> X) = lookup (pi\<bullet>\<theta>) (pi\<bullet>X)" | |
| 25832 | 49 | by (induct \<theta>) (auto simp add: eqvts) | 
| 22447 | 50 | |
| 51 | lemma lookup_fresh: | |
| 52 | fixes z::"name" | |
| 25832 | 53 | assumes a: "z\<sharp>\<theta>" and b: "z\<sharp>x" | 
| 22447 | 54 | shows "z \<sharp>lookup \<theta> x" | 
| 25832 | 55 | using a b | 
| 22447 | 56 | by (induct rule: lookup.induct) (auto simp add: fresh_list_cons) | 
| 57 | ||
| 58 | lemma lookup_fresh': | |
| 59 | assumes "z\<sharp>\<theta>" | |
| 60 | shows "lookup \<theta> z = Var z" | |
| 61 | using assms | |
| 62 | by (induct rule: lookup.induct) | |
| 63 | (auto simp add: fresh_list_cons fresh_prod fresh_atm) | |
| 64 | ||
| 25832 | 65 | (* parallel substitution *) | 
| 22447 | 66 | consts | 
| 67 |   psubst :: "(name\<times>trm) list \<Rightarrow> trm \<Rightarrow> trm"  ("_<_>" [95,95] 105)
 | |
| 68 | ||
| 69 | nominal_primrec | |
| 70 | "\<theta><(Var x)> = (lookup \<theta> x)" | |
| 71 | "\<theta><(App e\<^isub>1 e\<^isub>2)> = App (\<theta><e\<^isub>1>) (\<theta><e\<^isub>2>)" | |
| 72 | "x\<sharp>\<theta> \<Longrightarrow> \<theta><(Lam [x].e)> = Lam [x].(\<theta><e>)" | |
| 25832 | 73 | apply(finite_guess)+ | 
| 22472 | 74 | apply(rule TrueI)+ | 
| 75 | apply(simp add: abs_fresh)+ | |
| 25832 | 76 | apply(fresh_guess)+ | 
| 22472 | 77 | done | 
| 22447 | 78 | |
| 79 | lemma psubst_eqvt[eqvt]: | |
| 80 | fixes pi::"name prm" | |
| 81 | and t::"trm" | |
| 82 | shows "pi\<bullet>(\<theta><t>) = (pi\<bullet>\<theta>)<(pi\<bullet>t)>" | |
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26806diff
changeset | 83 | by (nominal_induct t avoiding: \<theta> rule: trm.strong_induct) | 
| 22472 | 84 | (perm_simp add: fresh_bij lookup_eqvt)+ | 
| 22447 | 85 | |
| 86 | lemma fresh_psubst: | |
| 87 | fixes z::"name" | |
| 88 | and t::"trm" | |
| 89 | assumes "z\<sharp>t" and "z\<sharp>\<theta>" | |
| 90 | shows "z\<sharp>(\<theta><t>)" | |
| 91 | using assms | |
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26806diff
changeset | 92 | by (nominal_induct t avoiding: z \<theta> t rule: trm.strong_induct) | 
| 22447 | 93 | (auto simp add: abs_fresh lookup_fresh) | 
| 94 | ||
| 25832 | 95 | lemma psubst_empty[simp]: | 
| 96 | shows "[]<t> = t" | |
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26806diff
changeset | 97 | by (nominal_induct t rule: trm.strong_induct) | 
| 25832 | 98 | (auto simp add: fresh_list_nil) | 
| 99 | ||
| 25867 | 100 | text {* Single substitution *}
 | 
| 22447 | 101 | abbreviation | 
| 25832 | 102 |   subst :: "trm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> trm" ("_[_::=_]" [100,100,100] 100)
 | 
| 103 | where | |
| 104 | "t[x::=t'] \<equiv> ([(x,t')])<t>" | |
| 22447 | 105 | |
| 106 | lemma subst[simp]: | |
| 107 | shows "(Var x)[y::=t'] = (if x=y then t' else (Var x))" | |
| 108 | and "(App t\<^isub>1 t\<^isub>2)[y::=t'] = App (t\<^isub>1[y::=t']) (t\<^isub>2[y::=t'])" | |
| 109 | and "x\<sharp>(y,t') \<Longrightarrow> (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])" | |
| 22472 | 110 | by (simp_all add: fresh_list_cons fresh_list_nil) | 
| 22447 | 111 | |
| 112 | lemma fresh_subst: | |
| 113 | fixes z::"name" | |
| 114 | and t\<^isub>1::"trm" | |
| 115 | and t2::"trm" | |
| 25867 | 116 | assumes a: "z\<sharp>t\<^isub>1" "z\<sharp>t\<^isub>2" | 
| 22447 | 117 | shows "z\<sharp>t\<^isub>1[y::=t\<^isub>2]" | 
| 25867 | 118 | using a | 
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26806diff
changeset | 119 | by (nominal_induct t\<^isub>1 avoiding: z y t\<^isub>2 rule: trm.strong_induct) | 
| 22447 | 120 | (auto simp add: abs_fresh fresh_atm) | 
| 121 | ||
| 122 | lemma fresh_subst': | |
| 25832 | 123 | assumes "x\<sharp>e'" | 
| 124 | shows "x\<sharp>e[x::=e']" | |
| 125 | using assms | |
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26806diff
changeset | 126 | by (nominal_induct e avoiding: x e' rule: trm.strong_induct) | 
| 25832 | 127 | (auto simp add: fresh_atm abs_fresh fresh_nat) | 
| 22447 | 128 | |
| 129 | lemma forget: | |
| 25867 | 130 | assumes a: "x\<sharp>e" | 
| 25832 | 131 | shows "e[x::=e'] = e" | 
| 25867 | 132 | using a | 
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26806diff
changeset | 133 | by (nominal_induct e avoiding: x e' rule: trm.strong_induct) | 
| 25867 | 134 | (auto simp add: fresh_atm abs_fresh) | 
| 22447 | 135 | |
| 136 | lemma psubst_subst_psubst: | |
| 25867 | 137 | assumes h: "x\<sharp>\<theta>" | 
| 25832 | 138 | shows "\<theta><e>[x::=e'] = ((x,e')#\<theta>)<e>" | 
| 139 | using h | |
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26806diff
changeset | 140 | by (nominal_induct e avoiding: \<theta> x e' rule: trm.strong_induct) | 
| 25867 | 141 | (auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh') | 
| 22447 | 142 | |
| 25832 | 143 | text {* Typing Judgements *}
 | 
| 22447 | 144 | |
| 23760 | 145 | inductive | 
| 25832 | 146 | valid :: "(name\<times>ty) list \<Rightarrow> bool" | 
| 22447 | 147 | where | 
| 25832 | 148 | v_nil[intro]: "valid []" | 
| 149 | | v_cons[intro]: "\<lbrakk>valid \<Gamma>;x\<sharp>\<Gamma>\<rbrakk> \<Longrightarrow> valid ((x,T)#\<Gamma>)" | |
| 22447 | 150 | |
| 22534 | 151 | equivariance valid | 
| 22447 | 152 | |
| 25832 | 153 | inductive_cases | 
| 154 | valid_cons_inv_auto[elim]: "valid ((x,T)#\<Gamma>)" | |
| 22447 | 155 | |
| 156 | lemma type_unicity_in_context: | |
| 157 | assumes asm1: "(x,t\<^isub>2) \<in> set ((x,t\<^isub>1)#\<Gamma>)" | |
| 158 | and asm2: "valid ((x,t\<^isub>1)#\<Gamma>)" | |
| 159 | shows "t\<^isub>1=t\<^isub>2" | |
| 160 | proof - | |
| 161 | from asm2 have "x\<sharp>\<Gamma>" by (cases, auto) | |
| 162 | then have "(x,t\<^isub>2) \<notin> set \<Gamma>" | |
| 163 | by (induct \<Gamma>) (auto simp add: fresh_list_cons fresh_prod fresh_atm) | |
| 164 | then have "(x,t\<^isub>2) = (x,t\<^isub>1)" using asm1 by auto | |
| 165 | then show "t\<^isub>1 = t\<^isub>2" by auto | |
| 166 | qed | |
| 167 | ||
| 168 | lemma case_distinction_on_context: | |
| 25832 | 169 | fixes \<Gamma>::"(name\<times>ty) list" | 
| 22447 | 170 | assumes asm1: "valid ((m,t)#\<Gamma>)" | 
| 171 | and asm2: "(n,U) \<in> set ((m,T)#\<Gamma>)" | |
| 172 | shows "(n,U) = (m,T) \<or> ((n,U) \<in> set \<Gamma> \<and> n \<noteq> m)" | |
| 173 | proof - | |
| 25832 | 174 | from asm2 have "(n,U) \<in> set [(m,T)] \<or> (n,U) \<in> set \<Gamma>" by auto | 
| 175 | moreover | |
| 176 |   { assume eq: "m=n"
 | |
| 177 | assume "(n,U) \<in> set \<Gamma>" | |
| 178 | then have "\<not> n\<sharp>\<Gamma>" | |
| 179 | by (induct \<Gamma>) (auto simp add: fresh_list_cons fresh_prod fresh_atm) | |
| 180 | moreover have "m\<sharp>\<Gamma>" using asm1 by auto | |
| 181 | ultimately have False using eq by auto | |
| 182 | } | |
| 183 | ultimately show ?thesis by auto | |
| 22447 | 184 | qed | 
| 185 | ||
| 25867 | 186 | text {* Typing Relation *}
 | 
| 187 | ||
| 23760 | 188 | inductive | 
| 22447 | 189 |   typing :: "(name\<times>ty) list\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ : _" [60,60,60] 60) 
 | 
| 190 | where | |
| 25832 | 191 | t_Var[intro]: "\<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T" | 
| 192 | | t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> e\<^isub>1 : T\<^isub>1\<rightarrow>T\<^isub>2; \<Gamma> \<turnstile> e\<^isub>2 : T\<^isub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App e\<^isub>1 e\<^isub>2 : T\<^isub>2" | |
| 22447 | 193 | | t_Lam[intro]: "\<lbrakk>x\<sharp>\<Gamma>; (x,T\<^isub>1)#\<Gamma> \<turnstile> e : T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].e : T\<^isub>1\<rightarrow>T\<^isub>2" | 
| 194 | ||
| 22730 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22650diff
changeset | 195 | equivariance typing | 
| 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22650diff
changeset | 196 | |
| 22531 | 197 | nominal_inductive typing | 
| 25832 | 198 | by (simp_all add: abs_fresh fresh_ty) | 
| 22531 | 199 | |
| 22472 | 200 | lemma typing_implies_valid: | 
| 25832 | 201 | assumes a: "\<Gamma> \<turnstile> t : T" | 
| 22447 | 202 | shows "valid \<Gamma>" | 
| 25832 | 203 | using a | 
| 22447 | 204 | by (induct) (auto) | 
| 205 | ||
| 25832 | 206 | lemma t_Lam_elim: | 
| 25867 | 207 | assumes a: "\<Gamma> \<turnstile> Lam [x].t : T" "x\<sharp>\<Gamma>" | 
| 22447 | 208 | obtains T\<^isub>1 and T\<^isub>2 where "(x,T\<^isub>1)#\<Gamma> \<turnstile> t : T\<^isub>2" and "T=T\<^isub>1\<rightarrow>T\<^isub>2" | 
| 25867 | 209 | using a | 
| 25832 | 210 | by (cases rule: typing.strong_cases [where x="x"]) | 
| 211 | (auto simp add: abs_fresh fresh_ty alpha trm.inject) | |
| 22447 | 212 | |
| 25832 | 213 | abbreviation | 
| 214 |   "sub_context" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" ("_ \<subseteq> _" [55,55] 55)
 | |
| 215 | where | |
| 216 | "\<Gamma>\<^isub>1 \<subseteq> \<Gamma>\<^isub>2 \<equiv> \<forall>x T. (x,T)\<in>set \<Gamma>\<^isub>1 \<longrightarrow> (x,T)\<in>set \<Gamma>\<^isub>2" | |
| 22447 | 217 | |
| 218 | lemma weakening: | |
| 25832 | 219 | fixes \<Gamma>\<^isub>1 \<Gamma>\<^isub>2::"(name\<times>ty) list" | 
| 22650 
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
 urbanc parents: 
22594diff
changeset | 220 | assumes "\<Gamma>\<^isub>1 \<turnstile> e: T" and "valid \<Gamma>\<^isub>2" and "\<Gamma>\<^isub>1 \<subseteq> \<Gamma>\<^isub>2" | 
| 22447 | 221 | shows "\<Gamma>\<^isub>2 \<turnstile> e: T" | 
| 222 | using assms | |
| 22534 | 223 | proof(nominal_induct \<Gamma>\<^isub>1 e T avoiding: \<Gamma>\<^isub>2 rule: typing.strong_induct) | 
| 22447 | 224 | case (t_Lam x \<Gamma>\<^isub>1 T\<^isub>1 t T\<^isub>2 \<Gamma>\<^isub>2) | 
| 25832 | 225 | have vc: "x\<sharp>\<Gamma>\<^isub>2" by fact | 
| 22650 
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
 urbanc parents: 
22594diff
changeset | 226 | have ih: "\<lbrakk>valid ((x,T\<^isub>1)#\<Gamma>\<^isub>2); (x,T\<^isub>1)#\<Gamma>\<^isub>1 \<subseteq> (x,T\<^isub>1)#\<Gamma>\<^isub>2\<rbrakk> \<Longrightarrow> (x,T\<^isub>1)#\<Gamma>\<^isub>2 \<turnstile> t : T\<^isub>2" by fact | 
| 25832 | 227 | have "valid \<Gamma>\<^isub>2" by fact | 
| 228 | then have "valid ((x,T\<^isub>1)#\<Gamma>\<^isub>2)" using vc by auto | |
| 229 | moreover | |
| 230 | have "\<Gamma>\<^isub>1 \<subseteq> \<Gamma>\<^isub>2" by fact | |
| 231 | then have "(x,T\<^isub>1)#\<Gamma>\<^isub>1 \<subseteq> (x,T\<^isub>1)#\<Gamma>\<^isub>2" by simp | |
| 22447 | 232 | ultimately have "(x,T\<^isub>1)#\<Gamma>\<^isub>2 \<turnstile> t : T\<^isub>2" using ih by simp | 
| 25832 | 233 | with vc show "\<Gamma>\<^isub>2 \<turnstile> Lam [x].t : T\<^isub>1\<rightarrow>T\<^isub>2" by auto | 
| 22447 | 234 | qed (auto) | 
| 235 | ||
| 236 | lemma context_exchange: | |
| 237 | assumes a: "(x\<^isub>1,T\<^isub>1)#(x\<^isub>2,T\<^isub>2)#\<Gamma> \<turnstile> e : T" | |
| 238 | shows "(x\<^isub>2,T\<^isub>2)#(x\<^isub>1,T\<^isub>1)#\<Gamma> \<turnstile> e : T" | |
| 239 | proof - | |
| 22472 | 240 | from a have "valid ((x\<^isub>1,T\<^isub>1)#(x\<^isub>2,T\<^isub>2)#\<Gamma>)" by (simp add: typing_implies_valid) | 
| 22447 | 241 | then have "x\<^isub>1\<noteq>x\<^isub>2" "x\<^isub>1\<sharp>\<Gamma>" "x\<^isub>2\<sharp>\<Gamma>" "valid \<Gamma>" | 
| 242 | by (auto simp: fresh_list_cons fresh_atm[symmetric]) | |
| 243 | then have "valid ((x\<^isub>2,T\<^isub>2)#(x\<^isub>1,T\<^isub>1)#\<Gamma>)" | |
| 25832 | 244 | by (auto simp: fresh_list_cons fresh_atm fresh_ty) | 
| 22447 | 245 | moreover | 
| 22650 
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
 urbanc parents: 
22594diff
changeset | 246 | have "(x\<^isub>1,T\<^isub>1)#(x\<^isub>2,T\<^isub>2)#\<Gamma> \<subseteq> (x\<^isub>2,T\<^isub>2)#(x\<^isub>1,T\<^isub>1)#\<Gamma>" by auto | 
| 22447 | 247 | ultimately show "(x\<^isub>2,T\<^isub>2)#(x\<^isub>1,T\<^isub>1)#\<Gamma> \<turnstile> e : T" using a by (auto intro: weakening) | 
| 248 | qed | |
| 249 | ||
| 250 | lemma typing_var_unicity: | |
| 25832 | 251 | assumes a: "(x,T\<^isub>1)#\<Gamma> \<turnstile> Var x : T\<^isub>2" | 
| 252 | shows "T\<^isub>1=T\<^isub>2" | |
| 22447 | 253 | proof - | 
| 25832 | 254 | have "(x,T\<^isub>2) \<in> set ((x,T\<^isub>1)#\<Gamma>)" using a | 
| 255 | by (cases) (auto simp add: trm.inject) | |
| 256 | moreover | |
| 257 | have "valid ((x,T\<^isub>1)#\<Gamma>)" using a by (simp add: typing_implies_valid) | |
| 258 | ultimately show "T\<^isub>1=T\<^isub>2" by (simp only: type_unicity_in_context) | |
| 22447 | 259 | qed | 
| 260 | ||
| 261 | lemma typing_substitution: | |
| 262 | fixes \<Gamma>::"(name \<times> ty) list" | |
| 263 | assumes "(x,T')#\<Gamma> \<turnstile> e : T" | |
| 264 | and "\<Gamma> \<turnstile> e': T'" | |
| 265 | shows "\<Gamma> \<turnstile> e[x::=e'] : T" | |
| 266 | using assms | |
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26806diff
changeset | 267 | proof (nominal_induct e avoiding: \<Gamma> e' x arbitrary: T rule: trm.strong_induct) | 
| 22447 | 268 | case (Var y \<Gamma> e' x T) | 
| 269 | have h1: "(x,T')#\<Gamma> \<turnstile> Var y : T" by fact | |
| 270 | have h2: "\<Gamma> \<turnstile> e' : T'" by fact | |
| 271 | show "\<Gamma> \<turnstile> (Var y)[x::=e'] : T" | |
| 272 | proof (cases "x=y") | |
| 273 | case True | |
| 274 | assume as: "x=y" | |
| 275 | then have "T=T'" using h1 typing_var_unicity by auto | |
| 276 | then show "\<Gamma> \<turnstile> (Var y)[x::=e'] : T" using as h2 by simp | |
| 277 | next | |
| 278 | case False | |
| 279 | assume as: "x\<noteq>y" | |
| 25832 | 280 | have "(y,T) \<in> set ((x,T')#\<Gamma>)" using h1 by (cases) (auto simp add: trm.inject) | 
| 22447 | 281 | then have "(y,T) \<in> set \<Gamma>" using as by auto | 
| 282 | moreover | |
| 22472 | 283 | have "valid \<Gamma>" using h2 by (simp only: typing_implies_valid) | 
| 22447 | 284 | ultimately show "\<Gamma> \<turnstile> (Var y)[x::=e'] : T" using as by (simp add: t_Var) | 
| 285 | qed | |
| 286 | next | |
| 287 | case (Lam y t \<Gamma> e' x T) | |
| 25832 | 288 | have vc: "y\<sharp>\<Gamma>" "y\<sharp>x" "y\<sharp>e'" by fact | 
| 22447 | 289 | have pr1: "\<Gamma> \<turnstile> e' : T'" by fact | 
| 290 | have pr2: "(x,T')#\<Gamma> \<turnstile> Lam [y].t : T" by fact | |
| 291 | then obtain T\<^isub>1 T\<^isub>2 where pr2': "(y,T\<^isub>1)#(x,T')#\<Gamma> \<turnstile> t : T\<^isub>2" and eq: "T = T\<^isub>1\<rightarrow>T\<^isub>2" | |
| 25832 | 292 | using vc by (auto elim: t_Lam_elim simp add: fresh_list_cons fresh_ty) | 
| 22447 | 293 | then have pr2'':"(x,T')#(y,T\<^isub>1)#\<Gamma> \<turnstile> t : T\<^isub>2" by (simp add: context_exchange) | 
| 294 | have ih: "\<lbrakk>(x,T')#(y,T\<^isub>1)#\<Gamma> \<turnstile> t : T\<^isub>2; (y,T\<^isub>1)#\<Gamma> \<turnstile> e' : T'\<rbrakk> \<Longrightarrow> (y,T\<^isub>1)#\<Gamma> \<turnstile> t[x::=e'] : T\<^isub>2" by fact | |
| 22472 | 295 | have "valid \<Gamma>" using pr1 by (simp add: typing_implies_valid) | 
| 22447 | 296 | then have "valid ((y,T\<^isub>1)#\<Gamma>)" using vc by auto | 
| 25832 | 297 | then have "(y,T\<^isub>1)#\<Gamma> \<turnstile> e' : T'" using pr1 by (simp add: weakening) | 
| 22447 | 298 | then have "(y,T\<^isub>1)#\<Gamma> \<turnstile> t[x::=e'] : T\<^isub>2" using ih pr2'' by simp | 
| 25832 | 299 | then have "\<Gamma> \<turnstile> Lam [y].(t[x::=e']) : T\<^isub>1\<rightarrow>T\<^isub>2" using vc by auto | 
| 300 | then show "\<Gamma> \<turnstile> (Lam [y].t)[x::=e'] : T" using vc eq by simp | |
| 22447 | 301 | next | 
| 25832 | 302 | case (App e1 e2 \<Gamma> e' x T) | 
| 303 | have "(x,T')#\<Gamma> \<turnstile> App e1 e2 : T" by fact | |
| 304 | then obtain Tn where a1: "(x,T')#\<Gamma> \<turnstile> e1 : Tn \<rightarrow> T" | |
| 305 | and a2: "(x,T')#\<Gamma> \<turnstile> e2 : Tn" | |
| 306 | by (cases) (auto simp add: trm.inject) | |
| 307 | have a3: "\<Gamma> \<turnstile> e' : T'" by fact | |
| 308 | have ih1: "\<lbrakk>(x,T')#\<Gamma> \<turnstile> e1 : Tn\<rightarrow>T; \<Gamma> \<turnstile> e' : T'\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> e1[x::=e'] : Tn\<rightarrow>T" by fact | |
| 309 | have ih2: "\<lbrakk>(x,T')#\<Gamma> \<turnstile> e2 : Tn; \<Gamma> \<turnstile> e' : T'\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> e2[x::=e'] : Tn" by fact | |
| 25867 | 310 | then show "\<Gamma> \<turnstile> (App e1 e2)[x::=e'] : T" using a1 a2 a3 ih1 ih2 by auto | 
| 25832 | 311 | qed | 
| 312 | ||
| 313 | text {* Values *}
 | |
| 314 | inductive | |
| 315 | val :: "trm\<Rightarrow>bool" | |
| 316 | where | |
| 317 | v_Lam[intro]: "val (Lam [x].e)" | |
| 318 | ||
| 319 | equivariance val | |
| 320 | ||
| 321 | lemma not_val_App[simp]: | |
| 322 | shows | |
| 323 | "\<not> val (App e\<^isub>1 e\<^isub>2)" | |
| 324 | "\<not> val (Var x)" | |
| 325 | by (auto elim: val.cases) | |
| 22447 | 326 | |
| 327 | text {* Big-Step Evaluation *}
 | |
| 328 | ||
| 23760 | 329 | inductive | 
| 22447 | 330 |   big :: "trm\<Rightarrow>trm\<Rightarrow>bool" ("_ \<Down> _" [80,80] 80) 
 | 
| 331 | where | |
| 332 | b_Lam[intro]: "Lam [x].e \<Down> Lam [x].e" | |
| 333 | | b_App[intro]: "\<lbrakk>x\<sharp>(e\<^isub>1,e\<^isub>2,e'); e\<^isub>1\<Down>Lam [x].e; e\<^isub>2\<Down>e\<^isub>2'; e[x::=e\<^isub>2']\<Down>e'\<rbrakk> \<Longrightarrow> App e\<^isub>1 e\<^isub>2 \<Down> e'" | |
| 334 | ||
| 22730 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22650diff
changeset | 335 | equivariance big | 
| 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22650diff
changeset | 336 | |
| 22447 | 337 | nominal_inductive big | 
| 25832 | 338 | by (simp_all add: abs_fresh) | 
| 22447 | 339 | |
| 25832 | 340 | lemma big_preserves_fresh: | 
| 341 | fixes x::"name" | |
| 342 | assumes a: "e \<Down> e'" "x\<sharp>e" | |
| 343 | shows "x\<sharp>e'" | |
| 344 | using a by (induct) (auto simp add: abs_fresh fresh_subst) | |
| 22447 | 345 | |
| 25832 | 346 | lemma b_App_elim: | 
| 347 | assumes a: "App e\<^isub>1 e\<^isub>2 \<Down> e'" "x\<sharp>(e\<^isub>1,e\<^isub>2,e')" | |
| 348 | obtains f\<^isub>1 and f\<^isub>2 where "e\<^isub>1 \<Down> Lam [x]. f\<^isub>1" "e\<^isub>2 \<Down> f\<^isub>2" "f\<^isub>1[x::=f\<^isub>2] \<Down> e'" | |
| 349 | using a | |
| 350 | by (cases rule: big.strong_cases[where x="x" and xa="x"]) | |
| 351 | (auto simp add: trm.inject) | |
| 22447 | 352 | |
| 353 | lemma subject_reduction: | |
| 25832 | 354 | assumes a: "e \<Down> e'" and b: "\<Gamma> \<turnstile> e : T" | 
| 22447 | 355 | shows "\<Gamma> \<turnstile> e' : T" | 
| 22472 | 356 | using a b | 
| 22534 | 357 | proof (nominal_induct avoiding: \<Gamma> arbitrary: T rule: big.strong_induct) | 
| 358 | case (b_App x e\<^isub>1 e\<^isub>2 e' e e\<^isub>2' \<Gamma> T) | |
| 22447 | 359 | have vc: "x\<sharp>\<Gamma>" by fact | 
| 360 | have "\<Gamma> \<turnstile> App e\<^isub>1 e\<^isub>2 : T" by fact | |
| 25832 | 361 | then obtain T' where a1: "\<Gamma> \<turnstile> e\<^isub>1 : T'\<rightarrow>T" and a2: "\<Gamma> \<turnstile> e\<^isub>2 : T'" | 
| 362 | by (cases) (auto simp add: trm.inject) | |
| 22472 | 363 | have ih1: "\<Gamma> \<turnstile> e\<^isub>1 : T' \<rightarrow> T \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].e : T' \<rightarrow> T" by fact | 
| 22447 | 364 | have ih2: "\<Gamma> \<turnstile> e\<^isub>2 : T' \<Longrightarrow> \<Gamma> \<turnstile> e\<^isub>2' : T'" by fact | 
| 365 | have ih3: "\<Gamma> \<turnstile> e[x::=e\<^isub>2'] : T \<Longrightarrow> \<Gamma> \<turnstile> e' : T" by fact | |
| 366 | have "\<Gamma> \<turnstile> Lam [x].e : T'\<rightarrow>T" using ih1 a1 by simp | |
| 25832 | 367 | then have "((x,T')#\<Gamma>) \<turnstile> e : T" using vc | 
| 368 | by (auto elim: t_Lam_elim simp add: ty.inject) | |
| 22447 | 369 | moreover | 
| 370 | have "\<Gamma> \<turnstile> e\<^isub>2': T'" using ih2 a2 by simp | |
| 371 | ultimately have "\<Gamma> \<turnstile> e[x::=e\<^isub>2'] : T" by (simp add: typing_substitution) | |
| 372 | thus "\<Gamma> \<turnstile> e' : T" using ih3 by simp | |
| 373 | qed (blast)+ | |
| 374 | ||
| 22472 | 375 | lemma unicity_of_evaluation: | 
| 376 | assumes a: "e \<Down> e\<^isub>1" | |
| 377 | and b: "e \<Down> e\<^isub>2" | |
| 22447 | 378 | shows "e\<^isub>1 = e\<^isub>2" | 
| 22472 | 379 | using a b | 
| 22534 | 380 | proof (nominal_induct e e\<^isub>1 avoiding: e\<^isub>2 rule: big.strong_induct) | 
| 22447 | 381 | case (b_Lam x e t\<^isub>2) | 
| 382 | have "Lam [x].e \<Down> t\<^isub>2" by fact | |
| 383 | thus "Lam [x].e = t\<^isub>2" by (cases, simp_all add: trm.inject) | |
| 384 | next | |
| 22534 | 385 | case (b_App x e\<^isub>1 e\<^isub>2 e' e\<^isub>1' e\<^isub>2' t\<^isub>2) | 
| 22447 | 386 | have ih1: "\<And>t. e\<^isub>1 \<Down> t \<Longrightarrow> Lam [x].e\<^isub>1' = t" by fact | 
| 387 | have ih2:"\<And>t. e\<^isub>2 \<Down> t \<Longrightarrow> e\<^isub>2' = t" by fact | |
| 388 | have ih3: "\<And>t. e\<^isub>1'[x::=e\<^isub>2'] \<Down> t \<Longrightarrow> e' = t" by fact | |
| 22472 | 389 | have app: "App e\<^isub>1 e\<^isub>2 \<Down> t\<^isub>2" by fact | 
| 25832 | 390 | have vc: "x\<sharp>e\<^isub>1" "x\<sharp>e\<^isub>2" "x\<sharp>t\<^isub>2" by fact | 
| 391 | then have "x\<sharp>App e\<^isub>1 e\<^isub>2" by auto | |
| 392 | from app vc obtain f\<^isub>1 f\<^isub>2 where x1: "e\<^isub>1 \<Down> Lam [x]. f\<^isub>1" and x2: "e\<^isub>2 \<Down> f\<^isub>2" and x3: "f\<^isub>1[x::=f\<^isub>2] \<Down> t\<^isub>2" | |
| 393 | by (cases rule: big.strong_cases[where x="x" and xa="x"]) | |
| 394 | (auto simp add: trm.inject) | |
| 22472 | 395 | then have "Lam [x]. f\<^isub>1 = Lam [x]. e\<^isub>1'" using ih1 by simp | 
| 396 | then | |
| 397 | have "f\<^isub>1 = e\<^isub>1'" by (auto simp add: trm.inject alpha) | |
| 398 | moreover | |
| 399 | have "f\<^isub>2 = e\<^isub>2'" using x2 ih2 by simp | |
| 22447 | 400 | ultimately have "e\<^isub>1'[x::=e\<^isub>2'] \<Down> t\<^isub>2" using x3 by simp | 
| 25832 | 401 | thus "e' = t\<^isub>2" using ih3 by simp | 
| 402 | qed | |
| 22447 | 403 | |
| 22472 | 404 | lemma reduces_evaluates_to_values: | 
| 22447 | 405 | assumes h:"t \<Down> t'" | 
| 406 | shows "val t'" | |
| 22472 | 407 | using h by (induct) (auto) | 
| 22447 | 408 | |
| 25832 | 409 | (* Valuation *) | 
| 22447 | 410 | consts | 
| 411 | V :: "ty \<Rightarrow> trm set" | |
| 412 | ||
| 413 | nominal_primrec | |
| 25832 | 414 |   "V (TVar x) = {e. val e}"
 | 
| 22447 | 415 |   "V (T\<^isub>1 \<rightarrow> T\<^isub>2) = {Lam [x].e | x e. \<forall> v \<in> (V T\<^isub>1). \<exists> v'. e[x::=v] \<Down> v' \<and> v' \<in> V T\<^isub>2}"
 | 
| 25832 | 416 | by (rule TrueI)+ | 
| 22447 | 417 | |
| 25832 | 418 | (* can go with strong inversion rules *) | 
| 22472 | 419 | lemma V_eqvt: | 
| 420 | fixes pi::"name prm" | |
| 421 | assumes a: "x\<in>V T" | |
| 422 | shows "(pi\<bullet>x)\<in>V T" | |
| 423 | using a | |
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26806diff
changeset | 424 | apply(nominal_induct T arbitrary: pi x rule: ty.strong_induct) | 
| 26806 | 425 | apply(auto simp add: trm.inject) | 
| 25832 | 426 | apply(simp add: eqvts) | 
| 22472 | 427 | apply(rule_tac x="pi\<bullet>xa" in exI) | 
| 428 | apply(rule_tac x="pi\<bullet>e" in exI) | |
| 429 | apply(simp) | |
| 430 | apply(auto) | |
| 431 | apply(drule_tac x="(rev pi)\<bullet>v" in bspec) | |
| 432 | apply(force) | |
| 433 | apply(auto) | |
| 434 | apply(rule_tac x="pi\<bullet>v'" in exI) | |
| 435 | apply(auto) | |
| 22542 | 436 | apply(drule_tac pi="pi" in big.eqvt) | 
| 22541 
c33b542394f3
the name for the collection of equivariance lemmas is now eqvts (changed from eqvt) in order to avoid clashes with eqvt-lemmas generated in nominal_inductive
 urbanc parents: 
22534diff
changeset | 437 | apply(perm_simp add: eqvts) | 
| 22472 | 438 | done | 
| 439 | ||
| 25832 | 440 | lemma V_arrow_elim_weak: | 
| 25867 | 441 | assumes h:"u \<in> V (T\<^isub>1 \<rightarrow> T\<^isub>2)" | 
| 22447 | 442 | obtains a t where "u = Lam[a].t" and "\<forall> v \<in> (V T\<^isub>1). \<exists> v'. t[a::=v] \<Down> v' \<and> v' \<in> V T\<^isub>2" | 
| 443 | using h by (auto) | |
| 444 | ||
| 25832 | 445 | lemma V_arrow_elim_strong: | 
| 22447 | 446 | fixes c::"'a::fs_name" | 
| 22472 | 447 | assumes h: "u \<in> V (T\<^isub>1 \<rightarrow> T\<^isub>2)" | 
| 448 | obtains a t where "a\<sharp>c" "u = Lam[a].t" "\<forall>v \<in> (V T\<^isub>1). \<exists> v'. t[a::=v] \<Down> v' \<and> v' \<in> V T\<^isub>2" | |
| 22447 | 449 | using h | 
| 450 | apply - | |
| 451 | apply(erule V_arrow_elim_weak) | |
| 22472 | 452 | apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(a,t,c)") (*A*) | 
| 22447 | 453 | apply(erule exE) | 
| 454 | apply(drule_tac x="a'" in meta_spec) | |
| 22472 | 455 | apply(drule_tac x="[(a,a')]\<bullet>t" in meta_spec) | 
| 456 | apply(drule meta_mp) | |
| 22447 | 457 | apply(simp) | 
| 22472 | 458 | apply(drule meta_mp) | 
| 459 | apply(simp add: trm.inject alpha fresh_left fresh_prod calc_atm fresh_atm) | |
| 22447 | 460 | apply(perm_simp) | 
| 22472 | 461 | apply(force) | 
| 462 | apply(drule meta_mp) | |
| 463 | apply(rule ballI) | |
| 464 | apply(drule_tac x="[(a,a')]\<bullet>v" in bspec) | |
| 465 | apply(simp add: V_eqvt) | |
| 22447 | 466 | apply(auto) | 
| 22472 | 467 | apply(rule_tac x="[(a,a')]\<bullet>v'" in exI) | 
| 468 | apply(auto) | |
| 22542 | 469 | apply(drule_tac pi="[(a,a')]" in big.eqvt) | 
| 22541 
c33b542394f3
the name for the collection of equivariance lemmas is now eqvts (changed from eqvt) in order to avoid clashes with eqvt-lemmas generated in nominal_inductive
 urbanc parents: 
22534diff
changeset | 470 | apply(perm_simp add: eqvts calc_atm) | 
| 22472 | 471 | apply(simp add: V_eqvt) | 
| 472 | (*A*) | |
| 22447 | 473 | apply(rule exists_fresh') | 
| 22472 | 474 | apply(simp add: fin_supp) | 
| 22447 | 475 | done | 
| 476 | ||
| 25832 | 477 | lemma Vs_are_values: | 
| 478 | assumes a: "e \<in> V T" | |
| 22447 | 479 | shows "val e" | 
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26806diff
changeset | 480 | using a by (nominal_induct T arbitrary: e rule: ty.strong_induct) (auto) | 
| 22447 | 481 | |
| 482 | lemma values_reduce_to_themselves: | |
| 25832 | 483 | assumes a: "val v" | 
| 22447 | 484 | shows "v \<Down> v" | 
| 25832 | 485 | using a by (induct) (auto) | 
| 22447 | 486 | |
| 25832 | 487 | lemma Vs_reduce_to_themselves: | 
| 488 | assumes a: "v \<in> V T" | |
| 489 | shows "v \<Down> v" | |
| 490 | using a by (simp add: values_reduce_to_themselves Vs_are_values) | |
| 22447 | 491 | |
| 25832 | 492 | text {* '\<theta> maps x to e' asserts that \<theta> substitutes x with e *}
 | 
| 22447 | 493 | abbreviation | 
| 25832 | 494 |   mapsto :: "(name\<times>trm) list \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> bool" ("_ maps _ to _" [55,55,55] 55) 
 | 
| 22447 | 495 | where | 
| 25832 | 496 | "\<theta> maps x to e \<equiv> (lookup \<theta> x) = e" | 
| 22447 | 497 | |
| 498 | abbreviation | |
| 499 |   v_closes :: "(name\<times>trm) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" ("_ Vcloses _" [55,55] 55) 
 | |
| 500 | where | |
| 25832 | 501 | "\<theta> Vcloses \<Gamma> \<equiv> \<forall>x T. (x,T) \<in> set \<Gamma> \<longrightarrow> (\<exists>v. \<theta> maps x to v \<and> v \<in> V T)" | 
| 22447 | 502 | |
| 503 | lemma monotonicity: | |
| 504 | fixes m::"name" | |
| 505 | fixes \<theta>::"(name \<times> trm) list" | |
| 506 | assumes h1: "\<theta> Vcloses \<Gamma>" | |
| 507 | and h2: "e \<in> V T" | |
| 508 | and h3: "valid ((x,T)#\<Gamma>)" | |
| 509 | shows "(x,e)#\<theta> Vcloses (x,T)#\<Gamma>" | |
| 510 | proof(intro strip) | |
| 511 | fix x' T' | |
| 512 | assume "(x',T') \<in> set ((x,T)#\<Gamma>)" | |
| 513 | then have "((x',T')=(x,T)) \<or> ((x',T')\<in>set \<Gamma> \<and> x'\<noteq>x)" using h3 | |
| 514 | by (rule_tac case_distinction_on_context) | |
| 515 | moreover | |
| 516 |   { (* first case *)
 | |
| 517 | assume "(x',T') = (x,T)" | |
| 518 | then have "\<exists>e'. ((x,e)#\<theta>) maps x to e' \<and> e' \<in> V T'" using h2 by auto | |
| 519 | } | |
| 520 | moreover | |
| 521 |   { (* second case *)
 | |
| 522 | assume "(x',T') \<in> set \<Gamma>" and neq:"x' \<noteq> x" | |
| 523 | then have "\<exists>e'. \<theta> maps x' to e' \<and> e' \<in> V T'" using h1 by auto | |
| 524 | then have "\<exists>e'. ((x,e)#\<theta>) maps x' to e' \<and> e' \<in> V T'" using neq by auto | |
| 525 | } | |
| 526 | ultimately show "\<exists>e'. ((x,e)#\<theta>) maps x' to e' \<and> e' \<in> V T'" by blast | |
| 527 | qed | |
| 528 | ||
| 529 | lemma termination_aux: | |
| 530 | assumes h1: "\<Gamma> \<turnstile> e : T" | |
| 531 | and h2: "\<theta> Vcloses \<Gamma>" | |
| 532 | shows "\<exists>v. \<theta><e> \<Down> v \<and> v \<in> V T" | |
| 533 | using h2 h1 | |
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26806diff
changeset | 534 | proof(nominal_induct e avoiding: \<Gamma> \<theta> arbitrary: T rule: trm.strong_induct) | 
| 22447 | 535 | case (App e\<^isub>1 e\<^isub>2 \<Gamma> \<theta> T) | 
| 536 | have ih\<^isub>1:"\<And>\<theta> \<Gamma> T. \<lbrakk>\<theta> Vcloses \<Gamma>; \<Gamma> \<turnstile> e\<^isub>1 : T\<rbrakk> \<Longrightarrow> \<exists>v. \<theta><e\<^isub>1> \<Down> v \<and> v \<in> V T" by fact | |
| 537 | have ih\<^isub>2:"\<And>\<theta> \<Gamma> T. \<lbrakk>\<theta> Vcloses \<Gamma>; \<Gamma> \<turnstile> e\<^isub>2 : T\<rbrakk> \<Longrightarrow> \<exists>v. \<theta><e\<^isub>2> \<Down> v \<and> v \<in> V T" by fact | |
| 538 | have as\<^isub>1:"\<theta> Vcloses \<Gamma>" by fact | |
| 539 | have as\<^isub>2: "\<Gamma> \<turnstile> App e\<^isub>1 e\<^isub>2 : T" by fact | |
| 25832 | 540 | then obtain T' where "\<Gamma> \<turnstile> e\<^isub>1 : T' \<rightarrow> T" and "\<Gamma> \<turnstile> e\<^isub>2 : T'" | 
| 541 | by (cases) (auto simp add: trm.inject) | |
| 22447 | 542 | then obtain v\<^isub>1 v\<^isub>2 where "(i)": "\<theta><e\<^isub>1> \<Down> v\<^isub>1" "v\<^isub>1 \<in> V (T' \<rightarrow> T)" | 
| 543 | and "(ii)":"\<theta><e\<^isub>2> \<Down> v\<^isub>2" "v\<^isub>2 \<in> V T'" using ih\<^isub>1 ih\<^isub>2 as\<^isub>1 by blast | |
| 544 | from "(i)" obtain x e' | |
| 545 | where "v\<^isub>1 = Lam[x].e'" | |
| 546 | and "(iii)": "(\<forall>v \<in> (V T').\<exists> v'. e'[x::=v] \<Down> v' \<and> v' \<in> V T)" | |
| 547 | and "(iv)": "\<theta><e\<^isub>1> \<Down> Lam [x].e'" | |
| 25832 | 548 | and fr: "x\<sharp>(\<theta>,e\<^isub>1,e\<^isub>2)" by (blast elim: V_arrow_elim_strong) | 
| 22447 | 549 | from fr have fr\<^isub>1: "x\<sharp>\<theta><e\<^isub>1>" and fr\<^isub>2: "x\<sharp>\<theta><e\<^isub>2>" by (simp_all add: fresh_psubst) | 
| 550 | from "(ii)" "(iii)" obtain v\<^isub>3 where "(v)": "e'[x::=v\<^isub>2] \<Down> v\<^isub>3" "v\<^isub>3 \<in> V T" by auto | |
| 25832 | 551 | from fr\<^isub>2 "(ii)" have "x\<sharp>v\<^isub>2" by (simp add: big_preserves_fresh) | 
| 552 | then have "x\<sharp>e'[x::=v\<^isub>2]" by (simp add: fresh_subst') | |
| 553 | then have fr\<^isub>3: "x\<sharp>v\<^isub>3" using "(v)" by (simp add: big_preserves_fresh) | |
| 22447 | 554 | from fr\<^isub>1 fr\<^isub>2 fr\<^isub>3 have "x\<sharp>(\<theta><e\<^isub>1>,\<theta><e\<^isub>2>,v\<^isub>3)" by simp | 
| 555 | with "(iv)" "(ii)" "(v)" have "App (\<theta><e\<^isub>1>) (\<theta><e\<^isub>2>) \<Down> v\<^isub>3" by auto | |
| 556 | then show "\<exists>v. \<theta><App e\<^isub>1 e\<^isub>2> \<Down> v \<and> v \<in> V T" using "(v)" by auto | |
| 557 | next | |
| 558 | case (Lam x e \<Gamma> \<theta> T) | |
| 559 | have ih:"\<And>\<theta> \<Gamma> T. \<lbrakk>\<theta> Vcloses \<Gamma>; \<Gamma> \<turnstile> e : T\<rbrakk> \<Longrightarrow> \<exists>v. \<theta><e> \<Down> v \<and> v \<in> V T" by fact | |
| 560 | have as\<^isub>1: "\<theta> Vcloses \<Gamma>" by fact | |
| 561 | have as\<^isub>2: "\<Gamma> \<turnstile> Lam [x].e : T" by fact | |
| 25832 | 562 | have fs: "x\<sharp>\<Gamma>" "x\<sharp>\<theta>" by fact | 
| 22447 | 563 | from as\<^isub>2 fs obtain T\<^isub>1 T\<^isub>2 | 
| 25832 | 564 | where "(i)": "(x,T\<^isub>1)#\<Gamma> \<turnstile> e:T\<^isub>2" and "(ii)": "T = T\<^isub>1 \<rightarrow> T\<^isub>2" using fs | 
| 565 | by (cases rule: typing.strong_cases[where x="x"]) | |
| 566 | (auto simp add: trm.inject alpha abs_fresh fresh_ty) | |
| 22472 | 567 | from "(i)" have "(iii)": "valid ((x,T\<^isub>1)#\<Gamma>)" by (simp add: typing_implies_valid) | 
| 22447 | 568 | have "\<forall>v \<in> (V T\<^isub>1). \<exists>v'. (\<theta><e>)[x::=v] \<Down> v' \<and> v' \<in> V T\<^isub>2" | 
| 569 | proof | |
| 570 | fix v | |
| 571 | assume "v \<in> (V T\<^isub>1)" | |
| 572 | with "(iii)" as\<^isub>1 have "(x,v)#\<theta> Vcloses (x,T\<^isub>1)#\<Gamma>" using monotonicity by auto | |
| 573 | with ih "(i)" obtain v' where "((x,v)#\<theta>)<e> \<Down> v' \<and> v' \<in> V T\<^isub>2" by blast | |
| 22472 | 574 | then have "\<theta><e>[x::=v] \<Down> v' \<and> v' \<in> V T\<^isub>2" using fs by (simp add: psubst_subst_psubst) | 
| 22447 | 575 | then show "\<exists>v'. \<theta><e>[x::=v] \<Down> v' \<and> v' \<in> V T\<^isub>2" by auto | 
| 576 | qed | |
| 577 | then have "Lam[x].\<theta><e> \<in> V (T\<^isub>1 \<rightarrow> T\<^isub>2)" by auto | |
| 578 | then have "\<theta><Lam [x].e> \<Down> Lam[x].\<theta><e> \<and> Lam[x].\<theta><e> \<in> V (T\<^isub>1\<rightarrow>T\<^isub>2)" using fs by auto | |
| 579 | thus "\<exists>v. \<theta><Lam [x].e> \<Down> v \<and> v \<in> V T" using "(ii)" by auto | |
| 580 | next | |
| 25832 | 581 | case (Var x \<Gamma> \<theta> T) | 
| 582 | have "\<Gamma> \<turnstile> (Var x) : T" by fact | |
| 583 | then have "(x,T)\<in>set \<Gamma>" by (cases) (auto simp add: trm.inject) | |
| 584 | with prems have "\<theta><Var x> \<Down> \<theta><Var x> \<and> \<theta><Var x>\<in> V T" | |
| 585 | by (auto intro!: Vs_reduce_to_themselves) | |
| 586 | then show "\<exists>v. \<theta><Var x> \<Down> v \<and> v \<in> V T" by auto | |
| 587 | qed | |
| 22447 | 588 | |
| 589 | theorem termination_of_evaluation: | |
| 590 | assumes a: "[] \<turnstile> e : T" | |
| 591 | shows "\<exists>v. e \<Down> v \<and> val v" | |
| 592 | proof - | |
| 25832 | 593 | from a have "\<exists>v. []<e> \<Down> v \<and> v \<in> V T" by (rule termination_aux) (auto) | 
| 594 | thus "\<exists>v. e \<Down> v \<and> val v" using Vs_are_values by auto | |
| 22447 | 595 | qed | 
| 596 | ||
| 597 | end |