| author | blanchet | 
| Sun, 01 Mar 2009 18:40:16 +0100 | |
| changeset 30238 | d8944fd4365e | 
| parent 29097 | 68245155eb58 | 
| child 30091 | 2fb0b721e9c2 | 
| permissions | -rw-r--r-- | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
1  | 
(*<*)  | 
| 19501 | 2  | 
theory Fsub  | 
3  | 
imports "../Nominal"  | 
|
| 18269 | 4  | 
begin  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
5  | 
(*>*)  | 
| 18269 | 6  | 
|
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
7  | 
text{* Authors: Christian Urban,
 | 
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
8  | 
Benjamin Pierce,  | 
| 18650 | 9  | 
Dimitrios Vytiniotis  | 
10  | 
Stephanie Weirich and  | 
|
11  | 
Steve Zdancewic  | 
|
| 
18266
 
55c201fe4c95
added an authors section (please let me know if somebody is left out or unhappy)
 
urbanc 
parents: 
18263 
diff
changeset
 | 
12  | 
|
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
13  | 
with great help from Stefan Berghofer and Markus Wenzel. *}  | 
| 18246 | 14  | 
|
| 18621 | 15  | 
section {* Types for Names, Nominal Datatype Declaration for Types and Terms *}
 | 
| 18424 | 16  | 
|
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
17  | 
text {* The main point of this solution is to use names everywhere (be they bound, 
 | 
| 18621 | 18  | 
  binding or free). In System \FSUB{} there are two kinds of names corresponding to 
 | 
19  | 
type-variables and to term-variables. These two kinds of names are represented in  | 
|
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
20  | 
  the nominal datatype package as atom-types @{text "tyvrs"} and @{text "vrs"}: *}
 | 
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
21  | 
|
| 18246 | 22  | 
atom_decl tyvrs vrs  | 
23  | 
||
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
24  | 
text{* There are numerous facts that come with this declaration: for example that 
 | 
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
25  | 
  there are infinitely many elements in @{text "tyvrs"} and @{text "vrs"}. *}
 | 
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
26  | 
|
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
27  | 
text{* The constructors for types and terms in System \FSUB{} contain abstractions 
 | 
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
28  | 
over type-variables and term-variables. The nominal datatype-package uses  | 
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
29  | 
  @{text "\<guillemotleft>\<dots>\<guillemotright>\<dots>"} to indicate where abstractions occur. *}
 | 
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
30  | 
|
| 18424 | 31  | 
nominal_datatype ty =  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
32  | 
Tvar "tyvrs"  | 
| 18424 | 33  | 
| Top  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
34  | 
  | Arrow  "ty" "ty"          ("_ \<rightarrow> _" [100,100] 100)
 | 
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
35  | 
| Forall "\<guillemotleft>tyvrs\<guillemotright>ty" "ty"  | 
| 18246 | 36  | 
|
| 18424 | 37  | 
nominal_datatype trm =  | 
38  | 
Var "vrs"  | 
|
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
39  | 
| Lam "\<guillemotleft>vrs\<guillemotright>trm" "ty"  | 
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
40  | 
| Tabs "\<guillemotleft>tyvrs\<guillemotright>trm" "ty"  | 
| 18424 | 41  | 
| App "trm" "trm"  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
42  | 
| Tapp "trm" "ty"  | 
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
43  | 
|
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
44  | 
text {* To be polite to the eye, some more familiar notation is introduced. 
 | 
| 18621 | 45  | 
Because of the change in the order of arguments, one needs to use  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
46  | 
translation rules, instead of syntax annotations at the term-constructors  | 
| 18650 | 47  | 
  as given above for @{term "Arrow"}. *}
 | 
| 18246 | 48  | 
|
49  | 
syntax  | 
|
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
50  | 
  Forall_syn :: "tyvrs \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> ty" ("\<forall>[_<:_]._" [100,100,100] 100)
 | 
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
51  | 
  Lam_syn    :: "vrs \<Rightarrow> ty \<Rightarrow> trm \<Rightarrow> trm"   ("Lam [_:_]._" [100,100,100] 100)
 | 
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
52  | 
  Tabs_syn   :: "tyvrs \<Rightarrow> ty \<Rightarrow> trm \<Rightarrow> trm" ("Tabs [_<:_]._" [100,100,100] 100)
 | 
| 18424 | 53  | 
|
| 18246 | 54  | 
translations  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
55  | 
"\<forall>[X<:T\<^isub>1].T\<^isub>2" \<rightleftharpoons> "ty.Forall X T\<^isub>2 T\<^isub>1"  | 
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
56  | 
"Lam [x:T].t" \<rightleftharpoons> "trm.Lam x t T"  | 
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
57  | 
"Tabs [X<:T].t" \<rightleftharpoons> "trm.Tabs X t T"  | 
| 18246 | 58  | 
|
| 18621 | 59  | 
text {* Again there are numerous facts that are proved automatically for @{typ "ty"} 
 | 
| 18650 | 60  | 
  and @{typ "trm"}: for example that the set of free variables, i.e.~the @{text "support"}, 
 | 
61  | 
  is finite. However note that nominal-datatype declarations do \emph{not} define 
 | 
|
62  | 
``classical" constructor-based datatypes, but rather define $\alpha$-equivalence  | 
|
| 18621 | 63  | 
  classes---we can for example show that $\alpha$-equivalent @{typ "ty"}s 
 | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
64  | 
  and @{typ "trm"}s are equal: *}
 | 
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
65  | 
|
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
66  | 
lemma alpha_illustration:  | 
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
67  | 
shows "\<forall>[X<:T].(Tvar X) = \<forall>[Y<:T].(Tvar Y)"  | 
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
68  | 
and "Lam [x:T].(Var x) = Lam [y:T].(Var y)"  | 
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
69  | 
by (simp_all add: ty.inject trm.inject alpha calc_atm fresh_atm)  | 
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
70  | 
|
| 18621 | 71  | 
section {* SubTyping Contexts *}
 | 
| 18246 | 72  | 
|
73  | 
types ty_context = "(tyvrs\<times>ty) list"  | 
|
74  | 
||
| 18650 | 75  | 
text {* Typing contexts are represented as lists that ``grow" on the left; we
 | 
| 18621 | 76  | 
thereby deviating from the convention in the POPLmark-paper. The lists contain  | 
| 18650 | 77  | 
pairs of type-variables and types (this is sufficient for Part 1A). *}  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
78  | 
|
| 18628 | 79  | 
text {* In order to state validity-conditions for typing-contexts, the notion of
 | 
| 18621 | 80  | 
  a @{text "domain"} of a typing-context is handy. *}
 | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
81  | 
|
| 18246 | 82  | 
consts  | 
83  | 
"domain" :: "ty_context \<Rightarrow> tyvrs set"  | 
|
84  | 
primrec  | 
|
85  | 
  "domain [] = {}"
 | 
|
86  | 
  "domain (X#\<Gamma>) = {fst X}\<union>(domain \<Gamma>)" 
 | 
|
87  | 
||
| 22537 | 88  | 
lemma domain_eqvt[eqvt]:  | 
| 18246 | 89  | 
fixes pi::"tyvrs prm"  | 
| 22537 | 90  | 
and pi'::"vrs prm"  | 
91  | 
shows "pi\<bullet>(domain \<Gamma>) = domain (pi\<bullet>\<Gamma>)"  | 
|
92  | 
and "pi'\<bullet>(domain \<Gamma>) = domain (pi'\<bullet>\<Gamma>)"  | 
|
| 
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: 
22537 
diff
changeset
 | 
93  | 
by (induct \<Gamma>) (simp_all add: eqvts)  | 
| 18246 | 94  | 
|
95  | 
lemma finite_domain:  | 
|
96  | 
shows "finite (domain \<Gamma>)"  | 
|
97  | 
by (induct \<Gamma>, auto)  | 
|
98  | 
||
| 18621 | 99  | 
lemma domain_supp:  | 
100  | 
shows "(supp (domain \<Gamma>)) = (domain \<Gamma>)"  | 
|
101  | 
by (simp only: at_fin_set_supp at_tyvrs_inst finite_domain)  | 
|
102  | 
||
| 18424 | 103  | 
lemma domain_inclusion:  | 
104  | 
assumes a: "(X,T)\<in>set \<Gamma>"  | 
|
105  | 
shows "X\<in>(domain \<Gamma>)"  | 
|
106  | 
using a by (induct \<Gamma>, auto)  | 
|
| 18246 | 107  | 
|
| 18424 | 108  | 
lemma domain_existence:  | 
109  | 
assumes a: "X\<in>(domain \<Gamma>)"  | 
|
110  | 
shows "\<exists>T.(X,T)\<in>set \<Gamma>"  | 
|
111  | 
using a by (induct \<Gamma>, auto)  | 
|
| 18246 | 112  | 
|
113  | 
lemma domain_append:  | 
|
114  | 
shows "domain (\<Gamma>@\<Delta>) = ((domain \<Gamma>) \<union> (domain \<Delta>))"  | 
|
115  | 
by (induct \<Gamma>, auto)  | 
|
116  | 
||
| 18621 | 117  | 
lemma fresh_domain_cons:  | 
118  | 
fixes X::"tyvrs"  | 
|
119  | 
shows "X\<sharp>(domain (Y#\<Gamma>)) = (X\<sharp>(fst Y) \<and> X\<sharp>(domain \<Gamma>))"  | 
|
120  | 
by (simp add: fresh_fin_insert pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst finite_domain)  | 
|
121  | 
||
122  | 
lemma fresh_domain:  | 
|
123  | 
fixes X::"tyvrs"  | 
|
124  | 
assumes a: "X\<sharp>\<Gamma>"  | 
|
125  | 
shows "X\<sharp>(domain \<Gamma>)"  | 
|
126  | 
using a  | 
|
127  | 
apply(induct \<Gamma>)  | 
|
128  | 
apply(simp add: fresh_set_empty)  | 
|
129  | 
apply(simp only: fresh_domain_cons)  | 
|
130  | 
apply(auto simp add: fresh_prod fresh_list_cons)  | 
|
131  | 
done  | 
|
132  | 
||
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
133  | 
text {* Not all lists of type @{typ "ty_context"} are well-formed. One condition
 | 
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
134  | 
  requires that in @{term "(X,S)#\<Gamma>"} all free variables of @{term "S"} must be 
 | 
| 18621 | 135  | 
  in the @{term "domain"} of @{term "\<Gamma>"}, that is @{term "S"} must be @{text "closed"} 
 | 
| 18650 | 136  | 
  in @{term "\<Gamma>"}. The set of free variables of @{term "S"} is the 
 | 
| 18621 | 137  | 
  @{text "support"} of @{term "S"}. *}
 | 
