1478
|
1 |
(* Title: ZF/Coind/Static.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 Static = Values + Types:
|
|
8 |
|
|
9 |
(*** Basic correspondence relation -- not completely specified, as it is a
|
|
10 |
parameter of the proof. A concrete version could be defined inductively.
|
|
11 |
***)
|
|
12 |
|
|
13 |
consts
|
|
14 |
isof :: "[i,i] => o"
|
|
15 |
|
|
16 |
axioms
|
|
17 |
isof_app: "[|isof(c1,t_fun(t1,t2)); isof(c2,t1)|] ==> isof(c_app(c1,c2),t2)"
|
|
18 |
|
|
19 |
(*Its extension to environments*)
|
|
20 |
|
|
21 |
constdefs
|
|
22 |
isofenv :: "[i,i] => o"
|
|
23 |
"isofenv(ve,te) ==
|
|
24 |
ve_dom(ve) = te_dom(te) &
|
|
25 |
(\<forall>x \<in> ve_dom(ve).
|
|
26 |
\<exists>c \<in> Const. ve_app(ve,x) = v_const(c) & isof(c,te_app(te,x)))"
|
|
27 |
|
|
28 |
|
|
29 |
(*** Elaboration ***)
|
916
|
30 |
|
6141
|
31 |
consts ElabRel :: i
|
|
32 |
|
916
|
33 |
inductive
|
|
34 |
domains "ElabRel" <= "TyEnv * Exp * Ty"
|
12595
|
35 |
intros
|
|
36 |
constI [intro!]:
|
|
37 |
"[| te \<in> TyEnv; c \<in> Const; t \<in> Ty; isof(c,t) |] ==>
|
|
38 |
<te,e_const(c),t> \<in> ElabRel"
|
|
39 |
varI [intro!]:
|
|
40 |
"[| te \<in> TyEnv; x \<in> ExVar; x \<in> te_dom(te) |] ==>
|
|
41 |
<te,e_var(x),te_app(te,x)> \<in> ElabRel"
|
|
42 |
fnI [intro!]:
|
|
43 |
"[| te \<in> TyEnv; x \<in> ExVar; e \<in> Exp; t1 \<in> Ty; t2 \<in> Ty;
|
|
44 |
<te_owr(te,x,t1),e,t2> \<in> ElabRel |] ==>
|
|
45 |
<te,e_fn(x,e),t_fun(t1,t2)> \<in> ElabRel"
|
|
46 |
fixI [intro!]:
|
|
47 |
"[| te \<in> TyEnv; f \<in> ExVar; x \<in> ExVar; t1 \<in> Ty; t2 \<in> Ty;
|
|
48 |
<te_owr(te_owr(te,f,t_fun(t1,t2)),x,t1),e,t2> \<in> ElabRel |] ==>
|
|
49 |
<te,e_fix(f,x,e),t_fun(t1,t2)> \<in> ElabRel"
|
|
50 |
appI [intro]:
|
|
51 |
"[| te \<in> TyEnv; e1 \<in> Exp; e2 \<in> Exp; t1 \<in> Ty; t2 \<in> Ty;
|
|
52 |
<te,e1,t_fun(t1,t2)> \<in> ElabRel;
|
|
53 |
<te,e2,t1> \<in> ElabRel |] ==> <te,e_app(e1,e2),t2> \<in> ElabRel"
|
|
54 |
type_intros te_appI Exp.intros Ty.intros
|
|
55 |
|
|
56 |
|
12610
|
57 |
inductive_cases
|
|
58 |
elab_constE [elim!]: "<te,e_const(c),t> \<in> ElabRel"
|
|
59 |
and elab_varE [elim!]: "<te,e_var(x),t> \<in> ElabRel"
|
|
60 |
and elab_fnE [elim]: "<te,e_fn(x,e),t> \<in> ElabRel"
|
|
61 |
and elab_fixE [elim!]: "<te,e_fix(f,x,e),t> \<in> ElabRel"
|
|
62 |
and elab_appE [elim]: "<te,e_app(e1,e2),t> \<in> ElabRel"
|
12595
|
63 |
|
|
64 |
declare ElabRel.dom_subset [THEN subsetD, dest]
|
916
|
65 |
|
|
66 |
end
|
|
67 |
|
|
68 |
|
|
69 |
|
|
70 |
|
|
71 |
|
|
72 |
|
|
73 |
|
|
74 |
|
|
75 |
|
|
76 |
|
|
77 |
|
|
78 |
|
|
79 |
|
|
80 |
|
|
81 |
|
|
82 |
|
|
83 |
|
|
84 |
|
|
85 |
|
|
86 |
|
|
87 |
|
|
88 |
|