| author | wenzelm | 
| Wed, 11 Sep 2024 23:26:25 +0200 | |
| changeset 80866 | 8c67b14fdd48 | 
| parent 74561 | 8e6c973003c8 | 
| 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 | 
| 67330 | 11 | val value_cmd: xstring -> string list -> string -> Toplevel.state -> unit | 
| 12 | val add_evaluator: binding * (Proof.context -> term -> term) | |
| 13 | -> theory -> string * theory | |
| 28227 | 14 | end; | 
| 15 | ||
| 63806 | 16 | structure Value_Command : VALUE_COMMAND = | 
| 28227 | 17 | struct | 
| 18 | ||
| 67330 | 19 | structure Evaluators = Theory_Data | 
| 20 | ( | |
| 21 | type T = (Proof.context -> term -> term) Name_Space.table; | |
| 22 | val empty = Name_Space.empty_table "evaluator"; | |
| 23 | val merge = Name_Space.merge_tables; | |
| 24 | ) | |
| 25 | ||
| 26 | fun add_evaluator (b, evaluator) thy = | |
| 27 | let | |
| 28 | val (name, tab') = Name_Space.define (Context.Theory thy) true | |
| 29 | (b, evaluator) (Evaluators.get thy); | |
| 30 | val thy' = Evaluators.put tab' thy; | |
| 31 | in (name, thy') end; | |
| 32 | ||
| 33 | fun intern_evaluator ctxt raw_name = | |
| 34 | if raw_name = "" then "" | |
| 35 | else Name_Space.intern (Name_Space.space_of_table | |
| 36 | (Evaluators.get (Proof_Context.theory_of ctxt))) raw_name; | |
| 37 | ||
| 58100 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 38 | fun default_value ctxt t = | 
| 56927 | 39 | if null (Term.add_frees t []) | 
| 66345 
882abe912da9
do not fall back on nbe if plain evaluation fails
 haftmann parents: 
63806diff
changeset | 40 | then Code_Evaluation.dynamic_value_strict ctxt t | 
| 56927 | 41 | else Nbe.dynamic_value ctxt t; | 
| 28227 | 42 | |
| 58100 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 43 | fun value_select name ctxt = | 
| 67330 | 44 | if name = "" | 
| 45 | then default_value ctxt | |
| 46 | else Name_Space.get (Evaluators.get (Proof_Context.theory_of ctxt)) name ctxt; | |
| 58100 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 47 | |
| 67330 | 48 | val value = value_select ""; | 
| 58100 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 49 | |
| 67330 | 50 | fun value_cmd raw_name modes raw_t state = | 
| 28227 | 51 | let | 
| 52 | val ctxt = Toplevel.context_of state; | |
| 67330 | 53 | val name = intern_evaluator ctxt raw_name; | 
| 28227 | 54 | val t = Syntax.read_term ctxt raw_t; | 
| 67330 | 55 | val t' = value_select name ctxt t; | 
| 28227 | 56 | val ty' = Term.type_of t'; | 
| 70308 | 57 | val ctxt' = Proof_Context.augment t' ctxt; | 
| 37146 
f652333bbf8e
renamed structure PrintMode to Print_Mode, keeping the old name as legacy alias for some time;
 wenzelm parents: 
36960diff
changeset | 58 | val p = Print_Mode.with_modes modes (fn () => | 
| 28227 | 59 | Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk, | 
| 60 | Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) (); | |
| 61 | in Pretty.writeln p end; | |
| 62 | ||
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
33522diff
changeset | 63 | val opt_modes = | 
| 67149 | 64 | Scan.optional (\<^keyword>\<open>(\<close> |-- Parse.!!! (Scan.repeat1 Parse.name --| \<^keyword>\<open>)\<close>)) []; | 
| 28227 | 65 | |
| 58100 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 66 | val opt_evaluator = | 
| 67330 | 67 | Scan.optional (\<^keyword>\<open>[\<close> |-- Parse.name --| \<^keyword>\<open>]\<close>) ""; | 
| 58100 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 68 | |
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
33522diff
changeset | 69 | val _ = | 
| 67149 | 70 | Outer_Syntax.command \<^command_keyword>\<open>value\<close> "evaluate and print term" | 
| 58100 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 71 | (opt_evaluator -- opt_modes -- Parse.term | 
| 67330 | 72 | >> (fn ((name, modes), t) => Toplevel.keep (value_cmd name modes t))); | 
| 28227 | 73 | |
| 59323 | 74 | val _ = Theory.setup | 
| 73761 | 75 | (Document_Output.antiquotation_pretty_source_embedded \<^binding>\<open>value\<close> | 
| 58100 
f54a8a4134d3
restored generic value slot, retaining default behaviour and separate approximate command
 haftmann parents: 
56975diff
changeset | 76 | (Scan.lift opt_evaluator -- Term_Style.parse -- Args.term) | 
| 67463 | 77 | (fn ctxt => fn ((name, style), t) => | 
| 73761 | 78 | Document_Output.pretty_term ctxt (style (value_select name ctxt t))) | 
| 67330 | 79 | #> add_evaluator (\<^binding>\<open>simp\<close>, Code_Simp.dynamic_value) #> snd | 
| 80 | #> add_evaluator (\<^binding>\<open>nbe\<close>, Nbe.dynamic_value) #> snd | |
| 81 | #> add_evaluator (\<^binding>\<open>code\<close>, Code_Evaluation.dynamic_value_strict) #> snd); | |
| 43612 | 82 | |
| 28227 | 83 | end; |