| 18246 | 138  | 
|
139  | 
constdefs  | 
|
140  | 
  "closed_in" :: "ty \<Rightarrow> ty_context \<Rightarrow> bool" ("_ closed'_in _" [100,100] 100)
 | 
|
141  | 
"S closed_in \<Gamma> \<equiv> (supp S)\<subseteq>(domain \<Gamma>)"  | 
|
142  | 
||
| 22537 | 143  | 
lemma closed_in_eqvt[eqvt]:  | 
| 18246 | 144  | 
fixes pi::"tyvrs prm"  | 
145  | 
assumes a: "S closed_in \<Gamma>"  | 
|
146  | 
shows "(pi\<bullet>S) closed_in (pi\<bullet>\<Gamma>)"  | 
|
147  | 
using a  | 
|
| 26091 | 148  | 
proof -  | 
149  | 
from a have "pi\<bullet>(S closed_in \<Gamma>)" by (simp add: perm_bool)  | 
|
150  | 
then show "(pi\<bullet>S) closed_in (pi\<bullet>\<Gamma>)" by (simp add: closed_in_def eqvts)  | 
|
| 18246 | 151  | 
qed  | 
152  | 
||
| 22537 | 153  | 
lemma ty_vrs_prm_simp:  | 
154  | 
fixes pi::"vrs prm"  | 
|
155  | 
and S::"ty"  | 
|
156  | 
shows "pi\<bullet>S = S"  | 
|
| 
26966
 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 
urbanc 
parents: 
26091 
diff
changeset
 | 
157  | 
by (induct S rule: ty.induct) (auto simp add: calc_atm)  | 
| 22537 | 158  | 
|
159  | 
lemma ty_context_vrs_prm_simp:  | 
|
160  | 
fixes pi::"vrs prm"  | 
|
161  | 
and \<Gamma>::"ty_context"  | 
|
162  | 
shows "pi\<bullet>\<Gamma> = \<Gamma>"  | 
|
163  | 
by (induct \<Gamma>)  | 
|
164  | 
(auto simp add: calc_atm ty_vrs_prm_simp)  | 
|
165  | 
||
166  | 
lemma closed_in_eqvt'[eqvt]:  | 
|
167  | 
fixes pi::"vrs prm"  | 
|
168  | 
assumes a: "S closed_in \<Gamma>"  | 
|
169  | 
shows "(pi\<bullet>S) closed_in (pi\<bullet>\<Gamma>)"  | 
|
170  | 
using a  | 
|
171  | 
by (simp add: ty_vrs_prm_simp ty_context_vrs_prm_simp)  | 
|
172  | 
||
| 18621 | 173  | 
text {* Now validity of a context is a straightforward inductive definition. *}
 | 
174  | 
||
| 23760 | 175  | 
inductive  | 
| 22537 | 176  | 
  valid_rel :: "ty_context \<Rightarrow> bool" ("\<turnstile> _ ok" [100] 100)
 | 
| 22436 | 177  | 
where  | 
178  | 
valid_nil[simp]: "\<turnstile> [] ok"  | 
|
179  | 
| valid_cons[simp]: "\<lbrakk>\<turnstile> \<Gamma> ok; X\<sharp>(domain \<Gamma>); T closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<turnstile> ((X,T)#\<Gamma>) ok"  | 
|
| 18246 | 180  | 
|
| 22537 | 181  | 
equivariance valid_rel  | 
| 18246 | 182  | 
|
183  | 
lemma validE:  | 
|
184  | 
assumes a: "\<turnstile> ((X,T)#\<Gamma>) ok"  | 
|
| 18621 | 185  | 
shows "\<turnstile> \<Gamma> ok \<and> X\<sharp>(domain \<Gamma>) \<and> T closed_in \<Gamma>"  | 
| 18246 | 186  | 
using a by (cases, auto)  | 
187  | 
||
| 18424 | 188  | 
lemma validE_append:  | 
189  | 
assumes a: "\<turnstile> (\<Delta>@\<Gamma>) ok"  | 
|
190  | 
shows "\<turnstile> \<Gamma> ok"  | 
|
191  | 
using a by (induct \<Delta>, auto dest: validE)  | 
|
| 18246 | 192  | 
|
| 18424 | 193  | 
lemma replace_type:  | 
194  | 
assumes a: "\<turnstile> (\<Delta>@(X,T)#\<Gamma>) ok"  | 
|
195  | 
and b: "S closed_in \<Gamma>"  | 
|
196  | 
shows "\<turnstile> (\<Delta>@(X,S)#\<Gamma>) ok"  | 
|
| 18621 | 197  | 
using a b  | 
| 18246 | 198  | 
apply(induct \<Delta>)  | 
| 18621 | 199  | 
apply(auto dest!: validE intro!: valid_cons simp add: domain_append closed_in_def)  | 
| 18246 | 200  | 
done  | 
201  | 
||
| 18650 | 202  | 
text {* Well-formed contexts have a unique type-binding for a type-variable. *} 
 | 
203  | 
||
| 18246 | 204  | 
lemma uniqueness_of_ctxt:  | 
205  | 
fixes \<Gamma>::"ty_context"  | 
|
| 18412 | 206  | 
assumes a: "\<turnstile> \<Gamma> ok"  | 
207  | 
and b: "(X,T)\<in>set \<Gamma>"  | 
|
208  | 
and c: "(X,S)\<in>set \<Gamma>"  | 
|
209  | 
shows "T=S"  | 
|
| 18621 | 210  | 
using a b c  | 
211  | 
proof (induct)  | 
|
212  | 
case valid_nil thus "T=S" by simp  | 
|
213  | 
next  | 
|
| 22436 | 214  | 
case valid_cons  | 
| 18621 | 215  | 
moreover  | 
216  | 
  { fix \<Gamma>::"ty_context"
 | 
|
217  | 
assume a: "X\<sharp>(domain \<Gamma>)"  | 
|
218  | 
have "\<not>(\<exists>T.(X,T)\<in>(set \<Gamma>))" using a  | 
|
219  | 
proof (induct \<Gamma>)  | 
|
220  | 
case (Cons Y \<Gamma>)  | 
|
221  | 
thus "\<not> (\<exists>T.(X,T)\<in>set(Y#\<Gamma>))"  | 
|
222  | 
by (simp only: fresh_domain_cons, auto simp add: fresh_atm)  | 
|
223  | 
qed (simp)  | 
|
224  | 
}  | 
|
225  | 
ultimately show "T=S" by auto  | 
|
226  | 
qed  | 
|
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
227  | 
|
| 18628 | 228  | 
section {* Size and Capture-Avoiding Substitution for Types *}
 | 
| 18621 | 229  | 
|
| 21554 | 230  | 
nominal_primrec  | 
| 
29097
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
26966 
diff
changeset
 | 
231  | 
size_ty :: "ty \<Rightarrow> nat"  | 
| 
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
26966 
diff
changeset
 | 
232  | 
where  | 
| 21554 | 233  | 
"size_ty (Tvar X) = 1"  | 
| 
29097
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
26966 
diff
changeset
 | 
234  | 
| "size_ty (Top) = 1"  | 
| 
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
26966 
diff
changeset
 | 
235  | 
| "size_ty (T1 \<rightarrow> T2) = (size_ty T1) + (size_ty T2) + 1"  | 
| 
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
26966 
diff
changeset
 | 
236  | 
| "X\<sharp>T1 \<Longrightarrow> size_ty (\<forall>[X<:T1].T2) = (size_ty T1) + (size_ty T2) + 1"  | 
| 
22418
 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 
urbanc 
parents: 
21554 
diff
changeset
 | 
237  | 
apply (finite_guess)+  | 
| 21554 | 238  | 
apply (rule TrueI)+  | 
| 
22418
 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 
urbanc 
parents: 
21554 
diff
changeset
 | 
239  | 
apply (simp add: fresh_nat)  | 
| 
 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 
urbanc 
parents: 
21554 
diff
changeset
 | 
240  | 
apply (fresh_guess)+  | 
| 21554 | 241  | 
done  | 
| 
20395
 
9a60e3151244
added definition for size and substitution using the recursion
 
urbanc 
parents: 
19972 
diff
changeset
 | 
242  | 
|
| 21554 | 243  | 
nominal_primrec  | 
| 
29097
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
26966 
diff
changeset
 | 
244  | 
  subst_ty :: "ty \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> ty" ("_[_:=_]\<^isub>t\<^isub>y" [100,100,100] 100)
 | 
| 
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
26966 
diff
changeset
 | 
245  | 
where  | 
| 21554 | 246  | 
"(Tvar X)[Y:=T]\<^isub>t\<^isub>y= (if X=Y then T else (Tvar X))"  | 
| 
29097
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
26966 
diff
changeset
 | 
247  | 
| "(Top)[Y:=T]\<^isub>t\<^isub>y = Top"  | 
| 
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
26966 
diff
changeset
 | 
248  | 
| "(T\<^isub>1 \<rightarrow> T\<^isub>2)[Y:=T]\<^isub>t\<^isub>y = (T\<^isub>1[Y:=T]\<^isub>t\<^isub>y) \<rightarrow> (T\<^isub>2[Y:=T]\<^isub>t\<^isub>y)"  | 
| 
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
26966 
diff
changeset
 | 
249  | 
| "\<lbrakk>X\<sharp>(Y,T); X\<sharp>T\<^isub>1\<rbrakk> \<Longrightarrow> (\<forall>[X<:T\<^isub>1].T\<^isub>2)[Y:=T]\<^isub>t\<^isub>y = (\<forall>[X<:(T\<^isub>1[Y:=T]\<^isub>t\<^isub>y)].(T\<^isub>2[Y:=T]\<^isub>t\<^isub>y))"  | 
| 
22418
 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 
urbanc 
parents: 
21554 
diff
changeset
 | 
250  | 
apply (finite_guess)+  | 
| 21554 | 251  | 
apply (rule TrueI)+  | 
252  | 
apply (simp add: abs_fresh)  | 
|
| 
22418
 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 
urbanc 
parents: 
21554 
diff
changeset
 | 
253  | 
apply (fresh_guess)+  | 
| 21554 | 254  | 
done  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
255  | 
|
| 22537 | 256  | 
consts  | 
257  | 
  subst_tyc :: "ty_context \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> ty_context" ("_[_:=_]\<^isub>t\<^isub>y\<^isub>c" [100,100,100] 100)
 | 
|
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
258  | 
primrec  | 
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
259  | 
"([])[Y:=T]\<^isub>t\<^isub>y\<^isub>c= []"  | 
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
260  | 
"(XT#\<Gamma>)[Y:=T]\<^isub>t\<^isub>y\<^isub>c = (fst XT,(snd XT)[Y:=T]\<^isub>t\<^isub>y)#(\<Gamma>[Y:=T]\<^isub>t\<^isub>y\<^isub>c)"  | 
| 18246 | 261  | 
|
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
262  | 
section {* Subtyping-Relation *}
 | 
