author | haftmann |
Thu, 26 May 2016 15:27:50 +0200 | |
changeset 63160 | 80a91e0e236e |
parent 63073 | 413184c7a2a2 |
child 63554 | d7c6a3a01b79 |
permissions | -rw-r--r-- |
48985 | 1 |
(* Title: Doc/more_antiquote.ML |
28440 | 2 |
Author: Florian Haftmann, TU Muenchen |
3 |
||
63026 | 4 |
More antiquotations (partly depending on Isabelle/HOL). |
28440 | 5 |
*) |
6 |
||
56061 | 7 |
structure More_Antiquote : sig end = |
28440 | 8 |
struct |
9 |
||
63026 | 10 |
(* class specifications *) |
11 |
||
12 |
local |
|
13 |
||
14 |
fun class_spec ctxt s = |
|
15 |
let |
|
16 |
val thy = Proof_Context.theory_of ctxt; |
|
17 |
val class = Sign.intern_class thy s; |
|
18 |
in Thy_Output.output ctxt (Class.pretty_specification thy class) end; |
|
19 |
||
20 |
in |
|
21 |
||
22 |
val _ = |
|
23 |
Theory.setup (Thy_Output.antiquotation @{binding class_spec} (Scan.lift Args.name) |
|
24 |
(fn {context, ...} => class_spec context)); |
|
25 |
||
26 |
end; |
|
27 |
||
29397 | 28 |
(* code theorem antiquotation *) |
29 |
||
30 |
local |
|
31 |
||
32 |
fun pretty_term ctxt t = Syntax.pretty_term (Variable.auto_fixes t ctxt) t; |
|
33 |
||
34 |
fun pretty_thm ctxt = pretty_term ctxt o Thm.full_prop_of; |
|
35 |
||
36 |
fun no_vars ctxt thm = |
|
37 |
let |
|
38 |
val ctxt' = Variable.set_body false ctxt; |
|
31794
71af1fd6a5e4
renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
wenzelm
parents:
31156
diff
changeset
|
39 |
val ((_, [thm]), _) = Variable.import true [thm] ctxt'; |
29397 | 40 |
in thm end; |
41 |
||
63026 | 42 |
fun pretty_code_thm ctxt source raw_const = |
29397 | 43 |
let |
42361 | 44 |
val thy = Proof_Context.theory_of ctxt; |
31156 | 45 |
val const = Code.check_const thy raw_const; |
63160 | 46 |
val { eqngr, ... } = Code_Preproc.obtain true { ctxt = ctxt, consts = [const], terms = [] }; |
29874 | 47 |
fun holize thm = @{thm meta_eq_to_obj_eq} OF [thm]; |
34896
a22b09addd78
adjusted to changes in code equation administration
haftmann
parents:
34072
diff
changeset
|
48 |
val thms = Code_Preproc.cert eqngr const |
35246 | 49 |
|> Code.equations_of_cert thy |
34896
a22b09addd78
adjusted to changes in code equation administration
haftmann
parents:
34072
diff
changeset
|
50 |
|> snd |
54889
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents:
51685
diff
changeset
|
51 |
|> these |
35246 | 52 |
|> map_filter (fn (_, (some_thm, proper)) => if proper then some_thm else NONE) |
63073
413184c7a2a2
clarified context, notably for internal use of Simplifier;
wenzelm
parents:
63026
diff
changeset
|
53 |
|> map (holize o no_vars ctxt o Axclass.overload ctxt); |
63026 | 54 |
in Thy_Output.output ctxt (Thy_Output.maybe_pretty_source pretty_thm ctxt source thms) end; |
29397 | 55 |
|
56 |
in |
|
57 |
||
56061 | 58 |
val _ = |
59 |
Theory.setup (Thy_Output.antiquotation @{binding code_thms} Args.term |
|
63026 | 60 |
(fn {source, context, ...} => pretty_code_thm context source)); |
29397 | 61 |
|
62 |
end; |
|
63 |
||
28440 | 64 |
end; |