| author | wenzelm | 
| Tue, 31 Dec 2013 14:29:16 +0100 | |
| changeset 54883 | dd04a8b654fc | 
| parent 54407 | e95831757903 | 
| child 56254 | a2dd9200854d | 
| 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 -> | 
| 
a294e4ebe0a3
clarified signature; simpler implementation in terms of function's tactic interface
 krauss parents: 
36521diff
changeset | 12 | (Attrib.binding * term) list -> Function_Common.function_config -> | 
| 
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 -> | 
| 
a294e4ebe0a3
clarified signature; simpler implementation in terms of function's tactic interface
 krauss parents: 
36521diff
changeset | 15 | (Attrib.binding * string) list -> 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 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 18 | val setup : theory -> theory | 
| 33098 | 19 | end | 
| 20 | ||
| 21 | structure Function_Fun : FUNCTION_FUN = | |
| 22 | struct | |
| 23 | ||
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33098diff
changeset | 24 | open Function_Lib | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33098diff
changeset | 25 | open Function_Common | 
| 33098 | 26 | |
| 27 | ||
| 28 | fun check_pats ctxt geq = | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 29 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 30 | fun err str = error (cat_lines ["Malformed definition:", | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 31 | str ^ " not allowed in sequential mode.", | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 32 | Syntax.string_of_term ctxt geq]) | 
| 42361 | 33 | val thy = Proof_Context.theory_of ctxt | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 34 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 35 | fun check_constr_pattern (Bound _) = () | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 36 | | check_constr_pattern t = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 37 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 38 | val (hd, args) = strip_comb t | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 39 | in | 
| 54407 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
49965diff
changeset | 40 | (case hd of | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
49965diff
changeset | 41 | Const (hd_s, hd_T) => | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
49965diff
changeset | 42 | (case body_type hd_T of | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
49965diff
changeset | 43 | Type (Tname, _) => | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
49965diff
changeset | 44 | (case Ctr_Sugar.ctr_sugar_of ctxt Tname of | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
49965diff
changeset | 45 |               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 | 46 | | NONE => false) | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
49965diff
changeset | 47 | | _ => false) | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
49965diff
changeset | 48 | | _ => false) orelse err "Non-constructor pattern"; | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
49965diff
changeset | 49 | map check_constr_pattern args; | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
49965diff
changeset | 50 | () | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 51 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 52 | |
| 39276 
2ad95934521f
improved error message for common mistake (fun f where "g x = ...")
 krauss parents: 
