| author | wenzelm | 
| Sun, 12 Aug 2007 18:53:03 +0200 | |
| changeset 24231 | 85fb973a8207 | 
| parent 23760 | aca2c7f80e2f | 
| child 25137 | 0835ac64834a | 
| permissions | -rw-r--r-- | 
| 18269 | 1  | 
(* $Id$ *)  | 
| 18105 | 2  | 
|
| 19496 | 3  | 
theory Weakening  | 
| 
22418
 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 
urbanc 
parents: 
22231 
diff
changeset
 | 
4  | 
imports "../Nominal"  | 
| 18105 | 5  | 
begin  | 
6  | 
||
| 
21052
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
7  | 
section {* Weakening Example for the Simply-Typed Lambda-Calculus *}
 | 
| 
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
8  | 
(*================================================================*)  | 
| 18105 | 9  | 
|
10  | 
atom_decl name  | 
|
11  | 
||
| 
21052
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
12  | 
nominal_datatype lam =  | 
| 
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
13  | 
Var "name"  | 
| 
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
14  | 
| App "lam" "lam"  | 
| 
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
15  | 
  | Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
 | 
| 18105 | 16  | 
|
| 18352 | 17  | 
nominal_datatype ty =  | 
18  | 
TVar "nat"  | 
|
| 18105 | 19  | 
| TArr "ty" "ty" (infix "\<rightarrow>" 200)  | 
20  | 
||
| 22533 | 21  | 
lemma ty_fresh:  | 
22  | 
fixes x::"name"  | 
|
23  | 
and T::"ty"  | 
|
24  | 
shows "x\<sharp>T"  | 
|
25  | 
by (nominal_induct T rule: ty.induct)  | 
|
26  | 
(auto simp add: fresh_nat)  | 
|
| 18105 | 27  | 
|
| 
21052
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
28  | 
text {* valid contexts *}
 | 
| 22533 | 29  | 
|
| 23760 | 30  | 
inductive  | 
| 18105 | 31  | 
valid :: "(name\<times>ty) list \<Rightarrow> bool"  | 
| 21364 | 32  | 
where  | 
33  | 
v1[intro]: "valid []"  | 
|
34  | 
| v2[intro]: "\<lbrakk>valid \<Gamma>;a\<sharp>\<Gamma>\<rbrakk>\<Longrightarrow> valid ((a,\<sigma>)#\<Gamma>)"  | 
|
| 18105 | 35  | 
|
| 22533 | 36  | 
equivariance valid  | 
| 
22231
 
f76f187c91f9
added an infrastructure that allows the user to declare lemmas to be equivariance lemmas; the intention is to use these lemmas in automated tools but also can be employed by the user
 
urbanc 
parents: 
21488 
diff
changeset
 | 
37  | 
|
| 
21052
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
38  | 
text{* typing judgements *}
 | 
| 23760 | 39  | 
inductive  | 
| 
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: 
22533 
diff
changeset
 | 
40  | 
  typing :: "(name\<times>ty) list\<Rightarrow>lam\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ : _" [60,60,60] 60) 
 | 
| 21364 | 41  | 
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: 
22533 
diff
changeset
 | 
42  | 
t_Var[intro]: "\<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"  | 
| 
 
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
 
urbanc 
parents: 
22533 
diff
changeset
 | 
43  | 
| t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : T1\<rightarrow>T2; \<Gamma> \<turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : T2"  | 
| 22533 | 44  | 
| t_Lam[intro]: "\<lbrakk>x\<sharp>\<Gamma>;((x,T1)#\<Gamma>) \<turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].t : T1\<rightarrow>T2"  | 
| 
21052
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
45  | 
|
| 
22730
 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 
berghofe 
parents: 
22650 
diff
changeset
 | 
46  | 
equivariance typing  | 
| 
 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 
berghofe 
parents: 
22650 
diff
changeset
 | 
47  | 
|
| 22533 | 48  | 
(* automatically deriving the strong induction principle *)  | 
49  | 
nominal_inductive typing  | 
|
50  | 
by (simp_all add: abs_fresh ty_fresh)  | 
|
| 18105 | 51  | 
|
| 
21052
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
52  | 
text {* definition of a subcontext *}
 | 
| 18105 | 53  | 
|
| 
21052
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
54  | 
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: 
22533 
diff
changeset
 | 