| 18246 | 263  | 
|
| 18650 | 264  | 
text {* The definition for the subtyping-relation follows quite closely what is written 
 | 
265  | 
in the POPLmark-paper, except for the premises dealing with well-formed contexts and  | 
|
266  | 
  the freshness constraint @{term "X\<sharp>\<Gamma>"} in the @{text "S_Forall"}-rule. (The freshness
 | 
|
267  | 
  constraint is specific to the \emph{nominal approach}. Note, however, that the constraint
 | 
|
268  | 
  does \emph{not} make the subtyping-relation ``partial"\ldots because we work over
 | 
|
269  | 
$\alpha$-equivalence classes.) *}  | 
|
| 18628 | 270  | 
|
| 23760 | 271  | 
inductive  | 
| 22537 | 272  | 
  subtype_of :: "ty_context \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool"   ("_\<turnstile>_<:_" [100,100,100] 100)
 | 
| 22436 | 273  | 
where  | 
274  | 
S_Top[intro]: "\<lbrakk>\<turnstile> \<Gamma> ok; S closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> S <: Top"  | 
|
275  | 
| S_Var[intro]: "\<lbrakk>(X,S) \<in> set \<Gamma>; \<Gamma> \<turnstile> S <: T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (Tvar X) <: T"  | 
|
276  | 
| S_Refl[intro]: "\<lbrakk>\<turnstile> \<Gamma> ok; X \<in> domain \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> Tvar X <: Tvar X"  | 
|
277  | 
| S_Arrow[intro]: "\<lbrakk>\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1; \<Gamma> \<turnstile> S\<^isub>2 <: T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (S\<^isub>1 \<rightarrow> S\<^isub>2) <: (T\<^isub>1 \<rightarrow> T\<^isub>2)"  | 
|
| 22537 | 278  | 
| S_Forall[intro]: "\<lbrakk>\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1; X\<sharp>\<Gamma>; ((X,T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> \<forall>[X<:S\<^isub>1].S\<^isub>2 <: \<forall>[X<:T\<^isub>1].T\<^isub>2"  | 
279  | 
||
280  | 
lemma subtype_implies_ok:  | 
|
281  | 
fixes X::"tyvrs"  | 
|
282  | 
assumes a: "\<Gamma> \<turnstile> S <: T"  | 
|
283  | 
shows "\<turnstile> \<Gamma> ok"  | 
|
284  | 
using a by (induct) (auto)  | 
|
| 18246 | 285  | 
|
286  | 
lemma subtype_implies_closed:  | 
|
287  | 
assumes a: "\<Gamma> \<turnstile> S <: T"  | 
|
288  | 
shows "S closed_in \<Gamma> \<and> T closed_in \<Gamma>"  | 
|
289  | 
using a  | 
|
290  | 
proof (induct)  | 
|
| 22436 | 291  | 
case (S_Top \<Gamma> S)  | 
| 18424 | 292  | 
have "Top closed_in \<Gamma>" by (simp add: closed_in_def ty.supp)  | 
| 18246 | 293  | 
moreover  | 
294  | 
have "S closed_in \<Gamma>" by fact  | 
|
295  | 
ultimately show "S closed_in \<Gamma> \<and> Top closed_in \<Gamma>" by simp  | 
|
296  | 
next  | 
|
| 22436 | 297  | 
case (S_Var X S \<Gamma> T)  | 
| 18246 | 298  | 
have "(X,S)\<in>set \<Gamma>" by fact  | 
| 18424 | 299  | 
hence "X \<in> domain \<Gamma>" by (rule domain_inclusion)  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
300  | 
hence "(Tvar X) closed_in \<Gamma>" by (simp add: closed_in_def ty.supp supp_atm)  | 
| 18246 | 301  | 
moreover  | 
302  | 
have "S closed_in \<Gamma> \<and> T closed_in \<Gamma>" by fact  | 
|
303  | 
hence "T closed_in \<Gamma>" by force  | 
|
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
304  | 
ultimately show "(Tvar X) closed_in \<Gamma> \<and> T closed_in \<Gamma>" by simp  | 
| 18424 | 305  | 
qed (auto simp add: closed_in_def ty.supp supp_atm abs_supp)  | 
| 18246 | 306  | 
|
307  | 
lemma subtype_implies_fresh:  | 
|
308  | 
fixes X::"tyvrs"  | 
|
309  | 
assumes a1: "\<Gamma> \<turnstile> S <: T"  | 
|
310  | 
and a2: "X\<sharp>\<Gamma>"  | 
|
| 18424 | 311  | 
shows "X\<sharp>S \<and> X\<sharp>T"  | 
| 18246 | 312  | 
proof -  | 
313  | 
from a1 have "\<turnstile> \<Gamma> ok" by (rule subtype_implies_ok)  | 
|
| 18621 | 314  | 
with a2 have "X\<sharp>domain(\<Gamma>)" by (simp add: fresh_domain)  | 
| 18424 | 315  | 
moreover  | 
| 18246 | 316  | 
from a1 have "S closed_in \<Gamma> \<and> T closed_in \<Gamma>" by (rule subtype_implies_closed)  | 
| 18424 | 317  | 
hence "supp S \<subseteq> ((supp (domain \<Gamma>))::tyvrs set)"  | 
| 18621 | 318  | 
and "supp T \<subseteq> ((supp (domain \<Gamma>))::tyvrs set)" by (simp_all add: domain_supp closed_in_def)  | 
| 18424 | 319  | 
ultimately show "X\<sharp>S \<and> X\<sharp>T" by (force simp add: supp_prod fresh_def)  | 
| 18246 | 320  | 
qed  | 
321  | 
||
| 
22730
 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 
berghofe 
parents: 
22542 
diff
changeset
 | 
322  | 
equivariance subtype_of  | 
| 
 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 
berghofe 
parents: 
22542 
diff
changeset
 | 
323  | 
|
| 22537 | 324  | 
nominal_inductive subtype_of  | 
325  | 
by (simp_all add: abs_fresh subtype_implies_fresh)  | 
|
| 18628 | 326  | 
|
| 22537 | 327  | 
thm subtype_of.strong_induct  | 
| 18246 | 328  | 
|
| 18621 | 329  | 
section {* Reflexivity of Subtyping *}
 | 
| 18246 | 330  | 
|
331  | 
lemma subtype_reflexivity:  | 
|
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
332  | 
assumes a: "\<turnstile> \<Gamma> ok"  | 
| 18424 | 333  | 
and b: "T closed_in \<Gamma>"  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
334  | 
shows "\<Gamma> \<turnstile> T <: T"  | 
| 
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
335  | 
using a b  | 
| 
26966
 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 
urbanc 
parents: 
26091 
diff
changeset
 | 
336  | 
proof(nominal_induct T avoiding: \<Gamma> rule: ty.strong_induct)  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
337  | 
case (Forall X T\<^isub>1 T\<^isub>2)  | 
| 18747 | 338  | 
have ih_T\<^isub>1: "\<And>\<Gamma>. \<lbrakk>\<turnstile> \<Gamma> ok; T\<^isub>1 closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> T\<^isub>1 <: T\<^isub>1" by fact  | 
339  | 
have ih_T\<^isub>2: "\<And>\<Gamma>. \<lbrakk>\<turnstile> \<Gamma> ok; T\<^isub>2 closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> T\<^isub>2 <: T\<^isub>2" by fact  | 
|
| 18424 | 340  | 
have fresh_cond: "X\<sharp>\<Gamma>" by fact  | 
| 18621 | 341  | 
hence fresh_domain: "X\<sharp>(domain \<Gamma>)" by (simp add: fresh_domain)  | 
| 18424 | 342  | 
have "(\<forall>[X<:T\<^isub>2].T\<^isub>1) closed_in \<Gamma>" by fact  | 
343  | 
hence closed\<^isub>T\<^isub>2: "T\<^isub>2 closed_in \<Gamma>" and closed\<^isub>T\<^isub>1: "T\<^isub>1 closed_in ((X,T\<^isub>2)#\<Gamma>)"  | 
|
344  | 
by (auto simp add: closed_in_def ty.supp abs_supp)  | 
|
345  | 
have ok: "\<turnstile> \<Gamma> ok" by fact  | 
|
| 18621 | 346  | 
hence ok': "\<turnstile> ((X,T\<^isub>2)#\<Gamma>) ok" using closed\<^isub>T\<^isub>2 fresh_domain by simp  | 
| 18424 | 347  | 
have "\<Gamma> \<turnstile> T\<^isub>2 <: T\<^isub>2" using ih_T\<^isub>2 closed\<^isub>T\<^isub>2 ok by simp  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
348  | 
moreover  | 
| 18424 | 349  | 
have "((X,T\<^isub>2)#\<Gamma>) \<turnstile> T\<^isub>1 <: T\<^isub>1" using ih_T\<^isub>1 closed\<^isub>T\<^isub>1 ok' by simp  | 
350  | 
ultimately show "\<Gamma> \<turnstile> \<forall>[X<:T\<^isub>2].T\<^isub>1 <: \<forall>[X<:T\<^isub>2].T\<^isub>1" using fresh_cond  | 
|
| 18621 | 351  | 
by (simp add: subtype_of.S_Forall)  | 
| 18246 | 352  | 
qed (auto simp add: closed_in_def ty.supp supp_atm)  | 
353  | 
||
| 18621 | 354  | 
lemma subtype_reflexivity_semiautomated:  | 
| 
18305
 
a780f9c1538b
changed everything until the interesting transitivity_narrowing
 
urbanc 
parents: 
18269 
diff
changeset
 | 
355  | 
assumes a: "\<turnstile> \<Gamma> ok"  | 
| 
 
a780f9c1538b
changed everything until the interesting transitivity_narrowing
 
urbanc 
parents: 
18269 
diff
changeset
 | 
356  | 
and b: "T closed_in \<Gamma>"  | 
| 
 
a780f9c1538b
changed everything until the interesting transitivity_narrowing
 
urbanc 
parents: 
18269 
diff
changeset
 | 
357  | 
shows "\<Gamma> \<turnstile> T <: T"  | 
| 
 
a780f9c1538b
changed everything until the interesting transitivity_narrowing
 
urbanc 
parents: 
18269 
diff
changeset
 | 
358  | 
using a b  | 
| 
26966
 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 
urbanc 
parents: 
26091 
diff
changeset
 | 
