| author | blanchet | 
| Mon, 12 Aug 2013 15:25:17 +0200 | |
| changeset 52968 | 2b430bbb5a1a | 
| parent 52384 | 80c00a851de5 | 
| child 53603 | 59ef06cda7b9 | 
| permissions | -rw-r--r-- | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 1 | (* Title: HOL/Tools/Function/function_common.ML | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 2 | Author: Alexander Krauss, TU Muenchen | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 3 | |
| 41114 | 4 | Common definitions and other infrastructure for the function package. | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 5 | *) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 6 | |
| 34230 | 7 | signature FUNCTION_DATA = | 
| 8 | sig | |
| 9 | ||
| 10 | type info = | |
| 11 |  {is_partial : bool,
 | |
| 12 | defname : string, | |
| 13 | (* contains no logical entities: invariant under morphisms: *) | |
| 14 | add_simps : (binding -> binding) -> string -> (binding -> binding) -> | |
| 15 | Attrib.src list -> thm list -> local_theory -> thm list * local_theory, | |
| 16 | case_names : string list, | |
| 17 | fs : term list, | |
| 18 | R : term, | |
| 52384 | 19 | dom: term, | 
| 34230 | 20 | psimps: thm list, | 
| 21 | pinducts: thm list, | |
| 34231 | 22 | simps : thm list option, | 
| 23 | inducts : thm list option, | |
| 52383 | 24 | termination : thm, | 
| 25 | cases : thm} | |
| 34230 | 26 | |
| 27 | end | |
| 28 | ||
| 29 | structure Function_Data : FUNCTION_DATA = | |
| 30 | struct | |
| 31 | ||
| 32 | type info = | |
| 33 |  {is_partial : bool,
 | |
| 34 | defname : string, | |
| 35 | (* contains no logical entities: invariant under morphisms: *) | |
| 36 | add_simps : (binding -> binding) -> string -> (binding -> binding) -> | |
| 37 | Attrib.src list -> thm list -> local_theory -> thm list * local_theory, | |
| 38 | case_names : string list, | |
| 39 | fs : term list, | |
| 40 | R : term, | |
| 52384 | 41 | dom: term, | 
| 34230 | 42 | psimps: thm list, | 
| 43 | pinducts: thm list, | |
| 34231 | 44 | simps : thm list option, | 
| 45 | inducts : thm list option, | |
| 52383 | 46 | termination : thm, | 
| 47 | cases : thm} | |
| 34230 | 48 | |
| 49 | end | |
| 50 | ||
| 49965 | 51 | signature FUNCTION_COMMON = | 
| 52 | sig | |
| 53 | include FUNCTION_DATA | |
| 54 | val profile : bool Unsynchronized.ref | |
| 55 |   val PROFILE : string -> ('a -> 'b) -> 'a -> 'b
 | |
| 56 | val mk_acc : typ -> term -> term | |
| 57 | val function_name : string -> string | |
| 58 | val graph_name : string -> string | |
| 59 | val rel_name : string -> string | |
| 60 | val dom_name : string -> string | |
| 61 | val apply_termination_rule : Proof.context -> int -> tactic | |
| 62 | datatype function_result = FunctionResult of | |
| 63 |    {fs: term list,
 | |
| 64 | G: term, | |
| 65 | R: term, | |
| 52384 | 66 | dom: term, | 
| 49965 | 67 | psimps : thm list, | 
| 68 | simple_pinducts : thm list, | |
| 69 | cases : thm, | |
| 70 | termination : thm, | |
| 71 | domintros : thm list option} | |
| 72 | val transform_function_data : info -> morphism -> info | |
| 73 | val get_function : Proof.context -> (term * info) Item_Net.T | |
| 74 | val import_function_data : term -> Proof.context -> info option | |
| 75 | val import_last_function : Proof.context -> info option | |
| 76 | val add_function_data : info -> Context.generic -> Context.generic | |
| 77 | structure Termination_Simps: NAMED_THMS | |
| 78 | val set_termination_prover : (Proof.context -> tactic) -> Context.generic -> Context.generic | |
| 79 | val get_termination_prover : Proof.context -> tactic | |
| 49979 
4de92b4aa74a
added function store_termination_rule to the signature, as it is used in Nominal2
 Christian Urban <urbanc@in.tum.de> parents: 
