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