34 |
34 |
35 yields: |
35 yields: |
36 [ |
36 [ |
37 [], |
37 [], |
38 [[], [], [(NONE, 0)]], |
38 [[], [], [(NONE, 0)]], |
39 [[], [], [(SOME (Const f), 0), (Some (Const g), 1, 2)]]] |
39 [[], [], [(SOME (Const f), 0), (Some (Const g), 1)]]] |
40 *) |
40 *) |
41 |
41 |
|
42 ML foldr1 |
|
43 ML fold |
42 ML {* |
44 ML {* |
43 val {descr, ...} = Datatype.the_info @{theory} "Fv.rlam"; |
45 val {descr, ...} = Datatype.the_info @{theory} "Fv.rlam"; |
44 val sorts = []; |
46 val sorts = []; |
45 val binds = [[], [[], []], [[], [(NONE, 0)]]]; |
47 val noatoms = @{term "{} :: atom set"} |
|
48 val binds = [[[]], [[], []], [[], [(NONE, 0)]]]; |
46 fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i); |
49 fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i); |
47 val fv_names = Datatype_Prop.indexify_names (map (fn (i, _) => |
50 val fv_names = Datatype_Prop.indexify_names (map (fn (i, _) => |
48 "fv_" ^ name_of_typ (nth_dtyp i)) descr); |
51 "fv_" ^ name_of_typ (nth_dtyp i)) descr); |
49 val fv_types = map (fn (i, _) => nth_dtyp i --> @{typ "atom set"}) descr; |
52 val fv_types = map (fn (i, _) => nth_dtyp i --> @{typ "atom set"}) descr; |
50 val fv_frees = map Free (fv_names ~~ fv_types); |
53 val fv_frees = map Free (fv_names ~~ fv_types); |
51 (* TODO: should be optimized to avoid {}+{}+{} *) |
54 fun mk_union sets = |
52 fun mk_union sets = foldr1 (HOLogic.mk_binop @{const_name union}) sets; |
55 fold (fn a => fn b => |
53 |
56 if a = noatoms then b else |
54 fun fv_eq_constr i (cname, dts) bind = |
57 if b = noatoms then a else |
|
58 HOLogic.mk_binop @{const_name union} (a, b)) sets noatoms; |
|
59 fun mk_diff a b = |
|
60 if b = noatoms then a else |
|
61 HOLogic.mk_binop @{const_name "Algebras.minus_class.minus"} (a, b); |
|
62 fun fv_eq_constr i (cname, dts) bindcs = |
55 let |
63 let |
56 val Ts = map (typ_of_dtyp descr sorts) dts; |
64 val Ts = map (typ_of_dtyp descr sorts) dts; |
57 val names = Name.variant_list ["pi"] (Datatype_Prop.make_tnames Ts); |
65 val names = Name.variant_list ["pi"] (Datatype_Prop.make_tnames Ts); |
58 val args = map Free (names ~~ Ts); |
66 val args = map Free (names ~~ Ts); |
59 val c = Const (cname, Ts ---> (nth_dtyp i)); |
67 val c = Const (cname, Ts ---> (nth_dtyp i)); |
60 val fv_c = Free (nth fv_names i, (nth_dtyp i) --> @{typ "atom set"}); |
68 val fv_c = Free (nth fv_names i, (nth_dtyp i) --> @{typ "atom set"}); |
61 fun fv_arg (dt, x) = |
69 fun fv_bind (_, i) = @{term "{} :: atom set"} |
|
70 fun fv_arg ((dt, x), bindxs) = |
62 let |
71 let |
63 val arg = |
72 val arg = |
64 if is_rec_type dt then nth fv_frees (body_index dt) $ x |
73 if is_rec_type dt then nth fv_frees (body_index dt) $ x |
65 (* TODO: we just assume everything is a 'name' *) |
74 (* TODO: we just assume everything is a 'name' *) |
66 else HOLogic.mk_set @{typ atom} [mk_atom (type_of x) $ x] |
75 else HOLogic.mk_set @{typ atom} [mk_atom (type_of x) $ x] |
|
76 val sub = mk_union (map fv_bind bindxs) |
67 in |
77 in |
68 arg |
78 mk_diff arg sub |
69 end |
79 end; |
70 in |
80 val _ = tracing (string_of_int (length dts)); |
|
81 val _ = tracing (string_of_int (length bindcs)); |
|
82 in |
71 (Attrib.empty_binding, HOLogic.mk_Trueprop (HOLogic.mk_eq |
83 (Attrib.empty_binding, HOLogic.mk_Trueprop (HOLogic.mk_eq |
72 (fv_c $ list_comb (c, args), mk_union (map fv_arg (dts ~~ args))))) |
84 (fv_c $ list_comb (c, args), mk_union (map fv_arg (dts ~~ args ~~ bindcs))))) |
73 end; |
85 end; |
74 fun fv_eq (i, (_, _, constrs)) = map2 (fv_eq_constr i) constrs binds; |
86 fun fv_eq (i, (_, _, constrs)) = map2 (fv_eq_constr i) constrs binds; |
75 val fv_eqs = maps fv_eq descr |
87 val fv_eqs = maps fv_eq descr |
76 *} |
88 *} |
|
89 |
77 |
90 |
78 local_setup {* |
91 local_setup {* |
79 snd o (Primrec.add_primrec |
92 snd o (Primrec.add_primrec |
80 (map (fn s => (Binding.name s, NONE, NoSyn)) fv_names) fv_eqs) |
93 (map (fn s => (Binding.name s, NONE, NoSyn)) fv_names) fv_eqs) |
81 *} |
94 *} |