55  | 
  "sub_context" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" ("_ \<subseteq> _" [60,60] 60) 
 | 
| 22533 | 56  | 
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: 
22533 
diff
changeset
 | 
57  | 
"\<Gamma>1 \<subseteq> \<Gamma>2 \<equiv> \<forall>x T. (x,T)\<in>set \<Gamma>1 \<longrightarrow> (x,T)\<in>set \<Gamma>2"  | 
| 
21052
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
58  | 
|
| 22533 | 59  | 
text {* Now it comes: The Weakening Lemma *}
 | 
| 18105 | 60  | 
|
| 
18311
 
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
 
urbanc 
parents: 
18296 
diff
changeset
 | 
61  | 
lemma weakening_version1:  | 
| 
24231
 
85fb973a8207
added type constraints to resolve syntax ambiguities;
 
wenzelm 
parents: 
23760 
diff
changeset
 | 
62  | 
fixes \<Gamma>1 \<Gamma>2 :: "(name\<times>ty) list"  | 
| 
22418
 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 
urbanc 
parents: 
22231 
diff
changeset
 | 
63  | 
assumes a: "\<Gamma>1 \<turnstile> t : T"  | 
| 
18311
 
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
 
urbanc 
parents: 
18296 
diff
changeset
 | 
64  | 
and b: "valid \<Gamma>2"  | 
| 
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: 
22533 
diff
changeset
 | 
65  | 
and c: "\<Gamma>1 \<subseteq> \<Gamma>2"  | 
| 
22418
 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 
urbanc 
parents: 
22231 
diff
changeset
 | 
66  | 
shows "\<Gamma>2 \<turnstile> t : T"  | 
| 
18311
 
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
 
urbanc 
parents: 
18296 
diff
changeset
 | 
67  | 
using a b c  | 
| 22533 | 68  | 
by (nominal_induct \<Gamma>1 t T avoiding: \<Gamma>2 rule: typing.strong_induct)  | 
| 
22418
 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 
urbanc 
parents: 
22231 
diff
changeset
 | 
69  | 
(auto | atomize)+  | 
| 18105 | 70  | 
|
| 
18311
 
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
 
urbanc 
parents: 
18296 
diff
changeset
 | 
71  | 
lemma weakening_version2:  | 
| 18105 | 72  | 
fixes \<Gamma>1::"(name\<times>ty) list"  | 
73  | 
and t ::"lam"  | 
|
74  | 
and \<tau> ::"ty"  | 
|
| 
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: 
22533 
diff
changeset
 | 
75  | 
assumes a: "\<Gamma>1 \<turnstile> t : T"  | 
| 
18311
 
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
 
urbanc 
parents: 
18296 
diff
changeset
 | 
76  | 
and b: "valid \<Gamma>2"  | 
| 
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: 
22533 
diff
changeset
 | 
77  | 
and c: "\<Gamma>1 \<subseteq> \<Gamma>2"  | 
| 
 
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
 
urbanc 
parents: 
22533 
diff
changeset
 | 
78  | 
shows "\<Gamma>2 \<turnstile> t : T"  | 
| 
18311
 
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
 
urbanc 
parents: 
18296 
diff
changeset
 | 
79  | 
using a b c  | 
| 22533 | 80  | 
proof (nominal_induct \<Gamma>1 t T avoiding: \<Gamma>2 rule: typing.strong_induct)  | 
81  | 
case (t_Var \<Gamma>1 x T) (* variable case *)  | 
|
| 
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: 
22533 
diff
changeset
 | 
82  | 
have "\<Gamma>1 \<subseteq> \<Gamma>2" by fact  | 
| 
21052
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
83  | 
moreover  | 
| 
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
84  | 
have "valid \<Gamma>2" by fact  | 
| 
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
85  | 
moreover  | 
| 22533 | 86  | 
have "(x,T)\<in> set \<Gamma>1" by fact  | 
87  | 
ultimately show "\<Gamma>2 \<turnstile> Var x : T" by auto  | 
|
| 18105 | 88  | 
next  | 
| 22533 | 89  | 
case (t_Lam x \<Gamma>1 T1 t T2) (* lambda case *)  | 
90  | 
have vc: "x\<sharp>\<Gamma>2" by fact (* variable convention *)  | 
|
| 
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: 
22533 
diff
changeset
 | 
