Quot/Nominal/Fv.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 17 Feb 2010 16:22:16 +0100
changeset 1180 3f36936f1280
parent 1178 275a1cb3f2ba
child 1185 7566b899ca6a
permissions -rw-r--r--
Testing Fv

theory Fv
imports "Nominal2_Atoms"
begin

(* 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)]]]
*)

ML {*
  open Datatype_Aux; (* typ_of_dtyp, DtRec, ... *);
  (* TODO: It is the same as one in 'nominal_atoms' *)
  fun mk_atom ty = Const (@{const_name atom}, ty --> @{typ atom});
  val noatoms = @{term "{} :: atom set"};
  fun mk_single_atom x = HOLogic.mk_set @{typ atom} [mk_atom (type_of x) $ x];
  fun mk_union sets =
    fold (fn a => fn b =>
      if a = noatoms then b else
      if b = noatoms then a else
      HOLogic.mk_binop @{const_name union} (a, b)) (rev sets) noatoms;
  fun mk_diff a b =
    if b = noatoms then a else
    if b = a then noatoms else
    HOLogic.mk_binop @{const_name minus} (a, b);
*}

ML {*
(* Currently needs just one full_tname to access Datatype *)
fun define_raw_fv full_tname bindsall lthy =
let
  val thy = ProofContext.theory_of lthy
  val {descr, ...} = Datatype.the_info thy full_tname;
  val sorts = []; (* TODO *)
  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);
  fun fv_eq_constr i (cname, dts) bindcs =
    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_bind (NONE, i) =
            if is_rec_type (nth dts i) then (nth fv_frees (body_index (nth dts i))) $ (nth args i) else
            (* TODO we assume that all can be 'atomized' *)
            mk_single_atom (nth args i)
        | fv_bind (SOME f, i) = f $ (nth args i);
      fun fv_arg ((dt, x), bindxs) =
        let
          val arg =
            if is_rec_type dt then nth fv_frees (body_index dt) $ x else
            (* TODO: we just assume everything can be 'atomized' *)
            HOLogic.mk_set @{typ atom} [mk_atom (type_of x) $ x]
          val sub = mk_union (map fv_bind bindxs)
        in
          mk_diff arg sub
        end;
        val _ = tracing ("d" ^ string_of_int (length dts));
        val _ = tracing (string_of_int (length args));
        val _ = tracing (string_of_int (length bindcs));
    in
      (Attrib.empty_binding, HOLogic.mk_Trueprop (HOLogic.mk_eq
        (fv_c $ list_comb (c, args), mk_union (map fv_arg (dts ~~ args ~~ bindcs)))))
    end;
  fun fv_eq (i, (_, _, constrs)) binds = map2 (fv_eq_constr i) constrs binds;
  val fv_eqs = flat (map2 fv_eq descr bindsall)
in
  snd (Primrec.add_primrec
    (map (fn s => (Binding.name s, NONE, NoSyn)) fv_names) fv_eqs lthy)
end
*}

(* test
atom_decl name

datatype rtrm1 =
  rVr1 "name"
| rAp1 "rtrm1" "rtrm1"
| rLm1 "name" "rtrm1"        --"name is bound in trm1"
| rLt1 "bp" "rtrm1" "rtrm1"   --"all variables in bp are bound in the 2nd trm1"
and bp =
  BUnit
| BVr "name"
| BPr "bp" "bp"

(* to be given by the user *)

primrec 
  bv1
where
  "bv1 (BUnit) = {}"
| "bv1 (BVr x) = {atom x}"
| "bv1 (BPr bp1 bp2) = (bv1 bp1) \<union> (bv1 bp1)"

local_setup {* define_raw_fv "Fv.rtrm1"
  [[[[]], [[], []], [[(NONE, 0)], [(NONE, 0)]], [[(NONE, 0)], [], [(SOME @{term bv1}, 0)]]],
   [[], [[]], [[], []]]] *}
print_theorems

*)

end