| author | wenzelm | 
| Sun, 27 Dec 2015 22:07:17 +0100 | |
| changeset 61943 | 7fba644ed827 | 
| parent 60697 | e266d5463e9d | 
| child 63026 | 9a9c2d846d4a | 
| permissions | -rw-r--r-- | 
| 48985 | 1  | 
(* Title: Doc/more_antiquote.ML  | 
| 28440 | 2  | 
Author: Florian Haftmann, TU Muenchen  | 
3  | 
||
| 56061 | 4  | 
More antiquotations (depending on Isabelle/HOL).  | 
| 28440 | 5  | 
*)  | 
6  | 
||
| 56061 | 7  | 
structure More_Antiquote : sig end =  | 
| 28440 | 8  | 
struct  | 
9  | 
||
| 29397 | 10  | 
(* code theorem antiquotation *)  | 
11  | 
||
12  | 
local  | 
|
13  | 
||
14  | 
fun pretty_term ctxt t = Syntax.pretty_term (Variable.auto_fixes t ctxt) t;  | 
|
15  | 
||
16  | 
fun pretty_thm ctxt = pretty_term ctxt o Thm.full_prop_of;  | 
|
17  | 
||
18  | 
fun no_vars ctxt thm =  | 
|
19  | 
let  | 
|
20  | 
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
 | 
21  | 
val ((_, [thm]), _) = Variable.import true [thm] ctxt';  | 
| 29397 | 22  | 
in thm end;  | 
23  | 
||
24  | 
fun pretty_code_thm src ctxt raw_const =  | 
|
25  | 
let  | 
|
| 42361 | 26  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 31156 | 27  | 
val const = Code.check_const thy raw_const;  | 
| 60697 | 28  | 
val (_, eqngr) = Code_Preproc.obtain true ctxt [const] [];  | 
| 29874 | 29  | 
    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
 | 
30  | 
val thms = Code_Preproc.cert eqngr const  | 
| 35246 | 31  | 
|> Code.equations_of_cert thy  | 
| 
34896
 
a22b09addd78
adjusted to changes in code equation administration
 
haftmann 
parents: 
34072 
diff
changeset
 | 
32  | 
|> snd  | 
| 
54889
 
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
 
haftmann 
parents: 
51685 
diff
changeset
 | 
33  | 
|> these  | 
| 35246 | 34  | 
|> map_filter (fn (_, (some_thm, proper)) => if proper then some_thm else NONE)  | 
| 
51685
 
385ef6706252
more standard module name Axclass (according to file name);
 
wenzelm 
parents: 
48985 
diff
changeset
 | 
35  | 
|> map (holize o no_vars ctxt o Axclass.overload thy);  | 
| 
38767
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
36  | 
in Thy_Output.output ctxt (Thy_Output.maybe_pretty_source pretty_thm ctxt src thms) end;  | 
| 29397 | 37  | 
|
38  | 
in  | 
|
39  | 
||
| 56061 | 40  | 
val _ =  | 
41  | 
  Theory.setup (Thy_Output.antiquotation @{binding code_thms} Args.term
 | 
|
42  | 
    (fn {source, context, ...} => pretty_code_thm source context));
 | 
|
| 29397 | 43  | 
|
44  | 
end;  | 
|
45  | 
||
| 28440 | 46  | 
end;  |