359  | 
apply(nominal_induct T avoiding: \<Gamma> rule: ty.strong_induct)  | 
| 18747 | 360  | 
apply(auto simp add: ty.supp abs_supp supp_atm closed_in_def)  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
361  | 
  --{* Too bad that this instantiation cannot be found automatically by
 | 
| 18621 | 362  | 
  \isakeyword{auto}; \isakeyword{blast} would find it if we had not used 
 | 
| 18628 | 363  | 
  an explicit definition for @{text "closed_in_def"}. *}
 | 
| 
18305
 
a780f9c1538b
changed everything until the interesting transitivity_narrowing
 
urbanc 
parents: 
18269 
diff
changeset
 | 
364  | 
apply(drule_tac x="(tyvrs, ty2)#\<Gamma>" in meta_spec)  | 
| 18747 | 365  | 
apply(force dest: fresh_domain simp add: closed_in_def)  | 
| 18246 | 366  | 
done  | 
367  | 
||
| 18747 | 368  | 
|
| 18628 | 369  | 
section {* Weakening *}
 | 
| 18246 | 370  | 
|
| 18628 | 371  | 
text {* In order to prove weakening we introduce the notion of a type-context extending 
 | 
372  | 
another. This generalization seems to make the proof for weakening to be  | 
|
373  | 
smoother than if we had strictly adhered to the version in the POPLmark-paper. *}  | 
|
| 18246 | 374  | 
|
375  | 
constdefs  | 
|
376  | 
  extends :: "ty_context \<Rightarrow> ty_context \<Rightarrow> bool" ("_ extends _" [100,100] 100)
 | 
|
377  | 
"\<Delta> extends \<Gamma> \<equiv> \<forall>X Q. (X,Q)\<in>set \<Gamma> \<longrightarrow> (X,Q)\<in>set \<Delta>"  | 
|
378  | 
||
379  | 
lemma extends_domain:  | 
|
380  | 
assumes a: "\<Delta> extends \<Gamma>"  | 
|
381  | 
shows "domain \<Gamma> \<subseteq> domain \<Delta>"  | 
|
382  | 
using a  | 
|
383  | 
apply (auto simp add: extends_def)  | 
|
384  | 
apply (drule domain_existence)  | 
|
385  | 
apply (force simp add: domain_inclusion)  | 
|
386  | 
done  | 
|
387  | 
||
388  | 
lemma extends_closed:  | 
|
389  | 
assumes a1: "T closed_in \<Gamma>"  | 
|
390  | 
and a2: "\<Delta> extends \<Gamma>"  | 
|
391  | 
shows "T closed_in \<Delta>"  | 
|
392  | 
using a1 a2  | 
|
393  | 
by (auto dest: extends_domain simp add: closed_in_def)  | 
|
394  | 
||
| 18424 | 395  | 
lemma extends_memb:  | 
396  | 
assumes a: "\<Delta> extends \<Gamma>"  | 
|
397  | 
and b: "(X,T) \<in> set \<Gamma>"  | 
|
398  | 
shows "(X,T) \<in> set \<Delta>"  | 
|
399  | 
using a b by (simp add: extends_def)  | 
|
400  | 
||
| 18246 | 401  | 
lemma weakening:  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
402  | 
assumes a: "\<Gamma> \<turnstile> S <: T"  | 
| 18424 | 403  | 
and b: "\<turnstile> \<Delta> ok"  | 
404  | 
and c: "\<Delta> extends \<Gamma>"  | 
|
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
405  | 
shows "\<Delta> \<turnstile> S <: T"  | 
| 
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
406  | 
using a b c  | 
| 22537 | 407  | 
proof (nominal_induct \<Gamma> S T avoiding: \<Delta> rule: subtype_of.strong_induct)  | 
| 
18305
 
a780f9c1538b
changed everything until the interesting transitivity_narrowing
 
urbanc 
parents: 
18269 
diff
changeset
 | 
408  | 
case (S_Top \<Gamma> S)  | 
| 18246 | 409  | 
have lh_drv_prem: "S closed_in \<Gamma>" by fact  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
410  | 
have "\<turnstile> \<Delta> ok" by fact  | 
| 
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
411  | 
moreover  | 
| 
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
412  | 
have "\<Delta> extends \<Gamma>" by fact  | 
| 18424 | 413  | 
hence "S closed_in \<Delta>" using lh_drv_prem by (simp only: extends_closed)  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
414  | 
ultimately show "\<Delta> \<turnstile> S <: Top" by force  | 
| 18246 | 415  | 
next  | 
| 22537 | 416  | 
case (S_Var X S \<Gamma> T)  | 
| 18246 | 417  | 
have lh_drv_prem: "(X,S) \<in> set \<Gamma>" by fact  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
418  | 
have ih: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends \<Gamma> \<Longrightarrow> \<Delta> \<turnstile> S <: T" by fact  | 
| 
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
419  | 
have ok: "\<turnstile> \<Delta> ok" by fact  | 
| 
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
420  | 
have extends: "\<Delta> extends \<Gamma>" by fact  | 
| 18424 | 421  | 
have "(X,S) \<in> set \<Delta>" using lh_drv_prem extends by (simp only: extends_memb)  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
422  | 
moreover  | 
| 
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
423  | 
have "\<Delta> \<turnstile> S <: T" using ok extends ih by simp  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
424  | 
ultimately show "\<Delta> \<turnstile> Tvar X <: T" using ok by force  | 
| 18246 | 425  | 
next  | 
| 
18305
 
a780f9c1538b
changed everything until the interesting transitivity_narrowing
 
urbanc 
parents: 
18269 
diff
changeset
 | 
