| author | wenzelm | 
| Sat, 08 Jun 2024 19:47:14 +0200 | |
| changeset 80302 | d1c7da4ff0f5 | 
| parent 80149 | 40a3fc07a587 | 
| child 80914 | d97fdabd9e2b | 
| 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  | 
| 
66453
 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 
wenzelm 
parents: 
63167 
diff
changeset
 | 
3  | 
imports "HOL-Nominal.Nominal"  | 
| 18269 | 4  | 
begin  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
5  | 
(*>*)  | 
| 18269 | 6  | 
|
| 63167 | 7  | 
text\<open>Authors: Christian Urban,  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
8  | 
Benjamin Pierce,  | 
| 18650 | 9  | 
Dimitrios Vytiniotis  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
10  | 
Stephanie Weirich  | 
| 18650 | 11  | 
Steve Zdancewic  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
12  | 
Julien Narboux  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
13  | 
Stefan Berghofer  | 
| 
18266
 
55c201fe4c95
added an authors section (please let me know if somebody is left out or unhappy)
 
urbanc 
parents: 
18263 
diff
changeset
 | 
14  | 
|
| 63167 | 15  | 
with great help from Markus Wenzel.\<close>  | 
| 18246 | 16  | 
|
| 63167 | 17  | 
section \<open>Types for Names, Nominal Datatype Declaration for Types and Terms\<close>  | 
| 18424 | 18  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
19  | 
no_syntax  | 
| 61069 | 20  | 
  "_Map" :: "maplets => 'a \<rightharpoonup> 'b"  ("(1[_])")
 | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
21  | 
|
| 63167 | 22  | 
text \<open>The main point of this solution is to use names everywhere (be they bound,  | 
| 18621 | 23  | 
  binding or free). In System \FSUB{} there are two kinds of names corresponding to 
 | 
24  | 
type-variables and to term-variables. These two kinds of names are represented in  | 
|
| 63167 | 25  | 
the nominal datatype package as atom-types \<open>tyvrs\<close> and \<open>vrs\<close>:\<close>  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
26  | 
|
| 18246 | 27  | 
atom_decl tyvrs vrs  | 
28  | 
||
| 63167 | 29  | 
text\<open>There are numerous facts that come with this declaration: for example that  | 
30  | 
there are infinitely many elements in \<open>tyvrs\<close> and \<open>vrs\<close>.\<close>  | 
|
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
31  | 
|
| 63167 | 32  | 
text\<open>The constructors for types and terms in System \FSUB{} contain abstractions 
 | 
| 
58305
 
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
 
blanchet 
parents: 
55417 
diff
changeset
 | 
33  | 
over type-variables and term-variables. The nominal datatype package uses  | 
| 63167 | 34  | 
\<open>\<guillemotleft>\<dots>\<guillemotright>\<dots>\<close> to indicate where abstractions occur.\<close>  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
35  | 
|
| 18424 | 36  | 
nominal_datatype ty =  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
37  | 
Tvar "tyvrs"  | 
| 18424 | 38  | 
| Top  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
39  | 
| Arrow "ty" "ty" (infixr "\<rightarrow>" 200)  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
40  | 
| Forall "\<guillemotleft>tyvrs\<guillemotright>ty" "ty"  | 
| 18246 | 41  | 
|
| 18424 | 42  | 
nominal_datatype trm =  | 
43  | 
Var "vrs"  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
44  | 
| Abs "\<guillemotleft>vrs\<guillemotright>trm" "ty"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
45  | 
| TAbs "\<guillemotleft>tyvrs\<guillemotright>trm" "ty"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
46  | 
| App "trm" "trm" (infixl "\<cdot>" 200)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
47  | 
| TApp "trm" "ty" (infixl "\<cdot>\<^sub>\<tau>" 200)  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
48  | 
|
| 63167 | 49  | 
text \<open>To be polite to the eye, some more familiar notation is introduced.  | 
| 18621 | 50  | 
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
 | 
51  | 
translation rules, instead of syntax annotations at the term-constructors  | 
| 69597 | 52  | 
as given above for \<^term>\<open>Arrow\<close>.\<close>  | 
| 18246 | 53  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
54  | 
abbreviation  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
55  | 
  Forall_syn :: "tyvrs \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> ty"  ("(3\<forall>_<:_./ _)" [0, 0, 10] 10) 
 | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
56  | 
where  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
57  | 
"\<forall>X<:T\<^sub>1. T\<^sub>2 \<equiv> ty.Forall X T\<^sub>2 T\<^sub>1"  | 
| 18424 | 58  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
59  | 
abbreviation  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
60  | 
  Abs_syn    :: "vrs \<Rightarrow> ty \<Rightarrow> trm \<Rightarrow> trm"  ("(3\<lambda>_:_./ _)" [0, 0, 10] 10) 
 | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
61  | 
where  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
62  | 
"\<lambda>x:T. t \<equiv> trm.Abs x t T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
63  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
64  | 
abbreviation  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
65  | 
  TAbs_syn   :: "tyvrs \<Rightarrow> ty \<Rightarrow> trm \<Rightarrow> trm" ("(3\<lambda>_<:_./ _)" [0, 0, 10] 10) 
 | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
66  | 
where  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
67  | 
"\<lambda>X<:T. t \<equiv> trm.TAbs X t T"  | 
| 18246 | 68  | 
|
| 69597 | 69  | 
text \<open>Again there are numerous facts that are proved automatically for \<^typ>\<open>ty\<close>  | 
70  | 
and \<^typ>\<open>trm\<close>: for example that the set of free variables, i.e.~the \<open>support\<close>,  | 
|
| 18650 | 71  | 
  is finite. However note that nominal-datatype declarations do \emph{not} define 
 | 
72  | 
``classical" constructor-based datatypes, but rather define $\alpha$-equivalence  | 
|
| 69597 | 73  | 
classes---we can for example show that $\alpha$-equivalent \<^typ>\<open>ty\<close>s  | 
74  | 
and \<^typ>\<open>trm\<close>s are equal:\<close>  | 
|
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
75  | 
|
| 
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
76  | 
lemma alpha_illustration:  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
77  | 
shows "(\<forall>X<:T. Tvar X) = (\<forall>Y<:T. Tvar Y)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
78  | 
and "(\<lambda>x:T. Var x) = (\<lambda>y:T. Var y)"  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
79  | 
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
 | 
80  | 
|
| 63167 | 81  | 
section \<open>SubTyping Contexts\<close>  | 
| 18246 | 82  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
83  | 
nominal_datatype binding =  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
84  | 
VarB vrs ty  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
85  | 
| TVarB tyvrs ty  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
86  | 
|
| 41798 | 87  | 
type_synonym env = "binding list"  | 
| 18246 | 88  | 
|
| 63167 | 89  | 
text \<open>Typing contexts are represented as lists that ``grow" on the left; we  | 
| 18621 | 90  | 
thereby deviating from the convention in the POPLmark-paper. The lists contain  | 
| 63167 | 91  | 
pairs of type-variables and types (this is sufficient for Part 1A).\<close>  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
92  | 
|
| 63167 | 93  | 
text \<open>In order to state validity-conditions for typing-contexts, the notion of  | 
94  | 
a \<open>dom\<close> of a typing-context is handy.\<close>  | 
|
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
95  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
96  | 
nominal_primrec  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
97  | 
"tyvrs_of" :: "binding \<Rightarrow> tyvrs set"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
98  | 
where  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
99  | 
  "tyvrs_of (VarB  x y) = {}"
 | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
100  | 
| "tyvrs_of (TVarB x y) = {x}"
 | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
101  | 
by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
102  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
103  | 
nominal_primrec  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
104  | 
"vrs_of" :: "binding \<Rightarrow> vrs set"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
105  | 
where  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
106  | 
  "vrs_of (VarB  x y) = {x}"
 | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
107  | 
| "vrs_of (TVarB x y) = {}"
 | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
108  | 
by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
109  | 
|
| 39246 | 110  | 
primrec  | 
| 32011 | 111  | 
"ty_dom" :: "env \<Rightarrow> tyvrs set"  | 
| 39246 | 112  | 
where  | 
| 32011 | 113  | 
  "ty_dom [] = {}"
 | 
| 39246 | 114  | 
| "ty_dom (X#\<Gamma>) = (tyvrs_of X)\<union>(ty_dom \<Gamma>)"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
115  | 
|
| 39246 | 116  | 
primrec  | 
| 32011 | 117  | 
"trm_dom" :: "env \<Rightarrow> vrs set"  | 
| 39246 | 118  | 
where  | 
| 32011 | 119  | 
  "trm_dom [] = {}"
 | 
| 39246 | 120  | 
| "trm_dom (X#\<Gamma>) = (vrs_of X)\<union>(trm_dom \<Gamma>)"  | 
| 18246 | 121  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
122  | 
lemma vrs_of_eqvt[eqvt]:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
123  | 
fixes pi ::"tyvrs prm"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
124  | 
and pi'::"vrs prm"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
125  | 
shows "pi \<bullet>(tyvrs_of x) = tyvrs_of (pi\<bullet>x)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
126  | 
and "pi'\<bullet>(tyvrs_of x) = tyvrs_of (pi'\<bullet>x)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
127  | 
and "pi \<bullet>(vrs_of x) = vrs_of (pi\<bullet>x)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
128  | 
and "pi'\<bullet>(vrs_of x) = vrs_of (pi'\<bullet>x)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
129  | 
by (nominal_induct x rule: binding.strong_induct) (simp_all add: tyvrs_of.simps eqvts)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
130  | 
|
| 32011 | 131  | 
lemma doms_eqvt[eqvt]:  | 
| 18246 | 132  | 
fixes pi::"tyvrs prm"  | 
| 22537 | 133  | 
and pi'::"vrs prm"  | 
| 32011 | 134  | 
shows "pi \<bullet>(ty_dom \<Gamma>) = ty_dom (pi\<bullet>\<Gamma>)"  | 
135  | 
and "pi'\<bullet>(ty_dom \<Gamma>) = ty_dom (pi'\<bullet>\<Gamma>)"  | 
|
136  | 
and "pi \<bullet>(trm_dom \<Gamma>) = trm_dom (pi\<bullet>\<Gamma>)"  | 
|
137  | 
and "pi'\<bullet>(trm_dom \<Gamma>) = trm_dom (pi'\<bullet>\<Gamma>)"  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
138  | 
by (induct \<Gamma>) (simp_all add: eqvts)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
139  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
140  | 
lemma finite_vrs:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
141  | 
shows "finite (tyvrs_of x)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
142  | 
and "finite (vrs_of x)"  | 
| 49171 | 143  | 
by (nominal_induct rule:binding.strong_induct) auto  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
144  | 
|
| 32011 | 145  | 
lemma finite_doms:  | 
146  | 
shows "finite (ty_dom \<Gamma>)"  | 
|
147  | 
and "finite (trm_dom \<Gamma>)"  | 
|
| 80142 | 148  | 
by (induct \<Gamma>) (auto simp: finite_vrs)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
149  | 
|
| 32011 | 150  | 
lemma ty_dom_supp:  | 
151  | 
shows "(supp (ty_dom \<Gamma>)) = (ty_dom \<Gamma>)"  | 
|
152  | 
and "(supp (trm_dom \<Gamma>)) = (trm_dom \<Gamma>)"  | 
|
153  | 
by (simp only: at_fin_set_supp at_tyvrs_inst at_vrs_inst finite_doms)+  | 
|
| 18246 | 154  | 
|
| 32011 | 155  | 
lemma ty_dom_inclusion:  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
156  | 
assumes a: "(TVarB X T)\<in>set \<Gamma>"  | 
| 32011 | 157  | 
shows "X\<in>(ty_dom \<Gamma>)"  | 
| 49171 | 158  | 
using a by (induct \<Gamma>) (auto)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
159  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
160  | 
lemma ty_binding_existence:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
161  | 
assumes "X \<in> (tyvrs_of a)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
162  | 
shows "\<exists>T.(TVarB X T=a)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
163  | 
using assms  | 
| 49171 | 164  | 
by (nominal_induct a rule: binding.strong_induct) (auto)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
165  | 
|
| 32011 | 166  | 
lemma ty_dom_existence:  | 
167  | 
assumes a: "X\<in>(ty_dom \<Gamma>)"  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
168  | 
shows "\<exists>T.(TVarB X T)\<in>set \<Gamma>"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
169  | 
using a  | 
| 80142 | 170  | 
proof (induction \<Gamma>)  | 
171  | 
case Nil then show ?case by simp  | 
|
172  | 
next  | 
|
173  | 
case (Cons a \<Gamma>) then show ?case  | 
|
174  | 
using ty_binding_existence by fastforce  | 
|
175  | 
qed  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
176  | 
|
| 32011 | 177  | 
lemma doms_append:  | 
178  | 
shows "ty_dom (\<Gamma>@\<Delta>) = ((ty_dom \<Gamma>) \<union> (ty_dom \<Delta>))"  | 
|
179  | 
and "trm_dom (\<Gamma>@\<Delta>) = ((trm_dom \<Gamma>) \<union> (trm_dom \<Delta>))"  | 
|
| 49171 | 180  | 
by (induct \<Gamma>) (auto)  | 
| 18246 | 181  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
182  | 
lemma ty_vrs_prm_simp:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
183  | 
fixes pi::"vrs prm"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
184  | 
and S::"ty"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
185  | 
shows "pi\<bullet>S = S"  | 
| 80142 | 186  | 
by (induct S rule: ty.induct) (auto simp: calc_atm)  | 
| 18246 | 187  | 
|
| 32011 | 188  | 
lemma fresh_ty_dom_cons:  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
189  | 
fixes X::"tyvrs"  | 
| 32011 | 190  | 
shows "X\<sharp>(ty_dom (Y#\<Gamma>)) = (X\<sharp>(tyvrs_of Y) \<and> X\<sharp>(ty_dom \<Gamma>))"  | 
| 80142 | 191  | 
proof (nominal_induct rule:binding.strong_induct)  | 
192  | 
case (VarB vrs ty)  | 
|
193  | 
then show ?case by auto  | 
|
194  | 
next  | 
|
195  | 
case (TVarB tyvrs ty)  | 
|
196  | 
then show ?case  | 
|
197  | 
by (simp add: at_fin_set_supp at_tyvrs_inst finite_doms(1) fresh_def supp_atm(1))  | 
|
198  | 
qed  | 
|
| 18246 | 199  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
200  | 
lemma tyvrs_fresh:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
201  | 
fixes X::"tyvrs"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
202  | 
assumes "X \<sharp> a"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
203  | 
shows "X \<sharp> tyvrs_of a"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
204  | 
and "X \<sharp> vrs_of a"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
205  | 
using assms  | 
| 80142 | 206  | 
by (nominal_induct a rule:binding.strong_induct) (force simp: fresh_singleton)+  | 
| 18621 | 207  | 
|
| 32011 | 208  | 
lemma fresh_dom:  | 
| 18621 | 209  | 
fixes X::"tyvrs"  | 
210  | 
assumes a: "X\<sharp>\<Gamma>"  | 
|
| 32011 | 211  | 
shows "X\<sharp>(ty_dom \<Gamma>)"  | 
| 18621 | 212  | 
using a  | 
| 80142 | 213  | 
proof (induct \<Gamma>)  | 
214  | 
case Nil then show ?case by auto  | 
|
215  | 
next  | 
|
216  | 
case (Cons a \<Gamma>) then show ?case  | 
|
217  | 
by (meson fresh_list_cons fresh_ty_dom_cons tyvrs_fresh(1))  | 
|
218  | 
qed  | 
|
| 18621 | 219  | 
|
| 69597 | 220  | 
text \<open>Not all lists of type \<^typ>\<open>env\<close> are well-formed. One condition  | 
221  | 
requires that in \<^term>\<open>TVarB X S#\<Gamma>\<close> all free variables of \<^term>\<open>S\<close> must be  | 
|
222  | 
in the \<^term>\<open>ty_dom\<close> of \<^term>\<open>\<Gamma>\<close>, that is \<^term>\<open>S\<close> must be \<open>closed\<close>  | 
|
223  | 
in \<^term>\<open>\<Gamma>\<close>. The set of free variables of \<^term>\<open>S\<close> is the  | 
|
224  | 
\<open>support\<close> of \<^term>\<open>S\<close>.\<close>  | 
|
| 18246 | 225  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
34915 
diff
changeset
 | 
226  | 
definition "closed_in" :: "ty \<Rightarrow> env \<Rightarrow> bool" ("_ closed'_in _" [100,100] 100) where
 | 
| 32011 | 227  | 
"S closed_in \<Gamma> \<equiv> (supp S)\<subseteq>(ty_dom \<Gamma>)"  | 
| 18246 | 228  | 
|
| 22537 | 229  | 
lemma closed_in_eqvt[eqvt]:  | 
| 18246 | 230  | 
fixes pi::"tyvrs prm"  | 
231  | 
assumes a: "S closed_in \<Gamma>"  | 
|
232  | 
shows "(pi\<bullet>S) closed_in (pi\<bullet>\<Gamma>)"  | 
|
233  | 
using a  | 
|
| 26091 | 234  | 
proof -  | 
235  | 
from a have "pi\<bullet>(S closed_in \<Gamma>)" by (simp add: perm_bool)  | 
|
236  | 
then show "(pi\<bullet>S) closed_in (pi\<bullet>\<Gamma>)" by (simp add: closed_in_def eqvts)  | 
|
| 18246 | 237  | 
qed  | 
238  | 
||
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
239  | 
lemma tyvrs_vrs_prm_simp:  | 
| 22537 | 240  | 
fixes pi::"vrs prm"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
241  | 
shows "tyvrs_of (pi\<bullet>a) = tyvrs_of a"  | 
| 80142 | 242  | 
by (nominal_induct rule:binding.strong_induct) (auto simp: vrs_prm_tyvrs_def)  | 
| 22537 | 243  | 
|
| 
30986
 
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
 
Christian Urban <urbanc@in.tum.de> 
parents: 
30091 
diff
changeset
 | 
244  | 
lemma ty_vrs_fresh:  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
245  | 
fixes x::"vrs"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
246  | 
and T::"ty"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
247  | 
shows "x \<sharp> T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
248  | 
by (simp add: fresh_def supp_def ty_vrs_prm_simp)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
249  | 
|
| 32011 | 250  | 
lemma ty_dom_vrs_prm_simp:  | 
| 22537 | 251  | 
fixes pi::"vrs prm"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
252  | 
and \<Gamma>::"env"  | 
| 32011 | 253  | 
shows "(ty_dom (pi\<bullet>\<Gamma>)) = (ty_dom \<Gamma>)"  | 
| 80142 | 254  | 
by (induct \<Gamma>) (auto simp: tyvrs_vrs_prm_simp)  | 
| 22537 | 255  | 
|
256  | 
lemma closed_in_eqvt'[eqvt]:  | 
|
257  | 
fixes pi::"vrs prm"  | 
|
258  | 
assumes a: "S closed_in \<Gamma>"  | 
|
259  | 
shows "(pi\<bullet>S) closed_in (pi\<bullet>\<Gamma>)"  | 
|
| 80142 | 260  | 
using assms closed_in_def ty_dom_vrs_prm_simp ty_vrs_prm_simp by force  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
261  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
262  | 
lemma fresh_vrs_of:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
263  | 
fixes x::"vrs"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
264  | 
shows "x\<sharp>vrs_of b = x\<sharp>b"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
265  | 
by (nominal_induct b rule: binding.strong_induct)  | 
| 
46182
 
b4aa5e39f944
Removed strange hack introduced in b27e93132603, since equivariance
 
berghofe 
parents: 
45971 
diff
changeset
 | 
266  | 
(simp_all add: fresh_singleton fresh_set_empty ty_vrs_fresh fresh_atm)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
267  | 
|
| 32011 | 268  | 
lemma fresh_trm_dom:  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
269  | 
fixes x::"vrs"  | 
| 32011 | 270  | 
shows "x\<sharp> trm_dom \<Gamma> = x\<sharp>\<Gamma>"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
271  | 
by (induct \<Gamma>)  | 
| 80142 | 272  | 
(simp_all add: fresh_list_cons  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
273  | 
fresh_fin_union [OF pt_vrs_inst at_vrs_inst fs_vrs_inst]  | 
| 80142 | 274  | 
finite_doms finite_vrs fresh_vrs_of)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
275  | 
|
| 32011 | 276  | 
lemma closed_in_fresh: "(X::tyvrs) \<sharp> ty_dom \<Gamma> \<Longrightarrow> T closed_in \<Gamma> \<Longrightarrow> X \<sharp> T"  | 
| 80142 | 277  | 
by (auto simp: closed_in_def fresh_def ty_dom_supp)  | 
| 22537 | 278  | 
|
| 63167 | 279  | 
text \<open>Now validity of a context is a straightforward inductive definition.\<close>  | 
| 18621 | 280  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
281  | 
inductive  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
282  | 
  valid_rel :: "env \<Rightarrow> bool" ("\<turnstile> _ ok" [100] 100)
 | 
