| author | haftmann | 
| Fri, 17 Apr 2009 08:34:54 +0200 | |
| changeset 30942 | 1e246776f876 | 
| parent 30487 | a14ff49d3083 | 
| child 31172 | 74d72ba262fb | 
| permissions | -rw-r--r-- | 
| 5719 | 1 | (* Title: HOL/Tools/primrec_package.ML | 
| 25557 | 2 | Author: Stefan Berghofer, TU Muenchen; Norbert Voelker, FernUni Hagen; | 
| 3 | Florian Haftmann, TU Muenchen | |
| 5178 | 4 | |
| 6359 | 5 | Package for defining functions on datatypes by primitive recursion. | 
| 5178 | 6 | *) | 
| 7 | ||
| 8 | signature PRIMREC_PACKAGE = | |
| 9 | sig | |
| 29581 | 10 | val add_primrec: (binding * typ option * mixfix) list -> | 
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 11 | (Attrib.binding * term) list -> local_theory -> thm list * local_theory | 
| 30487 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30455diff
changeset | 12 | val add_primrec_cmd: (binding * string option * mixfix) list -> | 
| 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30455diff
changeset | 13 | (Attrib.binding * string) list -> local_theory -> thm list * local_theory | 
| 29581 | 14 | val add_primrec_global: (binding * typ option * mixfix) list -> | 
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 15 | (Attrib.binding * term) list -> theory -> thm list * theory | 
| 26679 
447f4872b7ee
Added add_primrec_global and add_primrec_overloaded functions (thanks to Markus).
 berghofe parents: 
26556diff
changeset | 16 | val add_primrec_overloaded: (string * (string * typ) * bool) list -> | 
| 29581 | 17 | (binding * typ option * mixfix) list -> | 
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 18 | (Attrib.binding * term) list -> theory -> thm list * theory | 
| 5178 | 19 | end; | 
| 20 | ||
| 21 | structure PrimrecPackage : PRIMREC_PACKAGE = | |
| 22 | struct | |
| 23 | ||
| 24 | open DatatypeAux; | |
| 25 | ||
| 25557 | 26 | exception PrimrecError of string * term option; | 
| 6359 | 27 | |
| 25557 | 28 | fun primrec_error msg = raise PrimrecError (msg, NONE); | 
| 29 | 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 | 30 | |
| 30451 
11e5e8bb28f9
corrected type inference of primitive definitions
 haftmann parents: 
30364diff
changeset | 31 | fun message s = if ! Toplevel.debug then tracing s else (); | 
| 23765 
997e5fe47532
Function unify_consts moved from OldInductivePackage to PrimrecPackage.
 berghofe parents: 
22728diff
changeset | 32 | |
| 
997e5fe47532
Function unify_consts moved from OldInductivePackage to PrimrecPackage.
 berghofe parents: 
