| author | haftmann | 
| Tue, 04 May 2010 08:55:39 +0200 | |
| changeset 36634 | f9b43d197d16 | 
| parent 36521 | 73ed9f18fdd3 | 
| child 36692 | 54b64d4ad524 | 
| 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 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 4 | A package for general recursive function definitions. | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 5 | Common definitions and other infrastructure. | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 6 | *) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 7 | |
| 34230 | 8 | signature FUNCTION_DATA = | 
| 9 | sig | |
| 10 | ||
| 11 | type info = | |
| 12 |  {is_partial : bool,
 | |
| 13 | defname : string, | |
| 14 | (* contains no logical entities: invariant under morphisms: *) | |
| 15 | add_simps : (binding -> binding) -> string -> (binding -> binding) -> | |
| 16 | Attrib.src list -> thm list -> local_theory -> thm list * local_theory, | |
| 17 | case_names : string list, | |
| 18 | fs : term list, | |
| 19 | R : term, | |
| 20 | psimps: thm list, | |
| 21 | pinducts: thm list, | |
| 34231 | 22 | simps : thm list option, | 
| 23 | inducts : thm list option, | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 24 | termination: thm} | 
| 34230 | 25 | |
| 26 | end | |
| 27 | ||
| 28 | structure Function_Data : FUNCTION_DATA = | |
| 29 | struct | |
| 30 | ||
| 31 | type info = | |
| 32 |  {is_partial : bool,
 | |
| 33 | defname : string, | |
| 34 | (* contains no logical entities: invariant under morphisms: *) | |
| 35 | add_simps : (binding -> binding) -> string -> (binding -> binding) -> | |
| 36 | Attrib.src list -> thm list -> local_theory -> thm list * local_theory, | |
| 37 | case_names : string list, | |
| 38 | fs : term list, | |
| 39 | R : term, | |
| 40 | psimps: thm list, | |
| 41 | pinducts: thm list, | |
| 34231 | 42 | simps : thm list option, | 
| 43 | inducts : thm list option, | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 44 | termination: thm} | 
| 34230 | 45 | |
| 46 | end | |
| 47 | ||
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 48 | structure Function_Common = | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 49 | struct | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 50 | |
| 34230 | 51 | open Function_Data | 
| 52 | ||
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 53 | local open Function_Lib in | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 54 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 55 | (* Profiling *) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 56 | val profile = Unsynchronized.ref false; | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 57 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 58 | 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 | 59 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 60 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 61 | val acc_const_name = @{const_name accp}
 | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 62 | fun mk_acc domT R = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 63 | 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 | 64 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 65 | val function_name = suffix "C" | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 66 | val graph_name = suffix "_graph" | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 67 | val rel_name = suffix "_rel" | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 68 | val dom_name = suffix "_dom" | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 69 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 70 | (* Termination rules *) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 71 | |
