| author | wenzelm |
| Sun, 14 Oct 2001 20:02:59 +0200 | |
| changeset 11753 | 02b257ef0ee2 |
| parent 11539 | 0f17da240450 |
| child 11834 | 02825c735938 |
| permissions | -rw-r--r-- |
| 5719 | 1 |
(* Title: HOL/Tools/primrec_package.ML |
| 5178 | 2 |
ID: $Id$ |
| 11539 | 3 |
Author: Stefan Berghofer, TU Muenchen and Norbert Voelker, FernUni Hagen |
| 5178 | 4 |
|
| 6359 | 5 |
Package for defining functions on datatypes by primitive recursion. |
| 5178 | 6 |
*) |
7 |
||
8 |
signature PRIMREC_PACKAGE = |
|
9 |
sig |
|
| 6427 | 10 |
val quiet_mode: bool ref |
|
8480
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
11 |
val print_primrecs: theory -> unit |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
12 |
val get_primrec: theory -> string -> (string * thm list) list option |
| 6425 | 13 |
val add_primrec: string -> ((bstring * string) * Args.src list) list |
| 6359 | 14 |
-> theory -> theory * thm list |
| 6425 | 15 |
val add_primrec_i: string -> ((bstring * term) * theory attribute list) list |
| 6359 | 16 |
-> theory -> theory * thm list |
|
8480
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
17 |
val setup: (theory -> theory) list |
| 5178 | 18 |
end; |
19 |
||
20 |
structure PrimrecPackage : PRIMREC_PACKAGE = |
|
21 |
struct |
|
22 |
||
23 |
open DatatypeAux; |
|
24 |
||
25 |
exception RecError of string; |
|
26 |
||
27 |
fun primrec_err s = error ("Primrec definition error:\n" ^ s);
|
|
28 |
fun primrec_eq_err sign s eq = |
|
|
7544
dee529666dcd
Fixed bug in add_primrec which caused non-informative error message.
berghofe
parents:
7152
diff
changeset
|
29 |
primrec_err (s ^ "\nin\n" ^ quote (Sign.string_of_term sign eq)); |
| 6427 | 30 |
|
31 |
||
32 |
(* messages *) |
|
33 |
||
34 |
val quiet_mode = ref false; |
|
35 |
fun message s = if ! quiet_mode then () else writeln s; |
|
| 5178 | 36 |
|
| 6359 | 37 |
|
|
8480
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
38 |
(** theory data **) |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
39 |
|
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
40 |
(* data kind 'HOL/primrec' *) |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
41 |
|
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
42 |
structure PrimrecArgs = |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
43 |
struct |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
44 |
val name = "HOL/primrec"; |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
45 |
type T = (string * thm list) list Symtab.table; |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
46 |
|
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
47 |
val empty = Symtab.empty; |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
48 |
val copy = I; |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
49 |
val prep_ext = I; |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
50 |
val merge: T * T -> T = Symtab.merge (K true); |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
51 |
|
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
52 |
fun print sg tab = |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
53 |
Pretty.writeln (Pretty.strs ("primrecs:" ::
|
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
54 |
map #1 (Sign.cond_extern_table sg Sign.constK tab))); |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
55 |
end; |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
56 |
|
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
57 |
structure PrimrecData = TheoryDataFun(PrimrecArgs); |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
58 |
val print_primrecs = PrimrecData.print; |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
59 |
|
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
60 |
|
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
61 |
(* get and put data *) |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
62 |
|
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
63 |
fun get_primrec thy name = Symtab.lookup (PrimrecData.get thy, name); |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
64 |
|
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
65 |
fun put_primrec name info thy = |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
66 |
let val tab = PrimrecData.get thy |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
67 |
in |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
68 |
PrimrecData.put (case Symtab.lookup (tab, name) of |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
69 |
None => Symtab.update_new ((name, [info]), tab) |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
70 |
| Some infos => Symtab.update ((name, info::infos), tab)) thy end; |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
71 |
|
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
72 |
|
| 5178 | 73 |
(* preprocessing of equations *) |
74 |
||
75 |
fun process_eqn sign (eq, rec_fns) = |
|
76 |
let |
|
| 6032 | 77 |
val (lhs, rhs) = |
78 |
if null (term_vars eq) then |
|
79 |
HOLogic.dest_eq (HOLogic.dest_Trueprop eq) |
|
|
7016
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
80 |
handle TERM _ => raise RecError "not a proper equation" |
| 6032 | 81 |
else raise RecError "illegal schematic variable(s)"; |
| 5178 | 82 |
|
83 |
val (recfun, args) = strip_comb lhs; |
|
|
7016
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
84 |
val (fname, _) = dest_Const recfun handle TERM _ => |
| 5178 | 85 |
raise RecError "function is not declared as constant in theory"; |
86 |
||
87 |
val (ls', rest) = take_prefix is_Free args; |
|
88 |
val (middle, rs') = take_suffix is_Free rest; |
|
89 |
val rpos = length ls'; |
|
90 |
||
91 |
val (constr, cargs') = if null middle then raise RecError "constructor missing" |
|
92 |
else strip_comb (hd middle); |
|
93 |
val (cname, T) = dest_Const constr |
|
|
7016
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
94 |
handle TERM _ => raise RecError "ill-formed constructor"; |
|
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
95 |
val (tname, _) = dest_Type (body_type T) handle TYPE _ => |
| 5178 | 96 |
raise RecError "cannot determine datatype associated with function" |
97 |
||
| 6037 | 98 |
val (ls, cargs, rs) = (map dest_Free ls', |
99 |
map dest_Free cargs', |
|
100 |
map dest_Free rs') |
|
|
7016
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
101 |
handle TERM _ => raise RecError "illegal argument in pattern"; |
| 5178 | 102 |
val lfrees = ls @ rs @ cargs; |
103 |
||
| 8973 | 104 |
val _ = case duplicates lfrees of |
105 |
[] => () |
|
106 |
| vars => raise RecError("repeated variable names in pattern: " ^
|
|
107 |
commas_quote (map #1 vars)) |
|
108 |
||
109 |
val _ = case (map dest_Free (term_frees rhs)) \\ lfrees of |
|
110 |
[] => () |
|
111 |
| vars => raise RecError |
|
112 |
("extra variables on rhs: " ^
|
|
113 |
commas_quote (map #1 vars)) |
|
| 5178 | 114 |
in |
| 8973 | 115 |
if length middle > 1 then |
| 5178 | 116 |
raise RecError "more than one non-variable in pattern" |
117 |
else (case assoc (rec_fns, fname) of |
|
118 |
None => |
|
119 |
(fname, (tname, rpos, [(cname, (ls, cargs, rs, rhs, eq))]))::rec_fns |
|
120 |
| Some (_, rpos', eqns) => |
|
121 |
if is_some (assoc (eqns, cname)) then |
|
| 6037 | 122 |
raise RecError "constructor already occurred as pattern" |
| 5178 | 123 |
else if rpos <> rpos' then |
124 |
raise RecError "position of recursive argument inconsistent" |
|
125 |
else |
|
| 6037 | 126 |
overwrite (rec_fns, |
127 |
(fname, |
|
128 |
(tname, rpos, |
|
129 |
(cname, (ls, cargs, rs, rhs, eq))::eqns)))) |
|
| 5178 | 130 |
end |
131 |
handle RecError s => primrec_eq_err sign s eq; |
|
132 |
||
133 |
fun process_fun sign descr rec_eqns ((i, fname), (fnames, fnss)) = |
|
134 |
let |
|
135 |
val (_, (tname, _, constrs)) = nth_elem (i, descr); |
|
136 |
||
137 |
(* substitute "fname ls x rs" by "y ls rs" for (x, (_, y)) in subs *) |
|
138 |
||
139 |
fun subst [] x = x |
|
140 |
| subst subs (fs, Abs (a, T, t)) = |
|
141 |
let val (fs', t') = subst subs (fs, t) |
|
142 |
in (fs', Abs (a, T, t')) end |
|
143 |
| subst subs (fs, t as (_ $ _)) = |
|
144 |
let val (f, ts) = strip_comb t; |
|
145 |
in |
|
146 |
if is_Const f andalso (fst (dest_Const f)) mem (map fst rec_eqns) then |
|
147 |
let |
|
148 |
val (fname', _) = dest_Const f; |
|
149 |
val (_, rpos, _) = the (assoc (rec_eqns, fname')); |
|
150 |
val ls = take (rpos, ts); |
|
151 |
val rest = drop (rpos, ts); |
|
|
7016
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
152 |
val (x', rs) = (hd rest, tl rest) |
|
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
153 |
handle LIST _ => raise RecError ("not enough arguments\
|
|
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
154 |
\ in recursive application\nof function " ^ quote fname' ^ " on rhs"); |
|
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
155 |
val (x, xs) = strip_comb x' |
| 5178 | 156 |
in |
157 |
(case assoc (subs, x) of |
|
158 |
None => |
|
159 |
let |
|
160 |
val (fs', ts') = foldl_map (subst subs) (fs, ts) |
|
161 |
in (fs', list_comb (f, ts')) end |
|
162 |
| Some (i', y) => |
|
163 |
let |
|
|
7016
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
164 |
val (fs', ts') = foldl_map (subst subs) (fs, xs @ ls @ rs); |
| 5178 | 165 |
val fs'' = process_fun sign descr rec_eqns ((i', fname'), fs') |
166 |
in (fs'', list_comb (y, ts')) |
|
167 |
end) |
|
168 |
end |
|
169 |
else |
|
170 |
let |
|
171 |
val (fs', f'::ts') = foldl_map (subst subs) (fs, f::ts) |
|
172 |
in (fs', list_comb (f', ts')) end |
|
173 |
end |
|
174 |
| subst _ x = x; |
|
175 |
||
176 |
(* translate rec equations into function arguments suitable for rec comb *) |
|
177 |
||
178 |
fun trans eqns ((cname, cargs), (fnames', fnss', fns)) = |
|
179 |
(case assoc (eqns, cname) of |
|
| 6427 | 180 |
None => (warning ("no equation for constructor " ^ quote cname ^
|
181 |
"\nin definition of function " ^ quote fname); |
|
| 5178 | 182 |
(fnames', fnss', (Const ("arbitrary", dummyT))::fns))
|
183 |
| Some (ls, cargs', rs, rhs, eq) => |
|
184 |
let |
|
|
7016
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
185 |
fun rec_index (DtRec k) = k |
|
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
186 |
| rec_index (DtType ("fun", [_, DtRec k])) = k;
|
|
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
187 |
|
| 5178 | 188 |
val recs = filter (is_rec_type o snd) (cargs' ~~ cargs); |
189 |
val rargs = map fst recs; |
|
| 6037 | 190 |
val subs = map (rpair dummyT o fst) |
191 |
(rev (rename_wrt_term rhs rargs)); |
|
192 |
val ((fnames'', fnss''), rhs') = |
|
193 |
(subst (map (fn ((x, y), z) => |
|
|
7016
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
194 |
(Free x, (rec_index y, Free z))) |
| 6037 | 195 |
(recs ~~ subs)) |
196 |
((fnames', fnss'), rhs)) |
|
| 5178 | 197 |
handle RecError s => primrec_eq_err sign s eq |
| 6037 | 198 |
in (fnames'', fnss'', |
199 |
(list_abs_free (cargs' @ subs @ ls @ rs, rhs'))::fns) |
|
| 5178 | 200 |
end) |
201 |
||
202 |
in (case assoc (fnames, i) of |
|
203 |
None => |
|
204 |
if exists (equal fname o snd) fnames then |
|
| 6427 | 205 |
raise RecError ("inconsistent functions for datatype " ^ quote tname)
|
| 5178 | 206 |
else |
207 |
let |
|
208 |
val (_, _, eqns) = the (assoc (rec_eqns, fname)); |
|
209 |
val (fnames', fnss', fns) = foldr (trans eqns) |
|
210 |
(constrs, ((i, fname)::fnames, fnss, [])) |
|
211 |
in |
|
212 |
(fnames', (i, (fname, #1 (snd (hd eqns)), fns))::fnss') |
|
213 |
end |
|
214 |
| Some fname' => |
|
215 |
if fname = fname' then (fnames, fnss) |
|
| 6427 | 216 |
else raise RecError ("inconsistent functions for datatype " ^ quote tname))
|
| 5178 | 217 |
end; |
218 |
||
| 6359 | 219 |
|
| 5178 | 220 |
(* prepare functions needed for definitions *) |
221 |
||
222 |
fun get_fns fns (((i, (tname, _, constrs)), rec_name), (fs, defs)) = |
|
223 |
case assoc (fns, i) of |
|
224 |
None => |
|
225 |
let |
|
226 |
val dummy_fns = map (fn (_, cargs) => Const ("arbitrary",
|
|
227 |
replicate ((length cargs) + (length (filter is_rec_type cargs))) |
|
228 |
dummyT ---> HOLogic.unitT)) constrs; |
|
| 6427 | 229 |
val _ = warning ("No function definition for datatype " ^ quote tname)
|
| 5178 | 230 |
in |
231 |
(dummy_fns @ fs, defs) |
|
232 |
end |
|
233 |
| Some (fname, ls, fs') => (fs' @ fs, (fname, ls, rec_name, tname)::defs); |
|
234 |
||
| 6359 | 235 |
|
| 5178 | 236 |
(* make definition *) |
237 |
||
238 |
fun make_def sign fs (fname, ls, rec_name, tname) = |
|
239 |
let |
|
| 6037 | 240 |
val rhs = foldr (fn (T, t) => Abs ("", T, t))
|
241 |
((map snd ls) @ [dummyT], |
|
242 |
list_comb (Const (rec_name, dummyT), |
|
243 |
fs @ map Bound (0 ::(length ls downto 1)))); |
|
| 5178 | 244 |
val defpair = (Sign.base_name fname ^ "_" ^ Sign.base_name tname ^ "_def", |
| 6037 | 245 |
Logic.mk_equals (Const (fname, dummyT), rhs)) |
| 6394 | 246 |
in Theory.inferT_axm sign defpair end; |
| 5178 | 247 |
|
| 6359 | 248 |
|
| 5178 | 249 |
(* find datatypes which contain all datatypes in tnames' *) |
250 |
||
251 |
fun find_dts (dt_info : datatype_info Symtab.table) _ [] = [] |
|
252 |
| find_dts dt_info tnames' (tname::tnames) = |
|
253 |
(case Symtab.lookup (dt_info, tname) of |
|
| 6427 | 254 |
None => primrec_err (quote tname ^ " is not a datatype") |
| 5178 | 255 |
| Some dt => |
256 |
if tnames' subset (map (#1 o snd) (#descr dt)) then |
|
257 |
(tname, dt)::(find_dts dt_info tnames' tnames) |
|
258 |
else find_dts dt_info tnames' tnames); |
|
259 |
||
| 8432 | 260 |
fun prepare_induct ({descr, induction, ...}: datatype_info) rec_eqns =
|
261 |
let |
|
262 |
fun constrs_of (_, (_, _, cs)) = |
|
263 |
map (fn (cname:string, (_, cargs, _, _, _)) => (cname, map fst cargs)) cs; |
|
264 |
val params_of = Library.assocs (flat (map constrs_of rec_eqns)); |
|
265 |
in |
|
266 |
induction |
|
267 |
|> RuleCases.rename_params (map params_of (flat (map (map #1 o #3 o #2) descr))) |
|
| 10525 | 268 |
|> RuleCases.save induction |
| 8432 | 269 |
end; |
270 |
||
| 6359 | 271 |
fun add_primrec_i alt_name eqns_atts thy = |
| 5178 | 272 |
let |
| 6359 | 273 |
val (eqns, atts) = split_list eqns_atts; |
| 6394 | 274 |
val sg = Theory.sign_of thy; |
| 5178 | 275 |
val dt_info = DatatypePackage.get_datatypes thy; |
| 5719 | 276 |
val rec_eqns = foldr (process_eqn sg) (map snd eqns, []); |
| 5178 | 277 |
val tnames = distinct (map (#1 o snd) rec_eqns); |
278 |
val dts = find_dts dt_info tnames tnames; |
|
| 6037 | 279 |
val main_fns = |
280 |
map (fn (tname, {index, ...}) =>
|
|
281 |
(index, |
|
282 |
fst (the (find_first (fn f => #1 (snd f) = tname) rec_eqns)))) |
|
283 |
dts; |
|
284 |
val {descr, rec_names, rec_rewrites, ...} =
|
|
285 |
if null dts then |
|
| 6359 | 286 |
primrec_err ("datatypes " ^ commas_quote tnames ^
|
| 6037 | 287 |
"\nare not mutually recursive") |
288 |
else snd (hd dts); |
|
289 |
val (fnames, fnss) = foldr (process_fun sg descr rec_eqns) |
|
290 |
(main_fns, ([], [])); |
|
| 5178 | 291 |
val (fs, defs) = foldr (get_fns fnss) (descr ~~ rec_names, ([], [])); |
292 |
val defs' = map (make_def sg fs) defs; |
|
293 |
val names1 = map snd fnames; |
|
294 |
val names2 = map fst rec_eqns; |
|
| 8432 | 295 |
val primrec_name = |
296 |
if alt_name = "" then (space_implode "_" (map (Sign.base_name o #1) defs)) else alt_name; |
|
297 |
val (thy', defs_thms') = thy |> Theory.add_path primrec_name |> |
|
| 9315 | 298 |
(if eq_set (names1, names2) then (PureThy.add_defs_i false o map Thm.no_attributes) defs' |
| 6359 | 299 |
else primrec_err ("functions " ^ commas_quote names2 ^
|
| 5216 | 300 |
"\nare not mutually recursive")); |
| 8432 | 301 |
val rewrites = o_def :: (map mk_meta_eq rec_rewrites) @ defs_thms'; |
| 6427 | 302 |
val _ = message ("Proving equations for primrec function(s) " ^ commas_quote names1 ^ " ...");
|
|
8480
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
303 |
val simps = map (fn (_, t) => prove_goalw_cterm rewrites (cterm_of (Theory.sign_of thy') t) |
| 5178 | 304 |
(fn _ => [rtac refl 1])) eqns; |
|
9575
af71f5f4ca6b
Equations that are added to the simpset now have proper names.
berghofe
parents:
9315
diff
changeset
|
305 |
val (thy'', simps') = PureThy.add_thms ((map fst eqns ~~ simps) ~~ atts) thy'; |
|
af71f5f4ca6b
Equations that are added to the simpset now have proper names.
berghofe
parents:
9315
diff
changeset
|
306 |
val thy''' = thy'' |
|
af71f5f4ca6b
Equations that are added to the simpset now have proper names.
berghofe
parents:
9315
diff
changeset
|
307 |
|> (#1 o PureThy.add_thmss [(("simps", simps'), [Simplifier.simp_add_global])])
|
|
af71f5f4ca6b
Equations that are added to the simpset now have proper names.
berghofe
parents:
9315
diff
changeset
|
308 |
|> (#1 o PureThy.add_thms [(("induct", prepare_induct (#2 (hd dts)) rec_eqns), [])])
|
|
af71f5f4ca6b
Equations that are added to the simpset now have proper names.
berghofe
parents:
9315
diff
changeset
|
309 |
|> Theory.parent_path |
|
8480
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
310 |
in |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
311 |
(foldl (fn (thy, (fname, _, _, tname)) => |
|
9575
af71f5f4ca6b
Equations that are added to the simpset now have proper names.
berghofe
parents:
9315
diff
changeset
|
312 |
put_primrec fname (tname, simps') thy) (thy''', defs), simps') |
|
8480
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
313 |
end; |
| 6359 | 314 |
|
315 |
||
|
7016
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
316 |
fun add_primrec alt_name eqns thy = |
|
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
317 |
let |
|
7544
dee529666dcd
Fixed bug in add_primrec which caused non-informative error message.
berghofe
parents:
7152
diff
changeset
|
318 |
val sign = Theory.sign_of thy; |
|
7016
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
319 |
val ((names, strings), srcss) = apfst split_list (split_list eqns); |
|
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
320 |
val atts = map (map (Attrib.global_attribute thy)) srcss; |
| 7722 | 321 |
val eqn_ts = map (term_of o Thm.read_cterm sign o rpair propT) strings; |
|
7016
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
322 |
val rec_ts = map (fn eq => head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop eq))) |
|
7544
dee529666dcd
Fixed bug in add_primrec which caused non-informative error message.
berghofe
parents:
7152
diff
changeset
|
323 |
handle TERM _ => primrec_eq_err sign "not a proper equation" eq) eqn_ts; |
|
7016
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
324 |
val (_, eqn_ts') = InductivePackage.unify_consts (sign_of thy) rec_ts eqn_ts |
|
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
325 |
in |
|
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
326 |
add_primrec_i alt_name (names ~~ eqn_ts' ~~ atts) thy |
|
df54b5365477
- Now also supports arbitrarily branching datatypes.
berghofe
parents:
6729
diff
changeset
|
327 |
end; |
| 6359 | 328 |
|
| 5178 | 329 |
|
|
8480
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
330 |
(** package setup **) |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
331 |
|
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
332 |
val setup = [PrimrecData.init]; |
|
50266d517b0c
Added new theory data slot for primrec equations.
berghofe
parents:
8432
diff
changeset
|
333 |
|
| 6359 | 334 |
(* outer syntax *) |
335 |
||
| 6723 | 336 |
local structure P = OuterParse and K = OuterSyntax.Keyword in |
| 6359 | 337 |
|
338 |
val primrec_decl = |
|
| 6723 | 339 |
Scan.optional (P.$$$ "(" |-- P.name --| P.$$$ ")") "" --
|
| 7152 | 340 |
Scan.repeat1 (P.opt_thm_name ":" -- P.prop --| P.marg_comment); |
| 6359 | 341 |
|
342 |
val primrecP = |
|
| 6723 | 343 |
OuterSyntax.command "primrec" "define primitive recursive functions on datatypes" K.thy_decl |
| 6359 | 344 |
(primrec_decl >> (fn (alt_name, eqns) => |
| 6723 | 345 |
Toplevel.theory (#1 o add_primrec alt_name (map P.triple_swap eqns)))); |
| 6359 | 346 |
|
347 |
val _ = OuterSyntax.add_parsers [primrecP]; |
|
| 5178 | 348 |
|
349 |
end; |
|
| 6384 | 350 |
|
351 |
||
352 |
end; |