91  | 
have ih: "\<lbrakk>valid ((x,T1)#\<Gamma>2); (x,T1)#\<Gamma>1 \<subseteq> (x,T1)#\<Gamma>2\<rbrakk> \<Longrightarrow> (x,T1)#\<Gamma>2 \<turnstile> t:T2" by fact  | 
| 
 
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
 
urbanc 
parents: 
22533 
diff
changeset
 | 
92  | 
have "\<Gamma>1 \<subseteq> \<Gamma>2" by fact  | 
| 
 
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
 
urbanc 
parents: 
22533 
diff
changeset
 | 
93  | 
then have "(x,T1)#\<Gamma>1 \<subseteq> (x,T1)#\<Gamma>2" by simp  | 
| 18105 | 94  | 
moreover  | 
| 
21052
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
95  | 
have "valid \<Gamma>2" by fact  | 
| 22533 | 96  | 
then have "valid ((x,T1)#\<Gamma>2)" using vc by (simp add: v2)  | 
| 
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: 
22533 
diff
changeset
 | 
97  | 
ultimately have "(x,T1)#\<Gamma>2 \<turnstile> t : T2" using ih by simp  | 
| 
 
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
 
urbanc 
parents: 
22533 
diff
changeset
 | 
98  | 
with vc show "\<Gamma>2 \<turnstile> Lam [x].t : T1\<rightarrow>T2" by auto  | 
| 21488 | 99  | 
qed (auto) (* app case *)  | 
| 18105 | 100  | 
|
| 
24231
 
85fb973a8207
added type constraints to resolve syntax ambiguities;
 
wenzelm 
parents: 
23760 
diff
changeset
 | 
101  | 
lemma weakening_version3:  | 
| 
 
85fb973a8207
added type constraints to resolve syntax ambiguities;
 
wenzelm 
parents: 
23760 
diff
changeset
 | 
102  | 
fixes \<Gamma>1 \<Gamma>2 :: "(name\<times>ty) list"  | 
| 
22418
 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 
urbanc 
parents: 
22231 
diff
changeset
 | 
103  | 
assumes a: "\<Gamma>1 \<turnstile> t : T"  | 
| 
18311
 
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
 
urbanc 
parents: 
18296 
diff
changeset
 | 
104  | 
and b: "valid \<Gamma>2"  | 
| 
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: 
22533 
diff
changeset
 | 
105  | 
and c: "\<Gamma>1 \<subseteq> \<Gamma>2"  | 
| 
22418
 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 
urbanc 
parents: 
22231 
diff
changeset
 | 
106  | 
shows "\<Gamma>2 \<turnstile> t : T"  | 
| 
18311
 
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
 
urbanc 
parents: 
18296 
diff
changeset
 | 
107  | 
using a b c  | 
| 22533 | 108  | 
proof (nominal_induct \<Gamma>1 t T avoiding: \<Gamma>2 rule: typing.strong_induct)  | 
109  | 
case (t_Lam x \<Gamma>1 T1 t T2) (* lambda case *)  | 
|
110  | 
have vc: "x\<sharp>\<Gamma>2" by fact (* variable convention *)  | 
|
| 
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: 
22533 
diff
changeset
 | 
111  | 
have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; (x,T1)#\<Gamma>1 \<subseteq> \<Gamma>3\<rbrakk> \<Longrightarrow> \<Gamma>3 \<turnstile> t : T2" by fact  | 
| 
 
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
 
urbanc 
parents: 
22533 
diff
changeset
 | 
112  | 
have "\<Gamma>1 \<subseteq> \<Gamma>2" by fact  | 
| 
 
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
 
urbanc 
parents: 
22533 
diff
changeset
 | 
113  | 
then have "(x,T1)#\<Gamma>1 \<subseteq> (x,T1)#\<Gamma>2" by simp  | 
| 
18311
 
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
 
urbanc 
parents: 
18296 
diff
changeset
 | 
114  | 
moreover  | 
| 
21052
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
115  | 
have "valid \<Gamma>2" by fact  | 
| 22533 | 116  | 
then have "valid ((x,T1)#\<Gamma>2)" using vc by (simp add: v2)  | 
| 
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: 
22533 
diff
changeset
 | 
117  | 
ultimately have "(x,T1)#\<Gamma>2 \<turnstile> t : T2" using ih by simp  | 
| 
 
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
 
