| author | paulson <lp15@cam.ac.uk> | 
| Fri, 30 Aug 2024 10:16:48 +0100 | |
| changeset 80790 | 07c51801c2ea | 
| parent 73761 | ef1a18e20ace | 
| 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  | 
val _ =  | 
|
| 73761 | 13  | 
Theory.setup (Document_Output.antiquotation_pretty \<^binding>\<open>class_spec\<close> (Scan.lift Args.name)  | 
| 67463 | 14  | 
(fn ctxt => fn s =>  | 
15  | 
let  | 
|
16  | 
val thy = Proof_Context.theory_of ctxt;  | 
|
17  | 
val class = Sign.intern_class thy s;  | 
|
| 
67505
 
ceb324e34c14
clarified signature: items with \isasep are special;
 
wenzelm 
parents: 
67469 
diff
changeset
 | 
18  | 
in Pretty.chunks (Class.pretty_specification thy class) end));  | 
| 63026 | 19  | 
|
| 63554 | 20  | 
|
| 29397 | 21  | 
(* code theorem antiquotation *)  | 
22  | 
||
| 56061 | 23  | 
val _ =  | 
| 73761 | 24  | 
Theory.setup (Document_Output.antiquotation_pretty \<^binding>\<open>code_thms\<close> Args.term  | 
| 67463 | 25  | 
(fn ctxt => fn raw_const =>  | 
26  | 
let  | 
|
27  | 
val thy = Proof_Context.theory_of ctxt;  | 
|
28  | 
val const = Code.check_const thy raw_const;  | 
|
29  | 
        val { eqngr, ... } = Code_Preproc.obtain true { ctxt = ctxt, consts = [const], terms = [] };
 | 
|
30  | 
val thms = Code_Preproc.cert eqngr const  | 
|
31  | 
|> Code.equations_of_cert thy  | 
|
32  | 
|> snd  | 
|
33  | 
|> these  | 
|
34  | 
|> map_filter (fn (_, (some_thm, proper)) => if proper then some_thm else NONE)  | 
|
| 70304 | 35  | 
|> map (HOLogic.mk_obj_eq o Variable.import_vars ctxt o Axclass.overload ctxt);  | 
| 73761 | 36  | 
in Pretty.chunks (map (Document_Output.pretty_thm ctxt) thms) end));  | 
| 29397 | 37  | 
|
38  | 
end;  |