916
|
1 |
(* Title: ZF/Coind/Types.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Jacob Frost, Cambridge University Computer Laboratory
|
|
4 |
Copyright 1995 University of Cambridge
|
|
5 |
*)
|
|
6 |
|
|
7 |
Types = Language +
|
|
8 |
|
|
9 |
consts
|
934
|
10 |
Ty :: "i" (* Datatype of types *)
|
|
11 |
TyConst :: "i" (* Abstract type of type constants *)
|
|
12 |
datatype <= "univ(TyConst)"
|
|
13 |
"Ty" = t_const("tc:TyConst")
|
|
14 |
| t_fun("t1:Ty","t2:Ty")
|
916
|
15 |
|
|
16 |
|
|
17 |
(* Definition of type environments and associated operators *)
|
|
18 |
|
|
19 |
consts
|
|
20 |
TyEnv :: "i"
|
934
|
21 |
datatype <= "univ(Ty Un ExVar)"
|
|
22 |
"TyEnv" = te_emp
|
|
23 |
| te_owr("te:TyEnv","x:ExVar","t:Ty")
|
916
|
24 |
|
|
25 |
consts
|
|
26 |
te_rec :: "[i,i,[i,i,i,i]=>i] => i"
|
934
|
27 |
defs
|
916
|
28 |
te_rec_def
|
|
29 |
"te_rec(te,c,h) == \
|
|
30 |
\ Vrec(te,%te g.TyEnv_case(c,%tem x t.h(tem,x,t,g`tem),te))"
|
|
31 |
|
|
32 |
consts
|
|
33 |
te_dom :: "i => i"
|
|
34 |
te_app :: "[i,i] => i"
|
934
|
35 |
defs
|
916
|
36 |
te_dom_def "te_dom(te) == te_rec(te,0,% te x t r.r Un {x})"
|
|
37 |
te_app_def "te_app(te,x) == te_rec(te,0, % te y t r.if(x=y,t,r))"
|
|
38 |
|
|
39 |
|
|
40 |
end
|
|
41 |
|
|
42 |
|
|
43 |
|
|
44 |
|
|
45 |
|