| author | haftmann |
| Fri, 23 Nov 2007 21:09:32 +0100 | |
| changeset 25459 | d1dce7d0731c |
| parent 25138 | e453c480d599 |
| child 25722 | 0a104ddb72d9 |
| 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 |
||
|
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
|
12 |
text {* Terms and types *}
|
|
21052
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
urbanc
parents:
20955
diff
changeset
|
13 |
nominal_datatype lam = |
|
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
urbanc
parents:
20955
diff
changeset
|
14 |
Var "name" |
|
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
urbanc
parents:
20955
diff
changeset
|
15 |
| App "lam" "lam" |
|
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
urbanc
parents:
20955
diff
changeset
|
16 |
| Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
|
| 18105 | 17 |
|
| 18352 | 18 |
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
|
19 |
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
|
20 |
| TArr "ty" "ty" ("_ \<rightarrow> _" [100,100] 100)
|
| 18105 | 21 |
|
| 22533 | 22 |
lemma ty_fresh: |
23 |
fixes x::"name" |
|
24 |
and T::"ty" |
|
25 |
shows "x\<sharp>T" |
|
26 |
by (nominal_induct T rule: ty.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
|
27 |
(auto simp add: fresh_string) |
| 18105 | 28 |
|
|
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
|
29 |
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
|
30 |
Valid contexts (at the moment we have no type for finite |
|
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 |
sets yet; therefore typing-contexts are lists). *} |
| 23760 | 32 |
inductive |
| 18105 | 33 |
valid :: "(name\<times>ty) list \<Rightarrow> bool" |
| 21364 | 34 |
where |
35 |
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
|
36 |
| v2[intro]: "\<lbrakk>valid \<Gamma>;x\<sharp>\<Gamma>\<rbrakk>\<Longrightarrow> valid ((x,T)#\<Gamma>)" |
| 18105 | 37 |
|
| 22533 | 38 |
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
|
39 |
|
|
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
|
40 |
text{* Typing judgements *}
|
| 23760 | 41 |
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
|
42 |
typing :: "(name\<times>ty) list\<Rightarrow>lam\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ : _" [60,60,60] 60)
|
| 21364 | 43 |
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
|
44 |
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
|
45 |
| 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
|
46 |
| 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
|
47 |
|
|
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
|
48 |
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
|
49 |
We automatically derive the strong induction principle |
|
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
urbanc
parents:
24231
diff
changeset
|
50 |
for the typing relation (this induction principle has the |
|
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 |
variable convention already built in). *} |
|
22730
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22650
diff
changeset
|
52 |
equivariance typing |
| 22533 | 53 |
nominal_inductive typing |
54 |
by (simp_all add: abs_fresh ty_fresh) |
|
| 18105 | 55 |
|
|
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
|
56 |
text {* Abbreviation for the notion of subcontexts. *}
|
|
21052
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
urbanc
parents:
20955
diff
changeset
|
57 |
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
|
58 |
"sub_context" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" ("_ \<subseteq> _" [60,60] 60)
|
| 22533 | 59 |
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
|
60 |
"\<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
|
61 |
|
| 22533 | 62 |
text {* Now it comes: The Weakening Lemma *}
|
|
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
|
63 |
(*========================================*) |
| 18105 | 64 |
|
|
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
|
65 |
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
|
66 |
The first version is, after setting up the induction, |
|
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
urbanc
parents:
24231
diff
changeset
|
67 |
quite automatic except for use of atomize. *} |
|
18311
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents:
18296
diff
changeset
|
68 |
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
|
69 |
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
|
70 |
assumes a: "\<Gamma>1 \<turnstile> t : T" |
|
18311
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents:
18296
diff
changeset
|
71 |
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
|
72 |
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
|
73 |
shows "\<Gamma>2 \<turnstile> t : T" |
|
18311
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents:
18296
diff
changeset
|
74 |
using a b c |
| 22533 | 75 |
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
|
76 |
(auto | atomize)+ |
| 18105 | 77 |
|
|
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
|
78 |
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
|
79 |
The second version gives all details for the variable |
|
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
urbanc
parents:
24231
diff
changeset
|
80 |
and lambda case. *} |
|
18311
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents:
18296
diff
changeset
|
81 |
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
|
82 |
fixes \<Gamma>1 \<Gamma>2::"(name\<times>ty) list" |
| 18105 | 83 |
and t ::"lam" |
84 |
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
|
85 |
assumes a: "\<Gamma>1 \<turnstile> t : T" |
|
18311
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents:
18296
diff
changeset
|
86 |
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
|
87 |
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
|
88 |
shows "\<Gamma>2 \<turnstile> t : T" |
|
18311
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents:
18296
diff
changeset
|
89 |
using a b c |
| 22533 | 90 |
proof (nominal_induct \<Gamma>1 t T avoiding: \<Gamma>2 rule: typing.strong_induct) |
91 |
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
|
92 |
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
|
93 |
moreover |
|
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
urbanc
parents:
20955
diff
changeset
|
94 |
have "valid \<Gamma>2" by fact |
|
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
urbanc
parents:
20955
diff
changeset
|
95 |
moreover |
| 22533 | 96 |
have "(x,T)\<in> set \<Gamma>1" by fact |
97 |
ultimately show "\<Gamma>2 \<turnstile> Var x : T" by auto |
|
| 18105 | 98 |
next |
| 22533 | 99 |
case (t_Lam x \<Gamma>1 T1 t T2) (* lambda case *) |
100 |
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
|
101 |
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
|
102 |
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
|
103 |
then have "(x,T1)#\<Gamma>1 \<subseteq> (x,T1)#\<Gamma>2" by simp |
| 18105 | 104 |
moreover |
|
21052
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
urbanc
parents:
20955
diff
changeset
|
105 |
have "valid \<Gamma>2" by fact |
| 22533 | 106 |
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
|
107 |
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
|
108 |
with vc show "\<Gamma>2 \<turnstile> Lam [x].t : T1\<rightarrow>T2" by auto |
| 21488 | 109 |
qed (auto) (* app case *) |
| 18105 | 110 |
|
|
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
|
111 |
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
|
112 |
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
|
113 |
is not strong enough - even this simple lemma fails to be |
|
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
urbanc
parents:
24231
diff
changeset
|
114 |
simple ;o) *} |
|
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
urbanc
parents:
24231
diff
changeset
|
115 |
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
|
116 |
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
|
117 |
assumes a: "\<Gamma>1 \<turnstile> t : T" |
|
18311
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents:
18296
diff
changeset
|
118 |
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
|
119 |
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
|
120 |
shows "\<Gamma>2 \<turnstile> t : T" |
|
18311
b83b00cbaecf
changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents:
18296
diff
changeset
|
121 |
using a b c |
| 20503 | 122 |
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
|
123 |
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
|
124 |
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
|
125 |
moreover |
|
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
urbanc
parents:
20955
diff
changeset
|
126 |
have "valid \<Gamma>2" by fact |
|
ec5531061ed6
adapted to Stefan's new inductive package and cleaning up
urbanc
parents:
20955
diff
changeset
|
127 |
moreover |
| 22533 | 128 |
have "(x,T) \<in> (set \<Gamma>1)" by fact |
129 |
ultimately show "\<Gamma>2 \<turnstile> Var x : T" by auto |
|
| 18105 | 130 |
next |
| 22533 | 131 |
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
|
132 |
(* These are all assumptions available in this case*) |
| 22533 | 133 |
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
|
134 |
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
|
135 |
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
|
136 |
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
|
137 |
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
|
138 |
have "(x,T1)#\<Gamma>1 \<subseteq> (x,T1)#\<Gamma>2" using a2 by simp |
| 18105 | 139 |
moreover |
| 22533 | 140 |
have "valid ((x,T1)#\<Gamma>2)" using v2 (* fails *) |
| 19496 | 141 |
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
|
142 |
|
|
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
urbanc
parents:
24231
diff
changeset
|
143 |
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
|
144 |
The complete proof without using the variable convention. *} |
|
0835ac64834a
polished the proofs and added a version of the weakening lemma that does not use the variable convention
urbanc
parents:
24231
diff
changeset
|
145 |
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
|
146 |
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
|
147 |
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
|
148 |
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
|
149 |
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
|
150 |
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
|
151 |
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
|
152 |
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
|
153 |
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
|
154 |
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
|
155 |
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
|
156 |
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
|
157 |
by (rule exists_fresh) (auto simp add: fs_name1) |
| 25138 | 158 |
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
|
159 |
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
|
160 |
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
|
161 |
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
|
162 |
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
|
163 |
(* 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
|
164 |
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
|
165 |
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
|
166 |
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
|
167 |
then have "[(c,x)]\<bullet>((x,T1)#\<Gamma>1 \<subseteq> (x,T1)#([(c,x)]\<bullet>\<Gamma>2))" using fc0 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
|
168 |
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
|
169 |
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
|
170 |
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
|
171 |
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
|
172 |
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
|
173 |
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
|
174 |
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
|
175 |
then show "valid ((x,T1)#([(c,x)]\<bullet>\<Gamma>2))" 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
|
176 |
by (auto intro!: v2 simp add: fresh_left calc_atm eqvts) |
|
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 |
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
|
178 |
(* 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
|
179 |
ultimately have "(x,T1)#([(c,x)]\<bullet>\<Gamma>2) \<turnstile> t : T2" using ih by simp |
| 25138 | 180 |
(* 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
|
181 |
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
|
182 |
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
|
183 |
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
|
184 |
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
|
185 |
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
|
186 |
ultimately show "\<Gamma>2 \<turnstile> Lam [x].t : T1 \<rightarrow> T2" by simp |
|
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 |
qed (auto) |
| 18105 | 188 |
|
| 25138 | 189 |
text {*
|
190 |
Compare the proof with explicit renamings to version1 and version2, |
|
191 |
and imagine you are proving something more substantial than the |
|
192 |
weakening lemma. *} |
|
193 |
||
| 19496 | 194 |
end |