src/HOL/Tools/value_command.ML
author wenzelm
Thu, 25 Jan 2018 14:13:55 +0100
changeset 67505 ceb324e34c14
parent 67463 a5ca98950a91
child 69592 a80d8ec6c998
permissions -rw-r--r--
clarified signature: items with \isasep are special;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63806
c54a53ef1873 clarified modules;
wenzelm
parents: 62969
diff changeset
     1
(*  Title:      HOL/Tools/value_command.ML
28227
77221ee0f7b9 generic value command
haftmann
parents:
diff changeset
     2
    Author:     Florian Haftmann, TU Muenchen
77221ee0f7b9 generic value command
haftmann
parents:
diff changeset
     3
58100
f54a8a4134d3 restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents: 56975
diff changeset
     4
Generic value command for arbitrary evaluators, with default using nbe or SML.
28227
77221ee0f7b9 generic value command
haftmann
parents:
diff changeset
     5
*)
77221ee0f7b9 generic value command
haftmann
parents:
diff changeset
     6
63806
c54a53ef1873 clarified modules;
wenzelm
parents: 62969
diff changeset
     7
signature VALUE_COMMAND =
28227
77221ee0f7b9 generic value command
haftmann
parents:
diff changeset
     8
sig
77221ee0f7b9 generic value command
haftmann
parents:
diff changeset
     9
  val value: Proof.context -> term -> term
58100
f54a8a4134d3 restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents: 56975
diff changeset
    10
  val value_select: string -> Proof.context -> term -> term
67330
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    11
  val value_cmd: xstring -> string list -> string -> Toplevel.state -> unit
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    12
  val add_evaluator: binding * (Proof.context -> term -> term) 
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    13
    -> theory -> string * theory
28227
77221ee0f7b9 generic value command
haftmann
parents:
diff changeset
    14
end;
77221ee0f7b9 generic value command
haftmann
parents:
diff changeset
    15
63806
c54a53ef1873 clarified modules;
wenzelm
parents: 62969
diff changeset
    16
structure Value_Command : VALUE_COMMAND =
28227
77221ee0f7b9 generic value command
haftmann
parents:
diff changeset
    17
struct
77221ee0f7b9 generic value command
haftmann
parents:
diff changeset
    18
67330
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    19
structure Evaluators = Theory_Data
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    20
(
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    21
  type T = (Proof.context -> term -> term) Name_Space.table;
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    22
  val empty = Name_Space.empty_table "evaluator";
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    23
  val extend = I;
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    24
  val merge = Name_Space.merge_tables;
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    25
)
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    26
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    27
fun add_evaluator (b, evaluator) thy =
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    28
  let
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    29
    val (name, tab') = Name_Space.define (Context.Theory thy) true
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    30
      (b, evaluator) (Evaluators.get thy);
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    31
    val thy' = Evaluators.put tab' thy;
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    32
  in (name, thy') end;
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    33
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    34
fun intern_evaluator ctxt raw_name =
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    35
  if raw_name = "" then ""
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    36
  else Name_Space.intern (Name_Space.space_of_table
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    37
    (Evaluators.get (Proof_Context.theory_of ctxt))) raw_name;
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    38
58100
f54a8a4134d3 restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents: 56975
diff changeset
    39
fun default_value ctxt t =
56927
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 56926
diff changeset
    40
  if null (Term.add_frees t [])
66345
882abe912da9 do not fall back on nbe if plain evaluation fails
haftmann
parents: 63806
diff changeset
    41
  then Code_Evaluation.dynamic_value_strict ctxt t
56927
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 56926
diff changeset
    42
  else Nbe.dynamic_value ctxt t;
28227
77221ee0f7b9 generic value command
haftmann
parents:
diff changeset
    43
58100
f54a8a4134d3 restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents: 56975
diff changeset
    44
fun value_select name ctxt =
67330
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    45
  if name = ""
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    46
  then default_value ctxt
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    47
  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: 56975
diff changeset
    48
67330
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    49
val value = value_select "";
58100
f54a8a4134d3 restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents: 56975
diff changeset
    50
67330
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    51
fun value_cmd raw_name modes raw_t state =
28227
77221ee0f7b9 generic value command
haftmann
parents:
diff changeset
    52
  let
77221ee0f7b9 generic value command
haftmann
parents:
diff changeset
    53
    val ctxt = Toplevel.context_of state;
67330
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    54
    val name = intern_evaluator ctxt raw_name;
28227
77221ee0f7b9 generic value command
haftmann
parents:
diff changeset
    55
    val t = Syntax.read_term ctxt raw_t;
67330
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    56
    val t' = value_select name ctxt t;
28227
77221ee0f7b9 generic value command
haftmann
parents:
diff changeset
    57
    val ty' = Term.type_of t';
31218
fa54c1e614df fixed typo
haftmann
parents: 29288
diff changeset
    58
    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: 36960
diff changeset
    59
    val p = Print_Mode.with_modes modes (fn () =>
28227
77221ee0f7b9 generic value command
haftmann
parents:
diff changeset
    60
      Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk,
77221ee0f7b9 generic value command
haftmann
parents:
diff changeset
    61
        Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) ();
77221ee0f7b9 generic value command
haftmann
parents:
diff changeset
    62
  in Pretty.writeln p end;
77221ee0f7b9 generic value command
haftmann
parents:
diff changeset
    63
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 33522
diff changeset
    64
val opt_modes =
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66345
diff changeset
    65
  Scan.optional (\<^keyword>\<open>(\<close> |-- Parse.!!! (Scan.repeat1 Parse.name --| \<^keyword>\<open>)\<close>)) [];
28227
77221ee0f7b9 generic value command
haftmann
parents:
diff changeset
    66
58100
f54a8a4134d3 restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents: 56975
diff changeset
    67
val opt_evaluator =
67330
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    68
  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: 56975
diff changeset
    69
  
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 33522
diff changeset
    70
val _ =
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66345
diff changeset
    71
  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: 56975
diff changeset
    72
    (opt_evaluator -- opt_modes -- Parse.term
67330
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    73
      >> (fn ((name, modes), t) => Toplevel.keep (value_cmd name modes t)));
28227
77221ee0f7b9 generic value command
haftmann
parents:
diff changeset
    74
59323
468bd3aedfa1 modernized and more uniform style
haftmann
parents: 58893
diff changeset
    75
val _ = Theory.setup
67463
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67386
diff changeset
    76
  (Thy_Output.antiquotation_pretty_source \<^binding>\<open>value\<close>
58100
f54a8a4134d3 restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents: 56975
diff changeset
    77
    (Scan.lift opt_evaluator -- Term_Style.parse -- Args.term)
67463
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67386
diff changeset
    78
    (fn ctxt => fn ((name, style), t) =>
67505
ceb324e34c14 clarified signature: items with \isasep are special;
wenzelm
parents: 67463
diff changeset
    79
      Thy_Output.pretty_term ctxt (style (value_select name ctxt t)))
67330
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    80
  #> add_evaluator (\<^binding>\<open>simp\<close>, Code_Simp.dynamic_value) #> snd
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    81
  #> add_evaluator (\<^binding>\<open>nbe\<close>, Nbe.dynamic_value) #> snd
2505cabfc515 proper namespace for evaluators
haftmann
parents: 67149
diff changeset
    82
  #> add_evaluator (\<^binding>\<open>code\<close>, Code_Evaluation.dynamic_value_strict) #> snd);
43612
c32144b8baba adding a value antiquotation
bulwahn
parents: 42361
diff changeset
    83
28227
77221ee0f7b9 generic value command
haftmann
parents:
diff changeset
    84
end;