Code for generating the fv function, no bindings yet.
authorCezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 17 Feb 2010 10:12:01 +0100
changeset 1168 5c1e16806901
parent 1167 e09d148ccc94
child 1169 b9d02e0800e9
Code for generating the fv function, no bindings yet.
Quot/Nominal/Fv.thy
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Quot/Nominal/Fv.thy	Wed Feb 17 10:12:01 2010 +0100
@@ -0,0 +1,64 @@
+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} *}
+
+ML {*
+  val {descr, ...} = Datatype.the_info @{theory} "Fv.rlam";
+  val sorts = [];
+
+  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) =
+    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) =
+        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
+      (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)) = map (fv_eq_constr i) constrs;
+  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