| author | nipkow | 
| Mon, 02 Nov 2015 18:35:30 +0100 | |
| changeset 61534 | a88e07c8d0d5 | 
| parent 32960 | 69916a850301 | 
| child 62390 | 842917225d56 | 
| permissions | -rw-r--r-- | 
| 19496 | 1  | 
theory Weakening  | 
| 25722 | 2  | 
imports "../Nominal"  | 
| 18105 | 3  | 
begin  | 
4  | 
||
| 26055 | 5  | 
text {* 
 | 
6  | 
A simple proof of the Weakening Property in the simply-typed  | 
|
7  | 
lambda-calculus. The proof is simple, because we can make use  | 
|
8  | 
of the variable convention. *}  | 
|
| 18105 | 9  | 
|
10  | 
atom_decl name  | 
|
11  | 
||
| 25751 | 12  | 
text {* Terms and Types *}
 | 
| 26055 | 13  | 
|
| 
21052
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
14  | 
nominal_datatype lam =  | 
| 
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
15  | 
Var "name"  | 
| 
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
16  | 
| App "lam" "lam"  | 
| 
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
17  | 
  | Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
 | 
| 18105 | 18  | 
|
| 18352 | 19  | 
nominal_datatype ty =  | 
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
20  | 
TVar "string"  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
21  | 
  | TArr "ty" "ty" ("_ \<rightarrow> _" [100,100] 100)
 | 
| 18105 | 22  | 
|
| 22533 | 23  | 
lemma ty_fresh:  | 
24  | 
fixes x::"name"  | 
|
25  | 
and T::"ty"  | 
|
26  | 
shows "x\<sharp>T"  | 
|
| 
26966
 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 
urbanc 
parents: 
26262 
diff
changeset
 | 
27  | 
by (nominal_induct T rule: ty.strong_induct)  | 
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
28  | 
(auto simp add: fresh_string)  | 
| 18105 | 29  | 
|
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
30  | 
text {* 
 | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
31  | 
Valid contexts (at the moment we have no type for finite  | 
| 26055 | 32  | 
sets yet, therefore typing-contexts are lists). *}  | 
33  | 
||
| 23760 | 34  | 
inductive  | 
| 18105 | 35  | 
valid :: "(name\<times>ty) list \<Rightarrow> bool"  | 
| 21364 | 36  | 
where  | 
37  | 
v1[intro]: "valid []"  | 
|
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
38  | 
| v2[intro]: "\<lbrakk>valid \<Gamma>;x\<sharp>\<Gamma>\<rbrakk>\<Longrightarrow> valid ((x,T)#\<Gamma>)"  | 
| 18105 | 39  | 
|
| 22533 | 40  | 
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
 | 
41  | 
|
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
42  | 
text{* Typing judgements *}
 | 
| 26055 | 43  | 
|
| 23760 | 44  | 
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
 | 
45  | 
  typing :: "(name\<times>ty) list\<Rightarrow>lam\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ : _" [60,60,60] 60) 
 | 
| 21364 | 46  | 
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
 | 
47  | 
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
 | 
48  | 
| t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : T1\<rightarrow>T2; \<Gamma> \<turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : T2"  | 
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
49  | 
| 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
 | 
50  | 
|
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
51  | 
text {* 
 | 
| 25722 | 52  | 
We derive the strong induction principle for the typing  | 
53  | 
relation (this induction principle has the variable convention  | 
|
| 26055 | 54  | 
already built-in). *}  | 
55  | 
||
| 
22730
 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 
berghofe 
parents: 
22650 
diff
changeset
 | 
56  | 
equivariance typing  | 
| 22533 | 57  | 
nominal_inductive typing  | 
58  | 
by (simp_all add: abs_fresh ty_fresh)  | 
|
| 18105 | 59  | 
|
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
60  | 
text {* Abbreviation for the notion of subcontexts. *}
 | 
| 26055 | 61  | 
|
| 
21052
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
62  | 
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
 | 
63  | 
  "sub_context" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" ("_ \<subseteq> _" [60,60] 60) 
 | 
| 22533 | 64  | 
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
 | 
65  | 
"\<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
 | 
66  | 
|
| 22533 | 67  | 
text {* Now it comes: The Weakening Lemma *}
 | 
| 18105 | 68  | 
|
| 26055 | 69  | 
text {*
 | 
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
70  | 
The first version is, after setting up the induction,  | 
| 26055 | 71  | 
completely automatic except for use of atomize. *}  | 
72  | 
||
| 
18311
 
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
 
urbanc 
parents: 
18296 
diff
changeset
 | 
73  | 
lemma weakening_version1:  | 
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
74  | 
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
 | 
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"  | 
| 
22418
 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 
urbanc 
parents: 
22231 
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  | 
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
 | 
81  | 
(auto | atomize)+  | 
| 18105 | 82  | 
|
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
83  | 
text {* 
 | 
| 26055 | 84  | 
The second version gives the details for the variable  | 
85  | 
and lambda case. *}  | 
|
86  | 
||
| 
18311
 
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
 
