author | wenzelm |
Mon, 23 Sep 2024 22:33:37 +0200 | |
changeset 80935 | b5bdcdbf5ec1 |
parent 80914 | d97fdabd9e2b |
child 81125 | ec121999a9cb |
permissions | -rw-r--r-- |
63167 | 1 |
section \<open>Simply-typed lambda-calculus with let and tuple patterns\<close> |
33189 | 2 |
|
3 |
theory Pattern |
|
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
63167
diff
changeset
|
4 |
imports "HOL-Nominal.Nominal" |
33189 | 5 |
begin |
6 |
||
7 |
no_syntax |
|
80935 | 8 |
"_Map" :: "maplets => 'a \<rightharpoonup> 'b" (\<open>(\<open>indent=1 notation=\<open>mixfix maplets\<close>\<close>[_])\<close>) |
33189 | 9 |
|
10 |
atom_decl name |
|
11 |
||
12 |
nominal_datatype ty = |
|
13 |
Atom nat |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80142
diff
changeset
|
14 |
| Arrow ty ty (infixr \<open>\<rightarrow>\<close> 200) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80142
diff
changeset
|
15 |
| TupleT ty ty (infixr \<open>\<otimes>\<close> 210) |
33189 | 16 |
|
17 |
lemma fresh_type [simp]: "(a::name) \<sharp> (T::ty)" |
|
18 |
by (induct T rule: ty.induct) (simp_all add: fresh_nat) |
|
19 |
||
20 |
lemma supp_type [simp]: "supp (T::ty) = ({} :: name set)" |
|
21 |
by (induct T rule: ty.induct) (simp_all add: ty.supp supp_nat) |
|
22 |
||
23 |
lemma perm_type: "(pi::name prm) \<bullet> (T::ty) = T" |
|
24 |
by (induct T rule: ty.induct) (simp_all add: perm_nat_def) |
|
25 |
||
26 |
nominal_datatype trm = |
|
27 |
Var name |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80142
diff
changeset
|
28 |
| Tuple trm trm (\<open>(1'\<langle>_,/ _'\<rangle>)\<close>) |
33189 | 29 |
| Abs ty "\<guillemotleft>name\<guillemotright>trm" |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80142
diff
changeset
|
30 |
| App trm trm (infixl \<open>\<cdot>\<close> 200) |
33189 | 31 |
| Let ty trm btrm |
32 |
and btrm = |
|
33 |
Base trm |
|
34 |
| Bind ty "\<guillemotleft>name\<guillemotright>btrm" |
|
35 |
||
36 |
abbreviation |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80142
diff
changeset
|
37 |
Abs_syn :: "name \<Rightarrow> ty \<Rightarrow> trm \<Rightarrow> trm" (\<open>(3\<lambda>_:_./ _)\<close> [0, 0, 10] 10) |
33189 | 38 |
where |
39 |
"\<lambda>x:T. t \<equiv> Abs T x t" |
|
40 |
||
58310 | 41 |
datatype pat = |
33189 | 42 |
PVar name ty |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80142
diff
changeset
|
43 |
| PTuple pat pat (\<open>(1'\<langle>\<langle>_,/ _'\<rangle>\<rangle>)\<close>) |
33189 | 44 |
|
45 |
(* FIXME: The following should be done automatically by the nominal package *) |
|
46 |
overloading pat_perm \<equiv> "perm :: name prm \<Rightarrow> pat \<Rightarrow> pat" (unchecked) |
|
47 |
begin |
|
48 |
||
49 |
primrec pat_perm |
|
50 |
where |
|
51 |
"pat_perm pi (PVar x ty) = PVar (pi \<bullet> x) (pi \<bullet> ty)" |
|
52 |
| "pat_perm pi \<langle>\<langle>p, q\<rangle>\<rangle> = \<langle>\<langle>pat_perm pi p, pat_perm pi q\<rangle>\<rangle>" |
|
53 |
||
54 |
end |
|
55 |
||
56 |
declare pat_perm.simps [eqvt] |
|
57 |
||
58 |
lemma supp_PVar [simp]: "((supp (PVar x T))::name set) = supp x" |
|
59 |
by (simp add: supp_def perm_fresh_fresh) |
|
60 |
||
61 |
lemma supp_PTuple [simp]: "((supp \<langle>\<langle>p, q\<rangle>\<rangle>)::name set) = supp p \<union> supp q" |
|
62 |
by (simp add: supp_def Collect_disj_eq del: disj_not1) |
|
63 |
||
64 |
instance pat :: pt_name |
|
80140 | 65 |
proof |
66 |
fix x :: pat |
|
67 |
show "([]::(name \<times> _) list) \<bullet> x = x" |
|
68 |
by (induct x) simp_all |
|
69 |
fix pi1 pi2 :: "(name \<times> name) list" |
|
70 |
show "(pi1 @ pi2) \<bullet> x = pi1 \<bullet> pi2 \<bullet> x" |
|
71 |
by (induct x) (simp_all add: pt_name2) |
|
72 |
assume "pi1 \<triangleq> pi2" |
|
73 |
then show "pi1 \<bullet> x = pi2 \<bullet> x" |
|
74 |
by (induct x) (simp_all add: pt_name3) |
|
33189 | 75 |
qed |
76 |
||
77 |
instance pat :: fs_name |
|
80140 | 78 |
proof |
79 |
fix x :: pat |
|
80 |
show "finite (supp x::name set)" |
|
81 |
by (induct x) (simp_all add: fin_supp) |
|
33189 | 82 |
qed |
83 |
||
84 |
(* the following function cannot be defined using nominal_primrec, *) |
|
85 |
(* since variable parameters are currently not allowed. *) |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80142
diff
changeset
|
86 |
primrec abs_pat :: "pat \<Rightarrow> btrm \<Rightarrow> btrm" (\<open>(3\<lambda>[_]./ _)\<close> [0, 10] 10) |
33189 | 87 |
where |
88 |
"(\<lambda>[PVar x T]. t) = Bind T x t" |
|
89 |
| "(\<lambda>[\<langle>\<langle>p, q\<rangle>\<rangle>]. t) = (\<lambda>[p]. \<lambda>[q]. t)" |
|
90 |
||
91 |
lemma abs_pat_eqvt [eqvt]: |
|
92 |
"(pi :: name prm) \<bullet> (\<lambda>[p]. t) = (\<lambda>[pi \<bullet> p]. (pi \<bullet> t))" |
|
93 |
by (induct p arbitrary: t) simp_all |
|
94 |
||
95 |
lemma abs_pat_fresh [simp]: |
|
96 |
"(x::name) \<sharp> (\<lambda>[p]. t) = (x \<in> supp p \<or> x \<sharp> t)" |
|
97 |
by (induct p arbitrary: t) (simp_all add: abs_fresh supp_atm) |
|
98 |
||
99 |
lemma abs_pat_alpha: |
|
100 |
assumes fresh: "((pi::name prm) \<bullet> supp p::name set) \<sharp>* t" |
|
101 |
and pi: "set pi \<subseteq> supp p \<times> pi \<bullet> supp p" |
|
102 |
shows "(\<lambda>[p]. t) = (\<lambda>[pi \<bullet> p]. pi \<bullet> t)" |
|
103 |
proof - |
|
104 |
note pt_name_inst at_name_inst pi |
|
105 |
moreover have "(supp p::name set) \<sharp>* (\<lambda>[p]. t)" |
|
106 |
by (simp add: fresh_star_def) |
|
107 |
moreover from fresh |
|
108 |
have "(pi \<bullet> supp p::name set) \<sharp>* (\<lambda>[p]. t)" |
|
109 |
by (simp add: fresh_star_def) |
|
110 |
ultimately have "pi \<bullet> (\<lambda>[p]. t) = (\<lambda>[p]. t)" |
|
111 |
by (rule pt_freshs_freshs) |
|
112 |
then show ?thesis by (simp add: eqvts) |
|
113 |
qed |
|
114 |
||
115 |
primrec pat_vars :: "pat \<Rightarrow> name list" |
|
116 |
where |
|
117 |
"pat_vars (PVar x T) = [x]" |
|
118 |
| "pat_vars \<langle>\<langle>p, q\<rangle>\<rangle> = pat_vars q @ pat_vars p" |
|
119 |
||
120 |
lemma pat_vars_eqvt [eqvt]: |
|
121 |
"(pi :: name prm) \<bullet> (pat_vars p) = pat_vars (pi \<bullet> p)" |
|
122 |
by (induct p rule: pat.induct) (simp_all add: eqvts) |
|
123 |
||
124 |
lemma set_pat_vars_supp: "set (pat_vars p) = supp p" |
|
125 |
by (induct p) (auto simp add: supp_atm) |
|
126 |
||
127 |
lemma distinct_eqvt [eqvt]: |
|
128 |
"(pi :: name prm) \<bullet> (distinct (xs::name list)) = distinct (pi \<bullet> xs)" |
|
129 |
by (induct xs) (simp_all add: eqvts) |
|
130 |
||
131 |
primrec pat_type :: "pat \<Rightarrow> ty" |
|
132 |
where |
|
133 |
"pat_type (PVar x T) = T" |
|
134 |
| "pat_type \<langle>\<langle>p, q\<rangle>\<rangle> = pat_type p \<otimes> pat_type q" |
|
135 |
||
136 |
lemma pat_type_eqvt [eqvt]: |
|
137 |
"(pi :: name prm) \<bullet> (pat_type p) = pat_type (pi \<bullet> p)" |
|
138 |
by (induct p) simp_all |
|
139 |
||
140 |
lemma pat_type_perm_eq: "pat_type ((pi :: name prm) \<bullet> p) = pat_type p" |
|
141 |
by (induct p) (simp_all add: perm_type) |
|
142 |
||
41798 | 143 |
type_synonym ctx = "(name \<times> ty) list" |
33189 | 144 |
|
145 |
inductive |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80142
diff
changeset
|
146 |
ptyping :: "pat \<Rightarrow> ty \<Rightarrow> ctx \<Rightarrow> bool" (\<open>\<turnstile> _ : _ \<Rightarrow> _\<close> [60, 60, 60] 60) |
33189 | 147 |
where |
148 |
PVar: "\<turnstile> PVar x T : T \<Rightarrow> [(x, T)]" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
149 |
| PTuple: "\<turnstile> p : T \<Rightarrow> \<Delta>\<^sub>1 \<Longrightarrow> \<turnstile> q : U \<Rightarrow> \<Delta>\<^sub>2 \<Longrightarrow> \<turnstile> \<langle>\<langle>p, q\<rangle>\<rangle> : T \<otimes> U \<Rightarrow> \<Delta>\<^sub>2 @ \<Delta>\<^sub>1" |
33189 | 150 |
|
151 |
lemma pat_vars_ptyping: |
|
152 |
assumes "\<turnstile> p : T \<Rightarrow> \<Delta>" |
|
153 |
shows "pat_vars p = map fst \<Delta>" using assms |
|
154 |
by induct simp_all |
|
155 |
||
156 |
inductive |
|
157 |
valid :: "ctx \<Rightarrow> bool" |
|
158 |
where |
|
159 |
Nil [intro!]: "valid []" |
|
160 |
| Cons [intro!]: "valid \<Gamma> \<Longrightarrow> x \<sharp> \<Gamma> \<Longrightarrow> valid ((x, T) # \<Gamma>)" |
|
161 |
||
162 |
inductive_cases validE[elim!]: "valid ((x, T) # \<Gamma>)" |
|
163 |
||
164 |
lemma fresh_ctxt_set_eq: "((x::name) \<sharp> (\<Gamma>::ctx)) = (x \<notin> fst ` set \<Gamma>)" |
|
165 |
by (induct \<Gamma>) (auto simp add: fresh_list_nil fresh_list_cons fresh_prod fresh_atm) |
|
166 |
||
167 |
lemma valid_distinct: "valid \<Gamma> = distinct (map fst \<Gamma>)" |
|
168 |
by (induct \<Gamma>) (auto simp add: fresh_ctxt_set_eq [symmetric]) |
|
169 |
||
170 |
abbreviation |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80142
diff
changeset
|
171 |
"sub_ctx" :: "ctx \<Rightarrow> ctx \<Rightarrow> bool" (\<open>_ \<sqsubseteq> _\<close>) |
33189 | 172 |
where |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
173 |
"\<Gamma>\<^sub>1 \<sqsubseteq> \<Gamma>\<^sub>2 \<equiv> \<forall>x. x \<in> set \<Gamma>\<^sub>1 \<longrightarrow> x \<in> set \<Gamma>\<^sub>2" |
33189 | 174 |
|
175 |
abbreviation |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80142
diff
changeset
|
176 |
Let_syn :: "pat \<Rightarrow> trm \<Rightarrow> trm \<Rightarrow> trm" (\<open>(LET (_ =/ _)/ IN (_))\<close> 10) |
33189 | 177 |
where |
178 |
"LET p = t IN u \<equiv> Let (pat_type p) t (\<lambda>[p]. Base u)" |
|
179 |
||
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80142
diff
changeset
|
180 |
inductive typing :: "ctx \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool" (\<open>_ \<turnstile> _ : _\<close> [60, 60, 60] 60) |
33189 | 181 |
where |
182 |
Var [intro]: "valid \<Gamma> \<Longrightarrow> (x, T) \<in> set \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T" |
|
183 |
| Tuple [intro]: "\<Gamma> \<turnstile> t : T \<Longrightarrow> \<Gamma> \<turnstile> u : U \<Longrightarrow> \<Gamma> \<turnstile> \<langle>t, u\<rangle> : T \<otimes> U" |
|
184 |
| Abs [intro]: "(x, T) # \<Gamma> \<turnstile> t : U \<Longrightarrow> \<Gamma> \<turnstile> (\<lambda>x:T. t) : T \<rightarrow> U" |
|
185 |
| App [intro]: "\<Gamma> \<turnstile> t : T \<rightarrow> U \<Longrightarrow> \<Gamma> \<turnstile> u : T \<Longrightarrow> \<Gamma> \<turnstile> t \<cdot> u : U" |
|
186 |
| Let: "((supp p)::name set) \<sharp>* t \<Longrightarrow> |
|
187 |
\<Gamma> \<turnstile> t : T \<Longrightarrow> \<turnstile> p : T \<Rightarrow> \<Delta> \<Longrightarrow> \<Delta> @ \<Gamma> \<turnstile> u : U \<Longrightarrow> |
|
188 |
\<Gamma> \<turnstile> (LET p = t IN u) : U" |
|
189 |
||
190 |
equivariance ptyping |
|
191 |
||
192 |
equivariance valid |
|
193 |
||
194 |
equivariance typing |
|
195 |
||
196 |
lemma valid_typing: |
|
197 |
assumes "\<Gamma> \<turnstile> t : T" |
|
198 |
shows "valid \<Gamma>" using assms |
|
199 |
by induct auto |
|
200 |
||
201 |
lemma pat_var: |
|
202 |
assumes "\<turnstile> p : T \<Rightarrow> \<Delta>" |
|
203 |
shows "(supp p::name set) = supp \<Delta>" using assms |
|
204 |
by induct (auto simp add: supp_list_nil supp_list_cons supp_prod supp_list_append) |
|
205 |
||
206 |
lemma valid_app_fresh: |
|
207 |
assumes "valid (\<Delta> @ \<Gamma>)" and "(x::name) \<in> supp \<Delta>" |
|
208 |
shows "x \<sharp> \<Gamma>" using assms |
|
209 |
by (induct \<Delta>) |
|
210 |
(auto simp add: supp_list_nil supp_list_cons supp_prod supp_atm fresh_list_append) |
|
211 |
||
212 |
lemma pat_freshs: |
|
213 |
assumes "\<turnstile> p : T \<Rightarrow> \<Delta>" |
|
214 |
shows "(supp p::name set) \<sharp>* c = (supp \<Delta>::name set) \<sharp>* c" using assms |
|
215 |
by (auto simp add: fresh_star_def pat_var) |
|
216 |
||
217 |
lemma valid_app_mono: |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
218 |
assumes "valid (\<Delta> @ \<Gamma>\<^sub>1)" and "(supp \<Delta>::name set) \<sharp>* \<Gamma>\<^sub>2" and "valid \<Gamma>\<^sub>2" and "\<Gamma>\<^sub>1 \<sqsubseteq> \<Gamma>\<^sub>2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
219 |
shows "valid (\<Delta> @ \<Gamma>\<^sub>2)" using assms |
33189 | 220 |
by (induct \<Delta>) |
221 |
(auto simp add: supp_list_cons fresh_star_Un_elim supp_prod |
|
222 |
fresh_list_append supp_atm fresh_star_insert_elim fresh_star_empty_elim) |
|
223 |
||
224 |
nominal_inductive2 typing |
|
225 |
avoids |
|
226 |
Abs: "{x}" |
|
227 |
| Let: "(supp p)::name set" |
|
228 |
by (auto simp add: fresh_star_def abs_fresh fin_supp pat_var |
|
229 |
dest!: valid_typing valid_app_fresh) |
|
230 |
||
231 |
lemma better_T_Let [intro]: |
|
232 |
assumes t: "\<Gamma> \<turnstile> t : T" and p: "\<turnstile> p : T \<Rightarrow> \<Delta>" and u: "\<Delta> @ \<Gamma> \<turnstile> u : U" |
|
233 |
shows "\<Gamma> \<turnstile> (LET p = t IN u) : U" |
|
234 |
proof - |
|
235 |
obtain pi::"name prm" where pi: "(pi \<bullet> (supp p::name set)) \<sharp>* (t, Base u, \<Gamma>)" |
|
236 |
and pi': "set pi \<subseteq> supp p \<times> (pi \<bullet> supp p)" |
|
237 |
by (rule at_set_avoiding [OF at_name_inst fin_supp fin_supp]) |
|
238 |
from p u have p_fresh: "(supp p::name set) \<sharp>* \<Gamma>" |
|
239 |
by (auto simp add: fresh_star_def pat_var dest!: valid_typing valid_app_fresh) |
|
240 |
from pi have p_fresh': "(pi \<bullet> (supp p::name set)) \<sharp>* \<Gamma>" |
|
241 |
by (simp add: fresh_star_prod_elim) |
|
242 |
from pi have p_fresh'': "(pi \<bullet> (supp p::name set)) \<sharp>* Base u" |
|
243 |
by (simp add: fresh_star_prod_elim) |
|
244 |
from pi have "(supp (pi \<bullet> p)::name set) \<sharp>* t" |
|
245 |
by (simp add: fresh_star_prod_elim eqvts) |
|
246 |
moreover note t |
|
247 |
moreover from p have "pi \<bullet> (\<turnstile> p : T \<Rightarrow> \<Delta>)" by (rule perm_boolI) |
|
248 |
then have "\<turnstile> (pi \<bullet> p) : T \<Rightarrow> (pi \<bullet> \<Delta>)" by (simp add: eqvts perm_type) |
|
249 |
moreover from u have "pi \<bullet> (\<Delta> @ \<Gamma> \<turnstile> u : U)" by (rule perm_boolI) |
|
250 |
with pt_freshs_freshs [OF pt_name_inst at_name_inst pi' p_fresh p_fresh'] |
|
251 |
have "(pi \<bullet> \<Delta>) @ \<Gamma> \<turnstile> (pi \<bullet> u) : U" by (simp add: eqvts perm_type) |
|
252 |
ultimately have "\<Gamma> \<turnstile> (LET (pi \<bullet> p) = t IN (pi \<bullet> u)) : U" |
|
253 |
by (rule Let) |
|
254 |
then show ?thesis by (simp add: abs_pat_alpha [OF p_fresh'' pi'] pat_type_perm_eq) |
|
255 |
qed |
|
256 |
||
257 |
lemma weakening: |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
258 |
assumes "\<Gamma>\<^sub>1 \<turnstile> t : T" and "valid \<Gamma>\<^sub>2" and "\<Gamma>\<^sub>1 \<sqsubseteq> \<Gamma>\<^sub>2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
259 |
shows "\<Gamma>\<^sub>2 \<turnstile> t : T" using assms |
80140 | 260 |
proof (nominal_induct \<Gamma>\<^sub>1 t T avoiding: \<Gamma>\<^sub>2 rule: typing.strong_induct) |
261 |
case (Abs x T \<Gamma> t U) |
|
262 |
then show ?case |
|
263 |
by (simp add: typing.Abs valid.Cons) |
|
264 |
next |
|
265 |
case (Let p t \<Gamma> T \<Delta> u U) |
|
266 |
then show ?case |
|
267 |
by (smt (verit, ccfv_threshold) Un_iff pat_freshs set_append typing.simps valid_app_mono valid_typing) |
|
268 |
qed auto |
|
33189 | 269 |
|
270 |
inductive |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80142
diff
changeset
|
271 |
match :: "pat \<Rightarrow> trm \<Rightarrow> (name \<times> trm) list \<Rightarrow> bool" (\<open>\<turnstile> _ \<rhd> _ \<Rightarrow> _\<close> [50, 50, 50] 50) |
33189 | 272 |
where |
273 |
PVar: "\<turnstile> PVar x T \<rhd> t \<Rightarrow> [(x, t)]" |
|
274 |
| PProd: "\<turnstile> p \<rhd> t \<Rightarrow> \<theta> \<Longrightarrow> \<turnstile> q \<rhd> u \<Rightarrow> \<theta>' \<Longrightarrow> \<turnstile> \<langle>\<langle>p, q\<rangle>\<rangle> \<rhd> \<langle>t, u\<rangle> \<Rightarrow> \<theta> @ \<theta>'" |
|
275 |
||
276 |
fun |
|
277 |
lookup :: "(name \<times> trm) list \<Rightarrow> name \<Rightarrow> trm" |
|
278 |
where |
|
279 |
"lookup [] x = Var x" |
|
280 |
| "lookup ((y, e) # \<theta>) x = (if x = y then e else lookup \<theta> x)" |
|
281 |
||
282 |
lemma lookup_eqvt[eqvt]: |
|
283 |
fixes pi :: "name prm" |
|
284 |
and \<theta> :: "(name \<times> trm) list" |
|
285 |
and X :: "name" |
|
286 |
shows "pi \<bullet> (lookup \<theta> X) = lookup (pi \<bullet> \<theta>) (pi \<bullet> X)" |
|
287 |
by (induct \<theta>) (auto simp add: eqvts) |
|
288 |
||
289 |
nominal_primrec |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80142
diff
changeset
|
290 |
psubst :: "(name \<times> trm) list \<Rightarrow> trm \<Rightarrow> trm" (\<open>_\<lparr>_\<rparr>\<close> [95,0] 210) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80142
diff
changeset
|
291 |
and psubstb :: "(name \<times> trm) list \<Rightarrow> btrm \<Rightarrow> btrm" (\<open>_\<lparr>_\<rparr>\<^sub>b\<close> [95,0] 210) |
33189 | 292 |
where |
293 |
"\<theta>\<lparr>Var x\<rparr> = (lookup \<theta> x)" |
|
294 |
| "\<theta>\<lparr>t \<cdot> u\<rparr> = \<theta>\<lparr>t\<rparr> \<cdot> \<theta>\<lparr>u\<rparr>" |
|
295 |
| "\<theta>\<lparr>\<langle>t, u\<rangle>\<rparr> = \<langle>\<theta>\<lparr>t\<rparr>, \<theta>\<lparr>u\<rparr>\<rangle>" |
|
296 |
| "\<theta>\<lparr>Let T t u\<rparr> = Let T (\<theta>\<lparr>t\<rparr>) (\<theta>\<lparr>u\<rparr>\<^sub>b)" |
|
297 |
| "x \<sharp> \<theta> \<Longrightarrow> \<theta>\<lparr>\<lambda>x:T. t\<rparr> = (\<lambda>x:T. \<theta>\<lparr>t\<rparr>)" |
|
298 |
| "\<theta>\<lparr>Base t\<rparr>\<^sub>b = Base (\<theta>\<lparr>t\<rparr>)" |
|
299 |
| "x \<sharp> \<theta> \<Longrightarrow> \<theta>\<lparr>Bind T x t\<rparr>\<^sub>b = Bind T x (\<theta>\<lparr>t\<rparr>\<^sub>b)" |
|
80140 | 300 |
by (finite_guess | simp add: abs_fresh | fresh_guess)+ |
33189 | 301 |
|
302 |
lemma lookup_fresh: |
|
303 |
"x = y \<longrightarrow> x \<in> set (map fst \<theta>) \<Longrightarrow> \<forall>(y, t)\<in>set \<theta>. x \<sharp> t \<Longrightarrow> x \<sharp> lookup \<theta> y" |
|
80140 | 304 |
by (induct \<theta>) (use fresh_atm in force)+ |
33189 | 305 |
|
306 |
lemma psubst_fresh: |
|
307 |
assumes "x \<in> set (map fst \<theta>)" and "\<forall>(y, t)\<in>set \<theta>. x \<sharp> t" |
|
308 |
shows "x \<sharp> \<theta>\<lparr>t\<rparr>" and "x \<sharp> \<theta>\<lparr>t'\<rparr>\<^sub>b" using assms |
|
80140 | 309 |
proof (nominal_induct t and t' avoiding: \<theta> rule: trm_btrm.strong_inducts) |
310 |
case (Var name) |
|
311 |
then show ?case |
|
312 |
by (metis lookup_fresh simps(1)) |
|
313 |
qed (auto simp: abs_fresh) |
|
33189 | 314 |
|
315 |
lemma psubst_eqvt[eqvt]: |
|
316 |
fixes pi :: "name prm" |
|
317 |
shows "pi \<bullet> (\<theta>\<lparr>t\<rparr>) = (pi \<bullet> \<theta>)\<lparr>pi \<bullet> t\<rparr>" |
|
318 |
and "pi \<bullet> (\<theta>\<lparr>t'\<rparr>\<^sub>b) = (pi \<bullet> \<theta>)\<lparr>pi \<bullet> t'\<rparr>\<^sub>b" |
|
319 |
by (nominal_induct t and t' avoiding: \<theta> rule: trm_btrm.strong_inducts) |
|
320 |
(simp_all add: eqvts fresh_bij) |
|
321 |
||
322 |
abbreviation |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80142
diff
changeset
|
323 |
subst :: "trm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> trm" (\<open>_[_\<mapsto>_]\<close> [100,0,0] 100) |
33189 | 324 |
where |
325 |
"t[x\<mapsto>t'] \<equiv> [(x,t')]\<lparr>t\<rparr>" |
|
326 |
||
327 |
abbreviation |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80142
diff
changeset
|
328 |
substb :: "btrm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> btrm" (\<open>_[_\<mapsto>_]\<^sub>b\<close> [100,0,0] 100) |
33189 | 329 |
where |
330 |
"t[x\<mapsto>t']\<^sub>b \<equiv> [(x,t')]\<lparr>t\<rparr>\<^sub>b" |
|
331 |
||
332 |
lemma lookup_forget: |
|
333 |
"(supp (map fst \<theta>)::name set) \<sharp>* x \<Longrightarrow> lookup \<theta> x = Var x" |
|
334 |
by (induct \<theta>) (auto simp add: split_paired_all fresh_star_def fresh_atm supp_list_cons supp_atm) |
|
335 |
||
336 |
lemma supp_fst: "(x::name) \<in> supp (map fst (\<theta>::(name \<times> trm) list)) \<Longrightarrow> x \<in> supp \<theta>" |
|
337 |
by (induct \<theta>) (auto simp add: supp_list_nil supp_list_cons supp_prod) |
|
338 |
||
339 |
lemma psubst_forget: |
|
340 |
"(supp (map fst \<theta>)::name set) \<sharp>* t \<Longrightarrow> \<theta>\<lparr>t\<rparr> = t" |
|
341 |
"(supp (map fst \<theta>)::name set) \<sharp>* t' \<Longrightarrow> \<theta>\<lparr>t'\<rparr>\<^sub>b = t'" |
|
80140 | 342 |
proof (nominal_induct t and t' avoiding: \<theta> rule: trm_btrm.strong_inducts) |
343 |
case (Var name) |
|
344 |
then show ?case |
|
345 |
by (simp add: fresh_star_set lookup_forget) |
|
346 |
next |
|
347 |
case (Abs ty name trm) |
|
348 |
then show ?case |
|
349 |
apply (simp add: fresh_def) |
|
350 |
by (metis abs_fresh(1) fresh_star_set supp_fst trm.fresh(3)) |
|
351 |
next |
|
352 |
case (Bind ty name btrm) |
|
353 |
then show ?case |
|
354 |
apply (simp add: fresh_def) |
|
355 |
by (metis abs_fresh(1) btrm.fresh(2) fresh_star_set supp_fst) |
|
356 |
qed (auto simp: fresh_star_set) |
|
33189 | 357 |
|
80142 | 358 |
lemma psubst_nil[simp]: "[]\<lparr>t\<rparr> = t" "[]\<lparr>t'\<rparr>\<^sub>b = t'" |
33189 | 359 |
by (induct t and t' rule: trm_btrm.inducts) (simp_all add: fresh_list_nil) |
360 |
||
361 |
lemma psubst_cons: |
|
362 |
assumes "(supp (map fst \<theta>)::name set) \<sharp>* u" |
|
363 |
shows "((x, u) # \<theta>)\<lparr>t\<rparr> = \<theta>\<lparr>t[x\<mapsto>u]\<rparr>" and "((x, u) # \<theta>)\<lparr>t'\<rparr>\<^sub>b = \<theta>\<lparr>t'[x\<mapsto>u]\<^sub>b\<rparr>\<^sub>b" |
|
364 |
using assms |
|
365 |
by (nominal_induct t and t' avoiding: x u \<theta> rule: trm_btrm.strong_inducts) |
|
366 |
(simp_all add: fresh_list_nil fresh_list_cons psubst_forget) |
|
367 |
||
368 |
lemma psubst_append: |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
369 |
"(supp (map fst (\<theta>\<^sub>1 @ \<theta>\<^sub>2))::name set) \<sharp>* map snd (\<theta>\<^sub>1 @ \<theta>\<^sub>2) \<Longrightarrow> (\<theta>\<^sub>1 @ \<theta>\<^sub>2)\<lparr>t\<rparr> = \<theta>\<^sub>2\<lparr>\<theta>\<^sub>1\<lparr>t\<rparr>\<rparr>" |
80142 | 370 |
proof (induct \<theta>\<^sub>1 arbitrary: t) |
371 |
case Nil |
|
372 |
then show ?case |
|
373 |
by (auto simp: psubst_nil) |
|
374 |
next |
|
375 |
case (Cons a \<theta>\<^sub>1) |
|
376 |
then show ?case |
|
377 |
proof (cases a) |
|
378 |
case (Pair a b) |
|
379 |
with Cons show ?thesis |
|
380 |
apply (simp add: supp_list_cons fresh_star_set fresh_list_cons) |
|
381 |
by (metis Un_iff fresh_star_set map_append psubst_cons(1) supp_list_append) |
|
382 |
qed |
|
383 |
qed |
|
33189 | 384 |
|
385 |
lemma abs_pat_psubst [simp]: |
|
386 |
"(supp p::name set) \<sharp>* \<theta> \<Longrightarrow> \<theta>\<lparr>\<lambda>[p]. t\<rparr>\<^sub>b = (\<lambda>[p]. \<theta>\<lparr>t\<rparr>\<^sub>b)" |
|
387 |
by (induct p arbitrary: t) (auto simp add: fresh_star_def supp_atm) |
|
388 |
||
389 |
lemma valid_insert: |
|
390 |
assumes "valid (\<Delta> @ [(x, T)] @ \<Gamma>)" |
|
391 |
shows "valid (\<Delta> @ \<Gamma>)" using assms |
|
392 |
by (induct \<Delta>) |
|
393 |
(auto simp add: fresh_list_append fresh_list_cons) |
|
394 |
||
395 |
lemma fresh_set: |
|
396 |
shows "y \<sharp> xs = (\<forall>x\<in>set xs. y \<sharp> x)" |
|
397 |
by (induct xs) (simp_all add: fresh_list_nil fresh_list_cons) |
|
398 |
||
399 |
lemma context_unique: |
|
400 |
assumes "valid \<Gamma>" |
|
401 |
and "(x, T) \<in> set \<Gamma>" |
|
402 |
and "(x, U) \<in> set \<Gamma>" |
|
403 |
shows "T = U" using assms |
|
404 |
by induct (auto simp add: fresh_set fresh_prod fresh_atm) |
|
405 |
||
406 |
lemma subst_type_aux: |
|
407 |
assumes a: "\<Delta> @ [(x, U)] @ \<Gamma> \<turnstile> t : T" |
|
408 |
and b: "\<Gamma> \<turnstile> u : U" |
|
409 |
shows "\<Delta> @ \<Gamma> \<turnstile> t[x\<mapsto>u] : T" using a b |
|
410 |
proof (nominal_induct \<Gamma>'\<equiv>"\<Delta> @ [(x, U)] @ \<Gamma>" t T avoiding: x u \<Delta> rule: typing.strong_induct) |
|
34915 | 411 |
case (Var y T x u \<Delta>) |
63167 | 412 |
from \<open>valid (\<Delta> @ [(x, U)] @ \<Gamma>)\<close> |
34915 | 413 |
have valid: "valid (\<Delta> @ \<Gamma>)" by (rule valid_insert) |
33189 | 414 |
show "\<Delta> @ \<Gamma> \<turnstile> Var y[x\<mapsto>u] : T" |
415 |
proof cases |
|
416 |
assume eq: "x = y" |
|
34915 | 417 |
from Var eq have "T = U" by (auto intro: context_unique) |
418 |
with Var eq valid show "\<Delta> @ \<Gamma> \<turnstile> Var y[x\<mapsto>u] : T" by (auto intro: weakening) |
|
33189 | 419 |
next |
420 |
assume ineq: "x \<noteq> y" |
|
34915 | 421 |
from Var ineq have "(y, T) \<in> set (\<Delta> @ \<Gamma>)" by simp |
422 |
then show "\<Delta> @ \<Gamma> \<turnstile> Var y[x\<mapsto>u] : T" using ineq valid by auto |
|
33189 | 423 |
qed |
424 |
next |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
425 |
case (Tuple t\<^sub>1 T\<^sub>1 t\<^sub>2 T\<^sub>2) |
63167 | 426 |
from refl \<open>\<Gamma> \<turnstile> u : U\<close> |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
427 |
have "\<Delta> @ \<Gamma> \<turnstile> t\<^sub>1[x\<mapsto>u] : T\<^sub>1" by (rule Tuple) |
63167 | 428 |
moreover from refl \<open>\<Gamma> \<turnstile> u : U\<close> |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
429 |
have "\<Delta> @ \<Gamma> \<turnstile> t\<^sub>2[x\<mapsto>u] : T\<^sub>2" by (rule Tuple) |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
430 |
ultimately have "\<Delta> @ \<Gamma> \<turnstile> \<langle>t\<^sub>1[x\<mapsto>u], t\<^sub>2[x\<mapsto>u]\<rangle> : T\<^sub>1 \<otimes> T\<^sub>2" .. |
33189 | 431 |
then show ?case by simp |
432 |
next |
|
34915 | 433 |
case (Let p t T \<Delta>' s S) |
63167 | 434 |
from refl \<open>\<Gamma> \<turnstile> u : U\<close> |
33189 | 435 |
have "\<Delta> @ \<Gamma> \<turnstile> t[x\<mapsto>u] : T" by (rule Let) |
63167 | 436 |
moreover note \<open>\<turnstile> p : T \<Rightarrow> \<Delta>'\<close> |
34915 | 437 |
moreover have "\<Delta>' @ (\<Delta> @ [(x, U)] @ \<Gamma>) = (\<Delta>' @ \<Delta>) @ [(x, U)] @ \<Gamma>" by simp |
63167 | 438 |
then have "(\<Delta>' @ \<Delta>) @ \<Gamma> \<turnstile> s[x\<mapsto>u] : S" using \<open>\<Gamma> \<turnstile> u : U\<close> by (rule Let) |
33189 | 439 |
then have "\<Delta>' @ \<Delta> @ \<Gamma> \<turnstile> s[x\<mapsto>u] : S" by simp |
440 |
ultimately have "\<Delta> @ \<Gamma> \<turnstile> (LET p = t[x\<mapsto>u] IN s[x\<mapsto>u]) : S" |
|
441 |
by (rule better_T_Let) |
|
442 |
moreover from Let have "(supp p::name set) \<sharp>* [(x, u)]" |
|
443 |
by (simp add: fresh_star_def fresh_list_nil fresh_list_cons) |
|
444 |
ultimately show ?case by simp |
|
445 |
next |
|
34915 | 446 |
case (Abs y T t S) |
447 |
have "(y, T) # \<Delta> @ [(x, U)] @ \<Gamma> = ((y, T) # \<Delta>) @ [(x, U)] @ \<Gamma>" |
|
33189 | 448 |
by simp |
63167 | 449 |
then have "((y, T) # \<Delta>) @ \<Gamma> \<turnstile> t[x\<mapsto>u] : S" using \<open>\<Gamma> \<turnstile> u : U\<close> by (rule Abs) |
33189 | 450 |
then have "(y, T) # \<Delta> @ \<Gamma> \<turnstile> t[x\<mapsto>u] : S" by simp |
451 |
then have "\<Delta> @ \<Gamma> \<turnstile> (\<lambda>y:T. t[x\<mapsto>u]) : T \<rightarrow> S" |
|
452 |
by (rule typing.Abs) |
|
453 |
moreover from Abs have "y \<sharp> [(x, u)]" |
|
454 |
by (simp add: fresh_list_nil fresh_list_cons) |
|
455 |
ultimately show ?case by simp |
|
456 |
next |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
457 |
case (App t\<^sub>1 T S t\<^sub>2) |
63167 | 458 |
from refl \<open>\<Gamma> \<turnstile> u : U\<close> |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
459 |
have "\<Delta> @ \<Gamma> \<turnstile> t\<^sub>1[x\<mapsto>u] : T \<rightarrow> S" by (rule App) |
63167 | 460 |
moreover from refl \<open>\<Gamma> \<turnstile> u : U\<close> |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
461 |
have "\<Delta> @ \<Gamma> \<turnstile> t\<^sub>2[x\<mapsto>u] : T" by (rule App) |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
462 |
ultimately have "\<Delta> @ \<Gamma> \<turnstile> (t\<^sub>1[x\<mapsto>u]) \<cdot> (t\<^sub>2[x\<mapsto>u]) : S" |
33189 | 463 |
by (rule typing.App) |
464 |
then show ?case by simp |
|
465 |
qed |
|
466 |
||
467 |
lemmas subst_type = subst_type_aux [of "[]", simplified] |
|
468 |
||
469 |
lemma match_supp_fst: |
|
470 |
assumes "\<turnstile> p \<rhd> u \<Rightarrow> \<theta>" shows "(supp (map fst \<theta>)::name set) = supp p" using assms |
|
471 |
by induct (simp_all add: supp_list_nil supp_list_cons supp_list_append) |
|
472 |
||
473 |
lemma match_supp_snd: |
|
474 |
assumes "\<turnstile> p \<rhd> u \<Rightarrow> \<theta>" shows "(supp (map snd \<theta>)::name set) = supp u" using assms |
|
475 |
by induct (simp_all add: supp_list_nil supp_list_cons supp_list_append trm.supp) |
|
476 |
||
477 |
lemma match_fresh: "\<turnstile> p \<rhd> u \<Rightarrow> \<theta> \<Longrightarrow> (supp p::name set) \<sharp>* u \<Longrightarrow> |
|
478 |
(supp (map fst \<theta>)::name set) \<sharp>* map snd \<theta>" |
|
479 |
by (simp add: fresh_star_def fresh_def match_supp_fst match_supp_snd) |
|
480 |
||
481 |
lemma match_type_aux: |
|
482 |
assumes "\<turnstile> p : U \<Rightarrow> \<Delta>" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
483 |
and "\<Gamma>\<^sub>2 \<turnstile> u : U" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
484 |
and "\<Gamma>\<^sub>1 @ \<Delta> @ \<Gamma>\<^sub>2 \<turnstile> t : T" |
33189 | 485 |
and "\<turnstile> p \<rhd> u \<Rightarrow> \<theta>" |
486 |
and "(supp p::name set) \<sharp>* u" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
487 |
shows "\<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile> \<theta>\<lparr>t\<rparr> : T" using assms |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
488 |
proof (induct arbitrary: \<Gamma>\<^sub>1 \<Gamma>\<^sub>2 t u T \<theta>) |
33189 | 489 |
case (PVar x U) |
63167 | 490 |
from \<open>\<Gamma>\<^sub>1 @ [(x, U)] @ \<Gamma>\<^sub>2 \<turnstile> t : T\<close> \<open>\<Gamma>\<^sub>2 \<turnstile> u : U\<close> |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
491 |
have "\<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile> t[x\<mapsto>u] : T" by (rule subst_type_aux) |
63167 | 492 |
moreover from \<open>\<turnstile> PVar x U \<rhd> u \<Rightarrow> \<theta>\<close> have "\<theta> = [(x, u)]" |
33189 | 493 |
by cases simp_all |
494 |
ultimately show ?case by simp |
|
495 |
next |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
496 |
case (PTuple p S \<Delta>\<^sub>1 q U \<Delta>\<^sub>2) |
63167 | 497 |
from \<open>\<turnstile> \<langle>\<langle>p, q\<rangle>\<rangle> \<rhd> u \<Rightarrow> \<theta>\<close> obtain u\<^sub>1 u\<^sub>2 \<theta>\<^sub>1 \<theta>\<^sub>2 |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
498 |
where u: "u = \<langle>u\<^sub>1, u\<^sub>2\<rangle>" and \<theta>: "\<theta> = \<theta>\<^sub>1 @ \<theta>\<^sub>2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
499 |
and p: "\<turnstile> p \<rhd> u\<^sub>1 \<Rightarrow> \<theta>\<^sub>1" and q: "\<turnstile> q \<rhd> u\<^sub>2 \<Rightarrow> \<theta>\<^sub>2" |
33189 | 500 |
by cases simp_all |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
501 |
with PTuple have "\<Gamma>\<^sub>2 \<turnstile> \<langle>u\<^sub>1, u\<^sub>2\<rangle> : S \<otimes> U" by simp |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
502 |
then obtain u\<^sub>1: "\<Gamma>\<^sub>2 \<turnstile> u\<^sub>1 : S" and u\<^sub>2: "\<Gamma>\<^sub>2 \<turnstile> u\<^sub>2 : U" |
33189 | 503 |
by cases (simp_all add: ty.inject trm.inject) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
504 |
note u\<^sub>1 |
63167 | 505 |
moreover from \<open>\<Gamma>\<^sub>1 @ (\<Delta>\<^sub>2 @ \<Delta>\<^sub>1) @ \<Gamma>\<^sub>2 \<turnstile> t : T\<close> |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
506 |
have "(\<Gamma>\<^sub>1 @ \<Delta>\<^sub>2) @ \<Delta>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile> t : T" by simp |
33189 | 507 |
moreover note p |
63167 | 508 |
moreover from \<open>supp \<langle>\<langle>p, q\<rangle>\<rangle> \<sharp>* u\<close> and u |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
509 |
have "(supp p::name set) \<sharp>* u\<^sub>1" by (simp add: fresh_star_def) |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
510 |
ultimately have \<theta>\<^sub>1: "(\<Gamma>\<^sub>1 @ \<Delta>\<^sub>2) @ \<Gamma>\<^sub>2 \<turnstile> \<theta>\<^sub>1\<lparr>t\<rparr> : T" |
33189 | 511 |
by (rule PTuple) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
512 |
note u\<^sub>2 |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
513 |
moreover from \<theta>\<^sub>1 |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
514 |
have "\<Gamma>\<^sub>1 @ \<Delta>\<^sub>2 @ \<Gamma>\<^sub>2 \<turnstile> \<theta>\<^sub>1\<lparr>t\<rparr> : T" by simp |
33189 | 515 |
moreover note q |
63167 | 516 |
moreover from \<open>supp \<langle>\<langle>p, q\<rangle>\<rangle> \<sharp>* u\<close> and u |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
517 |
have "(supp q::name set) \<sharp>* u\<^sub>2" by (simp add: fresh_star_def) |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
518 |
ultimately have "\<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile> \<theta>\<^sub>2\<lparr>\<theta>\<^sub>1\<lparr>t\<rparr>\<rparr> : T" |
33189 | 519 |
by (rule PTuple) |
63167 | 520 |
moreover from \<open>\<turnstile> \<langle>\<langle>p, q\<rangle>\<rangle> \<rhd> u \<Rightarrow> \<theta>\<close> \<open>supp \<langle>\<langle>p, q\<rangle>\<rangle> \<sharp>* u\<close> |
33189 | 521 |
have "(supp (map fst \<theta>)::name set) \<sharp>* map snd \<theta>" |
522 |
by (rule match_fresh) |
|
523 |
ultimately show ?case using \<theta> by (simp add: psubst_append) |
|
524 |
qed |
|
525 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
526 |
lemmas match_type = match_type_aux [where \<Gamma>\<^sub>1="[]", simplified] |
33189 | 527 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80142
diff
changeset
|
528 |
inductive eval :: "trm \<Rightarrow> trm \<Rightarrow> bool" (\<open>_ \<longmapsto> _\<close> [60,60] 60) |
33189 | 529 |
where |
530 |
TupleL: "t \<longmapsto> t' \<Longrightarrow> \<langle>t, u\<rangle> \<longmapsto> \<langle>t', u\<rangle>" |
|
531 |
| TupleR: "u \<longmapsto> u' \<Longrightarrow> \<langle>t, u\<rangle> \<longmapsto> \<langle>t, u'\<rangle>" |
|
532 |
| Abs: "t \<longmapsto> t' \<Longrightarrow> (\<lambda>x:T. t) \<longmapsto> (\<lambda>x:T. t')" |
|
533 |
| AppL: "t \<longmapsto> t' \<Longrightarrow> t \<cdot> u \<longmapsto> t' \<cdot> u" |
|
534 |
| AppR: "u \<longmapsto> u' \<Longrightarrow> t \<cdot> u \<longmapsto> t \<cdot> u'" |
|
535 |
| Beta: "x \<sharp> u \<Longrightarrow> (\<lambda>x:T. t) \<cdot> u \<longmapsto> t[x\<mapsto>u]" |
|
536 |
| Let: "((supp p)::name set) \<sharp>* t \<Longrightarrow> distinct (pat_vars p) \<Longrightarrow> |
|
537 |
\<turnstile> p \<rhd> t \<Rightarrow> \<theta> \<Longrightarrow> (LET p = t IN u) \<longmapsto> \<theta>\<lparr>u\<rparr>" |
|
538 |
||
539 |
equivariance match |
|
540 |
||
541 |
equivariance eval |
|
542 |
||
543 |
lemma match_vars: |
|
544 |
assumes "\<turnstile> p \<rhd> t \<Rightarrow> \<theta>" and "x \<in> supp p" |
|
545 |
shows "x \<in> set (map fst \<theta>)" using assms |
|
546 |
by induct (auto simp add: supp_atm) |
|
547 |
||
548 |
lemma match_fresh_mono: |
|
549 |
assumes "\<turnstile> p \<rhd> t \<Rightarrow> \<theta>" and "(x::name) \<sharp> t" |
|
550 |
shows "\<forall>(y, t)\<in>set \<theta>. x \<sharp> t" using assms |
|
551 |
by induct auto |
|
552 |
||
553 |
nominal_inductive2 eval |
|
554 |
avoids |
|
555 |
Abs: "{x}" |
|
556 |
| Beta: "{x}" |
|
557 |
| Let: "(supp p)::name set" |
|
80140 | 558 |
proof (simp_all add: fresh_star_def abs_fresh fin_supp) |
559 |
show "x \<sharp> t[x\<mapsto>u]" if "x \<sharp> u" for x t u |
|
560 |
by (simp add: \<open>x \<sharp> u\<close> psubst_fresh(1)) |
|
561 |
next |
|
562 |
show "\<forall>x\<in>supp p. (x::name) \<sharp> \<theta>\<lparr>u\<rparr>" |
|
563 |
if "\<forall>x\<in>supp p. (x::name) \<sharp> t" and "\<turnstile> p \<rhd> t \<Rightarrow> \<theta>" |
|
564 |
for p t \<theta> u |
|
565 |
by (meson that match_fresh_mono match_vars psubst_fresh(1)) |
|
566 |
qed |
|
33189 | 567 |
|
568 |
lemma typing_case_Abs: |
|
569 |
assumes ty: "\<Gamma> \<turnstile> (\<lambda>x:T. t) : S" |
|
570 |
and fresh: "x \<sharp> \<Gamma>" |
|
571 |
and R: "\<And>U. S = T \<rightarrow> U \<Longrightarrow> (x, T) # \<Gamma> \<turnstile> t : U \<Longrightarrow> P" |
|
572 |
shows P using ty |
|
573 |
proof cases |
|
34990 | 574 |
case (Abs x' T' t' U) |
33189 | 575 |
obtain y::name where y: "y \<sharp> (x, \<Gamma>, \<lambda>x':T'. t')" |
576 |
by (rule exists_fresh) (auto intro: fin_supp) |
|
63167 | 577 |
from \<open>(\<lambda>x:T. t) = (\<lambda>x':T'. t')\<close> [symmetric] |
33189 | 578 |
have x: "x \<sharp> (\<lambda>x':T'. t')" by (simp add: abs_fresh) |
579 |
have x': "x' \<sharp> (\<lambda>x':T'. t')" by (simp add: abs_fresh) |
|
63167 | 580 |
from \<open>(x', T') # \<Gamma> \<turnstile> t' : U\<close> have x'': "x' \<sharp> \<Gamma>" |
33189 | 581 |
by (auto dest: valid_typing) |
582 |
have "(\<lambda>x:T. t) = (\<lambda>x':T'. t')" by fact |
|
583 |
also from x x' y have "\<dots> = [(x, y)] \<bullet> [(x', y)] \<bullet> (\<lambda>x':T'. t')" |
|
584 |
by (simp only: perm_fresh_fresh fresh_prod) |
|
585 |
also have "\<dots> = (\<lambda>x:T'. [(x, y)] \<bullet> [(x', y)] \<bullet> t')" |
|
586 |
by (simp add: swap_simps perm_fresh_fresh) |
|
587 |
finally have "(\<lambda>x:T. t) = (\<lambda>x:T'. [(x, y)] \<bullet> [(x', y)] \<bullet> t')" . |
|
588 |
then have T: "T = T'" and t: "[(x, y)] \<bullet> [(x', y)] \<bullet> t' = t" |
|
589 |
by (simp_all add: trm.inject alpha) |
|
590 |
from Abs T have "S = T \<rightarrow> U" by simp |
|
63167 | 591 |
moreover from \<open>(x', T') # \<Gamma> \<turnstile> t' : U\<close> |
34990 | 592 |
have "[(x, y)] \<bullet> [(x', y)] \<bullet> ((x', T') # \<Gamma> \<turnstile> t' : U)" |
33189 | 593 |
by (simp add: perm_bool) |
34990 | 594 |
with T t y x'' fresh have "(x, T) # \<Gamma> \<turnstile> t : U" |
33189 | 595 |
by (simp add: eqvts swap_simps perm_fresh_fresh fresh_prod) |
596 |
ultimately show ?thesis by (rule R) |
|
597 |
qed simp_all |
|
598 |
||
599 |
nominal_primrec ty_size :: "ty \<Rightarrow> nat" |
|
600 |
where |
|
601 |
"ty_size (Atom n) = 0" |
|
602 |
| "ty_size (T \<rightarrow> U) = ty_size T + ty_size U + 1" |
|
603 |
| "ty_size (T \<otimes> U) = ty_size T + ty_size U + 1" |
|
604 |
by (rule TrueI)+ |
|
605 |
||
606 |
lemma bind_tuple_ineq: |
|
607 |
"ty_size (pat_type p) < ty_size U \<Longrightarrow> Bind U x t \<noteq> (\<lambda>[p]. u)" |
|
608 |
by (induct p arbitrary: U x t u) (auto simp add: btrm.inject) |
|
609 |
||
610 |
lemma valid_appD: assumes "valid (\<Gamma> @ \<Delta>)" |
|
611 |
shows "valid \<Gamma>" "valid \<Delta>" using assms |
|
612 |
by (induct \<Gamma>'\<equiv>"\<Gamma> @ \<Delta>" arbitrary: \<Gamma> \<Delta>) |
|
613 |
(auto simp add: Cons_eq_append_conv fresh_list_append) |
|
614 |
||
615 |
lemma valid_app_freshs: assumes "valid (\<Gamma> @ \<Delta>)" |
|
616 |
shows "(supp \<Gamma>::name set) \<sharp>* \<Delta>" "(supp \<Delta>::name set) \<sharp>* \<Gamma>" using assms |
|
617 |
by (induct \<Gamma>'\<equiv>"\<Gamma> @ \<Delta>" arbitrary: \<Gamma> \<Delta>) |
|
618 |
(auto simp add: Cons_eq_append_conv fresh_star_def |
|
619 |
fresh_list_nil fresh_list_cons supp_list_nil supp_list_cons fresh_list_append |
|
620 |
supp_prod fresh_prod supp_atm fresh_atm |
|
44687 | 621 |
dest: notE [OF iffD1 [OF fresh_def]]) |
33189 | 622 |
|
623 |
lemma perm_mem_left: "(x::name) \<in> ((pi::name prm) \<bullet> A) \<Longrightarrow> (rev pi \<bullet> x) \<in> A" |
|
624 |
by (drule perm_boolI [of _ "rev pi"]) (simp add: eqvts perm_pi_simp) |
|
625 |
||
626 |
lemma perm_mem_right: "(rev (pi::name prm) \<bullet> (x::name)) \<in> A \<Longrightarrow> x \<in> (pi \<bullet> A)" |
|
627 |
by (drule perm_boolI [of _ pi]) (simp add: eqvts perm_pi_simp) |
|
628 |
||
629 |
lemma perm_cases: |
|
630 |
assumes pi: "set pi \<subseteq> A \<times> A" |
|
631 |
shows "((pi::name prm) \<bullet> B) \<subseteq> A \<union> B" |
|
632 |
proof |
|
633 |
fix x assume "x \<in> pi \<bullet> B" |
|
634 |
then show "x \<in> A \<union> B" using pi |
|
80140 | 635 |
proof (induct pi arbitrary: x B rule: rev_induct) |
636 |
case Nil |
|
637 |
then show ?case |
|
638 |
by simp |
|
639 |
next |
|
640 |
case (snoc y xs) |
|
641 |
then show ?case |
|
642 |
apply simp |
|
643 |
by (metis SigmaE perm_mem_left perm_pi_simp(2) pt_name2 swap_simps(3)) |
|
644 |
qed |
|
33189 | 645 |
qed |
646 |
||
647 |
lemma abs_pat_alpha': |
|
648 |
assumes eq: "(\<lambda>[p]. t) = (\<lambda>[q]. u)" |
|
649 |
and ty: "pat_type p = pat_type q" |
|
650 |
and pv: "distinct (pat_vars p)" |
|
651 |
and qv: "distinct (pat_vars q)" |
|
652 |
shows "\<exists>pi::name prm. p = pi \<bullet> q \<and> t = pi \<bullet> u \<and> |
|
653 |
set pi \<subseteq> (supp p \<union> supp q) \<times> (supp p \<union> supp q)" |
|
654 |
using assms |
|
45129
1fce03e3e8ad
tuned proofs -- eliminated vacuous "induct arbitrary: ..." situations;
wenzelm
parents:
44687
diff
changeset
|
655 |
proof (induct p arbitrary: q t u) |
33189 | 656 |
case (PVar x T) |
657 |
note PVar' = this |
|
658 |
show ?case |
|
659 |
proof (cases q) |
|
660 |
case (PVar x' T') |
|
63167 | 661 |
with \<open>(\<lambda>[PVar x T]. t) = (\<lambda>[q]. u)\<close> |
33189 | 662 |
have "x = x' \<and> t = u \<or> x \<noteq> x' \<and> t = [(x, x')] \<bullet> u \<and> x \<sharp> u" |
663 |
by (simp add: btrm.inject alpha) |
|
664 |
then show ?thesis |
|
665 |
proof |
|
666 |
assume "x = x' \<and> t = u" |
|
667 |
with PVar PVar' have "PVar x T = ([]::name prm) \<bullet> q \<and> |
|
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33189
diff
changeset
|
668 |
t = ([]::name prm) \<bullet> u \<and> |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33189
diff
changeset
|
669 |
set ([]::name prm) \<subseteq> (supp (PVar x T) \<union> supp q) \<times> |
33189 | 670 |
(supp (PVar x T) \<union> supp q)" by simp |
671 |
then show ?thesis .. |
|
672 |
next |
|
673 |
assume "x \<noteq> x' \<and> t = [(x, x')] \<bullet> u \<and> x \<sharp> u" |
|
674 |
with PVar PVar' have "PVar x T = [(x, x')] \<bullet> q \<and> |
|
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33189
diff
changeset
|
675 |
t = [(x, x')] \<bullet> u \<and> |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33189
diff
changeset
|
676 |
set [(x, x')] \<subseteq> (supp (PVar x T) \<union> supp q) \<times> |
33189 | 677 |
(supp (PVar x T) \<union> supp q)" |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33189
diff
changeset
|
678 |
by (simp add: perm_swap swap_simps supp_atm perm_type) |
33189 | 679 |
then show ?thesis .. |
680 |
qed |
|
681 |
next |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
682 |
case (PTuple p\<^sub>1 p\<^sub>2) |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
683 |
with PVar have "ty_size (pat_type p\<^sub>1) < ty_size T" by simp |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
684 |
then have "Bind T x t \<noteq> (\<lambda>[p\<^sub>1]. \<lambda>[p\<^sub>2]. u)" |
33189 | 685 |
by (rule bind_tuple_ineq) |
686 |
moreover from PTuple PVar |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
687 |
have "Bind T x t = (\<lambda>[p\<^sub>1]. \<lambda>[p\<^sub>2]. u)" by simp |
33189 | 688 |
ultimately show ?thesis .. |
689 |
qed |
|
690 |
next |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
691 |
case (PTuple p\<^sub>1 p\<^sub>2) |
33189 | 692 |
note PTuple' = this |
693 |
show ?case |
|
694 |
proof (cases q) |
|
695 |
case (PVar x T) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
696 |
with PTuple have "ty_size (pat_type p\<^sub>1) < ty_size T" by auto |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
697 |
then have "Bind T x u \<noteq> (\<lambda>[p\<^sub>1]. \<lambda>[p\<^sub>2]. t)" |
33189 | 698 |
by (rule bind_tuple_ineq) |
699 |
moreover from PTuple PVar |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
700 |
have "Bind T x u = (\<lambda>[p\<^sub>1]. \<lambda>[p\<^sub>2]. t)" by simp |
33189 | 701 |
ultimately show ?thesis .. |
702 |
next |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
703 |
case (PTuple p\<^sub>1' p\<^sub>2') |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
704 |
with PTuple' have "(\<lambda>[p\<^sub>1]. \<lambda>[p\<^sub>2]. t) = (\<lambda>[p\<^sub>1']. \<lambda>[p\<^sub>2']. u)" by simp |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
705 |
moreover from PTuple PTuple' have "pat_type p\<^sub>1 = pat_type p\<^sub>1'" |
33189 | 706 |
by (simp add: ty.inject) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
707 |
moreover from PTuple' have "distinct (pat_vars p\<^sub>1)" by simp |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
708 |
moreover from PTuple PTuple' have "distinct (pat_vars p\<^sub>1')" by simp |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
709 |
ultimately have "\<exists>pi::name prm. p\<^sub>1 = pi \<bullet> p\<^sub>1' \<and> |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
710 |
(\<lambda>[p\<^sub>2]. t) = pi \<bullet> (\<lambda>[p\<^sub>2']. u) \<and> |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
711 |
set pi \<subseteq> (supp p\<^sub>1 \<union> supp p\<^sub>1') \<times> (supp p\<^sub>1 \<union> supp p\<^sub>1')" |
33189 | 712 |
by (rule PTuple') |
713 |
then obtain pi::"name prm" where |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
714 |
"p\<^sub>1 = pi \<bullet> p\<^sub>1'" "(\<lambda>[p\<^sub>2]. t) = pi \<bullet> (\<lambda>[p\<^sub>2']. u)" and |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
715 |
pi: "set pi \<subseteq> (supp p\<^sub>1 \<union> supp p\<^sub>1') \<times> (supp p\<^sub>1 \<union> supp p\<^sub>1')" by auto |
63167 | 716 |
from \<open>(\<lambda>[p\<^sub>2]. t) = pi \<bullet> (\<lambda>[p\<^sub>2']. u)\<close> |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
717 |
have "(\<lambda>[p\<^sub>2]. t) = (\<lambda>[pi \<bullet> p\<^sub>2']. pi \<bullet> u)" |
33189 | 718 |
by (simp add: eqvts) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
719 |
moreover from PTuple PTuple' have "pat_type p\<^sub>2 = pat_type (pi \<bullet> p\<^sub>2')" |
33189 | 720 |
by (simp add: ty.inject pat_type_perm_eq) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
721 |
moreover from PTuple' have "distinct (pat_vars p\<^sub>2)" by simp |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
722 |
moreover from PTuple PTuple' have "distinct (pat_vars (pi \<bullet> p\<^sub>2'))" |
33189 | 723 |
by (simp add: pat_vars_eqvt [symmetric] distinct_eqvt [symmetric]) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
724 |
ultimately have "\<exists>pi'::name prm. p\<^sub>2 = pi' \<bullet> pi \<bullet> p\<^sub>2' \<and> |
33189 | 725 |
t = pi' \<bullet> pi \<bullet> u \<and> |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
726 |
set pi' \<subseteq> (supp p\<^sub>2 \<union> supp (pi \<bullet> p\<^sub>2')) \<times> (supp p\<^sub>2 \<union> supp (pi \<bullet> p\<^sub>2'))" |
33189 | 727 |
by (rule PTuple') |
728 |
then obtain pi'::"name prm" where |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
729 |
"p\<^sub>2 = pi' \<bullet> pi \<bullet> p\<^sub>2'" "t = pi' \<bullet> pi \<bullet> u" and |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
730 |
pi': "set pi' \<subseteq> (supp p\<^sub>2 \<union> supp (pi \<bullet> p\<^sub>2')) \<times> |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
731 |
(supp p\<^sub>2 \<union> supp (pi \<bullet> p\<^sub>2'))" by auto |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
732 |
from PTuple PTuple' have "pi \<bullet> distinct (pat_vars \<langle>\<langle>p\<^sub>1', p\<^sub>2'\<rangle>\<rangle>)" by simp |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
733 |
then have "distinct (pat_vars \<langle>\<langle>pi \<bullet> p\<^sub>1', pi \<bullet> p\<^sub>2'\<rangle>\<rangle>)" by (simp only: eqvts) |
63167 | 734 |
with \<open>p\<^sub>1 = pi \<bullet> p\<^sub>1'\<close> PTuple' |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
735 |
have fresh: "(supp p\<^sub>2 \<union> supp (pi \<bullet> p\<^sub>2') :: name set) \<sharp>* p\<^sub>1" |
33189 | 736 |
by (auto simp add: set_pat_vars_supp fresh_star_def fresh_def eqvts) |
63167 | 737 |
from \<open>p\<^sub>1 = pi \<bullet> p\<^sub>1'\<close> have "pi' \<bullet> (p\<^sub>1 = pi \<bullet> p\<^sub>1')" by (rule perm_boolI) |
33189 | 738 |
with pt_freshs_freshs [OF pt_name_inst at_name_inst pi' fresh fresh] |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
739 |
have "p\<^sub>1 = pi' \<bullet> pi \<bullet> p\<^sub>1'" by (simp add: eqvts) |
63167 | 740 |
with \<open>p\<^sub>2 = pi' \<bullet> pi \<bullet> p\<^sub>2'\<close> have "\<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> = (pi' @ pi) \<bullet> \<langle>\<langle>p\<^sub>1', p\<^sub>2'\<rangle>\<rangle>" |
33189 | 741 |
by (simp add: pt_name2) |
742 |
moreover |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
743 |
have "((supp p\<^sub>2 \<union> (pi \<bullet> supp p\<^sub>2')) \<times> (supp p\<^sub>2 \<union> (pi \<bullet> supp p\<^sub>2'))::(name \<times> name) set) \<subseteq> |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
744 |
(supp p\<^sub>2 \<union> (supp p\<^sub>1 \<union> supp p\<^sub>1' \<union> supp p\<^sub>2')) \<times> (supp p\<^sub>2 \<union> (supp p\<^sub>1 \<union> supp p\<^sub>1' \<union> supp p\<^sub>2'))" |
33189 | 745 |
by (rule subset_refl Sigma_mono Un_mono perm_cases [OF pi])+ |
746 |
with pi' have "set pi' \<subseteq> \<dots>" by (simp add: supp_eqvt [symmetric]) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
747 |
with pi have "set (pi' @ pi) \<subseteq> (supp \<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> \<union> supp \<langle>\<langle>p\<^sub>1', p\<^sub>2'\<rangle>\<rangle>) \<times> |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
748 |
(supp \<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> \<union> supp \<langle>\<langle>p\<^sub>1', p\<^sub>2'\<rangle>\<rangle>)" |
33189 | 749 |
by (simp add: Sigma_Un_distrib1 Sigma_Un_distrib2 Un_ac) blast |
63167 | 750 |
moreover note \<open>t = pi' \<bullet> pi \<bullet> u\<close> |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
751 |
ultimately have "\<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> = (pi' @ pi) \<bullet> q \<and> t = (pi' @ pi) \<bullet> u \<and> |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
752 |
set (pi' @ pi) \<subseteq> (supp \<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> \<union> supp q) \<times> |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
753 |
(supp \<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> \<union> supp q)" using PTuple |
33189 | 754 |
by (simp add: pt_name2) |
755 |
then show ?thesis .. |
|
756 |
qed |
|
757 |
qed |
|
758 |
||
759 |
lemma typing_case_Let: |
|
760 |
assumes ty: "\<Gamma> \<turnstile> (LET p = t IN u) : U" |
|
761 |
and fresh: "(supp p::name set) \<sharp>* \<Gamma>" |
|
762 |
and distinct: "distinct (pat_vars p)" |
|
763 |
and R: "\<And>T \<Delta>. \<Gamma> \<turnstile> t : T \<Longrightarrow> \<turnstile> p : T \<Rightarrow> \<Delta> \<Longrightarrow> \<Delta> @ \<Gamma> \<turnstile> u : U \<Longrightarrow> P" |
|
764 |
shows P using ty |
|
765 |
proof cases |
|
34990 | 766 |
case (Let p' t' T \<Delta> u') |
33189 | 767 |
then have "(supp \<Delta>::name set) \<sharp>* \<Gamma>" |
768 |
by (auto intro: valid_typing valid_app_freshs) |
|
769 |
with Let have "(supp p'::name set) \<sharp>* \<Gamma>" |
|
770 |
by (simp add: pat_var) |
|
771 |
with fresh have fresh': "(supp p \<union> supp p' :: name set) \<sharp>* \<Gamma>" |
|
772 |
by (auto simp add: fresh_star_def) |
|
773 |
from Let have "(\<lambda>[p]. Base u) = (\<lambda>[p']. Base u')" |
|
774 |
by (simp add: trm.inject) |
|
775 |
moreover from Let have "pat_type p = pat_type p'" |
|
776 |
by (simp add: trm.inject) |
|
777 |
moreover note distinct |
|
63167 | 778 |
moreover from \<open>\<Delta> @ \<Gamma> \<turnstile> u' : U\<close> have "valid (\<Delta> @ \<Gamma>)" |
33189 | 779 |
by (rule valid_typing) |
780 |
then have "valid \<Delta>" by (rule valid_appD) |
|
63167 | 781 |
with \<open>\<turnstile> p' : T \<Rightarrow> \<Delta>\<close> have "distinct (pat_vars p')" |
33189 | 782 |
by (simp add: valid_distinct pat_vars_ptyping) |
783 |
ultimately have "\<exists>pi::name prm. p = pi \<bullet> p' \<and> Base u = pi \<bullet> Base u' \<and> |
|
784 |
set pi \<subseteq> (supp p \<union> supp p') \<times> (supp p \<union> supp p')" |
|
785 |
by (rule abs_pat_alpha') |
|
786 |
then obtain pi::"name prm" where pi: "p = pi \<bullet> p'" "u = pi \<bullet> u'" |
|
787 |
and pi': "set pi \<subseteq> (supp p \<union> supp p') \<times> (supp p \<union> supp p')" |
|
788 |
by (auto simp add: btrm.inject) |
|
789 |
from Let have "\<Gamma> \<turnstile> t : T" by (simp add: trm.inject) |
|
63167 | 790 |
moreover from \<open>\<turnstile> p' : T \<Rightarrow> \<Delta>\<close> have "\<turnstile> (pi \<bullet> p') : (pi \<bullet> T) \<Rightarrow> (pi \<bullet> \<Delta>)" |
33189 | 791 |
by (simp add: ptyping.eqvt) |
792 |
with pi have "\<turnstile> p : T \<Rightarrow> (pi \<bullet> \<Delta>)" by (simp add: perm_type) |
|
793 |
moreover from Let |
|
794 |
have "(pi \<bullet> \<Delta>) @ (pi \<bullet> \<Gamma>) \<turnstile> (pi \<bullet> u') : (pi \<bullet> U)" |
|
795 |
by (simp add: append_eqvt [symmetric] typing.eqvt) |
|
796 |
with pi have "(pi \<bullet> \<Delta>) @ \<Gamma> \<turnstile> u : U" |
|
797 |
by (simp add: perm_type pt_freshs_freshs |
|
798 |
[OF pt_name_inst at_name_inst pi' fresh' fresh']) |
|
799 |
ultimately show ?thesis by (rule R) |
|
800 |
qed simp_all |
|
801 |
||
802 |
lemma preservation: |
|
803 |
assumes "t \<longmapsto> t'" and "\<Gamma> \<turnstile> t : T" |
|
804 |
shows "\<Gamma> \<turnstile> t' : T" using assms |
|
805 |
proof (nominal_induct avoiding: \<Gamma> T rule: eval.strong_induct) |
|
806 |
case (TupleL t t' u) |
|
63167 | 807 |
from \<open>\<Gamma> \<turnstile> \<langle>t, u\<rangle> : T\<close> obtain T\<^sub>1 T\<^sub>2 |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
808 |
where "T = T\<^sub>1 \<otimes> T\<^sub>2" "\<Gamma> \<turnstile> t : T\<^sub>1" "\<Gamma> \<turnstile> u : T\<^sub>2" |
33189 | 809 |
by cases (simp_all add: trm.inject) |
63167 | 810 |
from \<open>\<Gamma> \<turnstile> t : T\<^sub>1\<close> have "\<Gamma> \<turnstile> t' : T\<^sub>1" by (rule TupleL) |
811 |
then have "\<Gamma> \<turnstile> \<langle>t', u\<rangle> : T\<^sub>1 \<otimes> T\<^sub>2" using \<open>\<Gamma> \<turnstile> u : T\<^sub>2\<close> |
|
33189 | 812 |
by (rule Tuple) |
63167 | 813 |
with \<open>T = T\<^sub>1 \<otimes> T\<^sub>2\<close> show ?case by simp |
33189 | 814 |
next |
815 |
case (TupleR u u' t) |
|
63167 | 816 |
from \<open>\<Gamma> \<turnstile> \<langle>t, u\<rangle> : T\<close> obtain T\<^sub>1 T\<^sub>2 |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
45129
diff
changeset
|
817 |
where "T = T\<^sub>1 \<otimes> T\<^sub>2" "\<Gamma> \<turnstile> t : T\<^sub>1" "\<Gamma> \<turnstile> u : T\<^sub>2" |
33189 | 818 |
by cases (simp_all add: trm.inject) |
63167 | 819 |
from \<open>\<Gamma> \<turnstile> u : T\<^sub>2\<close> have "\<Gamma> \<turnstile> u' : T\<^sub>2" by (rule TupleR) |
820 |
with \<open>\<Gamma> \<turnstile> t : T\<^sub>1\<close> have "\<Gamma> \<turnstile> \<langle>t, u'\<rangle> : T\<^sub>1 \<otimes> T\<^sub>2" |
|
33189 | 821 |
by (rule Tuple) |
63167 | 822 |
with \<open>T = T\<^sub>1 \<otimes> T\<^sub>2\<close> show ?case by simp |
33189 | 823 |
next |
824 |
case (Abs t t' x S) |
|
63167 | 825 |
from \<open>\<Gamma> \<turnstile> (\<lambda>x:S. t) : T\<close> \<open>x \<sharp> \<Gamma>\<close> obtain U where |
33189 | 826 |
T: "T = S \<rightarrow> U" and U: "(x, S) # \<Gamma> \<turnstile> t : U" |
827 |
by (rule typing_case_Abs) |
|
828 |
from U have "(x, S) # \<Gamma> \<turnstile> t' : U" by (rule Abs) |
|
829 |
then have "\<Gamma> \<turnstile> (\<lambda>x:S. t') : S \<rightarrow> U" |
|
830 |
by (rule typing.Abs) |
|
831 |
with T show ?case by simp |
|
832 |
next |
|
833 |
case (Beta x u S t) |
|
63167 | 834 |
from \<open>\<Gamma> \<turnstile> (\<lambda>x:S. t) \<cdot> u : T\<close> \<open>x \<sharp> \<Gamma>\<close> |
33189 | 835 |
obtain "(x, S) # \<Gamma> \<turnstile> t : T" and "\<Gamma> \<turnstile> u : S" |
836 |
by cases (auto simp add: trm.inject ty.inject elim: typing_case_Abs) |
|
837 |
then show ?case by (rule subst_type) |
|
838 |
next |
|
839 |
case (Let p t \<theta> u) |
|
63167 | 840 |
from \<open>\<Gamma> \<turnstile> (LET p = t IN u) : T\<close> \<open>supp p \<sharp>* \<Gamma>\<close> \<open>distinct (pat_vars p)\<close> |
33189 | 841 |
obtain U \<Delta> where "\<turnstile> p : U \<Rightarrow> \<Delta>" "\<Gamma> \<turnstile> t : U" "\<Delta> @ \<Gamma> \<turnstile> u : T" |
842 |
by (rule typing_case_Let) |
|
63167 | 843 |
then show ?case using \<open>\<turnstile> p \<rhd> t \<Rightarrow> \<theta>\<close> \<open>supp p \<sharp>* t\<close> |
33189 | 844 |
by (rule match_type) |
845 |
next |
|
846 |
case (AppL t t' u) |
|
63167 | 847 |
from \<open>\<Gamma> \<turnstile> t \<cdot> u : T\<close> obtain U where |
33189 | 848 |
t: "\<Gamma> \<turnstile> t : U \<rightarrow> T" and u: "\<Gamma> \<turnstile> u : U" |
849 |
by cases (auto simp add: trm.inject) |
|
850 |
from t have "\<Gamma> \<turnstile> t' : U \<rightarrow> T" by (rule AppL) |
|
851 |
then show ?case using u by (rule typing.App) |
|
852 |
next |
|
853 |
case (AppR u u' t) |
|
63167 | 854 |
from \<open>\<Gamma> \<turnstile> t \<cdot> u : T\<close> obtain U where |
33189 | 855 |
t: "\<Gamma> \<turnstile> t : U \<rightarrow> T" and u: "\<Gamma> \<turnstile> u : U" |
856 |
by cases (auto simp add: trm.inject) |
|
857 |
from u have "\<Gamma> \<turnstile> u' : U" by (rule AppR) |
|
858 |
with t show ?case by (rule typing.App) |
|
859 |
qed |
|
860 |
||
861 |
end |