| 33519 | 72 | structure TerminationRule = Generic_Data | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 73 | ( | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 74 | type T = thm list | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 75 | val empty = [] | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 76 | val extend = I | 
| 33519 | 77 | val merge = Thm.merge_thms | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 78 | ); | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 79 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 80 | val get_termination_rules = TerminationRule.get | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 81 | val store_termination_rule = TerminationRule.map o cons | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 82 | 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 | 83 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 84 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 85 | (* Function definition result data *) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 86 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 87 | datatype function_result = FunctionResult of | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 88 |  {fs: term list,
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 89 | G: term, | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 90 | R: term, | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 91 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 92 | psimps : thm list, | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 93 | trsimps : thm list option, | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 94 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 95 | simple_pinducts : thm list, | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 96 | cases : thm, | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 97 | termination : thm, | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 98 | domintros : thm list option} | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 99 | |
| 34230 | 100 | fun morph_function_data ({add_simps, case_names, fs, R, psimps, pinducts,
 | 
| 34231 | 101 | simps, inducts, termination, defname, is_partial} : info) phi = | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 102 | let | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 103 | val term = Morphism.term phi val thm = Morphism.thm phi val fact = Morphism.fact phi | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 104 | 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 | 105 | in | 
| 34230 | 106 |       { add_simps = add_simps, case_names = case_names,
 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 107 | fs = map term fs, R = term R, psimps = fact psimps, | 
| 34231 | 108 | pinducts = fact pinducts, simps = Option.map fact simps, | 
| 109 | inducts = Option.map fact inducts, termination = thm termination, | |
| 34230 | 110 | defname = name defname, is_partial=is_partial } | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 111 | end | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 112 | |
| 33519 | 113 | structure FunctionData = Generic_Data | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 114 | ( | 
| 34230 | 115 | type T = (term * info) Item_Net.T; | 
| 33373 | 116 | 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 | 117 | val extend = I; | 
| 33519 | 118 | fun merge tabs : T = Item_Net.merge tabs; | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 119 | ) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 120 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 121 | val get_function = FunctionData.get o Context.Proof; | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 122 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 123 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 124 | fun lift_morphism thy f = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 125 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 126 | val term = Drule.term_rule thy f | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 127 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 128 | Morphism.thm_morphism f $> Morphism.term_morphism term | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 129 | $> Morphism.typ_morphism (Logic.type_map term) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 130 | end | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 131 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 132 | fun import_function_data t ctxt = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 133 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 134 | val thy = ProofContext.theory_of ctxt | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 135 | val ct = cterm_of thy t | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 136 | val inst_morph = lift_morphism thy o Thm.instantiate | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 137 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 138 | fun match (trm, data) = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 139 | SOME (morph_function_data data (inst_morph (Thm.match (cterm_of thy trm, ct)))) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 140 | handle Pattern.MATCH => NONE | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 141 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 142 | get_first match (Item_Net.retrieve (get_function ctxt) t) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 143 | end | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 144 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 145 | fun import_last_function ctxt = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 146 | case Item_Net.content (get_function ctxt) of | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 147 | [] => NONE | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 148 | | (t, data) :: _ => | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 149 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 150 | val ([t'], ctxt') = Variable.import_terms true [t] ctxt | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 151 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 152 | import_function_data t' ctxt' | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 153 | end | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 154 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 155 | val all_function_data = Item_Net.content o get_function | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 156 | |
| 34230 | 157 | fun add_function_data (data : info as {fs, termination, ...}) =
 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 158 | 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 | 159 | #> store_termination_rule termination | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 160 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 161 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 162 | (* Simp rules for termination proofs *) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 163 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 164 | structure Termination_Simps = Named_Thms | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 165 | ( | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 166 | val name = "termination_simp" | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 167 | val description = "Simplification rule for termination proofs" | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 168 | ) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 169 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 170 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 171 | (* Default Termination Prover *) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 172 | |
| 33519 | 173 | structure TerminationProver = Generic_Data | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 174 | ( | 
| 36521 | 175 | type T = Proof.context -> tactic | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 176 | val empty = (fn _ => error "Termination prover not configured") | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 177 | val extend = I | 
| 33519 | 178 | fun merge (a, b) = b (* FIXME ? *) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 179 | ) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 180 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 181 | val set_termination_prover = TerminationProver.put | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 182 | val get_termination_prover = TerminationProver.get o Context.Proof | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 183 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 184 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 185 | (* Configuration management *) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 186 | datatype function_opt | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 187 | = Sequential | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 188 | | Default of string | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 189 | | DomIntros | 
| 33101 | 190 | | No_Partials | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 191 | | Tailrec | 
| 
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 | datatype function_config = FunctionConfig of | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 194 |  {sequential: bool,
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 195 | default: string, | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 196 | domintros: bool, | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 197 | partials: bool, | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 198 | tailrec: bool} | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 199 | |
| 33101 | 200 | fun apply_opt Sequential (FunctionConfig {sequential, default, domintros, partials, tailrec}) =
 | 
| 201 |     FunctionConfig {sequential=true, default=default, domintros=domintros, partials=partials, tailrec=tailrec}
 | |
| 202 |   | apply_opt (Default d) (FunctionConfig {sequential, default, domintros, partials, tailrec}) =
 | |
| 203 |     FunctionConfig {sequential=sequential, default=d, domintros=domintros, partials=partials, tailrec=tailrec}
 | |
| 204 |   | apply_opt DomIntros (FunctionConfig {sequential, default, domintros, partials, tailrec}) =
 | |
| 205 |     FunctionConfig {sequential=sequential, default=default, domintros=true, partials=partials, tailrec=tailrec}
 | |
| 206 |   | apply_opt Tailrec (FunctionConfig {sequential, default, domintros, partials, tailrec}) =
 | |
| 207 |     FunctionConfig {sequential=sequential, default=default, domintros=domintros, partials=partials, tailrec=true}
 | |
| 208 |   | apply_opt No_Partials (FunctionConfig {sequential, default, domintros, partials, tailrec}) =
 | |
| 209 |     FunctionConfig {sequential=sequential, default=default, domintros=domintros, partials=false, tailrec=true}
 | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 210 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 211 | val default_config = | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 212 |   FunctionConfig { sequential=false, default="%x. undefined" (*FIXME dynamic scoping*), 
 | 
| 33101 | 213 | domintros=false, partials=true, tailrec=false } | 
| 33099 
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 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 216 | (* Analyzing function equations *) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 217 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 218 | fun split_def ctxt geq = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 219 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 220 | 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 | 221 | val qs = Term.strip_qnt_vars "all" geq | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 222 | val imp = Term.strip_qnt_body "all" geq | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 223 | val (gs, eq) = Logic.strip_horn imp | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 224 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 225 | 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 | 226 | handle TERM _ => error (input_error "Not an equation") | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 227 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 228 | val (head, args) = strip_comb f_args | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 229 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 230 | val fname = fst (dest_Free head) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 231 | handle TERM _ => error (input_error "Head symbol must not be a bound variable") | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 232 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 233 | (fname, qs, gs, args, rhs) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 234 | end | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 235 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 236 | (* Check for all sorts of errors in the input *) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 237 | fun check_defs ctxt fixes eqs = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 238 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 239 | val fnames = map (fst o fst) fixes | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 240 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 241 | fun check geq = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 242 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 243 | 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 | 244 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 245 | val fqgar as (fname, qs, gs, args, rhs) = split_def ctxt geq | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 246 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 247 | val _ = fname mem fnames | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 248 |           orelse input_error ("Head symbol of left hand side must be " ^
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 249 | plural "" "one out of " fnames ^ commas_quote fnames) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 250 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 251 | 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 | 252 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 253 | fun add_bvs t is = add_loose_bnos (t, 0, is) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 254 | val rvs = (subtract (op =) (fold add_bvs args []) (add_bvs rhs [])) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 255 | |> map (fst o nth (rev qs)) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 256 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 257 | val _ = null rvs orelse input_error | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 258 |           ("Variable" ^ plural " " "s " rvs ^ commas_quote rvs ^
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 259 | " occur" ^ plural "s" "" rvs ^ " on right hand side only:") | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 260 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 261 | val _ = forall (not o Term.exists_subterm | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 262 | (fn Free (n, _) => n mem fnames | _ => false)) (gs @ args) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 263 | 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 | 264 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 265 | 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 | 266 | 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 | 267 | val _ = null funvars orelse (warning (cat_lines | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 268 | ["Bound variable" ^ plural " " "s " funvars ^ | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 269 | commas_quote (map fst funvars) ^ " occur" ^ plural "s" "" funvars ^ | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 270 | " in function position.", "Misspelled constructor???"]); true) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 271 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 272 | (fname, length args) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 273 | end | 
| 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 grouped_args = AList.group (op =) (map check eqs) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 276 | val _ = grouped_args | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 277 | |> map (fn (fname, ars) => | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 278 | length (distinct (op =) ars) = 1 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 279 |         orelse error ("Function " ^ quote fname ^
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 280 | " has different numbers of arguments in different equations")) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 281 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 282 | val not_defined = subtract (op =) (map fst grouped_args) fnames | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 283 | val _ = null not_defined | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 284 |       orelse error ("No defining equations for function" ^
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 285 | 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 | 286 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 287 | fun check_sorts ((fname, fT), _) = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 288 | Sorts.of_sort (Sign.classes_of (ProofContext.theory_of ctxt)) (fT, HOLogic.typeS) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 289 | orelse error (cat_lines | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 290 | ["Type of " ^ quote fname ^ " is not of sort " ^ quote "type" ^ ":", | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 291 | setmp_CRITICAL show_sorts true (Syntax.string_of_typ ctxt) fT]) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 292 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 293 | val _ = map check_sorts fixes | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 294 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 295 | () | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 296 | end | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 297 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 298 | (* Preprocessors *) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 299 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 300 | type fixes = ((string * typ) * mixfix) list | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 301 | type 'a spec = (Attrib.binding * 'a list) list | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 302 | type preproc = function_config -> Proof.context -> fixes -> term spec -> | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 303 | (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 | 304 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 305 | 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 | 306 | 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 | 307 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 308 | 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 | 309 | | mk_case_names _ n 0 = [] | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 310 | | mk_case_names _ n 1 = [n] | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 311 | | 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 | 312 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 313 | fun empty_preproc check _ ctxt fixes spec = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 314 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 315 | val (bnds, tss) = split_list spec | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 316 | val ts = flat tss | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 317 | val _ = check ctxt fixes ts | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 318 | val fnames = map (fst o fst) fixes | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 319 | 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 | 320 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 321 | 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 | 322 | (indices ~~ xs) |> map (map snd) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 323 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 324 | (* using theorem names for case name currently disabled *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 325 | 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 | 326 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 327 | (ts, curry op ~~ bnds o Library.unflat tss, sort, cnames) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 328 | end | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 329 | |
| 33519 | 330 | structure Preprocessor = Generic_Data | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 331 | ( | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 332 | type T = preproc | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 333 | val empty : T = empty_preproc check_defs | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 334 | val extend = I | 
| 33519 | 335 | fun merge (a, _) = a | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 336 | ) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 337 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 338 | val get_preproc = Preprocessor.get o Context.Proof | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 339 | val set_preproc = Preprocessor.map o K | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 340 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 341 | |
| 
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 | local | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 344 | structure P = OuterParse and K = OuterKeyword | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 345 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 346 | val option_parser = P.group "option" | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 347 | ((P.reserved "sequential" >> K Sequential) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 348 | || ((P.reserved "default" |-- P.term) >> Default) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 349 | || (P.reserved "domintros" >> K DomIntros) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 350 | || (P.reserved "no_partials" >> K No_Partials) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 351 | || (P.reserved "tailrec" >> K Tailrec)) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 352 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 353 | fun config_parser default = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 354 |     (Scan.optional (P.$$$ "(" |-- P.!!! (P.list1 option_parser) --| P.$$$ ")") [])
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 355 | >> (fn opts => fold apply_opt opts default) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 356 | in | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34231diff
changeset | 357 | fun function_parser default_cfg = | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 358 | config_parser default_cfg -- P.fixes -- SpecParse.where_alt_specs | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 359 | end | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 360 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 361 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 362 | end | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 363 | end |