|
1 theory Test |
|
2 imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" |
|
3 begin |
|
4 |
|
5 atom_decl name |
|
6 |
|
7 |
|
8 section{* Interface for nominal_datatype *} |
|
9 |
|
10 text {* |
|
11 |
|
12 Nominal-Datatype-part: |
|
13 |
|
14 1st Arg: string list |
|
15 ^^^^^^^^^^^ |
|
16 strings of the types to be defined |
|
17 |
|
18 2nd Arg: (string list * binding * mixfix * (binding * typ list * mixfix) list) list |
|
19 ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^ ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^ |
|
20 type(s) to be defined constructors list |
|
21 (ty args, name, syn) (name, typs, syn) |
|
22 |
|
23 Binder-Function-part: |
|
24 |
|
25 3rd Arg: (binding * typ option * mixfix) list |
|
26 ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^ |
|
27 binding function(s) |
|
28 to be defined |
|
29 (name, type, syn) |
|
30 |
|
31 4th Arg: term list |
|
32 ^^^^^^^^^ |
|
33 the equations of the binding functions |
|
34 (Trueprop equations) |
|
35 *} |
|
36 |
|
37 text {*****************************************************} |
|
38 ML {* |
|
39 (* nominal datatype parser *) |
|
40 local |
|
41 structure P = OuterParse |
|
42 in |
|
43 |
|
44 val _ = OuterKeyword.keyword "bind" |
|
45 val anno_typ = Scan.option (P.name --| P.$$$ "::") -- P.typ |
|
46 |
|
47 (* binding specification *) |
|
48 (* should use and_list *) |
|
49 val bind_parser = |
|
50 P.enum "," ((P.$$$ "bind" |-- P.term) -- (P.$$$ "in" |-- P.name)) |
|
51 |
|
52 val constr_parser = |
|
53 P.binding -- Scan.repeat anno_typ |
|
54 |
|
55 (* datatype parser *) |
|
56 val dt_parser = |
|
57 ((P.type_args -- P.binding -- P.opt_infix) >> P.triple1) -- |
|
58 (P.$$$ "=" |-- P.enum1 "|" ((constr_parser -- bind_parser -- P.opt_mixfix) >> P.triple_swap)) |
|
59 |
|
60 (* function equation parser *) |
|
61 val fun_parser = |
|
62 Scan.optional (P.$$$ "binder" |-- P.fixes -- SpecParse.where_alt_specs) ([],[]) |
|
63 |
|
64 (* main parser *) |
|
65 val main_parser = |
|
66 (P.and_list1 dt_parser) -- fun_parser |
|
67 |
|
68 end |
|
69 *} |
|
70 |
|
71 (* adds "_raw" to the end of constants and types *) |
|
72 ML {* |
|
73 fun add_raw s = s ^ "_raw" |
|
74 fun add_raws ss = map add_raw ss |
|
75 fun raw_bind bn = Binding.suffix_name "_raw" bn |
|
76 |
|
77 fun replace_str ss s = |
|
78 case (AList.lookup (op=) ss s) of |
|
79 SOME s' => s' |
|
80 | NONE => s |
|
81 |
|
82 fun replace_typ ty_ss (Type (a, Ts)) = Type (replace_str ty_ss a, map (replace_typ ty_ss) Ts) |
|
83 | replace_typ ty_ss T = T |
|
84 |
|
85 fun raw_dts ty_ss dts = |
|
86 let |
|
87 val ty_ss' = ty_ss ~~ (add_raws ty_ss) |
|
88 |
|
89 fun raw_dts_aux1 (bind, tys, mx) = |
|
90 (raw_bind bind, map (replace_typ ty_ss') tys, mx) |
|
91 |
|
92 fun raw_dts_aux2 (ty_args, bind, mx, constrs) = |
|
93 (ty_args, raw_bind bind, mx, map raw_dts_aux1 constrs) |
|
94 in |
|
95 map raw_dts_aux2 dts |
|
96 end |
|
97 |
|
98 fun replace_aterm trm_ss (Const (a, T)) = Const (replace_str trm_ss a, T) |
|
99 | replace_aterm trm_ss (Free (a, T)) = Free (replace_str trm_ss a, T) |
|
100 | replace_aterm trm_ss trm = trm |
|
101 |
|
102 fun replace_term trm_ss ty_ss trm = |
|
103 trm |> Term.map_aterms (replace_aterm trm_ss) |> map_types (replace_typ ty_ss) |
|
104 *} |
|
105 |
|
106 ML {* |
|
107 fun get_constrs dts = |
|
108 flat (map (fn (_, _, _, constrs) => constrs) dts) |
|
109 |
|
110 fun get_typed_constrs dts = |
|
111 flat (map (fn (_, bn, _, constrs) => |
|
112 (map (fn (bn', _, _) => (Binding.name_of bn, Binding.name_of bn')) constrs)) dts) |
|
113 |
|
114 fun get_constr_strs dts = |
|
115 map (fn (bn, _, _) => Binding.name_of bn) (get_constrs dts) |
|
116 |
|
117 fun get_bn_fun_strs bn_funs = |
|
118 map (fn (bn_fun, _, _) => Binding.name_of bn_fun) bn_funs |
|
119 *} |
|
120 |
|
121 ML {* |
|
122 fun raw_dts_decl dt_names dts lthy = |
|
123 let |
|
124 val thy = ProofContext.theory_of lthy |
|
125 val conf = Datatype.default_config |
|
126 |
|
127 val dt_names' = add_raws dt_names |
|
128 val dt_full_names = map (Sign.full_bname thy) dt_names |
|
129 val dts' = raw_dts dt_full_names dts |
|
130 in |
|
131 lthy |
|
132 |> Local_Theory.theory_result (Datatype.add_datatype conf dt_names' dts') |
|
133 end |
|
134 *} |
|
135 |
|
136 ML {* |
|
137 fun raw_bn_fun_decl dt_names dts bn_funs bn_eqs lthy = |
|
138 let |
|
139 val thy = ProofContext.theory_of lthy |
|
140 |
|
141 val dt_names' = add_raws dt_names |
|
142 val dt_full_names = map (Sign.full_bname thy) dt_names |
|
143 val dt_full_names' = map (Sign.full_bname thy) dt_names' |
|
144 |
|
145 val ctrs_names = map (Sign.full_bname thy) (get_constr_strs dts) |
|
146 val ctrs_names' = map (fn (x, y) => (Sign.full_bname_path thy (add_raw x) (add_raw y))) |
|
147 (get_typed_constrs dts) |
|
148 |
|
149 val bn_fun_strs = get_bn_fun_strs bn_funs |
|
150 val bn_fun_strs' = add_raws bn_fun_strs |
|
151 |
|
152 val bn_funs' = map (fn (bn, opt_ty, mx) => |
|
153 (raw_bind bn, Option.map (replace_typ (dt_full_names ~~ dt_full_names')) opt_ty, mx)) bn_funs |
|
154 |
|
155 val bn_eqs' = map (fn trm => |
|
156 (Attrib.empty_binding, |
|
157 (replace_term ((ctrs_names ~~ ctrs_names') @ (bn_fun_strs ~~ bn_fun_strs')) (dt_full_names ~~ dt_full_names') trm))) bn_eqs |
|
158 in |
|
159 if null bn_eqs |
|
160 then (([], []), lthy) |
|
161 else Primrec.add_primrec bn_funs' bn_eqs' lthy |
|
162 end |
|
163 *} |
|
164 |
|
165 ML {* |
|
166 fun nominal_datatype2 dts bn_funs bn_eqs lthy = |
|
167 let |
|
168 val dts_names = map (fn (_, s, _, _) => Binding.name_of s) dts |
|
169 in |
|
170 lthy |
|
171 |> raw_dts_decl dts_names dts |
|
172 ||>> raw_bn_fun_decl dts_names dts bn_funs bn_eqs |
|
173 end |
|
174 *} |
|
175 |
|
176 ML {* |
|
177 (* makes a full named type out of a binding with tvars applied to it *) |
|
178 fun mk_type thy bind tvrs = |
|
179 Type (Sign.full_name thy bind, map (fn s => TVar ((s, 0), [])) tvrs) |
|
180 |
|
181 fun get_constrs2 thy dts = |
|
182 let |
|
183 val dts' = map (fn (tvrs, tname, _, constrs) => (mk_type thy tname tvrs, constrs)) dts |
|
184 in |
|
185 flat (map (fn (ty, constrs) => map (fn (bn, tys, mx) => (bn, tys ---> ty, mx)) constrs) dts') |
|
186 end |
|
187 *} |
|
188 |
|
189 ML {* |
|
190 fun nominal_datatype2_cmd (dt_strs, (bn_fun_strs, bn_eq_strs)) lthy = |
|
191 let |
|
192 val thy = ProofContext.theory_of lthy |
|
193 |
|
194 fun prep_typ ((tvs, tname, mx), _) = (tname, length tvs, mx); |
|
195 |
|
196 (* adding the types for parsing datatypes *) |
|
197 val lthy_tmp = lthy |
|
198 |> Local_Theory.theory (Sign.add_types (map prep_typ dt_strs)) |
|
199 |
|
200 fun prep_cnstr lthy (((cname, atys), mx), binders) = |
|
201 (cname, map (Syntax.read_typ lthy o snd) atys, mx) |
|
202 |
|
203 fun prep_dt lthy ((tvs, tname, mx), cnstrs) = |
|
204 (tvs, tname, mx, map (prep_cnstr lthy) cnstrs) |
|
205 |
|
206 (* parsing the datatypes *) |
|
207 val dts_prep = map (prep_dt lthy_tmp) dt_strs |
|
208 |
|
209 (* adding constructors for parsing functions *) |
|
210 val lthy_tmp2 = lthy_tmp |
|
211 |> Local_Theory.theory (Sign.add_consts_i (get_constrs2 thy dts_prep)) |
|
212 |
|
213 val (bn_fun_aux, bn_eq_aux) = fst (Specification.read_spec bn_fun_strs bn_eq_strs lthy_tmp2) |
|
214 |
|
215 fun prep_bn_fun ((bn, T), mx) = (bn, SOME T, mx) |
|
216 |
|
217 fun prep_bn_eq (attr, t) = t |
|
218 |
|
219 val bn_fun_prep = map prep_bn_fun bn_fun_aux |
|
220 val bn_eq_prep = map prep_bn_eq bn_eq_aux |
|
221 in |
|
222 nominal_datatype2 dts_prep bn_fun_prep bn_eq_prep lthy |> snd |
|
223 end |
|
224 *} |
|
225 |
|
226 (* Command Keyword *) |
|
227 ML {* |
|
228 let |
|
229 val kind = OuterKeyword.thy_decl |
|
230 in |
|
231 OuterSyntax.local_theory "nominal_datatype" "test" kind |
|
232 (main_parser >> nominal_datatype2_cmd) |
|
233 end |
|
234 *} |
|
235 |
|
236 text {* example 1 *} |
|
237 |
|
238 nominal_datatype lam = |
|
239 VAR "name" |
|
240 | APP "lam" "lam" |
|
241 | LET bp::"bp" t::"lam" bind "bi bp" in t ("Let _ in _" [100,100] 100) |
|
242 and bp = |
|
243 BP "name" "lam" ("_ ::= _" [100,100] 100) |
|
244 binder |
|
245 bi::"bp \<Rightarrow> name set" |
|
246 where |
|
247 "bi (BP x t) = {x}" |
|
248 |
|
249 typ lam_raw |
|
250 term VAR_raw |
|
251 term Test.BP_raw |
|
252 thm bi_raw.simps |
|
253 |
|
254 print_theorems |
|
255 |
|
256 text {* examples 2 *} |
|
257 nominal_datatype trm = |
|
258 Var "name" |
|
259 | App "trm" "trm" |
|
260 | Lam x::"name" t::"trm" bind x in t |
|
261 | Let p::"pat" "trm" t::"trm" bind "f p" in t |
|
262 and pat = |
|
263 PN |
|
264 | PS "name" |
|
265 | PD "name" "name" |
|
266 binder |
|
267 f::"pat \<Rightarrow> name set" |
|
268 where |
|
269 "f PN = {}" |
|
270 | "f (PS x) = {x}" |
|
271 | "f (PD x y) = {x,y}" |
|
272 |
|
273 thm f_raw.simps |
|
274 |
|
275 nominal_datatype trm0 = |
|
276 Var0 "name" |
|
277 | App0 "trm0" "trm0" |
|
278 | Lam0 x::"name" t::"trm0" bind x in t |
|
279 | Let0 p::"pat0" "trm0" t::"trm0" bind "f0 p" in t |
|
280 and pat0 = |
|
281 PN0 |
|
282 | PS0 "name" |
|
283 | PD0 "pat0" "pat0" |
|
284 binder |
|
285 f0::"pat0 \<Rightarrow> name set" |
|
286 where |
|
287 "f0 PN0 = {}" |
|
288 | "f0 (PS0 x) = {x}" |
|
289 | "f0 (PD0 p1 p2) = (f0 p1) \<union> (f0 p2)" |
|
290 |
|
291 thm f0_raw.simps |
|
292 |
|
293 text {* example type schemes *} |
|
294 datatype ty = |
|
295 Var "name" |
|
296 | Fun "ty" "ty" |
|
297 |
|
298 nominal_datatype tyS = |
|
299 All xs::"name list" ty::"ty" bind xs in ty |
|
300 |
|
301 |
|
302 |
|
303 (* example 1 from Terms.thy *) |
|
304 |
|
305 nominal_datatype trm1 = |
|
306 Vr1 "name" |
|
307 | Ap1 "trm1" "trm1" |
|
308 | Lm1 x::"name" t::"trm1" bind x in t |
|
309 | Lt1 p::"bp1" "trm1" t::"trm1" bind "bv1 p" in t |
|
310 and bp1 = |
|
311 BUnit1 |
|
312 | BV1 "name" |
|
313 | BP1 "bp1" "bp1" |
|
314 binder |
|
315 bv1 |
|
316 where |
|
317 "bv1 (BUnit1) = {}" |
|
318 | "bv1 (BV1 x) = {atom x}" |
|
319 | "bv1 (BP1 bp1 bp2) = (bv1 bp1) \<union> (bv1 bp2)" |
|
320 |
|
321 thm bv1_raw.simps |
|
322 |
|
323 (* example 2 from Terms.thy *) |
|
324 |
|
325 nominal_datatype trm2 = |
|
326 Vr2 "name" |
|
327 | Ap2 "trm2" "trm2" |
|
328 | Lm2 x::"name" t::"trm2" bind x in t |
|
329 | Lt2 r::"rassign" t::"trm2" bind "bv2 r" in t |
|
330 and rassign = |
|
331 As "name" "trm2" |
|
332 binder |
|
333 bv2 |
|
334 where |
|
335 "bv2 (As x t) = {atom x}" |
|
336 |
|
337 (* example 3 from Terms.thy *) |
|
338 |
|
339 nominal_datatype trm3 = |
|
340 Vr3 "name" |
|
341 | Ap3 "trm3" "trm3" |
|
342 | Lm3 x::"name" t::"trm3" bind x in t |
|
343 | Lt3 r::"rassigns3" t::"trm3" bind "bv3 r" in t |
|
344 and rassigns3 = |
|
345 ANil |
|
346 | ACons "name" "trm3" "rassigns3" |
|
347 binder |
|
348 bv3 |
|
349 where |
|
350 "bv3 ANil = {}" |
|
351 | "bv3 (ACons x t as) = {atom x} \<union> (bv3 as)" |
|
352 |
|
353 (* example 4 from Terms.thy *) |
|
354 |
|
355 nominal_datatype trm4 = |
|
356 Vr4 "name" |
|
357 | Ap4 "trm4" "trm4 list" |
|
358 | Lm4 x::"name" t::"trm4" bind x in t |
|
359 |
|
360 (* example 5 from Terms.thy *) |
|
361 |
|
362 nominal_datatype trm5 = |
|
363 Vr5 "name" |
|
364 | Ap5 "trm5" "trm5" |
|
365 | Lt5 l::"lts" t::"trm5" bind "bv5 l" in t |
|
366 and lts = |
|
367 Lnil |
|
368 | Lcons "name" "trm5" "lts" |
|
369 binder |
|
370 bv5 |
|
371 where |
|
372 "bv5 Lnil = {}" |
|
373 | "bv5 (Lcons n t ltl) = {atom n} \<union> (bv5 ltl)" |
|
374 |
|
375 (* example 6 from Terms.thy *) |
|
376 |
|
377 nominal_datatype trm6 = |
|
378 Vr6 "name" |
|
379 | Lm6 x::"name" t::"trm6" bind x in t |
|
380 | Lt6 left::"trm6" right::"trm6" bind "bv6 left" in right |
|
381 binder |
|
382 bv6 |
|
383 where |
|
384 "bv6 (Vr6 n) = {}" |
|
385 | "bv6 (Lm6 n t) = {atom n} \<union> bv6 t" |
|
386 | "bv6 (Lt6 l r) = bv6 l \<union> bv6 r" |
|
387 |
|
388 (* example 7 from Terms.thy *) |
|
389 |
|
390 nominal_datatype trm7 = |
|
391 Vr7 "name" |
|
392 | Lm7 l::"name" r::"trm7" bind l in r |
|
393 | Lt7 l::"trm7" r::"trm7" bind "bv7 l" in r |
|
394 binder |
|
395 bv7 |
|
396 where |
|
397 "bv7 (Vr7 n) = {atom n}" |
|
398 | "bv7 (Lm7 n t) = bv7 t - {atom n}" |
|
399 | "bv7 (Lt7 l r) = bv7 l \<union> bv7 r" |
|
400 |
|
401 (* example 8 from Terms.thy *) |
|
402 |
|
403 nominal_datatype foo8 = |
|
404 Foo0 "name" |
|
405 | Foo1 b::"bar8" f::"foo8" bind "bv8 b" in foo |
|
406 and bar8 = |
|
407 Bar0 "name" |
|
408 | Bar1 "name" s::"name" b::"bar8" bind s in b |
|
409 binder |
|
410 bv8 |
|
411 where |
|
412 "bv8 (Bar0 x) = {}" |
|
413 | "bv8 (Bar1 v x b) = {atom v}" |
|
414 |
|
415 (* example 9 from Terms.thy *) |
|
416 |
|
417 nominal_datatype lam9 = |
|
418 Var9 "name" |
|
419 | Lam9 n::"name" l::"lam9" bind n in l |
|
420 and bla9 = |
|
421 Bla9 f::"lam9" s::"lam9" bind "bv9 f" in s |
|
422 binder |
|
423 bv9 |
|
424 where |
|
425 "bv9 (Var9 x) = {}" |
|
426 | "bv9 (Lam9 x b) = {atom x}" |
|
427 |
|
428 end |
|
429 |
|
430 |
|
431 |