22728diff
changeset | 33 | |
| 5178 | 34 | (* preprocessing of equations *) | 
| 35 | ||
| 25559 | 36 | fun process_eqn is_fixed spec rec_fns = | 
| 5178 | 37 | let | 
| 25566 | 38 | val (vs, Ts) = split_list (strip_qnt_vars "all" spec); | 
| 25557 | 39 | val body = strip_qnt_body "all" spec; | 
| 25566 | 40 | val (vs', _) = Name.variants vs (Name.make_context (fold_aterms | 
| 41 | (fn Free (v, _) => insert (op =) v | _ => I) body [])); | |
| 42 | val eqn = curry subst_bounds (map2 (curry Free) vs' Ts |> rev) body; | |
| 25557 | 43 | val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop eqn) | 
| 44 | handle TERM _ => primrec_error "not a proper equation"; | |
| 5178 | 45 | val (recfun, args) = strip_comb lhs; | 
| 25557 | 46 | val fname = case recfun of Free (v, _) => if is_fixed v then v | 
| 47 | else primrec_error "illegal head of function equation" | |
| 48 | | _ => primrec_error "illegal head of function equation"; | |
| 5178 | 49 | |
| 50 | val (ls', rest) = take_prefix is_Free args; | |
| 51 | val (middle, rs') = take_suffix is_Free rest; | |
| 52 | val rpos = length ls'; | |
| 53 | ||
| 25557 | 54 | val (constr, cargs') = if null middle then primrec_error "constructor missing" | 
| 5178 | 55 | else strip_comb (hd middle); | 
| 56 | val (cname, T) = dest_Const constr | |
| 25557 | 57 | handle TERM _ => primrec_error "ill-formed constructor"; | 
| 7016 
df54b5365477
- Now also supports arbitrarily branching datatypes.
 berghofe parents: 
6729diff
changeset | 58 | val (tname, _) = dest_Type (body_type T) handle TYPE _ => | 
| 25557 | 59 | primrec_error "cannot determine datatype associated with function" | 
| 5178 | 60 | |
| 20176 | 61 | val (ls, cargs, rs) = | 
| 62 | (map dest_Free ls', map dest_Free cargs', map dest_Free rs') | |
| 25557 | 63 | handle TERM _ => primrec_error "illegal argument in pattern"; | 
| 5178 | 64 | val lfrees = ls @ rs @ cargs; | 
| 65 | ||
| 12474 | 66 | fun check_vars _ [] = () | 
| 25557 | 67 | | check_vars s vars = primrec_error (s ^ commas_quote (map fst vars)) eqn; | 
| 5178 | 68 | in | 
| 22692 | 69 | if length middle > 1 then | 
| 25557 | 70 | primrec_error "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: " | 
| 29265 
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
 wenzelm parents: 
29006diff
changeset | 74 | (map dest_Free (OldTerm.term_frees rhs) |> subtract (op =) lfrees | 
| 25559 | 75 | |> filter_out (is_fixed o fst)); | 
| 25557 | 76 | case AList.lookup (op =) rec_fns fname of | 
| 15531 | 77 | NONE => | 
| 25557 | 78 | (fname, (tname, rpos, [(cname, (ls, cargs, rs, rhs, eqn))]))::rec_fns | 
| 15531 | 79 | | SOME (_, rpos', eqns) => | 
| 17184 | 80 | if AList.defined (op =) eqns cname then | 
| 25557 | 81 | primrec_error "constructor already occurred as pattern" | 
| 5178 | 82 | else if rpos <> rpos' then | 
| 25557 | 83 | primrec_error "position of recursive argument inconsistent" | 
| 5178 | 84 | else | 
| 25557 | 85 | AList.update (op =) | 
| 86 | (fname, (tname, rpos, (cname, (ls, cargs, rs, rhs, eqn))::eqns)) | |
| 17314 | 87 | rec_fns) | 
| 25557 | 88 | end handle PrimrecError (msg, NONE) => primrec_error_eqn msg spec; | 
| 5178 | 89 | |
| 25557 | 90 | fun process_fun descr eqns (i, fname) (fnames, fnss) = | 
| 5178 | 91 | let | 
| 25557 | 92 | val (_, (tname, _, constrs)) = nth descr i; | 
| 5178 | 93 | |
| 94 | (* substitute "fname ls x rs" by "y ls rs" for (x, (_, y)) in subs *) | |
| 95 | ||
| 21064 | 96 | fun subst [] t fs = (t, fs) | 
| 97 | | subst subs (Abs (a, T, t)) fs = | |
| 98 | fs | |
| 99 | |> subst subs t | |
| 100 | |-> (fn t' => pair (Abs (a, T, t'))) | |
| 101 | | subst subs (t as (_ $ _)) fs = | |
| 102 | let | |
| 103 | val (f, ts) = strip_comb t; | |
| 5178 | 104 | in | 
| 25557 | 105 | if is_Free f | 
| 106 | andalso member (fn ((v, _), (w, _)) => v = w) eqns (dest_Free f) then | |
| 5178 | 107 | let | 
| 25557 | 108 | val (fname', _) = dest_Free f; | 
| 109 | val (_, rpos, _) = the (AList.lookup (op =) eqns fname'); | |
| 27301 
bf7d82193a2e
fixed bind error for malformed primrec specifications
 haftmann parents: 
26679diff
changeset | 110 | val (ls, rs) = chop rpos ts | 
| 
bf7d82193a2e
fixed bind error for malformed primrec specifications
 haftmann parents: 
26679diff
changeset | 111 | val (x', rs') = case rs | 
| 
bf7d82193a2e
fixed bind error for malformed primrec specifications
 haftmann parents: 
26679diff
changeset | 112 | of x' :: rs => (x', rs) | 
| 
bf7d82193a2e
fixed bind error for malformed primrec specifications
 haftmann parents: 
26679diff
changeset | 113 |                   | [] => primrec_error ("not enough arguments in recursive application\n"
 | 
| 
bf7d82193a2e
fixed bind error for malformed primrec specifications
 haftmann parents: 
26679diff
changeset | 114 | ^ "of function " ^ quote fname' ^ " on rhs"); | 
| 
bf7d82193a2e
fixed bind error for malformed primrec specifications
 haftmann parents: 
26679diff
changeset | 115 | val (x, xs) = strip_comb x'; | 
| 21064 | 116 | in case AList.lookup (op =) subs x | 
| 117 | of NONE => | |
| 118 | fs | |
| 119 | |> fold_map (subst subs) ts | |
| 120 | |-> (fn ts' => pair (list_comb (f, ts'))) | |
| 121 | | SOME (i', y) => | |
| 122 | fs | |
| 27301 
bf7d82193a2e
fixed bind error for malformed primrec specifications
 haftmann parents: 
26679diff
changeset | 123 | |> fold_map (subst subs) (xs @ ls @ rs') | 
| 25557 | 124 | ||> process_fun descr eqns (i', fname') | 
| 21064 | 125 | |-> (fn ts' => pair (list_comb (y, ts'))) | 
| 5178 | 126 | end | 
| 127 | else | |
| 21064 | 128 | fs | 
| 129 | |> fold_map (subst subs) (f :: ts) | |
| 130 | |-> (fn (f'::ts') => pair (list_comb (f', ts'))) | |
| 5178 | 131 | end | 
| 21064 | 132 | | subst _ t fs = (t, fs); | 
| 5178 | 133 | |
| 134 | (* translate rec equations into function arguments suitable for rec comb *) | |
| 135 | ||
| 25557 | 136 | fun trans eqns (cname, cargs) (fnames', fnss', fns) = | 
| 17184 | 137 | (case AList.lookup (op =) eqns cname of | 
| 15531 | 138 |           NONE => (warning ("No equation for constructor " ^ quote cname ^
 | 
| 6427 | 139 | "\nin definition of function " ^ quote fname); | 
| 25557 | 140 |               (fnames', fnss', (Const ("HOL.undefined", dummyT))::fns))
 | 
| 15531 | 141 | | SOME (ls, cargs', rs, rhs, eq) => | 
| 5178 | 142 | let | 
| 21064 | 143 | val recs = filter (is_rec_type o snd) (cargs' ~~ cargs); | 
| 5178 | 144 | val rargs = map fst recs; | 
| 22692 | 145 | val subs = map (rpair dummyT o fst) | 
| 29276 | 146 | (rev (Term.rename_wrt_term rhs rargs)); | 
| 27301 
bf7d82193a2e
fixed bind error for malformed primrec specifications
 haftmann parents: 
26679diff
changeset | 147 | val (rhs', (fnames'', fnss'')) = subst (map2 (fn (x, y) => fn z => | 
| 
bf7d82193a2e
fixed bind error for malformed primrec specifications
 haftmann parents: 
26679diff
changeset | 148 | (Free x, (body_index y, Free z))) recs subs) rhs (fnames', fnss') | 
| 25557 | 149 | handle PrimrecError (s, NONE) => primrec_error_eqn s eq | 
| 150 | in (fnames'', fnss'', | |
| 20176 | 151 | (list_abs_free (cargs' @ subs @ ls @ rs, rhs'))::fns) | 
| 5178 | 152 | end) | 
| 153 | ||
| 25557 | 154 | in (case AList.lookup (op =) fnames i of | 
| 15531 | 155 | NONE => | 
| 25557 | 156 | if exists (fn (_, v) => fname = v) fnames then | 
| 157 |           primrec_error ("inconsistent functions for datatype " ^ quote tname)
 | |
| 5178 | 158 | else | 
| 159 | let | |
| 25557 | 160 | val (_, _, eqns) = the (AList.lookup (op =) eqns fname); | 
| 161 | val (fnames', fnss', fns) = fold_rev (trans eqns) constrs | |
| 162 | ((i, fname)::fnames, fnss, []) | |
| 5178 | 163 | in | 
| 25557 | 164 | (fnames', (i, (fname, #1 (snd (hd eqns)), fns))::fnss') | 
| 5178 | 165 | end | 
| 25557 | 166 | | SOME fname' => | 
| 167 | if fname = fname' then (fnames, fnss) | |
| 168 |         else primrec_error ("inconsistent functions for datatype " ^ quote tname))
 | |
| 5178 | 169 | end; | 
| 170 | ||
| 6359 | 171 | |
| 5178 | 172 | (* prepare functions needed for definitions *) | 
| 173 | ||
| 21064 | 174 | fun get_fns fns ((i : int, (tname, _, constrs)), rec_name) (fs, defs) = | 
| 17184 | 175 | case AList.lookup (op =) fns i of | 
| 15531 | 176 | NONE => | 
| 5178 | 177 | let | 
| 22480 | 178 |          val dummy_fns = map (fn (_, cargs) => Const ("HOL.undefined",
 | 
| 15570 | 179 | replicate ((length cargs) + (length (List.filter is_rec_type cargs))) | 
| 5178 | 180 | dummyT ---> HOLogic.unitT)) constrs; | 
| 6427 | 181 |          val _ = warning ("No function definition for datatype " ^ quote tname)
 | 
| 5178 | 182 | in | 
| 183 | (dummy_fns @ fs, defs) | |
| 184 | end | |
| 21064 | 185 | | SOME (fname, ls, fs') => (fs' @ fs, (fname, ls, rec_name, tname) :: defs); | 
| 5178 | 186 | |
| 6359 | 187 | |
| 5178 | 188 | (* make definition *) | 
| 189 | ||
| 25557 | 190 | fun make_def ctxt fixes fs (fname, ls, rec_name, tname) = | 
| 5178 | 191 | let | 
| 30451 
11e5e8bb28f9
corrected type inference of primitive definitions
 haftmann parents: 
30364diff
changeset | 192 | val SOME (var, varT) = get_first (fn ((b, T), mx) => | 
| 
11e5e8bb28f9
corrected type inference of primitive definitions
 haftmann parents: 
30364diff
changeset | 193 | 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 | 194 | val def_name = Thm.def_name (Long_Name.base_name fname); | 
| 30451 
11e5e8bb28f9
corrected type inference of primitive definitions
 haftmann parents: 
30364diff
changeset | 195 |     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 | 196 | (list_comb (Const (rec_name, dummyT), fs @ map Bound (0 :: (length ls downto 1)))) | 
| 
11e5e8bb28f9
corrected type inference of primitive definitions
 haftmann parents: 
30364diff
changeset | 197 | val rhs = singleton (Syntax.check_terms ctxt) | 
| 
11e5e8bb28f9
corrected type inference of primitive definitions
 haftmann parents: 
30364diff
changeset | 198 | (TypeInfer.constrain varT raw_rhs); | 
| 28965 | 199 | in (var, ((Binding.name def_name, []), rhs)) 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 | 
| 25557 | 207 | NONE => primrec_error (quote tname ^ " is not a datatype") | 
| 15531 | 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 | ||
| 25557 | 213 | |
| 214 | (* primrec definition *) | |
| 215 | ||
| 20841 | 216 | local | 
| 217 | ||
| 30451 
11e5e8bb28f9
corrected type inference of primitive definitions
 haftmann parents: 
30364diff
changeset | 218 | fun prove_spec ctxt names rec_rewrites defs eqs = | 
| 25557 | 219 | let | 
| 220 | val rewrites = map mk_meta_eq rec_rewrites @ map (snd o snd) defs; | |
| 221 | fun tac _ = EVERY [rewrite_goals_tac rewrites, rtac refl 1]; | |
| 25566 | 222 |     val _ = message ("Proving equations for primrec function(s) " ^ commas_quote names);
 | 
| 30451 
11e5e8bb28f9
corrected type inference of primitive definitions
 haftmann parents: 
30364diff
changeset | 223 | in map (fn (a, t) => (a, [Goal.prove ctxt [] [] t tac])) eqs end; | 
| 25557 | 224 | |
| 26129 | 225 | fun gen_primrec set_group prep_spec raw_fixes raw_spec lthy = | 
| 5178 | 226 | let | 
| 30487 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30455diff
changeset | 227 | val (fixes, spec) = fst (prep_spec raw_fixes raw_spec lthy); | 
| 25559 | 228 | val eqns = fold_rev (process_eqn (fn v => Variable.is_fixed lthy v | 
| 30223 
24d975352879
renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
 wenzelm parents: 
29866diff
changeset | 229 | orelse exists (fn ((w, _), _) => v = Binding.name_of w) fixes) o snd) spec []; | 
| 25557 | 230 | val tnames = distinct (op =) (map (#1 o snd) eqns); | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
27301diff
changeset | 231 | val dts = find_dts (DatatypePackage.get_datatypes (ProofContext.theory_of lthy)) tnames tnames; | 
| 25557 | 232 |     val main_fns = map (fn (tname, {index, ...}) =>
 | 
| 233 | (index, (fst o the o find_first (fn (_, x) => #1 x = tname)) eqns)) dts; | |
| 22692 | 234 |     val {descr, rec_names, rec_rewrites, ...} =
 | 
| 25557 | 235 | if null dts then primrec_error | 
| 236 |         ("datatypes " ^ commas_quote tnames ^ "\nare not mutually recursive")
 | |
| 18362 
e8b7e0a22727
removed thms 'swap' and 'nth_map' from ML toplevel
 haftmann parents: 
18358diff
changeset | 237 | else snd (hd dts); | 
| 25557 | 238 | val (fnames, fnss) = fold_rev (process_fun descr eqns) main_fns ([], []); | 
| 21064 | 239 | val (fs, defs) = fold_rev (get_fns fnss) (descr ~~ rec_names) ([], []); | 
| 25566 | 240 | val names1 = map snd fnames; | 
| 241 | val names2 = map fst eqns; | |
| 242 | val _ = if gen_eq_set (op =) (names1, names2) then () | |
| 243 |       else primrec_error ("functions " ^ commas_quote names2 ^
 | |
| 25557 | 244 | "\nare not mutually recursive"); | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 245 | val prefix = space_implode "_" (map (Long_Name.base_name o #1) defs); | 
| 30223 
24d975352879
renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
 wenzelm parents: 
29866diff
changeset | 246 | val qualify = Binding.qualify false prefix; | 
| 29807 | 247 | val spec' = (map o apfst) | 
| 248 | (fn (b, attrs) => (qualify b, Code.add_default_eqn_attrib :: attrs)) spec; | |
| 25570 | 249 | val simp_atts = map (Attrib.internal o K) | 
| 29866 
6e93ae65c678
Added Nitpick_Const_Psimp attribute, dropped the 's' in Nitpick_Const_Simps, and killed the Nitpick_Ind_Intros attribute.
 blanchet parents: 
29864diff
changeset | 250 | [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add]; | 
| 8480 
50266d517b0c
Added new theory data slot for primrec equations.
 berghofe parents: 
8432diff
changeset | 251 | in | 
| 25557 | 252 | lthy | 
| 26129 | 253 | |> set_group ? LocalTheory.set_group (serial_string ()) | 
| 25557 | 254 | |> fold_map (LocalTheory.define Thm.definitionK o make_def lthy fixes fs) defs | 
| 26556 | 255 | |-> (fn defs => `(fn ctxt => prove_spec ctxt names1 rec_rewrites defs spec')) | 
| 25557 | 256 | |-> (fn simps => fold_map (LocalTheory.note Thm.theoremK) simps) | 
| 257 | |-> (fn simps' => LocalTheory.note Thm.theoremK | |
| 30435 
e62d6ecab6ad
explicit Binding.qualified_name -- prevents implicitly qualified bstring;
 wenzelm parents: 
30364diff
changeset | 258 | ((qualify (Binding.qualified_name "simps"), simp_atts), maps snd simps')) | 
| 25604 | 259 | |>> snd | 
| 25557 | 260 | end handle PrimrecError (msg, some_eqn) => | 
| 261 |     error ("Primrec definition error:\n" ^ msg ^ (case some_eqn
 | |
| 262 | of SOME eqn => "\nin\n" ^ quote (Syntax.string_of_term lthy eqn) | |
| 263 | | NONE => "")); | |
| 20841 | 264 | |
| 265 | in | |
| 266 | ||
| 30487 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30455diff
changeset | 267 | val add_primrec = gen_primrec false Specification.check_spec; | 
| 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30455diff
changeset | 268 | val add_primrec_cmd = gen_primrec true Specification.read_spec; | 
| 20841 | 269 | |
| 22692 | 270 | end; | 
| 19688 | 271 | |
| 26679 
447f4872b7ee
Added add_primrec_global and add_primrec_overloaded functions (thanks to Markus).
 berghofe parents: 
26556diff
changeset | 272 | fun add_primrec_global fixes specs thy = | 
| 
447f4872b7ee
Added add_primrec_global and add_primrec_overloaded functions (thanks to Markus).
 berghofe parents: 
26556diff
changeset | 273 | let | 
| 
447f4872b7ee
Added add_primrec_global and add_primrec_overloaded functions (thanks to Markus).
 berghofe parents: 
26556diff
changeset | 274 | val lthy = TheoryTarget.init NONE thy; | 
| 
447f4872b7ee
Added add_primrec_global and add_primrec_overloaded functions (thanks to Markus).
 berghofe parents: 
26556diff
changeset | 275 | val (simps, lthy') = add_primrec fixes specs lthy; | 
| 
447f4872b7ee
Added add_primrec_global and add_primrec_overloaded functions (thanks to Markus).
 berghofe parents: 
26556diff
changeset | 276 | val simps' = ProofContext.export lthy' lthy simps; | 
| 28394 | 277 | in (simps', LocalTheory.exit_global lthy') end; | 
| 26679 
447f4872b7ee
Added add_primrec_global and add_primrec_overloaded functions (thanks to Markus).
 berghofe parents: 
26556diff
changeset | 278 | |
| 
447f4872b7ee
Added add_primrec_global and add_primrec_overloaded functions (thanks to Markus).
 berghofe parents: 
26556diff
changeset | 279 | 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 | 280 | let | 
| 
447f4872b7ee
Added add_primrec_global and add_primrec_overloaded functions (thanks to Markus).
 berghofe parents: 
26556diff
changeset | 281 | val lthy = TheoryTarget.overloading ops thy; | 
| 
447f4872b7ee
Added add_primrec_global and add_primrec_overloaded functions (thanks to Markus).
 berghofe parents: 
26556diff
changeset | 282 | val (simps, lthy') = add_primrec fixes specs lthy; | 
| 
447f4872b7ee
Added add_primrec_global and add_primrec_overloaded functions (thanks to Markus).
 berghofe parents: 
26556diff
changeset | 283 | val simps' = ProofContext.export lthy' lthy simps; | 
| 28394 | 284 | in (simps', LocalTheory.exit_global lthy') end; | 
| 26679 
447f4872b7ee
Added add_primrec_global and add_primrec_overloaded functions (thanks to Markus).
 berghofe parents: 
26556diff
changeset | 285 | |
| 5178 | 286 | |
| 6359 | 287 | (* outer syntax *) | 
| 288 | ||
| 17057 | 289 | local structure P = OuterParse and K = OuterKeyword in | 
| 6359 | 290 | |
| 19688 | 291 | val opt_unchecked_name = | 
| 292 |   Scan.optional (P.$$$ "(" |-- P.!!!
 | |
| 293 | (((P.$$$ "unchecked" >> K true) -- Scan.optional P.name "" || | |
| 294 | P.name >> pair false) --| P.$$$ ")")) (false, ""); | |
| 295 | ||
| 25557 | 296 | val old_primrec_decl = | 
| 30223 
24d975352879
renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
 wenzelm parents: 
29866diff
changeset | 297 | opt_unchecked_name -- Scan.repeat1 ((SpecParse.opt_thm_name ":" >> apfst Binding.name_of) -- P.prop); | 
| 6359 | 298 | |
| 30487 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30455diff
changeset | 299 | val primrec_decl = P.opt_target -- P.fixes -- SpecParse.where_alt_specs; | 
| 25557 | 300 | |
| 24867 | 301 | val _ = | 
| 6723 | 302 | OuterSyntax.command "primrec" "define primitive recursive functions on datatypes" K.thy_decl | 
| 30487 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30455diff
changeset | 303 | ((primrec_decl >> (fn ((opt_target, fixes), specs) => | 
| 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30455diff
changeset | 304 | Toplevel.local_theory opt_target (add_primrec_cmd fixes specs #> snd))) | 
| 25557 | 305 | || (old_primrec_decl >> (fn ((unchecked, alt_name), eqns) => | 
| 20176 | 306 | Toplevel.theory (snd o | 
| 30487 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30455diff
changeset | 307 | (if unchecked then OldPrimrecPackage.add_primrec_unchecked else OldPrimrecPackage.add_primrec) | 
| 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30455diff
changeset | 308 | alt_name (map P.triple_swap eqns))))); | 
| 6359 | 309 | |
| 5178 | 310 | end; | 
| 6384 | 311 | |
| 312 | end; |