urbanc 
parents: 
18296 
diff
changeset
 | 
87  | 
lemma weakening_version2:  | 
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
88  | 
fixes \<Gamma>1 \<Gamma>2::"(name\<times>ty) list"  | 
| 18105 | 89  | 
and t ::"lam"  | 
90  | 
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
 | 
91  | 
assumes a: "\<Gamma>1 \<turnstile> t : T"  | 
| 
18311
 
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
 
urbanc 
parents: 
18296 
diff
changeset
 | 
92  | 
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
 | 
93  | 
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
 | 
94  | 
shows "\<Gamma>2 \<turnstile> t : T"  | 
| 
18311
 
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
 
urbanc 
parents: 
18296 
diff
changeset
 | 
95  | 
using a b c  | 
| 22533 | 96  | 
proof (nominal_induct \<Gamma>1 t T avoiding: \<Gamma>2 rule: typing.strong_induct)  | 
97  | 
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
 | 
98  | 
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
 | 
99  | 
moreover  | 
| 
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
100  | 
have "valid \<Gamma>2" by fact  | 
| 
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
101  | 
moreover  | 
| 22533 | 102  | 
have "(x,T)\<in> set \<Gamma>1" by fact  | 
103  | 
ultimately show "\<Gamma>2 \<turnstile> Var x : T" by auto  | 
|
| 18105 | 104  | 
next  | 
| 22533 | 105  | 
case (t_Lam x \<Gamma>1 T1 t T2) (* lambda case *)  | 
106  | 
have vc: "x\<sharp>\<Gamma>2" by fact (* variable convention *)  | 
|
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
107  | 
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  | 
| 
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
 | 
108  | 
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
 | 
109  | 
then have "(x,T1)#\<Gamma>1 \<subseteq> (x,T1)#\<Gamma>2" by simp  | 
| 18105 | 110  | 
moreover  | 
| 
21052
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
111  | 
have "valid \<Gamma>2" by fact  | 
| 22533 | 112  | 
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
 | 
113  | 
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
 | 
114  | 
with vc show "\<Gamma>2 \<turnstile> Lam [x].t : T1\<rightarrow>T2" by auto  | 
| 21488 | 115  | 
qed (auto) (* app case *)  | 
| 18105 | 116  | 
|
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
117  | 
text{* 
 | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
118  | 
The original induction principle for the typing relation  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
119  | 
is not strong enough - even this simple lemma fails to be  | 
| 26055 | 120  | 
simple ;o) *}  | 
121  | 
||
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
122  | 
lemma weakening_not_straigh_forward:  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
123  | 
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
 | 
124  | 
assumes a: "\<Gamma>1 \<turnstile> t : T"  | 
| 
18311
 
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
 
urbanc 
parents: 
18296 
diff
changeset
 | 
125  | 
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
 | 
126  | 
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
 | 
127  | 
shows "\<Gamma>2 \<turnstile> t : T"  | 
| 
18311
 
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
 
urbanc 
parents: 
18296 
diff
changeset
 | 
128  | 
using a b c  | 
| 20503 | 129  | 
proof (induct arbitrary: \<Gamma>2)  | 
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
130  | 
case (t_Var \<Gamma>1 x T) (* variable case still works *)  | 
| 
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
 | 
131  | 
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
 | 
132  | 
moreover  | 
| 
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
133  | 
have "valid \<Gamma>2" by fact  | 
| 
 
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
 
urbanc 
parents: 
20955 
diff
changeset
 | 
134  | 
moreover  | 
| 22533 | 135  | 
have "(x,T) \<in> (set \<Gamma>1)" by fact  | 
136  | 
ultimately show "\<Gamma>2 \<turnstile> Var x : T" by auto  | 
|
| 18105 | 137  | 
next  | 
| 22533 | 138  | 
case (t_Lam x \<Gamma>1 T1 t T2) (* lambda case *)  | 
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
139  | 
(* These are all assumptions available in this case*)  | 
| 22533 | 140  | 
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
 | 
141  | 
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
 | 
142  | 
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
 | 
143  | 
have a3: "valid \<Gamma>2" by fact  | 
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
144  | 
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  | 
| 
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
 | 