49966diff
changeset | 80 | val store_termination_rule : thm -> Context.generic -> Context.generic | 
| 49965 | 81 | datatype function_config = FunctionConfig of | 
| 82 |    {sequential: bool,
 | |
| 83 | default: string option, | |
| 84 | domintros: bool, | |
| 85 | partials: bool} | |
| 86 | val default_config : function_config | |
| 87 | val split_def : Proof.context -> (string -> 'a) -> term -> | |
| 88 | string * (string * typ) list * term list * term list * term | |
| 89 | val check_defs : Proof.context -> ((string * typ) * 'a) list -> term list -> unit | |
| 90 | type fixes = ((string * typ) * mixfix) list | |
| 91 | type 'a spec = (Attrib.binding * 'a list) list | |
| 92 | type preproc = function_config -> Proof.context -> fixes -> term spec -> | |
| 93 | (term list * (thm list -> thm spec) * (thm list -> thm list list) * string list) | |
| 94 | val fname_of : term -> string | |
| 95 | val mk_case_names : int -> string -> int -> string list | |
| 96 | val empty_preproc : (Proof.context -> ((string * typ) * mixfix) list -> term list -> 'c) -> preproc | |
| 97 | val get_preproc: Proof.context -> preproc | |
| 98 | val set_preproc: preproc -> Context.generic -> Context.generic | |
| 99 | val function_parser : function_config -> | |
| 100 | ((function_config * (binding * string option * mixfix) list) * (Attrib.binding * string) list) parser | |
| 101 | end | |
| 102 | ||
| 103 | structure Function_Common : FUNCTION_COMMON = | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 104 | struct | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 105 | |
| 34230 | 106 | open Function_Data | 
| 107 | ||
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 108 | local open Function_Lib in | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 109 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 110 | (* Profiling *) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 111 | val profile = Unsynchronized.ref false; | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 112 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 113 | fun PROFILE msg = if !profile then timeap_msg msg else I | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 114 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 115 | val acc_const_name = @{const_name accp}
 | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 116 | fun mk_acc domT R = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 117 | Const (acc_const_name, (domT --> domT --> HOLogic.boolT) --> domT --> HOLogic.boolT) $ R | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 118 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 119 | val function_name = suffix "C" | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 120 | val graph_name = suffix "_graph" | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 121 | val rel_name = suffix "_rel" | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 122 | val dom_name = suffix "_dom" | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 123 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 124 | (* Termination rules *) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 125 | |
| 46042 | 126 | (* FIXME just one data slot (record) per program unit *) | 
| 33519 | 127 | structure TerminationRule = Generic_Data | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 128 | ( | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 129 | type T = thm list | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 130 | val empty = [] | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 131 | val extend = I | 
| 33519 | 132 | val merge = Thm.merge_thms | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 133 | ); | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 134 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 135 | val get_termination_rules = TerminationRule.get | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 136 | val store_termination_rule = TerminationRule.map o cons | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 137 | val apply_termination_rule = resolve_tac o get_termination_rules o Context.Proof | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 138 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 139 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 140 | (* Function definition result data *) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 141 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 142 | datatype function_result = FunctionResult of | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 143 |  {fs: term list,
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 144 | G: term, | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 145 | R: term, | 
| 52384 | 146 | dom: term, | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 147 | psimps : thm list, | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 148 | simple_pinducts : thm list, | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 149 | cases : thm, | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 150 | termination : thm, | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 151 | domintros : thm list option} | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 152 | |
| 52384 | 153 | fun transform_function_data ({add_simps, case_names, fs, R, dom, psimps, pinducts,
 | 
| 52383 | 154 | simps, inducts, termination, defname, is_partial, cases} : info) phi = | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 155 | let | 
| 45290 | 156 | val term = Morphism.term phi | 
| 157 | val thm = Morphism.thm phi | |
| 158 | val fact = Morphism.fact phi | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 159 | val name = Binding.name_of o Morphism.binding phi o Binding.name | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 160 | in | 
| 34230 | 161 |       { add_simps = add_simps, case_names = case_names,
 | 
| 52384 | 162 | fs = map term fs, R = term R, dom = term dom, psimps = fact psimps, | 
| 34231 | 163 | pinducts = fact pinducts, simps = Option.map fact simps, | 
| 164 | inducts = Option.map fact inducts, termination = thm termination, | |
| 52383 | 165 | defname = name defname, is_partial=is_partial, cases = thm cases } | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 166 | end | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 167 | |
| 46042 | 168 | (* FIXME just one data slot (record) per program unit *) | 
| 33519 | 169 | structure FunctionData = Generic_Data | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 170 | ( | 
| 34230 | 171 | type T = (term * info) Item_Net.T; | 
| 33373 | 172 | val empty : T = Item_Net.init (op aconv o pairself fst) (single o fst); | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 173 | val extend = I; | 
| 33519 | 174 | fun merge tabs : T = Item_Net.merge tabs; | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 175 | ) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 176 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 177 | val get_function = FunctionData.get o Context.Proof; | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 178 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 179 | fun lift_morphism thy f = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 180 | let | 
| 46496 
b8920f3fd259
discontinued Drule.term_rule, which tends to cause confusion due to builtin cterm_of (NB: the standard morphisms already provide a separate term component);
 wenzelm parents: 
46042diff
changeset | 181 | fun term t = Thm.term_of (Drule.cterm_rule f (Thm.cterm_of thy t)) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 182 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 183 | Morphism.thm_morphism f $> Morphism.term_morphism term | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 184 | $> Morphism.typ_morphism (Logic.type_map term) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 185 | end | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 186 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 187 | fun import_function_data t ctxt = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 188 | let | 
| 42361 | 189 | val thy = Proof_Context.theory_of ctxt | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 190 | val ct = cterm_of thy t | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 191 | val inst_morph = lift_morphism thy o Thm.instantiate | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 192 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 193 | fun match (trm, data) = | 
| 45290 | 194 | SOME (transform_function_data data (inst_morph (Thm.match (cterm_of thy trm, ct)))) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 195 | handle Pattern.MATCH => NONE | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 196 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 197 | get_first match (Item_Net.retrieve (get_function ctxt) t) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 198 | end | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 199 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 200 | fun import_last_function ctxt = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 201 | case Item_Net.content (get_function ctxt) of | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 202 | [] => NONE | 
| 49966 | 203 | | (t, _) :: _ => | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 204 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 205 | val ([t'], ctxt') = Variable.import_terms true [t] ctxt | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 206 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 207 | import_function_data t' ctxt' | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 208 | end | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 209 | |
| 34230 | 210 | fun add_function_data (data : info as {fs, termination, ...}) =
 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 211 | FunctionData.map (fold (fn f => Item_Net.update (f, data)) fs) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 212 | #> store_termination_rule termination | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 213 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 214 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 215 | (* Simp rules for termination proofs *) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 216 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 217 | structure Termination_Simps = Named_Thms | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 218 | ( | 
| 45294 | 219 |   val name = @{binding termination_simp}
 | 
| 38764 
e6a18808873c
more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
 wenzelm parents: 
38761diff
changeset | 220 | val description = "simplification rules for termination proofs" | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 221 | ) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 222 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 223 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 224 | (* Default Termination Prover *) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 225 | |
| 46042 | 226 | (* FIXME just one data slot (record) per program unit *) | 
| 33519 | 227 | structure TerminationProver = Generic_Data | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 228 | ( | 
| 36521 | 229 | type T = Proof.context -> tactic | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 230 | val empty = (fn _ => error "Termination prover not configured") | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 231 | val extend = I | 
| 38761 | 232 | fun merge (a, _) = a | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 233 | ) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 234 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 235 | val set_termination_prover = TerminationProver.put | 
| 49965 | 236 | fun get_termination_prover ctxt = TerminationProver.get (Context.Proof ctxt) ctxt | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 237 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 238 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 239 | (* Configuration management *) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 240 | datatype function_opt | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 241 | = Sequential | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 242 | | Default of string | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 243 | | DomIntros | 
| 33101 | 244 | | No_Partials | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 245 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 246 | datatype function_config = FunctionConfig of | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 247 |  {sequential: bool,
 | 
| 41417 | 248 | default: string option, | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 249 | domintros: bool, | 
| 41846 
b368a7aee46a
removed support for tail-recursion from function package (now implemented by partial_function)
 krauss parents: 
41417diff
changeset | 250 | partials: bool} | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 251 | |
| 41846 
b368a7aee46a
removed support for tail-recursion from function package (now implemented by partial_function)
 krauss parents: 
41417diff
changeset | 252 | fun apply_opt Sequential (FunctionConfig {sequential, default, domintros, partials}) =
 | 
| 
b368a7aee46a
removed support for tail-recursion from function package (now implemented by partial_function)
 krauss parents: 
41417diff
changeset | 253 |     FunctionConfig {sequential=true, default=default, domintros=domintros, partials=partials}
 | 
| 
b368a7aee46a
removed support for tail-recursion from function package (now implemented by partial_function)
 krauss parents: 
41417diff
changeset | 254 |   | apply_opt (Default d) (FunctionConfig {sequential, default, domintros, partials}) =
 | 
| 
b368a7aee46a
removed support for tail-recursion from function package (now implemented by partial_function)
 krauss parents: 
41417diff
changeset | 255 |     FunctionConfig {sequential=sequential, default=SOME d, domintros=domintros, partials=partials}
 | 
| 
b368a7aee46a
removed support for tail-recursion from function package (now implemented by partial_function)
 krauss parents: 
41417diff
changeset | 256 |   | apply_opt DomIntros (FunctionConfig {sequential, default, domintros, partials}) =
 | 
| 
b368a7aee46a
removed support for tail-recursion from function package (now implemented by partial_function)
 krauss parents: 
41417diff
changeset | 257 |     FunctionConfig {sequential=sequential, default=default, domintros=true, partials=partials}
 | 
| 
b368a7aee46a
removed support for tail-recursion from function package (now implemented by partial_function)
 krauss parents: 
41417diff
changeset | 258 |   | apply_opt No_Partials (FunctionConfig {sequential, default, domintros, partials}) =
 | 
| 
b368a7aee46a
removed support for tail-recursion from function package (now implemented by partial_function)
 krauss parents: 
41417diff
changeset | 259 |     FunctionConfig {sequential=sequential, default=default, domintros=domintros, partials=false}
 | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 260 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 261 | val default_config = | 
| 41417 | 262 |   FunctionConfig { sequential=false, default=NONE,
 | 
| 41846 
b368a7aee46a
removed support for tail-recursion from function package (now implemented by partial_function)
 krauss parents: 
41417diff
changeset | 263 | domintros=false, partials=true} | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 264 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 265 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 266 | (* Analyzing function equations *) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 267 | |
| 39276 
2ad95934521f
improved error message for common mistake (fun f where "g x = ...")
 krauss parents: 
39134diff
changeset | 268 | fun split_def ctxt check_head geq = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 269 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 270 | fun input_error msg = cat_lines [msg, Syntax.string_of_term ctxt geq] | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 271 | val qs = Term.strip_qnt_vars "all" geq | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 272 | val imp = Term.strip_qnt_body "all" geq | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 273 | val (gs, eq) = Logic.strip_horn imp | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 274 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 275 | val (f_args, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop eq) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 276 | handle TERM _ => error (input_error "Not an equation") | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 277 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 278 | val (head, args) = strip_comb f_args | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 279 | |
| 39276 
2ad95934521f
improved error message for common mistake (fun f where "g x = ...")
 krauss parents: 
39134diff
changeset | 280 | val fname = fst (dest_Free head) handle TERM _ => "" | 
| 
2ad95934521f
improved error message for common mistake (fun f where "g x = ...")
 krauss parents: 
39134diff
changeset | 281 | val _ = check_head fname | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 282 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 283 | (fname, qs, gs, args, rhs) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 284 | end | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 285 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 286 | (* Check for all sorts of errors in the input *) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 287 | fun check_defs ctxt fixes eqs = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 288 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 289 | val fnames = map (fst o fst) fixes | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 290 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 291 | fun check geq = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 292 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 293 | fun input_error msg = error (cat_lines [msg, Syntax.string_of_term ctxt geq]) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 294 | |
| 39276 
2ad95934521f
improved error message for common mistake (fun f where "g x = ...")
 krauss parents: 
39134diff
changeset | 295 | fun check_head fname = | 
| 
2ad95934521f
improved error message for common mistake (fun f where "g x = ...")
 krauss parents: 
39134diff
changeset | 296 | member (op =) fnames fname orelse | 
| 
2ad95934521f
improved error message for common mistake (fun f where "g x = ...")
 krauss parents: 
39134diff
changeset | 297 |           input_error ("Illegal equation head. Expected " ^ commas_quote fnames)
 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 298 | |
| 39276 
2ad95934521f
improved error message for common mistake (fun f where "g x = ...")
 krauss parents: 
39134diff
changeset | 299 | val (fname, qs, gs, args, rhs) = split_def ctxt check_head geq | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 300 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 301 | val _ = length args > 0 orelse input_error "Function has no arguments:" | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 302 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 303 | fun add_bvs t is = add_loose_bnos (t, 0, is) | 
| 42081 | 304 | val rvs = (subtract (op =) (fold add_bvs args []) (add_bvs rhs [])) | 
| 305 | |> map (fst o nth (rev qs)) | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 306 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 307 | val _ = null rvs orelse input_error | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 308 |           ("Variable" ^ plural " " "s " rvs ^ commas_quote rvs ^
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 309 | " occur" ^ plural "s" "" rvs ^ " on right hand side only:") | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 310 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 311 | val _ = forall (not o Term.exists_subterm | 
| 36692 
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
 haftmann parents: 
36521diff
changeset | 312 | (fn Free (n, _) => member (op =) fnames n | _ => false)) (gs @ args) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 313 | orelse input_error "Defined function may not occur in premises or arguments" | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 314 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 315 | val freeargs = map (fn t => subst_bounds (rev (map Free qs), t)) args | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 316 | val funvars = filter (fn q => exists (exists_subterm (fn (Free q') $ _ => q = q' | _ => false)) freeargs) qs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 317 | val _ = null funvars orelse (warning (cat_lines | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 318 | ["Bound variable" ^ plural " " "s " funvars ^ | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 319 | commas_quote (map fst funvars) ^ " occur" ^ plural "s" "" funvars ^ | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 320 | " in function position.", "Misspelled constructor???"]); true) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 321 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 322 | (fname, length args) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 323 | end | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 324 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 325 | val grouped_args = AList.group (op =) (map check eqs) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 326 | val _ = grouped_args | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 327 | |> map (fn (fname, ars) => | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 328 | length (distinct (op =) ars) = 1 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 329 |         orelse error ("Function " ^ quote fname ^
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 330 | " has different numbers of arguments in different equations")) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 331 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 332 | val not_defined = subtract (op =) (map fst grouped_args) fnames | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 333 | val _ = null not_defined | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 334 |       orelse error ("No defining equations for function" ^
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 335 | plural " " "s " not_defined ^ commas_quote not_defined) | 
| 33751 
7ead0ccf6cbd
check if equations are present for all functions to avoid low-level exception later
 krauss parents: 
33519diff
changeset | 336 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 337 | fun check_sorts ((fname, fT), _) = | 
| 42361 | 338 | Sorts.of_sort (Sign.classes_of (Proof_Context.theory_of ctxt)) (fT, HOLogic.typeS) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 339 | orelse error (cat_lines | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 340 | ["Type of " ^ quote fname ^ " is not of sort " ^ quote "type" ^ ":", | 
| 39134 
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
 wenzelm parents: 
38764diff
changeset | 341 | Syntax.string_of_typ (Config.put show_sorts true ctxt) fT]) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 342 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 343 | val _ = map check_sorts fixes | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 344 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 345 | () | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 346 | end | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 347 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 348 | (* Preprocessors *) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 349 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 350 | type fixes = ((string * typ) * mixfix) list | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 351 | type 'a spec = (Attrib.binding * 'a list) list | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 352 | type preproc = function_config -> Proof.context -> fixes -> term spec -> | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 353 | (term list * (thm list -> thm spec) * (thm list -> thm list list) * string list) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 354 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 355 | val fname_of = fst o dest_Free o fst o strip_comb o fst o HOLogic.dest_eq o | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 356 | HOLogic.dest_Trueprop o Logic.strip_imp_concl o snd o dest_all_all | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 357 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 358 | fun mk_case_names i "" k = mk_case_names i (string_of_int (i + 1)) k | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 359 | | mk_case_names _ n 0 = [] | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 360 | | mk_case_names _ n 1 = [n] | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 361 | | mk_case_names _ n k = map (fn i => n ^ "_" ^ string_of_int i) (1 upto k) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 362 | |
| 49965 | 363 | fun empty_preproc check (_: function_config) (ctxt: Proof.context) (fixes: fixes) spec = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 364 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 365 | val (bnds, tss) = split_list spec | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 366 | val ts = flat tss | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 367 | val _ = check ctxt fixes ts | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 368 | val fnames = map (fst o fst) fixes | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 369 | val indices = map (fn eq => find_index (curry op = (fname_of eq)) fnames) ts | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 370 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 371 | fun sort xs = partition_list (fn i => fn (j,_) => i = j) 0 (length fnames - 1) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 372 | (indices ~~ xs) |> map (map snd) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 373 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 374 | (* using theorem names for case name currently disabled *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 375 | val cnames = map_index (fn (i, _) => mk_case_names i "" 1) bnds |> flat | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 376 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 377 | (ts, curry op ~~ bnds o Library.unflat tss, sort, cnames) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 378 | end | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 379 | |
| 46042 | 380 | (* FIXME just one data slot (record) per program unit *) | 
| 33519 | 381 | structure Preprocessor = Generic_Data | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 382 | ( | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 383 | type T = preproc | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 384 | val empty : T = empty_preproc check_defs | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 385 | val extend = I | 
| 33519 | 386 | fun merge (a, _) = a | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 387 | ) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 388 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 389 | val get_preproc = Preprocessor.get o Context.Proof | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 390 | val set_preproc = Preprocessor.map o K | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 391 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 392 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 393 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 394 | local | 
| 44357 | 395 | val option_parser = Parse.group (fn () => "option") | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 396 | ((Parse.reserved "sequential" >> K Sequential) | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 397 | || ((Parse.reserved "default" |-- Parse.term) >> Default) | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 398 | || (Parse.reserved "domintros" >> K DomIntros) | 
| 41846 
b368a7aee46a
removed support for tail-recursion from function package (now implemented by partial_function)
 krauss parents: 
41417diff
changeset | 399 | || (Parse.reserved "no_partials" >> K No_Partials)) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 400 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 401 | fun config_parser default = | 
| 46949 | 402 |     (Scan.optional (@{keyword "("} |-- Parse.!!! (Parse.list1 option_parser) --| @{keyword ")"}) [])
 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 403 | >> (fn opts => fold apply_opt opts default) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 404 | in | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 405 | fun function_parser default_cfg = | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 406 | config_parser default_cfg -- Parse.fixes -- Parse_Spec.where_alt_specs | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 407 | end | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 408 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 409 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 410 | end | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 411 | end |