| author | nipkow | 
| Mon, 10 Apr 2006 11:35:02 +0200 | |
| changeset 19402 | 742b7934ccfc | 
| parent 19046 | bc5c6c9b114e | 
| child 19636 | 50515882049e | 
| 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 | 
| 15703 | 11 | val add_primrec: string -> ((bstring * string) * Attrib.src list) list | 
| 6359 | 12 | -> theory -> theory * thm list | 
| 18728 | 13 | val add_primrec_i: string -> ((bstring * term) * attribute list) list | 
| 6359 | 14 | -> theory -> theory * thm list | 
| 5178 | 15 | end; | 
| 16 | ||
| 17 | structure PrimrecPackage : PRIMREC_PACKAGE = | |
| 18 | struct | |
| 19 | ||
| 20 | open DatatypeAux; | |
| 21 | ||
| 22 | exception RecError of string; | |
| 23 | ||
| 24 | fun primrec_err s = error ("Primrec definition error:\n" ^ s);
 | |
| 25 | fun primrec_eq_err sign s eq = | |
| 7544 
dee529666dcd
Fixed bug in add_primrec which caused non-informative error message.
 berghofe parents: 
7152diff
changeset | 26 | primrec_err (s ^ "\nin\n" ^ quote (Sign.string_of_term sign eq)); | 
| 6427 | 27 | |
| 28 | ||
| 29 | (* messages *) | |
| 30 | ||
| 31 | val quiet_mode = ref false; | |
| 32 | fun message s = if ! quiet_mode then () else writeln s; | |
| 5178 | 33 | |
| 6359 | 34 | |
| 5178 | 35 | (* preprocessing of equations *) | 
| 36 | ||
| 37 | fun process_eqn sign (eq, rec_fns) = | |
| 38 | let | |
| 6032 | 39 | val (lhs, rhs) = | 
| 40 | if null (term_vars eq) then | |
| 41 | HOLogic.dest_eq (HOLogic.dest_Trueprop eq) | |
| 7016 
df54b5365477
- Now also supports arbitrarily branching datatypes.
 berghofe parents: 
