Quot/Nominal/Fv.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 17 Feb 2010 14:17:02 +0100
changeset 1172 9a609fefcf24
parent 1169 b9d02e0800e9
child 1173 9cb99a28b40e
permissions -rw-r--r--
Simplified format of bindings.

theory Fv
imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs"
begin

atom_decl name

datatype rlam =
  rVar "name"
| rApp "rlam" "rlam"
| rLam "name" "rlam"

ML {*
  open Datatype_Aux (* typ_of_dtyp, DtRec, ... *);
  fun mk_atom ty = Const (@{const_name atom}, ty --> @{typ atom})
*}

ML {* dest_Type @{typ rlam} *}

(* Bindings are given as a list which has a length being equal
   to the length of the number of constructors.

   Each element is a list whose length is equal to the number
   of arguents.

   Every element specifies bindings of this argument given as
   a tuple: function, bound argument.

  Eg:
nominal_datatype

   C1
 | C2 x y z bind x in z
 | C3 x y z bind f x in z bind g y in z

yields:
[
 [],
 [[], [], [(NONE, 0)]],
 [[], [], [(SOME (Const f), 0), (Some (Const g), 1, 2)]]]
*)

ML {*
  val {descr, ...} = Datatype.the_info @{theory} "Fv.rlam";
  val sorts = [];
  val binds = [[], [[], []], [[], [(NONE, 0)]]];
  fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i);
  val fv_names = Datatype_Prop.indexify_names (map (fn (i, _) =>
    "fv_" ^ name_of_typ (nth_dtyp i)) descr);
  val fv_types = map (fn (i, _) => nth_dtyp i --> @{typ "atom set"}) descr;
  val fv_frees = map Free (fv_names ~~ fv_types);
  (* TODO: should be optimized to avoid {}+{}+{} *)
  fun mk_union sets = foldr1 (HOLogic.mk_binop @{const_name union}) sets;

  fun fv_eq_constr i (cname, dts) bind =
    let
      val Ts = map (typ_of_dtyp descr sorts) dts;
      val names = Name.variant_list ["pi"] (Datatype_Prop.make_tnames Ts);
      val args = map Free (names ~~ Ts);
      val c = Const (cname, Ts ---> (nth_dtyp i));
      val fv_c = Free (nth fv_names i, (nth_dtyp i) --> @{typ "atom set"});
      fun fv_arg (dt, x) =
        let
          val arg =
            if is_rec_type dt then nth fv_frees (body_index dt) $ x
            (* TODO: we just assume everything is a 'name' *)
            else HOLogic.mk_set @{typ atom} [mk_atom (type_of x) $ x]
        in
          arg
        end
     in
      (Attrib.empty_binding, HOLogic.mk_Trueprop (HOLogic.mk_eq
        (fv_c $ list_comb (c, args), mk_union (map fv_arg (dts ~~ args)))))
    end;
  fun fv_eq (i, (_, _, constrs)) = map2 (fv_eq_constr i) constrs binds;
  val fv_eqs = maps fv_eq descr
*}

local_setup {*
snd o (Primrec.add_primrec
  (map (fn s => (Binding.name s, NONE, NoSyn)) fv_names) fv_eqs)
*}
print_theorems

fun
  fv_rlam :: "rlam \<Rightarrow> atom set"
where
  "fv_rlam (rVar a) = {atom a}"
| "fv_rlam (rApp t1 t2) = (fv_rlam t1) \<union> (fv_rlam t2)"
| "fv_rlam (rLam a t) = (fv_rlam t) - {atom a}"

end