| author | wenzelm | 
| Sat, 20 Aug 2011 23:35:30 +0200 | |
| changeset 44338 | 700008399ee5 | 
| parent 44241 | 7943b69f0188 | 
| child 45592 | 8baa0b7f3f66 | 
| permissions | -rw-r--r-- | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31296diff
changeset | 1 | (* Title: HOL/Tools/primrec.ML | 
| 39812 | 2 | Author: Norbert Voelker, FernUni Hagen | 
| 3 | Author: Stefan Berghofer, TU Muenchen | |
| 4 | Author: Florian Haftmann, TU Muenchen | |
| 5178 | 5 | |
| 39812 | 6 | Primitive recursive functions on datatypes. | 
| 5178 | 7 | *) | 
| 8 | ||
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31296diff
changeset | 9 | signature PRIMREC = | 
| 5178 | 10 | sig | 
| 29581 | 11 | val add_primrec: (binding * typ option * mixfix) list -> | 
| 35166 | 12 | (Attrib.binding * term) list -> local_theory -> (term list * thm list) * local_theory | 
| 30487 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30455diff
changeset | 13 | val add_primrec_cmd: (binding * string option * mixfix) list -> | 
| 35166 | 14 | (Attrib.binding * string) list -> local_theory -> (term list * thm list) * local_theory | 
| 29581 | 15 | val add_primrec_global: (binding * typ option * mixfix) list -> | 
| 35166 | 16 | (Attrib.binding * term) list -> theory -> (term list * thm list) * theory | 
| 26679 
447f4872b7ee
Added add_primrec_global and add_primrec_overloaded functions (thanks to Markus).
 berghofe parents: 
26556diff
changeset | 17 | val add_primrec_overloaded: (string * (string * typ) * bool) list -> | 
| 29581 | 18 | (binding * typ option * mixfix) list -> | 
| 35166 | 19 | (Attrib.binding * term) list -> theory -> (term list * thm list) * theory | 
| 31269 | 20 | val add_primrec_simple: ((binding * typ) * mixfix) list -> term list -> | 
| 35166 | 21 | local_theory -> (string * (term list * thm list)) * local_theory | 
| 5178 | 22 | end; | 
| 23 | ||
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31296diff
changeset | 24 | structure Primrec : PRIMREC = | 
| 5178 | 25 | struct | 
| 26 | ||
| 33968 
f94fb13ecbb3
modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
 haftmann parents: 
33963diff
changeset | 27 | open Datatype_Aux; | 
| 5178 | 28 | |
| 25557 | 29 | exception PrimrecError of string * term option; | 
| 6359 | 30 | |
| 25557 | 31 | fun primrec_error msg = raise PrimrecError (msg, NONE); | 
| 32 | fun primrec_error_eqn msg eqn = raise PrimrecError (msg, SOME eqn); | |
| 23765 
997e5fe47532
Function unify_consts moved from OldInductivePackage to PrimrecPackage.
 berghofe parents: 
22728diff
changeset | 33 | |
| 
997e5fe47532
Function unify_consts moved from OldInductivePackage to PrimrecPackage.
 berghofe parents: 
