| author | haftmann | 
| Tue, 31 Aug 2010 13:55:54 +0200 | |
| changeset 38924 | fcd1d0457e27 | 
| parent 38923 | 79d7f2b4cf71 | 
| child 38928 | 0e6f54c9d201 | 
| permissions | -rw-r--r-- | 
| 37744 | 1 | (* Title: Tools/Code/code_eval.ML | 
| 28054 | 2 | Author: Florian Haftmann, TU Muenchen | 
| 3 | ||
| 34028 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 haftmann parents: 
33992diff
changeset | 4 | Runtime services building on code generation into implementation language SML. | 
| 28054 | 5 | *) | 
| 6 | ||
| 34028 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 haftmann parents: 
33992diff
changeset | 7 | signature CODE_EVAL = | 
| 28054 | 8 | sig | 
| 34028 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 haftmann parents: 
33992diff
changeset | 9 | val target: string | 
| 32740 | 10 | val eval: string option -> string * (unit -> 'a) option Unsynchronized.ref | 
| 30970 
3fe2e418a071
generic postprocessing scheme for term evaluations
 haftmann parents: 
30962diff
changeset | 11 | -> ((term -> term) -> 'a -> 'a) -> theory -> term -> string list -> 'a | 
| 36470 | 12 | val evaluation_code: theory -> string -> string list -> string list | 
| 34028 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 haftmann parents: 
33992diff
changeset | 13 | -> string * ((string * string) list * (string * string) list) | 
| 28054 | 14 | val setup: theory -> theory | 
| 15 | end; | |
| 16 | ||
| 34028 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 haftmann parents: 
33992diff
changeset | 17 | structure Code_Eval : CODE_EVAL = | 
| 28054 | 18 | struct | 
| 19 | ||
| 33992 | 20 | (** generic **) | 
| 21 | ||
| 34028 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 haftmann parents: 
33992diff
changeset | 22 | val target = "Eval"; | 
| 28054 | 23 | |
| 34032 | 24 | val eval_struct_name = "Code"; | 
| 33992 | 25 | |
| 36470 | 26 | fun evaluation_code thy struct_name_hint tycos consts = | 
| 33992 | 27 | let | 
| 36271 
2ef9dbddfcb8
optionally ignore errors during translation of equations
 haftmann parents: 
35360diff
changeset | 28 | val (consts', (naming, program)) = Code_Thingol.consts_program thy false consts; | 
| 34028 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 haftmann parents: 
33992diff
changeset | 29 | val tycos' = map (the o Code_Thingol.lookup_tyco naming) tycos; | 
| 36470 | 30 | val struct_name = if struct_name_hint = "" then eval_struct_name | 
| 31 | else struct_name_hint; | |
| 38921 | 32 | val (ml_code, target_names) = Code_Target.produce_code_for thy | 
| 33 | target NONE (SOME struct_name) [] naming program (consts' @ tycos', []); | |
| 34028 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 haftmann parents: 
33992diff
changeset | 34 | val (consts'', tycos'') = chop (length consts') target_names; | 
| 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 haftmann parents: 
33992diff
changeset | 35 | val consts_map = map2 (fn const => fn NONE => | 
| 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 haftmann parents: 
33992diff
changeset | 36 |         error ("Constant " ^ (quote o Code.string_of_const thy) const
 | 
| 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 haftmann parents: 
33992diff
changeset | 37 | ^ "\nhas a user-defined serialization") | 
| 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 haftmann parents: 
33992diff
changeset | 38 | | SOME const'' => (const, const'')) consts consts'' | 
| 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 haftmann parents: 
33992diff
changeset | 39 | val tycos_map = map2 (fn tyco => fn NONE => | 
| 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 haftmann parents: 
33992diff
changeset | 40 |         error ("Type " ^ (quote o Sign.extern_type thy) tyco
 | 
| 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 haftmann parents: 
33992diff
changeset | 41 | ^ "\nhas a user-defined serialization") | 
| 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 haftmann parents: 
33992diff
changeset | 42 | | SOME tyco'' => (tyco, tyco'')) tycos tycos''; | 
| 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 haftmann parents: 
33992diff
changeset | 43 | in (ml_code, (tycos_map, consts_map)) end; | 
| 28054 | 44 | |
| 45 | ||
| 34028 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 haftmann parents: 
33992diff
changeset | 46 | (** evaluation **) | 
| 28054 | 47 | |
| 30970 
3fe2e418a071
generic postprocessing scheme for term evaluations
 haftmann parents: 
30962diff
changeset | 48 | fun eval some_target reff postproc thy t args = | 
| 28054 | 49 | let | 
| 36610 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 wenzelm parents: 
36534diff
changeset | 50 | val ctxt = ProofContext.init_global thy; | 
| 31063 
88aaab83b6fc
dropped explicit suppport for frees in evaluation conversion stack
 haftmann parents: 
31054diff
changeset | 51 | fun evaluator naming program ((_, (_, ty)), t) deps = | 
| 28054 | 52 | let | 
| 53 | val _ = if Code_Thingol.contains_dictvar t then | |
| 28724 | 54 | error "Term to be evaluated contains free dictionaries" else (); | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28350diff
changeset | 55 | val value_name = "Value.VALUE.value" | 
| 28054 | 56 | val program' = program | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28350diff
changeset | 57 | |> Graph.new_node (value_name, | 
| 37437 | 58 | Code_Thingol.Fun (Term.dummy_patternN, ((([], ty), [(([], t), (NONE, true))]), NONE))) | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28350diff
changeset | 59 | |> fold (curry Graph.add_edge value_name) deps; | 
| 38921 | 60 | val (value_code, [SOME value_name']) = Code_Target.produce_code_for thy | 
| 61 | (the_default target some_target) NONE (SOME "") [] naming program ([value_name], []); | |
| 28054 | 62 | val sml_code = "let\n" ^ value_code ^ "\nin " ^ value_name' | 
| 63 |           ^ space_implode " " (map (enclose "(" ")") args) ^ " end";
 | |
| 30672 
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
 wenzelm parents: 
30648diff
changeset | 64 | in ML_Context.evaluate ctxt false reff sml_code end; | 
| 38669 
9ff76d0f0610
refined and unified naming convention for dynamic code evaluation techniques
 haftmann parents: 
37950diff
changeset | 65 | in Code_Thingol.dynamic_eval_value thy postproc evaluator t end; | 
| 28054 | 66 | |
| 67 | ||
| 34028 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 haftmann parents: 
33992diff
changeset | 68 | (** instrumentalization by antiquotation **) | 
| 28054 | 69 | |
| 70 | local | |
| 71 | ||
| 33519 | 72 | structure CodeAntiqData = Proof_Data | 
| 28054 | 73 | ( | 
| 30962 | 74 | type T = (string list * string list) * (bool * (string | 
| 75 | * (string * ((string * string) list * (string * string) list)) lazy)); | |
| 76 |   fun init _ = (([], []), (true, ("", Lazy.value ("", ([], [])))));
 | |
| 28054 | 77 | ); | 
| 78 | ||
| 79 | val is_first_occ = fst o snd o CodeAntiqData.get; | |
| 80 | ||
| 30962 | 81 | fun register_code new_tycos new_consts ctxt = | 
| 28054 | 82 | let | 
| 30962 | 83 | val ((tycos, consts), (_, (struct_name, _))) = CodeAntiqData.get ctxt; | 
| 84 | val tycos' = fold (insert (op =)) new_tycos tycos; | |
| 85 | val consts' = fold (insert (op =)) new_consts consts; | |
| 28054 | 86 | val (struct_name', ctxt') = if struct_name = "" | 
| 33992 | 87 | then ML_Antiquote.variant eval_struct_name ctxt | 
| 28054 | 88 | else (struct_name, ctxt); | 
| 36470 | 89 | val acc_code = Lazy.lazy | 
| 90 | (fn () => evaluation_code (ProofContext.theory_of ctxt) eval_struct_name tycos' consts'); | |
| 30962 | 91 | in CodeAntiqData.put ((tycos', consts'), (false, (struct_name', acc_code))) ctxt' end; | 
| 92 | ||
| 93 | fun register_const const = register_code [] [const]; | |
| 28054 | 94 | |
| 30962 | 95 | fun register_datatype tyco constrs = register_code [tyco] constrs; | 
| 96 | ||
| 97 | fun print_const const all_struct_name tycos_map consts_map = | |
| 98 | (Long_Name.append all_struct_name o the o AList.lookup (op =) consts_map) const; | |
| 99 | ||
| 35019 
1ec0a3ff229e
simplified interface for ML antiquotations, struct_name is always "Isabelle";
 wenzelm parents: 
34032diff
changeset | 100 | fun print_code is_first print_it ctxt = | 
| 30962 | 101 | let | 
| 102 | val (_, (_, (struct_code_name, acc_code))) = CodeAntiqData.get ctxt; | |
| 33992 | 103 | val (ml_code, (tycos_map, consts_map)) = Lazy.force acc_code; | 
| 104 | val ml_code = if is_first then ml_code | |
| 28054 | 105 | else ""; | 
| 35019 
1ec0a3ff229e
simplified interface for ML antiquotations, struct_name is always "Isabelle";
 wenzelm parents: 
34032diff
changeset | 106 | val all_struct_name = "Isabelle." ^ struct_code_name; | 
| 30962 | 107 | in (ml_code, print_it all_struct_name tycos_map consts_map) end; | 
| 28054 | 108 | |
| 109 | in | |
| 110 | ||
| 35019 
1ec0a3ff229e
simplified interface for ML antiquotations, struct_name is always "Isabelle";
 wenzelm parents: 
34032diff
changeset | 111 | fun ml_code_antiq raw_const background = | 
| 28054 | 112 | let | 
| 31156 | 113 | val const = Code.check_const (ProofContext.theory_of background) raw_const; | 
| 28054 | 114 | val is_first = is_first_occ background; | 
| 115 | val background' = register_const const background; | |
| 35019 
1ec0a3ff229e
simplified interface for ML antiquotations, struct_name is always "Isabelle";
 wenzelm parents: 
34032diff
changeset | 116 | in (print_code is_first (print_const const), background') end; | 
| 30962 | 117 | |
| 28054 | 118 | end; (*local*) | 
| 119 | ||
| 120 | ||
| 36470 | 121 | (** reflection support **) | 
| 122 | ||
| 123 | fun check_datatype thy tyco consts = | |
| 124 | let | |
| 37448 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37437diff
changeset | 125 | val constrs = (map (fst o fst) o snd o Code.get_type thy) tyco; | 
| 36470 | 126 | val missing_constrs = subtract (op =) consts constrs; | 
| 127 | val _ = if null missing_constrs then [] | |
| 128 |       else error ("Missing constructor(s) " ^ commas (map quote missing_constrs)
 | |
| 129 | ^ " for datatype " ^ quote tyco); | |
| 130 | val false_constrs = subtract (op =) constrs consts; | |
| 131 | val _ = if null false_constrs then [] | |
| 132 |       else error ("Non-constructor(s) " ^ commas (map quote false_constrs)
 | |
| 133 | ^ " for datatype " ^ quote tyco); | |
| 134 | in () end; | |
| 135 | ||
| 136 | fun add_eval_tyco (tyco, tyco') thy = | |
| 137 | let | |
| 138 | val k = Sign.arity_number thy tyco; | |
| 139 | fun pr pr' fxy [] = tyco' | |
| 140 | | pr pr' fxy [ty] = | |
| 141 | Code_Printer.concat [pr' Code_Printer.BR ty, tyco'] | |
| 142 | | pr pr' fxy tys = | |
| 143 |           Code_Printer.concat [Code_Printer.enum "," "(" ")" (map (pr' Code_Printer.BR) tys), tyco']
 | |
| 144 | in | |
| 145 | thy | |
| 38923 | 146 | |> Code_Target.add_tyco_syntax target tyco (SOME (k, pr)) | 
| 36470 | 147 | end; | 
| 148 | ||
| 36514 | 149 | fun add_eval_constr (const, const') thy = | 
| 150 | let | |
| 151 | val k = Code.args_number thy const; | |
| 152 | fun pr pr' fxy ts = Code_Printer.brackify fxy | |
| 38922 | 153 | (const' :: the_list (Code_Printer.tuplify pr' Code_Printer.BR (map fst ts))); | 
| 36514 | 154 | in | 
| 155 | thy | |
| 38923 | 156 | |> Code_Target.add_const_syntax target const (SOME (Code_Printer.simple_const_syntax (k, pr))) | 
| 36514 | 157 | end; | 
| 158 | ||
| 38923 | 159 | fun add_eval_const (const, const') = Code_Target.add_const_syntax target | 
| 36470 | 160 | const (SOME (Code_Printer.simple_const_syntax (0, (K o K o K) const'))); | 
| 161 | ||
| 36514 | 162 | fun process (code_body, (tyco_map, (constr_map, const_map))) module_name NONE thy = | 
| 36470 | 163 | let | 
| 164 | val pr = Code_Printer.str o Long_Name.append module_name; | |
| 165 | in | |
| 166 | thy | |
| 167 | |> Code_Target.add_reserved target module_name | |
| 37198 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
36960diff
changeset | 168 | |> Context.theory_map | 
| 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
36960diff
changeset | 169 | (ML_Context.exec (fn () => ML_Context.eval_text true Position.none code_body)) | 
| 36470 | 170 | |> fold (add_eval_tyco o apsnd pr) tyco_map | 
| 36514 | 171 | |> fold (add_eval_constr o apsnd pr) constr_map | 
| 36470 | 172 | |> fold (add_eval_const o apsnd pr) const_map | 
| 173 | end | |
| 174 | | process (code_body, _) _ (SOME file_name) thy = | |
| 175 | let | |
| 37950 
bc285d91041e
moved basic thy file name operations from Thy_Load to Thy_Header;
 wenzelm parents: 
37744diff
changeset | 176 | val preamble = | 
| 
bc285d91041e
moved basic thy file name operations from Thy_Load to Thy_Header;
 wenzelm parents: 
37744diff
changeset | 177 | "(* Generated from " ^ Path.implode (Thy_Header.thy_path (Context.theory_name thy)) | 
| 36470 | 178 | ^ "; DO NOT EDIT! *)"; | 
| 179 | val _ = File.write (Path.explode file_name) (preamble ^ "\n\n" ^ code_body); | |
| 180 | in | |
| 181 | thy | |
| 182 | end; | |
| 183 | ||
| 184 | fun gen_code_reflect prep_type prep_const raw_datatypes raw_functions module_name some_file thy = | |
| 185 | let | |
| 186 | val datatypes = map (fn (raw_tyco, raw_cos) => | |
| 187 | (prep_type thy raw_tyco, map (prep_const thy) raw_cos)) raw_datatypes; | |
| 36514 | 188 | val _ = map (uncurry (check_datatype thy)) datatypes; | 
| 189 | val tycos = map fst datatypes; | |
| 190 | val constrs = maps snd datatypes; | |
| 36470 | 191 | val functions = map (prep_const thy) raw_functions; | 
| 36514 | 192 | val result = evaluation_code thy module_name tycos (constrs @ functions) | 
| 193 | |> (apsnd o apsnd) (chop (length constrs)); | |
| 36470 | 194 | in | 
| 195 | thy | |
| 36514 | 196 | |> process result module_name some_file | 
| 36470 | 197 | end; | 
| 198 | ||
| 199 | val code_reflect = gen_code_reflect Code_Target.cert_tyco Code.check_const; | |
| 200 | val code_reflect_cmd = gen_code_reflect Code_Target.read_tyco Code.read_const; | |
| 201 | ||
| 202 | ||
| 28054 | 203 | (** Isar setup **) | 
| 204 | ||
| 205 | val _ = ML_Context.add_antiq "code" (fn _ => Args.term >> ml_code_antiq); | |
| 206 | ||
| 36470 | 207 | local | 
| 208 | ||
| 209 | val datatypesK = "datatypes"; | |
| 210 | val functionsK = "functions"; | |
| 211 | val fileK = "file"; | |
| 212 | val andK = "and" | |
| 213 | ||
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 214 | val _ = List.app Keyword.keyword [datatypesK, functionsK]; | 
| 36470 | 215 | |
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 216 | val parse_datatype = | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 217 | Parse.name --| Parse.$$$ "=" -- (Parse.term ::: (Scan.repeat (Parse.$$$ "|" |-- Parse.term))); | 
| 36470 | 218 | |
| 219 | in | |
| 220 | ||
| 221 | val _ = | |
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 222 | Outer_Syntax.command "code_reflect" "enrich runtime environment with generated code" | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 223 | Keyword.thy_decl (Parse.name -- Scan.optional (Parse.$$$ datatypesK |-- (parse_datatype | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 224 | ::: Scan.repeat (Parse.$$$ andK |-- parse_datatype))) [] | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 225 | -- Scan.optional (Parse.$$$ functionsK |-- Scan.repeat1 Parse.name) [] | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 226 | -- Scan.option (Parse.$$$ fileK |-- Parse.name) | 
| 36534 | 227 | >> (fn (((module_name, raw_datatypes), raw_functions), some_file) => Toplevel.theory | 
| 36470 | 228 | (code_reflect_cmd raw_datatypes raw_functions module_name some_file))); | 
| 229 | ||
| 230 | end; (*local*) | |
| 231 | ||
| 34028 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 haftmann parents: 
33992diff
changeset | 232 | val setup = Code_Target.extend_target (target, (Code_ML.target_SML, K I)); | 
| 28054 | 233 | |
| 234 | end; (*struct*) |