1476
|
1 |
(* Title: HOL/ex/mt.thy
|
969
|
2 |
ID: $Id$
|
1476
|
3 |
Author: Jacob Frost, Cambridge University Computer Laboratory
|
969
|
4 |
Copyright 1993 University of Cambridge
|
|
5 |
|
|
6 |
Based upon the article
|
|
7 |
Robin Milner and Mads Tofte,
|
|
8 |
Co-induction in Relational Semantics,
|
|
9 |
Theoretical Computer Science 87 (1991), pages 209-220.
|
|
10 |
|
|
11 |
Written up as
|
|
12 |
Jacob Frost, A Case Study of Co_induction in Isabelle/HOL
|
|
13 |
Report 308, Computer Lab, University of Cambridge (1993).
|
|
14 |
*)
|
|
15 |
|
17289
|
16 |
theory MT
|
|
17 |
imports Main
|
|
18 |
begin
|
969
|
19 |
|
17289
|
20 |
typedecl Const
|
969
|
21 |
|
17289
|
22 |
typedecl ExVar
|
|
23 |
typedecl Ex
|
969
|
24 |
|
17289
|
25 |
typedecl TyConst
|
|
26 |
typedecl Ty
|
969
|
27 |
|
17289
|
28 |
typedecl Clos
|
|
29 |
typedecl Val
|
969
|
30 |
|
17289
|
31 |
typedecl ValEnv
|
|
32 |
typedecl TyEnv
|
969
|
33 |
|
|
34 |
consts
|
15450
|
35 |
c_app :: "[Const, Const] => Const"
|
969
|
36 |
|
15450
|
37 |
e_const :: "Const => Ex"
|
|
38 |
e_var :: "ExVar => Ex"
|
|
39 |
e_fn :: "[ExVar, Ex] => Ex" ("fn _ => _" [0,51] 1000)
|
|
40 |
e_fix :: "[ExVar, ExVar, Ex] => Ex" ("fix _ ( _ ) = _" [0,51,51] 1000)
|
17289
|
41 |
e_app :: "[Ex, Ex] => Ex" ("_ @@ _" [51,51] 1000)
|
15450
|
42 |
e_const_fst :: "Ex => Const"
|
969
|
43 |
|
15450
|
44 |
t_const :: "TyConst => Ty"
|
|
45 |
t_fun :: "[Ty, Ty] => Ty" ("_ -> _" [51,51] 1000)
|
969
|
46 |
|
15450
|
47 |
v_const :: "Const => Val"
|
|
48 |
v_clos :: "Clos => Val"
|
17289
|
49 |
|
1376
|
50 |
ve_emp :: ValEnv
|
15450
|
51 |
ve_owr :: "[ValEnv, ExVar, Val] => ValEnv" ("_ + { _ |-> _ }" [36,0,0] 50)
|
|
52 |
ve_dom :: "ValEnv => ExVar set"
|
|
53 |
ve_app :: "[ValEnv, ExVar] => Val"
|
969
|
54 |
|
15450
|
55 |
clos_mk :: "[ExVar, Ex, ValEnv] => Clos" ("<| _ , _ , _ |>" [0,0,0] 1000)
|
969
|
56 |
|
1376
|
57 |
te_emp :: TyEnv
|
15450
|
58 |
te_owr :: "[TyEnv, ExVar, Ty] => TyEnv" ("_ + { _ |=> _ }" [36,0,0] 50)
|
|
59 |
te_app :: "[TyEnv, ExVar] => Ty"
|
|
60 |
te_dom :: "TyEnv => ExVar set"
|
969
|
61 |
|
|
62 |
eval_fun :: "((ValEnv * Ex) * Val) set => ((ValEnv * Ex) * Val) set"
|
|
63 |
eval_rel :: "((ValEnv * Ex) * Val) set"
|
15450
|
64 |
eval :: "[ValEnv, Ex, Val] => bool" ("_ |- _ ---> _" [36,0,36] 50)
|
969
|
65 |
|
|
66 |
elab_fun :: "((TyEnv * Ex) * Ty) set => ((TyEnv * Ex) * Ty) set"
|
|
67 |
elab_rel :: "((TyEnv * Ex) * Ty) set"
|
15450
|
68 |
elab :: "[TyEnv, Ex, Ty] => bool" ("_ |- _ ===> _" [36,0,36] 50)
|
17289
|
69 |
|
15450
|
70 |
isof :: "[Const, Ty] => bool" ("_ isof _" [36,36] 50)
|
|
71 |
isof_env :: "[ValEnv,TyEnv] => bool" ("_ isofenv _")
|
969
|
72 |
|
|
73 |
hasty_fun :: "(Val * Ty) set => (Val * Ty) set"
|
|
74 |
hasty_rel :: "(Val * Ty) set"
|
15450
|
75 |
hasty :: "[Val, Ty] => bool" ("_ hasty _" [36,36] 50)
|
|
76 |
hasty_env :: "[ValEnv,TyEnv] => bool" ("_ hastyenv _ " [36,36] 35)
|
969
|
77 |
|
17289
|
78 |
axioms
|
969
|
79 |
|
17289
|
80 |
(*
|
969
|
81 |
Expression constructors must be injective, distinct and it must be possible
|
|
82 |
to do induction over expressions.
|
|
83 |
*)
|
|
84 |
|
|
85 |
(* All the constructors are injective *)
|
|
86 |
|
17289
|
87 |
e_const_inj: "e_const(c1) = e_const(c2) ==> c1 = c2"
|
|
88 |
e_var_inj: "e_var(ev1) = e_var(ev2) ==> ev1 = ev2"
|
|
89 |
e_fn_inj: "fn ev1 => e1 = fn ev2 => e2 ==> ev1 = ev2 & e1 = e2"
|
|
90 |
e_fix_inj:
|
|
91 |
" fix ev11e(v12) = e1 = fix ev21(ev22) = e2 ==>
|
|
92 |
ev11 = ev21 & ev12 = ev22 & e1 = e2
|
1151
|
93 |
"
|
17289
|
94 |
e_app_inj: "e11 @@ e12 = e21 @@ e22 ==> e11 = e21 & e12 = e22"
|
969
|
95 |
|
|
96 |
(* All constructors are distinct *)
|
|
97 |
|
17289
|
98 |
e_disj_const_var: "~e_const(c) = e_var(ev)"
|
|
99 |
e_disj_const_fn: "~e_const(c) = fn ev => e"
|
|
100 |
e_disj_const_fix: "~e_const(c) = fix ev1(ev2) = e"
|
|
101 |
e_disj_const_app: "~e_const(c) = e1 @@ e2"
|
|
102 |
e_disj_var_fn: "~e_var(ev1) = fn ev2 => e"
|
|
103 |
e_disj_var_fix: "~e_var(ev) = fix ev1(ev2) = e"
|
|
104 |
e_disj_var_app: "~e_var(ev) = e1 @@ e2"
|
|
105 |
e_disj_fn_fix: "~fn ev1 => e1 = fix ev21(ev22) = e2"
|
|
106 |
e_disj_fn_app: "~fn ev1 => e1 = e21 @@ e22"
|
|
107 |
e_disj_fix_app: "~fix ev11(ev12) = e1 = e21 @@ e22"
|
969
|
108 |
|
|
109 |
(* Strong elimination, induction on expressions *)
|
|
110 |
|
17289
|
111 |
e_ind:
|
|
112 |
" [| !!ev. P(e_var(ev));
|
|
113 |
!!c. P(e_const(c));
|
|
114 |
!!ev e. P(e) ==> P(fn ev => e);
|
|
115 |
!!ev1 ev2 e. P(e) ==> P(fix ev1(ev2) = e);
|
|
116 |
!!e1 e2. P(e1) ==> P(e2) ==> P(e1 @@ e2)
|
|
117 |
|] ==>
|
|
118 |
P(e)
|
1151
|
119 |
"
|
969
|
120 |
|
|
121 |
(* Types - same scheme as for expressions *)
|
|
122 |
|
17289
|
123 |
(* All constructors are injective *)
|
969
|
124 |
|
17289
|
125 |
t_const_inj: "t_const(c1) = t_const(c2) ==> c1 = c2"
|
|
126 |
t_fun_inj: "t11 -> t12 = t21 -> t22 ==> t11 = t21 & t12 = t22"
|
969
|
127 |
|
|
128 |
(* All constructors are distinct, not needed so far ... *)
|
|
129 |
|
|
130 |
(* Strong elimination, induction on types *)
|
|
131 |
|
17289
|
132 |
t_ind:
|
|
133 |
"[| !!p. P(t_const p); !!t1 t2. P(t1) ==> P(t2) ==> P(t_fun t1 t2) |]
|
1151
|
134 |
==> P(t)"
|
969
|
135 |
|
|
136 |
|
|
137 |
(* Values - same scheme again *)
|
|
138 |
|
17289
|
139 |
(* All constructors are injective *)
|
969
|
140 |
|
17289
|
141 |
v_const_inj: "v_const(c1) = v_const(c2) ==> c1 = c2"
|
|
142 |
v_clos_inj:
|
|
143 |
" v_clos(<|ev1,e1,ve1|>) = v_clos(<|ev2,e2,ve2|>) ==>
|
1151
|
144 |
ev1 = ev2 & e1 = e2 & ve1 = ve2"
|
17289
|
145 |
|
969
|
146 |
(* All constructors are distinct *)
|
|
147 |
|
17289
|
148 |
v_disj_const_clos: "~v_const(c) = v_clos(cl)"
|
969
|
149 |
|
15450
|
150 |
(* No induction on values: they are a codatatype! ... *)
|
969
|
151 |
|
|
152 |
|
17289
|
153 |
(*
|
969
|
154 |
Value environments bind variables to values. Only the following trivial
|
|
155 |
properties are needed.
|
|
156 |
*)
|
|
157 |
|
17289
|
158 |
ve_dom_owr: "ve_dom(ve + {ev |-> v}) = ve_dom(ve) Un {ev}"
|
|
159 |
|
|
160 |
ve_app_owr1: "ve_app (ve + {ev |-> v}) ev=v"
|
|
161 |
ve_app_owr2: "~ev1=ev2 ==> ve_app (ve+{ev1 |-> v}) ev2=ve_app ve ev2"
|
969
|
162 |
|
|
163 |
|
|
164 |
(* Type Environments bind variables to types. The following trivial
|
|
165 |
properties are needed. *)
|
|
166 |
|
17289
|
167 |
te_dom_owr: "te_dom(te + {ev |=> t}) = te_dom(te) Un {ev}"
|
|
168 |
|
|
169 |
te_app_owr1: "te_app (te + {ev |=> t}) ev=t"
|
|
170 |
te_app_owr2: "~ev1=ev2 ==> te_app (te+{ev1 |=> t}) ev2=te_app te ev2"
|
969
|
171 |
|
|
172 |
|
|
173 |
(* The dynamic semantics is defined inductively by a set of inference
|
|
174 |
rules. These inference rules allows one to draw conclusions of the form ve
|
|
175 |
|- e ---> v, read the expression e evaluates to the value v in the value
|
|
176 |
environment ve. Therefore the relation _ |- _ ---> _ is defined in Isabelle
|
|
177 |
as the least fixpoint of the functor eval_fun below. From this definition
|
|
178 |
introduction rules and a strong elimination (induction) rule can be
|
17289
|
179 |
derived.
|
969
|
180 |
*)
|
|
181 |
|
17289
|
182 |
defs
|
|
183 |
eval_fun_def:
|
|
184 |
" eval_fun(s) ==
|
|
185 |
{ pp.
|
|
186 |
(? ve c. pp=((ve,e_const(c)),v_const(c))) |
|
1151
|
187 |
(? ve x. pp=((ve,e_var(x)),ve_app ve x) & x:ve_dom(ve)) |
|
17289
|
188 |
(? ve e x. pp=((ve,fn x => e),v_clos(<|x,e,ve|>)))|
|
|
189 |
( ? ve e x f cl.
|
|
190 |
pp=((ve,fix f(x) = e),v_clos(cl)) &
|
|
191 |
cl=<|x, e, ve+{f |-> v_clos(cl)} |>
|
|
192 |
) |
|
|
193 |
( ? ve e1 e2 c1 c2.
|
|
194 |
pp=((ve,e1 @@ e2),v_const(c_app c1 c2)) &
|
|
195 |
((ve,e1),v_const(c1)):s & ((ve,e2),v_const(c2)):s
|
|
196 |
) |
|
|
197 |
( ? ve vem e1 e2 em xm v v2.
|
|
198 |
pp=((ve,e1 @@ e2),v) &
|
|
199 |
((ve,e1),v_clos(<|xm,em,vem|>)):s &
|
|
200 |
((ve,e2),v2):s &
|
|
201 |
((vem+{xm |-> v2},em),v):s
|
|
202 |
)
|
1151
|
203 |
}"
|
969
|
204 |
|
17289
|
205 |
eval_rel_def: "eval_rel == lfp(eval_fun)"
|
|
206 |
eval_def: "ve |- e ---> v == ((ve,e),v):eval_rel"
|
969
|
207 |
|
|
208 |
(* The static semantics is defined in the same way as the dynamic
|
|
209 |
semantics. The relation te |- e ===> t express the expression e has the
|
|
210 |
type t in the type environment te.
|
|
211 |
*)
|
|
212 |
|
17289
|
213 |
elab_fun_def:
|
|
214 |
"elab_fun(s) ==
|
|
215 |
{ pp.
|
|
216 |
(? te c t. pp=((te,e_const(c)),t) & c isof t) |
|
|
217 |
(? te x. pp=((te,e_var(x)),te_app te x) & x:te_dom(te)) |
|
|
218 |
(? te x e t1 t2. pp=((te,fn x => e),t1->t2) & ((te+{x |=> t1},e),t2):s) |
|
|
219 |
(? te f x e t1 t2.
|
|
220 |
pp=((te,fix f(x)=e),t1->t2) & ((te+{f |=> t1->t2}+{x |=> t1},e),t2):s
|
|
221 |
) |
|
|
222 |
(? te e1 e2 t1 t2.
|
|
223 |
pp=((te,e1 @@ e2),t2) & ((te,e1),t1->t2):s & ((te,e2),t1):s
|
|
224 |
)
|
1151
|
225 |
}"
|
969
|
226 |
|
17289
|
227 |
elab_rel_def: "elab_rel == lfp(elab_fun)"
|
|
228 |
elab_def: "te |- e ===> t == ((te,e),t):elab_rel"
|
969
|
229 |
|
|
230 |
(* The original correspondence relation *)
|
|
231 |
|
17289
|
232 |
isof_env_def:
|
|
233 |
" ve isofenv te ==
|
|
234 |
ve_dom(ve) = te_dom(te) &
|
|
235 |
( ! x.
|
|
236 |
x:ve_dom(ve) -->
|
|
237 |
(? c. ve_app ve x = v_const(c) & c isof te_app te x)
|
|
238 |
)
|
1151
|
239 |
"
|
969
|
240 |
|
17289
|
241 |
axioms
|
|
242 |
isof_app: "[| c1 isof t1->t2; c2 isof t1 |] ==> c_app c1 c2 isof t2"
|
969
|
243 |
|
17289
|
244 |
defs
|
969
|
245 |
(* The extented correspondence relation *)
|
|
246 |
|
17289
|
247 |
hasty_fun_def:
|
|
248 |
" hasty_fun(r) ==
|
|
249 |
{ p.
|
|
250 |
( ? c t. p = (v_const(c),t) & c isof t) |
|
|
251 |
( ? ev e ve t te.
|
|
252 |
p = (v_clos(<|ev,e,ve|>),t) &
|
|
253 |
te |- fn ev => e ===> t &
|
|
254 |
ve_dom(ve) = te_dom(te) &
|
|
255 |
(! ev1. ev1:ve_dom(ve) --> (ve_app ve ev1,te_app te ev1) : r)
|
|
256 |
)
|
|
257 |
}
|
1151
|
258 |
"
|
969
|
259 |
|
17289
|
260 |
hasty_rel_def: "hasty_rel == gfp(hasty_fun)"
|
|
261 |
hasty_def: "v hasty t == (v,t) : hasty_rel"
|
|
262 |
hasty_env_def:
|
|
263 |
" ve hastyenv te ==
|
|
264 |
ve_dom(ve) = te_dom(te) &
|
1151
|
265 |
(! x. x: ve_dom(ve) --> ve_app ve x hasty te_app te x)"
|
969
|
266 |
|
17289
|
267 |
ML {* use_legacy_bindings (the_context ()) *}
|
|
268 |
|
969
|
269 |
end
|