author | wenzelm |
Tue, 15 Jan 2013 17:28:46 +0100 | |
changeset 50902 | cb2b940e2fdf |
parent 46961 | 5c6955f487e5 |
child 51658 | 21c10672633b |
permissions | -rw-r--r-- |
37744 | 1 |
(* Title: Tools/value.ML |
28227 | 2 |
Author: Florian Haftmann, TU Muenchen |
3 |
||
28952
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents:
28227
diff
changeset
|
4 |
Generic value command for arbitrary evaluators. |
28227 | 5 |
*) |
6 |
||
7 |
signature VALUE = |
|
8 |
sig |
|
9 |
val value: Proof.context -> term -> term |
|
10 |
val value_select: string -> Proof.context -> term -> term |
|
11 |
val value_cmd: string option -> string list -> string -> Toplevel.state -> unit |
|
12 |
val add_evaluator: string * (Proof.context -> term -> term) -> theory -> theory |
|
43612 | 13 |
val setup : theory -> theory |
28227 | 14 |
end; |
15 |
||
16 |
structure Value : VALUE = |
|
17 |
struct |
|
18 |
||
33522 | 19 |
structure Evaluator = Theory_Data |
20 |
( |
|
28227 | 21 |
type T = (string * (Proof.context -> term -> term)) list; |
22 |
val empty = []; |
|
23 |
val extend = I; |
|
33522 | 24 |
fun merge data : T = AList.merge (op =) (K true) data; |
28227 | 25 |
) |
26 |
||
27 |
val add_evaluator = Evaluator.map o AList.update (op =); |
|
28 |
||
29 |
fun value_select name ctxt = |
|
42361 | 30 |
case AList.lookup (op =) (Evaluator.get (Proof_Context.theory_of ctxt)) name |
28227 | 31 |
of NONE => error ("No such evaluator: " ^ name) |
32 |
| SOME f => f ctxt; |
|
33 |
||
42361 | 34 |
fun value ctxt t = let val evaluators = Evaluator.get (Proof_Context.theory_of ctxt) |
28227 | 35 |
in if null evaluators then error "No evaluators" |
36 |
else let val (evaluators, (_, evaluator)) = split_last evaluators |
|
37 |
in case get_first (fn (_, f) => try (f ctxt) t) evaluators |
|
38 |
of SOME t' => t' |
|
39 |
| NONE => evaluator ctxt t |
|
40 |
end end; |
|
41 |
||
43612 | 42 |
fun value_maybe_select some_name = |
43 |
case some_name |
|
44 |
of NONE => value |
|
45 |
| SOME name => value_select name; |
|
46 |
||
28227 | 47 |
fun value_cmd some_name modes raw_t state = |
48 |
let |
|
49 |
val ctxt = Toplevel.context_of state; |
|
50 |
val t = Syntax.read_term ctxt raw_t; |
|
43612 | 51 |
val t' = value_maybe_select some_name ctxt t; |
28227 | 52 |
val ty' = Term.type_of t'; |
31218 | 53 |
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
|
54 |
val p = Print_Mode.with_modes modes (fn () => |
28227 | 55 |
Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk, |
56 |
Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) (); |
|
57 |
in Pretty.writeln p end; |
|
58 |
||
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
33522
diff
changeset
|
59 |
val opt_modes = |
46949 | 60 |
Scan.optional (@{keyword "("} |-- Parse.!!! (Scan.repeat1 Parse.xname --| @{keyword ")"})) []; |
28227 | 61 |
|
43612 | 62 |
val opt_evaluator = |
46949 | 63 |
Scan.option (@{keyword "["} |-- Parse.xname --| @{keyword "]"}) |
43612 | 64 |
|
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
33522
diff
changeset
|
65 |
val _ = |
46961
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents:
46949
diff
changeset
|
66 |
Outer_Syntax.improper_command @{command_spec "value"} "evaluate and print term" |
43612 | 67 |
(opt_evaluator -- opt_modes -- Parse.term |
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
33522
diff
changeset
|
68 |
>> (fn ((some_name, modes), t) => Toplevel.no_timing o Toplevel.keep |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
33522
diff
changeset
|
69 |
(value_cmd some_name modes t))); |
28227 | 70 |
|
43612 | 71 |
val antiq_setup = |
43619
3803869014aa
proper @{binding} antiquotations (relevant for formal references);
wenzelm
parents:
43612
diff
changeset
|
72 |
Thy_Output.antiquotation @{binding value} |
43612 | 73 |
(Scan.lift opt_evaluator -- Term_Style.parse -- Args.term) |
74 |
(fn {source, context, ...} => fn ((some_name, style), t) => Thy_Output.output context |
|
75 |
(Thy_Output.maybe_pretty_source Thy_Output.pretty_term context source |
|
76 |
[style (value_maybe_select some_name context t)])); |
|
77 |
||
78 |
val setup = antiq_setup; |
|
79 |
||
28227 | 80 |
end; |