| author | smolkas | 
| Thu, 14 Feb 2013 22:49:22 +0100 | |
| changeset 51129 | 1edc2cc25f19 | 
| parent 48985 | 5386df44a037 | 
| child 51685 | 385ef6706252 | 
| permissions | -rw-r--r-- | 
| 48985 | 1  | 
(* Title: Doc/more_antiquote.ML  | 
| 28440 | 2  | 
Author: Florian Haftmann, TU Muenchen  | 
3  | 
||
4  | 
More antiquotations.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
signature MORE_ANTIQUOTE =  | 
|
8  | 
sig  | 
|
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
9  | 
val setup: theory -> theory  | 
| 28440 | 10  | 
end;  | 
11  | 
||
12  | 
structure More_Antiquote : MORE_ANTIQUOTE =  | 
|
13  | 
struct  | 
|
14  | 
||
| 29397 | 15  | 
(* code theorem antiquotation *)  | 
16  | 
||
17  | 
local  | 
|
18  | 
||
19  | 
fun pretty_term ctxt t = Syntax.pretty_term (Variable.auto_fixes t ctxt) t;  | 
|
20  | 
||
21  | 
fun pretty_thm ctxt = pretty_term ctxt o Thm.full_prop_of;  | 
|
22  | 
||
23  | 
fun no_vars ctxt thm =  | 
|
24  | 
let  | 
|
25  | 
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
 | 
26  | 
val ((_, [thm]), _) = Variable.import true [thm] ctxt';  | 
| 29397 | 27  | 
in thm end;  | 
28  | 
||
29  | 
fun pretty_code_thm src ctxt raw_const =  | 
|
30  | 
let  | 
|
| 42361 | 31  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 31156 | 32  | 
val const = Code.check_const thy raw_const;  | 
| 39540 | 33  | 
val (_, eqngr) = Code_Preproc.obtain true thy [const] [];  | 
| 29874 | 34  | 
    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
 | 
35  | 
val thms = Code_Preproc.cert eqngr const  | 
| 35246 | 36  | 
|> Code.equations_of_cert thy  | 
| 
34896
 
a22b09addd78
adjusted to changes in code equation administration
 
haftmann 
parents: 
34072 
diff
changeset
 | 
37  | 
|> snd  | 
| 35246 | 38  | 
|> map_filter (fn (_, (some_thm, proper)) => if proper then some_thm else NONE)  | 
| 29874 | 39  | 
|> 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
 | 
40  | 
in Thy_Output.output ctxt (Thy_Output.maybe_pretty_source pretty_thm ctxt src thms) end;  | 
| 29397 | 41  | 
|
42  | 
in  | 
|
43  | 
||
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
44  | 
val setup =  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
45  | 
  Thy_Output.antiquotation @{binding code_thms} Args.term
 | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
46  | 
    (fn {source, context, ...} => pretty_code_thm source context);
 | 
| 29397 | 47  | 
|
48  | 
end;  | 
|
49  | 
||
| 28440 | 50  | 
end;  |