| author | paulson <lp15@cam.ac.uk> | 
| Thu, 19 Jul 2018 23:23:10 +0200 | |
| changeset 68664 | bd0df72c16d5 | 
| parent 67149 | e61557884799 | 
| child 81507 | 08574da77b4a | 
| permissions | -rw-r--r-- | 
| 33098 | 1 | (* Title: HOL/Tools/Function/fun.ML | 
| 2 | Author: Alexander Krauss, TU Muenchen | |
| 3 | ||
| 41114 | 4 | Command "fun": Function definitions with pattern splitting/completion | 
| 5 | and automated termination proofs. | |
| 33098 | 6 | *) | 
| 7 | ||
| 8 | signature FUNCTION_FUN = | |
| 9 | sig | |
| 44237 
2a2040c9d898
export Function_Fun.fun_config for user convenience;
 wenzelm parents: 
43329diff
changeset | 10 | val fun_config : Function_Common.function_config | 
| 36523 
a294e4ebe0a3
clarified signature; simpler implementation in terms of function's tactic interface
 krauss parents: 
36521diff
changeset | 11 | val add_fun : (binding * typ option * mixfix) list -> | 
| 63064 
2f18172214c8
support 'assumes' in specifications, e.g. 'definition', 'inductive';
 wenzelm parents: 
60682diff
changeset | 12 | Specification.multi_specs -> Function_Common.function_config -> | 
| 36523 
a294e4ebe0a3
clarified signature; simpler implementation in terms of function's tactic interface
 krauss parents: 
36521diff
changeset | 13 | local_theory -> Proof.context | 
| 
a294e4ebe0a3
clarified signature; simpler implementation in terms of function's tactic interface
 krauss parents: 
36521diff
changeset | 14 | val add_fun_cmd : (binding * string option * mixfix) list -> | 
| 63064 
2f18172214c8
support 'assumes' in specifications, e.g. 'definition', 'inductive';
 wenzelm parents: 
60682diff
changeset | 15 | Specification.multi_specs_cmd -> Function_Common.function_config -> | 
| 44239 
47ecd30e018d
less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
 wenzelm parents: 
44237diff
changeset | 16 | bool -> local_theory -> Proof.context | 
| 33098 | 17 | end | 
| 18 | ||
| 19 | structure Function_Fun : FUNCTION_FUN = | |
| 20 | struct | |
| 21 | ||
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33098diff
changeset | 22 | open Function_Lib | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33098diff
changeset | 23 | open Function_Common | 
| 33098 | 24 | |
| 25 | ||
| 26 | fun check_pats ctxt geq = | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 27 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 28 | fun err str = error (cat_lines ["Malformed definition:", | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 29 | str ^ " not allowed in sequential mode.", | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 30 | Syntax.string_of_term ctxt geq]) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 31 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 32 | fun check_constr_pattern (Bound _) = () | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 33 | | check_constr_pattern t = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 34 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 35 | val (hd, args) = strip_comb t | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 36 | in | 
| 54407 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
49965diff
changeset | 37 | (case hd of | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
49965diff
changeset | 38 | Const (hd_s, hd_T) => | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
49965diff
changeset | 39 | (case body_type hd_T of | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
49965diff
changeset | 40 | Type (Tname, _) => | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
49965diff
changeset | 41 | (case Ctr_Sugar.ctr_sugar_of ctxt Tname of | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
49965diff
changeset | 42 |               SOME {ctrs, ...} => exists (fn Const (s, _) => s = hd_s) ctrs
 | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
49965diff
changeset | 43 | | NONE => false) | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
49965diff
changeset | 44 | | _ => false) | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
49965diff
changeset | 45 | | _ => false) orelse err "Non-constructor pattern"; | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
49965diff
changeset | 46 | map check_constr_pattern args; | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
49965diff
changeset | 47 | () | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 48 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 49 | |
| 39276 
2ad95934521f
improved error message for common mistake (fun f where "g x = ...")
 krauss parents: 
