| author | wenzelm | 
| Sat, 09 Aug 2008 22:43:53 +0200 | |
| changeset 27811 | 44bc67675210 | 
| parent 27272 | 75b251e9cdb7 | 
| child 29097 | 68245155eb58 | 
| permissions | -rw-r--r-- | 
| 18263 
7f75925498da
cleaned up all examples so that they work with the
 urbanc parents: 
18106diff
changeset | 1 | (* $Id$ *) | 
| 18106 | 2 | |
| 19496 | 3 | theory SN | 
| 25867 | 4 | imports Lam_Funs | 
| 18106 | 5 | begin | 
| 6 | ||
| 18269 | 7 | text {* Strong Normalisation proof from the Proofs and Types book *}
 | 
| 18106 | 8 | |
| 9 | section {* Beta Reduction *}
 | |
| 10 | ||
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 11 | lemma subst_rename: | 
| 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 12 | assumes a: "c\<sharp>t1" | 
| 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 13 | shows "t1[a::=t2] = ([(c,a)]\<bullet>t1)[c::=t2]" | 
| 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 14 | using a | 
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26932diff
changeset | 15 | by (nominal_induct t1 avoiding: a c t2 rule: lam.strong_induct) | 
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 16 | (auto simp add: calc_atm fresh_atm abs_fresh) | 
| 18106 | 17 | |
| 18313 | 18 | lemma forget: | 
| 19 | assumes a: "a\<sharp>t1" | |
| 20 | shows "t1[a::=t2] = t1" | |
| 21 | using a | |
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26932diff
changeset | 22 | by (nominal_induct t1 avoiding: a t2 rule: lam.strong_induct) | 
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 23 | (auto simp add: abs_fresh fresh_atm) | 
| 18106 | 24 | |
| 18313 | 25 | lemma fresh_fact: | 
| 26 | fixes a::"name" | |
| 23970 | 27 | assumes a: "a\<sharp>t1" "a\<sharp>t2" | 
| 22540 | 28 | shows "a\<sharp>t1[b::=t2]" | 
| 23970 | 29 | using a | 
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26932diff
changeset | 30 | by (nominal_induct t1 avoiding: a b t2 rule: lam.strong_induct) | 
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 31 | (auto simp add: abs_fresh fresh_atm) | 
| 18106 | 32 | |
| 22540 | 33 | lemma fresh_fact': | 
| 34 | fixes a::"name" | |
| 35 | assumes a: "a\<sharp>t2" | |
| 36 | shows "a\<sharp>t1[a::=t2]" | |
| 37 | using a | |
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26932diff
changeset | 38 | by (nominal_induct t1 avoiding: a t2 rule: lam.strong_induct) | 
| 22540 | 39 | (auto simp add: abs_fresh fresh_atm) | 
| 40 | ||
| 18383 | 41 | lemma subst_lemma: | 
| 18313 | 42 | assumes a: "x\<noteq>y" | 
| 43 | and b: "x\<sharp>L" | |
| 44 | shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]" | |
| 45 | using a b | |
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26932diff
changeset | 46 | by (nominal_induct M avoiding: x y N L rule: lam.strong_induct) | 
| 18313 | 47 | (auto simp add: fresh_fact forget) | 
| 18106 | 48 | |
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 49 | lemma id_subs: | 
| 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 50 | shows "t[x::=Var x] = t" | 
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26932diff
changeset | 51 | by (nominal_induct t avoiding: x rule: lam.strong_induct) | 
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 52 | (simp_all add: fresh_atm) | 
| 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 53 | |
| 26772 | 54 | lemma lookup_fresh: | 
| 55 | fixes z::"name" | |
| 56 | assumes "z\<sharp>\<theta>" "z\<sharp>x" | |
| 57 | shows "z\<sharp> lookup \<theta> x" | |
| 58 | using assms | |
| 59 | by (induct rule: lookup.induct) (auto simp add: fresh_list_cons) | |
| 60 | ||
| 61 | lemma lookup_fresh': | |
| 62 | assumes "z\<sharp>\<theta>" | |
| 63 | shows "lookup \<theta> z = Var z" | |
| 64 | using assms | |
| 65 | by (induct rule: lookup.induct) | |
| 66 | (auto simp add: fresh_list_cons fresh_prod fresh_atm) | |
| 67 | ||
| 23142 | 68 | lemma psubst_subst: | 
| 69 | assumes h:"c\<sharp>\<theta>" | |
| 70 | shows "(\<theta><t>)[c::=s] = ((c,s)#\<theta>)<t>" | |
| 71 | using h | |
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26932diff
changeset | 72 | by (nominal_induct t avoiding: \<theta> c s rule: lam.strong_induct) | 
| 23142 | 73 | (auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh') | 
| 74 | ||
| 23760 | 75 | inductive | 
| 23142 | 76 |   Beta :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>\<beta> _" [80,80] 80)
 | 
| 22271 | 77 | where | 
| 23970 | 78 | b1[intro!]: "s1 \<longrightarrow>\<^isub>\<beta> s2 \<Longrightarrow> App s1 t \<longrightarrow>\<^isub>\<beta> App s2 t" | 
| 79 | | b2[intro!]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> App t s1 \<longrightarrow>\<^isub>\<beta> App t s2" | |
| 80 | | b3[intro!]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> Lam [a].s1 \<longrightarrow>\<^isub>\<beta> Lam [a].s2" | |
| 81 | | b4[intro!]: "a\<sharp>s2 \<Longrightarrow> App (Lam [a].s1) s2\<longrightarrow>\<^isub>\<beta> (s1[a::=s2])" | |
| 22538 
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
 urbanc parents: 
22531diff
changeset | 82 | |
| 22730 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22650diff
changeset | 83 | equivariance Beta | 
| 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22650diff
changeset | 84 | |
| 22538 
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
 urbanc parents: 
22531diff
changeset | 85 | nominal_inductive Beta | 
| 22540 | 86 | by (simp_all add: abs_fresh fresh_fact') | 
| 18106 | 87 | |
| 25831 | 88 | lemma beta_preserves_fresh: | 
| 89 | fixes a::"name" | |
| 18378 | 90 | assumes a: "t\<longrightarrow>\<^isub>\<beta> s" | 
| 25831 | 91 | shows "a\<sharp>t \<Longrightarrow> a\<sharp>s" | 
| 18378 | 92 | using a | 
| 25831 | 93 | apply(nominal_induct t s avoiding: a rule: Beta.strong_induct) | 
| 94 | apply(auto simp add: abs_fresh fresh_fact fresh_atm) | |
| 95 | done | |
| 18378 | 96 | |
| 23970 | 97 | lemma beta_abs: | 
| 25831 | 98 | assumes a: "Lam [a].t\<longrightarrow>\<^isub>\<beta> t'" | 
| 23970 | 99 | shows "\<exists>t''. t'=Lam [a].t'' \<and> t\<longrightarrow>\<^isub>\<beta> t''" | 
| 25831 | 100 | proof - | 
| 101 | have "a\<sharp>Lam [a].t" by (simp add: abs_fresh) | |
| 102 | with a have "a\<sharp>t'" by (simp add: beta_preserves_fresh) | |
| 103 | with a show ?thesis | |
| 104 | by (cases rule: Beta.strong_cases[where a="a" and aa="a"]) | |
| 105 | (auto simp add: lam.inject abs_fresh alpha) | |
| 106 | qed | |
| 18106 | 107 | |
| 18313 | 108 | lemma beta_subst: | 
| 18106 | 109 | assumes a: "M \<longrightarrow>\<^isub>\<beta> M'" | 
| 110 | shows "M[x::=N]\<longrightarrow>\<^isub>\<beta> M'[x::=N]" | |
| 111 | using a | |
| 23142 | 112 | by (nominal_induct M M' avoiding: x N rule: Beta.strong_induct) | 
| 113 | (auto simp add: fresh_atm subst_lemma fresh_fact) | |
| 18106 | 114 | |
| 18383 | 115 | section {* types *}
 | 
| 116 | ||
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 117 | nominal_datatype ty = | 
| 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 118 | TVar "nat" | 
| 18106 | 119 | | TArr "ty" "ty" (infix "\<rightarrow>" 200) | 
| 120 | ||
| 121 | lemma fresh_ty: | |
| 122 | fixes a ::"name" | |
| 123 | and \<tau> ::"ty" | |
| 124 | shows "a\<sharp>\<tau>" | |
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26932diff
changeset | 125 | by (nominal_induct \<tau> rule: ty.strong_induct) | 
| 25867 | 126 | (auto simp add: fresh_nat) | 
| 18106 | 127 | |
| 128 | (* valid contexts *) | |
| 129 | ||
| 23760 | 130 | inductive | 
| 23142 | 131 | valid :: "(name\<times>ty) list \<Rightarrow> bool" | 
| 22271 | 132 | where | 
| 133 | v1[intro]: "valid []" | |
| 134 | | v2[intro]: "\<lbrakk>valid \<Gamma>;a\<sharp>\<Gamma>\<rbrakk>\<Longrightarrow> valid ((a,\<sigma>)#\<Gamma>)" | |
| 18106 | 135 | |
| 22538 
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
 urbanc parents: 
22531diff
changeset | 136 | equivariance valid | 
| 18106 | 137 | |
| 138 | (* typing judgements *) | |
| 139 | ||
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 140 | lemma fresh_context: | 
| 18106 | 141 | fixes \<Gamma> :: "(name\<times>ty)list" | 
| 142 | and a :: "name" | |
| 18378 | 143 | assumes a: "a\<sharp>\<Gamma>" | 
| 144 | shows "\<not>(\<exists>\<tau>::ty. (a,\<tau>)\<in>set \<Gamma>)" | |
| 145 | using a | |
| 23970 | 146 | by (induct \<Gamma>) | 
| 147 | (auto simp add: fresh_prod fresh_list_cons fresh_atm) | |
| 18106 | 148 | |
| 23760 | 149 | inductive | 
| 23142 | 150 |   typing :: "(name\<times>ty) list\<Rightarrow>lam\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ : _" [60,60,60] 60)
 | 
| 22271 | 151 | where | 
| 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: 
22542diff
changeset | 152 | t1[intro]: "\<lbrakk>valid \<Gamma>; (a,\<tau>)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var a : \<tau>" | 
| 
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
 urbanc parents: 
22542diff
changeset | 153 | | t2[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : \<tau>\<rightarrow>\<sigma>; \<Gamma> \<turnstile> t2 : \<tau>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : \<sigma>" | 
| 22271 | 154 | | t3[intro]: "\<lbrakk>a\<sharp>\<Gamma>;((a,\<tau>)#\<Gamma>) \<turnstile> t : \<sigma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [a].t : \<tau>\<rightarrow>\<sigma>" | 
| 18106 | 155 | |
| 22730 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22650diff
changeset | 156 | equivariance typing | 
| 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22650diff
changeset | 157 | |
| 22538 
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
 urbanc parents: 
22531diff
changeset | 158 | nominal_inductive typing | 
| 
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
 urbanc parents: 
22531diff
changeset | 159 | by (simp_all add: abs_fresh fresh_ty) | 
| 18106 | 160 | |
| 25867 | 161 | subsection {* a fact about beta *}
 | 
| 18106 | 162 | |
| 163 | constdefs | |
| 164 | "NORMAL" :: "lam \<Rightarrow> bool" | |
| 165 | "NORMAL t \<equiv> \<not>(\<exists>t'. t\<longrightarrow>\<^isub>\<beta> t')" | |
| 166 | ||
| 18383 | 167 | lemma NORMAL_Var: | 
| 168 | shows "NORMAL (Var a)" | |
| 169 | proof - | |
| 170 |   { assume "\<exists>t'. (Var a) \<longrightarrow>\<^isub>\<beta> t'"
 | |
| 171 | then obtain t' where "(Var a) \<longrightarrow>\<^isub>\<beta> t'" by blast | |
| 25867 | 172 | hence False by (cases) (auto) | 
| 18383 | 173 | } | 
| 25867 | 174 | thus "NORMAL (Var a)" by (auto simp add: NORMAL_def) | 
| 18383 | 175 | qed | 
| 176 | ||
| 25867 | 177 | text {* Inductive version of Strong Normalisation *}
 | 
| 23970 | 178 | inductive | 
| 179 | SN :: "lam \<Rightarrow> bool" | |
| 180 | where | |
| 181 | SN_intro: "(\<And>t'. t \<longrightarrow>\<^isub>\<beta> t' \<Longrightarrow> SN t') \<Longrightarrow> SN t" | |
| 182 | ||
| 183 | lemma SN_preserved: | |
| 184 | assumes a: "SN t1" "t1\<longrightarrow>\<^isub>\<beta> t2" | |
| 185 | shows "SN t2" | |
| 186 | using a | |
| 187 | by (cases) (auto) | |
| 18106 | 188 | |
| 23970 | 189 | lemma double_SN_aux: | 
| 190 | assumes a: "SN a" | |
| 191 | and b: "SN b" | |
| 192 | and hyp: "\<And>x z. | |
| 24899 | 193 | \<lbrakk>\<And>y. x \<longrightarrow>\<^isub>\<beta> y \<Longrightarrow> SN y; \<And>y. x \<longrightarrow>\<^isub>\<beta> y \<Longrightarrow> P y z; | 
| 194 | \<And>u. z \<longrightarrow>\<^isub>\<beta> u \<Longrightarrow> SN u; \<And>u. z \<longrightarrow>\<^isub>\<beta> u \<Longrightarrow> P x u\<rbrakk> \<Longrightarrow> P x z" | |
| 23970 | 195 | shows "P a b" | 
| 196 | proof - | |
| 197 | from a | |
| 198 | have r: "\<And>b. SN b \<Longrightarrow> P a b" | |
| 199 | proof (induct a rule: SN.SN.induct) | |
| 200 | case (SN_intro x) | |
| 201 | note SNI' = SN_intro | |
| 202 | have "SN b" by fact | |
| 203 | thus ?case | |
| 204 | proof (induct b rule: SN.SN.induct) | |
| 205 | case (SN_intro y) | |
| 206 | show ?case | |
| 207 | apply (rule hyp) | |
| 208 | apply (erule SNI') | |
| 209 | apply (erule SNI') | |
| 210 | apply (rule SN.SN_intro) | |
| 211 | apply (erule SN_intro)+ | |
| 212 | done | |
| 213 | qed | |
| 214 | qed | |
| 215 | from b show ?thesis by (rule r) | |
| 216 | qed | |
| 18106 | 217 | |
| 23970 | 218 | lemma double_SN[consumes 2]: | 
| 219 | assumes a: "SN a" | |
| 220 | and b: "SN b" | |
| 221 | and c: "\<And>x z. \<lbrakk>\<And>y. x \<longrightarrow>\<^isub>\<beta> y \<Longrightarrow> P y z; \<And>u. z \<longrightarrow>\<^isub>\<beta> u \<Longrightarrow> P x u\<rbrakk> \<Longrightarrow> P x z" | |
| 222 | shows "P a b" | |
| 223 | using a b c | |
| 224 | apply(rule_tac double_SN_aux) | |
| 225 | apply(assumption)+ | |
| 226 | apply(blast) | |
| 18106 | 227 | done | 
| 228 | ||
| 229 | section {* Candidates *}
 | |
| 230 | ||
| 22440 
7e4f4f19002f
deleted function for defining candidates and used nominal_primrec instead
 urbanc parents: 
22420diff
changeset | 231 | consts | 
| 18106 | 232 | RED :: "ty \<Rightarrow> lam set" | 
| 22440 
7e4f4f19002f
deleted function for defining candidates and used nominal_primrec instead
 urbanc parents: 
22420diff
changeset | 233 | |
| 
7e4f4f19002f
deleted function for defining candidates and used nominal_primrec instead
 urbanc parents: 
22420diff
changeset | 234 | nominal_primrec | 
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 235 |   "RED (TVar X) = {t. SN(t)}"
 | 
| 22440 
7e4f4f19002f
deleted function for defining candidates and used nominal_primrec instead
 urbanc parents: 
22420diff
changeset | 236 |   "RED (\<tau>\<rightarrow>\<sigma>) =   {t. \<forall>u. (u\<in>RED \<tau> \<longrightarrow> (App t u)\<in>RED \<sigma>)}"
 | 
| 23970 | 237 | by (rule TrueI)+ | 
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 238 | |
| 25867 | 239 | text {* neutral terms *}
 | 
| 18106 | 240 | constdefs | 
| 241 | NEUT :: "lam \<Rightarrow> bool" | |
| 23970 | 242 | "NEUT t \<equiv> (\<exists>a. t = Var a) \<or> (\<exists>t1 t2. t = App t1 t2)" | 
| 18106 | 243 | |
| 244 | (* a slight hack to get the first element of applications *) | |
| 23970 | 245 | (* this is needed to get (SN t) from SN (App t s) *) | 
| 23760 | 246 | inductive | 
| 23142 | 247 |   FST :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<guillemotright> _" [80,80] 80)
 | 
| 22271 | 248 | where | 
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 249 | fst[intro!]: "(App t s) \<guillemotright> t" | 
| 18106 | 250 | |
| 24899 | 251 | consts | 
| 252 | fst_app_aux::"lam\<Rightarrow>lam option" | |
| 253 | ||
| 254 | nominal_primrec | |
| 255 | "fst_app_aux (Var a) = None" | |
| 256 | "fst_app_aux (App t1 t2) = Some t1" | |
| 257 | "fst_app_aux (Lam [x].t) = None" | |
| 258 | apply(finite_guess)+ | |
| 259 | apply(rule TrueI)+ | |
| 260 | apply(simp add: fresh_none) | |
| 261 | apply(fresh_guess)+ | |
| 262 | done | |
| 263 | ||
| 264 | definition | |
| 265 | fst_app_def[simp]: "fst_app t = the (fst_app_aux t)" | |
| 266 | ||
| 23970 | 267 | lemma SN_of_FST_of_App: | 
| 268 | assumes a: "SN (App t s)" | |
| 24899 | 269 | shows "SN (fst_app (App t s))" | 
| 23970 | 270 | using a | 
| 271 | proof - | |
| 272 | from a have "\<forall>z. (App t s \<guillemotright> z) \<longrightarrow> SN z" | |
| 273 | by (induct rule: SN.SN.induct) | |
| 274 | (blast elim: FST.cases intro: SN_intro) | |
| 24899 | 275 | then have "SN t" by blast | 
| 276 | then show "SN (fst_app (App t s))" by simp | |
| 23970 | 277 | qed | 
| 18106 | 278 | |
| 18383 | 279 | section {* Candidates *}
 | 
| 280 | ||
| 18106 | 281 | constdefs | 
| 18383 | 282 | "CR1" :: "ty \<Rightarrow> bool" | 
| 23970 | 283 | "CR1 \<tau> \<equiv> \<forall>t. (t\<in>RED \<tau> \<longrightarrow> SN t)" | 
| 18106 | 284 | |
| 18383 | 285 | "CR2" :: "ty \<Rightarrow> bool" | 
| 286 | "CR2 \<tau> \<equiv> \<forall>t t'. (t\<in>RED \<tau> \<and> t \<longrightarrow>\<^isub>\<beta> t') \<longrightarrow> t'\<in>RED \<tau>" | |
| 18106 | 287 | |
| 18383 | 288 | "CR3_RED" :: "lam \<Rightarrow> ty \<Rightarrow> bool" | 
| 289 | "CR3_RED t \<tau> \<equiv> \<forall>t'. t\<longrightarrow>\<^isub>\<beta> t' \<longrightarrow> t'\<in>RED \<tau>" | |
| 18106 | 290 | |
| 18383 | 291 | "CR3" :: "ty \<Rightarrow> bool" | 
| 292 | "CR3 \<tau> \<equiv> \<forall>t. (NEUT t \<and> CR3_RED t \<tau>) \<longrightarrow> t\<in>RED \<tau>" | |
| 18106 | 293 | |
| 18383 | 294 | "CR4" :: "ty \<Rightarrow> bool" | 
| 295 | "CR4 \<tau> \<equiv> \<forall>t. (NEUT t \<and> NORMAL t) \<longrightarrow>t\<in>RED \<tau>" | |
| 18106 | 296 | |
| 23970 | 297 | lemma CR3_implies_CR4: | 
| 298 | assumes a: "CR3 \<tau>" | |
| 299 | shows "CR4 \<tau>" | |
| 300 | using a by (auto simp add: CR3_def CR3_RED_def CR4_def NORMAL_def) | |
| 18106 | 301 | |
| 23970 | 302 | (* sub_induction in the arrow-type case for the next proof *) | 
| 303 | lemma sub_induction: | |
| 304 | assumes a: "SN(u)" | |
| 305 | and b: "u\<in>RED \<tau>" | |
| 306 | and c1: "NEUT t" | |
| 307 | and c2: "CR2 \<tau>" | |
| 308 | and c3: "CR3 \<sigma>" | |
| 309 | and c4: "CR3_RED t (\<tau>\<rightarrow>\<sigma>)" | |
| 310 | shows "(App t u)\<in>RED \<sigma>" | |
| 311 | using a b | |
| 312 | proof (induct) | |
| 313 | fix u | |
| 314 | assume as: "u\<in>RED \<tau>" | |
| 315 | assume ih: " \<And>u'. \<lbrakk>u \<longrightarrow>\<^isub>\<beta> u'; u' \<in> RED \<tau>\<rbrakk> \<Longrightarrow> App t u' \<in> RED \<sigma>" | |
| 316 | have "NEUT (App t u)" using c1 by (auto simp add: NEUT_def) | |
| 317 | moreover | |
| 318 | have "CR3_RED (App t u) \<sigma>" unfolding CR3_RED_def | |
| 319 | proof (intro strip) | |
| 320 | fix r | |
| 321 | assume red: "App t u \<longrightarrow>\<^isub>\<beta> r" | |
| 322 | moreover | |
| 323 |     { assume "\<exists>t'. t \<longrightarrow>\<^isub>\<beta> t' \<and> r = App t' u"
 | |
| 324 | then obtain t' where a1: "t \<longrightarrow>\<^isub>\<beta> t'" and a2: "r = App t' u" by blast | |
| 325 | have "t'\<in>RED (\<tau>\<rightarrow>\<sigma>)" using c4 a1 by (simp add: CR3_RED_def) | |
| 326 | then have "App t' u\<in>RED \<sigma>" using as by simp | |
| 327 | then have "r\<in>RED \<sigma>" using a2 by simp | |
| 328 | } | |
| 329 | moreover | |
| 330 |     { assume "\<exists>u'. u \<longrightarrow>\<^isub>\<beta> u' \<and> r = App t u'"
 | |
| 331 | then obtain u' where b1: "u \<longrightarrow>\<^isub>\<beta> u'" and b2: "r = App t u'" by blast | |
| 332 | have "u'\<in>RED \<tau>" using as b1 c2 by (auto simp add: CR2_def) | |
| 333 | with ih have "App t u' \<in> RED \<sigma>" using b1 by simp | |
| 334 | then have "r\<in>RED \<sigma>" using b2 by simp | |
| 335 | } | |
| 336 | moreover | |
| 337 |     { assume "\<exists>x t'. t = Lam [x].t'"
 | |
| 338 | then obtain x t' where "t = Lam [x].t'" by blast | |
| 339 | then have "NEUT (Lam [x].t')" using c1 by simp | |
| 340 | then have "False" by (simp add: NEUT_def) | |
| 341 | then have "r\<in>RED \<sigma>" by simp | |
| 342 | } | |
| 343 | ultimately show "r \<in> RED \<sigma>" by (cases) (auto simp add: lam.inject) | |
| 344 | qed | |
| 345 | ultimately show "App t u \<in> RED \<sigma>" using c3 by (simp add: CR3_def) | |
| 346 | qed | |
| 18106 | 347 | |
| 25867 | 348 | text {* properties of the candiadates *}
 | 
| 18383 | 349 | lemma RED_props: | 
| 350 | shows "CR1 \<tau>" and "CR2 \<tau>" and "CR3 \<tau>" | |
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26932diff
changeset | 351 | proof (nominal_induct \<tau> rule: ty.strong_induct) | 
| 18611 | 352 | case (TVar a) | 
| 353 |   { case 1 show "CR1 (TVar a)" by (simp add: CR1_def)
 | |
| 354 | next | |
| 23970 | 355 | case 2 show "CR2 (TVar a)" by (auto intro: SN_preserved simp add: CR2_def) | 
| 18611 | 356 | next | 
| 23970 | 357 | case 3 show "CR3 (TVar a)" by (auto intro: SN_intro simp add: CR3_def CR3_RED_def) | 
| 18611 | 358 | } | 
| 18599 
e01112713fdc
changed PRO_RED proof to conform with the new induction rules
 urbanc parents: 
18383diff
changeset | 359 | next | 
| 18611 | 360 | case (TArr \<tau>1 \<tau>2) | 
| 361 |   { case 1
 | |
| 362 | have ih_CR3_\<tau>1: "CR3 \<tau>1" by fact | |
| 363 | have ih_CR1_\<tau>2: "CR1 \<tau>2" by fact | |
| 25867 | 364 | have "\<And>t. t \<in> RED (\<tau>1 \<rightarrow> \<tau>2) \<Longrightarrow> SN t" | 
| 365 | proof - | |
| 18611 | 366 | fix t | 
| 25867 | 367 | assume "t \<in> RED (\<tau>1 \<rightarrow> \<tau>2)" | 
| 368 | then have a: "\<forall>u. u \<in> RED \<tau>1 \<longrightarrow> App t u \<in> RED \<tau>2" by simp | |
| 23970 | 369 | from ih_CR3_\<tau>1 have "CR4 \<tau>1" by (simp add: CR3_implies_CR4) | 
| 18611 | 370 | moreover | 
| 26932 | 371 | fix a have "NEUT (Var a)" by (force simp add: NEUT_def) | 
| 18611 | 372 | moreover | 
| 373 | have "NORMAL (Var a)" by (rule NORMAL_Var) | |
| 374 | ultimately have "(Var a)\<in> RED \<tau>1" by (simp add: CR4_def) | |
| 375 | with a have "App t (Var a) \<in> RED \<tau>2" by simp | |
| 376 | hence "SN (App t (Var a))" using ih_CR1_\<tau>2 by (simp add: CR1_def) | |
| 25867 | 377 | thus "SN t" by (auto dest: SN_of_FST_of_App) | 
| 18611 | 378 | qed | 
| 25867 | 379 | then show "CR1 (\<tau>1 \<rightarrow> \<tau>2)" unfolding CR1_def by simp | 
| 18611 | 380 | next | 
| 381 | case 2 | |
| 382 | have ih_CR2_\<tau>2: "CR2 \<tau>2" by fact | |
| 25867 | 383 | then show "CR2 (\<tau>1 \<rightarrow> \<tau>2)" unfolding CR2_def by auto | 
| 18611 | 384 | next | 
| 385 | case 3 | |
| 386 | have ih_CR1_\<tau>1: "CR1 \<tau>1" by fact | |
| 387 | have ih_CR2_\<tau>1: "CR2 \<tau>1" by fact | |
| 388 | have ih_CR3_\<tau>2: "CR3 \<tau>2" by fact | |
| 23970 | 389 | show "CR3 (\<tau>1 \<rightarrow> \<tau>2)" unfolding CR3_def | 
| 390 | proof (simp, intro strip) | |
| 18611 | 391 | fix t u | 
| 392 | assume a1: "u \<in> RED \<tau>1" | |
| 393 | assume a2: "NEUT t \<and> CR3_RED t (\<tau>1 \<rightarrow> \<tau>2)" | |
| 23970 | 394 | have "SN(u)" using a1 ih_CR1_\<tau>1 by (simp add: CR1_def) | 
| 395 | then show "(App t u)\<in>RED \<tau>2" using ih_CR2_\<tau>1 ih_CR3_\<tau>2 a1 a2 by (blast intro: sub_induction) | |
| 18611 | 396 | qed | 
| 397 | } | |
| 18383 | 398 | qed | 
| 23970 | 399 | |
| 25867 | 400 | text {* 
 | 
| 401 | the next lemma not as simple as on paper, probably because of | |
| 402 | the stronger double_SN induction | |
| 403 | *} | |
| 23970 | 404 | lemma abs_RED: | 
| 405 | assumes asm: "\<forall>s\<in>RED \<tau>. t[x::=s]\<in>RED \<sigma>" | |
| 406 | shows "Lam [x].t\<in>RED (\<tau>\<rightarrow>\<sigma>)" | |
| 18106 | 407 | proof - | 
| 23970 | 408 | have b1: "SN t" | 
| 409 | proof - | |
| 410 | have "Var x\<in>RED \<tau>" | |
| 411 | proof - | |
| 412 | have "CR4 \<tau>" by (simp add: RED_props CR3_implies_CR4) | |
| 413 | moreover | |
| 414 | have "NEUT (Var x)" by (auto simp add: NEUT_def) | |
| 415 | moreover | |
| 416 | have "NORMAL (Var x)" by (auto elim: Beta.cases simp add: NORMAL_def) | |
| 417 | ultimately show "Var x\<in>RED \<tau>" by (simp add: CR4_def) | |
| 418 | qed | |
| 419 | then have "t[x::=Var x]\<in>RED \<sigma>" using asm by simp | |
| 420 | then have "t\<in>RED \<sigma>" by (simp add: id_subs) | |
| 421 | moreover | |
| 422 | have "CR1 \<sigma>" by (simp add: RED_props) | |
| 423 | ultimately show "SN t" by (simp add: CR1_def) | |
| 424 | qed | |
| 425 | show "Lam [x].t\<in>RED (\<tau>\<rightarrow>\<sigma>)" | |
| 426 | proof (simp, intro strip) | |
| 427 | fix u | |
| 428 | assume b2: "u\<in>RED \<tau>" | |
| 429 | then have b3: "SN u" using RED_props by (auto simp add: CR1_def) | |
| 430 | show "App (Lam [x].t) u \<in> RED \<sigma>" using b1 b3 b2 asm | |
| 431 | proof(induct t u rule: double_SN) | |
| 432 | fix t u | |
| 433 | assume ih1: "\<And>t'. \<lbrakk>t \<longrightarrow>\<^isub>\<beta> t'; u\<in>RED \<tau>; \<forall>s\<in>RED \<tau>. t'[x::=s]\<in>RED \<sigma>\<rbrakk> \<Longrightarrow> App (Lam [x].t') u \<in> RED \<sigma>" | |
| 434 | assume ih2: "\<And>u'. \<lbrakk>u \<longrightarrow>\<^isub>\<beta> u'; u'\<in>RED \<tau>; \<forall>s\<in>RED \<tau>. t[x::=s]\<in>RED \<sigma>\<rbrakk> \<Longrightarrow> App (Lam [x].t) u' \<in> RED \<sigma>" | |
| 435 | assume as1: "u \<in> RED \<tau>" | |
| 436 | assume as2: "\<forall>s\<in>RED \<tau>. t[x::=s]\<in>RED \<sigma>" | |
| 437 | have "CR3_RED (App (Lam [x].t) u) \<sigma>" unfolding CR3_RED_def | |
| 438 | proof(intro strip) | |
| 439 | fix r | |
| 440 | assume red: "App (Lam [x].t) u \<longrightarrow>\<^isub>\<beta> r" | |
| 441 | moreover | |
| 442 | 	{ assume "\<exists>t'. t \<longrightarrow>\<^isub>\<beta> t' \<and> r = App (Lam [x].t') u"
 | |
| 443 | then obtain t' where a1: "t \<longrightarrow>\<^isub>\<beta> t'" and a2: "r = App (Lam [x].t') u" by blast | |
| 444 | have "App (Lam [x].t') u\<in>RED \<sigma>" using ih1 a1 as1 as2 | |
| 445 | apply(auto) | |
| 446 | apply(drule_tac x="t'" in meta_spec) | |
| 447 | apply(simp) | |
| 448 | apply(drule meta_mp) | |
| 25867 | 449 | prefer 2 | 
| 450 | apply(auto)[1] | |
| 451 | apply(rule ballI) | |
| 23970 | 452 | apply(drule_tac x="s" in bspec) | 
| 453 | apply(simp) | |
| 25867 | 454 | apply(subgoal_tac "CR2 \<sigma>")(*A*) | 
| 23970 | 455 | apply(unfold CR2_def)[1] | 
| 456 | apply(drule_tac x="t[x::=s]" in spec) | |
| 457 | apply(drule_tac x="t'[x::=s]" in spec) | |
| 458 | apply(simp add: beta_subst) | |
| 25867 | 459 | (*A*) | 
| 23970 | 460 | apply(simp add: RED_props) | 
| 461 | done | |
| 462 | then have "r\<in>RED \<sigma>" using a2 by simp | |
| 463 | } | |
| 464 | moreover | |
| 465 | 	{ assume "\<exists>u'. u \<longrightarrow>\<^isub>\<beta> u' \<and> r = App (Lam [x].t) u'"
 | |
| 466 | then obtain u' where b1: "u \<longrightarrow>\<^isub>\<beta> u'" and b2: "r = App (Lam [x].t) u'" by blast | |
| 467 | have "App (Lam [x].t) u'\<in>RED \<sigma>" using ih2 b1 as1 as2 | |
| 468 | apply(auto) | |
| 469 | apply(drule_tac x="u'" in meta_spec) | |
| 470 | apply(simp) | |
| 471 | apply(drule meta_mp) | |
| 472 | apply(subgoal_tac "CR2 \<tau>") | |
| 473 | apply(unfold CR2_def)[1] | |
| 474 | apply(drule_tac x="u" in spec) | |
| 475 | apply(drule_tac x="u'" in spec) | |
| 476 | apply(simp) | |
| 477 | apply(simp add: RED_props) | |
| 478 | apply(simp) | |
| 479 | done | |
| 480 | then have "r\<in>RED \<sigma>" using b2 by simp | |
| 481 | } | |
| 482 | moreover | |
| 483 | 	{ assume "r = t[x::=u]"
 | |
| 484 | then have "r\<in>RED \<sigma>" using as1 as2 by auto | |
| 485 | } | |
| 486 | ultimately show "r \<in> RED \<sigma>" | |
| 25867 | 487 | (* one wants to use the strong elimination principle; for this one | 
| 488 | has to know that x\<sharp>u *) | |
| 23970 | 489 | apply(cases) | 
| 490 | apply(auto simp add: lam.inject) | |
| 491 | apply(drule beta_abs) | |
| 25867 | 492 | apply(auto)[1] | 
| 23970 | 493 | apply(auto simp add: alpha subst_rename) | 
| 18106 | 494 | done | 
| 495 | qed | |
| 23970 | 496 | moreover | 
| 497 | have "NEUT (App (Lam [x].t) u)" unfolding NEUT_def by (auto) | |
| 498 | ultimately show "App (Lam [x].t) u \<in> RED \<sigma>" using RED_props by (simp add: CR3_def) | |
| 18106 | 499 | qed | 
| 500 | qed | |
| 23970 | 501 | qed | 
| 18106 | 502 | |
| 22420 | 503 | abbreviation | 
| 504 |  mapsto :: "(name\<times>lam) list \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> bool" ("_ maps _ to _" [55,55,55] 55) 
 | |
| 505 | where | |
| 25867 | 506 | "\<theta> maps x to e \<equiv> (lookup \<theta> x) = e" | 
| 22420 | 507 | |
| 508 | abbreviation | |
| 509 |   closes :: "(name\<times>lam) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" ("_ closes _" [55,55] 55) 
 | |
| 510 | where | |
| 511 | "\<theta> closes \<Gamma> \<equiv> \<forall>x T. ((x,T) \<in> set \<Gamma> \<longrightarrow> (\<exists>t. \<theta> maps x to t \<and> t \<in> RED T))" | |
| 21107 
e69c0e82955a
new file for defining functions in the lambda-calculus
 urbanc parents: 
19972diff
changeset | 512 | |
| 18106 | 513 | lemma all_RED: | 
| 22420 | 514 | assumes a: "\<Gamma> \<turnstile> t : \<tau>" | 
| 515 | and b: "\<theta> closes \<Gamma>" | |
| 516 | shows "\<theta><t> \<in> RED \<tau>" | |
| 18345 | 517 | using a b | 
| 23142 | 518 | proof(nominal_induct avoiding: \<theta> rule: typing.strong_induct) | 
| 519 | case (t3 a \<Gamma> \<sigma> t \<tau> \<theta>) --"lambda case" | |
| 520 | have ih: "\<And>\<theta>. \<theta> closes ((a,\<sigma>)#\<Gamma>) \<Longrightarrow> \<theta><t> \<in> RED \<tau>" by fact | |
| 521 | have \<theta>_cond: "\<theta> closes \<Gamma>" by fact | |
| 23393 | 522 | have fresh: "a\<sharp>\<Gamma>" "a\<sharp>\<theta>" by fact+ | 
| 24899 | 523 | from ih have "\<forall>s\<in>RED \<sigma>. ((a,s)#\<theta>)<t> \<in> RED \<tau>" using fresh \<theta>_cond fresh_context by simp | 
| 524 | then have "\<forall>s\<in>RED \<sigma>. \<theta><t>[a::=s] \<in> RED \<tau>" using fresh by (simp add: psubst_subst) | |
| 23970 | 525 | then have "Lam [a].(\<theta><t>) \<in> RED (\<sigma> \<rightarrow> \<tau>)" by (simp only: abs_RED) | 
| 23142 | 526 | then show "\<theta><(Lam [a].t)> \<in> RED (\<sigma> \<rightarrow> \<tau>)" using fresh by simp | 
| 527 | qed auto | |
| 18345 | 528 | |
| 23142 | 529 | section {* identity substitution generated from a context \<Gamma> *}
 | 
| 530 | fun | |
| 18382 | 531 | "id" :: "(name\<times>ty) list \<Rightarrow> (name\<times>lam) list" | 
| 23142 | 532 | where | 
| 18382 | 533 | "id [] = []" | 
| 23142 | 534 | | "id ((x,\<tau>)#\<Gamma>) = (x,Var x)#(id \<Gamma>)" | 
| 18382 | 535 | |
| 23142 | 536 | lemma id_maps: | 
| 537 | shows "(id \<Gamma>) maps a to (Var a)" | |
| 538 | by (induct \<Gamma>) (auto) | |
| 18382 | 539 | |
| 540 | lemma id_fresh: | |
| 541 | fixes a::"name" | |
| 542 | assumes a: "a\<sharp>\<Gamma>" | |
| 543 | shows "a\<sharp>(id \<Gamma>)" | |
| 544 | using a | |
| 23142 | 545 | by (induct \<Gamma>) | 
| 546 | (auto simp add: fresh_list_nil fresh_list_cons) | |
| 18382 | 547 | |
| 548 | lemma id_apply: | |
| 22420 | 549 | shows "(id \<Gamma>)<t> = t" | 
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26932diff
changeset | 550 | by (nominal_induct t avoiding: \<Gamma> rule: lam.strong_induct) | 
| 23970 | 551 | (auto simp add: id_maps id_fresh) | 
| 18382 | 552 | |
| 23142 | 553 | lemma id_closes: | 
| 22420 | 554 | shows "(id \<Gamma>) closes \<Gamma>" | 
| 18383 | 555 | apply(auto) | 
| 23142 | 556 | apply(simp add: id_maps) | 
| 22420 | 557 | apply(subgoal_tac "CR3 T") --"A" | 
| 23970 | 558 | apply(drule CR3_implies_CR4) | 
| 18382 | 559 | apply(simp add: CR4_def) | 
| 22420 | 560 | apply(drule_tac x="Var x" in spec) | 
| 18383 | 561 | apply(force simp add: NEUT_def NORMAL_Var) | 
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 562 | --"A" | 
| 18383 | 563 | apply(rule RED_props) | 
| 18382 | 564 | done | 
| 565 | ||
| 18383 | 566 | lemma typing_implies_RED: | 
| 23142 | 567 | assumes a: "\<Gamma> \<turnstile> t : \<tau>" | 
| 18383 | 568 | shows "t \<in> RED \<tau>" | 
| 569 | proof - | |
| 22420 | 570 | have "(id \<Gamma>)<t>\<in>RED \<tau>" | 
| 18383 | 571 | proof - | 
| 23142 | 572 | have "(id \<Gamma>) closes \<Gamma>" by (rule id_closes) | 
| 18383 | 573 | with a show ?thesis by (rule all_RED) | 
| 574 | qed | |
| 575 | thus"t \<in> RED \<tau>" by (simp add: id_apply) | |
| 576 | qed | |
| 577 | ||
| 578 | lemma typing_implies_SN: | |
| 23142 | 579 | assumes a: "\<Gamma> \<turnstile> t : \<tau>" | 
| 18383 | 580 | shows "SN(t)" | 
| 581 | proof - | |
| 582 | from a have "t \<in> RED \<tau>" by (rule typing_implies_RED) | |
| 583 | moreover | |
| 584 | have "CR1 \<tau>" by (rule RED_props) | |
| 585 | ultimately show "SN(t)" by (simp add: CR1_def) | |
| 586 | qed | |
| 18382 | 587 | |
| 588 | end |