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