|
1 theory Fv |
|
2 imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs" |
|
3 begin |
|
4 |
|
5 atom_decl name |
|
6 |
|
7 datatype rlam = |
|
8 rVar "name" |
|
9 | rApp "rlam" "rlam" |
|
10 | rLam "name" "rlam" |
|
11 |
|
12 ML {* |
|
13 open Datatype_Aux (* typ_of_dtyp, DtRec, ... *); |
|
14 fun mk_atom ty = Const (@{const_name atom}, ty --> @{typ atom}) |
|
15 *} |
|
16 |
|
17 ML {* dest_Type @{typ rlam} *} |
|
18 |
|
19 ML {* |
|
20 val {descr, ...} = Datatype.the_info @{theory} "Fv.rlam"; |
|
21 val sorts = []; |
|
22 |
|
23 fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i); |
|
24 val fv_names = Datatype_Prop.indexify_names (map (fn (i, _) => |
|
25 "fv_" ^ name_of_typ (nth_dtyp i)) descr); |
|
26 val fv_types = map (fn (i, _) => nth_dtyp i --> @{typ "atom set"}) descr; |
|
27 val fv_frees = map Free (fv_names ~~ fv_types); |
|
28 (* TODO: should be optimized to avoid {}+{}+{} *) |
|
29 fun mk_union sets = foldr1 (HOLogic.mk_binop @{const_name union}) sets; |
|
30 |
|
31 fun fv_eq_constr i (cname, dts) = |
|
32 let |
|
33 val Ts = map (typ_of_dtyp descr sorts) dts; |
|
34 val names = Name.variant_list ["pi"] (Datatype_Prop.make_tnames Ts); |
|
35 val args = map Free (names ~~ Ts); |
|
36 val c = Const (cname, Ts ---> (nth_dtyp i)); |
|
37 val fv_c = Free (nth fv_names i, (nth_dtyp i) --> @{typ "atom set"}); |
|
38 fun fv_arg (dt, x) = |
|
39 if is_rec_type dt then |
|
40 nth fv_frees (body_index dt) $ x |
|
41 (* TODO: we just assume everything is a 'name' *) |
|
42 else HOLogic.mk_set @{typ atom} [mk_atom (type_of x) $ x] |
|
43 in |
|
44 (Attrib.empty_binding, HOLogic.mk_Trueprop (HOLogic.mk_eq |
|
45 (fv_c $ list_comb (c, args), mk_union (map fv_arg (dts ~~ args))))) |
|
46 end; |
|
47 fun fv_eq (i, (_, _, constrs)) = map (fv_eq_constr i) constrs; |
|
48 val fv_eqs = maps fv_eq descr |
|
49 *} |
|
50 |
|
51 local_setup {* |
|
52 snd o (Primrec.add_primrec |
|
53 (map (fn s => (Binding.name s, NONE, NoSyn)) fv_names) fv_eqs) |
|
54 *} |
|
55 print_theorems |
|
56 |
|
57 fun |
|
58 fv_rlam :: "rlam \<Rightarrow> atom set" |
|
59 where |
|
60 "fv_rlam (rVar a) = {atom a}" |
|
61 | "fv_rlam (rApp t1 t2) = (fv_rlam t1) \<union> (fv_rlam t2)" |
|
62 | "fv_rlam (rLam a t) = (fv_rlam t) - {atom a}" |
|
63 |
|
64 end |