36960diff
changeset | 50 | val (_, qs, gs, args, _) = split_def ctxt (K true) geq | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 51 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 52 | val _ = if not (null gs) then err "Conditional equations" else () | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 53 | val _ = map check_constr_pattern args | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 54 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 55 | (* just count occurrences to check linearity *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 56 | val _ = if fold (fold_aterms (fn Bound _ => Integer.add 1 | _ => I)) args 0 > length qs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 57 | then err "Nonlinear patterns" else () | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 58 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 59 | () | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 60 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 61 | |
| 33098 | 62 | fun mk_catchall fixes arity_of = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 63 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 64 | fun mk_eqn ((fname, fT), _) = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 65 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 66 | val n = arity_of fname | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 67 | val (argTs, rT) = chop n (binder_types fT) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 68 | |> apsnd (fn Ts => Ts ---> body_type fT) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 69 | |
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43323diff
changeset | 70 | val qs = map Free (Name.invent Name.context "a" n ~~ argTs) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 71 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 72 | HOLogic.mk_eq(list_comb (Free (fname, fT), qs), | 
| 67149 | 73 | Const (\<^const_name>\<open>undefined\<close>, rT)) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 74 | |> HOLogic.mk_Trueprop | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 75 | |> fold_rev Logic.all qs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 76 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 77 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 78 | map mk_eqn fixes | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 79 | end | 
| 33098 | 80 | |
| 81 | fun add_catchall ctxt fixes spec = | |
| 39276 
2ad95934521f
improved error message for common mistake (fun f where "g x = ...")
 krauss parents: 
36960diff
changeset | 82 | let val fqgars = map (split_def ctxt (K true)) spec | 
| 33098 | 83 | val arity_of = map (fn (fname,_,_,args,_) => (fname, length args)) fqgars | 
| 84 | |> AList.lookup (op =) #> the | |
| 85 | in | |
| 86 | spec @ mk_catchall fixes arity_of | |
| 87 | end | |
| 88 | ||
| 48099 
e7e647949c95
fun command: produce hard failure when equations do not contribute to the specification (i.e., are covered by preceding clauses), to avoid confusing inexperienced users
 krauss parents: 
46961diff
changeset | 89 | fun further_checks ctxt origs tss = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 90 | let | 
| 48099 
e7e647949c95
fun command: produce hard failure when equations do not contribute to the specification (i.e., are covered by preceding clauses), to avoid confusing inexperienced users
 krauss parents: 
46961diff
changeset | 91 | fun fail_redundant t = | 
| 
e7e647949c95
fun command: produce hard failure when equations do not contribute to the specification (i.e., are covered by preceding clauses), to avoid confusing inexperienced users
 krauss parents: 
46961diff
changeset | 92 | error (cat_lines ["Equation is redundant (covered by preceding clauses):", Syntax.string_of_term ctxt t]) | 
| 42947 
fcb6250bf6b4
fun command produces warning when patterns are incomplete (somewhat analogous to primrec)
 krauss parents: 
42793diff
changeset | 93 | fun warn_missing strs = | 
| 43277 | 94 |       warning (cat_lines ("Missing patterns in function definition:" :: strs))
 | 
| 42947 
fcb6250bf6b4
fun command produces warning when patterns are incomplete (somewhat analogous to primrec)
 krauss parents: 
42793diff
changeset | 95 | |
| 
fcb6250bf6b4
fun command produces warning when patterns are incomplete (somewhat analogous to primrec)
 krauss parents: 
42793diff
changeset | 96 | val (tss', added) = chop (length origs) tss | 
| 33098 | 97 | |
| 42947 
fcb6250bf6b4
fun command produces warning when patterns are incomplete (somewhat analogous to primrec)
 krauss parents: 
42793diff
changeset | 98 | val _ = case chop 3 (flat added) of | 
| 
fcb6250bf6b4
fun command produces warning when patterns are incomplete (somewhat analogous to primrec)
 krauss parents: 
42793diff
changeset | 99 | ([], []) => () | 
| 
fcb6250bf6b4
fun command produces warning when patterns are incomplete (somewhat analogous to primrec)
 krauss parents: 
42793diff
changeset | 100 | | (eqs, []) => warn_missing (map (Syntax.string_of_term ctxt) eqs) | 
| 
fcb6250bf6b4
fun command produces warning when patterns are incomplete (somewhat analogous to primrec)
 krauss parents: 
42793diff
changeset | 101 | | (eqs, rest) => warn_missing (map (Syntax.string_of_term ctxt) eqs | 
| 
fcb6250bf6b4
fun command produces warning when patterns are incomplete (somewhat analogous to primrec)
 krauss parents: 
42793diff
changeset | 102 |          @ ["(" ^ string_of_int (length rest) ^ " more)"])
 | 
| 
fcb6250bf6b4
fun command produces warning when patterns are incomplete (somewhat analogous to primrec)
 krauss parents: 
42793diff
changeset | 103 | |
| 
fcb6250bf6b4
fun command produces warning when patterns are incomplete (somewhat analogous to primrec)
 krauss parents: 
42793diff
changeset | 104 | val _ = (origs ~~ tss') | 
| 48099 
e7e647949c95
fun command: produce hard failure when equations do not contribute to the specification (i.e., are covered by preceding clauses), to avoid confusing inexperienced users
 krauss parents: 
46961diff
changeset | 105 | |> map (fn (t, ts) => if null ts then fail_redundant t else ()) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 106 | in | 
| 42947 
fcb6250bf6b4
fun command produces warning when patterns are incomplete (somewhat analogous to primrec)
 krauss parents: 
42793diff
changeset | 107 | () | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 108 | end | 
| 33098 | 109 | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33098diff
changeset | 110 | fun sequential_preproc (config as FunctionConfig {sequential, ...}) ctxt fixes spec =
 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 111 | if sequential then | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 112 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 113 | val (bnds, eqss) = split_list spec | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 114 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 115 | val eqs = map the_single eqss | 
| 33098 | 116 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 117 | val feqs = eqs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 118 | |> tap (check_defs ctxt fixes) (* Standard checks *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 119 | |> tap (map (check_pats ctxt)) (* More checks for sequential mode *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 120 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 121 | val compleqs = add_catchall ctxt fixes feqs (* Completion *) | 
| 33098 | 122 | |
| 42947 
fcb6250bf6b4
fun command produces warning when patterns are incomplete (somewhat analogous to primrec)
 krauss parents: 
42793diff
changeset | 123 | val spliteqs = Function_Split.split_all_equations ctxt compleqs | 
| 48099 
e7e647949c95
fun command: produce hard failure when equations do not contribute to the specification (i.e., are covered by preceding clauses), to avoid confusing inexperienced users
 krauss parents: 
46961diff
changeset | 124 | |> tap (further_checks ctxt feqs) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 125 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 126 | fun restore_spec thms = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 127 | bnds ~~ take (length bnds) (unflat spliteqs thms) | 
| 33098 | 128 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 129 | val spliteqs' = flat (take (length bnds) spliteqs) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 130 | val fnames = map (fst o fst) fixes | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 131 | val indices = map (fn eq => find_index (curry op = (fname_of eq)) fnames) spliteqs' | 
| 33098 | 132 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 133 | fun sort xs = partition_list (fn i => fn (j,_) => i = j) 0 (length fnames - 1) (indices ~~ xs) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 134 | |> map (map snd) | 
| 33098 | 135 | |
| 136 | ||
| 63352 | 137 | val bnds' = bnds @ replicate (length spliteqs - length bnds) Binding.empty_atts | 
| 33098 | 138 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 139 | (* using theorem names for case name currently disabled *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 140 | val case_names = map_index (fn (i, (_, es)) => mk_case_names i "" (length es)) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 141 | (bnds' ~~ spliteqs) |> flat | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 142 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 143 | (flat spliteqs, restore_spec, sort, case_names) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 144 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 145 | else | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 146 | Function_Common.empty_preproc check_defs config ctxt fixes spec | 
| 33098 | 147 | |
| 58826 | 148 | val _ = Theory.setup (Context.theory_map (Function_Common.set_preproc sequential_preproc)) | 
| 149 | ||
| 33098 | 150 | |
| 151 | ||
| 41417 | 152 | val fun_config = FunctionConfig { sequential=true, default=NONE,
 | 
| 41846 
b368a7aee46a
removed support for tail-recursion from function package (now implemented by partial_function)
 krauss parents: 
41417diff
changeset | 153 | domintros=false, partials=false } | 
| 33098 | 154 | |
| 44239 
47ecd30e018d
less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
 wenzelm parents: 
44237diff
changeset | 155 | fun gen_add_fun add lthy = | 
| 36523 
a294e4ebe0a3
clarified signature; simpler implementation in terms of function's tactic interface
 krauss parents: 
36521diff
changeset | 156 | let | 
| 
a294e4ebe0a3
clarified signature; simpler implementation in terms of function's tactic interface
 krauss parents: 
36521diff
changeset | 157 | fun pat_completeness_auto ctxt = | 
| 
a294e4ebe0a3
clarified signature; simpler implementation in terms of function's tactic interface
 krauss parents: 
36521diff
changeset | 158 | Pat_Completeness.pat_completeness_tac ctxt 1 | 
| 42793 | 159 | THEN auto_tac ctxt | 
| 36523 
a294e4ebe0a3
clarified signature; simpler implementation in terms of function's tactic interface
 krauss parents: 
36521diff
changeset | 160 | fun prove_termination lthy = | 
| 60682 
5a6cd2560549
have the installed termination prover take a 'quiet' flag
 blanchet parents: 
59936diff
changeset | 161 | Function.prove_termination NONE (Function_Common.termination_prover_tac false lthy) lthy | 
| 36523 
a294e4ebe0a3
clarified signature; simpler implementation in terms of function's tactic interface
 krauss parents: 
36521diff
changeset | 162 | in | 
| 
a294e4ebe0a3
clarified signature; simpler implementation in terms of function's tactic interface
 krauss parents: 
36521diff
changeset | 163 | lthy | 
| 44239 
47ecd30e018d
less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
 wenzelm parents: 
44237diff
changeset | 164 | |> add pat_completeness_auto |> snd | 
| 36547 
2a9d0ec8c10d
return updated info record after termination proof
 krauss parents: 
36523diff
changeset | 165 | |> prove_termination |> snd | 
| 36523 
a294e4ebe0a3
clarified signature; simpler implementation in terms of function's tactic interface
 krauss parents: 
36521diff
changeset | 166 | end | 
| 33098 | 167 | |
| 44239 
47ecd30e018d
less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
 wenzelm parents: 
44237diff
changeset | 168 | fun add_fun a b c = gen_add_fun (Function.add_function a b c) | 
| 
47ecd30e018d
less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
 wenzelm parents: 
44237diff
changeset | 169 | fun add_fun_cmd a b c int = gen_add_fun (fn tac => Function.add_function_cmd a b c tac int) | 
| 33098 | 170 | |
| 171 | ||
| 172 | ||
| 173 | val _ = | |
| 67149 | 174 | Outer_Syntax.local_theory' \<^command_keyword>\<open>fun\<close> | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
45639diff
changeset | 175 | "define general recursive functions (short version)" | 
| 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
45639diff
changeset | 176 | (function_parser fun_config | 
| 63183 | 177 | >> (fn (config, (fixes, specs)) => add_fun_cmd fixes specs config)) | 
| 33098 | 178 | |
| 179 | end |