22728diff
changeset | 34 | |
| 5178 | 35 | (* preprocessing of equations *) | 
| 36 | ||
| 25559 | 37 | fun process_eqn is_fixed spec rec_fns = | 
| 5178 | 38 | let | 
| 25566 | 39 | val (vs, Ts) = split_list (strip_qnt_vars "all" spec); | 
| 25557 | 40 | val body = strip_qnt_body "all" spec; | 
| 43326 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 wenzelm parents: 
42482diff
changeset | 41 | val (vs', _) = fold_map Name.variant vs (Name.make_context (fold_aterms | 
| 25566 | 42 | (fn Free (v, _) => insert (op =) v | _ => I) body [])); | 
| 43 | val eqn = curry subst_bounds (map2 (curry Free) vs' Ts |> rev) body; | |
| 25557 | 44 | val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop eqn) | 
| 45 | handle TERM _ => primrec_error "not a proper equation"; | |
| 5178 | 46 | val (recfun, args) = strip_comb lhs; | 
| 39812 | 47 | val fname = | 
| 48 | (case recfun of | |
| 49 | Free (v, _) => | |
| 50 | if is_fixed v then v | |
| 25557 | 51 | else primrec_error "illegal head of function equation" | 
| 39812 | 52 | | _ => primrec_error "illegal head of function equation"); | 
| 5178 | 53 | |
| 54 | val (ls', rest) = take_prefix is_Free args; | |
| 55 | val (middle, rs') = take_suffix is_Free rest; | |
| 56 | val rpos = length ls'; | |
| 57 | ||
| 39812 | 58 | val (constr, cargs') = | 
| 59 | if null middle then primrec_error "constructor missing" | |
| 5178 | 60 | else strip_comb (hd middle); | 
| 61 | val (cname, T) = dest_Const constr | |
| 25557 | 62 | handle TERM _ => primrec_error "ill-formed constructor"; | 
| 7016 
df54b5365477
- Now also supports arbitrarily branching datatypes.
 berghofe parents: 
6729diff
changeset | 63 | val (tname, _) = dest_Type (body_type T) handle TYPE _ => | 
| 25557 | 64 | primrec_error "cannot determine datatype associated with function" | 
| 5178 | 65 | |
| 20176 | 66 | val (ls, cargs, rs) = | 
| 67 | (map dest_Free ls', map dest_Free cargs', map dest_Free rs') | |
| 25557 | 68 | handle TERM _ => primrec_error "illegal argument in pattern"; | 
| 5178 | 69 | val lfrees = ls @ rs @ cargs; | 
| 70 | ||
| 12474 | 71 | fun check_vars _ [] = () | 
| 25557 | 72 | | check_vars s vars = primrec_error (s ^ commas_quote (map fst vars)) eqn; | 
| 5178 | 73 | in | 
| 22692 | 74 | if length middle > 1 then | 
| 25557 | 75 | primrec_error "more than one non-variable in pattern" | 
| 12474 | 76 | else | 
| 18964 | 77 | (check_vars "repeated variable names in pattern: " (duplicates (op =) lfrees); | 
| 12474 | 78 | check_vars "extra variables on rhs: " | 
| 39814 | 79 | (Term.add_frees rhs [] |> subtract (op =) lfrees | 
| 25559 | 80 | |> filter_out (is_fixed o fst)); | 
| 39812 | 81 | (case AList.lookup (op =) rec_fns fname of | 
| 15531 | 82 | NONE => | 
| 39812 | 83 | (fname, (tname, rpos, [(cname, (ls, cargs, rs, rhs, eqn))])) :: rec_fns | 
| 15531 | 84 | | SOME (_, rpos', eqns) => | 
| 17184 | 85 | if AList.defined (op =) eqns cname then | 
| 25557 | 86 | primrec_error "constructor already occurred as pattern" | 
| 5178 | 87 | else if rpos <> rpos' then | 
| 25557 | 88 | primrec_error "position of recursive argument inconsistent" | 
| 5178 | 89 | else | 
| 25557 | 90 | AList.update (op =) | 
| 39812 | 91 | (fname, (tname, rpos, (cname, (ls, cargs, rs, rhs, eqn)) :: eqns)) | 
| 92 | rec_fns)) | |
| 25557 | 93 | end handle PrimrecError (msg, NONE) => primrec_error_eqn msg spec; | 
| 5178 | 94 | |
| 25557 | 95 | fun process_fun descr eqns (i, fname) (fnames, fnss) = | 
| 5178 | 96 | let | 
| 25557 | 97 | val (_, (tname, _, constrs)) = nth descr i; | 
| 5178 | 98 | |
| 99 | (* substitute "fname ls x rs" by "y ls rs" for (x, (_, y)) in subs *) | |
| 100 | ||
| 21064 | 101 | fun subst [] t fs = (t, fs) | 
| 102 | | subst subs (Abs (a, T, t)) fs = | |
| 103 | fs | |
| 104 | |> subst subs t | |
| 105 | |-> (fn t' => pair (Abs (a, T, t'))) | |
| 106 | | subst subs (t as (_ $ _)) fs = | |
| 107 | let | |
| 108 | val (f, ts) = strip_comb t; | |
| 5178 | 109 | in | 
| 25557 | 110 | if is_Free f | 
| 111 | andalso member (fn ((v, _), (w, _)) => v = w) eqns (dest_Free f) then | |
| 5178 | 112 | let | 
| 25557 | 113 | val (fname', _) = dest_Free f; | 
| 114 | val (_, rpos, _) = the (AList.lookup (op =) eqns fname'); | |
| 27301 
bf7d82193a2e
fixed bind error for malformed primrec specifications
 haftmann parents: 
26679diff
changeset | 115 | val (ls, rs) = chop rpos ts | 
| 39812 | 116 | val (x', rs') = | 
| 117 | (case rs of | |
| 118 | x' :: rs => (x', rs) | |
| 119 |                   | [] => primrec_error ("not enough arguments in recursive application\n" ^
 | |
| 120 | "of function " ^ quote fname' ^ " on rhs")); | |
| 27301 
bf7d82193a2e
fixed bind error for malformed primrec specifications
 haftmann parents: 
26679diff
changeset | 121 | val (x, xs) = strip_comb x'; | 
| 39812 | 122 | in | 
| 123 | (case AList.lookup (op =) subs x of | |
| 124 | NONE => | |
| 21064 | 125 | fs | 
| 126 | |> fold_map (subst subs) ts | |
| 127 | |-> (fn ts' => pair (list_comb (f, ts'))) | |
| 128 | | SOME (i', y) => | |
| 129 | fs | |
| 27301 
bf7d82193a2e
fixed bind error for malformed primrec specifications
 haftmann parents: 
26679diff
changeset | 130 | |> fold_map (subst subs) (xs @ ls @ rs') | 
| 25557 | 131 | ||> process_fun descr eqns (i', fname') | 
| 39812 | 132 | |-> (fn ts' => pair (list_comb (y, ts')))) | 
| 5178 | 133 | end | 
| 134 | else | |
| 21064 | 135 | fs | 
| 136 | |> fold_map (subst subs) (f :: ts) | |
| 39812 | 137 | |-> (fn f' :: ts' => pair (list_comb (f', ts'))) | 
| 5178 | 138 | end | 
| 21064 | 139 | | subst _ t fs = (t, fs); | 
| 5178 | 140 | |
| 141 | (* translate rec equations into function arguments suitable for rec comb *) | |
| 142 | ||
| 25557 | 143 | fun trans eqns (cname, cargs) (fnames', fnss', fns) = | 
| 17184 | 144 | (case AList.lookup (op =) eqns cname of | 
| 39812 | 145 |         NONE => (warning ("No equation for constructor " ^ quote cname ^
 | 
| 146 | "\nin definition of function " ^ quote fname); | |
| 39813 | 147 |             (fnames', fnss', (Const (@{const_name undefined}, dummyT)) :: fns))
 | 
| 39812 | 148 | | SOME (ls, cargs', rs, rhs, eq) => | 
| 149 | let | |
| 150 | val recs = filter (is_rec_type o snd) (cargs' ~~ cargs); | |
| 151 | val rargs = map fst recs; | |
| 152 | val subs = map (rpair dummyT o fst) | |
| 153 | (rev (Term.rename_wrt_term rhs rargs)); | |
| 154 | val (rhs', (fnames'', fnss'')) = subst (map2 (fn (x, y) => fn z => | |
| 155 | (Free x, (body_index y, Free z))) recs subs) rhs (fnames', fnss') | |
| 156 | handle PrimrecError (s, NONE) => primrec_error_eqn s eq | |
| 44241 | 157 | in | 
| 158 | (fnames'', fnss'', fold_rev absfree (cargs' @ subs @ ls @ rs) rhs' :: fns) | |
| 39812 | 159 | end) | 
| 5178 | 160 | |
| 39812 | 161 | in | 
| 162 | (case AList.lookup (op =) fnames i of | |
| 15531 | 163 | NONE => | 
| 25557 | 164 | if exists (fn (_, v) => fname = v) fnames then | 
| 165 |           primrec_error ("inconsistent functions for datatype " ^ quote tname)
 | |
| 5178 | 166 | else | 
| 167 | let | |
| 25557 | 168 | val (_, _, eqns) = the (AList.lookup (op =) eqns fname); | 
| 169 | val (fnames', fnss', fns) = fold_rev (trans eqns) constrs | |
| 39812 | 170 | ((i, fname) :: fnames, fnss, []) | 
| 5178 | 171 | in | 
| 39812 | 172 | (fnames', (i, (fname, #1 (snd (hd eqns)), fns)) :: fnss') | 
| 5178 | 173 | end | 
| 25557 | 174 | | SOME fname' => | 
| 175 | if fname = fname' then (fnames, fnss) | |
| 176 |         else primrec_error ("inconsistent functions for datatype " ^ quote tname))
 | |
| 5178 | 177 | end; | 
| 178 | ||
| 6359 | 179 | |
| 5178 | 180 | (* prepare functions needed for definitions *) | 
| 181 | ||
| 21064 | 182 | fun get_fns fns ((i : int, (tname, _, constrs)), rec_name) (fs, defs) = | 
| 39812 | 183 | (case AList.lookup (op =) fns i of | 
| 184 | NONE => | |
| 185 | let | |
| 39813 | 186 |         val dummy_fns = map (fn (_, cargs) => Const (@{const_name undefined},
 | 
| 39812 | 187 | replicate (length cargs + length (filter is_rec_type cargs)) | 
| 188 | dummyT ---> HOLogic.unitT)) constrs; | |
| 189 |         val _ = warning ("No function definition for datatype " ^ quote tname)
 | |
| 190 | in | |
| 191 | (dummy_fns @ fs, defs) | |
| 192 | end | |
| 193 | | SOME (fname, ls, fs') => (fs' @ fs, (fname, ls, rec_name, tname) :: defs)); | |
| 5178 | 194 | |
| 6359 | 195 | |
| 5178 | 196 | (* make definition *) | 
| 197 | ||
| 25557 | 198 | fun make_def ctxt fixes fs (fname, ls, rec_name, tname) = | 
| 5178 | 199 | let | 
| 30451 
11e5e8bb28f9
corrected type inference of primitive definitions
 haftmann parents: 
30364diff
changeset | 200 | val SOME (var, varT) = get_first (fn ((b, T), mx) => | 
| 
11e5e8bb28f9
corrected type inference of primitive definitions
 haftmann parents: 
30364diff
changeset | 201 | if Binding.name_of b = fname then SOME ((b, mx), T) else NONE) fixes; | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 202 | val def_name = Thm.def_name (Long_Name.base_name fname); | 
| 30451 
11e5e8bb28f9
corrected type inference of primitive definitions
 haftmann parents: 
30364diff
changeset | 203 |     val raw_rhs = fold_rev (fn T => fn t => Abs ("", T, t)) (map snd ls @ [dummyT])
 | 
| 
11e5e8bb28f9
corrected type inference of primitive definitions
 haftmann parents: 
30364diff
changeset | 204 | (list_comb (Const (rec_name, dummyT), fs @ map Bound (0 :: (length ls downto 1)))) | 
| 39288 | 205 | val rhs = singleton (Syntax.check_terms ctxt) (Type.constraint varT raw_rhs); | 
| 33755 
6dc1b67f2127
concealing internal definitions of primrec specifications
 bulwahn parents: 
33726diff
changeset | 206 | in (var, ((Binding.conceal (Binding.name def_name), []), rhs)) end; | 
| 5178 | 207 | |
| 6359 | 208 | |
| 5178 | 209 | (* find datatypes which contain all datatypes in tnames' *) | 
| 210 | ||
| 31737 | 211 | fun find_dts (dt_info : info Symtab.table) _ [] = [] | 
| 39812 | 212 | | find_dts dt_info tnames' (tname :: tnames) = | 
| 17412 | 213 | (case Symtab.lookup dt_info tname of | 
| 39812 | 214 | NONE => primrec_error (quote tname ^ " is not a datatype") | 
| 215 | | SOME dt => | |
| 216 | if subset (op =) (tnames', map (#1 o snd) (#descr dt)) then | |
| 217 | (tname, dt) :: (find_dts dt_info tnames' tnames) | |
| 218 | else find_dts dt_info tnames' tnames); | |
| 5178 | 219 | |
| 25557 | 220 | |
| 31262 | 221 | (* distill primitive definition(s) from primrec specification *) | 
| 25557 | 222 | |
| 31262 | 223 | fun distill lthy fixes eqs = | 
| 25557 | 224 | let | 
| 25559 | 225 | val eqns = fold_rev (process_eqn (fn v => Variable.is_fixed lthy v | 
| 31262 | 226 | orelse exists (fn ((w, _), _) => v = Binding.name_of w) fixes)) eqs []; | 
| 25557 | 227 | val tnames = distinct (op =) (map (#1 o snd) eqns); | 
| 42361 | 228 | val dts = find_dts (Datatype_Data.get_all (Proof_Context.theory_of lthy)) tnames tnames; | 
| 25557 | 229 |     val main_fns = map (fn (tname, {index, ...}) =>
 | 
| 230 | (index, (fst o the o find_first (fn (_, x) => #1 x = tname)) eqns)) dts; | |
| 22692 | 231 |     val {descr, rec_names, rec_rewrites, ...} =
 | 
| 25557 | 232 | if null dts then primrec_error | 
| 233 |         ("datatypes " ^ commas_quote tnames ^ "\nare not mutually recursive")
 | |
| 18362 
e8b7e0a22727
removed thms 'swap' and 'nth_map' from ML toplevel
 haftmann parents: 
18358diff
changeset | 234 | else snd (hd dts); | 
| 25557 | 235 | val (fnames, fnss) = fold_rev (process_fun descr eqns) main_fns ([], []); | 
| 31262 | 236 | val (fs, raw_defs) = fold_rev (get_fns fnss) (descr ~~ rec_names) ([], []); | 
| 237 | val defs = map (make_def lthy fixes fs) raw_defs; | |
| 238 | val names = map snd fnames; | |
| 239 | val names_eqns = map fst eqns; | |
| 39812 | 240 | val _ = | 
| 241 | if eq_set (op =) (names, names_eqns) then () | |
| 31262 | 242 |       else primrec_error ("functions " ^ commas_quote names_eqns ^
 | 
| 25557 | 243 | "\nare not mutually recursive"); | 
| 31262 | 244 | val rec_rewrites' = map mk_meta_eq rec_rewrites; | 
| 245 | val prefix = space_implode "_" (map (Long_Name.base_name o #1) raw_defs); | |
| 246 | fun prove lthy defs = | |
| 247 | let | |
| 42482 | 248 | val frees = fold (Variable.add_free_names lthy) eqs []; | 
| 31262 | 249 | val rewrites = rec_rewrites' @ map (snd o snd) defs; | 
| 250 | fun tac _ = EVERY [rewrite_goals_tac rewrites, rtac refl 1]; | |
| 35166 | 251 | in map (fn eq => Goal.prove lthy frees [] eq tac) eqs end; | 
| 31262 | 252 | in ((prefix, (fs, defs)), prove) end | 
| 253 | handle PrimrecError (msg, some_eqn) => | |
| 39812 | 254 |     error ("Primrec definition error:\n" ^ msg ^
 | 
| 255 | (case some_eqn of | |
| 256 | SOME eqn => "\nin\n" ^ quote (Syntax.string_of_term lthy eqn) | |
| 31262 | 257 | | NONE => "")); | 
| 258 | ||
| 259 | ||
| 260 | (* primrec definition *) | |
| 261 | ||
| 31269 | 262 | fun add_primrec_simple fixes ts lthy = | 
| 31262 | 263 | let | 
| 31269 | 264 | val ((prefix, (fs, defs)), prove) = distill lthy fixes ts; | 
| 31262 | 265 | in | 
| 266 | lthy | |
| 33766 
c679f05600cd
adapted Local_Theory.define -- eliminated odd thm kind;
 wenzelm parents: 
33755diff
changeset | 267 | |> fold_map Local_Theory.define defs | 
| 35166 | 268 | |-> (fn defs => `(fn lthy => (prefix, (map fst defs, prove lthy defs)))) | 
| 31262 | 269 | end; | 
| 270 | ||
| 271 | local | |
| 272 | ||
| 33726 
0878aecbf119
eliminated slightly odd name space grouping -- now managed by Isar toplevel;
 wenzelm parents: 
33671diff
changeset | 273 | fun gen_primrec prep_spec raw_fixes raw_spec lthy = | 
| 31262 | 274 | let | 
| 275 | val (fixes, spec) = fst (prep_spec raw_fixes raw_spec lthy); | |
| 276 | fun attr_bindings prefix = map (fn ((b, attrs), _) => | |
| 277 | (Binding.qualify false prefix b, Code.add_default_eqn_attrib :: attrs)) spec; | |
| 33552 | 278 | fun simp_attr_binding prefix = | 
| 279 | (Binding.qualify true prefix (Binding.name "simps"), | |
| 280 | map (Attrib.internal o K) [Simplifier.simp_add, Nitpick_Simps.add]); | |
| 8480 
50266d517b0c
Added new theory data slot for primrec equations.
 berghofe parents: 
8432diff
changeset | 281 | in | 
| 25557 | 282 | lthy | 
| 31269 | 283 | |> add_primrec_simple fixes (map snd spec) | 
| 35756 
cfde251d03a5
refining and adding Spec_Rules to definitional packages old_primrec, primrec, recdef, size and function
 bulwahn parents: 
35166diff
changeset | 284 | |-> (fn (prefix, (ts, simps)) => | 
| 
cfde251d03a5
refining and adding Spec_Rules to definitional packages old_primrec, primrec, recdef, size and function
 bulwahn parents: 
35166diff
changeset | 285 | Spec_Rules.add Spec_Rules.Equational (ts, simps) | 
| 
cfde251d03a5
refining and adding Spec_Rules to definitional packages old_primrec, primrec, recdef, size and function
 bulwahn parents: 
35166diff
changeset | 286 | #> fold_map Local_Theory.note (attr_bindings prefix ~~ map single simps) | 
| 34952 
bd7e347eb768
added registration of equational theorems from prim_rec and rec_def to Spec_Rules
 bulwahn parents: 
33968diff
changeset | 287 | #-> (fn simps' => Local_Theory.note (simp_attr_binding prefix, maps snd simps') | 
| 35756 
cfde251d03a5
refining and adding Spec_Rules to definitional packages old_primrec, primrec, recdef, size and function
 bulwahn parents: 
35166diff
changeset | 288 | #>> (fn (_, simps'') => (ts, simps'')))) | 
| 31262 | 289 | end; | 
| 20841 | 290 | |
| 291 | in | |
| 292 | ||
| 33726 
0878aecbf119
eliminated slightly odd name space grouping -- now managed by Isar toplevel;
 wenzelm parents: 
33671diff
changeset | 293 | val add_primrec = gen_primrec Specification.check_spec; | 
| 
0878aecbf119
eliminated slightly odd name space grouping -- now managed by Isar toplevel;
 wenzelm parents: 
33671diff
changeset | 294 | val add_primrec_cmd = gen_primrec Specification.read_spec; | 
| 20841 | 295 | |
| 22692 | 296 | end; | 
| 19688 | 297 | |
| 26679 
447f4872b7ee
Added add_primrec_global and add_primrec_overloaded functions (thanks to Markus).
 berghofe parents: 
26556diff
changeset | 298 | fun add_primrec_global fixes specs thy = | 
| 
447f4872b7ee
Added add_primrec_global and add_primrec_overloaded functions (thanks to Markus).
 berghofe parents: 
26556diff
changeset | 299 | let | 
| 38388 | 300 | val lthy = Named_Target.theory_init thy; | 
| 35166 | 301 | val ((ts, simps), lthy') = add_primrec fixes specs lthy; | 
| 42361 | 302 | val simps' = Proof_Context.export lthy' lthy simps; | 
| 35166 | 303 | in ((ts, simps'), Local_Theory.exit_global lthy') end; | 
| 26679 
447f4872b7ee
Added add_primrec_global and add_primrec_overloaded functions (thanks to Markus).
 berghofe parents: 
26556diff
changeset | 304 | |
| 
447f4872b7ee
Added add_primrec_global and add_primrec_overloaded functions (thanks to Markus).
 berghofe parents: 
26556diff
changeset | 305 | fun add_primrec_overloaded ops fixes specs thy = | 
| 
447f4872b7ee
Added add_primrec_global and add_primrec_overloaded functions (thanks to Markus).
 berghofe parents: 
26556diff
changeset | 306 | let | 
| 38342 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 haftmann parents: 
37145diff
changeset | 307 | val lthy = Overloading.overloading ops thy; | 
| 35166 | 308 | val ((ts, simps), lthy') = add_primrec fixes specs lthy; | 
| 42361 | 309 | val simps' = Proof_Context.export lthy' lthy simps; | 
| 35166 | 310 | in ((ts, simps'), Local_Theory.exit_global lthy') end; | 
| 26679 
447f4872b7ee
Added add_primrec_global and add_primrec_overloaded functions (thanks to Markus).
 berghofe parents: 
26556diff
changeset | 311 | |
| 5178 | 312 | |
| 6359 | 313 | (* outer syntax *) | 
| 314 | ||
| 24867 | 315 | val _ = | 
| 39812 | 316 | Outer_Syntax.local_theory "primrec" "define primitive recursive functions on datatypes" | 
| 317 | Keyword.thy_decl | |
| 318 | (Parse.fixes -- Parse_Spec.where_alt_specs | |
| 319 | >> (fn (fixes, specs) => add_primrec_cmd fixes specs #> snd)); | |
| 6359 | 320 | |
| 5178 | 321 | end; |