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)]]]
A SOME binding has to have a function returning an atom set,
and a NONE binding has to be on an argument that is an atom
or an atom set.
How the procedure works:
For each of the defined datatypes,
For each of the constructors,
It creates a union of free variables for each argument.
For an argument the free variables are the variables minus
bound variables.
The variables are:
For an atom, a singleton set with the atom itself.
For an atom set, the atom set itself.
For a recursive argument, the appropriate fv function applied to it.
(* TODO: This one is not implemented *)
For other arguments it should be an appropriate fv function stored
in the database.
The bound variables are a union of results of all bindings that
involve the given argument. For a paricular binding the result is:
For a function applied to an argument this function with the argument.
For an atom, a singleton set with the atom itself.
For an atom set, the atom set itself.
For a recursive argument, the appropriate fv function applied to it.
(* TODO: This one is not implemented *)
For other arguments it should be an appropriate fv function stored
in the database.
*)
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);
fun mk_atoms t =
let
val ty = fastype_of t;
val atom_ty = HOLogic.dest_setT ty --> @{typ atom};
val img_ty = atom_ty --> ty --> @{typ "atom set"};
in
(Const (@{const_name image}, img_ty) $ Const (@{const_name atom}, atom_ty) $ t)
end;
(* Copy from Term *)
fun is_funtype (Type ("fun", [_, _])) = true
| is_funtype _ = false;
*}
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 = nth fv_frees i;
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' *)
if (is_funtype o fastype_of) (nth args i) then mk_atoms (nth args i) else
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' *)
if (is_funtype o fastype_of) x then mk_atoms x else
HOLogic.mk_set @{typ atom} [mk_atom (fastype_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
(* The snd will be removed later *)
snd (Primrec.add_primrec
(map (fn s => (Binding.name s, NONE, NoSyn)) fv_names) fv_eqs lthy)
end
*}
ML {*
fun define_alpha 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 alpha_names = Datatype_Prop.indexify_names (map (fn (i, _) =>
"alpha_" ^ name_of_typ (nth_dtyp i)) descr);
val alpha_types = map (fn (i, _) => nth_dtyp i --> nth_dtyp i --> @{typ bool}) descr;
val alpha_frees = map Free (alpha_names ~~ alpha_types);
fun alpha_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 names2 = Name.variant_list ("pi" :: names) (Datatype_Prop.make_tnames Ts);
val args = map Free (names ~~ Ts);
val args2 = map Free (names2 ~~ Ts);
val c = Const (cname, Ts ---> (nth_dtyp i));
val alpha = nth alpha_frees i;
in
(Attrib.empty_binding, HOLogic.mk_Trueprop (alpha $ (list_comb (c, args)) $ (list_comb (c, args2))))
end;
fun alpha_eq (i, (_, _, constrs)) binds = map2 (alpha_eq_constr i) constrs binds;
val alpha_eqs = flat (map2 alpha_eq descr bindsall)
in
(* The snd will be removed later *)
snd (Inductive.add_inductive_i
{quiet_mode = false, verbose = true, alt_name = Binding.empty,
coind = false, no_elim = false, no_ind = false, skip_mono = true, fork_mono = false}
(map2 (fn x => fn y => ((Binding.name x, y), NoSyn)) alpha_names alpha_types) [] (alpha_eqs) [] lthy)
end
*}
atom_decl name
(*datatype ty =
Var "name set"
ML {* Syntax.check_term @{context} (mk_atoms @{term "a :: name set"}) *}
local_setup {* define_raw_fv "Fv.ty" [[[[]]]] *}
print_theorems
*)
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
local_setup {* define_alpha "Fv.rtrm1"
[[[[]], [[], []], [[(NONE, 0)], [(NONE, 0)]], [[(NONE, 0)], [], [(SOME @{term bv1}, 0)]]],
[[], [[]], [[], []]]] *}
print_theorems
end