| author | wenzelm | 
| Mon, 02 Jan 2017 18:08:04 +0100 | |
| changeset 64757 | 7e3924224769 | 
| parent 63806 | c54a53ef1873 | 
| child 66345 | 882abe912da9 | 
| permissions | -rw-r--r-- | 
| 63806 | 1 | (* Title: HOL/Tools/value_command.ML | 
| 28227 | 2 | Author: Florian Haftmann, TU Muenchen | 
| 3 | ||
| 58100 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 4 | Generic value command for arbitrary evaluators, with default using nbe or SML. | 
| 28227 | 5 | *) | 
| 6 | ||
| 63806 | 7 | signature VALUE_COMMAND = | 
| 28227 | 8 | sig | 
| 9 | val value: Proof.context -> term -> term | |
| 58100 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 10 | val value_select: string -> Proof.context -> term -> term | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 11 | val value_cmd: string option -> string list -> string -> Toplevel.state -> unit | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 12 | val add_evaluator: string * (Proof.context -> term -> term) -> theory -> theory | 
| 28227 | 13 | end; | 
| 14 | ||
| 63806 | 15 | structure Value_Command : VALUE_COMMAND = | 
| 28227 | 16 | struct | 
| 17 | ||
| 58100 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 18 | fun default_value ctxt t = | 
| 56927 | 19 | if null (Term.add_frees t []) | 
| 20 | then case try (Code_Evaluation.dynamic_value_strict ctxt) t of | |
| 21 | SOME t' => t' | |
| 22 | | NONE => Nbe.dynamic_value ctxt t | |
| 23 | else Nbe.dynamic_value ctxt t; | |
| 28227 | 24 | |
| 58100 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 25 | structure Evaluator = Theory_Data | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 26 | ( | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 27 | type T = (string * (Proof.context -> term -> term)) list; | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 28 |   val empty = [("default", default_value)];
 | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 29 | val extend = I; | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 30 | fun merge data : T = AList.merge (op =) (K true) data; | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 31 | ) | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 32 | |
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 33 | val add_evaluator = Evaluator.map o AList.update (op =); | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 34 | |
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 35 | fun value_select name ctxt = | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 36 | case AList.lookup (op =) (Evaluator.get (Proof_Context.theory_of ctxt)) name | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 37 |    of NONE => error ("No such evaluator: " ^ name)
 | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 38 | | SOME f => f ctxt; | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 39 | |
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 40 | fun value ctxt = | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 41 | let | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 42 | val evaluators = Evaluator.get (Proof_Context.theory_of ctxt) | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 43 | in | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 44 | if null evaluators | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 45 | then error "No evaluators" | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 46 | else (snd o snd o split_last) evaluators ctxt | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 47 | end; | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 48 | |
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 49 | fun value_maybe_select some_name = | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 50 | case some_name | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 51 | of NONE => value | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 52 | | SOME name => value_select name; | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 53 | |
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 54 | fun value_cmd some_name modes raw_t state = | 
| 28227 | 55 | let | 
| 56 | val ctxt = Toplevel.context_of state; | |
| 57 | val t = Syntax.read_term ctxt raw_t; | |
| 58100 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 58 | val t' = value_maybe_select some_name ctxt t; | 
| 28227 | 59 | val ty' = Term.type_of t'; | 
| 31218 | 60 | val ctxt' = Variable.auto_fixes t' ctxt; | 
| 37146 
f652333bbf8e
renamed structure PrintMode to Print_Mode, keeping the old name as legacy alias for some time;
 wenzelm parents: 
36960diff
changeset | 61 | val p = Print_Mode.with_modes modes (fn () => | 
| 28227 | 62 | Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk, | 
| 63 | Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) (); | |
| 64 | in Pretty.writeln p end; | |
| 65 | ||
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
33522diff
changeset | 66 | val opt_modes = | 
| 62969 | 67 |   Scan.optional (@{keyword "("} |-- Parse.!!! (Scan.repeat1 Parse.name --| @{keyword ")"})) [];
 | 
| 28227 | 68 | |
| 58100 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 69 | val opt_evaluator = | 
| 62969 | 70 |   Scan.option (@{keyword "["} |-- Parse.name --| @{keyword "]"})
 | 
| 58100 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 71 | |
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
33522diff
changeset | 72 | val _ = | 
| 59936 
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
 wenzelm parents: 
59323diff
changeset | 73 |   Outer_Syntax.command @{command_keyword value} "evaluate and print term"
 | 
| 58100 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 74 | (opt_evaluator -- opt_modes -- Parse.term | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 75 | >> (fn ((some_name, modes), t) => Toplevel.keep (value_cmd some_name modes t))); | 
| 28227 | 76 | |
| 59323 | 77 | val _ = Theory.setup | 
| 56926 | 78 |   (Thy_Output.antiquotation @{binding value}
 | 
| 58100 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 79 | (Scan.lift opt_evaluator -- Term_Style.parse -- Args.term) | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 80 |     (fn {source, context, ...} => fn ((some_name, style), t) => Thy_Output.output context
 | 
| 43612 | 81 | (Thy_Output.maybe_pretty_source Thy_Output.pretty_term context source | 
| 58100 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 82 | [style (value_maybe_select some_name context t)])) | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 83 |   #> add_evaluator ("simp", Code_Simp.dynamic_value)
 | 
| 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 84 |   #> add_evaluator ("nbe", Nbe.dynamic_value)
 | 
| 59323 | 85 |   #> add_evaluator ("code", Code_Evaluation.dynamic_value_strict));
 | 
| 43612 | 86 | |
| 28227 | 87 | end; |