| author | Lars Hupel <lars.hupel@mytum.de> | 
| Wed, 28 Sep 2016 14:48:46 +0200 | |
| changeset 63951 | 8739c1cd2851 | 
| parent 63554 | d7c6a3a01b79 | 
| child 67386 | 998e01d6f8fd | 
| 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  | 
||
| 63554 | 28  | 
|
| 29397 | 29  | 
(* code theorem antiquotation *)  | 
30  | 
||
31  | 
local  | 
|
32  | 
||
33  | 
fun no_vars ctxt thm =  | 
|
34  | 
let  | 
|
35  | 
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
 | 
36  | 
val ((_, [thm]), _) = Variable.import true [thm] ctxt';  | 
| 29397 | 37  | 
in thm end;  | 
38  | 
||
| 63554 | 39  | 
fun pretty_code_thm ctxt raw_const =  | 
| 29397 | 40  | 
let  | 
| 42361 | 41  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 31156 | 42  | 
val const = Code.check_const thy raw_const;  | 
| 63160 | 43  | 
    val { eqngr, ... } = Code_Preproc.obtain true { ctxt = ctxt, consts = [const], terms = [] };
 | 
| 29874 | 44  | 
    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
 | 
45  | 
val thms = Code_Preproc.cert eqngr const  | 
| 35246 | 46  | 
|> Code.equations_of_cert thy  | 
| 
34896
 
a22b09addd78
adjusted to changes in code equation administration
 
haftmann 
parents: 
34072 
diff
changeset
 | 
47  | 
|> snd  | 
| 
54889
 
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
 
haftmann 
parents: 
51685 
diff
changeset
 | 
48  | 
|> these  | 
| 35246 | 49  | 
|> 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
 | 
50  | 
|> map (holize o no_vars ctxt o Axclass.overload ctxt);  | 
| 63554 | 51  | 
in Thy_Output.output ctxt (map (Thy_Output.pretty_thm ctxt) thms) end;  | 
| 29397 | 52  | 
|
53  | 
in  | 
|
54  | 
||
| 56061 | 55  | 
val _ =  | 
56  | 
  Theory.setup (Thy_Output.antiquotation @{binding code_thms} Args.term
 | 
|
| 63554 | 57  | 
    (fn {context, ...} => pretty_code_thm context));
 | 
| 29397 | 58  | 
|
59  | 
end;  | 
|
60  | 
||
| 28440 | 61  | 
end;  |