| 22436 | 283  | 
where  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
284  | 
valid_nil[simp]: "\<turnstile> [] ok"  | 
| 32011 | 285  | 
| valid_consT[simp]: "\<lbrakk>\<turnstile> \<Gamma> ok; X\<sharp>(ty_dom \<Gamma>); T closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<turnstile> (TVarB X T#\<Gamma>) ok"  | 
286  | 
| valid_cons [simp]: "\<lbrakk>\<turnstile> \<Gamma> ok; x\<sharp>(trm_dom \<Gamma>); T closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<turnstile> (VarB x T#\<Gamma>) ok"  | 
|
| 18246 | 287  | 
|
| 22537 | 288  | 
equivariance valid_rel  | 
| 18246 | 289  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
290  | 
declare binding.inject [simp add]  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
291  | 
declare trm.inject [simp add]  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
292  | 
|
| 32011 | 293  | 
inductive_cases validE[elim]:  | 
294  | 
"\<turnstile> (TVarB X T#\<Gamma>) ok"  | 
|
295  | 
"\<turnstile> (VarB x T#\<Gamma>) ok"  | 
|
296  | 
"\<turnstile> (b#\<Gamma>) ok"  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
297  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
298  | 
declare binding.inject [simp del]  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
299  | 
declare trm.inject [simp del]  | 
| 18246 | 300  | 
|
| 18424 | 301  | 
lemma validE_append:  | 
302  | 
assumes a: "\<turnstile> (\<Delta>@\<Gamma>) ok"  | 
|
303  | 
shows "\<turnstile> \<Gamma> ok"  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
304  | 
using a  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
305  | 
proof (induct \<Delta>)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
306  | 
case (Cons a \<Gamma>')  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
307  | 
then show ?case  | 
| 80142 | 308  | 
by (nominal_induct a rule:binding.strong_induct) auto  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
309  | 
qed (auto)  | 
| 18246 | 310  | 
|
| 18424 | 311  | 
lemma replace_type:  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
312  | 
assumes a: "\<turnstile> (\<Delta>@(TVarB X T)#\<Gamma>) ok"  | 
| 18424 | 313  | 
and b: "S closed_in \<Gamma>"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
314  | 
shows "\<turnstile> (\<Delta>@(TVarB X S)#\<Gamma>) ok"  | 
| 18621 | 315  | 
using a b  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
316  | 
proof(induct \<Delta>)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
317  | 
case Nil  | 
| 80142 | 318  | 
then show ?case by (auto intro: valid_cons simp add: doms_append closed_in_def)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
319  | 
next  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
320  | 
case (Cons a \<Gamma>')  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
321  | 
then show ?case  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
322  | 
by (nominal_induct a rule:binding.strong_induct)  | 
| 80142 | 323  | 
(auto intro!: valid_cons simp add: doms_append closed_in_def)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
324  | 
qed  | 
| 18246 | 325  | 
|
| 63167 | 326  | 
text \<open>Well-formed contexts have a unique type-binding for a type-variable.\<close>  | 
| 18650 | 327  | 
|
| 18246 | 328  | 
lemma uniqueness_of_ctxt:  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
329  | 
fixes \<Gamma>::"env"  | 
| 18412 | 330  | 
assumes a: "\<turnstile> \<Gamma> ok"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
331  | 
and b: "(TVarB X T)\<in>set \<Gamma>"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
332  | 
and c: "(TVarB X S)\<in>set \<Gamma>"  | 
| 18412 | 333  | 
shows "T=S"  | 
| 18621 | 334  | 
using a b c  | 
335  | 
proof (induct)  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
336  | 
case (valid_consT \<Gamma> X' T')  | 
| 18621 | 337  | 
moreover  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
338  | 
  { fix \<Gamma>'::"env"
 | 
| 32011 | 339  | 
assume a: "X'\<sharp>(ty_dom \<Gamma>')"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
340  | 
have "\<not>(\<exists>T.(TVarB X' T)\<in>(set \<Gamma>'))" using a  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
341  | 
proof (induct \<Gamma>')  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
342  | 
case (Cons Y \<Gamma>')  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
343  | 
thus "\<not> (\<exists>T.(TVarB X' T)\<in>set(Y#\<Gamma>'))"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
344  | 
by (simp add: fresh_ty_dom_cons  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
345  | 
fresh_fin_union[OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst]  | 
| 32011 | 346  | 
finite_vrs finite_doms,  | 
| 80142 | 347  | 
auto simp: fresh_atm fresh_singleton)  | 
| 18621 | 348  | 
qed (simp)  | 
349  | 
}  | 
|
| 80142 | 350  | 
ultimately show "T=S" by (auto simp: binding.inject)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
351  | 
qed (auto)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
352  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
353  | 
lemma uniqueness_of_ctxt':  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
354  | 
fixes \<Gamma>::"env"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
355  | 
assumes a: "\<turnstile> \<Gamma> ok"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
356  | 
and b: "(VarB x T)\<in>set \<Gamma>"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
357  | 
and c: "(VarB x S)\<in>set \<Gamma>"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
358  | 
shows "T=S"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
359  | 
using a b c  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
360  | 
proof (induct)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
361  | 
case (valid_cons \<Gamma> x' T')  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
362  | 
moreover  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
363  | 
  { fix \<Gamma>'::"env"
 | 
| 32011 | 364  | 
assume a: "x'\<sharp>(trm_dom \<Gamma>')"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
365  | 
have "\<not>(\<exists>T.(VarB x' T)\<in>(set \<Gamma>'))" using a  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
366  | 
proof (induct \<Gamma>')  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
367  | 
case (Cons y \<Gamma>')  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
368  | 
thus "\<not> (\<exists>T.(VarB x' T)\<in>set(y#\<Gamma>'))"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
369  | 
by (simp add: fresh_fin_union[OF pt_vrs_inst at_vrs_inst fs_vrs_inst]  | 
| 32011 | 370  | 
finite_vrs finite_doms,  | 
| 80142 | 371  | 
auto simp: fresh_atm fresh_singleton)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
372  | 
qed (simp)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
373  | 
}  | 
| 80142 | 374  | 
ultimately show "T=S" by (auto simp: binding.inject)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
375  | 
qed (auto)  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
376  | 
|
| 63167 | 377  | 
section \<open>Size and Capture-Avoiding Substitution for Types\<close>  | 
| 18621 | 378  | 
|
| 21554 | 379  | 
nominal_primrec  | 
| 
29097
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
26966 
diff
changeset
 | 
380  | 
size_ty :: "ty \<Rightarrow> nat"  | 
| 
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
26966 
diff
changeset
 | 
381  | 
where  | 
| 21554 | 382  | 
"size_ty (Tvar X) = 1"  | 
| 
29097
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
26966 
diff
changeset
 | 
383  | 
| "size_ty (Top) = 1"  | 
| 
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
26966 
diff
changeset
 | 
384  | 
| "size_ty (T1 \<rightarrow> T2) = (size_ty T1) + (size_ty T2) + 1"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
385  | 
| "X \<sharp> T1 \<Longrightarrow> size_ty (\<forall>X<:T1. T2) = (size_ty T1) + (size_ty T2) + 1"  | 
| 80142 | 386  | 
by (finite_guess | fresh_guess | simp)+  | 
| 
20395
 
9a60e3151244
added definition for size and substitution using the recursion
 
urbanc 
parents: 
19972 
diff
changeset
 | 
387  | 
|
| 21554 | 388  | 
nominal_primrec  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
389  | 
  subst_ty :: "ty \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> ty" ("_[_ \<mapsto> _]\<^sub>\<tau>" [300, 0, 0] 300)
 | 
| 
29097
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
26966 
diff
changeset
 | 
390  | 
where  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
391  | 
"(Tvar X)[Y \<mapsto> T]\<^sub>\<tau> = (if X=Y then T else Tvar X)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
392  | 
| "(Top)[Y \<mapsto> T]\<^sub>\<tau> = Top"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
393  | 
| "(T\<^sub>1 \<rightarrow> T\<^sub>2)[Y \<mapsto> T]\<^sub>\<tau> = T\<^sub>1[Y \<mapsto> T]\<^sub>\<tau> \<rightarrow> T\<^sub>2[Y \<mapsto> T]\<^sub>\<tau>"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
394  | 
| "X\<sharp>(Y,T,T\<^sub>1) \<Longrightarrow> (\<forall>X<:T\<^sub>1. T\<^sub>2)[Y \<mapsto> T]\<^sub>\<tau> = (\<forall>X<:T\<^sub>1[Y \<mapsto> T]\<^sub>\<tau>. T\<^sub>2[Y \<mapsto> T]\<^sub>\<tau>)"  | 
| 80142 | 395  | 
by (finite_guess | fresh_guess | simp add: abs_fresh)+  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
396  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
397  | 
lemma subst_eqvt[eqvt]:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
398  | 
fixes pi::"tyvrs prm"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
399  | 
and T::"ty"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
400  | 
shows "pi\<bullet>(T[X \<mapsto> T']\<^sub>\<tau>) = (pi\<bullet>T)[(pi\<bullet>X) \<mapsto> (pi\<bullet>T')]\<^sub>\<tau>"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
401  | 
by (nominal_induct T avoiding: X T' rule: ty.strong_induct)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
402  | 
(perm_simp add: fresh_bij)+  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
403  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
404  | 
lemma subst_eqvt'[eqvt]:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
405  | 
fixes pi::"vrs prm"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
406  | 
and T::"ty"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
407  | 
shows "pi\<bullet>(T[X \<mapsto> T']\<^sub>\<tau>) = (pi\<bullet>T)[(pi\<bullet>X) \<mapsto> (pi\<bullet>T')]\<^sub>\<tau>"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
408  | 
by (nominal_induct T avoiding: X T' rule: ty.strong_induct)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
409  | 
(perm_simp add: fresh_left)+  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
410  | 
|
| 
30986
 
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
 
Christian Urban <urbanc@in.tum.de> 
parents: 
30091 
diff
changeset
 | 
411  | 
lemma type_subst_fresh:  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
412  | 
fixes X::"tyvrs"  | 
| 32011 | 413  | 
assumes "X\<sharp>T" and "X\<sharp>P"  | 
414  | 
shows "X\<sharp>T[Y \<mapsto> P]\<^sub>\<tau>"  | 
|
| 80142 | 415  | 
using assms  | 
416  | 
by (nominal_induct T avoiding: X Y P rule:ty.strong_induct)  | 
|
417  | 
(auto simp: abs_fresh)  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
418  | 
|
| 
30986
 
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
 
Christian Urban <urbanc@in.tum.de> 
parents: 
30091 
diff
changeset
 | 
419  | 
lemma fresh_type_subst_fresh:  | 
| 80142 | 420  | 
assumes "X\<sharp>T'"  | 
421  | 
shows "X\<sharp>T[X \<mapsto> T']\<^sub>\<tau>"  | 
|
422  | 
using assms  | 
|
423  | 
by (nominal_induct T avoiding: X T' rule: ty.strong_induct)  | 
|
424  | 
(auto simp: fresh_atm abs_fresh)  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
425  | 
|
| 32011 | 426  | 
lemma type_subst_identity:  | 
427  | 
"X\<sharp>T \<Longrightarrow> T[X \<mapsto> U]\<^sub>\<tau> = T"  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
428  | 
by (nominal_induct T avoiding: X U rule: ty.strong_induct)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
429  | 
(simp_all add: fresh_atm abs_fresh)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
430  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
431  | 
lemma type_substitution_lemma:  | 
| 32011 | 432  | 
"X \<noteq> Y \<Longrightarrow> X\<sharp>L \<Longrightarrow> M[X \<mapsto> N]\<^sub>\<tau>[Y \<mapsto> L]\<^sub>\<tau> = M[Y \<mapsto> L]\<^sub>\<tau>[X \<mapsto> N[Y \<mapsto> L]\<^sub>\<tau>]\<^sub>\<tau>"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
433  | 
by (nominal_induct M avoiding: X Y N L rule: ty.strong_induct)  | 
| 80142 | 434  | 
(auto simp: type_subst_fresh type_subst_identity)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
435  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
436  | 
lemma type_subst_rename:  | 
| 32011 | 437  | 
"Y\<sharp>T \<Longrightarrow> ([(Y,X)]\<bullet>T)[Y \<mapsto> U]\<^sub>\<tau> = T[X \<mapsto> U]\<^sub>\<tau>"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
438  | 
by (nominal_induct T avoiding: X Y U rule: ty.strong_induct)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
439  | 
(simp_all add: fresh_atm calc_atm abs_fresh fresh_aux)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
440  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
441  | 
nominal_primrec  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
442  | 
  subst_tyb :: "binding \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> binding" ("_[_ \<mapsto> _]\<^sub>b" [100,100,100] 100)
 | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
443  | 
where  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
444  | 
"(TVarB X U)[Y \<mapsto> T]\<^sub>b = TVarB X (U[Y \<mapsto> T]\<^sub>\<tau>)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
445  | 
| "(VarB X U)[Y \<mapsto> T]\<^sub>b = VarB X (U[Y \<mapsto> T]\<^sub>\<tau>)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
446  | 
by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
447  | 
|
| 
30986
 
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
 
Christian Urban <urbanc@in.tum.de> 
parents: 
30091 
diff
changeset
 | 
448  | 
lemma binding_subst_fresh:  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
449  | 
fixes X::"tyvrs"  | 
| 32011 | 450  | 
assumes "X\<sharp>a"  | 
451  | 
and "X\<sharp>P"  | 
|
452  | 
shows "X\<sharp>a[Y \<mapsto> P]\<^sub>b"  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
453  | 
using assms  | 
| 
30986
 
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
 
Christian Urban <urbanc@in.tum.de> 
parents: 
30091 
diff
changeset
 | 
454  | 
by (nominal_induct a rule: binding.strong_induct)  | 
| 80142 | 455  | 
(auto simp: type_subst_fresh)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
456  | 
|
| 
30986
 
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
 
Christian Urban <urbanc@in.tum.de> 
parents: 
30091 
diff
changeset
 | 
457  | 
lemma binding_subst_identity:  | 
| 32011 | 458  | 
shows "X\<sharp>B \<Longrightarrow> B[X \<mapsto> U]\<^sub>b = B"  | 
| 
30986
 
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
 
Christian Urban <urbanc@in.tum.de> 
parents: 
30091 
diff
changeset
 | 
459  | 
by (induct B rule: binding.induct)  | 
| 
 
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
 
Christian Urban <urbanc@in.tum.de> 
parents: 
30091 
diff
changeset
 | 
460  | 
(simp_all add: fresh_atm type_subst_identity)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
461  | 
|
| 39246 | 462  | 
primrec subst_tyc :: "env \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> env" ("_[_ \<mapsto> _]\<^sub>e" [100,100,100] 100) where
 | 
463  | 
"([])[Y \<mapsto> T]\<^sub>e= []"  | 
|
464  | 
| "(B#\<Gamma>)[Y \<mapsto> T]\<^sub>e = (B[Y \<mapsto> T]\<^sub>b)#(\<Gamma>[Y \<mapsto> T]\<^sub>e)"  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
465  | 
|
| 
30986
 
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
 
Christian Urban <urbanc@in.tum.de> 
parents: 
30091 
diff
changeset
 | 
466  | 
lemma ctxt_subst_fresh':  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
467  | 
fixes X::"tyvrs"  | 
| 32011 | 468  | 
assumes "X\<sharp>\<Gamma>"  | 
469  | 
and "X\<sharp>P"  | 
|
470  | 
shows "X\<sharp>\<Gamma>[Y \<mapsto> P]\<^sub>e"  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
471  | 
using assms  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
472  | 
by (induct \<Gamma>)  | 
| 80142 | 473  | 
(auto simp: fresh_list_cons binding_subst_fresh)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
474  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
475  | 
lemma ctxt_subst_mem_TVarB: "TVarB X T \<in> set \<Gamma> \<Longrightarrow> TVarB X (T[Y \<mapsto> U]\<^sub>\<tau>) \<in> set (\<Gamma>[Y \<mapsto> U]\<^sub>e)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
476  | 
by (induct \<Gamma>) auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
477  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
478  | 
lemma ctxt_subst_mem_VarB: "VarB x T \<in> set \<Gamma> \<Longrightarrow> VarB x (T[Y \<mapsto> U]\<^sub>\<tau>) \<in> set (\<Gamma>[Y \<mapsto> U]\<^sub>e)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
479  | 
by (induct \<Gamma>) auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
480  | 
|
| 32011 | 481  | 
lemma ctxt_subst_identity: "X\<sharp>\<Gamma> \<Longrightarrow> \<Gamma>[X \<mapsto> U]\<^sub>e = \<Gamma>"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
482  | 
by (induct \<Gamma>) (simp_all add: fresh_list_cons binding_subst_identity)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
483  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
484  | 
lemma ctxt_subst_append: "(\<Delta> @ \<Gamma>)[X \<mapsto> T]\<^sub>e = \<Delta>[X \<mapsto> T]\<^sub>e @ \<Gamma>[X \<mapsto> T]\<^sub>e"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
485  | 
by (induct \<Delta>) simp_all  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
486  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
487  | 
nominal_primrec  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
488  | 
   subst_trm :: "trm \<Rightarrow> vrs \<Rightarrow> trm \<Rightarrow> trm"  ("_[_ \<mapsto> _]" [300, 0, 0] 300)
 | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
489  | 
where  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
490  | 
"(Var x)[y \<mapsto> t'] = (if x=y then t' else (Var x))"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
491  | 
| "(t1 \<cdot> t2)[y \<mapsto> t'] = t1[y \<mapsto> t'] \<cdot> t2[y \<mapsto> t']"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
492  | 
| "(t \<cdot>\<^sub>\<tau> T)[y \<mapsto> t'] = t[y \<mapsto> t'] \<cdot>\<^sub>\<tau> T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
493  | 
| "X\<sharp>(T,t') \<Longrightarrow> (\<lambda>X<:T. t)[y \<mapsto> t'] = (\<lambda>X<:T. t[y \<mapsto> t'])"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
494  | 
| "x\<sharp>(y,t') \<Longrightarrow> (\<lambda>x:T. t)[y \<mapsto> t'] = (\<lambda>x:T. t[y \<mapsto> t'])"  | 
| 80149 | 495  | 
by (finite_guess | simp add: abs_fresh | fresh_guess add: ty_vrs_fresh abs_fresh)+  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
496  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
497  | 
lemma subst_trm_fresh_tyvar:  | 
| 32011 | 498  | 
fixes X::"tyvrs"  | 
499  | 
shows "X\<sharp>t \<Longrightarrow> X\<sharp>u \<Longrightarrow> X\<sharp>t[x \<mapsto> u]"  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
500  | 
by (nominal_induct t avoiding: x u rule: trm.strong_induct)  | 
| 80142 | 501  | 
(auto simp: abs_fresh)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
502  | 
|
| 32011 | 503  | 
lemma subst_trm_fresh_var:  | 
504  | 
"x\<sharp>u \<Longrightarrow> x\<sharp>t[x \<mapsto> u]"  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
505  | 
by (nominal_induct t avoiding: x u rule: trm.strong_induct)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
506  | 
(simp_all add: abs_fresh fresh_atm ty_vrs_fresh)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
507  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
508  | 
lemma subst_trm_eqvt[eqvt]:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
509  | 
fixes pi::"tyvrs prm"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
510  | 
and t::"trm"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
511  | 
shows "pi\<bullet>(t[x \<mapsto> u]) = (pi\<bullet>t)[(pi\<bullet>x) \<mapsto> (pi\<bullet>u)]"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
512  | 
by (nominal_induct t avoiding: x u rule: trm.strong_induct)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
513  | 
(perm_simp add: fresh_left)+  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
514  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
515  | 
lemma subst_trm_eqvt'[eqvt]:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
516  | 
fixes pi::"vrs prm"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
517  | 
and t::"trm"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
518  | 
shows "pi\<bullet>(t[x \<mapsto> u]) = (pi\<bullet>t)[(pi\<bullet>x) \<mapsto> (pi\<bullet>u)]"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
519  | 
by (nominal_induct t avoiding: x u rule: trm.strong_induct)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
520  | 
(perm_simp add: fresh_left)+  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
521  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
522  | 
lemma subst_trm_rename:  | 
| 32011 | 523  | 
"y\<sharp>t \<Longrightarrow> ([(y, x)] \<bullet> t)[y \<mapsto> u] = t[x \<mapsto> u]"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
524  | 
by (nominal_induct t avoiding: x y u rule: trm.strong_induct)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
525  | 
(simp_all add: fresh_atm calc_atm abs_fresh fresh_aux ty_vrs_fresh perm_fresh_fresh)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
526  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
527  | 
nominal_primrec (freshness_context: "T2::ty")  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
528  | 
  subst_trm_ty :: "trm \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> trm"  ("_[_ \<mapsto>\<^sub>\<tau> _]" [300, 0, 0] 300)
 | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
529  | 
where  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
530  | 
"(Var x)[Y \<mapsto>\<^sub>\<tau> T2] = Var x"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
531  | 
| "(t1 \<cdot> t2)[Y \<mapsto>\<^sub>\<tau> T2] = t1[Y \<mapsto>\<^sub>\<tau> T2] \<cdot> t2[Y \<mapsto>\<^sub>\<tau> T2]"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
532  | 
| "(t1 \<cdot>\<^sub>\<tau> T)[Y \<mapsto>\<^sub>\<tau> T2] = t1[Y \<mapsto>\<^sub>\<tau> T2] \<cdot>\<^sub>\<tau> T[Y \<mapsto> T2]\<^sub>\<tau>"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
533  | 
| "X\<sharp>(Y,T,T2) \<Longrightarrow> (\<lambda>X<:T. t)[Y \<mapsto>\<^sub>\<tau> T2] = (\<lambda>X<:T[Y \<mapsto> T2]\<^sub>\<tau>. t[Y \<mapsto>\<^sub>\<tau> T2])"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
534  | 
| "(\<lambda>x:T. t)[Y \<mapsto>\<^sub>\<tau> T2] = (\<lambda>x:T[Y \<mapsto> T2]\<^sub>\<tau>. t[Y \<mapsto>\<^sub>\<tau> T2])"  | 
| 80142 | 535  | 
apply(finite_guess | simp add: abs_fresh ty_vrs_fresh type_subst_fresh)+  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
536  | 
apply(fresh_guess add: ty_vrs_fresh abs_fresh)+  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
537  | 
done  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
538  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
539  | 
lemma subst_trm_ty_fresh:  | 
| 32011 | 540  | 
fixes X::"tyvrs"  | 
541  | 
shows "X\<sharp>t \<Longrightarrow> X\<sharp>T \<Longrightarrow> X\<sharp>t[Y \<mapsto>\<^sub>\<tau> T]"  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
542  | 
by (nominal_induct t avoiding: Y T rule: trm.strong_induct)  | 
| 80142 | 543  | 
(auto simp: abs_fresh type_subst_fresh)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
544  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
545  | 
lemma subst_trm_ty_fresh':  | 
| 32011 | 546  | 
"X\<sharp>T \<Longrightarrow> X\<sharp>t[X \<mapsto>\<^sub>\<tau> T]"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
547  | 
by (nominal_induct t avoiding: X T rule: trm.strong_induct)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
548  | 
(simp_all add: abs_fresh fresh_type_subst_fresh fresh_atm)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
549  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
550  | 
lemma subst_trm_ty_eqvt[eqvt]:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
551  | 
fixes pi::"tyvrs prm"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
552  | 
and t::"trm"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
553  | 
shows "pi\<bullet>(t[X \<mapsto>\<^sub>\<tau> T]) = (pi\<bullet>t)[(pi\<bullet>X) \<mapsto>\<^sub>\<tau> (pi\<bullet>T)]"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
554  | 
by (nominal_induct t avoiding: X T rule: trm.strong_induct)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
555  | 
(perm_simp add: fresh_bij subst_eqvt)+  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
556  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
557  | 
lemma subst_trm_ty_eqvt'[eqvt]:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
558  | 
fixes pi::"vrs prm"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
559  | 
and t::"trm"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
560  | 
shows "pi\<bullet>(t[X \<mapsto>\<^sub>\<tau> T]) = (pi\<bullet>t)[(pi\<bullet>X) \<mapsto>\<^sub>\<tau> (pi\<bullet>T)]"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
561  | 
by (nominal_induct t avoiding: X T rule: trm.strong_induct)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
562  | 
(perm_simp add: fresh_left subst_eqvt')+  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
563  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
564  | 
lemma subst_trm_ty_rename:  | 
| 32011 | 565  | 
"Y\<sharp>t \<Longrightarrow> ([(Y, X)] \<bullet> t)[Y \<mapsto>\<^sub>\<tau> U] = t[X \<mapsto>\<^sub>\<tau> U]"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
566  | 
by (nominal_induct t avoiding: X Y U rule: trm.strong_induct)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
567  | 
(simp_all add: fresh_atm calc_atm abs_fresh fresh_aux type_subst_rename)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
568  | 
|
| 63167 | 569  | 
section \<open>Subtyping-Relation\<close>  | 
| 18246 | 570  | 
|
| 63167 | 571  | 
text \<open>The definition for the subtyping-relation follows quite closely what is written  | 
| 18650 | 572  | 
in the POPLmark-paper, except for the premises dealing with well-formed contexts and  | 
| 69597 | 573  | 
the freshness constraint \<^term>\<open>X\<sharp>\<Gamma>\<close> in the \<open>S_Forall\<close>-rule. (The freshness  | 
| 18650 | 574  | 
  constraint is specific to the \emph{nominal approach}. Note, however, that the constraint
 | 
575  | 
  does \emph{not} make the subtyping-relation ``partial"\ldots because we work over
 | 
|
| 63167 | 576  | 
$\alpha$-equivalence classes.)\<close>  | 
| 18628 | 577  | 
|
| 23760 | 578  | 
inductive  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
579  | 
  subtype_of :: "env \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool"   ("_\<turnstile>_<:_" [100,100,100] 100)
 | 
| 22436 | 580  | 
where  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
581  | 
SA_Top[intro]: "\<lbrakk>\<turnstile> \<Gamma> ok; S closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> S <: Top"  | 
| 32011 | 582  | 
| SA_refl_TVar[intro]: "\<lbrakk>\<turnstile> \<Gamma> ok; X \<in> ty_dom \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> Tvar X <: Tvar X"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
583  | 
| SA_trans_TVar[intro]: "\<lbrakk>(TVarB X S) \<in> set \<Gamma>; \<Gamma> \<turnstile> S <: T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (Tvar X) <: T"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
584  | 
| SA_arrow[intro]: "\<lbrakk>\<Gamma> \<turnstile> T\<^sub>1 <: S\<^sub>1; \<Gamma> \<turnstile> S\<^sub>2 <: T\<^sub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (S\<^sub>1 \<rightarrow> S\<^sub>2) <: (T\<^sub>1 \<rightarrow> T\<^sub>2)"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
585  | 
| SA_all[intro]: "\<lbrakk>\<Gamma> \<turnstile> T\<^sub>1 <: S\<^sub>1; ((TVarB X T\<^sub>1)#\<Gamma>) \<turnstile> S\<^sub>2 <: T\<^sub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (\<forall>X<:S\<^sub>1. S\<^sub>2) <: (\<forall>X<:T\<^sub>1. T\<^sub>2)"  | 
| 22537 | 586  | 
|
587  | 
lemma subtype_implies_ok:  | 
|
588  | 
fixes X::"tyvrs"  | 
|
589  | 
assumes a: "\<Gamma> \<turnstile> S <: T"  | 
|
590  | 
shows "\<turnstile> \<Gamma> ok"  | 
|
591  | 
using a by (induct) (auto)  | 
|
| 18246 | 592  | 
|
593  | 
lemma subtype_implies_closed:  | 
|
594  | 
assumes a: "\<Gamma> \<turnstile> S <: T"  | 
|
595  | 
shows "S closed_in \<Gamma> \<and> T closed_in \<Gamma>"  | 
|
596  | 
using a  | 
|
597  | 
proof (induct)  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
598  | 
case (SA_Top \<Gamma> S)  | 
| 18424 | 599  | 
have "Top closed_in \<Gamma>" by (simp add: closed_in_def ty.supp)  | 
| 18246 | 600  | 
moreover  | 
601  | 
have "S closed_in \<Gamma>" by fact  | 
|
602  | 
ultimately show "S closed_in \<Gamma> \<and> Top closed_in \<Gamma>" by simp  | 
|
603  | 
next  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
604  | 
case (SA_trans_TVar X S \<Gamma> T)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
605  | 
have "(TVarB X S)\<in>set \<Gamma>" by fact  | 
| 32011 | 606  | 
hence "X \<in> ty_dom \<Gamma>" by (rule ty_dom_inclusion)  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
607  | 
hence "(Tvar X) closed_in \<Gamma>" by (simp add: closed_in_def ty.supp supp_atm)  | 
| 18246 | 608  | 
moreover  | 
609  | 
have "S closed_in \<Gamma> \<and> T closed_in \<Gamma>" by fact  | 
|
610  | 
hence "T closed_in \<Gamma>" by force  | 
|
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
611  | 
ultimately show "(Tvar X) closed_in \<Gamma> \<and> T closed_in \<Gamma>" by simp  | 
| 80142 | 612  | 
qed (auto simp: closed_in_def ty.supp supp_atm abs_supp)  | 
| 18246 | 613  | 
|
614  | 
lemma subtype_implies_fresh:  | 
|
615  | 
fixes X::"tyvrs"  | 
|
616  | 
assumes a1: "\<Gamma> \<turnstile> S <: T"  | 
|
617  | 
and a2: "X\<sharp>\<Gamma>"  | 
|
| 18424 | 618  | 
shows "X\<sharp>S \<and> X\<sharp>T"  | 
| 18246 | 619  | 
proof -  | 
620  | 
from a1 have "\<turnstile> \<Gamma> ok" by (rule subtype_implies_ok)  | 
|
| 32011 | 621  | 
with a2 have "X\<sharp>ty_dom(\<Gamma>)" by (simp add: fresh_dom)  | 
| 18424 | 622  | 
moreover  | 
| 18246 | 623  | 
from a1 have "S closed_in \<Gamma> \<and> T closed_in \<Gamma>" by (rule subtype_implies_closed)  | 
| 32011 | 624  | 
hence "supp S \<subseteq> ((supp (ty_dom \<Gamma>))::tyvrs set)"  | 
625  | 
and "supp T \<subseteq> ((supp (ty_dom \<Gamma>))::tyvrs set)" by (simp_all add: ty_dom_supp closed_in_def)  | 
|
| 80142 | 626  | 
ultimately show "X\<sharp>S \<and> X\<sharp>T" by (force simp: supp_prod fresh_def)  | 
| 18246 | 627  | 
qed  | 
628  | 
||
| 32011 | 629  | 
lemma valid_ty_dom_fresh:  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
630  | 
fixes X::"tyvrs"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
631  | 
assumes valid: "\<turnstile> \<Gamma> ok"  | 
| 32011 | 632  | 
shows "X\<sharp>(ty_dom \<Gamma>) = X\<sharp>\<Gamma>"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
633  | 
using valid  | 
| 80142 | 634  | 
proof induct  | 
635  | 
case valid_nil  | 
|
636  | 
then show ?case by auto  | 
|
637  | 
next  | 
|
638  | 
case (valid_consT \<Gamma> X T)  | 
|
639  | 
then show ?case  | 
|
640  | 
by (auto simp: fresh_list_cons closed_in_fresh  | 
|
641  | 
fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_doms)  | 
|
642  | 
next  | 
|
643  | 
case (valid_cons \<Gamma> x T)  | 
|
644  | 
then show ?case  | 
|
645  | 
using fresh_atm by (auto simp: fresh_list_cons closed_in_fresh)  | 
|
646  | 
qed  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
647  | 
|
| 
22730
 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 
berghofe 
parents: 
22542 
diff
changeset
 | 
648  | 
equivariance subtype_of  | 
| 
 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 
berghofe 
parents: 
22542 
diff
changeset
 | 
649  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
650  | 
nominal_inductive subtype_of  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
651  | 
apply (simp_all add: abs_fresh)  | 
| 80142 | 652  | 
apply (fastforce simp: valid_ty_dom_fresh dest: subtype_implies_ok)  | 
653  | 
apply (force simp: closed_in_fresh dest: subtype_implies_closed subtype_implies_ok)+  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
654  | 
done  | 
| 18246 | 655  | 
|
| 63167 | 656  | 
section \<open>Reflexivity of Subtyping\<close>  | 
| 18246 | 657  | 
|
658  | 
lemma subtype_reflexivity:  | 
|
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
659  | 
assumes a: "\<turnstile> \<Gamma> ok"  | 
| 18424 | 660  | 
and b: "T closed_in \<Gamma>"  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
661  | 
shows "\<Gamma> \<turnstile> T <: T"  | 
| 
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
662  | 
using a b  | 
| 
26966
 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 
urbanc 
parents: 
26091 
diff
changeset
 | 
663  | 
proof(nominal_induct T avoiding: \<Gamma> rule: ty.strong_induct)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
664  | 
case (Forall X T\<^sub>1 T\<^sub>2)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
665  | 
have ih_T\<^sub>1: "\<And>\<Gamma>. \<lbrakk>\<turnstile> \<Gamma> ok; T\<^sub>1 closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> T\<^sub>1 <: T\<^sub>1" by fact  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
666  | 
have ih_T\<^sub>2: "\<And>\<Gamma>. \<lbrakk>\<turnstile> \<Gamma> ok; T\<^sub>2 closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> T\<^sub>2 <: T\<^sub>2" by fact  | 
| 18424 | 667  | 
have fresh_cond: "X\<sharp>\<Gamma>" by fact  | 
| 32011 | 668  | 
hence fresh_ty_dom: "X\<sharp>(ty_dom \<Gamma>)" by (simp add: fresh_dom)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
669  | 
have "(\<forall>X<:T\<^sub>2. T\<^sub>1) closed_in \<Gamma>" by fact  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
670  | 
hence closed\<^sub>T2: "T\<^sub>2 closed_in \<Gamma>" and closed\<^sub>T1: "T\<^sub>1 closed_in ((TVarB X T\<^sub>2)#\<Gamma>)"  | 
| 80142 | 671  | 
by (auto simp: closed_in_def ty.supp abs_supp)  | 
| 18424 | 672  | 
have ok: "\<turnstile> \<Gamma> ok" by fact  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
673  | 
hence ok': "\<turnstile> ((TVarB X T\<^sub>2)#\<Gamma>) ok" using closed\<^sub>T2 fresh_ty_dom by simp  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
674  | 
have "\<Gamma> \<turnstile> T\<^sub>2 <: T\<^sub>2" using ih_T\<^sub>2 closed\<^sub>T2 ok by simp  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
675  | 
moreover  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
676  | 
have "((TVarB X T\<^sub>2)#\<Gamma>) \<turnstile> T\<^sub>1 <: T\<^sub>1" using ih_T\<^sub>1 closed\<^sub>T1 ok' by simp  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
677  | 
ultimately show "\<Gamma> \<turnstile> (\<forall>X<:T\<^sub>2. T\<^sub>1) <: (\<forall>X<:T\<^sub>2. T\<^sub>1)" using fresh_cond  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
678  | 
by (simp add: subtype_of.SA_all)  | 
| 80142 | 679  | 
qed (auto simp: closed_in_def ty.supp supp_atm)  | 
| 18246 | 680  | 
|
| 18621 | 681  | 
lemma subtype_reflexivity_semiautomated:  | 
| 
18305
 
a780f9c1538b
changed everything until the interesting transitivity_narrowing
 
urbanc 
parents: 
18269 
diff
changeset
 | 
682  | 
assumes a: "\<turnstile> \<Gamma> ok"  | 
| 
 
a780f9c1538b
changed everything until the interesting transitivity_narrowing
 
urbanc 
parents: 
18269 
diff
changeset
 | 
683  | 
and b: "T closed_in \<Gamma>"  | 
| 
 
a780f9c1538b
changed everything until the interesting transitivity_narrowing
 
urbanc 
parents: 
18269 
diff
changeset
 | 
684  | 
shows "\<Gamma> \<turnstile> T <: T"  | 
| 
 
a780f9c1538b
changed everything until the interesting transitivity_narrowing
 
urbanc 
parents: 
18269 
diff
changeset
 | 
685  | 
using a b  | 
| 
26966
 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 
urbanc 
parents: 
26091 
diff
changeset
 | 
686  | 
apply(nominal_induct T avoiding: \<Gamma> rule: ty.strong_induct)  | 
| 80142 | 687  | 
apply(auto simp: ty.supp abs_supp supp_atm closed_in_def)  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
66453 
diff
changeset
 | 
688  | 
\<comment> \<open>Too bad that this instantiation cannot be found automatically by  | 
| 18621 | 689  | 
  \isakeyword{auto}; \isakeyword{blast} would find it if we had not used 
 | 
| 63167 | 690  | 
an explicit definition for \<open>closed_in_def\<close>.\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
691  | 
apply(drule_tac x="(TVarB tyvrs ty2)#\<Gamma>" in meta_spec)  | 
| 32011 | 692  | 
apply(force dest: fresh_dom simp add: closed_in_def)  | 
| 18246 | 693  | 
done  | 
694  | 
||
| 63167 | 695  | 
section \<open>Weakening\<close>  | 
| 18246 | 696  | 
|
| 63167 | 697  | 
text \<open>In order to prove weakening we introduce the notion of a type-context extending  | 
| 18628 | 698  | 
another. This generalization seems to make the proof for weakening to be  | 
| 63167 | 699  | 
smoother than if we had strictly adhered to the version in the POPLmark-paper.\<close>  | 
| 18246 | 700  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
34915 
diff
changeset
 | 
701  | 
definition extends :: "env \<Rightarrow> env \<Rightarrow> bool" ("_ extends _" [100,100] 100) where
 | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
702  | 
"\<Delta> extends \<Gamma> \<equiv> \<forall>X Q. (TVarB X Q)\<in>set \<Gamma> \<longrightarrow> (TVarB X Q)\<in>set \<Delta>"  | 
| 18246 | 703  | 
|
| 32011 | 704  | 
lemma extends_ty_dom:  | 
| 80142 | 705  | 
assumes "\<Delta> extends \<Gamma>"  | 
| 32011 | 706  | 
shows "ty_dom \<Gamma> \<subseteq> ty_dom \<Delta>"  | 
| 80142 | 707  | 
using assms  | 
708  | 
by (meson extends_def subsetI ty_dom_existence ty_dom_inclusion)  | 
|
| 18246 | 709  | 
|
710  | 
lemma extends_closed:  | 
|
| 80142 | 711  | 
assumes "T closed_in \<Gamma>" and "\<Delta> extends \<Gamma>"  | 
| 18246 | 712  | 
shows "T closed_in \<Delta>"  | 
| 80142 | 713  | 
by (meson assms closed_in_def extends_ty_dom order.trans)  | 
| 18246 | 714  | 
|
| 18424 | 715  | 
lemma extends_memb:  | 
716  | 
assumes a: "\<Delta> extends \<Gamma>"  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
717  | 
and b: "(TVarB X T) \<in> set \<Gamma>"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
718  | 
shows "(TVarB X T) \<in> set \<Delta>"  | 
| 18424 | 719  | 
using a b by (simp add: extends_def)  | 
720  | 
||
| 18246 | 721  | 
lemma weakening:  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
722  | 
assumes a: "\<Gamma> \<turnstile> S <: T"  | 
| 18424 | 723  | 
and b: "\<turnstile> \<Delta> ok"  | 
724  | 
and c: "\<Delta> extends \<Gamma>"  | 
|
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
725  | 
shows "\<Delta> \<turnstile> S <: T"  | 
| 
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
726  | 
using a b c  | 
| 22537 | 727  | 
proof (nominal_induct \<Gamma> S T avoiding: \<Delta> rule: subtype_of.strong_induct)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
728  | 
case (SA_Top \<Gamma> S)  | 
| 18246 | 729  | 
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
 | 
730  | 
have "\<turnstile> \<Delta> ok" by fact  | 
| 
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
731  | 
moreover  | 
| 
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
732  | 
have "\<Delta> extends \<Gamma>" by fact  | 
| 18424 | 733  | 
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
 | 
734  | 
ultimately show "\<Delta> \<turnstile> S <: Top" by force  | 
| 18246 | 735  | 
next  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
736  | 
case (SA_trans_TVar X S \<Gamma> T)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
737  | 
have lh_drv_prem: "(TVarB X S) \<in> set \<Gamma>" by fact  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
738  | 
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
 | 
739  | 
have ok: "\<turnstile> \<Delta> ok" by fact  | 
| 
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
740  | 
have extends: "\<Delta> extends \<Gamma>" by fact  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
741  | 
have "(TVarB 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
 | 
742  | 
moreover  | 
| 
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
743  | 
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
 | 
744  | 
ultimately show "\<Delta> \<turnstile> Tvar X <: T" using ok by force  | 
| 18246 | 745  | 
next  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
746  | 
case (SA_refl_TVar \<Gamma> X)  | 
| 32011 | 747  | 
have lh_drv_prem: "X \<in> ty_dom \<Gamma>" by fact  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
748  | 
have "\<turnstile> \<Delta> ok" by fact  | 
| 
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
749  | 
moreover  | 
| 
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
750  | 
have "\<Delta> extends \<Gamma>" by fact  | 
| 32011 | 751  | 
hence "X \<in> ty_dom \<Delta>" using lh_drv_prem by (force dest: extends_ty_dom)  | 
| 
18577
 
a636846a02c7
added more documentation; will now try out a modification
 
urbanc 
parents: 
18424 
diff
changeset
 | 
752  | 
ultimately show "\<Delta> \<turnstile> Tvar X <: Tvar X" by force  | 
| 18246 | 753  | 
next  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
754  | 
case (SA_arrow \<Gamma> T\<^sub>1 S\<^sub>1 S\<^sub>2 T\<^sub>2) thus "\<Delta> \<turnstile> S\<^sub>1 \<rightarrow> S\<^sub>2 <: T\<^sub>1 \<rightarrow> T\<^sub>2" by blast  | 
| 18246 | 755  | 
next  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
756  | 
case (SA_all \<Gamma> T\<^sub>1 S\<^sub>1 X S\<^sub>2 T\<^sub>2)  | 
| 18424 | 757  | 
have fresh_cond: "X\<sharp>\<Delta>" by fact  | 
| 32011 | 758  | 
hence fresh_dom: "X\<sharp>(ty_dom \<Delta>)" by (simp add: fresh_dom)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
759  | 
have ih\<^sub>1: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends \<Gamma> \<Longrightarrow> \<Delta> \<turnstile> T\<^sub>1 <: S\<^sub>1" by fact  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
760  | 
have ih\<^sub>2: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends ((TVarB X T\<^sub>1)#\<Gamma>) \<Longrightarrow> \<Delta> \<turnstile> S\<^sub>2 <: T\<^sub>2" by fact  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
761  | 
have lh_drv_prem: "\<Gamma> \<turnstile> T\<^sub>1 <: S\<^sub>1" by fact  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
762  | 
hence closed\<^sub>T1: "T\<^sub>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
 | 
763  | 
have ok: "\<turnstile> \<Delta> ok" by fact  | 
| 18424 | 764  | 
have ext: "\<Delta> extends \<Gamma>" by fact  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
765  | 
have "T\<^sub>1 closed_in \<Delta>" using ext closed\<^sub>T1 by (simp only: extends_closed)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
766  | 
hence "\<turnstile> ((TVarB X T\<^sub>1)#\<Delta>) ok" using fresh_dom ok by force  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
767  | 
moreover  | 
| 80142 | 768  | 
have "((TVarB X T\<^sub>1)#\<Delta>) extends ((TVarB X T\<^sub>1)#\<Gamma>)" using ext by (force simp: extends_def)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
769  | 
ultimately have "((TVarB X T\<^sub>1)#\<Delta>) \<turnstile> S\<^sub>2 <: T\<^sub>2" using ih\<^sub>2 by simp  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
770  | 
moreover  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
771  | 
have "\<Delta> \<turnstile> T\<^sub>1 <: S\<^sub>1" using ok ext ih\<^sub>1 by simp  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
772  | 
ultimately show "\<Delta> \<turnstile> (\<forall>X<:S\<^sub>1. S\<^sub>2) <: (\<forall>X<:T\<^sub>1. T\<^sub>2)" using ok by (force intro: SA_all)  | 
| 18246 | 773  | 
qed  | 
774  | 
||
| 63167 | 775  | 
text \<open>In fact all ``non-binding" cases can be solved automatically:\<close>  | 
| 18246 | 776  | 
|
| 18628 | 777  | 
lemma weakening_more_automated:  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
778  | 
assumes a: "\<Gamma> \<turnstile> S <: T"  | 
| 18424 | 779  | 
and b: "\<turnstile> \<Delta> ok"  | 
780  | 
and c: "\<Delta> extends \<Gamma>"  | 
|
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
781  | 
shows "\<Delta> \<turnstile> S <: T"  | 
| 
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
782  | 
using a b c  | 
| 22537 | 783  | 
proof (nominal_induct \<Gamma> S T avoiding: \<Delta> rule: subtype_of.strong_induct)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
784  | 
case (SA_all \<Gamma> T\<^sub>1 S\<^sub>1 X S\<^sub>2 T\<^sub>2)  | 
| 18424 | 785  | 
have fresh_cond: "X\<sharp>\<Delta>" by fact  | 
| 32011 | 786  | 
hence fresh_dom: "X\<sharp>(ty_dom \<Delta>)" by (simp add: fresh_dom)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
787  | 
have ih\<^sub>1: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends \<Gamma> \<Longrightarrow> \<Delta> \<turnstile> T\<^sub>1 <: S\<^sub>1" by fact  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
788  | 
have ih\<^sub>2: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends ((TVarB X T\<^sub>1)#\<Gamma>) \<Longrightarrow> \<Delta> \<turnstile> S\<^sub>2 <: T\<^sub>2" by fact  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
789  | 
have lh_drv_prem: "\<Gamma> \<turnstile> T\<^sub>1 <: S\<^sub>1" by fact  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
790  | 
hence closed\<^sub>T1: "T\<^sub>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
 | 
791  | 
have ok: "\<turnstile> \<Delta> ok" by fact  | 
| 18424 | 792  | 
have ext: "\<Delta> extends \<Gamma>" by fact  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
793  | 
have "T\<^sub>1 closed_in \<Delta>" using ext closed\<^sub>T1 by (simp only: extends_closed)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
794  | 
hence "\<turnstile> ((TVarB X T\<^sub>1)#\<Delta>) ok" using fresh_dom ok by force  | 
| 18628 | 795  | 
moreover  | 
| 80142 | 796  | 
have "((TVarB X T\<^sub>1)#\<Delta>) extends ((TVarB X T\<^sub>1)#\<Gamma>)" using ext by (force simp: extends_def)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
797  | 
ultimately have "((TVarB X T\<^sub>1)#\<Delta>) \<turnstile> S\<^sub>2 <: T\<^sub>2" using ih\<^sub>2 by simp  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
798  | 
moreover  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
799  | 
have "\<Delta> \<turnstile> T\<^sub>1 <: S\<^sub>1" using ok ext ih\<^sub>1 by simp  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
800  | 
ultimately show "\<Delta> \<turnstile> (\<forall>X<:S\<^sub>1. S\<^sub>2) <: (\<forall>X<:T\<^sub>1. T\<^sub>2)" using ok by (force intro: SA_all)  | 
| 32011 | 801  | 
qed (blast intro: extends_closed extends_memb dest: extends_ty_dom)+  | 
| 18246 | 802  | 
|
| 63167 | 803  | 
section \<open>Transitivity and Narrowing\<close>  | 
| 18628 | 804  | 
|
| 63167 | 805  | 
text \<open>Some inversion lemmas that are needed in the transitivity and narrowing proof.\<close>  | 
| 18650 | 806  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
807  | 
declare ty.inject [simp add]  | 
| 18650 | 808  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
809  | 
inductive_cases S_TopE: "\<Gamma> \<turnstile> Top <: T"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
810  | 
inductive_cases S_ArrowE_left: "\<Gamma> \<turnstile> S\<^sub>1 \<rightarrow> S\<^sub>2 <: T"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
811  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
812  | 
declare ty.inject [simp del]  | 
| 18650 | 813  | 
|
814  | 
lemma S_ForallE_left:  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
815  | 
shows "\<lbrakk>\<Gamma> \<turnstile> (\<forall>X<:S\<^sub>1. S\<^sub>2) <: T; X\<sharp>\<Gamma>; X\<sharp>S\<^sub>1; X\<sharp>T\<rbrakk>  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
816  | 
\<Longrightarrow> T = Top \<or> (\<exists>T\<^sub>1 T\<^sub>2. T = (\<forall>X<:T\<^sub>1. T\<^sub>2) \<and> \<Gamma> \<turnstile> T\<^sub>1 <: S\<^sub>1 \<and> ((TVarB X T\<^sub>1)#\<Gamma>) \<turnstile> S\<^sub>2 <: T\<^sub>2)"  | 
| 80142 | 817  | 
using subtype_of.strong_cases[where X="X"]  | 
818  | 
by(force simp: abs_fresh ty.inject alpha)  | 
|
| 18650 | 819  | 
|
| 63167 | 820  | 
text \<open>Next we prove the transitivity and narrowing for the subtyping-relation.  | 
| 18621 | 821  | 
The POPLmark-paper says the following:  | 
822  | 
||
| 18650 | 823  | 
\begin{quote}
 | 
| 18621 | 824  | 
\begin{lemma}[Transitivity and Narrowing] \
 | 
825  | 
\begin{enumerate}
 | 
|
| 69597 | 826  | 
\item If \<^term>\<open>\<Gamma> \<turnstile> S<:Q\<close> and \<^term>\<open>\<Gamma> \<turnstile> Q<:T\<close>, then \<^term>\<open>\<Gamma> \<turnstile> S<:T\<close>.  | 
827  | 
\item If \<open>\<Gamma>,X<:Q,\<Delta> \<turnstile> M<:N\<close> and \<^term>\<open>\<Gamma> \<turnstile> P<:Q\<close> then \<open>\<Gamma>,X<:P,\<Delta> \<turnstile> M<:N\<close>.  | 
|
| 18621 | 828  | 
\end{enumerate}
 | 
829  | 
\end{lemma}
 | 
|
830  | 
||
831  | 
The two parts are proved simultaneously, by induction on the size  | 
|
| 69597 | 832  | 
of \<^term>\<open>Q\<close>. The argument for part (2) assumes that part (1) has  | 
833  | 
been established already for the \<^term>\<open>Q\<close> in question; part (1) uses  | 
|
834  | 
part (2) only for strictly smaller \<^term>\<open>Q\<close>.  | 
|
| 18650 | 835  | 
\end{quote}
 | 
| 18621 | 836  | 
|
| 69597 | 837  | 
For the induction on the size of \<^term>\<open>Q\<close>, we use the induction-rule  | 
| 63167 | 838  | 
\<open>measure_induct_rule\<close>:  | 
| 18621 | 839  | 
|
840  | 
\begin{center}
 | 
|
841  | 
@{thm measure_induct_rule[of "size_ty",no_vars]}
 | 
|
842  | 
\end{center}
 | 
|
| 18410 | 843  | 
|
| 69597 | 844  | 
That means in order to show a property \<^term>\<open>P a\<close> for all \<^term>\<open>a\<close>,  | 
845  | 
the induct-rule requires to prove that for all \<^term>\<open>x\<close> \<^term>\<open>P x\<close> holds using the  | 
|
846  | 
assumption that for all \<^term>\<open>y\<close> whose size is strictly smaller than  | 
|
847  | 
that of \<^term>\<open>x\<close> the property \<^term>\<open>P y\<close> holds.\<close>  | 
|
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
848  | 
|
| 18621 | 849  | 
lemma  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
850  | 
shows subtype_transitivity: "\<Gamma>\<turnstile>S<:Q \<Longrightarrow> \<Gamma>\<turnstile>Q<:T \<Longrightarrow> \<Gamma>\<turnstile>S<:T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
851  | 
and subtype_narrow: "(\<Delta>@[(TVarB X Q)]@\<Gamma>)\<turnstile>M<:N \<Longrightarrow> \<Gamma>\<turnstile>P<:Q \<Longrightarrow> (\<Delta>@[(TVarB X P)]@\<Gamma>)\<turnstile>M<:N"  | 
| 20503 | 852  | 
proof (induct Q arbitrary: \<Gamma> S T \<Delta> X P M N taking: "size_ty" rule: measure_induct_rule)  | 
| 18621 | 853  | 
case (less Q)  | 
854  | 
have IH_trans:  | 
|
855  | 
"\<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  | 
|
856  | 
have IH_narrow:  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
857  | 
"\<And>Q' \<Delta> \<Gamma> X M N P. \<lbrakk>size_ty Q' < size_ty Q; (\<Delta>@[(TVarB X Q')]@\<Gamma>)\<turnstile>M<:N; \<Gamma>\<turnstile>P<:Q'\<rbrakk>  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
858  | 
\<Longrightarrow> (\<Delta>@[(TVarB X P)]@\<Gamma>)\<turnstile>M<:N" by fact  | 
| 32011 | 859  | 
|
860  | 
  { fix \<Gamma> S T
 | 
|
861  | 
have "\<lbrakk>\<Gamma> \<turnstile> S <: Q; \<Gamma> \<turnstile> Q <: T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> S <: T"  | 
|
862  | 
proof (induct \<Gamma> S Q\<equiv>Q rule: subtype_of.induct)  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
863  | 
case (SA_Top \<Gamma> S)  | 
| 32011 | 864  | 
then have rh_drv: "\<Gamma> \<turnstile> Top <: T" by simp  | 
865  | 
then have T_inst: "T = Top" by (auto elim: S_TopE)  | 
|
| 63167 | 866  | 
from \<open>\<turnstile> \<Gamma> ok\<close> and \<open>S closed_in \<Gamma>\<close>  | 
| 32011 | 867  | 
have "\<Gamma> \<turnstile> S <: Top" by auto  | 
868  | 
then show "\<Gamma> \<turnstile> S <: T" using T_inst by simp  | 
|
| 18246 | 869  | 
next  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
870  | 
case (SA_trans_TVar Y U \<Gamma>)  | 
| 32011 | 871  | 
then have IH_inner: "\<Gamma> \<turnstile> U <: T" by simp  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
872  | 
have "(TVarB Y U) \<in> set \<Gamma>" by fact  | 
| 32011 | 873  | 
with IH_inner show "\<Gamma> \<turnstile> Tvar Y <: T" by auto  | 
| 18246 | 874  | 
next  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
875  | 
case (SA_refl_TVar \<Gamma> X)  | 
| 32011 | 876  | 
then show "\<Gamma> \<turnstile> Tvar X <: T" by simp  | 
| 18246 | 877  | 
next  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
878  | 
case (SA_arrow \<Gamma> Q\<^sub>1 S\<^sub>1 S\<^sub>2 Q\<^sub>2)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
879  | 
then have rh_drv: "\<Gamma> \<turnstile> Q\<^sub>1 \<rightarrow> Q\<^sub>2 <: T" by simp  | 
| 63167 | 880  | 
from \<open>Q\<^sub>1 \<rightarrow> Q\<^sub>2 = Q\<close>  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
881  | 
have Q\<^sub>12_less: "size_ty Q\<^sub>1 < size_ty Q" "size_ty Q\<^sub>2 < size_ty Q" by auto  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
882  | 
have lh_drv_prm\<^sub>1: "\<Gamma> \<turnstile> Q\<^sub>1 <: S\<^sub>1" by fact  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
883  | 
have lh_drv_prm\<^sub>2: "\<Gamma> \<turnstile> S\<^sub>2 <: Q\<^sub>2" by fact  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
884  | 
from rh_drv have "T=Top \<or> (\<exists>T\<^sub>1 T\<^sub>2. T=T\<^sub>1\<rightarrow>T\<^sub>2 \<and> \<Gamma>\<turnstile>T\<^sub>1<:Q\<^sub>1 \<and> \<Gamma>\<turnstile>Q\<^sub>2<:T\<^sub>2)"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
885  | 
by (auto elim: S_ArrowE_left)  | 
| 18621 | 886  | 
moreover  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
887  | 
have "S\<^sub>1 closed_in \<Gamma>" and "S\<^sub>2 closed_in \<Gamma>"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
888  | 
using lh_drv_prm\<^sub>1 lh_drv_prm\<^sub>2 by (simp_all add: subtype_implies_closed)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
889  | 
hence "(S\<^sub>1 \<rightarrow> S\<^sub>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
 | 
890  | 
moreover  | 
| 18424 | 891  | 
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
 | 
892  | 
moreover  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
893  | 
      { assume "\<exists>T\<^sub>1 T\<^sub>2. T = T\<^sub>1\<rightarrow>T\<^sub>2 \<and> \<Gamma> \<turnstile> T\<^sub>1 <: Q\<^sub>1 \<and> \<Gamma> \<turnstile> Q\<^sub>2 <: T\<^sub>2"
 | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
894  | 
then obtain T\<^sub>1 T\<^sub>2  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
895  | 
where T_inst: "T = T\<^sub>1 \<rightarrow> T\<^sub>2"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
896  | 
and rh_drv_prm\<^sub>1: "\<Gamma> \<turnstile> T\<^sub>1 <: Q\<^sub>1"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
897  | 
and rh_drv_prm\<^sub>2: "\<Gamma> \<turnstile> Q\<^sub>2 <: T\<^sub>2" by force  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
898  | 
from IH_trans[of "Q\<^sub>1"]  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
899  | 
have "\<Gamma> \<turnstile> T\<^sub>1 <: S\<^sub>1" using Q\<^sub>12_less rh_drv_prm\<^sub>1 lh_drv_prm\<^sub>1 by simp  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
900  | 
moreover  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
901  | 
from IH_trans[of "Q\<^sub>2"]  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
902  | 
have "\<Gamma> \<turnstile> S\<^sub>2 <: T\<^sub>2" using Q\<^sub>12_less rh_drv_prm\<^sub>2 lh_drv_prm\<^sub>2 by simp  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
903  | 
ultimately have "\<Gamma> \<turnstile> S\<^sub>1 \<rightarrow> S\<^sub>2 <: T\<^sub>1 \<rightarrow> T\<^sub>2" by auto  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
904  | 
then have "\<Gamma> \<turnstile> S\<^sub>1 \<rightarrow> S\<^sub>2 <: T" using T_inst by simp  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
905  | 
}  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
906  | 
ultimately show "\<Gamma> \<turnstile> S\<^sub>1 \<rightarrow> S\<^sub>2 <: T" by blast  | 
| 18246 | 907  | 
next  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
908  | 
case (SA_all \<Gamma> Q\<^sub>1 S\<^sub>1 X S\<^sub>2 Q\<^sub>2)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
909  | 
then have rh_drv: "\<Gamma> \<turnstile> (\<forall>X<:Q\<^sub>1. Q\<^sub>2) <: T" by simp  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
910  | 
have lh_drv_prm\<^sub>1: "\<Gamma> \<turnstile> Q\<^sub>1 <: S\<^sub>1" by fact  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
911  | 
have lh_drv_prm\<^sub>2: "((TVarB X Q\<^sub>1)#\<Gamma>) \<turnstile> S\<^sub>2 <: Q\<^sub>2" by fact  | 
| 32011 | 912  | 
then have "X\<sharp>\<Gamma>" by (force dest: subtype_implies_ok simp add: valid_ty_dom_fresh)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
913  | 
then have fresh_cond: "X\<sharp>\<Gamma>" "X\<sharp>Q\<^sub>1" "X\<sharp>T" using rh_drv lh_drv_prm\<^sub>1  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
914  | 
by (simp_all add: subtype_implies_fresh)  | 
| 18621 | 915  | 
from rh_drv  | 
| 32011 | 916  | 
have "T = Top \<or>  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
917  | 
(\<exists>T\<^sub>1 T\<^sub>2. T = (\<forall>X<:T\<^sub>1. T\<^sub>2) \<and> \<Gamma> \<turnstile> T\<^sub>1 <: Q\<^sub>1 \<and> ((TVarB X T\<^sub>1)#\<Gamma>) \<turnstile> Q\<^sub>2 <: T\<^sub>2)"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
918  | 
using fresh_cond by (simp add: S_ForallE_left)  | 
| 18621 | 919  | 
moreover  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
920  | 
have "S\<^sub>1 closed_in \<Gamma>" and "S\<^sub>2 closed_in ((TVarB X Q\<^sub>1)#\<Gamma>)"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
921  | 
using lh_drv_prm\<^sub>1 lh_drv_prm\<^sub>2 by (simp_all add: subtype_implies_closed)  | 
| 80142 | 922  | 
then have "(\<forall>X<:S\<^sub>1. S\<^sub>2) closed_in \<Gamma>" by (force simp: closed_in_def ty.supp abs_supp)  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
923  | 
moreover  | 
| 18424 | 924  | 
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
 | 
925  | 
moreover  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
926  | 
      { assume "\<exists>T\<^sub>1 T\<^sub>2. T=(\<forall>X<:T\<^sub>1. T\<^sub>2) \<and> \<Gamma>\<turnstile>T\<^sub>1<:Q\<^sub>1 \<and> ((TVarB X T\<^sub>1)#\<Gamma>)\<turnstile>Q\<^sub>2<:T\<^sub>2"
 | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
927  | 
then obtain T\<^sub>1 T\<^sub>2  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
928  | 
where T_inst: "T = (\<forall>X<:T\<^sub>1. T\<^sub>2)"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
929  | 
and rh_drv_prm\<^sub>1: "\<Gamma> \<turnstile> T\<^sub>1 <: Q\<^sub>1"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
930  | 
and rh_drv_prm\<^sub>2:"((TVarB X T\<^sub>1)#\<Gamma>) \<turnstile> Q\<^sub>2 <: T\<^sub>2" by force  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
931  | 
have "(\<forall>X<:Q\<^sub>1. Q\<^sub>2) = Q" by fact  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
932  | 
then have Q\<^sub>12_less: "size_ty Q\<^sub>1 < size_ty Q" "size_ty Q\<^sub>2 < size_ty Q"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
933  | 
using fresh_cond by auto  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
934  | 
from IH_trans[of "Q\<^sub>1"]  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
935  | 
have "\<Gamma> \<turnstile> T\<^sub>1 <: S\<^sub>1" using lh_drv_prm\<^sub>1 rh_drv_prm\<^sub>1 Q\<^sub>12_less by blast  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
936  | 
moreover  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
937  | 
from IH_narrow[of "Q\<^sub>1" "[]"]  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
938  | 
have "((TVarB X T\<^sub>1)#\<Gamma>) \<turnstile> S\<^sub>2 <: Q\<^sub>2" using Q\<^sub>12_less lh_drv_prm\<^sub>2 rh_drv_prm\<^sub>1 by simp  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
939  | 
with IH_trans[of "Q\<^sub>2"]  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
940  | 
have "((TVarB X T\<^sub>1)#\<Gamma>) \<turnstile> S\<^sub>2 <: T\<^sub>2" using Q\<^sub>12_less rh_drv_prm\<^sub>2 by simp  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
941  | 
ultimately have "\<Gamma> \<turnstile> (\<forall>X<:S\<^sub>1. S\<^sub>2) <: (\<forall>X<:T\<^sub>1. T\<^sub>2)"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
942  | 
using fresh_cond by (simp add: subtype_of.SA_all)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
943  | 
hence "\<Gamma> \<turnstile> (\<forall>X<:S\<^sub>1. S\<^sub>2) <: T" using T_inst by simp  | 
| 
18353
 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 
urbanc 
parents: 
18306 
diff
changeset
 | 
944  | 
}  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
945  | 
ultimately show "\<Gamma> \<turnstile> (\<forall>X<:S\<^sub>1. S\<^sub>2) <: T" by blast  | 
| 18246 | 946  | 
qed  | 
| 32011 | 947  | 
} note transitivity_lemma = this  | 
| 18246 | 948  | 
|
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
66453 
diff
changeset
 | 
949  | 
  { \<comment> \<open>The transitivity proof is now by the auxiliary lemma.\<close>
 | 
| 18621 | 950  | 
case 1  | 
| 63167 | 951  | 
from \<open>\<Gamma> \<turnstile> S <: Q\<close> and \<open>\<Gamma> \<turnstile> Q <: T\<close>  | 
| 32011 | 952  | 
show "\<Gamma> \<turnstile> S <: T" by (rule transitivity_lemma)  | 
| 18621 | 953  | 
next  | 
954  | 
case 2  | 
|
| 63167 | 955  | 
from \<open>(\<Delta>@[(TVarB X Q)]@\<Gamma>) \<turnstile> M <: N\<close>  | 
956  | 
and \<open>\<Gamma> \<turnstile> P<:Q\<close>  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
957  | 
show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> M <: N"  | 
| 34915 | 958  | 
proof (induct "\<Delta>@[(TVarB X Q)]@\<Gamma>" M N arbitrary: \<Gamma> X \<Delta> rule: subtype_of.induct)  | 
959  | 
case (SA_Top S \<Gamma> X \<Delta>)  | 
|
| 63167 | 960  | 
from \<open>\<Gamma> \<turnstile> P <: Q\<close>  | 
| 34915 | 961  | 
have "P closed_in \<Gamma>" by (simp add: subtype_implies_closed)  | 
| 63167 | 962  | 
with \<open>\<turnstile> (\<Delta>@[(TVarB X Q)]@\<Gamma>) ok\<close> have "\<turnstile> (\<Delta>@[(TVarB X P)]@\<Gamma>) ok"  | 
| 34915 | 963  | 
by (simp add: replace_type)  | 
| 18412 | 964  | 
moreover  | 
| 63167 | 965  | 
from \<open>S closed_in (\<Delta>@[(TVarB X Q)]@\<Gamma>)\<close> have "S closed_in (\<Delta>@[(TVarB X P)]@\<Gamma>)"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
966  | 
by (simp add: closed_in_def doms_append)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
967  | 
ultimately show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> S <: Top" by (simp add: subtype_of.SA_Top)  | 
| 18246 | 968  | 
next  | 
| 34915 | 969  | 
case (SA_trans_TVar Y S N \<Gamma> X \<Delta>)  | 
| 32011 | 970  | 
then have IH_inner: "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> S <: N"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
971  | 
and lh_drv_prm: "(TVarB Y S) \<in> set (\<Delta>@[(TVarB X Q)]@\<Gamma>)"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
972  | 
and rh_drv: "\<Gamma> \<turnstile> P<:Q"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
973  | 
and ok\<^sub>Q: "\<turnstile> (\<Delta>@[(TVarB X Q)]@\<Gamma>) ok" by (simp_all add: subtype_implies_ok)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
974  | 
then have ok\<^sub>P: "\<turnstile> (\<Delta>@[(TVarB X P)]@\<Gamma>) ok" by (simp add: subtype_implies_ok)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
975  | 
show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: N"  | 
| 18621 | 976  | 
proof (cases "X=Y")  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
977  | 
case False  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
978  | 
have "X\<noteq>Y" by fact  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
979  | 
hence "(TVarB Y S)\<in>set (\<Delta>@[(TVarB X P)]@\<Gamma>)" using lh_drv_prm by (simp add:binding.inject)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
980  | 
with IH_inner show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: N" by (simp add: subtype_of.SA_trans_TVar)  | 
| 18621 | 981  | 
next  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
982  | 
case True  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
983  | 
have memb\<^sub>XQ: "(TVarB X Q)\<in>set (\<Delta>@[(TVarB X Q)]@\<Gamma>)" by simp  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
984  | 
have memb\<^sub>XP: "(TVarB X P)\<in>set (\<Delta>@[(TVarB X P)]@\<Gamma>)" by simp  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
985  | 
have eq: "X=Y" by fact  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
986  | 
hence "S=Q" using ok\<^sub>Q lh_drv_prm memb\<^sub>XQ by (simp only: uniqueness_of_ctxt)  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
987  | 
hence "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Q <: N" using IH_inner by simp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
988  | 
moreover  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
989  | 
have "(\<Delta>@[(TVarB X P)]@\<Gamma>) extends \<Gamma>" by (simp add: extends_def)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
990  | 
hence "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> P <: Q" using rh_drv ok\<^sub>P by (simp only: weakening)  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
991  | 
ultimately have "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> P <: N" by (simp add: transitivity_lemma)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
992  | 
then show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: N" using memb\<^sub>XP eq by auto  | 
| 18621 | 993  | 
qed  | 
| 18246 | 994  | 
next  | 
| 34915 | 995  | 
case (SA_refl_TVar Y \<Gamma> X \<Delta>)  | 
| 63167 | 996  | 
from \<open>\<Gamma> \<turnstile> P <: Q\<close>  | 
| 34915 | 997  | 
have "P closed_in \<Gamma>" by (simp add: subtype_implies_closed)  | 
| 63167 | 998  | 
with \<open>\<turnstile> (\<Delta>@[(TVarB X Q)]@\<Gamma>) ok\<close> have "\<turnstile> (\<Delta>@[(TVarB X P)]@\<Gamma>) ok"  | 
| 34915 | 999  | 
by (simp add: replace_type)  | 
| 18424 | 1000  | 
moreover  | 
| 63167 | 1001  | 
from \<open>Y \<in> ty_dom (\<Delta>@[(TVarB X Q)]@\<Gamma>)\<close> have "Y \<in> ty_dom (\<Delta>@[(TVarB X P)]@\<Gamma>)"  | 
| 34915 | 1002  | 
by (simp add: doms_append)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1003  | 
ultimately show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: Tvar Y" by (simp add: subtype_of.SA_refl_TVar)  | 
| 18246 | 1004  | 
next  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1005  | 
case (SA_arrow S\<^sub>1 Q\<^sub>1 Q\<^sub>2 S\<^sub>2 \<Gamma> X \<Delta>)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1006  | 
then show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Q\<^sub>1 \<rightarrow> Q\<^sub>2 <: S\<^sub>1 \<rightarrow> S\<^sub>2" by blast  | 
| 18424 | 1007  | 
next  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1008  | 
case (SA_all T\<^sub>1 S\<^sub>1 Y S\<^sub>2 T\<^sub>2 \<Gamma> X \<Delta>)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1009  | 
have IH_inner\<^sub>1: "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> T\<^sub>1 <: S\<^sub>1"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1010  | 
and IH_inner\<^sub>2: "(((TVarB Y T\<^sub>1)#\<Delta>)@[(TVarB X P)]@\<Gamma>) \<turnstile> S\<^sub>2 <: T\<^sub>2"  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
41798 
diff
changeset
 | 
1011  | 
by (fastforce intro: SA_all)+  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1012  | 
then show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> (\<forall>Y<:S\<^sub>1. S\<^sub>2) <: (\<forall>Y<:T\<^sub>1. T\<^sub>2)" by auto  | 
| 18246 | 1013  | 
qed  | 
| 18621 | 1014  | 
}  | 
| 18246 | 1015  | 
qed  | 
1016  | 
||
| 63167 | 1017  | 
section \<open>Typing\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1018  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1019  | 
inductive  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1020  | 
  typing :: "env \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _" [60,60,60] 60) 
 | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1021  | 
where  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1022  | 
T_Var[intro]: "\<lbrakk> VarB x T \<in> set \<Gamma>; \<turnstile> \<Gamma> ok \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1023  | 
| T_App[intro]: "\<lbrakk> \<Gamma> \<turnstile> t\<^sub>1 : T\<^sub>1 \<rightarrow> T\<^sub>2; \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>1 \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> t\<^sub>1 \<cdot> t\<^sub>2 : T\<^sub>2"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1024  | 
| T_Abs[intro]: "\<lbrakk> VarB x T\<^sub>1 # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2 \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (\<lambda>x:T\<^sub>1. t\<^sub>2) : T\<^sub>1 \<rightarrow> T\<^sub>2"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1025  | 
| T_Sub[intro]: "\<lbrakk> \<Gamma> \<turnstile> t : S; \<Gamma> \<turnstile> S <: T \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> t : T"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1026  | 
| T_TAbs[intro]:"\<lbrakk> TVarB X T\<^sub>1 # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2 \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (\<lambda>X<:T\<^sub>1. t\<^sub>2) : (\<forall>X<:T\<^sub>1. T\<^sub>2)"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1027  | 
| T_TApp[intro]:"\<lbrakk>X\<sharp>(\<Gamma>,t\<^sub>1,T\<^sub>2); \<Gamma> \<turnstile> t\<^sub>1 : (\<forall>X<:T\<^sub>11. T\<^sub>12); \<Gamma> \<turnstile> T\<^sub>2 <: T\<^sub>11\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> t\<^sub>1 \<cdot>\<^sub>\<tau> T\<^sub>2 : (T\<^sub>12[X \<mapsto> T\<^sub>2]\<^sub>\<tau>)"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1028  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1029  | 
equivariance typing  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1030  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1031  | 
lemma better_T_TApp:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1032  | 
assumes H1: "\<Gamma> \<turnstile> t\<^sub>1 : (\<forall>X<:T11. T12)"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1033  | 
and H2: "\<Gamma> \<turnstile> T2 <: T11"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1034  | 
shows "\<Gamma> \<turnstile> t\<^sub>1 \<cdot>\<^sub>\<tau> T2 : (T12[X \<mapsto> T2]\<^sub>\<tau>)"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1035  | 
proof -  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1036  | 
obtain Y::tyvrs where Y: "Y \<sharp> (X, T12, \<Gamma>, t\<^sub>1, T2)"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1037  | 
by (rule exists_fresh) (rule fin_supp)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1038  | 
then have "Y \<sharp> (\<Gamma>, t\<^sub>1, T2)" by simp  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1039  | 
moreover from Y have "(\<forall>X<:T11. T12) = (\<forall>Y<:T11. [(Y, X)] \<bullet> T12)"  | 
| 80142 | 1040  | 
by (auto simp: ty.inject alpha' fresh_prod fresh_atm)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1041  | 
with H1 have "\<Gamma> \<turnstile> t\<^sub>1 : (\<forall>Y<:T11. [(Y, X)] \<bullet> T12)" by simp  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1042  | 
ultimately have "\<Gamma> \<turnstile> t\<^sub>1 \<cdot>\<^sub>\<tau> T2 : (([(Y, X)] \<bullet> T12)[Y \<mapsto> T2]\<^sub>\<tau>)" using H2  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1043  | 
by (rule T_TApp)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1044  | 
with Y show ?thesis by (simp add: type_subst_rename)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1045  | 
qed  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1046  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1047  | 
lemma typing_ok:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1048  | 
assumes "\<Gamma> \<turnstile> t : T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1049  | 
shows "\<turnstile> \<Gamma> ok"  | 
| 49171 | 1050  | 
using assms by (induct) (auto)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1051  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1052  | 
nominal_inductive typing  | 
| 32011 | 1053  | 
by (auto dest!: typing_ok intro: closed_in_fresh fresh_dom type_subst_fresh  | 
1054  | 
simp: abs_fresh fresh_type_subst_fresh ty_vrs_fresh valid_ty_dom_fresh fresh_trm_dom)  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1055  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1056  | 
lemma ok_imp_VarB_closed_in:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1057  | 
assumes ok: "\<turnstile> \<Gamma> ok"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1058  | 
shows "VarB x T \<in> set \<Gamma> \<Longrightarrow> T closed_in \<Gamma>" using ok  | 
| 80142 | 1059  | 
by induct (auto simp: binding.inject closed_in_def)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1060  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1061  | 
lemma tyvrs_of_subst: "tyvrs_of (B[X \<mapsto> T]\<^sub>b) = tyvrs_of B"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1062  | 
by (nominal_induct B rule: binding.strong_induct) simp_all  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1063  | 
|
| 32011 | 1064  | 
lemma ty_dom_subst: "ty_dom (\<Gamma>[X \<mapsto> T]\<^sub>e) = ty_dom \<Gamma>"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1065  | 
by (induct \<Gamma>) (simp_all add: tyvrs_of_subst)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1066  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1067  | 
lemma vrs_of_subst: "vrs_of (B[X \<mapsto> T]\<^sub>b) = vrs_of B"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1068  | 
by (nominal_induct B rule: binding.strong_induct) simp_all  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1069  | 
|
| 32011 | 1070  | 
lemma trm_dom_subst: "trm_dom (\<Gamma>[X \<mapsto> T]\<^sub>e) = trm_dom \<Gamma>"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1071  | 
by (induct \<Gamma>) (simp_all add: vrs_of_subst)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1072  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1073  | 
lemma subst_closed_in:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1074  | 
"T closed_in (\<Delta> @ TVarB X S # \<Gamma>) \<Longrightarrow> U closed_in \<Gamma> \<Longrightarrow> T[X \<mapsto> U]\<^sub>\<tau> closed_in (\<Delta>[X \<mapsto> U]\<^sub>e @ \<Gamma>)"  | 
| 80142 | 1075  | 
proof (nominal_induct T avoiding: X U \<Gamma> rule: ty.strong_induct)  | 
1076  | 
case (Tvar tyvrs)  | 
|
1077  | 
then show ?case  | 
|
1078  | 
by (auto simp add: closed_in_def ty.supp supp_atm doms_append ty_dom_subst)  | 
|
1079  | 
next  | 
|
1080  | 
case Top  | 
|
1081  | 
then show ?case  | 
|
1082  | 
using closed_in_def fresh_def by fastforce  | 
|
1083  | 
next  | 
|
1084  | 
case (Arrow ty1 ty2)  | 
|
1085  | 
then show ?case  | 
|
1086  | 
by (simp add: closed_in_def ty.supp)  | 
|
1087  | 
next  | 
|
1088  | 
case (Forall tyvrs ty1 ty2)  | 
|
1089  | 
then have "supp (ty1[X \<mapsto> U]\<^sub>\<tau>) \<subseteq> ty_dom (\<Delta>[X \<mapsto> U]\<^sub>e @ TVarB tyvrs ty2 # \<Gamma>)"  | 
|
1090  | 
apply (simp add: closed_in_def ty.supp abs_supp)  | 
|
1091  | 
by (metis Diff_subset_conv Un_left_commute doms_append(1) le_supI2 ty_dom.simps(2) tyvrs_of.simps(2))  | 
|
1092  | 
with Forall show ?case  | 
|
1093  | 
by (auto simp add: closed_in_def ty.supp abs_supp doms_append)  | 
|
1094  | 
qed  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1095  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1096  | 
lemmas subst_closed_in' = subst_closed_in [where \<Delta>="[]", simplified]  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1097  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1098  | 
lemma typing_closed_in:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1099  | 
assumes "\<Gamma> \<turnstile> t : T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1100  | 
shows "T closed_in \<Gamma>"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1101  | 
using assms  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1102  | 
proof induct  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1103  | 
case (T_Var x T \<Gamma>)  | 
| 63167 | 1104  | 
from \<open>\<turnstile> \<Gamma> ok\<close> and \<open>VarB x T \<in> set \<Gamma>\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1105  | 
show ?case by (rule ok_imp_VarB_closed_in)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1106  | 
next  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1107  | 
case (T_App \<Gamma> t\<^sub>1 T\<^sub>1 T\<^sub>2 t\<^sub>2)  | 
| 80142 | 1108  | 
then show ?case by (auto simp: ty.supp closed_in_def)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1109  | 
next  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1110  | 
case (T_Abs x T\<^sub>1 \<Gamma> t\<^sub>2 T\<^sub>2)  | 
| 63167 | 1111  | 
from \<open>VarB x T\<^sub>1 # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2\<close>  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1112  | 
have "T\<^sub>1 closed_in \<Gamma>" by (auto dest: typing_ok)  | 
| 80142 | 1113  | 
with T_Abs show ?case by (auto simp: ty.supp closed_in_def)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1114  | 
next  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1115  | 
case (T_Sub \<Gamma> t S T)  | 
| 63167 | 1116  | 
from \<open>\<Gamma> \<turnstile> S <: T\<close> show ?case by (simp add: subtype_implies_closed)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1117  | 
next  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1118  | 
case (T_TAbs X T\<^sub>1 \<Gamma> t\<^sub>2 T\<^sub>2)  | 
| 63167 | 1119  | 
from \<open>TVarB X T\<^sub>1 # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2\<close>  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1120  | 
have "T\<^sub>1 closed_in \<Gamma>" by (auto dest: typing_ok)  | 
| 80142 | 1121  | 
with T_TAbs show ?case by (auto simp: ty.supp closed_in_def abs_supp)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1122  | 
next  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1123  | 
case (T_TApp X \<Gamma> t\<^sub>1 T2 T11 T12)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1124  | 
then have "T12 closed_in (TVarB X T11 # \<Gamma>)"  | 
| 80142 | 1125  | 
by (auto simp: closed_in_def ty.supp abs_supp)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1126  | 
moreover from T_TApp have "T2 closed_in \<Gamma>"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1127  | 
by (simp add: subtype_implies_closed)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1128  | 
ultimately show ?case by (rule subst_closed_in')  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1129  | 
qed  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1130  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1131  | 
|
| 63167 | 1132  | 
subsection \<open>Evaluation\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1133  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1134  | 
inductive  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1135  | 
val :: "trm \<Rightarrow> bool"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1136  | 
where  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1137  | 
Abs[intro]: "val (\<lambda>x:T. t)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1138  | 
| TAbs[intro]: "val (\<lambda>X<:T. t)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1139  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1140  | 
equivariance val  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1141  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1142  | 
inductive_cases val_inv_auto[elim]:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1143  | 
"val (Var x)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1144  | 
"val (t1 \<cdot> t2)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1145  | 
"val (t1 \<cdot>\<^sub>\<tau> t2)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1146  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1147  | 
inductive  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1148  | 
  eval :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longmapsto> _" [60,60] 60)
 | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1149  | 
where  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1150  | 
E_Abs : "\<lbrakk> x \<sharp> v\<^sub>2; val v\<^sub>2 \<rbrakk> \<Longrightarrow> (\<lambda>x:T\<^sub>11. t\<^sub>12) \<cdot> v\<^sub>2 \<longmapsto> t\<^sub>12[x \<mapsto> v\<^sub>2]"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1151  | 
| E_App1 [intro]: "t \<longmapsto> t' \<Longrightarrow> t \<cdot> u \<longmapsto> t' \<cdot> u"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1152  | 
| E_App2 [intro]: "\<lbrakk> val v; t \<longmapsto> t' \<rbrakk> \<Longrightarrow> v \<cdot> t \<longmapsto> v \<cdot> t'"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1153  | 
| E_TAbs : "X \<sharp> (T\<^sub>11, T\<^sub>2) \<Longrightarrow> (\<lambda>X<:T\<^sub>11. t\<^sub>12) \<cdot>\<^sub>\<tau> T\<^sub>2 \<longmapsto> t\<^sub>12[X \<mapsto>\<^sub>\<tau> T\<^sub>2]"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1154  | 
| E_TApp [intro]: "t \<longmapsto> t' \<Longrightarrow> t \<cdot>\<^sub>\<tau> T \<longmapsto> t' \<cdot>\<^sub>\<tau> T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1155  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1156  | 
lemma better_E_Abs[intro]:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1157  | 
assumes H: "val v2"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1158  | 
shows "(\<lambda>x:T11. t12) \<cdot> v2 \<longmapsto> t12[x \<mapsto> v2]"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1159  | 
proof -  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1160  | 
obtain y::vrs where y: "y \<sharp> (x, t12, v2)" by (rule exists_fresh) (rule fin_supp)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1161  | 
then have "y \<sharp> v2" by simp  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1162  | 
then have "(\<lambda>y:T11. [(y, x)] \<bullet> t12) \<cdot> v2 \<longmapsto> ([(y, x)] \<bullet> t12)[y \<mapsto> v2]" using H  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1163  | 
by (rule E_Abs)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1164  | 
moreover from y have "(\<lambda>x:T11. t12) \<cdot> v2 = (\<lambda>y:T11. [(y, x)] \<bullet> t12) \<cdot> v2"  | 
| 80142 | 1165  | 
by (auto simp: trm.inject alpha' fresh_prod fresh_atm)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1166  | 
ultimately have "(\<lambda>x:T11. t12) \<cdot> v2 \<longmapsto> ([(y, x)] \<bullet> t12)[y \<mapsto> v2]"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1167  | 
by simp  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1168  | 
with y show ?thesis by (simp add: subst_trm_rename)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1169  | 
qed  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1170  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1171  | 
lemma better_E_TAbs[intro]: "(\<lambda>X<:T11. t12) \<cdot>\<^sub>\<tau> T2 \<longmapsto> t12[X \<mapsto>\<^sub>\<tau> T2]"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1172  | 
proof -  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1173  | 
obtain Y::tyvrs where Y: "Y \<sharp> (X, t12, T11, T2)" by (rule exists_fresh) (rule fin_supp)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1174  | 
then have "Y \<sharp> (T11, T2)" by simp  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1175  | 
then have "(\<lambda>Y<:T11. [(Y, X)] \<bullet> t12) \<cdot>\<^sub>\<tau> T2 \<longmapsto> ([(Y, X)] \<bullet> t12)[Y \<mapsto>\<^sub>\<tau> T2]"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1176  | 
by (rule E_TAbs)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1177  | 
moreover from Y have "(\<lambda>X<:T11. t12) \<cdot>\<^sub>\<tau> T2 = (\<lambda>Y<:T11. [(Y, X)] \<bullet> t12) \<cdot>\<^sub>\<tau> T2"  | 
| 80142 | 1178  | 
by (auto simp: trm.inject alpha' fresh_prod fresh_atm)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1179  | 
ultimately have "(\<lambda>X<:T11. t12) \<cdot>\<^sub>\<tau> T2 \<longmapsto> ([(Y, X)] \<bullet> t12)[Y \<mapsto>\<^sub>\<tau> T2]"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1180  | 
by simp  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1181  | 
with Y show ?thesis by (simp add: subst_trm_ty_rename)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1182  | 
qed  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1183  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1184  | 
equivariance eval  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1185  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1186  | 
nominal_inductive eval  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1187  | 
by (simp_all add: abs_fresh ty_vrs_fresh subst_trm_fresh_tyvar  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1188  | 
subst_trm_fresh_var subst_trm_ty_fresh')  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1189  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1190  | 
inductive_cases eval_inv_auto[elim]:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1191  | 
"Var x \<longmapsto> t'"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1192  | 
"(\<lambda>x:T. t) \<longmapsto> t'"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1193  | 
"(\<lambda>X<:T. t) \<longmapsto> t'"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1194  | 
|
| 32011 | 1195  | 
lemma ty_dom_cons:  | 
1196  | 
shows "ty_dom (\<Gamma>@[VarB X Q]@\<Delta>) = ty_dom (\<Gamma>@\<Delta>)"  | 
|
| 49171 | 1197  | 
by (induct \<Gamma>) (auto)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1198  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1199  | 
lemma closed_in_cons:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1200  | 
assumes "S closed_in (\<Gamma> @ VarB X Q # \<Delta>)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1201  | 
shows "S closed_in (\<Gamma>@\<Delta>)"  | 
| 32011 | 1202  | 
using assms ty_dom_cons closed_in_def by auto  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1203  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1204  | 
lemma closed_in_weaken: "T closed_in (\<Delta> @ \<Gamma>) \<Longrightarrow> T closed_in (\<Delta> @ B # \<Gamma>)"  | 
| 80142 | 1205  | 
by (auto simp: closed_in_def doms_append)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1206  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1207  | 
lemma closed_in_weaken': "T closed_in \<Gamma> \<Longrightarrow> T closed_in (\<Delta> @ \<Gamma>)"  | 
| 80142 | 1208  | 
by (auto simp: closed_in_def doms_append)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1209  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1210  | 
lemma valid_subst:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1211  | 
assumes ok: "\<turnstile> (\<Delta> @ TVarB X Q # \<Gamma>) ok"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1212  | 
and closed: "P closed_in \<Gamma>"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1213  | 
shows "\<turnstile> (\<Delta>[X \<mapsto> P]\<^sub>e @ \<Gamma>) ok" using ok closed  | 
| 80142 | 1214  | 
proof (induct \<Delta>)  | 
1215  | 
case Nil  | 
|
1216  | 
then show ?case  | 
|
1217  | 
by auto  | 
|
1218  | 
next  | 
|
1219  | 
case (Cons a \<Delta>)  | 
|
1220  | 
then have *: "\<turnstile> (a # \<Delta> @ TVarB X Q # \<Gamma>) ok"  | 
|
1221  | 
by fastforce  | 
|
1222  | 
then show ?case  | 
|
1223  | 
apply (rule validE)  | 
|
1224  | 
using Cons  | 
|
1225  | 
apply (simp add: at_tyvrs_inst closed doms_append(1) finite_doms(1) fresh_fin_insert fs_tyvrs_inst pt_tyvrs_inst subst_closed_in ty_dom_subst)  | 
|
1226  | 
by (simp add: doms_append(2) subst_closed_in Cons.hyps closed trm_dom_subst)  | 
|
1227  | 
qed  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1228  | 
|
| 32011 | 1229  | 
lemma ty_dom_vrs:  | 
1230  | 
shows "ty_dom (G @ [VarB x Q] @ D) = ty_dom (G @ D)"  | 
|
| 80142 | 1231  | 
by (induct G) (auto)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1232  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1233  | 
lemma valid_cons':  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1234  | 
assumes "\<turnstile> (\<Gamma> @ VarB x Q # \<Delta>) ok"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1235  | 
shows "\<turnstile> (\<Gamma> @ \<Delta>) ok"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1236  | 
using assms  | 
| 34915 | 1237  | 
proof (induct "\<Gamma> @ VarB x Q # \<Delta>" arbitrary: \<Gamma> \<Delta>)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1238  | 
case valid_nil  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1239  | 
have "[] = \<Gamma> @ VarB x Q # \<Delta>" by fact  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1240  | 
then have "False" by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1241  | 
then show ?case by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1242  | 
next  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1243  | 
case (valid_consT G X T)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1244  | 
then show ?case  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1245  | 
proof (cases \<Gamma>)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1246  | 
case Nil  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1247  | 
with valid_consT show ?thesis by simp  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1248  | 
next  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1249  | 
case (Cons b bs)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1250  | 
with valid_consT  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1251  | 
have "\<turnstile> (bs @ \<Delta>) ok" by simp  | 
| 32011 | 1252  | 
moreover from Cons and valid_consT have "X \<sharp> ty_dom (bs @ \<Delta>)"  | 
1253  | 
by (simp add: doms_append)  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1254  | 
moreover from Cons and valid_consT have "T closed_in (bs @ \<Delta>)"  | 
| 32011 | 1255  | 
by (simp add: closed_in_def doms_append)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1256  | 
ultimately have "\<turnstile> (TVarB X T # bs @ \<Delta>) ok"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1257  | 
by (rule valid_rel.valid_consT)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1258  | 
with Cons and valid_consT show ?thesis by simp  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1259  | 
qed  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1260  | 
next  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1261  | 
case (valid_cons G x T)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1262  | 
then show ?case  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1263  | 
proof (cases \<Gamma>)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1264  | 
case Nil  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1265  | 
with valid_cons show ?thesis by simp  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1266  | 
next  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1267  | 
case (Cons b bs)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1268  | 
with valid_cons  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1269  | 
have "\<turnstile> (bs @ \<Delta>) ok" by simp  | 
| 32011 | 1270  | 
moreover from Cons and valid_cons have "x \<sharp> trm_dom (bs @ \<Delta>)"  | 
1271  | 
by (simp add: doms_append finite_doms  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
1272  | 
fresh_fin_insert [OF pt_vrs_inst at_vrs_inst fs_vrs_inst])  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1273  | 
moreover from Cons and valid_cons have "T closed_in (bs @ \<Delta>)"  | 
| 32011 | 1274  | 
by (simp add: closed_in_def doms_append)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1275  | 
ultimately have "\<turnstile> (VarB x T # bs @ \<Delta>) ok"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1276  | 
by (rule valid_rel.valid_cons)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1277  | 
with Cons and valid_cons show ?thesis by simp  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1278  | 
qed  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1279  | 
qed  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1280  | 
|
| 63167 | 1281  | 
text \<open>A.5(6)\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1282  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1283  | 
lemma type_weaken:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1284  | 
assumes "(\<Delta>@\<Gamma>) \<turnstile> t : T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1285  | 
and "\<turnstile> (\<Delta> @ B # \<Gamma>) ok"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1286  | 
shows "(\<Delta> @ B # \<Gamma>) \<turnstile> t : T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1287  | 
using assms  | 
| 34915 | 1288  | 
proof(nominal_induct "\<Delta> @ \<Gamma>" t T avoiding: \<Delta> \<Gamma> B rule: typing.strong_induct)  | 
1289  | 
case (T_Var x T)  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1290  | 
then show ?case by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1291  | 
next  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1292  | 
case (T_App X t\<^sub>1 T\<^sub>2 T\<^sub>11 T\<^sub>12)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1293  | 
then show ?case by force  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1294  | 
next  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1295  | 
case (T_Abs y T\<^sub>1 t\<^sub>2 T\<^sub>2 \<Delta> \<Gamma>)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1296  | 
then have "VarB y T\<^sub>1 # \<Delta> @ \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2" by simp  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1297  | 
then have closed: "T\<^sub>1 closed_in (\<Delta> @ \<Gamma>)"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1298  | 
by (auto dest: typing_ok)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1299  | 
have "\<turnstile> (VarB y T\<^sub>1 # \<Delta> @ B # \<Gamma>) ok"  | 
| 80142 | 1300  | 
by (simp add: T_Abs closed closed_in_weaken fresh_list_append fresh_list_cons fresh_trm_dom)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1301  | 
then have "\<turnstile> ((VarB y T\<^sub>1 # \<Delta>) @ B # \<Gamma>) ok" by simp  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1302  | 
with _ have "(VarB y T\<^sub>1 # \<Delta>) @ B # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2"  | 
| 34915 | 1303  | 
by (rule T_Abs) simp  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1304  | 
then have "VarB y T\<^sub>1 # \<Delta> @ B # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2" by simp  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1305  | 
then show ?case by (rule typing.T_Abs)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1306  | 
next  | 
| 34915 | 1307  | 
case (T_Sub t S T \<Delta> \<Gamma>)  | 
| 63167 | 1308  | 
from refl and \<open>\<turnstile> (\<Delta> @ B # \<Gamma>) ok\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1309  | 
have "\<Delta> @ B # \<Gamma> \<turnstile> t : S" by (rule T_Sub)  | 
| 63167 | 1310  | 
moreover from \<open>(\<Delta> @ \<Gamma>)\<turnstile>S<:T\<close> and \<open>\<turnstile> (\<Delta> @ B # \<Gamma>) ok\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1311  | 
have "(\<Delta> @ B # \<Gamma>)\<turnstile>S<:T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1312  | 
by (rule weakening) (simp add: extends_def T_Sub)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1313  | 
ultimately show ?case by (rule typing.T_Sub)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1314  | 
next  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1315  | 
case (T_TAbs X T\<^sub>1 t\<^sub>2 T\<^sub>2 \<Delta> \<Gamma>)  | 
| 63167 | 1316  | 
from \<open>TVarB X T\<^sub>1 # \<Delta> @ \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2\<close>  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1317  | 
have closed: "T\<^sub>1 closed_in (\<Delta> @ \<Gamma>)"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1318  | 
by (auto dest: typing_ok)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1319  | 
have "\<turnstile> (TVarB X T\<^sub>1 # \<Delta> @ B # \<Gamma>) ok"  | 
| 80142 | 1320  | 
by (simp add: T_TAbs at_tyvrs_inst closed closed_in_weaken doms_append finite_doms finite_vrs  | 
1321  | 
fresh_dom fresh_fin_union fs_tyvrs_inst pt_tyvrs_inst tyvrs_fresh)  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1322  | 
then have "\<turnstile> ((TVarB X T\<^sub>1 # \<Delta>) @ B # \<Gamma>) ok" by simp  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1323  | 
with _ have "(TVarB X T\<^sub>1 # \<Delta>) @ B # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2"  | 
| 34915 | 1324  | 
by (rule T_TAbs) simp  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1325  | 
then have "TVarB X T\<^sub>1 # \<Delta> @ B # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2" by simp  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1326  | 
then show ?case by (rule typing.T_TAbs)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1327  | 
next  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1328  | 
case (T_TApp X t\<^sub>1 T2 T11 T12 \<Delta> \<Gamma>)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1329  | 
have "\<Delta> @ B # \<Gamma> \<turnstile> t\<^sub>1 : (\<forall>X<:T11. T12)"  | 
| 34915 | 1330  | 
by (rule T_TApp refl)+  | 
| 63167 | 1331  | 
moreover from \<open>(\<Delta> @ \<Gamma>)\<turnstile>T2<:T11\<close> and \<open>\<turnstile> (\<Delta> @ B # \<Gamma>) ok\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1332  | 
have "(\<Delta> @ B # \<Gamma>)\<turnstile>T2<:T11"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1333  | 
by (rule weakening) (simp add: extends_def T_TApp)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1334  | 
ultimately show ?case by (rule better_T_TApp)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1335  | 
qed  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1336  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1337  | 
lemma type_weaken':  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1338  | 
"\<Gamma> \<turnstile> t : T \<Longrightarrow> \<turnstile> (\<Delta>@\<Gamma>) ok \<Longrightarrow> (\<Delta>@\<Gamma>) \<turnstile> t : T"  | 
| 80142 | 1339  | 
proof (induct \<Delta>)  | 
1340  | 
case Nil  | 
|
1341  | 
then show ?case by auto  | 
|
1342  | 
next  | 
|
1343  | 
case (Cons a \<Delta>)  | 
|
1344  | 
then show ?case  | 
|
1345  | 
by (metis append_Cons append_Nil type_weaken validE(3))  | 
|
1346  | 
qed  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1347  | 
|
| 63167 | 1348  | 
text \<open>A.6\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1349  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1350  | 
lemma strengthening:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1351  | 
assumes "(\<Gamma> @ VarB x Q # \<Delta>) \<turnstile> S <: T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1352  | 
shows "(\<Gamma>@\<Delta>) \<turnstile> S <: T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1353  | 
using assms  | 
| 34915 | 1354  | 
proof (induct "\<Gamma> @ VarB x Q # \<Delta>" S T arbitrary: \<Gamma>)  | 
1355  | 
case (SA_Top S)  | 
|
1356  | 
then have "\<turnstile> (\<Gamma> @ \<Delta>) ok" by (auto dest: valid_cons')  | 
|
1357  | 
moreover have "S closed_in (\<Gamma> @ \<Delta>)" using SA_Top by (auto dest: closed_in_cons)  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1358  | 
ultimately show ?case using subtype_of.SA_Top by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1359  | 
next  | 
| 34915 | 1360  | 
case (SA_refl_TVar X)  | 
| 63167 | 1361  | 
from \<open>\<turnstile> (\<Gamma> @ VarB x Q # \<Delta>) ok\<close>  | 
| 34915 | 1362  | 
have h1:"\<turnstile> (\<Gamma> @ \<Delta>) ok" by (auto dest: valid_cons')  | 
1363  | 
have "X \<in> ty_dom (\<Gamma> @ VarB x Q # \<Delta>)" using SA_refl_TVar by auto  | 
|
1364  | 
then have h2:"X \<in> ty_dom (\<Gamma> @ \<Delta>)" using ty_dom_vrs by auto  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1365  | 
show ?case using h1 h2 by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1366  | 
next  | 
| 34915 | 1367  | 
case (SA_all T1 S1 X S2 T2)  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
41798 
diff
changeset
 | 
1368  | 
have h1:"((TVarB X T1 # \<Gamma>) @ \<Delta>)\<turnstile>S2<:T2" by (fastforce intro: SA_all)  | 
| 34915 | 1369  | 
have h2:"(\<Gamma> @ \<Delta>)\<turnstile>T1<:S1" using SA_all by auto  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1370  | 
then show ?case using h1 h2 by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1371  | 
qed (auto)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1372  | 
|
| 63167 | 1373  | 
lemma narrow_type: \<comment> \<open>A.7\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1374  | 
assumes H: "\<Delta> @ (TVarB X Q) # \<Gamma> \<turnstile> t : T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1375  | 
shows "\<Gamma> \<turnstile> P <: Q \<Longrightarrow> \<Delta> @ (TVarB X P) # \<Gamma> \<turnstile> t : T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1376  | 
using H  | 
| 34915 | 1377  | 
proof (nominal_induct "\<Delta> @ (TVarB X Q) # \<Gamma>" t T avoiding: P arbitrary: \<Delta> rule: typing.strong_induct)  | 
1378  | 
case (T_Var x T P D)  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1379  | 
then have "VarB x T \<in> set (D @ TVarB X P # \<Gamma>)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1380  | 
and "\<turnstile> (D @ TVarB X P # \<Gamma>) ok"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1381  | 
by (auto intro: replace_type dest!: subtype_implies_closed)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1382  | 
then show ?case by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1383  | 
next  | 
| 34915 | 1384  | 
case (T_App t1 T1 T2 t2 P D)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1385  | 
then show ?case by force  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1386  | 
next  | 
| 34915 | 1387  | 
case (T_Abs x T1 t2 T2 P D)  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
41798 
diff
changeset
 | 
1388  | 
then show ?case by (fastforce dest: typing_ok)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1389  | 
next  | 
| 34915 | 1390  | 
case (T_Sub t S T P D)  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
41798 
diff
changeset
 | 
1391  | 
then show ?case using subtype_narrow by fastforce  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1392  | 
next  | 
| 34915 | 1393  | 
case (T_TAbs X' T1 t2 T2 P D)  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
41798 
diff
changeset
 | 
1394  | 
then show ?case by (fastforce dest: typing_ok)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1395  | 
next  | 
| 34915 | 1396  | 
case (T_TApp X' t1 T2 T11 T12 P D)  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
41798 
diff
changeset
 | 
1397  | 
then have "D @ TVarB X P # \<Gamma> \<turnstile> t1 : Forall X' T12 T11" by fastforce  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1398  | 
moreover have "(D @ [TVarB X Q] @ \<Gamma>) \<turnstile> T2<:T11" using T_TApp by auto  | 
| 63167 | 1399  | 
then have "(D @ [TVarB X P] @ \<Gamma>) \<turnstile> T2<:T11" using \<open>\<Gamma>\<turnstile>P<:Q\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1400  | 
by (rule subtype_narrow)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1401  | 
moreover from T_TApp have "X' \<sharp> (D @ TVarB X P # \<Gamma>, t1, T2)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1402  | 
by (simp add: fresh_list_append fresh_list_cons fresh_prod)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1403  | 
ultimately show ?case by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1404  | 
qed  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1405  | 
|
| 63167 | 1406  | 
subsection \<open>Substitution lemmas\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1407  | 
|
| 63167 | 1408  | 
subsubsection \<open>Substition Preserves Typing\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1409  | 
|
| 63167 | 1410  | 
theorem subst_type: \<comment> \<open>A.8\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1411  | 
assumes H: "(\<Delta> @ (VarB x U) # \<Gamma>) \<turnstile> t : T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1412  | 
shows "\<Gamma> \<turnstile> u : U \<Longrightarrow> \<Delta> @ \<Gamma> \<turnstile> t[x \<mapsto> u] : T" using H  | 
| 34915 | 1413  | 
proof (nominal_induct "\<Delta> @ (VarB x U) # \<Gamma>" t T avoiding: x u arbitrary: \<Delta> rule: typing.strong_induct)  | 
1414  | 
case (T_Var y T x u D)  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1415  | 
show ?case  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1416  | 
proof (cases "x = y")  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1417  | 
assume eq:"x=y"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1418  | 
then have "T=U" using T_Var uniqueness_of_ctxt' by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1419  | 
then show ?case using eq T_Var  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1420  | 
by (auto intro: type_weaken' dest: valid_cons')  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1421  | 
next  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1422  | 
assume "x\<noteq>y"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1423  | 
then show ?case using T_Var  | 
| 80142 | 1424  | 
by (auto simp:binding.inject dest: valid_cons')  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1425  | 
qed  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1426  | 
next  | 
| 34915 | 1427  | 
case (T_App t1 T1 T2 t2 x u D)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1428  | 
then show ?case by force  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1429  | 
next  | 
| 34915 | 1430  | 
case (T_Abs y T1 t2 T2 x u D)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1431  | 
then show ?case by force  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1432  | 
next  | 
| 34915 | 1433  | 
case (T_Sub t S T x u D)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1434  | 
then have "D @ \<Gamma> \<turnstile> t[x \<mapsto> u] : S" by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1435  | 
moreover have "(D @ \<Gamma>) \<turnstile> S<:T" using T_Sub by (auto dest: strengthening)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1436  | 
ultimately show ?case by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1437  | 
next  | 
| 34915 | 1438  | 
case (T_TAbs X T1 t2 T2 x u D)  | 
| 63167 | 1439  | 
from \<open>TVarB X T1 # D @ VarB x U # \<Gamma> \<turnstile> t2 : T2\<close> have "X \<sharp> T1"  | 
| 80142 | 1440  | 
by (auto simp: valid_ty_dom_fresh dest: typing_ok intro!: closed_in_fresh)  | 
| 63167 | 1441  | 
with \<open>X \<sharp> u\<close> and T_TAbs show ?case by fastforce  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1442  | 
next  | 
| 34915 | 1443  | 
case (T_TApp X t1 T2 T11 T12 x u D)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1444  | 
then have "(D@\<Gamma>) \<turnstile>T2<:T11" using T_TApp by (auto dest: strengthening)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1445  | 
then show "((D @ \<Gamma>) \<turnstile> ((t1 \<cdot>\<^sub>\<tau> T2)[x \<mapsto> u]) : (T12[X \<mapsto> T2]\<^sub>\<tau>))" using T_TApp  | 
| 80142 | 1446  | 
by (force simp: fresh_prod fresh_list_append fresh_list_cons subst_trm_fresh_tyvar)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1447  | 
qed  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1448  | 
|
| 63167 | 1449  | 
subsubsection \<open>Type Substitution Preserves Subtyping\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1450  | 
|
| 63167 | 1451  | 
lemma substT_subtype: \<comment> \<open>A.10\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1452  | 
assumes H: "(\<Delta> @ ((TVarB X Q) # \<Gamma>)) \<turnstile> S <: T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1453  | 
shows "\<Gamma> \<turnstile> P <: Q \<Longrightarrow> (\<Delta>[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> S[X \<mapsto> P]\<^sub>\<tau> <: T[X \<mapsto> P]\<^sub>\<tau>"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1454  | 
using H  | 
| 34915 | 1455  | 
proof (nominal_induct "\<Delta> @ TVarB X Q # \<Gamma>" S T avoiding: X P arbitrary: \<Delta> rule: subtype_of.strong_induct)  | 
1456  | 
case (SA_Top S X P D)  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1457  | 
then have "\<turnstile> (D @ TVarB X Q # \<Gamma>) ok" by simp  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1458  | 
moreover have closed: "P closed_in \<Gamma>" using SA_Top subtype_implies_closed by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1459  | 
ultimately have "\<turnstile> (D[X \<mapsto> P]\<^sub>e @ \<Gamma>) ok" by (rule valid_subst)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1460  | 
moreover from SA_Top have "S closed_in (D @ TVarB X Q # \<Gamma>)" by simp  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1461  | 
then have "S[X \<mapsto> P]\<^sub>\<tau> closed_in (D[X \<mapsto> P]\<^sub>e @ \<Gamma>)" using closed by (rule subst_closed_in)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1462  | 
ultimately show ?case by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1463  | 
next  | 
| 34915 | 1464  | 
case (SA_trans_TVar Y S T X P D)  | 
1465  | 
have h:"(D @ TVarB X Q # \<Gamma>)\<turnstile>S<:T" by fact  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1466  | 
then have ST: "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> S[X \<mapsto> P]\<^sub>\<tau> <: T[X \<mapsto> P]\<^sub>\<tau>" using SA_trans_TVar by auto  | 
| 34915 | 1467  | 
from h have G_ok: "\<turnstile> (D @ TVarB X Q # \<Gamma>) ok" by (rule subtype_implies_ok)  | 
| 50252 | 1468  | 
from G_ok and SA_trans_TVar have X_\<Gamma>_ok: "\<turnstile> (TVarB X Q # \<Gamma>) ok"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1469  | 
by (auto intro: validE_append)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1470  | 
show "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> Tvar Y[X \<mapsto> P]\<^sub>\<tau><:T[X \<mapsto> P]\<^sub>\<tau>"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1471  | 
proof (cases "X = Y")  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1472  | 
assume eq: "X = Y"  | 
| 34915 | 1473  | 
from eq and SA_trans_TVar have "TVarB Y Q \<in> set (D @ TVarB X Q # \<Gamma>)" by simp  | 
| 63167 | 1474  | 
with G_ok have QS: "Q = S" using \<open>TVarB Y S \<in> set (D @ TVarB X Q # \<Gamma>)\<close>  | 
| 34915 | 1475  | 
by (rule uniqueness_of_ctxt)  | 
| 50252 | 1476  | 
from X_\<Gamma>_ok have "X \<sharp> ty_dom \<Gamma>" and "Q closed_in \<Gamma>" by auto  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1477  | 
then have XQ: "X \<sharp> Q" by (rule closed_in_fresh)  | 
| 63167 | 1478  | 
note \<open>\<Gamma>\<turnstile>P<:Q\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1479  | 
moreover from ST have "\<turnstile> (D[X \<mapsto> P]\<^sub>e @ \<Gamma>) ok" by (rule subtype_implies_ok)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1480  | 
moreover have "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) extends \<Gamma>" by (simp add: extends_def)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1481  | 
ultimately have "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> P<:Q" by (rule weakening)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1482  | 
with QS have "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> P<:S" by simp  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1483  | 
moreover from XQ and ST and QS have "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> S<:T[X \<mapsto> P]\<^sub>\<tau>"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1484  | 
by (simp add: type_subst_identity)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1485  | 
ultimately have "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> P<:T[X \<mapsto> P]\<^sub>\<tau>"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1486  | 
by (rule subtype_transitivity)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1487  | 
with eq show ?case by simp  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1488  | 
next  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1489  | 
assume neq: "X \<noteq> Y"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1490  | 
with SA_trans_TVar have "TVarB Y S \<in> set D \<or> TVarB Y S \<in> set \<Gamma>"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1491  | 
by (simp add: binding.inject)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1492  | 
then show ?case  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1493  | 
proof  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1494  | 
assume "TVarB Y S \<in> set D"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1495  | 
then have "TVarB Y (S[X \<mapsto> P]\<^sub>\<tau>) \<in> set (D[X \<mapsto> P]\<^sub>e)"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
1496  | 
by (rule ctxt_subst_mem_TVarB)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1497  | 
then have "TVarB Y (S[X \<mapsto> P]\<^sub>\<tau>) \<in> set (D[X \<mapsto> P]\<^sub>e @ \<Gamma>)" by simp  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1498  | 
with neq and ST show ?thesis by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1499  | 
next  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1500  | 
assume Y: "TVarB Y S \<in> set \<Gamma>"  | 
| 50252 | 1501  | 
from X_\<Gamma>_ok have "X \<sharp> ty_dom \<Gamma>" and "\<turnstile> \<Gamma> ok" by auto  | 
| 32011 | 1502  | 
then have "X \<sharp> \<Gamma>" by (simp add: valid_ty_dom_fresh)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1503  | 
with Y have "X \<sharp> S"  | 
| 80142 | 1504  | 
by (induct \<Gamma>) (auto simp: fresh_list_nil fresh_list_cons)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1505  | 
with ST have "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>)\<turnstile>S<:T[X \<mapsto> P]\<^sub>\<tau>"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32011 
diff
changeset
 | 
1506  | 
by (simp add: type_subst_identity)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1507  | 
moreover from Y have "TVarB Y S \<in> set (D[X \<mapsto> P]\<^sub>e @ \<Gamma>)" by simp  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1508  | 
ultimately show ?thesis using neq by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1509  | 
qed  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1510  | 
qed  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1511  | 
next  | 
| 34915 | 1512  | 
case (SA_refl_TVar Y X P D)  | 
| 63167 | 1513  | 
note \<open>\<turnstile> (D @ TVarB X Q # \<Gamma>) ok\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1514  | 
moreover from SA_refl_TVar have closed: "P closed_in \<Gamma>"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1515  | 
by (auto dest: subtype_implies_closed)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1516  | 
ultimately have ok: "\<turnstile> (D[X \<mapsto> P]\<^sub>e @ \<Gamma>) ok" using valid_subst by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1517  | 
from closed have closed': "P closed_in (D[X \<mapsto> P]\<^sub>e @ \<Gamma>)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1518  | 
by (simp add: closed_in_weaken')  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1519  | 
show ?case  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1520  | 
proof (cases "X = Y")  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1521  | 
assume "X = Y"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1522  | 
with closed' and ok show ?thesis  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1523  | 
by (auto intro: subtype_reflexivity)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1524  | 
next  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1525  | 
assume neq: "X \<noteq> Y"  | 
| 32011 | 1526  | 
with SA_refl_TVar have "Y \<in> ty_dom (D[X \<mapsto> P]\<^sub>e @ \<Gamma>)"  | 
1527  | 
by (simp add: ty_dom_subst doms_append)  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1528  | 
with neq and ok show ?thesis by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1529  | 
qed  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1530  | 
next  | 
| 34915 | 1531  | 
case (SA_arrow T1 S1 S2 T2 X P D)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1532  | 
then have h1:"(D[X \<mapsto> P]\<^sub>e @ \<Gamma>)\<turnstile>T1[X \<mapsto> P]\<^sub>\<tau><:S1[X \<mapsto> P]\<^sub>\<tau>" using SA_arrow by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1533  | 
from SA_arrow have h2:"(D[X \<mapsto> P]\<^sub>e @ \<Gamma>)\<turnstile>S2[X \<mapsto> P]\<^sub>\<tau><:T2[X \<mapsto> P]\<^sub>\<tau>" using SA_arrow by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1534  | 
show ?case using subtype_of.SA_arrow h1 h2 by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1535  | 
next  | 
| 34915 | 1536  | 
case (SA_all T1 S1 Y S2 T2 X P D)  | 
| 32011 | 1537  | 
then have Y: "Y \<sharp> ty_dom (D @ TVarB X Q # \<Gamma>)"  | 
1538  | 
by (auto dest: subtype_implies_ok intro: fresh_dom)  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1539  | 
moreover from SA_all have "S1 closed_in (D @ TVarB X Q # \<Gamma>)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1540  | 
by (auto dest: subtype_implies_closed)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1541  | 
ultimately have S1: "Y \<sharp> S1" by (rule closed_in_fresh)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1542  | 
from SA_all have "T1 closed_in (D @ TVarB X Q # \<Gamma>)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1543  | 
by (auto dest: subtype_implies_closed)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1544  | 
with Y have T1: "Y \<sharp> T1" by (rule closed_in_fresh)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1545  | 
with SA_all and S1 show ?case by force  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1546  | 
qed  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1547  | 
|
| 63167 | 1548  | 
subsubsection \<open>Type Substitution Preserves Typing\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1549  | 
|
| 63167 | 1550  | 
theorem substT_type: \<comment> \<open>A.11\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1551  | 
assumes H: "(D @ TVarB X Q # G) \<turnstile> t : T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1552  | 
shows "G \<turnstile> P <: Q \<Longrightarrow>  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1553  | 
(D[X \<mapsto> P]\<^sub>e @ G) \<turnstile> t[X \<mapsto>\<^sub>\<tau> P] : T[X \<mapsto> P]\<^sub>\<tau>" using H  | 
| 34915 | 1554  | 
proof (nominal_induct "D @ TVarB X Q # G" t T avoiding: X P arbitrary: D rule: typing.strong_induct)  | 
1555  | 
case (T_Var x T X P D')  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1556  | 
have "G\<turnstile>P<:Q" by fact  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1557  | 
then have "P closed_in G" using subtype_implies_closed by auto  | 
| 63167 | 1558  | 
moreover note \<open>\<turnstile> (D' @ TVarB X Q # G) ok\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1559  | 
ultimately have "\<turnstile> (D'[X \<mapsto> P]\<^sub>e @ G) ok" using valid_subst by auto  | 
| 63167 | 1560  | 
moreover note \<open>VarB x T \<in> set (D' @ TVarB X Q # G)\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1561  | 
then have "VarB x T \<in> set D' \<or> VarB x T \<in> set G" by simp  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1562  | 
then have "(VarB x (T[X \<mapsto> P]\<^sub>\<tau>)) \<in> set (D'[X \<mapsto> P]\<^sub>e @ G)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1563  | 
proof  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1564  | 
assume "VarB x T \<in> set D'"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1565  | 
then have "VarB x (T[X \<mapsto> P]\<^sub>\<tau>) \<in> set (D'[X \<mapsto> P]\<^sub>e)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1566  | 
by (rule ctxt_subst_mem_VarB)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1567  | 
then show ?thesis by simp  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1568  | 
next  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1569  | 
assume x: "VarB x T \<in> set G"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1570  | 
from T_Var have ok: "\<turnstile> G ok" by (auto dest: subtype_implies_ok)  | 
| 32011 | 1571  | 
then have "X \<sharp> ty_dom G" using T_Var by (auto dest: validE_append)  | 
1572  | 
with ok have "X \<sharp> G" by (simp add: valid_ty_dom_fresh)  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1573  | 
moreover from x have "VarB x T \<in> set (D' @ G)" by simp  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1574  | 
then have "VarB x (T[X \<mapsto> P]\<^sub>\<tau>) \<in> set ((D' @ G)[X \<mapsto> P]\<^sub>e)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1575  | 
by (rule ctxt_subst_mem_VarB)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1576  | 
ultimately show ?thesis  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1577  | 
by (simp add: ctxt_subst_append ctxt_subst_identity)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1578  | 
qed  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1579  | 
ultimately show ?case by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1580  | 
next  | 
| 34915 | 1581  | 
case (T_App t1 T1 T2 t2 X P D')  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1582  | 
then have "D'[X \<mapsto> P]\<^sub>e @ G \<turnstile> t1[X \<mapsto>\<^sub>\<tau> P] : (T1 \<rightarrow> T2)[X \<mapsto> P]\<^sub>\<tau>" by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1583  | 
moreover from T_App have "D'[X \<mapsto> P]\<^sub>e @ G \<turnstile> t2[X \<mapsto>\<^sub>\<tau> P] : T1[X \<mapsto> P]\<^sub>\<tau>" by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1584  | 
ultimately show ?case by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1585  | 
next  | 
| 34915 | 1586  | 
case (T_Abs x T1 t2 T2 X P D')  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1587  | 
then show ?case by force  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1588  | 
next  | 
| 34915 | 1589  | 
case (T_Sub t S T X P D')  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1590  | 
then show ?case using substT_subtype by force  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1591  | 
next  | 
| 34915 | 1592  | 
case (T_TAbs X' T1 t2 T2 X P D')  | 
| 32011 | 1593  | 
then have "X' \<sharp> ty_dom (D' @ TVarB X Q # G)"  | 
| 34915 | 1594  | 
and "T1 closed_in (D' @ TVarB X Q # G)"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1595  | 
by (auto dest: typing_ok)  | 
| 34915 | 1596  | 
then have "X' \<sharp> T1" by (rule closed_in_fresh)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1597  | 
with T_TAbs show ?case by force  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1598  | 
next  | 
| 34915 | 1599  | 
case (T_TApp X' t1 T2 T11 T12 X P D')  | 
| 32011 | 1600  | 
then have "X' \<sharp> ty_dom (D' @ TVarB X Q # G)"  | 
1601  | 
by (simp add: fresh_dom)  | 
|
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1602  | 
moreover from T_TApp have "T11 closed_in (D' @ TVarB X Q # G)"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1603  | 
by (auto dest: subtype_implies_closed)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1604  | 
ultimately have X': "X' \<sharp> T11" by (rule closed_in_fresh)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1605  | 
from T_TApp have "D'[X \<mapsto> P]\<^sub>e @ G \<turnstile> t1[X \<mapsto>\<^sub>\<tau> P] : (\<forall>X'<:T11. T12)[X \<mapsto> P]\<^sub>\<tau>"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1606  | 
by simp  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1607  | 
with X' and T_TApp show ?case  | 
| 80142 | 1608  | 
by (auto simp: fresh_atm type_substitution_lemma  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1609  | 
fresh_list_append fresh_list_cons  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1610  | 
ctxt_subst_fresh' type_subst_fresh subst_trm_ty_fresh  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1611  | 
intro: substT_subtype)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1612  | 
qed  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1613  | 
|
| 63167 | 1614  | 
lemma Abs_type: \<comment> \<open>A.13(1)\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1615  | 
assumes H: "\<Gamma> \<turnstile> (\<lambda>x:S. s) : T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1616  | 
and H': "\<Gamma> \<turnstile> T <: U \<rightarrow> U'"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1617  | 
and H'': "x \<sharp> \<Gamma>"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1618  | 
obtains S' where "\<Gamma> \<turnstile> U <: S"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1619  | 
and "(VarB x S) # \<Gamma> \<turnstile> s : S'"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1620  | 
and "\<Gamma> \<turnstile> S' <: U'"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1621  | 
using H H' H''  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1622  | 
proof (nominal_induct \<Gamma> t \<equiv> "\<lambda>x:S. s" T avoiding: x arbitrary: U U' S s rule: typing.strong_induct)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1623  | 
case (T_Abs y T\<^sub>1 \<Gamma> t\<^sub>2 T\<^sub>2)  | 
| 63167 | 1624  | 
from \<open>\<Gamma> \<turnstile> T\<^sub>1 \<rightarrow> T\<^sub>2 <: U \<rightarrow> U'\<close>  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1625  | 
obtain ty1: "\<Gamma> \<turnstile> U <: S" and ty2: "\<Gamma> \<turnstile> T\<^sub>2 <: U'" using T_Abs  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1626  | 
by cases (simp_all add: ty.inject trm.inject alpha fresh_atm)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1627  | 
from T_Abs have "VarB y S # \<Gamma> \<turnstile> [(y, x)] \<bullet> s : T\<^sub>2"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1628  | 
by (simp add: trm.inject alpha fresh_atm)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1629  | 
then have "[(y, x)] \<bullet> (VarB y S # \<Gamma>) \<turnstile> [(y, x)] \<bullet> [(y, x)] \<bullet> s : [(y, x)] \<bullet> T\<^sub>2"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1630  | 
by (rule typing.eqvt)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1631  | 
moreover from T_Abs have "y \<sharp> \<Gamma>"  | 
| 32011 | 1632  | 
by (auto dest!: typing_ok simp add: fresh_trm_dom)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1633  | 
ultimately have "VarB x S # \<Gamma> \<turnstile> s : T\<^sub>2" using T_Abs  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1634  | 
by (perm_simp add: ty_vrs_prm_simp)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1635  | 
with ty1 show ?case using ty2 by (rule T_Abs)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1636  | 
next  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1637  | 
case (T_Sub \<Gamma> t S T)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1638  | 
then show ?case using subtype_transitivity by blast  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1639  | 
qed simp_all  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1640  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1641  | 
lemma subtype_reflexivity_from_typing:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1642  | 
assumes "\<Gamma> \<turnstile> t : T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1643  | 
shows "\<Gamma> \<turnstile> T <: T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1644  | 
using assms subtype_reflexivity typing_ok typing_closed_in by simp  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1645  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1646  | 
lemma Abs_type':  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1647  | 
assumes H: "\<Gamma> \<turnstile> (\<lambda>x:S. s) : U \<rightarrow> U'"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1648  | 
and H': "x \<sharp> \<Gamma>"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1649  | 
obtains S'  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1650  | 
where "\<Gamma> \<turnstile> U <: S"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1651  | 
and "(VarB x S) # \<Gamma> \<turnstile> s : S'"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1652  | 
and "\<Gamma> \<turnstile> S' <: U'"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1653  | 
using H subtype_reflexivity_from_typing [OF H] H'  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1654  | 
by (rule Abs_type)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1655  | 
|
| 63167 | 1656  | 
lemma TAbs_type: \<comment> \<open>A.13(2)\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1657  | 
assumes H: "\<Gamma> \<turnstile> (\<lambda>X<:S. s) : T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1658  | 
and H': "\<Gamma> \<turnstile> T <: (\<forall>X<:U. U')"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1659  | 
and fresh: "X \<sharp> \<Gamma>" "X \<sharp> S" "X \<sharp> U"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1660  | 
obtains S'  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1661  | 
where "\<Gamma> \<turnstile> U <: S"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1662  | 
and "(TVarB X U # \<Gamma>) \<turnstile> s : S'"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1663  | 
and "(TVarB X U # \<Gamma>) \<turnstile> S' <: U'"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1664  | 
using H H' fresh  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1665  | 
proof (nominal_induct \<Gamma> t \<equiv> "\<lambda>X<:S. s" T avoiding: X U U' S arbitrary: s rule: typing.strong_induct)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1666  | 
case (T_TAbs Y T\<^sub>1 \<Gamma> t\<^sub>2 T\<^sub>2)  | 
| 63167 | 1667  | 
from \<open>TVarB Y T\<^sub>1 # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2\<close> have Y: "Y \<sharp> \<Gamma>"  | 
| 32011 | 1668  | 
by (auto dest!: typing_ok simp add: valid_ty_dom_fresh)  | 
| 63167 | 1669  | 
from \<open>Y \<sharp> U'\<close> and \<open>Y \<sharp> X\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1670  | 
have "(\<forall>X<:U. U') = (\<forall>Y<:U. [(Y, X)] \<bullet> U')"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1671  | 
by (simp add: ty.inject alpha' fresh_atm)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1672  | 
with T_TAbs have "\<Gamma> \<turnstile> (\<forall>Y<:S. T\<^sub>2) <: (\<forall>Y<:U. [(Y, X)] \<bullet> U')" by (simp add: trm.inject)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1673  | 
then obtain ty1: "\<Gamma> \<turnstile> U <: S" and ty2: "(TVarB Y U # \<Gamma>) \<turnstile> T\<^sub>2 <: ([(Y, X)] \<bullet> U')" using T_TAbs Y  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1674  | 
by (cases rule: subtype_of.strong_cases [where X=Y]) (simp_all add: ty.inject alpha abs_fresh)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1675  | 
note ty1  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1676  | 
moreover from T_TAbs have "TVarB Y S # \<Gamma> \<turnstile> ([(Y, X)] \<bullet> s) : T\<^sub>2"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1677  | 
by (simp add: trm.inject alpha fresh_atm)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1678  | 
then have "[(Y, X)] \<bullet> (TVarB Y S # \<Gamma>) \<turnstile> [(Y, X)] \<bullet> [(Y, X)] \<bullet> s : [(Y, X)] \<bullet> T\<^sub>2"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1679  | 
by (rule typing.eqvt)  | 
| 63167 | 1680  | 
with \<open>X \<sharp> \<Gamma>\<close> \<open>X \<sharp> S\<close> Y \<open>Y \<sharp> S\<close> have "TVarB X S # \<Gamma> \<turnstile> s : [(Y, X)] \<bullet> T\<^sub>2"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1681  | 
by perm_simp  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1682  | 
then have "TVarB X U # \<Gamma> \<turnstile> s : [(Y, X)] \<bullet> T\<^sub>2" using ty1  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1683  | 
by (rule narrow_type [of "[]", simplified])  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1684  | 
moreover from ty2 have "([(Y, X)] \<bullet> (TVarB Y U # \<Gamma>)) \<turnstile> ([(Y, X)] \<bullet> T\<^sub>2) <: ([(Y, X)] \<bullet> [(Y, X)] \<bullet> U')"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1685  | 
by (rule subtype_of.eqvt)  | 
| 63167 | 1686  | 
with \<open>X \<sharp> \<Gamma>\<close> \<open>X \<sharp> U\<close> Y \<open>Y \<sharp> U\<close> have "(TVarB X U # \<Gamma>) \<turnstile> ([(Y, X)] \<bullet> T\<^sub>2) <: U'"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1687  | 
by perm_simp  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1688  | 
ultimately show ?case by (rule T_TAbs)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1689  | 
next  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1690  | 
case (T_Sub \<Gamma> t S T)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1691  | 
then show ?case using subtype_transitivity by blast  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1692  | 
qed simp_all  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1693  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1694  | 
lemma TAbs_type':  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1695  | 
assumes H: "\<Gamma> \<turnstile> (\<lambda>X<:S. s) : (\<forall>X<:U. U')"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1696  | 
and fresh: "X \<sharp> \<Gamma>" "X \<sharp> S" "X \<sharp> U"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1697  | 
obtains S'  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1698  | 
where "\<Gamma> \<turnstile> U <: S"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1699  | 
and "(TVarB X U # \<Gamma>) \<turnstile> s : S'"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1700  | 
and "(TVarB X U # \<Gamma>) \<turnstile> S' <: U'"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1701  | 
using H subtype_reflexivity_from_typing [OF H] fresh  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1702  | 
by (rule TAbs_type)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1703  | 
|
| 63167 | 1704  | 
theorem preservation: \<comment> \<open>A.20\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1705  | 
assumes H: "\<Gamma> \<turnstile> t : T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1706  | 
shows "t \<longmapsto> t' \<Longrightarrow> \<Gamma> \<turnstile> t' : T" using H  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1707  | 
proof (nominal_induct avoiding: t' rule: typing.strong_induct)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1708  | 
case (T_App \<Gamma> t\<^sub>1 T\<^sub>11 T\<^sub>12 t\<^sub>2 t')  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1709  | 
obtain x::vrs where x_fresh: "x \<sharp> (\<Gamma>, t\<^sub>1 \<cdot> t\<^sub>2, t')"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1710  | 
by (rule exists_fresh) (rule fin_supp)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1711  | 
obtain X::tyvrs where "X \<sharp> (t\<^sub>1 \<cdot> t\<^sub>2, t')"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1712  | 
by (rule exists_fresh) (rule fin_supp)  | 
| 63167 | 1713  | 
with \<open>t\<^sub>1 \<cdot> t\<^sub>2 \<longmapsto> t'\<close> show ?case  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1714  | 
proof (cases rule: eval.strong_cases [where x=x and X=X])  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1715  | 
case (E_Abs v\<^sub>2 T\<^sub>11' t\<^sub>12)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1716  | 
with T_App and x_fresh have h: "\<Gamma> \<turnstile> (\<lambda>x:T\<^sub>11'. t\<^sub>12) : T\<^sub>11 \<rightarrow> T\<^sub>12"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1717  | 
by (simp add: trm.inject fresh_prod)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1718  | 
moreover from x_fresh have "x \<sharp> \<Gamma>" by simp  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1719  | 
ultimately obtain S'  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1720  | 
where T\<^sub>11: "\<Gamma> \<turnstile> T\<^sub>11 <: T\<^sub>11'"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1721  | 
and t\<^sub>12: "(VarB x T\<^sub>11') # \<Gamma> \<turnstile> t\<^sub>12 : S'"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1722  | 
and S': "\<Gamma> \<turnstile> S' <: T\<^sub>12"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1723  | 
by (rule Abs_type') blast  | 
| 63167 | 1724  | 
from \<open>\<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>11\<close>  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1725  | 
have "\<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>11'" using T\<^sub>11 by (rule T_Sub)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1726  | 
with t\<^sub>12 have "\<Gamma> \<turnstile> t\<^sub>12[x \<mapsto> t\<^sub>2] : S'"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1727  | 
by (rule subst_type [where \<Delta>="[]", simplified])  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1728  | 
hence "\<Gamma> \<turnstile> t\<^sub>12[x \<mapsto> t\<^sub>2] : T\<^sub>12" using S' by (rule T_Sub)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1729  | 
with E_Abs and x_fresh show ?thesis by (simp add: trm.inject fresh_prod)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1730  | 
next  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1731  | 
case (E_App1 t''' t'' u)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1732  | 
hence "t\<^sub>1 \<longmapsto> t''" by (simp add:trm.inject)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1733  | 
hence "\<Gamma> \<turnstile> t'' : T\<^sub>11 \<rightarrow> T\<^sub>12" by (rule T_App)  | 
| 63167 | 1734  | 
hence "\<Gamma> \<turnstile> t'' \<cdot> t\<^sub>2 : T\<^sub>12" using \<open>\<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>11\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1735  | 
by (rule typing.T_App)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1736  | 
with E_App1 show ?thesis by (simp add:trm.inject)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1737  | 
next  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1738  | 
case (E_App2 v t''' t'')  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1739  | 
hence "t\<^sub>2 \<longmapsto> t''" by (simp add:trm.inject)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1740  | 
hence "\<Gamma> \<turnstile> t'' : T\<^sub>11" by (rule T_App)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1741  | 
with T_App(1) have "\<Gamma> \<turnstile> t\<^sub>1 \<cdot> t'' : T\<^sub>12"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1742  | 
by (rule typing.T_App)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1743  | 
with E_App2 show ?thesis by (simp add:trm.inject)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1744  | 
qed (simp_all add: fresh_prod)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1745  | 
next  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1746  | 
case (T_TApp X \<Gamma> t\<^sub>1 T\<^sub>2 T\<^sub>11 T\<^sub>12 t')  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1747  | 
obtain x::vrs where "x \<sharp> (t\<^sub>1 \<cdot>\<^sub>\<tau> T\<^sub>2, t')"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1748  | 
by (rule exists_fresh) (rule fin_supp)  | 
| 63167 | 1749  | 
with \<open>t\<^sub>1 \<cdot>\<^sub>\<tau> T\<^sub>2 \<longmapsto> t'\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1750  | 
show ?case  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1751  | 
proof (cases rule: eval.strong_cases [where X=X and x=x])  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1752  | 
case (E_TAbs T\<^sub>11' T\<^sub>2' t\<^sub>12)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1753  | 
with T_TApp have "\<Gamma> \<turnstile> (\<lambda>X<:T\<^sub>11'. t\<^sub>12) : (\<forall>X<:T\<^sub>11. T\<^sub>12)" and "X \<sharp> \<Gamma>" and "X \<sharp> T\<^sub>11'"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1754  | 
by (simp_all add: trm.inject)  | 
| 63167 | 1755  | 
moreover from \<open>\<Gamma>\<turnstile>T\<^sub>2<:T\<^sub>11\<close> and \<open>X \<sharp> \<Gamma>\<close> have "X \<sharp> T\<^sub>11"  | 
| 32011 | 1756  | 
by (blast intro: closed_in_fresh fresh_dom dest: subtype_implies_closed)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1757  | 
ultimately obtain S'  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1758  | 
where "TVarB X T\<^sub>11 # \<Gamma> \<turnstile> t\<^sub>12 : S'"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1759  | 
and "(TVarB X T\<^sub>11 # \<Gamma>) \<turnstile> S' <: T\<^sub>12"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1760  | 
by (rule TAbs_type') blast  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1761  | 
hence "TVarB X T\<^sub>11 # \<Gamma> \<turnstile> t\<^sub>12 : T\<^sub>12" by (rule T_Sub)  | 
| 63167 | 1762  | 
hence "\<Gamma> \<turnstile> t\<^sub>12[X \<mapsto>\<^sub>\<tau> T\<^sub>2] : T\<^sub>12[X \<mapsto> T\<^sub>2]\<^sub>\<tau>" using \<open>\<Gamma> \<turnstile> T\<^sub>2 <: T\<^sub>11\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1763  | 
by (rule substT_type [where D="[]", simplified])  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1764  | 
with T_TApp and E_TAbs show ?thesis by (simp add: trm.inject)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1765  | 
next  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1766  | 
case (E_TApp t''' t'' T)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1767  | 
from E_TApp have "t\<^sub>1 \<longmapsto> t''" by (simp add: trm.inject)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1768  | 
then have "\<Gamma> \<turnstile> t'' : (\<forall>X<:T\<^sub>11. T\<^sub>12)" by (rule T_TApp)  | 
| 63167 | 1769  | 
then have "\<Gamma> \<turnstile> t'' \<cdot>\<^sub>\<tau> T\<^sub>2 : T\<^sub>12[X \<mapsto> T\<^sub>2]\<^sub>\<tau>" using \<open>\<Gamma> \<turnstile> T\<^sub>2 <: T\<^sub>11\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1770  | 
by (rule better_T_TApp)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1771  | 
with E_TApp show ?thesis by (simp add: trm.inject)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1772  | 
qed (simp_all add: fresh_prod)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1773  | 
next  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1774  | 
case (T_Sub \<Gamma> t S T t')  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1775  | 
have "t \<longmapsto> t'" by fact  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1776  | 
hence "\<Gamma> \<turnstile> t' : S" by (rule T_Sub)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1777  | 
moreover have "\<Gamma> \<turnstile> S <: T" by fact  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1778  | 
ultimately show ?case by (rule typing.T_Sub)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1779  | 
qed (auto)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1780  | 
|
| 63167 | 1781  | 
lemma Fun_canonical: \<comment> \<open>A.14(1)\<close>  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1782  | 
assumes ty: "[] \<turnstile> v : T\<^sub>1 \<rightarrow> T\<^sub>2"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1783  | 
shows "val v \<Longrightarrow> \<exists>x t S. v = (\<lambda>x:S. t)" using ty  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1784  | 
proof (induct "[]::env" v "T\<^sub>1 \<rightarrow> T\<^sub>2" arbitrary: T\<^sub>1 T\<^sub>2)  | 
| 34915 | 1785  | 
case (T_Sub t S)  | 
| 63167 | 1786  | 
from \<open>[] \<turnstile> S <: T\<^sub>1 \<rightarrow> T\<^sub>2\<close>  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1787  | 
obtain S\<^sub>1 S\<^sub>2 where S: "S = S\<^sub>1 \<rightarrow> S\<^sub>2"  | 
| 80142 | 1788  | 
by cases (auto simp: T_Sub)  | 
| 63167 | 1789  | 
then show ?case using \<open>val t\<close> by (rule T_Sub)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1790  | 
qed (auto)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1791  | 
|
| 63167 | 1792  | 
lemma TyAll_canonical: \<comment> \<open>A.14(3)\<close>  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1793  | 
fixes X::tyvrs  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1794  | 
assumes ty: "[] \<turnstile> v : (\<forall>X<:T\<^sub>1. T\<^sub>2)"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1795  | 
shows "val v \<Longrightarrow> \<exists>X t S. v = (\<lambda>X<:S. t)" using ty  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1796  | 
proof (induct "[]::env" v "\<forall>X<:T\<^sub>1. T\<^sub>2" arbitrary: X T\<^sub>1 T\<^sub>2)  | 
| 34915 | 1797  | 
case (T_Sub t S)  | 
| 63167 | 1798  | 
from \<open>[] \<turnstile> S <: (\<forall>X<:T\<^sub>1. T\<^sub>2)\<close>  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1799  | 
obtain X S\<^sub>1 S\<^sub>2 where S: "S = (\<forall>X<:S\<^sub>1. S\<^sub>2)"  | 
| 80142 | 1800  | 
by cases (auto simp: T_Sub)  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1801  | 
then show ?case using T_Sub by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1802  | 
qed (auto)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1803  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1804  | 
theorem progress:  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1805  | 
assumes "[] \<turnstile> t : T"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1806  | 
shows "val t \<or> (\<exists>t'. t \<longmapsto> t')"  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1807  | 
using assms  | 
| 34915 | 1808  | 
proof (induct "[]::env" t T)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1809  | 
case (T_App t\<^sub>1 T\<^sub>11 T\<^sub>12 t\<^sub>2)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1810  | 
hence "val t\<^sub>1 \<or> (\<exists>t'. t\<^sub>1 \<longmapsto> t')" by simp  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1811  | 
thus ?case  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1812  | 
proof  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1813  | 
assume t\<^sub>1_val: "val t\<^sub>1"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1814  | 
with T_App obtain x t3 S where t\<^sub>1: "t\<^sub>1 = (\<lambda>x:S. t3)"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1815  | 
by (auto dest!: Fun_canonical)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1816  | 
from T_App have "val t\<^sub>2 \<or> (\<exists>t'. t\<^sub>2 \<longmapsto> t')" by simp  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1817  | 
thus ?case  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1818  | 
proof  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1819  | 
assume "val t\<^sub>2"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1820  | 
with t\<^sub>1 have "t\<^sub>1 \<cdot> t\<^sub>2 \<longmapsto> t3[x \<mapsto> t\<^sub>2]" by auto  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1821  | 
thus ?case by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1822  | 
next  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1823  | 
assume "\<exists>t'. t\<^sub>2 \<longmapsto> t'"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1824  | 
then obtain t' where "t\<^sub>2 \<longmapsto> t'" by auto  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1825  | 
with t\<^sub>1_val have "t\<^sub>1 \<cdot> t\<^sub>2 \<longmapsto> t\<^sub>1 \<cdot> t'" by auto  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1826  | 
thus ?case by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1827  | 
qed  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1828  | 
next  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1829  | 
assume "\<exists>t'. t\<^sub>1 \<longmapsto> t'"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1830  | 
then obtain t' where "t\<^sub>1 \<longmapsto> t'" by auto  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1831  | 
hence "t\<^sub>1 \<cdot> t\<^sub>2 \<longmapsto> t' \<cdot> t\<^sub>2" by auto  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1832  | 
thus ?case by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1833  | 
qed  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1834  | 
next  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1835  | 
case (T_TApp X t\<^sub>1 T\<^sub>2 T\<^sub>11 T\<^sub>12)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1836  | 
hence "val t\<^sub>1 \<or> (\<exists>t'. t\<^sub>1 \<longmapsto> t')" by simp  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1837  | 
thus ?case  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1838  | 
proof  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1839  | 
assume "val t\<^sub>1"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1840  | 
with T_TApp obtain x t S where "t\<^sub>1 = (\<lambda>x<:S. t)"  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1841  | 
by (auto dest!: TyAll_canonical)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1842  | 
hence "t\<^sub>1 \<cdot>\<^sub>\<tau> T\<^sub>2 \<longmapsto> t[x \<mapsto>\<^sub>\<tau> T\<^sub>2]" by auto  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1843  | 
thus ?case by auto  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1844  | 
next  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50252 
diff
changeset
 | 
1845  | 
assume "\<exists>t'. t\<^sub>1 \<longmapsto> t'" thus ?case by auto  | 
| 
30091
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1846  | 
qed  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1847  | 
qed (auto)  | 
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1848  | 
|
| 
 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 
berghofe 
parents: 
29097 
diff
changeset
 | 
1849  | 
end  |