urbanc 
parents: 
22533 
diff
changeset
 | 
118  | 
with vc show "\<Gamma>2 \<turnstile> Lam [x].t : T1 \<rightarrow> T2" by auto  | 
| 
21052
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
119  | 
qed (auto) (* app and var case *)  | 
| 18105 | 120  | 
|
| 
21052
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
121  | 
text{* The original induction principle for the typing relation
 | 
| 
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: 
22533 
diff
changeset
 | 
122  | 
is not strong enough - even this simple lemma fails to be simple ;o) *}  | 
| 
 
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
 
urbanc 
parents: 
22533 
diff
changeset
 | 
123  | 
|
| 
18311
 
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
 
urbanc 
parents: 
18296 
diff
changeset
 | 
124  | 
lemma weakening_too_weak:  | 
| 
24231
 
85fb973a8207
added type constraints to resolve syntax ambiguities;
 
wenzelm 
parents: 
23760 
diff
changeset
 | 
125  | 
fixes \<Gamma>1 \<Gamma>2 :: "(name\<times>ty) list"  | 
| 
22418
 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 
urbanc 
parents: 
22231 
diff
changeset
 | 
126  | 
assumes a: "\<Gamma>1 \<turnstile> t : T"  | 
| 
18311
 
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
 
urbanc 
parents: 
18296 
diff
changeset
 | 
127  | 
and b: "valid \<Gamma>2"  | 
| 
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: 
22533 
diff
changeset
 | 
128  | 
and c: "\<Gamma>1 \<subseteq> \<Gamma>2"  | 
| 
22418
 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 
urbanc 
parents: 
22231 
diff
changeset
 | 
129  | 
shows "\<Gamma>2 \<turnstile> t : T"  | 
| 
18311
 
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
 
urbanc 
parents: 
18296 
diff
changeset
 | 
130  | 
using a b c  | 
| 20503 | 131  | 
proof (induct arbitrary: \<Gamma>2)  | 
| 
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: 
22533 
diff
changeset
 | 
132  | 
case (t_Var \<Gamma>1 x T) (* variable case works *)  | 
| 
 
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
 
urbanc 
parents: 
22533 
diff
changeset
 | 
133  | 
have "\<Gamma>1 \<subseteq> \<Gamma>2" by fact  | 
| 
21052
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
134  | 
moreover  | 
| 
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
135  | 
have "valid \<Gamma>2" by fact  | 
| 
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
136  | 
moreover  | 
| 22533 | 137  | 
have "(x,T) \<in> (set \<Gamma>1)" by fact  | 
138  | 
ultimately show "\<Gamma>2 \<turnstile> Var x : T" by auto  | 
|
| 18105 | 139  | 
next  | 
| 22533 | 140  | 
case (t_Lam x \<Gamma>1 T1 t T2) (* lambda case *)  | 
| 21488 | 141  | 
(* all assumptions available in this case*)  | 
| 22533 | 142  | 
have a0: "x\<sharp>\<Gamma>1" 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: 
22533 
diff
changeset
 | 
143  | 
have a1: "(x,T1)#\<Gamma>1 \<turnstile> t : T2" by fact  | 
| 
 
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
 
urbanc 
parents: 
22533 
diff
changeset
 | 
144  | 
have a2: "\<Gamma>1 \<subseteq> \<Gamma>2" by fact  | 
| 
18311
 
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
 
urbanc 
parents: 
18296 
diff
changeset
 | 
145  | 
have a3: "valid \<Gamma>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: 
22533 
diff
changeset
 | 
146  | 
have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; (x,T1)#\<Gamma>1 \<subseteq> \<Gamma>3\<rbrakk> \<Longrightarrow> \<Gamma>3 \<turnstile> t : T2" by fact  | 
| 
 
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
 
urbanc 
parents: 
22533 
diff
changeset
 | 
147  | 
have "(x,T1)#\<Gamma>1 \<subseteq> (x,T1)#\<Gamma>2" using a2 by simp  | 
| 18105 | 148  | 
moreover  | 
| 22533 | 149  | 
have "valid ((x,T1)#\<Gamma>2)" using v2 (* fails *)  | 
| 19496 | 150  | 
oops  | 
| 18105 | 151  | 
|
| 19496 | 152  | 
end  |