36960diff
changeset | 53 | 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 | 54 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 55 | val _ = if not (null gs) then err "Conditional equations" else () | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 56 | val _ = map check_constr_pattern args | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 57 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 58 | (* just count occurrences to check linearity *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 59 | 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 | 60 | then err "Nonlinear patterns" else () | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 61 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 62 | () | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 63 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 64 | |
| 33098 | 65 | fun mk_catchall fixes arity_of = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 66 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 67 | fun mk_eqn ((fname, fT), _) = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 68 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 69 | val n = arity_of fname | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 70 | val (argTs, rT) = chop n (binder_types fT) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 71 | |> apsnd (fn Ts => Ts ---> body_type fT) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 72 | |
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43323diff
changeset | 73 | 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 | 74 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 75 | HOLogic.mk_eq(list_comb (Free (fname, fT), qs), | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 76 |           Const ("HOL.undefined", rT))
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 77 | |> HOLogic.mk_Trueprop | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 78 | |> fold_rev Logic.all qs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 79 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 80 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 81 | map mk_eqn fixes | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 82 | end | 
| 33098 | 83 | |
| 84 | fun add_catchall ctxt fixes spec = | |
| 39276 
2ad95934521f
improved error message for common mistake (fun f where "g x = ...")
 krauss parents: 
36960diff
changeset | 85 | let val fqgars = map (split_def ctxt (K true)) spec | 
| 33098 | 86 | val arity_of = map (fn (fname,_,_,args,_) => (fname, length args)) fqgars | 
| 87 | |> AList.lookup (op =) #> the | |
| 88 | in | |
| 89 | spec @ mk_catchall fixes arity_of | |
| 90 | end | |
| 91 | ||
| 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 | 92 | fun further_checks ctxt origs tss = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 93 | 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 | 94 | 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 | 95 | 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 | 96 | fun warn_missing strs = | 
| 43277 | 97 |       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 | 98 | |
| 
fcb6250bf6b4
fun command produces warning when patterns are incomplete (somewhat analogous to primrec)
 krauss parents: 
42793diff
changeset | 99 | val (tss', added) = chop (length origs) tss | 
| 33098 | 100 | |
| 42947 
fcb6250bf6b4
fun command produces warning when patterns are incomplete (somewhat analogous to primrec)
 krauss parents: 
42793diff
changeset | 101 | val _ = case chop 3 (flat added) of | 
| 
fcb6250bf6b4
fun command produces warning when patterns are incomplete (somewhat analogous to primrec)
 krauss parents: 
42793diff
changeset | 102 | ([], []) => () | 
| 
fcb6250bf6b4
fun command produces warning when patterns are incomplete (somewhat analogous to primrec)
 krauss parents: 
42793diff
changeset | 103 | | (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 | 104 | | (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 | 105 |          @ ["(" ^ string_of_int (length rest) ^ " more)"])
 | 
| 
fcb6250bf6b4
fun command produces warning when patterns are incomplete (somewhat analogous to primrec)
 krauss parents: 
42793diff
changeset | 106 | |
| 
fcb6250bf6b4
fun command produces warning when patterns are incomplete (somewhat analogous to primrec)
 krauss parents: 
42793diff
changeset | 107 | 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 | 108 | |> 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 | 109 | in | 
| 42947 
fcb6250bf6b4
fun command produces warning when patterns are incomplete (somewhat analogous to primrec)
 krauss parents: 
42793diff
changeset | 110 | () | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 111 | end | 
| 33098 | 112 | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33098diff
changeset | 113 | 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 | 114 | if sequential then | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 115 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 116 | val (bnds, eqss) = split_list spec | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 117 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 118 | val eqs = map the_single eqss | 
| 33098 | 119 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 120 | val feqs = eqs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 121 | |> tap (check_defs ctxt fixes) (* Standard checks *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 122 | |> tap (map (check_pats ctxt)) (* More checks for sequential mode *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 123 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 124 | val compleqs = add_catchall ctxt fixes feqs (* Completion *) | 
| 33098 | 125 | |
| 42947 
fcb6250bf6b4
fun command produces warning when patterns are incomplete (somewhat analogous to primrec)
 krauss parents: 
42793diff
changeset | 126 | 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 | 127 | |> tap (further_checks ctxt feqs) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 128 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 129 | fun restore_spec thms = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 130 | bnds ~~ take (length bnds) (unflat spliteqs thms) | 
| 33098 | 131 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 132 | val spliteqs' = flat (take (length bnds) spliteqs) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 133 | val fnames = map (fst o fst) fixes | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 134 | val indices = map (fn eq => find_index (curry op = (fname_of eq)) fnames) spliteqs' | 
| 33098 | 135 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 136 | 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 | 137 | |> map (map snd) | 
| 33098 | 138 | |
| 139 | ||
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 140 | val bnds' = bnds @ replicate (length spliteqs - length bnds) Attrib.empty_binding | 
| 33098 | 141 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 142 | (* using theorem names for case name currently disabled *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 143 | 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 | 144 | (bnds' ~~ spliteqs) |> flat | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 145 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 146 | (flat spliteqs, restore_spec, sort, case_names) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 147 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 148 | else | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 149 | Function_Common.empty_preproc check_defs config ctxt fixes spec | 
| 33098 | 150 | |
| 151 | val setup = | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33098diff
changeset | 152 | Context.theory_map (Function_Common.set_preproc sequential_preproc) | 
| 33098 | 153 | |
| 154 | ||
| 41417 | 155 | 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 | 156 | domintros=false, partials=false } | 
| 33098 | 157 | |
| 44239 
47ecd30e018d
less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
 wenzelm parents: 
44237diff
changeset | 158 | fun gen_add_fun add lthy = | 
| 36523 
a294e4ebe0a3
clarified signature; simpler implementation in terms of function's tactic interface
 krauss parents: 
36521diff
changeset | 159 | let | 
| 
a294e4ebe0a3
clarified signature; simpler implementation in terms of function's tactic interface
 krauss parents: 
36521diff
changeset | 160 | fun pat_completeness_auto ctxt = | 
| 
a294e4ebe0a3
clarified signature; simpler implementation in terms of function's tactic interface
 krauss parents: 
36521diff
changeset | 161 | Pat_Completeness.pat_completeness_tac ctxt 1 | 
| 42793 | 162 | THEN auto_tac ctxt | 
| 36523 
a294e4ebe0a3
clarified signature; simpler implementation in terms of function's tactic interface
 krauss parents: 
36521diff
changeset | 163 | fun prove_termination lthy = | 
| 
a294e4ebe0a3
clarified signature; simpler implementation in terms of function's tactic interface
 krauss parents: 
36521diff
changeset | 164 | Function.prove_termination NONE | 
| 49965 | 165 | (Function_Common.get_termination_prover lthy) lthy | 
| 36523 
a294e4ebe0a3
clarified signature; simpler implementation in terms of function's tactic interface
 krauss parents: 
36521diff
changeset | 166 | in | 
| 
a294e4ebe0a3
clarified signature; simpler implementation in terms of function's tactic interface
 krauss parents: 
36521diff
changeset | 167 | lthy | 
| 44239 
47ecd30e018d
less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
 wenzelm parents: 
44237diff
changeset | 168 | |> add pat_completeness_auto |> snd | 
| 36547 
2a9d0ec8c10d
return updated info record after termination proof
 krauss parents: 
36523diff
changeset | 169 | |> prove_termination |> snd | 
| 36523 
a294e4ebe0a3
clarified signature; simpler implementation in terms of function's tactic interface
 krauss parents: 
36521diff
changeset | 170 | end | 
| 33098 | 171 | |
| 44239 
47ecd30e018d
less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
 wenzelm parents: 
44237diff
changeset | 172 | 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 | 173 | fun add_fun_cmd a b c int = gen_add_fun (fn tac => Function.add_function_cmd a b c tac int) | 
| 33098 | 174 | |
| 175 | ||
| 176 | ||
| 177 | val _ = | |
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
45639diff
changeset | 178 |   Outer_Syntax.local_theory' @{command_spec "fun"}
 | 
| 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
45639diff
changeset | 179 | "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 | 180 | (function_parser fun_config | 
| 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
45639diff
changeset | 181 | >> (fn ((config, fixes), statements) => add_fun_cmd fixes statements config)) | 
| 33098 | 182 | |
| 183 | end |