81 *} |
81 *} |
82 |
82 |
83 (* adds "_raw" to the end of constants and types *) |
83 (* adds "_raw" to the end of constants and types *) |
84 ML {* |
84 ML {* |
85 fun add_raw s = s ^ "_raw" |
85 fun add_raw s = s ^ "_raw" |
|
86 fun add_raws ss = map add_raw ss |
86 fun raw_bind bn = Binding.suffix_name "_raw" bn |
87 fun raw_bind bn = Binding.suffix_name "_raw" bn |
87 |
88 |
88 fun raw_str [] s = s |
89 fun replace_str ss s = |
89 | raw_str (s'::ss) s = |
90 case (AList.lookup (op=) ss s) of |
90 if (Long_Name.base_name s) = s' |
91 SOME s' => s' |
91 then add_raw s |
92 | NONE => s |
92 else raw_str ss s |
93 |
93 |
94 fun replace_typ ty_ss (Type (a, Ts)) = Type (replace_str ty_ss a, map (replace_typ ty_ss) Ts) |
94 fun raw_typ ty_strs (Type (a, Ts)) = Type (raw_str ty_strs a, map (raw_typ ty_strs) Ts) |
95 | replace_typ ty_ss T = T |
95 | raw_typ ty_strs T = T |
96 |
96 |
97 fun raw_dts ty_ss dts = |
97 fun raw_aterm trm_strs (Const (a, T)) = Const (raw_str trm_strs a, T) |
98 let |
98 | raw_aterm trm_strs (Free (a, T)) = Free (raw_str trm_strs a, T) |
99 val ty_ss' = ty_ss ~~ (add_raws ty_ss) |
99 | raw_aterm trm_strs trm = trm |
100 |
100 |
|
101 fun raw_term trm_strs ty_strs trm = |
|
102 trm |> Term.map_aterms (raw_aterm trm_strs) |> map_types (raw_typ ty_strs) |
|
103 |
|
104 fun raw_dts ty_strs dts = |
|
105 let |
|
106 fun raw_dts_aux1 (bind, tys, mx) = |
101 fun raw_dts_aux1 (bind, tys, mx) = |
107 (raw_bind bind, map (raw_typ ty_strs) tys, mx) |
102 (raw_bind bind, map (replace_typ ty_ss') tys, mx) |
108 |
103 |
109 fun raw_dts_aux2 (ty_args, bind, mx, constrs) = |
104 fun raw_dts_aux2 (ty_args, bind, mx, constrs) = |
110 (ty_args, raw_bind bind, mx, map raw_dts_aux1 constrs) |
105 (ty_args, raw_bind bind, mx, map raw_dts_aux1 constrs) |
111 in |
106 in |
112 map raw_dts_aux2 dts |
107 map raw_dts_aux2 dts |
113 end |
108 end |
114 |
109 |
|
110 fun replace_aterm trm_ss (Const (a, T)) = Const (replace_str trm_ss a, T) |
|
111 | replace_aterm trm_ss (Free (a, T)) = Free (replace_str trm_ss a, T) |
|
112 | replace_aterm trm_ss trm = trm |
|
113 |
|
114 fun replace_term trm_ss ty_ss trm = |
|
115 trm |> Term.map_aterms (replace_aterm trm_ss) |> map_types (replace_typ ty_ss) |
|
116 *} |
|
117 |
|
118 ML {* |
115 fun get_constrs dts = |
119 fun get_constrs dts = |
116 flat (map (fn (_, _, _, constrs) => map (fn (bn, tys, mx) => (bn, tys, mx)) constrs) dts) |
120 flat (map (fn (_, _, _, constrs) => constrs) dts) |
|
121 |
|
122 fun get_typed_constrs dts = |
|
123 flat (map (fn (_, bn, _, constrs) => |
|
124 (map (fn (bn', _, _) => (Binding.name_of bn, Binding.name_of bn')) constrs)) dts) |
117 |
125 |
118 fun get_constr_strs dts = |
126 fun get_constr_strs dts = |
119 map (fn (bn, _, _) => Binding.name_of bn) (get_constrs dts) |
127 map (fn (bn, _, _) => Binding.name_of bn) (get_constrs dts) |
120 |
128 |
121 fun get_constrs2 dts = |
|
122 flat (map (fn (tvrs, tname, _, constrs) => |
|
123 map (fn (bn, tys, mx) => (bn, tys ---> Type ("Test." ^ Binding.name_of tname, map (fn s => TVar ((s,0),[])) tvrs), mx)) constrs) dts) |
|
124 |
|
125 fun get_bn_fun_strs bn_funs = |
129 fun get_bn_fun_strs bn_funs = |
126 map (fn (bn_fun, _, _) => Binding.name_of bn_fun) bn_funs |
130 map (fn (bn_fun, _, _) => Binding.name_of bn_fun) bn_funs |
127 *} |
131 *} |
128 |
132 |
129 ML {* |
133 ML {* |
130 fun raw_dts_decl dt_names dts lthy = |
134 fun raw_dts_decl dt_names dts lthy = |
131 let |
135 let |
|
136 val thy = ProofContext.theory_of lthy |
132 val conf = Datatype.default_config |
137 val conf = Datatype.default_config |
133 |
138 |
134 val dt_names' = map add_raw dt_names |
139 val dt_names' = add_raws dt_names |
135 val dts' = raw_dts dt_names dts |
140 val dt_full_names = map (Sign.full_bname thy) dt_names |
136 in |
141 val dts' = raw_dts dt_full_names dts |
137 Local_Theory.theory_result (Datatype.add_datatype conf dt_names' dts') lthy |
142 in |
|
143 lthy |
|
144 |> Local_Theory.theory_result (Datatype.add_datatype conf dt_names' dts') |
138 end |
145 end |
139 *} |
146 *} |
140 |
147 |
141 ML {* |
148 ML {* |
142 fun raw_bn_fun_decl dt_names cnstr_names bn_funs bn_eqs lthy = |
149 fun raw_bn_fun_decl dt_names dts bn_funs bn_eqs lthy = |
143 let |
150 let |
|
151 val thy = ProofContext.theory_of lthy |
|
152 |
|
153 val dt_names' = add_raws dt_names |
|
154 val dt_full_names = map (Sign.full_bname thy) dt_names |
|
155 val dt_full_names' = map (Sign.full_bname thy) dt_names' |
|
156 |
|
157 val ctrs_names = map (Sign.full_bname thy) (get_constr_strs dts) |
|
158 val ctrs_names' = map (fn (x, y) => (Sign.full_bname_path thy (add_raw x) (add_raw y))) |
|
159 (get_typed_constrs dts) |
|
160 |
144 val bn_fun_strs = get_bn_fun_strs bn_funs |
161 val bn_fun_strs = get_bn_fun_strs bn_funs |
|
162 val bn_fun_strs' = add_raws bn_fun_strs |
145 |
163 |
146 val bn_funs' = map (fn (bn, opt_ty, mx) => |
164 val bn_funs' = map (fn (bn, opt_ty, mx) => |
147 (raw_bind bn, Option.map (raw_typ dt_names) opt_ty, mx)) bn_funs |
165 (raw_bind bn, Option.map (replace_typ (dt_full_names ~~ dt_full_names')) opt_ty, mx)) bn_funs |
|
166 |
148 val bn_eqs' = map (fn trm => |
167 val bn_eqs' = map (fn trm => |
149 (Attrib.empty_binding, raw_term (cnstr_names @ bn_fun_strs) dt_names trm)) bn_eqs |
168 (Attrib.empty_binding, |
150 in |
169 (replace_term ((ctrs_names ~~ ctrs_names') @ (bn_fun_strs ~~ bn_fun_strs')) (dt_full_names ~~ dt_full_names') trm))) bn_eqs |
151 Primrec.add_primrec bn_funs' bn_eqs' lthy |
170 in |
|
171 if null bn_eqs |
|
172 then (([], []), lthy) |
|
173 else Primrec.add_primrec bn_funs' bn_eqs' lthy |
152 end |
174 end |
153 *} |
175 *} |
154 |
176 |
155 ML {* |
177 ML {* |
156 fun nominal_datatype2 dts bn_funs bn_eqs lthy = |
178 fun nominal_datatype2 dts bn_funs bn_eqs lthy = |
157 let |
179 let |
158 val dts_names = map (fn (_, s, _, _) => Binding.name_of s) dts |
180 val dts_names = map (fn (_, s, _, _) => Binding.name_of s) dts |
159 val ctrs_names = get_constr_strs dts |
|
160 in |
181 in |
161 lthy |
182 lthy |
162 |> raw_dts_decl dts_names dts |
183 |> raw_dts_decl dts_names dts |
163 ||>> raw_bn_fun_decl dts_names ctrs_names bn_funs bn_eqs |
184 ||>> raw_bn_fun_decl dts_names dts bn_funs bn_eqs |
164 end |
185 end |
165 *} |
186 *} |
166 |
187 |
|
188 ML {* |
|
189 (* makes a full named type out of a binding with tvars applied to it *) |
|
190 fun mk_type thy bind tvrs = |
|
191 Type (Sign.full_name thy bind, map (fn s => TVar ((s, 0), [])) tvrs) |
|
192 |
|
193 fun get_constrs2 thy dts = |
|
194 let |
|
195 val dts' = map (fn (tvrs, tname, _, constrs) => (mk_type thy tname tvrs, constrs)) dts |
|
196 in |
|
197 flat (map (fn (ty, constrs) => map (fn (bn, tys, mx) => (bn, tys ---> ty, mx)) constrs) dts') |
|
198 end |
|
199 *} |
167 |
200 |
168 ML {* |
201 ML {* |
169 fun nominal_datatype2_cmd (dt_strs, (bn_fun_strs, bn_eq_strs)) lthy = |
202 fun nominal_datatype2_cmd (dt_strs, (bn_fun_strs, bn_eq_strs)) lthy = |
170 let |
203 let |
|
204 val thy = ProofContext.theory_of lthy |
|
205 |
|
206 fun prep_typ ((tvs, tname, mx), _) = (tname, length tvs, mx); |
|
207 |
|
208 (* adding the types for parsing datatypes *) |
171 val lthy_tmp = lthy |
209 val lthy_tmp = lthy |
172 |> Local_Theory.theory |
210 |> Local_Theory.theory (Sign.add_types (map prep_typ dt_strs)) |
173 (Sign.add_types (map (fn ((tvs, tname, mx), _) => |
|
174 (tname, length tvs, mx)) dt_strs)) |
|
175 |
211 |
176 fun prep_cnstr lthy (((cname, atys), mx), binders) = |
212 fun prep_cnstr lthy (((cname, atys), mx), binders) = |
177 (cname, map (Syntax.read_typ lthy o snd) atys, mx) |
213 (cname, map (Syntax.read_typ lthy o snd) atys, mx) |
178 |
214 |
179 fun prep_dt lthy ((tvs, tname, mx), cnstrs) = |
215 fun prep_dt lthy ((tvs, tname, mx), cnstrs) = |
180 (tvs, tname, mx, map (prep_cnstr lthy) cnstrs) |
216 (tvs, tname, mx, map (prep_cnstr lthy) cnstrs) |
181 |
217 |
|
218 (* parsing the datatypes *) |
182 val dts_prep = map (prep_dt lthy_tmp) dt_strs |
219 val dts_prep = map (prep_dt lthy_tmp) dt_strs |
183 |
220 |
|
221 (* adding constructors for parsing functions *) |
184 val lthy_tmp2 = lthy_tmp |
222 val lthy_tmp2 = lthy_tmp |
185 |> Local_Theory.theory (Sign.add_consts_i (get_constrs2 dts_prep)) |
223 |> Local_Theory.theory (Sign.add_consts_i (get_constrs2 thy dts_prep)) |
|
224 |
|
225 val (bn_fun_aux, bn_eq_aux) = fst (Specification.read_spec bn_fun_strs bn_eq_strs lthy_tmp2) |
186 |
226 |
187 fun prep_bn_fun ((bn, T), mx) = (bn, SOME T, mx) |
227 fun prep_bn_fun ((bn, T), mx) = (bn, SOME T, mx) |
188 |
228 |
189 fun prep_bn_eq (attr, t) = t |
229 fun prep_bn_eq (attr, t) = t |
190 |
230 |
191 val (bn_fun_aux, bn_eq_aux) = fst (Specification.read_spec bn_fun_strs bn_eq_strs lthy_tmp2) |
|
192 |
|
193 val bn_fun_prep = map prep_bn_fun bn_fun_aux |
231 val bn_fun_prep = map prep_bn_fun bn_fun_aux |
194 val bn_eq_prep = map prep_bn_eq bn_eq_aux |
232 val bn_eq_prep = map prep_bn_eq bn_eq_aux |
195 |
233 |
196 val _ = tracing (cat_lines (map (Syntax.string_of_term lthy_tmp2) bn_eq_prep)) |
|
197 in |
234 in |
198 nominal_datatype2 dts_prep bn_fun_prep bn_eq_prep lthy |> snd |
235 nominal_datatype2 dts_prep bn_fun_prep bn_eq_prep lthy |> snd |
199 end |
236 end |
200 *} |
237 *} |
201 |
238 |
268 nominal_datatype tyS = |
311 nominal_datatype tyS = |
269 All xs::"name list" ty::"ty" bind xs in ty |
312 All xs::"name list" ty::"ty" bind xs in ty |
270 |
313 |
271 |
314 |
272 |
315 |
273 |
316 (* example 1 from Terms.thy *) |
274 end |
317 |
275 |
318 nominal_datatype trm1 = |
276 |
319 Vr1 "name" |
277 (* printing stuff *) |
320 | Ap1 "trm1" "trm1" |
278 |
321 | Lm1 x::"name" t::"trm1" bind x in t |
279 ML {* |
322 | Lt1 p::"bp1" "trm1" t::"trm1" bind "bv1 p" in t |
280 fun print_str_option NONE = "NONE" |
323 and bp1 = |
281 | print_str_option (SOME s) = (s:bstring) |
324 BUnit1 |
282 |
325 | BV1 "name" |
283 fun print_anno_typ lthy (NONE, ty) = Syntax.string_of_typ lthy ty |
326 | BP1 "bp1" "bp1" |
284 | print_anno_typ lthy (SOME x, ty) = x ^ ":" ^ Syntax.string_of_typ lthy ty |
327 binder |
285 |
328 bv1 |
286 fun print_bind (s1, s2) = "bind " ^ s1 ^ " in " ^ s2 |
329 where |
287 |
330 "bv1 (BUnit1) = {}" |
288 fun print_constr lthy (((name, anno_tys), bds), syn) = |
331 | "bv1 (BV1 x) = {atom x}" |
289 (Binding.name_of name ^ " of " ^ (commas (map (print_anno_typ lthy) anno_tys)) ^ " " |
332 | "bv1 (BP1 bp1 bp2) = (bv1 bp1) \<union> (bv1 bp2)" |
290 ^ (commas (map print_bind bds)) ^ " " |
333 |
291 ^ (Pretty.str_of (Syntax.pretty_mixfix syn))) |
334 thm bv1_raw.simps |
292 |
335 |
293 fun print_single_dt lthy (((s2, s3), syn), constrs) = |
336 (* example 2 from Terms.thy *) |
294 ["constructor declaration"] |
337 |
295 @ ["type arguments: "] @ s2 |
338 nominal_datatype trm2 = |
296 @ ["datatype name: ", Binding.name_of s3] |
339 Vr2 "name" |
297 @ ["typ mixfix: ", Pretty.string_of (Syntax.pretty_mixfix syn)] |
340 | Ap2 "trm2" "trm2" |
298 @ ["constructors: "] @ (map (print_constr lthy) constrs) |
341 | Lm2 x::"name" t::"trm2" bind x in t |
299 |> separate "\n" |
342 | Lt2 r::"rassign" t::"trm2" bind "bv2 r" in t |
300 |> implode |
343 and rassign = |
301 |
344 As "name" "trm2" |
302 fun print_single_fun_eq lthy (at, s) = |
345 binder |
303 ["function equations: ", (Syntax.string_of_term lthy s)] |
346 bv2 |
304 |> separate "\n" |
347 where |
305 |> implode |
348 "bv2 (As x t) = {atom x}" |
306 |
349 |
307 fun print_single_fun_fix lthy (b, _, _) = |
350 (* example 3 from Terms.thy *) |
308 ["functions: ", Binding.name_of b] |
351 |
309 |> separate "\n" |
352 nominal_datatype trm3 = |
310 |> implode |
353 Vr3 "name" |
311 |
354 | Ap3 "trm3" "trm3" |
312 fun dt_name (((s2, s3), syn), constrs) = Binding.name_of s3 |
355 | Lm3 x::"name" t::"trm3" bind x in t |
313 |
356 | Lt3 r::"rassigns3" t::"trm3" bind "bv3 r" in t |
314 fun print_dts (dts, (funs, feqs)) lthy = |
357 and rassigns3 = |
315 let |
358 ANil |
316 val _ = warning (implode (separate "\n" (map (print_single_dt lthy) dts))); |
359 | ACons "name" "trm3" "rassigns3" |
317 val _ = warning (implode (separate "\n" (map (print_single_fun_fix lthy) funs))); |
360 binder |
318 val _ = warning (implode (separate "\n" (map (print_single_fun_eq lthy) feqs))); |
361 bv3 |
319 in |
362 where |
320 () |
363 "bv3 ANil = {}" |
321 end |
364 | "bv3 (ACons x t as) = {atom x} \<union> (bv3 as)" |
322 *} |
365 |
323 |
366 (* example 4 from Terms.thy *) |
|
367 |
|
368 nominal_datatype trm4 = |
|
369 Vr4 "name" |
|
370 | Ap4 "trm4" "trm4 list" |
|
371 | Lm4 x::"name" t::"trm4" bind x in t |
|
372 |
|
373 (* example 5 from Terms.thy *) |
|
374 |
|
375 nominal_datatype trm5 = |
|
376 Vr5 "name" |
|
377 | Ap5 "trm5" "trm5" |
|
378 | Lt5 l::"lts" t::"trm5" bind "bv5 l" in t |
|
379 and lts = |
|
380 Lnil |
|
381 | Lcons "name" "trm5" "lts" |
|
382 binder |
|
383 bv5 |
|
384 where |
|
385 "bv5 Lnil = {}" |
|
386 | "bv5 (Lcons n t ltl) = {atom n} \<union> (bv5 ltl)" |
|
387 |
|
388 (* example 6 from Terms.thy *) |
|
389 |
|
390 nominal_datatype trm6 = |
|
391 Vr6 "name" |
|
392 | Lm6 x::"name" t::"trm6" bind x in t |
|
393 | Lt6 left::"trm6" right::"trm6" bind "bv6 left" in right |
|
394 binder |
|
395 bv6 |
|
396 where |
|
397 "bv6 (Vr6 n) = {}" |
|
398 | "bv6 (Lm6 n t) = {atom n} \<union> bv6 t" |
|
399 | "bv6 (Lt6 l r) = bv6 l \<union> bv6 r" |
|
400 |
|
401 (* example 7 from Terms.thy *) |
|
402 |
|
403 nominal_datatype trm7 = |
|
404 Vr7 "name" |
|
405 | Lm7 l::"name" r::"trm7" bind l in r |
|
406 | Lt7 l::"trm7" r::"trm7" bind "bv7 l" in r |
|
407 binder |
|
408 bv7 |
|
409 where |
|
410 "bv7 (Vr7 n) = {atom n}" |
|
411 | "bv7 (Lm7 n t) = bv7 t - {atom n}" |
|
412 | "bv7 (Lt7 l r) = bv7 l \<union> bv7 r" |
|
413 |
|
414 (* example 8 from Terms.thy *) |
|
415 |
|
416 nominal_datatype foo8 = |
|
417 Foo0 "name" |
|
418 | Foo1 b::"bar8" f::"foo8" bind "bv8 b" in foo |
|
419 and bar8 = |
|
420 Bar0 "name" |
|
421 | Bar1 "name" s::"name" b::"bar8" bind s in b |
|
422 binder |
|
423 bv8 |
|
424 where |
|
425 "bv8 (Bar0 x) = {}" |
|
426 | "bv8 (Bar1 v x b) = {atom v}" |
|
427 |
|
428 (* example 9 from Terms.thy *) |
|
429 |
|
430 nominal_datatype lam9 = |
|
431 Var9 "name" |
|
432 | Lam9 n::"name" l::"lam9" bind n in l |
|
433 and bla9 = |
|
434 Bla9 f::"lam9" s::"lam9" bind "bv9 f" in s |
|
435 binder |
|
436 bv9 |
|
437 where |
|
438 "bv9 (Var9 x) = {}" |
|
439 | "bv9 (Lam9 x b) = {atom x}" |
|
440 |
|
441 end |
|
442 |
|
443 |
|
444 |