| author | haftmann | 
| Mon, 26 Apr 2010 11:34:15 +0200 | |
| changeset 36348 | 89c54f51f55a | 
| parent 34232 | 36a2a3029fd3 | 
| child 36519 | 46bf776a81e0 | 
| permissions | -rw-r--r-- | 
| 33098 | 1 | (* Title: HOL/Tools/Function/fun.ML | 
| 2 | Author: Alexander Krauss, TU Muenchen | |
| 3 | ||
| 4 | Sequential mode for function definitions | |
| 5 | Command "fun" for fully automated function definitions | |
| 6 | *) | |
| 7 | ||
| 8 | signature FUNCTION_FUN = | |
| 9 | sig | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 10 | val add_fun : Function_Common.function_config -> | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 11 | (binding * typ option * mixfix) list -> (Attrib.binding * term) list -> | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 12 | bool -> local_theory -> Proof.context | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 13 | val add_fun_cmd : Function_Common.function_config -> | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 14 | (binding * string option * mixfix) list -> (Attrib.binding * string) list -> | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 15 | bool -> local_theory -> Proof.context | 
| 33098 | 16 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 17 | val setup : theory -> theory | 
| 33098 | 18 | end | 
| 19 | ||
| 20 | structure Function_Fun : FUNCTION_FUN = | |
| 21 | struct | |
| 22 | ||
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33098diff
changeset | 23 | open Function_Lib | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33098diff
changeset | 24 | open Function_Common | 
| 33098 | 25 | |
| 26 | ||
| 27 | fun check_pats ctxt geq = | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 28 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 29 | fun err str = error (cat_lines ["Malformed definition:", | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 30 | str ^ " not allowed in sequential mode.", | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 31 | Syntax.string_of_term ctxt geq]) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 32 | val thy = ProofContext.theory_of ctxt | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 33 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 34 | fun check_constr_pattern (Bound _) = () | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 35 | | check_constr_pattern t = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 36 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 37 | val (hd, args) = strip_comb t | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 38 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 39 | (((case Datatype.info_of_constr thy (dest_Const hd) of | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 40 | SOME _ => () | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 41 | | NONE => err "Non-constructor pattern") | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 42 |           handle TERM ("dest_Const", _) => err "Non-constructor patterns");
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 43 | map check_constr_pattern args; | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 44 | ()) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 45 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 46 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 47 | val (_, qs, gs, args, _) = split_def ctxt geq | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 48 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 49 | val _ = if not (null gs) then err "Conditional equations" else () | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 50 | val _ = map check_constr_pattern args | 
| 
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 | (* just count occurrences to check linearity *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 53 | 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 | 54 | then err "Nonlinear patterns" else () | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 55 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 56 | () | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 57 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 58 | |
| 33098 | 59 | val by_pat_completeness_auto = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 60 | Proof.global_future_terminal_proof | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 61 | (Method.Basic Pat_Completeness.pat_completeness, | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 62 |      SOME (Method.Source_i (Args.src (("HOL.auto", []), Position.none))))
 | 
| 33098 | 63 | |
| 64 | fun termination_by method int = | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 65 | Function.termination_proof NONE | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 66 | #> Proof.global_future_terminal_proof (Method.Basic method, NONE) int | 
| 33098 | 67 | |
| 68 | fun mk_catchall fixes arity_of = | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 69 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 70 | fun mk_eqn ((fname, fT), _) = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 71 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 72 | val n = arity_of fname | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 73 | val (argTs, rT) = chop n (binder_types fT) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 74 | |> apsnd (fn Ts => Ts ---> body_type fT) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 75 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 76 | val qs = map Free (Name.invent_list [] "a" n ~~ argTs) | 
| 
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 | HOLogic.mk_eq(list_comb (Free (fname, fT), qs), | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 79 |           Const ("HOL.undefined", rT))
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 80 | |> HOLogic.mk_Trueprop | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 81 | |> fold_rev Logic.all qs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 82 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 83 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 84 | map mk_eqn fixes | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 85 | end | 
| 33098 | 86 | |
| 87 | fun add_catchall ctxt fixes spec = | |
| 88 | let val fqgars = map (split_def ctxt) spec | |
| 89 | val arity_of = map (fn (fname,_,_,args,_) => (fname, length args)) fqgars | |
| 90 | |> AList.lookup (op =) #> the | |
| 91 | in | |
| 92 | spec @ mk_catchall fixes arity_of | |
| 93 | end | |
| 94 | ||
| 95 | fun warn_if_redundant ctxt origs tss = | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 96 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 97 | fun msg t = "Ignoring redundant equation: " ^ quote (Syntax.string_of_term ctxt t) | 
| 33098 | 98 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 99 | val (tss', _) = chop (length origs) tss | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 100 | fun check (t, []) = (warning (msg t); []) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 101 | | check (t, s) = s | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 102 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 103 | (map check (origs ~~ tss'); tss) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 104 | end | 
| 33098 | 105 | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33098diff
changeset | 106 | 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 | 107 | if sequential then | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 108 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 109 | val (bnds, eqss) = split_list spec | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 110 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 111 | val eqs = map the_single eqss | 
| 33098 | 112 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 113 | val feqs = eqs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 114 | |> tap (check_defs ctxt fixes) (* Standard checks *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 115 | |> tap (map (check_pats ctxt)) (* More checks for sequential mode *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 116 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 117 | val compleqs = add_catchall ctxt fixes feqs (* Completion *) | 
| 33098 | 118 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 119 | val spliteqs = warn_if_redundant ctxt feqs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 120 | (Function_Split.split_all_equations ctxt compleqs) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 121 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 122 | fun restore_spec thms = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 123 | bnds ~~ take (length bnds) (unflat spliteqs thms) | 
| 33098 | 124 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 125 | val spliteqs' = flat (take (length bnds) spliteqs) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 126 | val fnames = map (fst o fst) fixes | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 127 | val indices = map (fn eq => find_index (curry op = (fname_of eq)) fnames) spliteqs' | 
| 33098 | 128 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 129 | 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 | 130 | |> map (map snd) | 
| 33098 | 131 | |
| 132 | ||
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 133 | val bnds' = bnds @ replicate (length spliteqs - length bnds) Attrib.empty_binding | 
| 33098 | 134 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 135 | (* using theorem names for case name currently disabled *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 136 | 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 | 137 | (bnds' ~~ spliteqs) |> flat | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 138 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 139 | (flat spliteqs, restore_spec, sort, case_names) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 140 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 141 | else | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 142 | Function_Common.empty_preproc check_defs config ctxt fixes spec | 
| 33098 | 143 | |
| 144 | val setup = | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33098diff
changeset | 145 | Context.theory_map (Function_Common.set_preproc sequential_preproc) | 
| 33098 | 146 | |
| 147 | ||
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33098diff
changeset | 148 | val fun_config = FunctionConfig { sequential=true, default="%x. undefined" (*FIXME dynamic scoping*), 
 | 
| 33101 | 149 | domintros=false, partials=false, tailrec=false } | 
| 33098 | 150 | |
| 151 | fun gen_fun add config fixes statements int lthy = | |
| 33726 
0878aecbf119
eliminated slightly odd name space grouping -- now managed by Isar toplevel;
 wenzelm parents: 
33671diff
changeset | 152 | lthy | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 153 | |> add fixes statements config | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 154 | |> by_pat_completeness_auto int | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 155 | |> Local_Theory.restore | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33957diff
changeset | 156 | |> termination_by (Function_Common.get_termination_prover lthy) int | 
| 33098 | 157 | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33098diff
changeset | 158 | val add_fun = gen_fun Function.add_function | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33098diff
changeset | 159 | val add_fun_cmd = gen_fun Function.add_function_cmd | 
| 33098 | 160 | |
| 161 | ||
| 162 | ||
| 163 | local structure P = OuterParse and K = OuterKeyword in | |
| 164 | ||
| 165 | val _ = | |
| 166 | OuterSyntax.local_theory' "fun" "define general recursive functions (short version)" K.thy_decl | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33098diff
changeset | 167 | (function_parser fun_config | 
| 33098 | 168 | >> (fn ((config, fixes), statements) => add_fun_cmd config fixes statements)); | 
| 169 | ||
| 170 | end | |
| 171 | ||
| 172 | end |