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