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