6729diff
changeset | 42 | handle TERM _ => raise RecError "not a proper equation" | 
| 6032 | 43 | else raise RecError "illegal schematic variable(s)"; | 
| 5178 | 44 | |
| 45 | val (recfun, args) = strip_comb lhs; | |
| 16765 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 46 | val fnameT = dest_Const recfun handle TERM _ => | 
| 5178 | 47 | raise RecError "function is not declared as constant in theory"; | 
| 48 | ||
| 49 | val (ls', rest) = take_prefix is_Free args; | |
| 50 | val (middle, rs') = take_suffix is_Free rest; | |
| 51 | val rpos = length ls'; | |
| 52 | ||
| 53 | val (constr, cargs') = if null middle then raise RecError "constructor missing" | |
| 54 | else strip_comb (hd middle); | |
| 55 | val (cname, T) = dest_Const constr | |
| 7016 
df54b5365477
- Now also supports arbitrarily branching datatypes.
 berghofe parents: 
6729diff
changeset | 56 | handle TERM _ => raise RecError "ill-formed constructor"; | 
| 
df54b5365477
- Now also supports arbitrarily branching datatypes.
 berghofe parents: 
6729diff
changeset | 57 | val (tname, _) = dest_Type (body_type T) handle TYPE _ => | 
| 5178 | 58 | raise RecError "cannot determine datatype associated with function" | 
| 59 | ||
| 6037 | 60 | val (ls, cargs, rs) = (map dest_Free ls', | 
| 61 | map dest_Free cargs', | |
| 62 | map dest_Free rs') | |
| 7016 
df54b5365477
- Now also supports arbitrarily branching datatypes.
 berghofe parents: 
6729diff
changeset | 63 | handle TERM _ => raise RecError "illegal argument in pattern"; | 
| 5178 | 64 | val lfrees = ls @ rs @ cargs; | 
| 65 | ||
| 12474 | 66 | fun check_vars _ [] = () | 
| 67 | | check_vars s vars = raise RecError (s ^ commas_quote (map fst vars)) | |
| 5178 | 68 | in | 
| 8973 | 69 | if length middle > 1 then | 
| 5178 | 70 | raise RecError "more than one non-variable in pattern" | 
| 12474 | 71 | else | 
| 18964 | 72 | (check_vars "repeated variable names in pattern: " (duplicates (op =) lfrees); | 
| 12474 | 73 | check_vars "extra variables on rhs: " | 
| 74 | (map dest_Free (term_frees rhs) \\ lfrees); | |
| 17184 | 75 | case AList.lookup (op =) rec_fns fnameT of | 
| 15531 | 76 | NONE => | 
| 16765 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 77 | (fnameT, (tname, rpos, [(cname, (ls, cargs, rs, rhs, eq))]))::rec_fns | 
| 15531 | 78 | | SOME (_, rpos', eqns) => | 
| 17184 | 79 | if AList.defined (op =) eqns cname then | 
| 6037 | 80 | raise RecError "constructor already occurred as pattern" | 
| 5178 | 81 | else if rpos <> rpos' then | 
| 82 | raise RecError "position of recursive argument inconsistent" | |
| 83 | else | |
| 17314 | 84 | AList.update (op =) (fnameT, (tname, rpos, (cname, (ls, cargs, rs, rhs, eq))::eqns)) | 
| 85 | rec_fns) | |
| 5178 | 86 | end | 
| 87 | handle RecError s => primrec_eq_err sign s eq; | |
| 88 | ||
| 16765 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 89 | fun process_fun sign descr rec_eqns ((i, fnameT as (fname, _)), (fnameTs, fnss)) = | 
| 5178 | 90 | let | 
| 15570 | 91 | val (_, (tname, _, constrs)) = List.nth (descr, i); | 
| 5178 | 92 | |
| 93 | (* substitute "fname ls x rs" by "y ls rs" for (x, (_, y)) in subs *) | |
| 94 | ||
| 95 | fun subst [] x = x | |
| 96 | | subst subs (fs, Abs (a, T, t)) = | |
| 97 | let val (fs', t') = subst subs (fs, t) | |
| 98 | in (fs', Abs (a, T, t')) end | |
| 99 | | subst subs (fs, t as (_ $ _)) = | |
| 100 | let val (f, ts) = strip_comb t; | |
| 101 | in | |
| 16765 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 102 | if is_Const f andalso dest_Const f mem map fst rec_eqns then | 
| 5178 | 103 | let | 
| 16765 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 104 | val fnameT' as (fname', _) = dest_Const f; | 
| 17184 | 105 | val (_, rpos, _) = the (AList.lookup (op =) rec_eqns fnameT'); | 
| 15570 | 106 | val ls = Library.take (rpos, ts); | 
| 107 | val rest = Library.drop (rpos, ts); | |
| 7016 
df54b5365477
- Now also supports arbitrarily branching datatypes.
 berghofe parents: 
6729diff
changeset | 108 | val (x', rs) = (hd rest, tl rest) | 
| 15570 | 109 |                   handle Empty => raise RecError ("not enough arguments\
 | 
| 7016 
df54b5365477
- Now also supports arbitrarily branching datatypes.
 berghofe parents: 
6729diff
changeset | 110 | \ in recursive application\nof function " ^ quote fname' ^ " on rhs"); | 
| 
df54b5365477
- Now also supports arbitrarily branching datatypes.
 berghofe parents: 
6729diff
changeset | 111 | val (x, xs) = strip_comb x' | 
| 5178 | 112 | in | 
| 17184 | 113 | (case AList.lookup (op =) subs x of | 
| 15531 | 114 | NONE => | 
| 5178 | 115 | let | 
| 116 | val (fs', ts') = foldl_map (subst subs) (fs, ts) | |
| 117 | in (fs', list_comb (f, ts')) end | |
| 15531 | 118 | | SOME (i', y) => | 
| 5178 | 119 | let | 
| 7016 
df54b5365477
- Now also supports arbitrarily branching datatypes.
 berghofe parents: 
6729diff
changeset | 120 | val (fs', ts') = foldl_map (subst subs) (fs, xs @ ls @ rs); | 
| 16765 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 121 | val fs'' = process_fun sign descr rec_eqns ((i', fnameT'), fs') | 
| 5178 | 122 | in (fs'', list_comb (y, ts')) | 
| 123 | end) | |
| 124 | end | |
| 125 | else | |
| 126 | let | |
| 127 | val (fs', f'::ts') = foldl_map (subst subs) (fs, f::ts) | |
| 128 | in (fs', list_comb (f', ts')) end | |
| 129 | end | |
| 130 | | subst _ x = x; | |
| 131 | ||
| 132 | (* translate rec equations into function arguments suitable for rec comb *) | |
| 133 | ||
| 16765 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 134 | fun trans eqns ((cname, cargs), (fnameTs', fnss', fns)) = | 
| 17184 | 135 | (case AList.lookup (op =) eqns cname of | 
| 15531 | 136 |           NONE => (warning ("No equation for constructor " ^ quote cname ^
 | 
| 6427 | 137 | "\nin definition of function " ^ quote fname); | 
| 16765 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 138 |               (fnameTs', fnss', (Const ("arbitrary", dummyT))::fns))
 | 
| 15531 | 139 | | SOME (ls, cargs', rs, rhs, eq) => | 
| 5178 | 140 | let | 
| 15570 | 141 | val recs = List.filter (is_rec_type o snd) (cargs' ~~ cargs); | 
| 5178 | 142 | val rargs = map fst recs; | 
| 6037 | 143 | val subs = map (rpair dummyT o fst) | 
| 144 | (rev (rename_wrt_term rhs rargs)); | |
| 16765 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 145 | val ((fnameTs'', fnss''), rhs') = | 
| 6037 | 146 | (subst (map (fn ((x, y), z) => | 
| 13641 
63d1790a43ed
Reimplemented parts of datatype package dealing with datatypes involving
 berghofe parents: 
12876diff
changeset | 147 | (Free x, (body_index y, Free z))) | 
| 6037 | 148 | (recs ~~ subs)) | 
| 16765 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 149 | ((fnameTs', fnss'), rhs)) | 
| 5178 | 150 | handle RecError s => primrec_eq_err sign s eq | 
| 16765 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 151 | in (fnameTs'', fnss'', | 
| 6037 | 152 | (list_abs_free (cargs' @ subs @ ls @ rs, rhs'))::fns) | 
| 5178 | 153 | end) | 
| 154 | ||
| 17184 | 155 | in (case AList.lookup (op =) fnameTs i of | 
| 15531 | 156 | NONE => | 
| 16765 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 157 | if exists (equal fnameT o snd) fnameTs then | 
| 6427 | 158 |           raise RecError ("inconsistent functions for datatype " ^ quote tname)
 | 
| 5178 | 159 | else | 
| 160 | let | |
| 17184 | 161 | val (_, _, eqns) = the (AList.lookup (op =) rec_eqns fnameT); | 
| 16765 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 162 | val (fnameTs', fnss', fns) = foldr (trans eqns) | 
| 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 163 | ((i, fnameT)::fnameTs, fnss, []) constrs | 
| 5178 | 164 | in | 
| 16765 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 165 | (fnameTs', (i, (fname, #1 (snd (hd eqns)), fns))::fnss') | 
| 5178 | 166 | end | 
| 16765 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 167 | | SOME fnameT' => | 
| 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 168 | if fnameT = fnameT' then (fnameTs, fnss) | 
| 6427 | 169 |         else raise RecError ("inconsistent functions for datatype " ^ quote tname))
 | 
| 5178 | 170 | end; | 
| 171 | ||
| 6359 | 172 | |
| 5178 | 173 | (* prepare functions needed for definitions *) | 
| 174 | ||
| 175 | fun get_fns fns (((i, (tname, _, constrs)), rec_name), (fs, defs)) = | |
| 17184 | 176 | case AList.lookup (op =) fns i of | 
| 15531 | 177 | NONE => | 
| 5178 | 178 | let | 
| 179 |          val dummy_fns = map (fn (_, cargs) => Const ("arbitrary",
 | |
| 15570 | 180 | replicate ((length cargs) + (length (List.filter is_rec_type cargs))) | 
| 5178 | 181 | dummyT ---> HOLogic.unitT)) constrs; | 
| 6427 | 182 |          val _ = warning ("No function definition for datatype " ^ quote tname)
 | 
| 5178 | 183 | in | 
| 184 | (dummy_fns @ fs, defs) | |
| 185 | end | |
| 15531 | 186 | | SOME (fname, ls, fs') => (fs' @ fs, (fname, ls, rec_name, tname)::defs); | 
| 5178 | 187 | |
| 6359 | 188 | |
| 5178 | 189 | (* make definition *) | 
| 190 | ||
| 191 | fun make_def sign fs (fname, ls, rec_name, tname) = | |
| 192 | let | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 193 |     val rhs = foldr (fn (T, t) => Abs ("", T, t)) 
 | 
| 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 194 | (list_comb (Const (rec_name, dummyT), | 
| 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 195 | fs @ map Bound (0 ::(length ls downto 1)))) | 
| 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 196 | ((map snd ls) @ [dummyT]); | 
| 5178 | 197 | val defpair = (Sign.base_name fname ^ "_" ^ Sign.base_name tname ^ "_def", | 
| 6037 | 198 | Logic.mk_equals (Const (fname, dummyT), rhs)) | 
| 6394 | 199 | in Theory.inferT_axm sign defpair end; | 
| 5178 | 200 | |
| 6359 | 201 | |
| 5178 | 202 | (* find datatypes which contain all datatypes in tnames' *) | 
| 203 | ||
| 204 | fun find_dts (dt_info : datatype_info Symtab.table) _ [] = [] | |
| 205 | | find_dts dt_info tnames' (tname::tnames) = | |
| 17412 | 206 | (case Symtab.lookup dt_info tname of | 
| 15531 | 207 | NONE => primrec_err (quote tname ^ " is not a datatype") | 
| 208 | | SOME dt => | |
| 5178 | 209 | if tnames' subset (map (#1 o snd) (#descr dt)) then | 
| 210 | (tname, dt)::(find_dts dt_info tnames' tnames) | |
| 211 | else find_dts dt_info tnames' tnames); | |
| 212 | ||
| 8432 | 213 | fun prepare_induct ({descr, induction, ...}: datatype_info) rec_eqns =
 | 
| 214 | let | |
| 215 | fun constrs_of (_, (_, _, cs)) = | |
| 216 | map (fn (cname:string, (_, cargs, _, _, _)) => (cname, map fst cargs)) cs; | |
| 17184 | 217 | val params_of = these o AList.lookup (op =) (List.concat (map constrs_of rec_eqns)); | 
| 8432 | 218 | in | 
| 219 | induction | |
| 15570 | 220 | |> RuleCases.rename_params (map params_of (List.concat (map (map #1 o #3 o #2) descr))) | 
| 10525 | 221 | |> RuleCases.save induction | 
| 8432 | 222 | end; | 
| 223 | ||
| 6359 | 224 | fun add_primrec_i alt_name eqns_atts thy = | 
| 5178 | 225 | let | 
| 6359 | 226 | val (eqns, atts) = split_list eqns_atts; | 
| 6394 | 227 | val sg = Theory.sign_of thy; | 
| 5178 | 228 | val dt_info = DatatypePackage.get_datatypes thy; | 
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 229 | val rec_eqns = foldr (process_eqn sg) [] (map snd eqns); | 
| 19046 
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
 wenzelm parents: 
18964diff
changeset | 230 | val tnames = distinct (op =) (map (#1 o snd) rec_eqns); | 
| 5178 | 231 | val dts = find_dts dt_info tnames tnames; | 
| 6037 | 232 | val main_fns = | 
| 18362 
e8b7e0a22727
removed thms 'swap' and 'nth_map' from ML toplevel
 haftmann parents: 
18358diff
changeset | 233 |       map (fn (tname, {index, ...}) =>
 | 
| 
e8b7e0a22727
removed thms 'swap' and 'nth_map' from ML toplevel
 haftmann parents: 
18358diff
changeset | 234 | (index, | 
| 
e8b7e0a22727
removed thms 'swap' and 'nth_map' from ML toplevel
 haftmann parents: 
18358diff
changeset | 235 | (fst o the o find_first (fn f => (#1 o snd) f = tname)) rec_eqns)) | 
| 
e8b7e0a22727
removed thms 'swap' and 'nth_map' from ML toplevel
 haftmann parents: 
18358diff
changeset | 236 | dts; | 
| 6037 | 237 |     val {descr, rec_names, rec_rewrites, ...} = 
 | 
| 18362 
e8b7e0a22727
removed thms 'swap' and 'nth_map' from ML toplevel
 haftmann parents: 
18358diff
changeset | 238 | if null dts then | 
| 
e8b7e0a22727
removed thms 'swap' and 'nth_map' from ML toplevel
 haftmann parents: 
18358diff
changeset | 239 |         primrec_err ("datatypes " ^ commas_quote tnames ^ "\nare not mutually recursive")
 | 
| 
e8b7e0a22727
removed thms 'swap' and 'nth_map' from ML toplevel
 haftmann parents: 
18358diff
changeset | 240 | else snd (hd dts); | 
| 16765 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 241 | val (fnameTs, fnss) = foldr (process_fun sg descr rec_eqns) | 
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 242 | ([], []) main_fns; | 
| 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 243 | val (fs, defs) = foldr (get_fns fnss) ([], []) (descr ~~ rec_names); | 
| 5178 | 244 | val defs' = map (make_def sg fs) defs; | 
| 16765 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 245 | val nameTs1 = map snd fnameTs; | 
| 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 246 | val nameTs2 = map fst rec_eqns; | 
| 8432 | 247 | val primrec_name = | 
| 248 | if alt_name = "" then (space_implode "_" (map (Sign.base_name o #1) defs)) else alt_name; | |
| 18358 | 249 | val (defs_thms', thy') = thy |> Theory.add_path primrec_name |> | 
| 16765 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 250 | (if eq_set (nameTs1, nameTs2) then (PureThy.add_defs_i false o map Thm.no_attributes) defs' | 
| 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 251 |        else primrec_err ("functions " ^ commas_quote (map fst nameTs2) ^
 | 
| 5216 | 252 | "\nare not mutually recursive")); | 
| 13641 
63d1790a43ed
Reimplemented parts of datatype package dealing with datatypes involving
 berghofe parents: 
12876diff
changeset | 253 | val rewrites = (map mk_meta_eq rec_rewrites) @ defs_thms'; | 
| 16765 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 254 |     val _ = message ("Proving equations for primrec function(s) " ^
 | 
| 
b8b1f310877f
Some changes to allow mutually recursive, overloaded functions with same name.
 berghofe parents: 
16646diff
changeset | 255 | commas_quote (map fst nameTs1) ^ " ..."); | 
| 17985 | 256 | val simps = map (fn (_, t) => standard (Goal.prove thy' [] [] t | 
| 257 | (fn _ => EVERY [rewrite_goals_tac rewrites, rtac refl 1]))) eqns; | |
| 18377 | 258 | val (simps', thy'') = PureThy.add_thms ((map fst eqns ~~ simps) ~~ atts) thy'; | 
| 9575 
af71f5f4ca6b
Equations that are added to the simpset now have proper names.
 berghofe parents: 
9315diff
changeset | 259 | val thy''' = thy'' | 
| 18377 | 260 |       |> (snd o PureThy.add_thmss [(("simps", simps'),
 | 
| 18728 | 261 | [Simplifier.simp_add, RecfunCodegen.add NONE])]) | 
| 18377 | 262 |       |> (snd o PureThy.add_thms [(("induct", prepare_induct (#2 (hd dts)) rec_eqns), [])])
 | 
| 9575 
af71f5f4ca6b
Equations that are added to the simpset now have proper names.
 berghofe parents: 
9315diff
changeset | 263 | |> Theory.parent_path | 
| 8480 
50266d517b0c
Added new theory data slot for primrec equations.
 berghofe parents: 
8432diff
changeset | 264 | in | 
| 12448 
473cb9f9e237
Recursive equations to be used for code generation are now registered
 berghofe parents: 
12311diff
changeset | 265 | (thy''', simps') | 
| 8480 
50266d517b0c
Added new theory data slot for primrec equations.
 berghofe parents: 
8432diff
changeset | 266 | end; | 
| 6359 | 267 | |
| 268 | ||
| 7016 
df54b5365477
- Now also supports arbitrarily branching datatypes.
 berghofe parents: 
6729diff
changeset | 269 | fun add_primrec alt_name eqns thy = | 
| 
df54b5365477
- Now also supports arbitrarily branching datatypes.
 berghofe parents: 
6729diff
changeset | 270 | let | 
| 7544 
dee529666dcd
Fixed bug in add_primrec which caused non-informative error message.
 berghofe parents: 
7152diff
changeset | 271 | val sign = Theory.sign_of thy; | 
| 7016 
df54b5365477
- Now also supports arbitrarily branching datatypes.
 berghofe parents: 
6729diff
changeset | 272 | val ((names, strings), srcss) = apfst split_list (split_list eqns); | 
| 18728 | 273 | val atts = map (map (Attrib.attribute thy)) srcss; | 
| 14258 
9bd184c007f0
Improved error handling: add_primrec now prints out ill-formed equation
 berghofe parents: 
13641diff
changeset | 274 | val eqn_ts = map (fn s => term_of (Thm.read_cterm sign (s, propT)) | 
| 18678 | 275 |       handle ERROR msg => cat_error msg ("The error(s) above occurred for " ^ s)) strings;
 | 
| 7016 
df54b5365477
- Now also supports arbitrarily branching datatypes.
 berghofe parents: 
6729diff
changeset | 276 | 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: 
7152diff
changeset | 277 | handle TERM _ => primrec_eq_err sign "not a proper equation" eq) eqn_ts; | 
| 7016 
df54b5365477
- Now also supports arbitrarily branching datatypes.
 berghofe parents: 
6729diff
changeset | 278 | val (_, eqn_ts') = InductivePackage.unify_consts (sign_of thy) rec_ts eqn_ts | 
| 
df54b5365477
- Now also supports arbitrarily branching datatypes.
 berghofe parents: 
6729diff
changeset | 279 | in | 
| 
df54b5365477
- Now also supports arbitrarily branching datatypes.
 berghofe parents: 
6729diff
changeset | 280 | add_primrec_i alt_name (names ~~ eqn_ts' ~~ atts) thy | 
| 
df54b5365477
- Now also supports arbitrarily branching datatypes.
 berghofe parents: 
6729diff
changeset | 281 | end; | 
| 6359 | 282 | |
| 5178 | 283 | |
| 6359 | 284 | (* outer syntax *) | 
| 285 | ||
| 17057 | 286 | local structure P = OuterParse and K = OuterKeyword in | 
| 6359 | 287 | |
| 288 | val primrec_decl = | |
| 6723 | 289 |   Scan.optional (P.$$$ "(" |-- P.name --| P.$$$ ")") "" --
 | 
| 12876 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 wenzelm parents: 
12474diff
changeset | 290 | Scan.repeat1 (P.opt_thm_name ":" -- P.prop); | 
| 6359 | 291 | |
| 292 | val primrecP = | |
| 6723 | 293 | OuterSyntax.command "primrec" "define primitive recursive functions on datatypes" K.thy_decl | 
| 6359 | 294 | (primrec_decl >> (fn (alt_name, eqns) => | 
| 6723 | 295 | Toplevel.theory (#1 o add_primrec alt_name (map P.triple_swap eqns)))); | 
| 6359 | 296 | |
| 297 | val _ = OuterSyntax.add_parsers [primrecP]; | |
| 5178 | 298 | |
| 299 | end; | |
| 6384 | 300 | |
| 301 | ||
| 302 | end; | |
| 15705 | 303 |