426  | 
case (S_Refl \<Gamma> X)  | 
| 18246 | 427  | 
have lh_drv_prem: "X \<in> domain \<Gamma>" by fact  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
428  | 
have "\<turnstile> \<Delta> ok" by fact  | 
| 
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
429  | 
moreover  | 
| 
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
430  | 
have "\<Delta> extends \<Gamma>" by fact  | 
| 
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
431  | 
hence "X \<in> domain \<Delta>" using lh_drv_prem by (force dest: extends_domain)  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
432  | 
ultimately show "\<Delta> \<turnstile> Tvar X <: Tvar X" by force  | 
| 18246 | 433  | 
next  | 
| 22537 | 434  | 
case (S_Arrow \<Gamma> T\<^isub>1 S\<^isub>1 S\<^isub>2 T\<^isub>2) thus "\<Delta> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T\<^isub>1 \<rightarrow> T\<^isub>2" by blast  | 
| 18246 | 435  | 
next  | 
| 22537 | 436  | 
case (S_Forall \<Gamma> T\<^isub>1 S\<^isub>1 X S\<^isub>2 T\<^isub>2)  | 
| 18424 | 437  | 
have fresh_cond: "X\<sharp>\<Delta>" by fact  | 
| 18621 | 438  | 
hence fresh_domain: "X\<sharp>(domain \<Delta>)" by (simp add: fresh_domain)  | 
| 18424 | 439  | 
have ih\<^isub>1: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends \<Gamma> \<Longrightarrow> \<Delta> \<turnstile> T\<^isub>1 <: S\<^isub>1" by fact  | 
440  | 
have ih\<^isub>2: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends ((X,T\<^isub>1)#\<Gamma>) \<Longrightarrow> \<Delta> \<turnstile> S\<^isub>2 <: T\<^isub>2" by fact  | 
|
441  | 
have lh_drv_prem: "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1" by fact  | 
|
442  | 
hence closed\<^isub>T\<^isub>1: "T\<^isub>1 closed_in \<Gamma>" by (simp add: subtype_implies_closed)  | 
|
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
443  | 
have ok: "\<turnstile> \<Delta> ok" by fact  | 
| 18424 | 444  | 
have ext: "\<Delta> extends \<Gamma>" by fact  | 
445  | 
have "T\<^isub>1 closed_in \<Delta>" using ext closed\<^isub>T\<^isub>1 by (simp only: extends_closed)  | 
|
| 18621 | 446  | 
hence "\<turnstile> ((X,T\<^isub>1)#\<Delta>) ok" using fresh_domain ok by force  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
447  | 
moreover  | 
| 18424 | 448  | 
have "((X,T\<^isub>1)#\<Delta>) extends ((X,T\<^isub>1)#\<Gamma>)" using ext by (force simp add: extends_def)  | 
449  | 
ultimately have "((X,T\<^isub>1)#\<Delta>) \<turnstile> S\<^isub>2 <: T\<^isub>2" using ih\<^isub>2 by simp  | 
|
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
450  | 
moreover  | 
| 18424 | 451  | 
have "\<Delta> \<turnstile> T\<^isub>1 <: S\<^isub>1" using ok ext ih\<^isub>1 by simp  | 
| 18621 | 452  | 
ultimately show "\<Delta> \<turnstile> \<forall>[X<:S\<^isub>1].S\<^isub>2 <: \<forall>[X<:T\<^isub>1].T\<^isub>2" using ok by (force intro: S_Forall)  | 
| 18246 | 453  | 
qed  | 
454  | 
||
| 18650 | 455  | 
text {* In fact all ``non-binding" cases can be solved automatically: *}
 | 
| 18246 | 456  | 
|
| 18628 | 457  | 
lemma weakening_more_automated:  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
458  | 
assumes a: "\<Gamma> \<turnstile> S <: T"  | 
| 18424 | 459  | 
and b: "\<turnstile> \<Delta> ok"  | 
460  | 
and c: "\<Delta> extends \<Gamma>"  | 
|
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
461  | 
shows "\<Delta> \<turnstile> S <: T"  | 
| 
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
462  | 
using a b c  | 
| 22537 | 463  | 
proof (nominal_induct \<Gamma> S T avoiding: \<Delta> rule: subtype_of.strong_induct)  | 
464  | 
case (S_Forall \<Gamma> T\<^isub>1 S\<^isub>1 X S\<^isub>2 T\<^isub>2)  | 
|
| 18424 | 465  | 
have fresh_cond: "X\<sharp>\<Delta>" by fact  | 
| 18621 | 466  | 
hence fresh_domain: "X\<sharp>(domain \<Delta>)" by (simp add: fresh_domain)  | 
| 18424 | 467  | 
have ih\<^isub>1: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends \<Gamma> \<Longrightarrow> \<Delta> \<turnstile> T\<^isub>1 <: S\<^isub>1" by fact  | 
468  | 
have ih\<^isub>2: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends ((X,T\<^isub>1)#\<Gamma>) \<Longrightarrow> \<Delta> \<turnstile> S\<^isub>2 <: T\<^isub>2" by fact  | 
|
469  | 
have lh_drv_prem: "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1" by fact  | 
|
470  | 
hence closed\<^isub>T\<^isub>1: "T\<^isub>1 closed_in \<Gamma>" by (simp add: subtype_implies_closed)  | 
|
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
471  | 
have ok: "\<turnstile> \<Delta> ok" by fact  | 
| 18424 | 472  | 
have ext: "\<Delta> extends \<Gamma>" by fact  | 
473  | 
have "T\<^isub>1 closed_in \<Delta>" using ext closed\<^isub>T\<^isub>1 by (simp only: extends_closed)  | 
|
| 18621 | 474  | 
hence "\<turnstile> ((X,T\<^isub>1)#\<Delta>) ok" using fresh_domain ok by force  | 
| 18628 | 475  | 
moreover  | 
| 18424 | 476  | 
have "((X,T\<^isub>1)#\<Delta>) extends ((X,T\<^isub>1)#\<Gamma>)" using ext by (force simp add: extends_def)  | 
477  | 
ultimately have "((X,T\<^isub>1)#\<Delta>) \<turnstile> S\<^isub>2 <: T\<^isub>2" using ih\<^isub>2 by simp  | 
|
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
478  | 
moreover  | 
| 18424 | 479  | 
have "\<Delta> \<turnstile> T\<^isub>1 <: S\<^isub>1" using ok ext ih\<^isub>1 by simp  | 
| 18621 | 480  | 
ultimately show "\<Delta> \<turnstile> \<forall>[X<:S\<^isub>1].S\<^isub>2 <: \<forall>[X<:T\<^isub>1].T\<^isub>2" using ok by (force intro: S_Forall)  | 
| 18424 | 481  | 
qed (blast intro: extends_closed extends_memb dest: extends_domain)+  | 
| 18246 | 482  | 
|
| 18628 | 483  | 
section {* Transitivity and Narrowing *}
 | 
484  | 
||
| 18650 | 485  | 
text {* Some inversion lemmas that are needed in the transitivity and narrowing proof.*}
 | 
486  | 
||
487  | 
lemma S_TopE:  | 
|
488  | 
assumes a: "\<Gamma> \<turnstile> Top <: T"  | 
|
489  | 
shows "T = Top"  | 
|
490  | 
using a by (cases, auto)  | 
|
491  | 
||
492  | 
lemma S_ArrowE_left:  | 
|
493  | 
assumes a: "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T"  | 
|
494  | 
shows "T = Top \<or> (\<exists>T\<^isub>1 T\<^isub>2. T = T\<^isub>1 \<rightarrow> T\<^isub>2 \<and> \<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1 \<and> \<Gamma> \<turnstile> S\<^isub>2 <: T\<^isub>2)"  | 
|
495  | 
using a by (cases, auto simp add: ty.inject)  | 
|
496  | 
||
497  | 
lemma S_ForallE_left:  | 
|
498  | 
shows "\<lbrakk>\<Gamma> \<turnstile> \<forall>[X<:S\<^isub>1].S\<^isub>2 <: T; X\<sharp>\<Gamma>; X\<sharp>S\<^isub>1\<rbrakk>  | 
|
499  | 
\<Longrightarrow> T = Top \<or> (\<exists>T\<^isub>1 T\<^isub>2. T = \<forall>[X<:T\<^isub>1].T\<^isub>2 \<and> \<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1 \<and> ((X,T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2)"  | 
|
500  | 
apply(frule subtype_implies_ok)  | 
|
| 23760 | 501  | 
apply(ind_cases "\<Gamma> \<turnstile> \<forall>[X<:S\<^isub>1].S\<^isub>2 <: T")  | 
| 18650 | 502  | 
apply(auto simp add: ty.inject alpha)  | 
503  | 
apply(rule_tac x="[(X,Xa)]\<bullet>T\<^isub>2" in exI)  | 
|
504  | 
apply(rule conjI)  | 
|
505  | 
apply(rule sym)  | 
|
506  | 
apply(rule pt_bij2[OF pt_tyvrs_inst, OF at_tyvrs_inst])  | 
|
507  | 
apply(rule pt_tyvrs3)  | 
|
508  | 
apply(simp)  | 
|
509  | 
apply(rule at_ds5[OF at_tyvrs_inst])  | 
|
510  | 
apply(rule conjI)  | 
|
511  | 
apply(simp add: pt_fresh_left[OF pt_tyvrs_inst, OF at_tyvrs_inst] calc_atm)  | 
|
512  | 
apply(drule_tac \<Gamma>="((Xa,T\<^isub>1)#\<Gamma>)" in subtype_implies_closed)+  | 
|
513  | 
apply(simp add: closed_in_def)  | 
|
514  | 
apply(drule fresh_domain)+  | 
|
515  | 
apply(simp add: fresh_def)  | 
|
516  | 
apply(subgoal_tac "X \<notin> (insert Xa (domain \<Gamma>))")(*A*)  | 
|
517  | 
apply(force)  | 
|
518  | 
(*A*)apply(simp add: at_fin_set_supp[OF at_tyvrs_inst, OF finite_domain])  | 
|
519  | 
(* 2nd conjunct *)apply(frule_tac X="X" in subtype_implies_fresh)  | 
|
520  | 
apply(assumption)  | 
|
521  | 
apply(drule_tac X="Xa" in subtype_implies_fresh)  | 
|
522  | 
apply(assumption)  | 
|
523  | 
apply(simp add: fresh_prod)  | 
|
| 22542 | 524  | 
apply(drule_tac pi="[(X,Xa)]" in subtype_of.eqvt(2))  | 
| 18650 | 525  | 
apply(simp add: calc_atm)  | 
526  | 
apply(simp add: pt_fresh_fresh[OF pt_tyvrs_inst, OF at_tyvrs_inst])  | 
|
527  | 
done  | 
|
528  | 
||
529  | 
text {* Next we prove the transitivity and narrowing for the subtyping-relation. 
 | 
|
| 18621 | 530  | 
The POPLmark-paper says the following:  | 
531  | 
||
| 18650 | 532  | 
\begin{quote}
 | 
| 18621 | 533  | 
\begin{lemma}[Transitivity and Narrowing] \
 | 
534  | 
\begin{enumerate}
 | 
|
535  | 
\item If @{term "\<Gamma> \<turnstile> S<:Q"} and @{term "\<Gamma> \<turnstile> Q<:T"}, then @{term "\<Gamma> \<turnstile> S<:T"}.
 | 
|
536  | 
\item If @{text "\<Gamma>,X<:Q,\<Delta> \<turnstile> M<:N"} and @{term "\<Gamma> \<turnstile> P<:Q"} then @{text "\<Gamma>,X<:P,\<Delta> \<turnstile> M<:N"}.
 | 
|
537  | 
\end{enumerate}
 | 
|
538  | 
\end{lemma}
 | 
|
539  | 
||
540  | 
The two parts are proved simultaneously, by induction on the size  | 
|
541  | 
of @{term "Q"}.  The argument for part (2) assumes that part (1) has 
 | 
|
542  | 
been established already for the @{term "Q"} in question; part (1) uses 
 | 
|
543  | 
part (2) only for strictly smaller @{term "Q"}.
 | 
|
| 18650 | 544  | 
\end{quote}
 | 
| 18621 | 545  | 
|
546  | 
For the induction on the size of @{term "Q"}, we use the induction-rule 
 | 
|
547  | 
@{text "measure_induct_rule"}:
 | 
|
548  | 
||
549  | 
\begin{center}
 | 
|
550  | 
@{thm measure_induct_rule[of "size_ty",no_vars]}
 | 
|
551  | 
\end{center}
 | 
|
| 18410 | 552  | 
|
| 18628 | 553  | 
That means in order to show a property @{term "P a"} for all @{term "a"}, 
 | 
| 18650 | 554  | 
the induct-rule requires to prove that for all @{term x} @{term "P x"} holds using the 
 | 
| 18621 | 555  | 
assumption that for all @{term y} whose size is strictly smaller than 
 | 
556  | 
that of @{term x} the property @{term "P y"} holds. *}
 | 
|
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
557  | 
|
| 18621 | 558  | 
lemma  | 
559  | 
shows trans: "\<Gamma>\<turnstile>S<:Q \<Longrightarrow> \<Gamma>\<turnstile>Q<:T \<Longrightarrow> \<Gamma>\<turnstile>S<:T"  | 
|
560  | 
and narrow: "(\<Delta>@[(X,Q)]@\<Gamma>)\<turnstile>M<:N \<Longrightarrow> \<Gamma>\<turnstile>P<:Q \<Longrightarrow> (\<Delta>@[(X,P)]@\<Gamma>)\<turnstile>M<:N"  | 
|
| 20503 | 561  | 
proof (induct Q arbitrary: \<Gamma> S T \<Delta> X P M N taking: "size_ty" rule: measure_induct_rule)  | 
| 18621 | 562  | 
case (less Q)  | 
563  | 
    --{* \begin{minipage}[t]{0.9\textwidth}
 | 
|
564  | 
First we mention the induction hypotheses of the outer induction for later  | 
|
565  | 
    reference:\end{minipage}*}
 | 
|
566  | 
have IH_trans:  | 
|
567  | 
"\<And>Q' \<Gamma> S T. \<lbrakk>size_ty Q' < size_ty Q; \<Gamma>\<turnstile>S<:Q'; \<Gamma>\<turnstile>Q'<:T\<rbrakk> \<Longrightarrow> \<Gamma>\<turnstile>S<:T" by fact  | 
|
568  | 
have IH_narrow:  | 
|
569  | 
"\<And>Q' \<Delta> \<Gamma> X M N P. \<lbrakk>size_ty Q' < size_ty Q; (\<Delta>@[(X,Q')]@\<Gamma>)\<turnstile>M<:N; \<Gamma>\<turnstile>P<:Q'\<rbrakk>  | 
|
570  | 
\<Longrightarrow> (\<Delta>@[(X,P)]@\<Gamma>)\<turnstile>M<:N" by fact  | 
|
571  | 
    --{* \begin{minipage}[t]{0.9\textwidth}
 | 
|
572  | 
We proceed with the transitivity proof as an auxiliary lemma, because it needs  | 
|
573  | 
    to be referenced in the narrowing proof.\end{minipage}*}
 | 
|
574  | 
have transitivity_aux:  | 
|
575  | 
"\<And>\<Gamma> S T. \<lbrakk>\<Gamma> \<turnstile> S <: Q; \<Gamma> \<turnstile> Q <: T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> S <: T"  | 
|
| 18246 | 576  | 
proof -  | 
| 18424 | 577  | 
fix \<Gamma>' S' T  | 
| 18621 | 578  | 
    assume "\<Gamma>' \<turnstile> S' <: Q" --{* left-hand derivation *}
 | 
579  | 
      and  "\<Gamma>' \<turnstile> Q <: T"  --{* right-hand derivation *}
 | 
|
580  | 
thus "\<Gamma>' \<turnstile> S' <: T"  | 
|
| 22537 | 581  | 
proof (nominal_induct \<Gamma>' S' Q\<equiv>Q rule: subtype_of.strong_induct)  | 
| 18424 | 582  | 
case (S_Top \<Gamma> S)  | 
| 18621 | 583  | 
	--{* \begin{minipage}[t]{0.9\textwidth}
 | 
584  | 
	In this case the left-hand derivation is @{term "\<Gamma> \<turnstile> S <: Top"}, giving
 | 
|
585  | 
	us @{term "\<turnstile> \<Gamma> ok"} and @{term "S closed_in \<Gamma>"}. This case is straightforward, 
 | 
|
586  | 
	because the right-hand derivation must be of the form @{term "\<Gamma> \<turnstile> Top <: Top"} 
 | 
|
587  | 
	giving us the equation @{term "T = Top"}.\end{minipage}*}
 | 
|
| 18424 | 588  | 
hence rh_drv: "\<Gamma> \<turnstile> Top <: T" by simp  | 
589  | 
hence T_inst: "T = Top" by (simp add: S_TopE)  | 
|
| 18621 | 590  | 
have "\<turnstile> \<Gamma> ok"  | 
| 23393 | 591  | 
and "S closed_in \<Gamma>" by fact+  | 
| 18621 | 592  | 
hence "\<Gamma> \<turnstile> S <: Top" by (simp add: subtype_of.S_Top)  | 
| 18424 | 593  | 
thus "\<Gamma> \<turnstile> S <: T" using T_inst by simp  | 
| 18246 | 594  | 
next  | 
| 22537 | 595  | 
case (S_Var Y U \<Gamma>)  | 
| 18621 | 596  | 
	-- {* \begin{minipage}[t]{0.9\textwidth}
 | 
597  | 
	In this case the left-hand derivation is @{term "\<Gamma> \<turnstile> Tvar Y <: Q"} 
 | 
|
598  | 
	with @{term "S = Tvar Y"}. We have therefore @{term "(Y,U)"} 
 | 
|
| 18650 | 599  | 
	is in @{term "\<Gamma>"} and by inner induction hypothesis that @{term "\<Gamma> \<turnstile> U <: T"}. 
 | 
| 18621 | 600  | 
	By @{text "S_Var"} follows @{term "\<Gamma> \<turnstile> Tvar Y <: T"}.\end{minipage}*}
 | 
| 18424 | 601  | 
hence IH_inner: "\<Gamma> \<turnstile> U <: T" by simp  | 
| 18621 | 602  | 
have "(Y,U) \<in> set \<Gamma>" by fact  | 
603  | 
with IH_inner show "\<Gamma> \<turnstile> Tvar Y <: T" by (simp add: subtype_of.S_Var)  | 
|
| 18246 | 604  | 
next  | 
| 18424 | 605  | 
case (S_Refl \<Gamma> X)  | 
| 18621 | 606  | 
	--{* \begin{minipage}[t]{0.9\textwidth}
 | 
607  | 
        In this case the left-hand derivation is @{term "\<Gamma>\<turnstile>(Tvar X) <: (Tvar X)"} with
 | 
|
608  | 
        @{term "Q=Tvar X"}. The goal then follows immediately from the right-hand 
 | 
|
609  | 
	derivation.\end{minipage}*}
 | 
|
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
610  | 
thus "\<Gamma> \<turnstile> Tvar X <: T" by simp  | 
| 18246 | 611  | 
next  | 
| 22537 | 612  | 
case (S_Arrow \<Gamma> Q\<^isub>1 S\<^isub>1 S\<^isub>2 Q\<^isub>2)  | 
| 18621 | 613  | 
	--{* \begin{minipage}[t]{0.9\textwidth}
 | 
614  | 
	In this case the left-hand derivation is @{term "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: Q\<^isub>1 \<rightarrow> Q\<^isub>2"} with
 | 
|
615  | 
        @{term "S\<^isub>1\<rightarrow>S\<^isub>2=S"} and @{term "Q\<^isub>1\<rightarrow>Q\<^isub>2=Q"}. We know that the @{text "size_ty"} of 
 | 
|
616  | 
	@{term Q\<^isub>1} and @{term Q\<^isub>2} is smaller than that of @{term Q};
 | 
|
617  | 
	so we can apply the outer induction hypotheses for @{term Q\<^isub>1} and @{term Q\<^isub>2}. 
 | 
|
618  | 
	We also have the sub-derivations  @{term "\<Gamma>\<turnstile>Q\<^isub>1<:S\<^isub>1"} and @{term "\<Gamma>\<turnstile>S\<^isub>2<:Q\<^isub>2"}.
 | 
|
| 18628 | 619  | 
	The right-hand derivation is @{term "\<Gamma> \<turnstile> Q\<^isub>1 \<rightarrow> Q\<^isub>2 <: T"}. There exist types 
 | 
| 18621 | 620  | 
	@{text "T\<^isub>1,T\<^isub>2"} such that @{term "T=Top \<or> T=T\<^isub>1\<rightarrow>T\<^isub>2"}. The @{term "Top"}-case is 
 | 
621  | 
	straightforward once we know @{term "(S\<^isub>1 \<rightarrow> S\<^isub>2) closed_in \<Gamma>"} and @{term "\<turnstile> \<Gamma> ok"}. 
 | 
|
622  | 
	In the other case we have the sub-derivations @{term "\<Gamma>\<turnstile>T\<^isub>1<:Q\<^isub>1"} and @{term "\<Gamma>\<turnstile>Q\<^isub>2<:T\<^isub>2"}. 
 | 
|
623  | 
	Using the outer induction hypothesis for transitivity we can derive @{term "\<Gamma>\<turnstile>T\<^isub>1<:S\<^isub>1"} 
 | 
|
624  | 
	and @{term "\<Gamma>\<turnstile>S\<^isub>2<:T\<^isub>2"}. By rule @{text "S_Arrow"} follows @{term "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T\<^isub>1 \<rightarrow> T\<^isub>2"},
 | 
|
625  | 
	which is @{term "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T\<^isub>"}.\end{minipage}*}
 | 
|
626  | 
hence rh_drv: "\<Gamma> \<turnstile> Q\<^isub>1 \<rightarrow> Q\<^isub>2 <: T" by simp  | 
|
627  | 
from `Q\<^isub>1 \<rightarrow> Q\<^isub>2 = Q`  | 
|
628  | 
have Q\<^isub>1\<^isub>2_less: "size_ty Q\<^isub>1 < size_ty Q" "size_ty Q\<^isub>2 < size_ty Q" by auto  | 
|
629  | 
have lh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> Q\<^isub>1 <: S\<^isub>1" by fact  | 
|
630  | 
have lh_drv_prm\<^isub>2: "\<Gamma> \<turnstile> S\<^isub>2 <: Q\<^isub>2" by fact  | 
|
631  | 
from rh_drv have "T=Top \<or> (\<exists>T\<^isub>1 T\<^isub>2. T=T\<^isub>1\<rightarrow>T\<^isub>2 \<and> \<Gamma>\<turnstile>T\<^isub>1<:Q\<^isub>1 \<and> \<Gamma>\<turnstile>Q\<^isub>2<:T\<^isub>2)"  | 
|
632  | 
by (simp add: S_ArrowE_left)  | 
|
633  | 
moreover  | 
|
634  | 
have "S\<^isub>1 closed_in \<Gamma>" and "S\<^isub>2 closed_in \<Gamma>"  | 
|
635  | 
using lh_drv_prm\<^isub>1 lh_drv_prm\<^isub>2 by (simp_all add: subtype_implies_closed)  | 
|
636  | 
hence "(S\<^isub>1 \<rightarrow> S\<^isub>2) closed_in \<Gamma>" by (simp add: closed_in_def ty.supp)  | 
|
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
637  | 
moreover  | 
| 18424 | 638  | 
have "\<turnstile> \<Gamma> ok" using rh_drv by (rule subtype_implies_ok)  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
639  | 
moreover  | 
| 18621 | 640  | 
      { assume "\<exists>T\<^isub>1 T\<^isub>2. T=T\<^isub>1\<rightarrow>T\<^isub>2 \<and> \<Gamma>\<turnstile>T\<^isub>1<:Q\<^isub>1 \<and> \<Gamma>\<turnstile>Q\<^isub>2<:T\<^isub>2"
 | 
641  | 
then obtain T\<^isub>1 T\<^isub>2  | 
|
642  | 
where T_inst: "T = T\<^isub>1 \<rightarrow> T\<^isub>2"  | 
|
643  | 
and rh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> T\<^isub>1 <: Q\<^isub>1"  | 
|
644  | 
and rh_drv_prm\<^isub>2: "\<Gamma> \<turnstile> Q\<^isub>2 <: T\<^isub>2" by force  | 
|
645  | 
from IH_trans[of "Q\<^isub>1"]  | 
|
646  | 
have "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1" using Q\<^isub>1\<^isub>2_less rh_drv_prm\<^isub>1 lh_drv_prm\<^isub>1 by simp  | 
|
| 18246 | 647  | 
moreover  | 
| 18621 | 648  | 
from IH_trans[of "Q\<^isub>2"]  | 
649  | 
have "\<Gamma> \<turnstile> S\<^isub>2 <: T\<^isub>2" using Q\<^isub>1\<^isub>2_less rh_drv_prm\<^isub>2 lh_drv_prm\<^isub>2 by simp  | 
|
650  | 
ultimately have "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T\<^isub>1 \<rightarrow> T\<^isub>2" by (simp add: subtype_of.S_Arrow)  | 
|
651  | 
hence "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T" using T_inst by simp  | 
|
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
652  | 
}  | 
| 18621 | 653  | 
ultimately show "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T" by blast  | 
| 18246 | 654  | 
next  | 
| 22537 | 655  | 
case (S_Forall \<Gamma> Q\<^isub>1 S\<^isub>1 X S\<^isub>2 Q\<^isub>2)  | 
| 18621 | 656  | 
	--{* \begin{minipage}[t]{0.9\textwidth}
 | 
657  | 
	In this case the left-hand derivation is @{text "\<Gamma>\<turnstile>\<forall>[X<:S\<^isub>1].S\<^isub>2 <: \<forall>[X<:Q\<^isub>1].Q\<^isub>2"} with 
 | 
|
658  | 
	@{text "\<forall>[X<:S\<^isub>1].S\<^isub>2=S"} and @{text "\<forall>[X<:Q\<^isub>1].Q\<^isub>2=Q"}. We therefore have the sub-derivations  
 | 
|
659  | 
	@{term "\<Gamma>\<turnstile>Q\<^isub>1<:S\<^isub>1"} and @{term "((X,Q\<^isub>1)#\<Gamma>)\<turnstile>S\<^isub>2<:Q\<^isub>2"}. Since @{term "X"} is a binder, we
 | 
|
660  | 
assume that it is sufficiently fresh; in particular we have the freshness conditions  | 
|
| 18650 | 661  | 
	@{term "X\<sharp>\<Gamma>"} and @{term "X\<sharp>Q\<^isub>1"} (these assumptions are provided by the strong 
 | 
662  | 
	induction-rule @{text "subtype_of_induct"}). We know that the @{text "size_ty"} of 
 | 
|
| 18621 | 663  | 
	@{term Q\<^isub>1} and @{term Q\<^isub>2} is smaller than that of @{term Q};
 | 
664  | 
	so we can apply the outer induction hypotheses for @{term Q\<^isub>1} and @{term Q\<^isub>2}. 
 | 
|
665  | 
	The right-hand derivation is @{text "\<Gamma> \<turnstile> \<forall>[X<:Q\<^isub>1].Q\<^isub>2 <: T"}. Since @{term "X\<sharp>\<Gamma>"} 
 | 
|
666  | 
	and @{term "X\<sharp>Q\<^isub>1"} there exists types @{text "T\<^isub>1,T\<^isub>2"} such that 
 | 
|
667  | 
	@{text "T=Top \<or> T=\<forall>[X<:T\<^isub>1].T\<^isub>2"}. The @{term "Top"}-case is straightforward once we know 
 | 
|
668  | 
	@{text "(\<forall>[X<:S\<^isub>1].S\<^isub>2) closed_in \<Gamma>"} and @{term "\<turnstile> \<Gamma> ok"}. In the other case we have 
 | 
|
| 18628 | 669  | 
	the sub-derivations @{term "\<Gamma>\<turnstile>T\<^isub>1<:Q\<^isub>1"} and @{term "((X,T\<^isub>1)#\<Gamma>)\<turnstile>Q\<^isub>2<:T\<^isub>2"}. Using the outer 
 | 
| 18621 | 670  | 
	induction hypothesis for transitivity we can derive @{term "\<Gamma>\<turnstile>T\<^isub>1<:S\<^isub>1"}. From the outer 
 | 
| 18628 | 671  | 
	induction for narrowing we get @{term "((X,T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: Q\<^isub>2"} and then using again 
 | 
672  | 
	induction for transitivity we obtain @{term "((X,T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2"}. By rule 
 | 
|
673  | 
	@{text "S_Forall"} and the freshness condition @{term "X\<sharp>\<Gamma>"} follows 
 | 
|
| 18650 | 674  | 
	@{text "\<Gamma> \<turnstile> \<forall>[X<:S\<^isub>1].S\<^isub>2 <: \<forall>[X<:T\<^isub>1].T\<^isub>2"}, which is @{text "\<Gamma> \<turnstile>  \<forall>[X<:S\<^isub>1].S\<^isub>2 <: T\<^isub>"}.
 | 
| 18628 | 675  | 
	\end{minipage}*}
 | 
| 18621 | 676  | 
hence rh_drv: "\<Gamma> \<turnstile> \<forall>[X<:Q\<^isub>1].Q\<^isub>2 <: T" by simp  | 
677  | 
have lh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> Q\<^isub>1 <: S\<^isub>1" by fact  | 
|
678  | 
have lh_drv_prm\<^isub>2: "((X,Q\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: Q\<^isub>2" by fact  | 
|
| 22537 | 679  | 
have "X\<sharp>\<Gamma>" by fact  | 
680  | 
then have fresh_cond: "X\<sharp>\<Gamma>" "X\<sharp>Q\<^isub>1" using lh_drv_prm\<^isub>1 by (simp_all add: subtype_implies_fresh)  | 
|
| 18621 | 681  | 
from `\<forall>[X<:Q\<^isub>1].Q\<^isub>2 = Q`  | 
| 
20395
 
9a60e3151244
added definition for size and substitution using the recursion
 
urbanc 
parents: 
19972 
diff
changeset
 | 
682  | 
have Q\<^isub>1\<^isub>2_less: "size_ty Q\<^isub>1 < size_ty Q" "size_ty Q\<^isub>2 < size_ty Q " using fresh_cond by auto  | 
| 18621 | 683  | 
from rh_drv  | 
684  | 
have "T=Top \<or> (\<exists>T\<^isub>1 T\<^isub>2. T=\<forall>[X<:T\<^isub>1].T\<^isub>2 \<and> \<Gamma>\<turnstile>T\<^isub>1<:Q\<^isub>1 \<and> ((X,T\<^isub>1)#\<Gamma>)\<turnstile>Q\<^isub>2<:T\<^isub>2)"  | 
|
685  | 
using fresh_cond by (simp add: S_ForallE_left)  | 
|
686  | 
moreover  | 
|
687  | 
have "S\<^isub>1 closed_in \<Gamma>" and "S\<^isub>2 closed_in ((X,Q\<^isub>1)#\<Gamma>)"  | 
|
688  | 
using lh_drv_prm\<^isub>1 lh_drv_prm\<^isub>2 by (simp_all add: subtype_implies_closed)  | 
|
689  | 
hence "(\<forall>[X<:S\<^isub>1].S\<^isub>2) closed_in \<Gamma>" by (force simp add: closed_in_def ty.supp abs_supp)  | 
|
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
690  | 
moreover  | 
| 18424 | 691  | 
have "\<turnstile> \<Gamma> ok" using rh_drv by (rule subtype_implies_ok)  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
692  | 
moreover  | 
| 18621 | 693  | 
      { assume "\<exists>T\<^isub>1 T\<^isub>2. T=\<forall>[X<:T\<^isub>1].T\<^isub>2 \<and> \<Gamma>\<turnstile>T\<^isub>1<:Q\<^isub>1 \<and> ((X,T\<^isub>1)#\<Gamma>)\<turnstile>Q\<^isub>2<:T\<^isub>2"
 | 
694  | 
then obtain T\<^isub>1 T\<^isub>2  | 
|
695  | 
where T_inst: "T = \<forall>[X<:T\<^isub>1].T\<^isub>2"  | 
|
696  | 
and rh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> T\<^isub>1 <: Q\<^isub>1"  | 
|
697  | 
and rh_drv_prm\<^isub>2:"((X,T\<^isub>1)#\<Gamma>) \<turnstile> Q\<^isub>2 <: T\<^isub>2" by force  | 
|
698  | 
from IH_trans[of "Q\<^isub>1"]  | 
|
699  | 
have "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1" using lh_drv_prm\<^isub>1 rh_drv_prm\<^isub>1 Q\<^isub>1\<^isub>2_less by blast  | 
|
700  | 
moreover  | 
|
701  | 
from IH_narrow[of "Q\<^isub>1" "[]"]  | 
|
702  | 
have "((X,T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: Q\<^isub>2" using Q\<^isub>1\<^isub>2_less lh_drv_prm\<^isub>2 rh_drv_prm\<^isub>1 by simp  | 
|
703  | 
with IH_trans[of "Q\<^isub>2"]  | 
|
704  | 
have "((X,T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2" using Q\<^isub>1\<^isub>2_less rh_drv_prm\<^isub>2 by simp  | 
|
705  | 
ultimately have "\<Gamma> \<turnstile> \<forall>[X<:S\<^isub>1].S\<^isub>2 <: \<forall>[X<:T\<^isub>1].T\<^isub>2"  | 
|
706  | 
using fresh_cond by (simp add: subtype_of.S_Forall)  | 
|
707  | 
hence "\<Gamma> \<turnstile> \<forall>[X<:S\<^isub>1].S\<^isub>2 <: T" using T_inst by simp  | 
|
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
708  | 
}  | 
| 18621 | 709  | 
ultimately show "\<Gamma> \<turnstile> \<forall>[X<:S\<^isub>1].S\<^isub>2 <: T" by blast  | 
| 18246 | 710  | 
qed  | 
711  | 
qed  | 
|
712  | 
||
| 18621 | 713  | 
  { --{* The transitivity proof is now by the auxiliary lemma. *}
 | 
714  | 
case 1  | 
|
715  | 
have "\<Gamma> \<turnstile> S <: Q"  | 
|
| 23393 | 716  | 
and "\<Gamma> \<turnstile> Q <: T" by fact+  | 
| 18621 | 717  | 
thus "\<Gamma> \<turnstile> S <: T" by (rule transitivity_aux)  | 
718  | 
next  | 
|
719  | 
    --{* The narrowing proof proceeds by an induction over @{term "(\<Delta>@[(X,Q)]@\<Gamma>) \<turnstile> M <: N"}. *}
 | 
|
720  | 
case 2  | 
|
721  | 
    have  "(\<Delta>@[(X,Q)]@\<Gamma>) \<turnstile> M <: N" --{* left-hand derivation *}
 | 
|
| 23393 | 722  | 
      and "\<Gamma> \<turnstile> P<:Q" by fact+ --{* right-hand derivation *}
 | 
| 18621 | 723  | 
thus "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> M <: N"  | 
| 22537 | 724  | 
proof (nominal_induct \<Gamma>\<equiv>"\<Delta>@[(X,Q)]@\<Gamma>" M N avoiding: \<Delta> \<Gamma> X rule: subtype_of.strong_induct)  | 
| 18424 | 725  | 
case (S_Top _ S \<Delta> \<Gamma> X)  | 
| 18621 | 726  | 
	--{* \begin{minipage}[t]{0.9\textwidth}
 | 
727  | 
	In this case the left-hand derivation is @{term "(\<Delta>@[(X,Q)]@\<Gamma>) \<turnstile> S <: Top"}. We show
 | 
|
728  | 
	that the context @{term "\<Delta>@[(X,P)]@\<Gamma>"} is ok and that @{term S} is closed in 
 | 
|
729  | 
	@{term "\<Delta>@[(X,P)]@\<Gamma>"}. Then we can apply the @{text "S_Top"}-rule.\end{minipage}*}
 | 
|
730  | 
hence lh_drv_prm\<^isub>1: "\<turnstile> (\<Delta>@[(X,Q)]@\<Gamma>) ok"  | 
|
731  | 
and lh_drv_prm\<^isub>2: "S closed_in (\<Delta>@[(X,Q)]@\<Gamma>)" by simp_all  | 
|
| 18424 | 732  | 
have rh_drv: "\<Gamma> \<turnstile> P <: Q" by fact  | 
733  | 
hence "P closed_in \<Gamma>" by (simp add: subtype_implies_closed)  | 
|
| 18621 | 734  | 
with lh_drv_prm\<^isub>1 have "\<turnstile> (\<Delta>@[(X,P)]@\<Gamma>) ok" by (simp add: replace_type)  | 
| 18412 | 735  | 
moreover  | 
| 18621 | 736  | 
from lh_drv_prm\<^isub>2 have "S closed_in (\<Delta>@[(X,P)]@\<Gamma>)"  | 
737  | 
by (simp add: closed_in_def domain_append)  | 
|
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
738  | 
ultimately show "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> S <: Top" by (simp add: subtype_of.S_Top)  | 
| 18246 | 739  | 
next  | 
| 22537 | 740  | 
case (S_Var Y S _ N \<Delta> \<Gamma> X)  | 
| 18621 | 741  | 
	--{* \begin{minipage}[t]{0.9\textwidth}
 | 
| 18628 | 742  | 
	In this case the left-hand derivation is @{term "(\<Delta>@[(X,Q)]@\<Gamma>) \<turnstile> Tvar Y <: N"} and
 | 
743  | 
	by inner induction hypothesis we have @{term "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> S <: N"}. We therefore 
 | 
|
| 18621 | 744  | 
	know that the contexts @{term "\<Delta>@[(X,Q)]@\<Gamma>"} and @{term "\<Delta>@[(X,P)]@\<Gamma>"} are ok, and that 
 | 
| 18628 | 745  | 
	@{term "(Y,S)"} is in @{term "\<Delta>@[(X,Q)]@\<Gamma>"}. We need to show that 
 | 
746  | 
	@{term "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> Tvar Y <: N"}  holds. In case @{term "X\<noteq>Y"} we know that 
 | 
|
747  | 
	@{term "(Y,S)"} is in @{term "\<Delta>@[(X,P)]@\<Gamma>"} and can use the inner induction hypothesis 
 | 
|
748  | 
	and rule @{text "S_Var"} to conclude. In case @{term "X=Y"} we can infer that 
 | 
|
749  | 
	@{term "S=Q"}; moreover we have that  @{term "(\<Delta>@[(X,P)]@\<Gamma>) extends \<Gamma>"} and therefore 
 | 
|
750  | 
	by @{text "weakening"} that @{term "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> P <: Q"} holds. By transitivity we 
 | 
|
751  | 
	obtain then @{term "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> P <: N"} and can conclude by applying rule 
 | 
|
752  | 
	@{text "S_Var"}.\end{minipage}*}
 | 
|
| 18621 | 753  | 
hence IH_inner: "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> S <: N"  | 
754  | 
and lh_drv_prm: "(Y,S) \<in> set (\<Delta>@[(X,Q)]@\<Gamma>)"  | 
|
755  | 
and rh_drv: "\<Gamma> \<turnstile> P<:Q"  | 
|
756  | 
and ok\<^isub>Q: "\<turnstile> (\<Delta>@[(X,Q)]@\<Gamma>) ok" by (simp_all add: subtype_implies_ok)  | 
|
757  | 
hence ok\<^isub>P: "\<turnstile> (\<Delta>@[(X,P)]@\<Gamma>) ok" by (simp add: subtype_implies_ok)  | 
|
758  | 
show "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> Tvar Y <: N"  | 
|
759  | 
proof (cases "X=Y")  | 
|
760  | 
case False  | 
|
761  | 
have "X\<noteq>Y" by fact  | 
|
762  | 
hence "(Y,S)\<in>set (\<Delta>@[(X,P)]@\<Gamma>)" using lh_drv_prm by simp  | 
|
763  | 
with IH_inner show "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> Tvar Y <: N" by (simp add: subtype_of.S_Var)  | 
|
764  | 
next  | 
|
765  | 
case True  | 
|
766  | 
have memb\<^isub>X\<^isub>Q: "(X,Q)\<in>set (\<Delta>@[(X,Q)]@\<Gamma>)" by simp  | 
|
767  | 
have memb\<^isub>X\<^isub>P: "(X,P)\<in>set (\<Delta>@[(X,P)]@\<Gamma>)" by simp  | 
|
768  | 
have eq: "X=Y" by fact  | 
|
769  | 
hence "S=Q" using ok\<^isub>Q lh_drv_prm memb\<^isub>X\<^isub>Q by (simp only: uniqueness_of_ctxt)  | 
|
| 18424 | 770  | 
hence "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> Q <: N" using IH_inner by simp  | 
771  | 
moreover  | 
|
772  | 
have "(\<Delta>@[(X,P)]@\<Gamma>) extends \<Gamma>" by (simp add: extends_def)  | 
|
| 18621 | 773  | 
hence "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> P <: Q" using rh_drv ok\<^isub>P by (simp only: weakening)  | 
774  | 
ultimately have "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> P <: N" by (simp add: transitivity_aux)  | 
|
775  | 
thus "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> Tvar Y <: N" using memb\<^isub>X\<^isub>P eq by (simp only: subtype_of.S_Var)  | 
|
776  | 
qed  | 
|
| 18246 | 777  | 
next  | 
| 18424 | 778  | 
case (S_Refl _ Y \<Delta> \<Gamma> X)  | 
| 18621 | 779  | 
	--{* \begin{minipage}[t]{0.9\textwidth}
 | 
780  | 
	In this case the left-hand derivation is @{term "(\<Delta>@[(X,Q)]@\<Gamma>) \<turnstile> Tvar Y <: Tvar Y"} and we
 | 
|
| 18628 | 781  | 
	therefore know that @{term "\<Delta>@[(X,Q)]@\<Gamma>"} is ok and that @{term "Y"} is in 
 | 
| 18621 | 782  | 
	the domain of @{term "\<Delta>@[(X,Q)]@\<Gamma>"}. We therefore know that @{term "\<Delta>@[(X,P)]@\<Gamma>"} is ok
 | 
783  | 
	and that @{term Y} is in the domain of @{term "\<Delta>@[(X,P)]@\<Gamma>"}. We can conclude by applying 
 | 
|
784  | 
	rule @{text "S_Refl"}.\end{minipage}*}
 | 
|
785  | 
hence lh_drv_prm\<^isub>1: "\<turnstile> (\<Delta>@[(X,Q)]@\<Gamma>) ok"  | 
|
786  | 
and lh_drv_prm\<^isub>2: "Y \<in> domain (\<Delta>@[(X,Q)]@\<Gamma>)" by simp_all  | 
|
| 18424 | 787  | 
have "\<Gamma> \<turnstile> P <: Q" by fact  | 
788  | 
hence "P closed_in \<Gamma>" by (simp add: subtype_implies_closed)  | 
|
| 18621 | 789  | 
with lh_drv_prm\<^isub>1 have "\<turnstile> (\<Delta>@[(X,P)]@\<Gamma>) ok" by (simp add: replace_type)  | 
| 18424 | 790  | 
moreover  | 
| 18621 | 791  | 
from lh_drv_prm\<^isub>2 have "Y \<in> domain (\<Delta>@[(X,P)]@\<Gamma>)" by (simp add: domain_append)  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
792  | 
ultimately show "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> Tvar Y <: Tvar Y" by (simp add: subtype_of.S_Refl)  | 
| 18246 | 793  | 
next  | 
| 22537 | 794  | 
case (S_Arrow _ S\<^isub>1 Q\<^isub>1 Q\<^isub>2 S\<^isub>2 \<Delta> \<Gamma> X)  | 
| 18621 | 795  | 
	--{* \begin{minipage}[t]{0.9\textwidth}
 | 
796  | 
	In this case the left-hand derivation is @{term "(\<Delta>@[(X,Q)]@\<Gamma>) \<turnstile> Q\<^isub>1 \<rightarrow> Q\<^isub>2 <: S\<^isub>1 \<rightarrow> S\<^isub>2"} 
 | 
|
797  | 
	and the proof is trivial.\end{minipage}*}
 | 
|
798  | 
thus "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> Q\<^isub>1 \<rightarrow> Q\<^isub>2 <: S\<^isub>1 \<rightarrow> S\<^isub>2" by blast  | 
|
| 18424 | 799  | 
next  | 
| 22537 | 800  | 
case (S_Forall _ T\<^isub>1 S\<^isub>1 Y S\<^isub>2 T\<^isub>2 \<Delta> \<Gamma> X)  | 
| 18621 | 801  | 
	--{* \begin{minipage}[t]{0.9\textwidth}
 | 
| 18628 | 802  | 
	In this case the left-hand derivation is @{text "(\<Delta>@[(X,Q)]@\<Gamma>) \<turnstile> \<forall>[Y<:S\<^isub>1].S\<^isub>2 <: \<forall>[Y<:T\<^isub>1].T\<^isub>2"}
 | 
| 18621 | 803  | 
	and therfore we know that the binder @{term Y} is fresh for @{term "\<Delta>@[(X,Q)]@\<Gamma>"}. By
 | 
804  | 
	the inner induction hypothesis we have that @{term "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> T\<^isub>1 <: S\<^isub>1"} and 
 | 
|
805  | 
	@{term "((Y,T\<^isub>1)#\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2"}. Since @{term P} is a subtype of @{term Q}
 | 
|
806  | 
	we can infer that @{term Y} is fresh for @{term P} and thus also fresh for 
 | 
|
807  | 
	@{term "\<Delta>@[(X,P)]@\<Gamma>"}. We can then conclude by applying rule @{text "S_Forall"}.
 | 
|
808  | 
	\end{minipage}*}
 | 
|
809  | 
hence IH_inner\<^isub>1: "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> T\<^isub>1 <: S\<^isub>1"  | 
|
810  | 
and IH_inner\<^isub>2: "((Y,T\<^isub>1)#\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2"  | 
|
811  | 
and lh_drv_prm: "Y\<sharp>(\<Delta>@[(X,Q)]@\<Gamma>)" by force+  | 
|
| 18424 | 812  | 
have rh_drv: "\<Gamma> \<turnstile> P <: Q" by fact  | 
| 18621 | 813  | 
hence "Y\<sharp>P" using lh_drv_prm by (simp only: fresh_list_append subtype_implies_fresh)  | 
814  | 
hence "Y\<sharp>(\<Delta>@[(X,P)]@\<Gamma>)" using lh_drv_prm  | 
|
| 18424 | 815  | 
by (simp add: fresh_list_append fresh_list_cons fresh_prod)  | 
| 18621 | 816  | 
with IH_inner\<^isub>1 IH_inner\<^isub>2  | 
817  | 
show "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> \<forall>[Y<:S\<^isub>1].S\<^isub>2 <: \<forall>[Y<:T\<^isub>1].T\<^isub>2" by (simp add: subtype_of.S_Forall)  | 
|
| 18246 | 818  | 
qed  | 
| 18621 | 819  | 
}  | 
| 18246 | 820  | 
qed  | 
821  | 
||
| 18416 | 822  | 
end  |