author | nipkow |
Sun, 05 Jan 2003 21:03:14 +0100 | |
changeset 13771 | 6cd59cc885a1 |
parent 13339 | 0f89104dd377 |
child 16417 | 9bc16273c2d4 |
permissions | -rw-r--r-- |
1478 | 1 |
(* Title: ZF/Coind/Types.thy |
916 | 2 |
ID: $Id$ |
1478 | 3 |
Author: Jacob Frost, Cambridge University Computer Laboratory |
916 | 4 |
Copyright 1995 University of Cambridge |
5 |
*) |
|
6 |
||
12595 | 7 |
theory Types = Language: |
916 | 8 |
|
9 |
consts |
|
12595 | 10 |
Ty :: i (* Datatype of types *) |
1478 | 11 |
TyConst :: i (* Abstract type of type constants *) |
6112 | 12 |
|
13 |
datatype |
|
12595 | 14 |
"Ty" = t_const ("tc \<in> TyConst") |
15 |
| t_fun ("t1 \<in> Ty","t2 \<in> Ty") |
|
916 | 16 |
|
17 |
||
18 |
(* Definition of type environments and associated operators *) |
|
19 |
||
20 |
consts |
|
1401 | 21 |
TyEnv :: i |
6112 | 22 |
|
23 |
datatype |
|
934 | 24 |
"TyEnv" = te_emp |
12595 | 25 |
| te_owr ("te \<in> TyEnv","x \<in> ExVar","t \<in> Ty") |
916 | 26 |
|
27 |
consts |
|
12595 | 28 |
te_dom :: "i => i" |
29 |
te_app :: "[i,i] => i" |
|
6046 | 30 |
|
31 |
||
32 |
primrec (*domain of the type environment*) |
|
33 |
"te_dom (te_emp) = 0" |
|
12595 | 34 |
"te_dom (te_owr(te,x,v)) = te_dom(te) Un {x}" |
6046 | 35 |
|
36 |
primrec (*lookup up identifiers in the type environment*) |
|
37 |
"te_app (te_emp,x) = 0" |
|
6068 | 38 |
"te_app (te_owr(te,y,t),x) = (if x=y then t else te_app(te,x))" |
916 | 39 |
|
12595 | 40 |
inductive_cases te_owrE [elim!]: "te_owr(te,f,t) \<in> TyEnv" |
41 |
||
42 |
(*redundant??*) |
|
43 |
lemma te_app_owr1: "te_app(te_owr(te,x,t),x) = t" |
|
44 |
by simp |
|
45 |
||
46 |
(*redundant??*) |
|
47 |
lemma te_app_owr2: "x \<noteq> y ==> te_app(te_owr(te,x,t),y) = te_app(te,y)" |
|
48 |
by auto |
|
49 |
||
50 |
lemma te_app_owr [simp]: |
|
51 |
"te_app(te_owr(te,x,t),y) = (if x=y then t else te_app(te,y))" |
|
52 |
by auto |
|
53 |
||
54 |
lemma te_appI: |
|
55 |
"[| te \<in> TyEnv; x \<in> ExVar; x \<in> te_dom(te) |] ==> te_app(te,x) \<in> Ty" |
|
56 |
apply (erule_tac P = "x \<in> te_dom (te) " in rev_mp) |
|
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
12595
diff
changeset
|
57 |
apply (erule TyEnv.induct, auto) |
12595 | 58 |
done |
916 | 59 |
|
60 |
||
61 |
||
62 |
||
63 |
||
12595 | 64 |
|
65 |
||
66 |
||
67 |
||
68 |
||
69 |
||
70 |
||
71 |
||
72 |
||
73 |
||
74 |
||
75 |
||
76 |
||
77 |
end |