author | kleing |
Tue, 02 Mar 2004 01:32:23 +0100 | |
changeset 14422 | b8da5f258b04 |
parent 13537 | f506eb568121 |
permissions | -rw-r--r-- |
1300 | 1 |
(* Title: HOL/MiniML/Type.thy |
2 |
ID: $Id$ |
|
2525 | 3 |
Author: Wolfgang Naraschewski and Tobias Nipkow |
4 |
Copyright 1996 TU Muenchen |
|
1300 | 5 |
|
6 |
MiniML-types and type substitutions. |
|
7 |
*) |
|
8 |
||
14422 | 9 |
theory Type = Maybe: |
1300 | 10 |
|
11 |
(* new class for structures containing type variables *) |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
5184
diff
changeset
|
12 |
axclass type_struct < type |
2525 | 13 |
|
1300 | 14 |
|
15 |
(* type expressions *) |
|
14422 | 16 |
datatype "typ" = TVar nat | "->" "typ" "typ" (infixr 70) |
1300 | 17 |
|
2525 | 18 |
(* type schemata *) |
14422 | 19 |
datatype type_scheme = FVar nat | BVar nat | "=->" type_scheme type_scheme (infixr 70) |
2525 | 20 |
|
21 |
||
22 |
(* embedding types into type schemata *) |
|
23 |
consts |
|
14422 | 24 |
mk_scheme :: "typ => type_scheme" |
2525 | 25 |
|
5184 | 26 |
primrec |
2525 | 27 |
"mk_scheme (TVar n) = (FVar n)" |
28 |
"mk_scheme (t1 -> t2) = ((mk_scheme t1) =-> (mk_scheme t2))" |
|
29 |
||
30 |
||
14422 | 31 |
instance "typ"::type_struct .. |
32 |
instance type_scheme::type_struct .. |
|
33 |
instance list::(type_struct)type_struct .. |
|
34 |
instance fun::(type,type_struct)type_struct .. |
|
2525 | 35 |
|
36 |
||
37 |
(* free_tv s: the type variables occuring freely in the type structure s *) |
|
38 |
||
39 |
consts |
|
14422 | 40 |
free_tv :: "['a::type_struct] => nat set" |
2525 | 41 |
|
13537 | 42 |
primrec (free_tv_typ) |
14422 | 43 |
free_tv_TVar: "free_tv (TVar m) = {m}" |
44 |
free_tv_Fun: "free_tv (t1 -> t2) = (free_tv t1) Un (free_tv t2)" |
|
2525 | 45 |
|
13537 | 46 |
primrec (free_tv_type_scheme) |
2525 | 47 |
"free_tv (FVar m) = {m}" |
48 |
"free_tv (BVar m) = {}" |
|
49 |
"free_tv (S1 =-> S2) = (free_tv S1) Un (free_tv S2)" |
|
50 |
||
13537 | 51 |
primrec (free_tv_list) |
2525 | 52 |
"free_tv [] = {}" |
53 |
"free_tv (x#l) = (free_tv x) Un (free_tv l)" |
|
54 |
||
2625 | 55 |
|
2525 | 56 |
(* executable version of free_tv: Implementation with lists *) |
57 |
consts |
|
14422 | 58 |
free_tv_ML :: "['a::type_struct] => nat list" |
2525 | 59 |
|
13537 | 60 |
primrec (free_tv_ML_type_scheme) |
2525 | 61 |
"free_tv_ML (FVar m) = [m]" |
62 |
"free_tv_ML (BVar m) = []" |
|
63 |
"free_tv_ML (S1 =-> S2) = (free_tv_ML S1) @ (free_tv_ML S2)" |
|
64 |
||
13537 | 65 |
primrec (free_tv_ML_list) |
2525 | 66 |
"free_tv_ML [] = []" |
67 |
"free_tv_ML (x#l) = (free_tv_ML x) @ (free_tv_ML l)" |
|
68 |
||
69 |
||
70 |
(* new_tv s n computes whether n is a new type variable w.r.t. a type |
|
71 |
structure s, i.e. whether n is greater than any type variable |
|
72 |
occuring in the type structure *) |
|
73 |
constdefs |
|
14422 | 74 |
new_tv :: "[nat,'a::type_struct] => bool" |
2525 | 75 |
"new_tv n ts == ! m. m:(free_tv ts) --> m<n" |
76 |
||
77 |
||
78 |
(* bound_tv s: the type variables occuring bound in the type structure s *) |
|
79 |
||
80 |
consts |
|
14422 | 81 |
bound_tv :: "['a::type_struct] => nat set" |
2525 | 82 |
|
13537 | 83 |
primrec (bound_tv_type_scheme) |
2525 | 84 |
"bound_tv (FVar m) = {}" |
85 |
"bound_tv (BVar m) = {m}" |
|
86 |
"bound_tv (S1 =-> S2) = (bound_tv S1) Un (bound_tv S2)" |
|
87 |
||
13537 | 88 |
primrec (bound_tv_list) |
2525 | 89 |
"bound_tv [] = {}" |
90 |
"bound_tv (x#l) = (bound_tv x) Un (bound_tv l)" |
|
91 |
||
92 |
||
93 |
(* minimal new free / bound variable *) |
|
94 |
||
95 |
consts |
|
14422 | 96 |
min_new_bound_tv :: "'a::type_struct => nat" |
2525 | 97 |
|
13537 | 98 |
primrec (min_new_bound_tv_type_scheme) |
2525 | 99 |
"min_new_bound_tv (FVar n) = 0" |
100 |
"min_new_bound_tv (BVar n) = Suc n" |
|
101 |
"min_new_bound_tv (sch1 =-> sch2) = max (min_new_bound_tv sch1) (min_new_bound_tv sch2)" |
|
102 |
||
103 |
||
104 |
(* substitutions *) |
|
105 |
||
106 |
(* type variable substitution *) |
|
1300 | 107 |
types |
14422 | 108 |
subst = "nat => typ" |
1300 | 109 |
|
110 |
(* identity *) |
|
1557 | 111 |
constdefs |
1476 | 112 |
id_subst :: subst |
3842 | 113 |
"id_subst == (%n. TVar n)" |
1300 | 114 |
|
115 |
(* extension of substitution to type structures *) |
|
116 |
consts |
|
14422 | 117 |
app_subst :: "[subst, 'a::type_struct] => 'a::type_struct" ("$") |
1300 | 118 |
|
13537 | 119 |
primrec (app_subst_typ) |
14422 | 120 |
app_subst_TVar: "$ S (TVar n) = S n" |
121 |
app_subst_Fun: "$ S (t1 -> t2) = ($ S t1) -> ($ S t2)" |
|
2525 | 122 |
|
13537 | 123 |
primrec (app_subst_type_scheme) |
2525 | 124 |
"$ S (FVar n) = mk_scheme (S n)" |
125 |
"$ S (BVar n) = (BVar n)" |
|
126 |
"$ S (sch1 =-> sch2) = ($ S sch1) =-> ($ S sch2)" |
|
1604 | 127 |
|
1300 | 128 |
defs |
14422 | 129 |
app_subst_list: "$ S == map ($ S)" |
1300 | 130 |
|
131 |
(* domain of a substitution *) |
|
1557 | 132 |
constdefs |
14422 | 133 |
dom :: "subst => nat set" |
2525 | 134 |
"dom S == {n. S n ~= TVar n}" |
1300 | 135 |
|
2525 | 136 |
(* codomain of a substitution: the introduced variables *) |
137 |
||
1557 | 138 |
constdefs |
14422 | 139 |
cod :: "subst => nat set" |
2525 | 140 |
"cod S == (UN m:dom S. (free_tv (S m)))" |
1300 | 141 |
|
142 |
defs |
|
14422 | 143 |
free_tv_subst: "free_tv S == (dom S) Un (cod S)" |
1300 | 144 |
|
2525 | 145 |
|
1300 | 146 |
(* unification algorithm mgu *) |
147 |
consts |
|
14422 | 148 |
mgu :: "[typ,typ] => subst option" |
149 |
axioms |
|
150 |
mgu_eq: "mgu t1 t2 = Some U ==> $U t1 = $U t2" |
|
151 |
mgu_mg: "[| (mgu t1 t2) = Some U; $S t1 = $S t2 |] ==> ? R. S = $R o U" |
|
152 |
mgu_Some: "$S t1 = $S t2 ==> ? U. mgu t1 t2 = Some U" |
|
153 |
mgu_free: "mgu t1 t2 = Some U ==> (free_tv U) <= (free_tv t1) Un (free_tv t2)" |
|
154 |
||
155 |
||
156 |
declare mgu_eq [simp] mgu_mg [simp] mgu_free [simp] |
|
157 |
||
158 |
||
159 |
(* lemmata for make scheme *) |
|
160 |
||
161 |
lemma mk_scheme_Fun [rule_format (no_asm)]: "mk_scheme t = sch1 =-> sch2 --> (? t1 t2. sch1 = mk_scheme t1 & sch2 = mk_scheme t2)" |
|
162 |
apply (induct_tac "t") |
|
163 |
apply (simp (no_asm)) |
|
164 |
apply simp |
|
165 |
apply fast |
|
166 |
done |
|
167 |
||
168 |
lemma mk_scheme_injective [rule_format (no_asm)]: "!t'. mk_scheme t = mk_scheme t' --> t=t'" |
|
169 |
apply (induct_tac "t") |
|
170 |
apply (rule allI) |
|
171 |
apply (induct_tac "t'") |
|
172 |
apply (simp (no_asm)) |
|
173 |
apply simp |
|
174 |
apply (rule allI) |
|
175 |
apply (induct_tac "t'") |
|
176 |
apply (simp (no_asm)) |
|
177 |
apply simp |
|
178 |
done |
|
179 |
||
180 |
lemma free_tv_mk_scheme: "free_tv (mk_scheme t) = free_tv t" |
|
181 |
apply (induct_tac "t") |
|
182 |
apply (simp_all (no_asm_simp)) |
|
183 |
done |
|
184 |
||
185 |
declare free_tv_mk_scheme [simp] |
|
186 |
||
187 |
lemma subst_mk_scheme: "$ S (mk_scheme t) = mk_scheme ($ S t)" |
|
188 |
apply (induct_tac "t") |
|
189 |
apply (simp_all (no_asm_simp)) |
|
190 |
done |
|
191 |
||
192 |
declare subst_mk_scheme [simp] |
|
193 |
||
194 |
||
195 |
(* constructor laws for app_subst *) |
|
196 |
||
197 |
lemma app_subst_Nil: |
|
198 |
"$ S [] = []" |
|
199 |
||
200 |
apply (unfold app_subst_list) |
|
201 |
apply (simp (no_asm)) |
|
202 |
done |
|
203 |
||
204 |
lemma app_subst_Cons: |
|
205 |
"$ S (x#l) = ($ S x)#($ S l)" |
|
206 |
apply (unfold app_subst_list) |
|
207 |
apply (simp (no_asm)) |
|
208 |
done |
|
209 |
||
210 |
declare app_subst_Nil [simp] app_subst_Cons [simp] |
|
211 |
||
212 |
||
213 |
(* constructor laws for new_tv *) |
|
214 |
||
215 |
lemma new_tv_TVar: |
|
216 |
"new_tv n (TVar m) = (m<n)" |
|
217 |
||
218 |
apply (unfold new_tv_def) |
|
219 |
apply (fastsimp) |
|
220 |
done |
|
221 |
||
222 |
lemma new_tv_FVar: |
|
223 |
"new_tv n (FVar m) = (m<n)" |
|
224 |
apply (unfold new_tv_def) |
|
225 |
apply (fastsimp) |
|
226 |
done |
|
227 |
||
228 |
lemma new_tv_BVar: |
|
229 |
"new_tv n (BVar m) = True" |
|
230 |
apply (unfold new_tv_def) |
|
231 |
apply (simp (no_asm)) |
|
232 |
done |
|
233 |
||
234 |
lemma new_tv_Fun: |
|
235 |
"new_tv n (t1 -> t2) = (new_tv n t1 & new_tv n t2)" |
|
236 |
apply (unfold new_tv_def) |
|
237 |
apply (fastsimp) |
|
238 |
done |
|
239 |
||
240 |
lemma new_tv_Fun2: |
|
241 |
"new_tv n (t1 =-> t2) = (new_tv n t1 & new_tv n t2)" |
|
242 |
apply (unfold new_tv_def) |
|
243 |
apply (fastsimp) |
|
244 |
done |
|
245 |
||
246 |
lemma new_tv_Nil: |
|
247 |
"new_tv n []" |
|
248 |
apply (unfold new_tv_def) |
|
249 |
apply (simp (no_asm)) |
|
250 |
done |
|
251 |
||
252 |
lemma new_tv_Cons: |
|
253 |
"new_tv n (x#l) = (new_tv n x & new_tv n l)" |
|
254 |
apply (unfold new_tv_def) |
|
255 |
apply (fastsimp) |
|
256 |
done |
|
257 |
||
258 |
lemma new_tv_TVar_subst: "new_tv n TVar" |
|
259 |
apply (unfold new_tv_def) |
|
260 |
apply (intro strip) |
|
261 |
apply (simp add: free_tv_subst dom_def cod_def) |
|
262 |
done |
|
263 |
||
264 |
declare new_tv_TVar [simp] new_tv_FVar [simp] new_tv_BVar [simp] new_tv_Fun [simp] new_tv_Fun2 [simp] new_tv_Nil [simp] new_tv_Cons [simp] new_tv_TVar_subst [simp] |
|
265 |
||
266 |
lemma new_tv_id_subst: |
|
267 |
"new_tv n id_subst" |
|
268 |
apply (unfold id_subst_def new_tv_def free_tv_subst dom_def cod_def) |
|
269 |
apply (simp (no_asm)) |
|
270 |
done |
|
271 |
declare new_tv_id_subst [simp] |
|
272 |
||
273 |
lemma new_if_subst_type_scheme: "new_tv n (sch::type_scheme) --> |
|
274 |
$(%k. if k<n then S k else S' k) sch = $S sch" |
|
275 |
apply (induct_tac "sch") |
|
276 |
apply (simp_all (no_asm_simp)) |
|
277 |
done |
|
278 |
declare new_if_subst_type_scheme [simp] |
|
279 |
||
280 |
lemma new_if_subst_type_scheme_list: "new_tv n (A::type_scheme list) --> |
|
281 |
$(%k. if k<n then S k else S' k) A = $S A" |
|
282 |
apply (induct_tac "A") |
|
283 |
apply (simp_all (no_asm_simp)) |
|
284 |
done |
|
285 |
declare new_if_subst_type_scheme_list [simp] |
|
286 |
||
287 |
||
288 |
(* constructor laws for dom and cod *) |
|
289 |
||
290 |
lemma dom_id_subst: |
|
291 |
"dom id_subst = {}" |
|
292 |
||
293 |
apply (unfold dom_def id_subst_def empty_def) |
|
294 |
apply (simp (no_asm)) |
|
295 |
done |
|
296 |
declare dom_id_subst [simp] |
|
297 |
||
298 |
lemma cod_id_subst: |
|
299 |
"cod id_subst = {}" |
|
300 |
apply (unfold cod_def) |
|
301 |
apply (simp (no_asm)) |
|
302 |
done |
|
303 |
declare cod_id_subst [simp] |
|
304 |
||
305 |
||
306 |
(* lemmata for free_tv *) |
|
307 |
||
308 |
lemma free_tv_id_subst: |
|
309 |
"free_tv id_subst = {}" |
|
310 |
||
311 |
apply (unfold free_tv_subst) |
|
312 |
apply (simp (no_asm)) |
|
313 |
done |
|
314 |
declare free_tv_id_subst [simp] |
|
315 |
||
316 |
lemma free_tv_nth_A_impl_free_tv_A [rule_format (no_asm)]: "!n. n < length A --> x : free_tv (A!n) --> x : free_tv A" |
|
317 |
apply (induct_tac "A") |
|
318 |
apply simp |
|
319 |
apply (rule allI) |
|
320 |
apply (induct_tac "n" rule: nat_induct) |
|
321 |
apply simp |
|
322 |
apply simp |
|
323 |
done |
|
324 |
||
325 |
declare free_tv_nth_A_impl_free_tv_A [simp] |
|
326 |
||
327 |
lemma free_tv_nth_nat_A [rule_format (no_asm)]: "!nat. nat < length A --> x : free_tv (A!nat) --> x : free_tv A" |
|
328 |
apply (induct_tac "A") |
|
329 |
apply simp |
|
330 |
apply (rule allI) |
|
331 |
apply (induct_tac "nat" rule: nat_induct) |
|
332 |
apply (intro strip) |
|
333 |
apply simp |
|
334 |
apply (simp (no_asm)) |
|
335 |
done |
|
336 |
||
337 |
||
338 |
(* if two substitutions yield the same result if applied to a type |
|
339 |
structure the substitutions coincide on the free type variables |
|
340 |
occurring in the type structure *) |
|
341 |
||
342 |
lemma typ_substitutions_only_on_free_variables [rule_format (no_asm)]: "(!x:free_tv t. (S x) = (S' x)) --> $ S (t::typ) = $ S' t" |
|
343 |
apply (induct_tac "t") |
|
344 |
apply (simp (no_asm)) |
|
345 |
apply simp |
|
346 |
done |
|
347 |
||
348 |
lemma eq_free_eq_subst_te: "(!n. n:(free_tv t) --> S1 n = S2 n) ==> $ S1 (t::typ) = $ S2 t" |
|
349 |
apply (rule typ_substitutions_only_on_free_variables) |
|
350 |
apply (simp (no_asm) add: Ball_def) |
|
351 |
done |
|
352 |
||
353 |
lemma scheme_substitutions_only_on_free_variables [rule_format (no_asm)]: "(!x:free_tv sch. (S x) = (S' x)) --> $ S (sch::type_scheme) = $ S' sch" |
|
354 |
apply (induct_tac "sch") |
|
355 |
apply (simp (no_asm)) |
|
356 |
apply (simp (no_asm)) |
|
357 |
apply simp |
|
358 |
done |
|
359 |
||
360 |
lemma eq_free_eq_subst_type_scheme: |
|
361 |
"(!n. n:(free_tv sch) --> S1 n = S2 n) ==> $ S1 (sch::type_scheme) = $ S2 sch" |
|
362 |
apply (rule scheme_substitutions_only_on_free_variables) |
|
363 |
apply (simp (no_asm) add: Ball_def) |
|
364 |
done |
|
365 |
||
366 |
lemma eq_free_eq_subst_scheme_list [rule_format (no_asm)]: |
|
367 |
"(!n. n:(free_tv A) --> S1 n = S2 n) --> $S1 (A::type_scheme list) = $S2 A" |
|
368 |
apply (induct_tac "A") |
|
369 |
(* case [] *) |
|
370 |
apply (fastsimp) |
|
371 |
(* case x#xl *) |
|
372 |
apply (fastsimp intro: eq_free_eq_subst_type_scheme) |
|
373 |
done |
|
374 |
||
375 |
lemma weaken_asm_Un: "((!x:A. (P x)) --> Q) ==> ((!x:A Un B. (P x)) --> Q)" |
|
376 |
apply fast |
|
377 |
done |
|
378 |
||
379 |
lemma scheme_list_substitutions_only_on_free_variables [rule_format (no_asm)]: "(!x:free_tv A. (S x) = (S' x)) --> $ S (A::type_scheme list) = $ S' A" |
|
380 |
apply (induct_tac "A") |
|
381 |
apply (simp (no_asm)) |
|
382 |
apply simp |
|
383 |
apply (rule weaken_asm_Un) |
|
384 |
apply (intro strip) |
|
385 |
apply (erule scheme_substitutions_only_on_free_variables) |
|
386 |
done |
|
387 |
||
388 |
lemma eq_subst_te_eq_free [rule_format (no_asm)]: |
|
389 |
"$ S1 (t::typ) = $ S2 t --> n:(free_tv t) --> S1 n = S2 n" |
|
390 |
apply (induct_tac "t") |
|
391 |
(* case TVar n *) |
|
392 |
apply (fastsimp) |
|
393 |
(* case Fun t1 t2 *) |
|
394 |
apply (fastsimp) |
|
395 |
done |
|
396 |
||
397 |
lemma eq_subst_type_scheme_eq_free [rule_format (no_asm)]: |
|
398 |
"$ S1 (sch::type_scheme) = $ S2 sch --> n:(free_tv sch) --> S1 n = S2 n" |
|
399 |
apply (induct_tac "sch") |
|
400 |
(* case TVar n *) |
|
401 |
apply (simp (no_asm_simp)) |
|
402 |
(* case BVar n *) |
|
403 |
apply (intro strip) |
|
404 |
apply (erule mk_scheme_injective) |
|
405 |
apply (simp (no_asm_simp)) |
|
406 |
(* case Fun t1 t2 *) |
|
407 |
apply simp |
|
408 |
done |
|
409 |
||
410 |
lemma eq_subst_scheme_list_eq_free [rule_format (no_asm)]: |
|
411 |
"$S1 (A::type_scheme list) = $S2 A --> n:(free_tv A) --> S1 n = S2 n" |
|
412 |
apply (induct_tac "A") |
|
413 |
(* case [] *) |
|
414 |
apply (fastsimp) |
|
415 |
(* case x#xl *) |
|
416 |
apply (fastsimp intro: eq_subst_type_scheme_eq_free) |
|
417 |
done |
|
418 |
||
419 |
lemma codD: |
|
420 |
"v : cod S ==> v : free_tv S" |
|
421 |
apply (unfold free_tv_subst) |
|
422 |
apply (fast) |
|
423 |
done |
|
424 |
||
425 |
lemma not_free_impl_id: |
|
426 |
"x ~: free_tv S ==> S x = TVar x" |
|
427 |
||
428 |
apply (unfold free_tv_subst dom_def) |
|
429 |
apply (fast) |
|
430 |
done |
|
431 |
||
432 |
lemma free_tv_le_new_tv: |
|
433 |
"[| new_tv n t; m:free_tv t |] ==> m<n" |
|
434 |
apply (unfold new_tv_def) |
|
435 |
apply (fast) |
|
436 |
done |
|
437 |
||
438 |
lemma cod_app_subst: |
|
439 |
"[| v : free_tv(S n); v~=n |] ==> v : cod S" |
|
440 |
apply (unfold dom_def cod_def UNION_def Bex_def) |
|
441 |
apply (simp (no_asm)) |
|
442 |
apply (safe intro!: exI conjI notI) |
|
443 |
prefer 2 apply (assumption) |
|
444 |
apply simp |
|
445 |
done |
|
446 |
declare cod_app_subst [simp] |
|
447 |
||
448 |
lemma free_tv_subst_var: "free_tv (S (v::nat)) <= insert v (cod S)" |
|
449 |
apply (case_tac "v:dom S") |
|
450 |
apply (fastsimp simp add: cod_def) |
|
451 |
apply (fastsimp simp add: dom_def) |
|
452 |
done |
|
453 |
||
454 |
lemma free_tv_app_subst_te: "free_tv ($ S (t::typ)) <= cod S Un free_tv t" |
|
455 |
apply (induct_tac "t") |
|
456 |
(* case TVar n *) |
|
457 |
apply (simp (no_asm) add: free_tv_subst_var) |
|
458 |
(* case Fun t1 t2 *) |
|
459 |
apply (fastsimp) |
|
460 |
done |
|
461 |
||
462 |
lemma free_tv_app_subst_type_scheme: "free_tv ($ S (sch::type_scheme)) <= cod S Un free_tv sch" |
|
463 |
apply (induct_tac "sch") |
|
464 |
(* case FVar n *) |
|
465 |
apply (simp (no_asm) add: free_tv_subst_var) |
|
466 |
(* case BVar n *) |
|
467 |
apply (simp (no_asm)) |
|
468 |
(* case Fun t1 t2 *) |
|
469 |
apply (fastsimp) |
|
470 |
done |
|
471 |
||
472 |
lemma free_tv_app_subst_scheme_list: "free_tv ($ S (A::type_scheme list)) <= cod S Un free_tv A" |
|
473 |
apply (induct_tac "A") |
|
474 |
(* case [] *) |
|
475 |
apply (simp (no_asm)) |
|
476 |
(* case a#al *) |
|
477 |
apply (cut_tac free_tv_app_subst_type_scheme) |
|
478 |
apply (fastsimp) |
|
479 |
done |
|
480 |
||
481 |
lemma free_tv_comp_subst: |
|
482 |
"free_tv (%u::nat. $ s1 (s2 u) :: typ) <= |
|
483 |
free_tv s1 Un free_tv s2" |
|
484 |
apply (unfold free_tv_subst dom_def) |
|
485 |
apply (clarsimp simp add: cod_def dom_def) |
|
486 |
apply (drule free_tv_app_subst_te [THEN subsetD]) |
|
487 |
apply clarsimp |
|
488 |
apply (auto simp add: cod_def dom_def) |
|
489 |
apply (drule free_tv_subst_var [THEN subsetD]) |
|
490 |
apply (auto simp add: cod_def dom_def) |
|
491 |
done |
|
492 |
||
493 |
lemma free_tv_o_subst: |
|
494 |
"free_tv ($ S1 o S2) <= free_tv S1 Un free_tv (S2 :: nat => typ)" |
|
495 |
apply (unfold o_def) |
|
496 |
apply (rule free_tv_comp_subst) |
|
497 |
done |
|
498 |
||
499 |
lemma free_tv_of_substitutions_extend_to_types [rule_format (no_asm)]: "n : free_tv t --> free_tv (S n) <= free_tv ($ S t::typ)" |
|
500 |
apply (induct_tac "t") |
|
501 |
apply (simp (no_asm)) |
|
502 |
apply (simp (no_asm)) |
|
503 |
apply fast |
|
504 |
done |
|
505 |
||
506 |
lemma free_tv_of_substitutions_extend_to_schemes [rule_format (no_asm)]: "n : free_tv sch --> free_tv (S n) <= free_tv ($ S sch::type_scheme)" |
|
507 |
apply (induct_tac "sch") |
|
508 |
apply (simp (no_asm)) |
|
509 |
apply (simp (no_asm)) |
|
510 |
apply (simp (no_asm)) |
|
511 |
apply fast |
|
512 |
done |
|
513 |
||
514 |
lemma free_tv_of_substitutions_extend_to_scheme_lists [rule_format (no_asm)]: "n : free_tv A --> free_tv (S n) <= free_tv ($ S A::type_scheme list)" |
|
515 |
apply (induct_tac "A") |
|
516 |
apply (simp (no_asm)) |
|
517 |
apply (simp (no_asm)) |
|
518 |
apply (fast dest: free_tv_of_substitutions_extend_to_schemes) |
|
519 |
done |
|
520 |
||
521 |
declare free_tv_of_substitutions_extend_to_scheme_lists [simp] |
|
522 |
||
523 |
lemma free_tv_ML_scheme: "!!sch::type_scheme. (n : free_tv sch) = (n: set (free_tv_ML sch))" |
|
524 |
apply (induct_tac "sch") |
|
525 |
apply (simp_all (no_asm_simp)) |
|
526 |
done |
|
527 |
||
528 |
lemma free_tv_ML_scheme_list: "!!A::type_scheme list. (n : free_tv A) = (n: set (free_tv_ML A))" |
|
529 |
apply (induct_tac "A") |
|
530 |
apply (simp (no_asm)) |
|
531 |
apply (simp (no_asm_simp) add: free_tv_ML_scheme) |
|
532 |
done |
|
533 |
||
534 |
||
535 |
(* lemmata for bound_tv *) |
|
536 |
||
537 |
lemma bound_tv_mk_scheme: "bound_tv (mk_scheme t) = {}" |
|
538 |
apply (induct_tac "t") |
|
539 |
apply (simp_all (no_asm_simp)) |
|
540 |
done |
|
541 |
||
542 |
declare bound_tv_mk_scheme [simp] |
|
543 |
||
544 |
lemma bound_tv_subst_scheme: "!!sch::type_scheme. bound_tv ($ S sch) = bound_tv sch" |
|
545 |
apply (induct_tac "sch") |
|
546 |
apply (simp_all (no_asm_simp)) |
|
547 |
done |
|
548 |
||
549 |
declare bound_tv_subst_scheme [simp] |
|
550 |
||
551 |
lemma bound_tv_subst_scheme_list: "!!A::type_scheme list. bound_tv ($ S A) = bound_tv A" |
|
552 |
apply (rule list.induct) |
|
553 |
apply (simp (no_asm)) |
|
554 |
apply (simp (no_asm_simp)) |
|
555 |
done |
|
556 |
||
557 |
declare bound_tv_subst_scheme_list [simp] |
|
558 |
||
559 |
||
560 |
(* lemmata for new_tv *) |
|
561 |
||
562 |
lemma new_tv_subst: |
|
563 |
"new_tv n S = ((!m. n <= m --> (S m = TVar m)) & |
|
564 |
(! l. l < n --> new_tv n (S l) ))" |
|
565 |
||
566 |
apply (unfold new_tv_def) |
|
567 |
apply (safe) |
|
568 |
(* ==> *) |
|
569 |
apply (fastsimp dest: leD simp add: free_tv_subst dom_def) |
|
570 |
apply (subgoal_tac "m:cod S | S l = TVar l") |
|
571 |
apply safe |
|
572 |
apply (fastsimp dest: UnI2 simp add: free_tv_subst) |
|
573 |
apply (drule_tac P = "%x. m:free_tv x" in subst , assumption) |
|
574 |
apply simp |
|
575 |
apply (fastsimp simp add: free_tv_subst cod_def dom_def) |
|
576 |
(* <== *) |
|
577 |
apply (unfold free_tv_subst cod_def dom_def) |
|
578 |
apply safe |
|
579 |
apply (cut_tac m = "m" and n = "n" in less_linear) |
|
580 |
apply (fast intro!: less_or_eq_imp_le) |
|
581 |
apply (cut_tac m = "ma" and n = "n" in less_linear) |
|
582 |
apply (fast intro!: less_or_eq_imp_le) |
|
583 |
done |
|
584 |
||
585 |
lemma new_tv_list: "new_tv n x = (!y:set x. new_tv n y)" |
|
586 |
apply (induct_tac "x") |
|
587 |
apply (simp_all (no_asm_simp)) |
|
588 |
done |
|
589 |
||
590 |
(* substitution affects only variables occurring freely *) |
|
591 |
lemma subst_te_new_tv: |
|
592 |
"new_tv n (t::typ) --> $(%x. if x=n then t' else S x) t = $S t" |
|
593 |
apply (induct_tac "t") |
|
594 |
apply (simp_all (no_asm_simp)) |
|
595 |
done |
|
596 |
declare subst_te_new_tv [simp] |
|
597 |
||
598 |
lemma subst_te_new_type_scheme [rule_format (no_asm)]: |
|
599 |
"new_tv n (sch::type_scheme) --> $(%x. if x=n then sch' else S x) sch = $S sch" |
|
600 |
apply (induct_tac "sch") |
|
601 |
apply (simp_all (no_asm_simp)) |
|
602 |
done |
|
603 |
declare subst_te_new_type_scheme [simp] |
|
604 |
||
605 |
lemma subst_tel_new_scheme_list [rule_format (no_asm)]: |
|
606 |
"new_tv n (A::type_scheme list) --> $(%x. if x=n then t else S x) A = $S A" |
|
607 |
apply (induct_tac "A") |
|
608 |
apply simp_all |
|
609 |
done |
|
610 |
declare subst_tel_new_scheme_list [simp] |
|
611 |
||
612 |
(* all greater variables are also new *) |
|
613 |
lemma new_tv_le: |
|
614 |
"n<=m ==> new_tv n t ==> new_tv m t" |
|
615 |
apply (unfold new_tv_def) |
|
616 |
apply safe |
|
617 |
apply (drule spec) |
|
618 |
apply (erule (1) notE impE) |
|
619 |
apply (simp (no_asm)) |
|
620 |
done |
|
621 |
declare lessI [THEN less_imp_le [THEN new_tv_le], simp] |
|
622 |
||
623 |
lemma new_tv_typ_le: "n<=m ==> new_tv n (t::typ) ==> new_tv m t" |
|
624 |
apply (simp (no_asm_simp) add: new_tv_le) |
|
625 |
done |
|
626 |
||
627 |
lemma new_scheme_list_le: "n<=m ==> new_tv n (A::type_scheme list) ==> new_tv m A" |
|
628 |
apply (simp (no_asm_simp) add: new_tv_le) |
|
629 |
done |
|
630 |
||
631 |
lemma new_tv_subst_le: "n<=m ==> new_tv n (S::subst) ==> new_tv m S" |
|
632 |
apply (simp (no_asm_simp) add: new_tv_le) |
|
633 |
done |
|
634 |
||
635 |
(* new_tv property remains if a substitution is applied *) |
|
636 |
lemma new_tv_subst_var: |
|
637 |
"[| n<m; new_tv m (S::subst) |] ==> new_tv m (S n)" |
|
638 |
apply (simp add: new_tv_subst) |
|
639 |
done |
|
640 |
||
641 |
lemma new_tv_subst_te [rule_format (no_asm)]: |
|
642 |
"new_tv n S --> new_tv n (t::typ) --> new_tv n ($ S t)" |
|
643 |
apply (induct_tac "t") |
|
644 |
apply (fastsimp simp add: new_tv_subst)+ |
|
645 |
done |
|
646 |
declare new_tv_subst_te [simp] |
|
647 |
||
648 |
lemma new_tv_subst_type_scheme [rule_format (no_asm)]: "new_tv n S --> new_tv n (sch::type_scheme) --> new_tv n ($ S sch)" |
|
649 |
apply (induct_tac "sch") |
|
650 |
apply (simp_all) |
|
651 |
apply (unfold new_tv_def) |
|
652 |
apply (simp (no_asm) add: free_tv_subst dom_def cod_def) |
|
653 |
apply (intro strip) |
|
654 |
apply (case_tac "S nat = TVar nat") |
|
655 |
apply simp |
|
656 |
apply (drule_tac x = "m" in spec) |
|
657 |
apply fast |
|
658 |
done |
|
659 |
declare new_tv_subst_type_scheme [simp] |
|
660 |
||
661 |
lemma new_tv_subst_scheme_list [rule_format (no_asm)]: |
|
662 |
"new_tv n S --> new_tv n (A::type_scheme list) --> new_tv n ($ S A)" |
|
663 |
apply (induct_tac "A") |
|
664 |
apply (fastsimp)+ |
|
665 |
done |
|
666 |
declare new_tv_subst_scheme_list [simp] |
|
667 |
||
668 |
lemma new_tv_Suc_list: "new_tv n A --> new_tv (Suc n) ((TVar n)#A)" |
|
669 |
apply (simp (no_asm) add: new_tv_list) |
|
670 |
done |
|
671 |
||
672 |
lemma new_tv_only_depends_on_free_tv_type_scheme [rule_format (no_asm)]: "!!sch::type_scheme. (free_tv sch) = (free_tv sch') --> (new_tv n sch --> new_tv n sch')" |
|
673 |
apply (unfold new_tv_def) |
|
674 |
apply (simp (no_asm_simp)) |
|
675 |
done |
|
676 |
||
677 |
lemma new_tv_only_depends_on_free_tv_scheme_list [rule_format (no_asm)]: "!!A::type_scheme list. (free_tv A) = (free_tv A') --> (new_tv n A --> new_tv n A')" |
|
678 |
apply (unfold new_tv_def) |
|
679 |
apply (simp (no_asm_simp)) |
|
680 |
done |
|
681 |
||
682 |
lemma new_tv_nth_nat_A [rule_format (no_asm)]: |
|
683 |
"!nat. nat < length A --> new_tv n A --> (new_tv n (A!nat))" |
|
684 |
apply (unfold new_tv_def) |
|
685 |
apply (induct_tac "A") |
|
686 |
apply simp |
|
687 |
apply (rule allI) |
|
688 |
apply (induct_tac "nat" rule: nat_induct) |
|
689 |
apply (intro strip) |
|
690 |
apply simp |
|
691 |
apply (simp (no_asm)) |
|
692 |
done |
|
693 |
||
694 |
||
695 |
(* composition of substitutions preserves new_tv proposition *) |
|
696 |
lemma new_tv_subst_comp_1: "[| new_tv n (S::subst); new_tv n R |] ==> new_tv n (($ R) o S)" |
|
697 |
apply (simp add: new_tv_subst) |
|
698 |
done |
|
699 |
||
700 |
lemma new_tv_subst_comp_2: "[| new_tv n (S::subst); new_tv n R |] ==> new_tv n (%v.$ R (S v))" |
|
701 |
apply (simp add: new_tv_subst) |
|
702 |
done |
|
703 |
||
704 |
declare new_tv_subst_comp_1 [simp] new_tv_subst_comp_2 [simp] |
|
705 |
||
706 |
(* new type variables do not occur freely in a type structure *) |
|
707 |
lemma new_tv_not_free_tv: |
|
708 |
"new_tv n A ==> n~:(free_tv A)" |
|
709 |
apply (unfold new_tv_def) |
|
710 |
apply (fast elim: less_irrefl) |
|
711 |
done |
|
712 |
declare new_tv_not_free_tv [simp] |
|
713 |
||
714 |
lemma fresh_variable_types: "!!t::typ. ? n. (new_tv n t)" |
|
715 |
apply (unfold new_tv_def) |
|
716 |
apply (induct_tac "t") |
|
717 |
apply (rule_tac x = "Suc nat" in exI) |
|
718 |
apply (simp (no_asm_simp)) |
|
719 |
apply (erule exE)+ |
|
720 |
apply (rename_tac "n'") |
|
721 |
apply (rule_tac x = "max n n'" in exI) |
|
722 |
apply (simp (no_asm) add: less_max_iff_disj) |
|
723 |
apply blast |
|
724 |
done |
|
725 |
||
726 |
declare fresh_variable_types [simp] |
|
727 |
||
728 |
lemma fresh_variable_type_schemes: "!!sch::type_scheme. ? n. (new_tv n sch)" |
|
729 |
apply (unfold new_tv_def) |
|
730 |
apply (induct_tac "sch") |
|
731 |
apply (rule_tac x = "Suc nat" in exI) |
|
732 |
apply (simp (no_asm)) |
|
733 |
apply (rule_tac x = "Suc nat" in exI) |
|
734 |
apply (simp (no_asm)) |
|
735 |
apply (erule exE)+ |
|
736 |
apply (rename_tac "n'") |
|
737 |
apply (rule_tac x = "max n n'" in exI) |
|
738 |
apply (simp (no_asm) add: less_max_iff_disj) |
|
739 |
apply blast |
|
740 |
done |
|
741 |
||
742 |
declare fresh_variable_type_schemes [simp] |
|
743 |
||
744 |
lemma fresh_variable_type_scheme_lists: "!!A::type_scheme list. ? n. (new_tv n A)" |
|
745 |
apply (induct_tac "A") |
|
746 |
apply (simp (no_asm)) |
|
747 |
apply (simp (no_asm)) |
|
748 |
apply (erule exE) |
|
749 |
apply (cut_tac sch = "a" in fresh_variable_type_schemes) |
|
750 |
apply (erule exE) |
|
751 |
apply (rename_tac "n'") |
|
752 |
apply (rule_tac x = " (max n n') " in exI) |
|
753 |
apply (subgoal_tac "n <= (max n n') ") |
|
754 |
apply (subgoal_tac "n' <= (max n n') ") |
|
755 |
apply (fast dest: new_tv_le) |
|
756 |
apply (simp_all (no_asm) add: le_max_iff_disj) |
|
757 |
done |
|
758 |
||
759 |
declare fresh_variable_type_scheme_lists [simp] |
|
760 |
||
761 |
lemma make_one_new_out_of_two: "[| ? n1. (new_tv n1 x); ? n2. (new_tv n2 y)|] ==> ? n. (new_tv n x) & (new_tv n y)" |
|
762 |
apply (erule exE)+ |
|
763 |
apply (rename_tac "n1" "n2") |
|
764 |
apply (rule_tac x = " (max n1 n2) " in exI) |
|
765 |
apply (subgoal_tac "n1 <= max n1 n2") |
|
766 |
apply (subgoal_tac "n2 <= max n1 n2") |
|
767 |
apply (fast dest: new_tv_le) |
|
768 |
apply (simp_all (no_asm) add: le_max_iff_disj) |
|
769 |
done |
|
770 |
||
771 |
lemma ex_fresh_variable: "!!(A::type_scheme list) (A'::type_scheme list) (t::typ) (t'::typ). |
|
772 |
? n. (new_tv n A) & (new_tv n A') & (new_tv n t) & (new_tv n t')" |
|
773 |
apply (cut_tac t = "t" in fresh_variable_types) |
|
774 |
apply (cut_tac t = "t'" in fresh_variable_types) |
|
775 |
apply (drule make_one_new_out_of_two) |
|
776 |
apply assumption |
|
777 |
apply (erule_tac V = "? n. new_tv n t'" in thin_rl) |
|
778 |
apply (cut_tac A = "A" in fresh_variable_type_scheme_lists) |
|
779 |
apply (cut_tac A = "A'" in fresh_variable_type_scheme_lists) |
|
780 |
apply (rotate_tac 1) |
|
781 |
apply (drule make_one_new_out_of_two) |
|
782 |
apply assumption |
|
783 |
apply (erule_tac V = "? n. new_tv n A'" in thin_rl) |
|
784 |
apply (erule exE)+ |
|
785 |
apply (rename_tac n2 n1) |
|
786 |
apply (rule_tac x = " (max n1 n2) " in exI) |
|
787 |
apply (unfold new_tv_def) |
|
788 |
apply (simp (no_asm) add: less_max_iff_disj) |
|
789 |
apply blast |
|
790 |
done |
|
791 |
||
792 |
(* mgu does not introduce new type variables *) |
|
793 |
lemma mgu_new: |
|
794 |
"[|mgu t1 t2 = Some u; new_tv n t1; new_tv n t2|] ==> new_tv n u" |
|
795 |
apply (unfold new_tv_def) |
|
796 |
apply (fast dest: mgu_free) |
|
797 |
done |
|
798 |
||
799 |
||
800 |
(* lemmata for substitutions *) |
|
801 |
||
802 |
lemma length_app_subst_list: |
|
803 |
"!!A:: ('a::type_struct) list. length ($ S A) = length A" |
|
804 |
||
805 |
apply (unfold app_subst_list) |
|
806 |
apply (simp (no_asm)) |
|
807 |
done |
|
808 |
declare length_app_subst_list [simp] |
|
809 |
||
810 |
lemma subst_TVar_scheme: "!!sch::type_scheme. $ TVar sch = sch" |
|
811 |
apply (induct_tac "sch") |
|
812 |
apply (simp_all (no_asm_simp)) |
|
813 |
done |
|
814 |
||
815 |
declare subst_TVar_scheme [simp] |
|
816 |
||
817 |
lemma subst_TVar_scheme_list: "!!A::type_scheme list. $ TVar A = A" |
|
818 |
apply (rule list.induct) |
|
819 |
apply (simp_all add: app_subst_list) |
|
820 |
done |
|
821 |
||
822 |
declare subst_TVar_scheme_list [simp] |
|
823 |
||
824 |
(* application of id_subst does not change type expression *) |
|
825 |
lemma app_subst_id_te: |
|
826 |
"$ id_subst = (%t::typ. t)" |
|
827 |
apply (unfold id_subst_def) |
|
828 |
apply (rule ext) |
|
829 |
apply (induct_tac "t") |
|
830 |
apply (simp_all (no_asm_simp)) |
|
831 |
done |
|
832 |
declare app_subst_id_te [simp] |
|
833 |
||
834 |
lemma app_subst_id_type_scheme: |
|
835 |
"$ id_subst = (%sch::type_scheme. sch)" |
|
836 |
apply (unfold id_subst_def) |
|
837 |
apply (rule ext) |
|
838 |
apply (induct_tac "sch") |
|
839 |
apply (simp_all (no_asm_simp)) |
|
840 |
done |
|
841 |
declare app_subst_id_type_scheme [simp] |
|
842 |
||
843 |
(* application of id_subst does not change list of type expressions *) |
|
844 |
lemma app_subst_id_tel: |
|
845 |
"$ id_subst = (%A::type_scheme list. A)" |
|
846 |
apply (unfold app_subst_list) |
|
847 |
apply (rule ext) |
|
848 |
apply (induct_tac "A") |
|
849 |
apply (simp_all (no_asm_simp)) |
|
850 |
done |
|
851 |
declare app_subst_id_tel [simp] |
|
852 |
||
853 |
lemma id_subst_sch: "!!sch::type_scheme. $ id_subst sch = sch" |
|
854 |
apply (induct_tac "sch") |
|
855 |
apply (simp_all add: id_subst_def) |
|
856 |
done |
|
857 |
||
858 |
declare id_subst_sch [simp] |
|
859 |
||
860 |
lemma id_subst_A: "!!A::type_scheme list. $ id_subst A = A" |
|
861 |
apply (induct_tac "A") |
|
862 |
apply simp |
|
863 |
apply simp |
|
864 |
done |
|
865 |
||
866 |
declare id_subst_A [simp] |
|
867 |
||
868 |
(* composition of substitutions *) |
|
869 |
lemma o_id_subst: "$S o id_subst = S" |
|
870 |
apply (unfold id_subst_def o_def) |
|
871 |
apply (rule ext) |
|
872 |
apply (simp (no_asm)) |
|
873 |
done |
|
874 |
declare o_id_subst [simp] |
|
875 |
||
876 |
lemma subst_comp_te: "$ R ($ S t::typ) = $ (%x. $ R (S x) ) t" |
|
877 |
apply (induct_tac "t") |
|
878 |
apply (simp_all (no_asm_simp)) |
|
879 |
done |
|
880 |
||
881 |
lemma subst_comp_type_scheme: "$ R ($ S sch::type_scheme) = $ (%x. $ R (S x) ) sch" |
|
882 |
apply (induct_tac "sch") |
|
883 |
apply simp_all |
|
884 |
done |
|
885 |
||
886 |
lemma subst_comp_scheme_list: |
|
887 |
"$ R ($ S A::type_scheme list) = $ (%x. $ R (S x)) A" |
|
888 |
apply (unfold app_subst_list) |
|
889 |
apply (induct_tac "A") |
|
890 |
(* case [] *) |
|
891 |
apply (simp (no_asm)) |
|
892 |
(* case x#xl *) |
|
893 |
apply (simp add: subst_comp_type_scheme) |
|
894 |
done |
|
895 |
||
896 |
lemma subst_id_on_type_scheme_list': "!!A::type_scheme list. !x : free_tv A. S x = (TVar x) ==> $ S A = $ id_subst A" |
|
897 |
apply (rule scheme_list_substitutions_only_on_free_variables) |
|
898 |
apply (simp add: id_subst_def) |
|
899 |
done |
|
900 |
||
901 |
lemma subst_id_on_type_scheme_list: "!!A::type_scheme list. !x : free_tv A. S x = (TVar x) ==> $ S A = A" |
|
902 |
apply (subst subst_id_on_type_scheme_list') |
|
903 |
apply assumption |
|
904 |
apply (simp (no_asm)) |
|
905 |
done |
|
906 |
||
907 |
lemma nth_subst [rule_format (no_asm)]: "!n. n < length A --> ($ S A)!n = $S (A!n)" |
|
908 |
apply (induct_tac "A") |
|
909 |
apply simp |
|
910 |
apply (rule allI) |
|
911 |
apply (rename_tac "n1") |
|
912 |
apply (induct_tac "n1" rule: nat_induct) |
|
913 |
apply simp |
|
914 |
apply simp |
|
915 |
done |
|
916 |
||
1300 | 917 |
|
918 |
end |