145  | 
have "(x,T1)#\<Gamma>1 \<subseteq> (x,T1)#\<Gamma>2" using a2 by simp  | 
| 18105 | 146  | 
moreover  | 
| 22533 | 147  | 
have "valid ((x,T1)#\<Gamma>2)" using v2 (* fails *)  | 
| 19496 | 148  | 
oops  | 
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
149  | 
|
| 26055 | 150  | 
text{* 
 | 
| 26262 | 151  | 
To show that the proof with explicit renaming is not simple,  | 
152  | 
here is the proof without using the variable convention. It  | 
|
153  | 
crucially depends on the equivariance property of the typing  | 
|
154  | 
relation.  | 
|
| 26055 | 155  | 
|
156  | 
*}  | 
|
| 25751 | 157  | 
|
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
158  | 
lemma weakening_with_explicit_renaming:  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
159  | 
fixes \<Gamma>1 \<Gamma>2::"(name\<times>ty) list"  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
160  | 
assumes a: "\<Gamma>1 \<turnstile> t : T"  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
161  | 
and b: "valid \<Gamma>2"  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
162  | 
and c: "\<Gamma>1 \<subseteq> \<Gamma>2"  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
163  | 
shows "\<Gamma>2 \<turnstile> t : T"  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
164  | 
using a b c  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
165  | 
proof (induct arbitrary: \<Gamma>2)  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
166  | 
case (t_Lam x \<Gamma>1 T1 t T2) (* lambda case *)  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
167  | 
have fc0: "x\<sharp>\<Gamma>1" by fact  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
168  | 
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  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
169  | 
obtain c::"name" where fc1: "c\<sharp>(x,t,\<Gamma>1,\<Gamma>2)" (* we obtain a fresh name *)  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
170  | 
by (rule exists_fresh) (auto simp add: fs_name1)  | 
| 25138 | 171  | 
have "Lam [c].([(c,x)]\<bullet>t) = Lam [x].t" using fc1 (* we then alpha-rename the lambda-abstraction *)  | 
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
172  | 
by (auto simp add: lam.inject alpha fresh_prod fresh_atm)  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
173  | 
moreover  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
174  | 
have "\<Gamma>2 \<turnstile> Lam [c].([(c,x)]\<bullet>t) : T1 \<rightarrow> T2" (* we can then alpha-rename our original goal *)  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
175  | 
proof -  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
176  | 
(* we establish (x,T1)#\<Gamma>1 \<subseteq> (x,T1)#([(c,x)]\<bullet>\<Gamma>2) and valid ((x,T1)#([(c,x)]\<bullet>\<Gamma>2)) *)  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
177  | 
have "(x,T1)#\<Gamma>1 \<subseteq> (x,T1)#([(c,x)]\<bullet>\<Gamma>2)"  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
178  | 
proof -  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
179  | 
have "\<Gamma>1 \<subseteq> \<Gamma>2" by fact  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
180  | 
then have "[(c,x)]\<bullet>((x,T1)#\<Gamma>1 \<subseteq> (x,T1)#([(c,x)]\<bullet>\<Gamma>2))" using fc0 fc1  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
26966 
diff
changeset
 | 
181  | 
by (perm_simp add: eqvts calc_atm perm_fresh_fresh ty_fresh)  | 
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
182  | 
then show "(x,T1)#\<Gamma>1 \<subseteq> (x,T1)#([(c,x)]\<bullet>\<Gamma>2)" by (rule perm_boolE)  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
183  | 
qed  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
184  | 
moreover  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
185  | 
have "valid ((x,T1)#([(c,x)]\<bullet>\<Gamma>2))"  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
186  | 
proof -  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
187  | 
have "valid \<Gamma>2" by fact  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
188  | 
then show "valid ((x,T1)#([(c,x)]\<bullet>\<Gamma>2))" using fc1  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
26966 
diff
changeset
 | 
189  | 
by (auto intro!: v2 simp add: fresh_left calc_atm eqvts)  | 
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
190  | 
qed  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
191  | 
(* these two facts give us by induction hypothesis the following *)  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
192  | 
ultimately have "(x,T1)#([(c,x)]\<bullet>\<Gamma>2) \<turnstile> t : T2" using ih by simp  | 
| 25138 | 193  | 
(* we now apply renamings to get to our goal *)  | 
| 
25137
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
194  | 
then have "[(c,x)]\<bullet>((x,T1)#([(c,x)]\<bullet>\<Gamma>2) \<turnstile> t : T2)" by (rule perm_boolI)  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
195  | 
then have "(c,T1)#\<Gamma>2 \<turnstile> ([(c,x)]\<bullet>t) : T2" using fc1  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
196  | 
by (perm_simp add: eqvts calc_atm perm_fresh_fresh ty_fresh)  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
197  | 
then show "\<Gamma>2 \<turnstile> Lam [c].([(c,x)]\<bullet>t) : T1 \<rightarrow> T2" using fc1 by auto  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
198  | 
qed  | 
| 
 
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
 
urbanc 
parents: 
24231 
diff
changeset
 | 
199  | 
ultimately show "\<Gamma>2 \<turnstile> Lam [x].t : T1 \<rightarrow> T2" by simp  | 
| 26262 | 200  | 
qed (auto) (* var and app cases *)  | 
| 18105 | 201  | 
|
| 25138 | 202  | 
text {*
 | 
| 26055 | 203  | 
Moral: compare the proof with explicit renamings to weakening_version1  | 
204  | 
and weakening_version2, and imagine you are proving something more substantial  | 
|
205  | 
than the weakening lemma. *}  | 
|
| 25138 | 206  | 
|
| 19496 | 207  | 
end  |