411 |
411 |
412 ML {* |
412 ML {* |
413 fun setify x = |
413 fun setify x = |
414 if fastype_of x = @{typ "atom list"} then |
414 if fastype_of x = @{typ "atom list"} then |
415 Const (@{const_name set}, @{typ "atom list \<Rightarrow> atom set"}) $ x else x |
415 Const (@{const_name set}, @{typ "atom list \<Rightarrow> atom set"}) $ x else x |
|
416 *} |
|
417 |
|
418 ML {* |
|
419 fun define_fv (dt_info : Datatype_Aux.info) bindsall bns lthy = |
|
420 let |
|
421 val thy = ProofContext.theory_of lthy; |
|
422 val {descr, sorts, ...} = dt_info; |
|
423 fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i); |
|
424 val fv_names = Datatype_Prop.indexify_names (map (fn (i, _) => |
|
425 "fv_" ^ name_of_typ (nth_dtyp i)) descr); |
|
426 val fv_types = map (fn (i, _) => nth_dtyp i --> @{typ "atom set"}) descr; |
|
427 val fv_frees = map Free (fv_names ~~ fv_types); |
|
428 (* TODO: We need a transitive closure, but instead we do this hack considering |
|
429 all binding functions as recursive or not *) |
|
430 val nr_bns = |
|
431 if (non_rec_binds bindsall) = [] then [] |
|
432 else map (fn (bn, _, _) => bn) bns; |
|
433 val rel_bns = filter (fn (bn, _, _) => bn mem nr_bns) bns; |
|
434 val (bn_fv_bns, fv_bn_names_eqs) = fv_bns thy dt_info fv_frees rel_bns; |
|
435 val fvbns = map snd bn_fv_bns; |
|
436 val (fv_bn_names, fv_bn_eqs) = split_list fv_bn_names_eqs; |
|
437 |
|
438 fun fv_constr ith_dtyp (cname, dts) bindcs = |
|
439 let |
|
440 val Ts = map (typ_of_dtyp descr sorts) dts; |
|
441 val bindslen = length bindcs |
|
442 val pi_strs_same = replicate bindslen "pi" |
|
443 val pi_strs = Name.variant_list [] pi_strs_same; |
|
444 val pis = map (fn ps => Free (ps, @{typ perm})) pi_strs; |
|
445 val bind_pis_gath = bindcs ~~ pis; |
|
446 val bind_pis = un_gather_binds_cons bind_pis_gath; |
|
447 val bindcs = map fst bind_pis; |
|
448 val names = Name.variant_list pi_strs (Datatype_Prop.make_tnames Ts); |
|
449 val args = map Free (names ~~ Ts); |
|
450 val c = Const (cname, Ts ---> (nth_dtyp ith_dtyp)); |
|
451 val fv_c = nth fv_frees ith_dtyp; |
|
452 val arg_nos = 0 upto (length dts - 1) |
|
453 fun fv_bind args (NONE, i, _, _) = |
|
454 if is_rec_type (nth dts i) then (nth fv_frees (body_index (nth dts i))) $ (nth args i) else |
|
455 if ((is_atom thy) o fastype_of) (nth args i) then mk_single_atom (nth args i) else |
|
456 if ((is_atom_set thy) o fastype_of) (nth args i) then mk_atom_set (nth args i) else |
|
457 if ((is_atom_fset thy) o fastype_of) (nth args i) then mk_atom_fset (nth args i) else |
|
458 (* TODO goes the code for preiously defined nominal datatypes *) |
|
459 @{term "{} :: atom set"} |
|
460 | fv_bind args (SOME (f, _), i, _, _) = f $ (nth args i) |
|
461 fun fv_binds_as_set args relevant = mk_union (map (setify o fv_bind args) relevant) |
|
462 fun find_nonrec_binder j (SOME (f, false), i, _, _) = if i = j then SOME f else NONE |
|
463 | find_nonrec_binder _ _ = NONE |
|
464 fun fv_arg ((dt, x), arg_no) = |
|
465 case get_first (find_nonrec_binder arg_no) bindcs of |
|
466 SOME f => |
|
467 (case get_first (fn (x, y) => if x = f then SOME y else NONE) bn_fv_bns of |
|
468 SOME fv_bn => fv_bn $ x |
|
469 | NONE => error "bn specified in a non-rec binding but not in bn list") |
|
470 | NONE => |
|
471 let |
|
472 val arg = |
|
473 if is_rec_type dt then nth fv_frees (body_index dt) $ x else |
|
474 if ((is_atom thy) o fastype_of) x then mk_single_atom x else |
|
475 if ((is_atom_set thy) o fastype_of) x then mk_atom_set x else |
|
476 if ((is_atom_fset thy) o fastype_of) x then mk_atom_fset x else |
|
477 (* TODO goes the code for preiously defined nominal datatypes *) |
|
478 @{term "{} :: atom set"}; |
|
479 (* If i = j then we generate it only once *) |
|
480 val relevant = filter (fn (_, i, j, _) => ((i = arg_no) orelse (j = arg_no))) bindcs; |
|
481 val sub = fv_binds_as_set args relevant |
|
482 in |
|
483 mk_diff arg sub |
|
484 end; |
|
485 val fv_eq = HOLogic.mk_Trueprop (HOLogic.mk_eq |
|
486 (fv_c $ list_comb (c, args), mk_union (map fv_arg (dts ~~ args ~~ arg_nos)))) |
|
487 in |
|
488 fv_eq |
|
489 end; |
|
490 fun fv_eq (i, (_, _, constrs)) binds = map2i (fv_constr i) constrs binds; |
|
491 val fveqs = map2i fv_eq descr (gather_binds bindsall) |
|
492 val fv_eqs_perfv = fveqs |
|
493 val rel_bns_nos = map (fn (_, i, _) => i) rel_bns; |
|
494 fun filter_fun (_, b) = b mem rel_bns_nos; |
|
495 val all_fvs = (fv_names ~~ fv_eqs_perfv) ~~ (0 upto (length fv_names - 1)) |
|
496 val (fv_names_fst, fv_eqs_fst) = apsnd flat (split_list (map fst (filter_out filter_fun all_fvs))) |
|
497 val (fv_names_snd, fv_eqs_snd) = apsnd flat (split_list (map fst (filter filter_fun all_fvs))) |
|
498 val fv_eqs_all = fv_eqs_fst @ (flat fv_bn_eqs); |
|
499 val fv_names_all = fv_names_fst @ fv_bn_names; |
|
500 val add_binds = map (fn x => (Attrib.empty_binding, x)) |
|
501 (* Function_Fun.add_fun Function_Common.default_config ... true *) |
|
502 val (fvs, lthy') = (Primrec.add_primrec |
|
503 (map (fn s => (Binding.name s, NONE, NoSyn)) fv_names_all) (add_binds fv_eqs_all) lthy) |
|
504 val (fvs2, lthy'') = |
|
505 if fv_eqs_snd = [] then (([], []), lthy') else |
|
506 (Primrec.add_primrec |
|
507 (map (fn s => (Binding.name s, NONE, NoSyn)) fv_names_snd) (add_binds fv_eqs_snd) lthy') |
|
508 val ordered_fvs = fv_frees @ fvbns; |
|
509 val all_fvs = (fst fvs @ fst fvs2, snd fvs @ snd fvs2) |
|
510 in |
|
511 ((all_fvs, ordered_fvs), lthy'') |
|
512 end |
416 *} |
513 *} |
417 |
514 |
418 ML {* |
515 ML {* |
419 fun define_fv_alpha (dt_info : Datatype_Aux.info) bindsall bns lthy = |
516 fun define_fv_alpha (dt_info : Datatype_Aux.info) bindsall bns lthy = |
420 let |
517 let |