| author | sultana | 
| Wed, 07 Mar 2012 13:00:30 +0000 | |
| changeset 46824 | 1257c80988cd | 
| parent 46811 | 03a2dc9e0624 | 
| child 46922 | 3717f3878714 | 
| permissions | -rw-r--r-- | 
| 22124 | 1  | 
(* Title: Pure/Thy/thy_output.ML  | 
2  | 
Author: Markus Wenzel, TU Muenchen  | 
|
3  | 
||
| 
30390
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
4  | 
Theory document output with antiquotations.  | 
| 22124 | 5  | 
*)  | 
6  | 
||
7  | 
signature THY_OUTPUT =  | 
|
8  | 
sig  | 
|
| 
38767
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
9  | 
val display_default: bool Unsynchronized.ref  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
10  | 
val quotes_default: bool Unsynchronized.ref  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
11  | 
val indent_default: int Unsynchronized.ref  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
12  | 
val source_default: bool Unsynchronized.ref  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
13  | 
val break_default: bool Unsynchronized.ref  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
14  | 
val display: bool Config.T  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
15  | 
val quotes: bool Config.T  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
16  | 
val indent: int Config.T  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
17  | 
val source: bool Config.T  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
18  | 
val break: bool Config.T  | 
| 
38766
 
8891f4520d16
Thy_Output: options based on proper context, although Thy_Output.add_wrapper still allows to maintain old-style wrapper combinators (setmp_CRITICAL etc.);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
19  | 
val add_wrapper: ((unit -> string) -> unit -> string) -> Proof.context -> Proof.context  | 
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
20  | 
val add_option: binding -> (string -> Proof.context -> Proof.context) -> theory -> theory  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
21  | 
val intern_command: theory -> xstring -> string  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
22  | 
val defined_command: theory -> string -> bool  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
23  | 
val intern_option: theory -> xstring -> string  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
24  | 
val defined_option: theory -> string -> bool  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
25  | 
val print_antiquotations: Proof.context -> unit  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
26  | 
val antiquotation: binding -> 'a context_parser ->  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
27  | 
    ({source: Args.src, state: Toplevel.state, context: Proof.context} -> 'a -> string) ->
 | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
28  | 
theory -> theory  | 
| 22124 | 29  | 
val boolean: string -> bool  | 
30  | 
val integer: string -> int  | 
|
31  | 
datatype markup = Markup | MarkupEnv | Verbatim  | 
|
| 32738 | 32  | 
val modes: string list Unsynchronized.ref  | 
| 
42508
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42399 
diff
changeset
 | 
33  | 
val eval_antiq: Scan.lexicon -> Toplevel.state -> Symbol_Pos.T list * Position.range -> string  | 
| 30573 | 34  | 
val eval_antiquote: Scan.lexicon -> Toplevel.state -> Symbol_Pos.text * Position.T -> string  | 
| 
28427
 
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
 
wenzelm 
parents: 
28273 
diff
changeset
 | 
35  | 
val present_thy: Scan.lexicon -> (string -> string list) -> (markup -> string -> bool) ->  | 
| 
46811
 
03a2dc9e0624
clarified command span: include trailing whitespace/comments and thus reduce number of ignored spans with associated transactions and states (factor 2);
 
wenzelm 
parents: 
46261 
diff
changeset
 | 
36  | 
(Toplevel.transition * Toplevel.state) list -> Token.T list -> Buffer.T  | 
| 
38767
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
37  | 
val pretty_text: Proof.context -> string -> Pretty.T  | 
| 28644 | 38  | 
val pretty_term: Proof.context -> term -> Pretty.T  | 
39  | 
val pretty_thm: Proof.context -> thm -> Pretty.T  | 
|
| 
30390
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
40  | 
val str_of_source: Args.src -> string  | 
| 
38767
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
41  | 
val maybe_pretty_source: (Proof.context -> 'a -> Pretty.T) -> Proof.context ->  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
42  | 
Args.src -> 'a list -> Pretty.T list  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
43  | 
val output: Proof.context -> Pretty.T list -> string  | 
| 22124 | 44  | 
end;  | 
45  | 
||
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
46  | 
structure Thy_Output: THY_OUTPUT =  | 
| 22124 | 47  | 
struct  | 
48  | 
||
49  | 
(** global options **)  | 
|
50  | 
||
| 
38767
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
51  | 
val display_default = Unsynchronized.ref false;  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
52  | 
val quotes_default = Unsynchronized.ref false;  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
53  | 
val indent_default = Unsynchronized.ref 0;  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
54  | 
val source_default = Unsynchronized.ref false;  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
55  | 
val break_default = Unsynchronized.ref false;  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
56  | 
|
| 
42616
 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 
wenzelm 
parents: 
42508 
diff
changeset
 | 
57  | 
val display = Attrib.setup_config_bool (Binding.name "thy_output_display") (fn _ => ! display_default);  | 
| 
 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 
wenzelm 
parents: 
42508 
diff
changeset
 | 
58  | 
val quotes = Attrib.setup_config_bool (Binding.name "thy_output_quotes") (fn _ => ! quotes_default);  | 
| 
 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 
wenzelm 
parents: 
42508 
diff
changeset
 | 
59  | 
val indent = Attrib.setup_config_int (Binding.name "thy_output_indent") (fn _ => ! indent_default);  | 
| 
 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 
wenzelm 
parents: 
42508 
diff
changeset
 | 
60  | 
val source = Attrib.setup_config_bool (Binding.name "thy_output_source") (fn _ => ! source_default);  | 
| 
 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 
wenzelm 
parents: 
42508 
diff
changeset
 | 
61  | 
val break = Attrib.setup_config_bool (Binding.name "thy_output_break") (fn _ => ! break_default);  | 
| 22124 | 62  | 
|
63  | 
||
| 
38766
 
8891f4520d16
Thy_Output: options based on proper context, although Thy_Output.add_wrapper still allows to maintain old-style wrapper combinators (setmp_CRITICAL etc.);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
64  | 
structure Wrappers = Proof_Data  | 
| 
 
8891f4520d16
Thy_Output: options based on proper context, although Thy_Output.add_wrapper still allows to maintain old-style wrapper combinators (setmp_CRITICAL etc.);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
65  | 
(  | 
| 
 
8891f4520d16
Thy_Output: options based on proper context, although Thy_Output.add_wrapper still allows to maintain old-style wrapper combinators (setmp_CRITICAL etc.);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
66  | 
type T = ((unit -> string) -> unit -> string) list;  | 
| 
 
8891f4520d16
Thy_Output: options based on proper context, although Thy_Output.add_wrapper still allows to maintain old-style wrapper combinators (setmp_CRITICAL etc.);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
67  | 
fun init _ = [];  | 
| 
 
8891f4520d16
Thy_Output: options based on proper context, although Thy_Output.add_wrapper still allows to maintain old-style wrapper combinators (setmp_CRITICAL etc.);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
68  | 
);  | 
| 
 
8891f4520d16
Thy_Output: options based on proper context, although Thy_Output.add_wrapper still allows to maintain old-style wrapper combinators (setmp_CRITICAL etc.);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
69  | 
|
| 
 
8891f4520d16
Thy_Output: options based on proper context, although Thy_Output.add_wrapper still allows to maintain old-style wrapper combinators (setmp_CRITICAL etc.);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
70  | 
fun add_wrapper wrapper = Wrappers.map (cons wrapper);  | 
| 
 
8891f4520d16
Thy_Output: options based on proper context, although Thy_Output.add_wrapper still allows to maintain old-style wrapper combinators (setmp_CRITICAL etc.);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
71  | 
|
| 
 
8891f4520d16
Thy_Output: options based on proper context, although Thy_Output.add_wrapper still allows to maintain old-style wrapper combinators (setmp_CRITICAL etc.);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
72  | 
val wrap = Wrappers.get #> fold (fn wrapper => fn f => wrapper f);  | 
| 
 
8891f4520d16
Thy_Output: options based on proper context, although Thy_Output.add_wrapper still allows to maintain old-style wrapper combinators (setmp_CRITICAL etc.);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
73  | 
|
| 
 
8891f4520d16
Thy_Output: options based on proper context, although Thy_Output.add_wrapper still allows to maintain old-style wrapper combinators (setmp_CRITICAL etc.);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
74  | 
|
| 22124 | 75  | 
|
| 
30390
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
76  | 
(** maintain global antiquotations **)  | 
| 22124 | 77  | 
|
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
78  | 
structure Antiquotations = Theory_Data  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
79  | 
(  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
80  | 
type T =  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
81  | 
(Args.src -> Toplevel.state -> Proof.context -> string) Name_Space.table *  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
82  | 
(string -> Proof.context -> Proof.context) Name_Space.table;  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
83  | 
val empty : T =  | 
| 45666 | 84  | 
(Name_Space.empty_table Isabelle_Markup.doc_antiquotationN,  | 
85  | 
Name_Space.empty_table Isabelle_Markup.doc_antiquotation_optionN);  | 
|
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
86  | 
val extend = I;  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
87  | 
fun merge ((commands1, options1), (commands2, options2)) : T =  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
88  | 
(Name_Space.merge_tables (commands1, commands2),  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
89  | 
Name_Space.merge_tables (options1, options2));  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
90  | 
);  | 
| 22124 | 91  | 
|
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
92  | 
fun add_command name cmd thy =  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
93  | 
thy |> Antiquotations.map  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
94  | 
(apfst  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
95  | 
(Name_Space.define  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
96  | 
(Proof_Context.init_global thy) true (Sign.naming_of thy) (name, cmd) #> snd));  | 
| 22124 | 97  | 
|
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
98  | 
fun add_option name opt thy =  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
99  | 
thy |> Antiquotations.map  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
100  | 
(apsnd  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
101  | 
(Name_Space.define  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
102  | 
(Proof_Context.init_global thy) true (Sign.naming_of thy) (name, opt) #> snd));  | 
| 22124 | 103  | 
|
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
104  | 
val intern_command = Name_Space.intern o #1 o #1 o Antiquotations.get;  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
105  | 
val defined_command = Symtab.defined o #2 o #1 o Antiquotations.get;  | 
| 26893 | 106  | 
|
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
107  | 
val intern_option = Name_Space.intern o #1 o #2 o Antiquotations.get;  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
108  | 
val defined_option = Symtab.defined o #2 o #2 o Antiquotations.get;  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
109  | 
|
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
110  | 
fun command src state ctxt =  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
111  | 
let  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
112  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
113  | 
val ((xname, _), pos) = Args.dest_src src;  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
114  | 
val (name, f) = Name_Space.check ctxt (#1 (Antiquotations.get thy)) (xname, pos);  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
115  | 
in  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
116  | 
f src state ctxt handle ERROR msg =>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
117  | 
      cat_error msg ("The error(s) above occurred in document antiquotation: " ^
 | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
118  | 
quote name ^ Position.str_of pos)  | 
| 22124 | 119  | 
end;  | 
120  | 
||
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
121  | 
fun option ((xname, pos), s) ctxt =  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
122  | 
let  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
123  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
124  | 
val (_, opt) = Name_Space.check ctxt (#2 (Antiquotations.get thy)) (xname, pos);  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
125  | 
in opt s ctxt end;  | 
| 22124 | 126  | 
|
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
127  | 
fun print_antiquotations ctxt =  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
128  | 
let  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
129  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
130  | 
val (commands, options) = Antiquotations.get thy;  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
131  | 
val command_names = map #1 (Name_Space.extern_table ctxt commands);  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
132  | 
val option_names = map #1 (Name_Space.extern_table ctxt options);  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
133  | 
in  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
134  | 
[Pretty.big_list "document antiquotations:" (map Pretty.str command_names),  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
135  | 
Pretty.big_list "document antiquotation options:" (map Pretty.str option_names)]  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
136  | 
|> Pretty.chunks |> Pretty.writeln  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
137  | 
end;  | 
| 22124 | 138  | 
|
| 
38766
 
8891f4520d16
Thy_Output: options based on proper context, although Thy_Output.add_wrapper still allows to maintain old-style wrapper combinators (setmp_CRITICAL etc.);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
139  | 
fun antiquotation name scan out =  | 
| 
 
8891f4520d16
Thy_Output: options based on proper context, although Thy_Output.add_wrapper still allows to maintain old-style wrapper combinators (setmp_CRITICAL etc.);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
140  | 
add_command name  | 
| 
 
8891f4520d16
Thy_Output: options based on proper context, although Thy_Output.add_wrapper still allows to maintain old-style wrapper combinators (setmp_CRITICAL etc.);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
141  | 
(fn src => fn state => fn ctxt =>  | 
| 
 
8891f4520d16
Thy_Output: options based on proper context, although Thy_Output.add_wrapper still allows to maintain old-style wrapper combinators (setmp_CRITICAL etc.);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
142  | 
let val (x, ctxt') = Args.context_syntax "document antiquotation" scan src ctxt  | 
| 
 
8891f4520d16
Thy_Output: options based on proper context, although Thy_Output.add_wrapper still allows to maintain old-style wrapper combinators (setmp_CRITICAL etc.);
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
143  | 
      in out {source = src, state = state, context = ctxt'} x end);
 | 
| 
30390
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
144  | 
|
| 22124 | 145  | 
|
146  | 
||
147  | 
(** syntax of antiquotations **)  | 
|
148  | 
||
149  | 
(* option values *)  | 
|
150  | 
||
151  | 
fun boolean "" = true  | 
|
152  | 
| boolean "true" = true  | 
|
153  | 
| boolean "false" = false  | 
|
154  | 
  | boolean s = error ("Bad boolean value: " ^ quote s);
 | 
|
155  | 
||
156  | 
fun integer s =  | 
|
157  | 
let  | 
|
158  | 
fun int ss =  | 
|
159  | 
(case Library.read_int ss of (i, []) => i  | 
|
160  | 
      | _ => error ("Bad integer value: " ^ quote s));
 | 
|
161  | 
in (case Symbol.explode s of "-" :: ss => ~ (int ss) | ss => int ss) end;  | 
|
162  | 
||
163  | 
||
164  | 
(* outer syntax *)  | 
|
165  | 
||
166  | 
local  | 
|
167  | 
||
| 36950 | 168  | 
val property =  | 
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
169  | 
Parse.position Parse.xname -- Scan.optional (Parse.$$$ "=" |-- Parse.!!! Parse.xname) "";  | 
| 36950 | 170  | 
|
171  | 
val properties =  | 
|
172  | 
Scan.optional (Parse.$$$ "[" |-- Parse.!!! (Parse.enum "," property --| Parse.$$$ "]")) [];  | 
|
| 22124 | 173  | 
|
174  | 
in  | 
|
175  | 
||
| 36950 | 176  | 
val antiq =  | 
| 
40800
 
330eb65c9469
Parse.liberal_name for document antiquotations and attributes;
 
wenzelm 
parents: 
39689 
diff
changeset
 | 
177  | 
Parse.!!!  | 
| 
 
330eb65c9469
Parse.liberal_name for document antiquotations and attributes;
 
wenzelm 
parents: 
39689 
diff
changeset
 | 
178  | 
(Parse.position Parse.liberal_name -- properties -- Args.parse --| Scan.ahead Parse.eof)  | 
| 22124 | 179  | 
>> (fn (((x, pos), y), z) => (y, Args.src ((x, z), pos)));  | 
180  | 
||
181  | 
end;  | 
|
182  | 
||
183  | 
||
184  | 
(* eval_antiquote *)  | 
|
185  | 
||
| 32738 | 186  | 
val modes = Unsynchronized.ref ([]: string list);  | 
| 22124 | 187  | 
|
| 
42508
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42399 
diff
changeset
 | 
188  | 
fun eval_antiq lex state (ss, (pos, _)) =  | 
| 
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42399 
diff
changeset
 | 
189  | 
let  | 
| 
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42399 
diff
changeset
 | 
190  | 
val (opts, src) = Token.read_antiq lex antiq (ss, pos);  | 
| 
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42399 
diff
changeset
 | 
191  | 
fun cmd ctxt = wrap ctxt (fn () => command src state ctxt) ();  | 
| 
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42399 
diff
changeset
 | 
192  | 
val preview_ctxt = fold option opts (Toplevel.presentation_context_of state);  | 
| 
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42399 
diff
changeset
 | 
193  | 
val print_ctxt = Context_Position.set_visible false preview_ctxt;  | 
| 
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42399 
diff
changeset
 | 
194  | 
val _ = cmd preview_ctxt;  | 
| 
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42399 
diff
changeset
 | 
195  | 
in Print_Mode.with_modes (! modes @ Latex.modes) (fn () => cmd print_ctxt) () end;  | 
| 
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42399 
diff
changeset
 | 
196  | 
|
| 30367 | 197  | 
fun eval_antiquote lex state (txt, pos) =  | 
| 22124 | 198  | 
let  | 
| 30589 | 199  | 
fun expand (Antiquote.Text ss) = Symbol_Pos.content ss  | 
| 
42508
 
e21362bf1d93
allow nested @{antiq} (nonterminal) and @@{antiq} terminal;
 
wenzelm 
parents: 
42399 
diff
changeset
 | 
200  | 
| expand (Antiquote.Antiq antiq) = eval_antiq lex state antiq  | 
| 27344 | 201  | 
| expand (Antiquote.Open _) = ""  | 
202  | 
| expand (Antiquote.Close _) = "";  | 
|
| 30642 | 203  | 
val ants = Antiquote.read (Symbol_Pos.explode (txt, pos), pos);  | 
| 22124 | 204  | 
in  | 
| 
30635
 
a7d175c48228
replaced Antiquote.is_antiq by Antiquote.is_text;
 
wenzelm 
parents: 
30613 
diff
changeset
 | 
205  | 
if Toplevel.is_toplevel state andalso not (forall Antiquote.is_text ants) then  | 
| 27897 | 206  | 
      error ("Unknown context -- cannot expand document antiquotations" ^ Position.str_of pos)
 | 
| 22124 | 207  | 
else implode (map expand ants)  | 
208  | 
end;  | 
|
209  | 
||
210  | 
||
211  | 
||
212  | 
(** present theory source **)  | 
|
213  | 
||
214  | 
(*NB: arranging white space around command spans is a black art.*)  | 
|
215  | 
||
216  | 
(* presentation tokens *)  | 
|
217  | 
||
218  | 
datatype token =  | 
|
219  | 
NoToken  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
220  | 
| BasicToken of Token.T  | 
| 22124 | 221  | 
| MarkupToken of string * (string * Position.T)  | 
222  | 
| MarkupEnvToken of string * (string * Position.T)  | 
|
223  | 
| VerbatimToken of string * Position.T;  | 
|
224  | 
||
225  | 
fun output_token lex state =  | 
|
| 30367 | 226  | 
let val eval = eval_antiquote lex state in  | 
| 22124 | 227  | 
fn NoToken => ""  | 
228  | 
| BasicToken tok => Latex.output_basic tok  | 
|
229  | 
| MarkupToken (cmd, txt) => Latex.output_markup cmd (eval txt)  | 
|
230  | 
| MarkupEnvToken (cmd, txt) => Latex.output_markup_env cmd (eval txt)  | 
|
231  | 
| VerbatimToken txt => Latex.output_verbatim (eval txt)  | 
|
232  | 
end;  | 
|
233  | 
||
234  | 
fun basic_token pred (BasicToken tok) = pred tok  | 
|
235  | 
| basic_token _ _ = false;  | 
|
236  | 
||
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
237  | 
val improper_token = basic_token (not o Token.is_proper);  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
238  | 
val comment_token = basic_token Token.is_comment;  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
239  | 
val blank_token = basic_token Token.is_blank;  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
240  | 
val newline_token = basic_token Token.is_newline;  | 
| 22124 | 241  | 
|
242  | 
||
243  | 
(* command spans *)  | 
|
244  | 
||
245  | 
type command = string * Position.T * string list; (*name, position, tags*)  | 
|
246  | 
type source = (token * (string * int)) list; (*token, markup flag, meta-comment depth*)  | 
|
247  | 
||
248  | 
datatype span = Span of command * (source * source * source * source) * bool;  | 
|
249  | 
||
250  | 
fun make_span cmd src =  | 
|
251  | 
let  | 
|
252  | 
fun take_newline (tok :: toks) =  | 
|
253  | 
if newline_token (fst tok) then ([tok], toks, true)  | 
|
254  | 
else ([], tok :: toks, false)  | 
|
255  | 
| take_newline [] = ([], [], false);  | 
|
256  | 
val (((src_prefix, src_main), src_suffix1), (src_suffix2, src_appendix, newline)) =  | 
|
257  | 
src  | 
|
258  | 
|> take_prefix (improper_token o fst)  | 
|
259  | 
||>> take_suffix (improper_token o fst)  | 
|
260  | 
||>> take_prefix (comment_token o fst)  | 
|
261  | 
||> take_newline;  | 
|
262  | 
in Span (cmd, (src_prefix, src_main, src_suffix1 @ src_suffix2, src_appendix), newline) end;  | 
|
263  | 
||
264  | 
||
265  | 
(* present spans *)  | 
|
266  | 
||
267  | 
local  | 
|
268  | 
||
269  | 
fun err_bad_nesting pos =  | 
|
270  | 
  error ("Bad nesting of commands in presentation" ^ pos);
 | 
|
271  | 
||
272  | 
fun edge which f (x: string option, y) =  | 
|
273  | 
if x = y then I  | 
|
274  | 
else (case which (x, y) of NONE => I | SOME txt => Buffer.add (f txt));  | 
|
275  | 
||
276  | 
val begin_tag = edge #2 Latex.begin_tag;  | 
|
277  | 
val end_tag = edge #1 Latex.end_tag;  | 
|
278  | 
fun open_delim delim e = edge #2 Latex.begin_delim e #> delim #> edge #2 Latex.end_delim e;  | 
|
279  | 
fun close_delim delim e = edge #1 Latex.begin_delim e #> delim #> edge #1 Latex.end_delim e;  | 
|
280  | 
||
281  | 
in  | 
|
282  | 
||
283  | 
fun present_span lex default_tags span state state'  | 
|
284  | 
(tag_stack, active_tag, newline, buffer, present_cont) =  | 
|
285  | 
let  | 
|
286  | 
val present = fold (fn (tok, (flag, 0)) =>  | 
|
287  | 
Buffer.add (output_token lex state' tok)  | 
|
288  | 
#> Buffer.add flag  | 
|
289  | 
| _ => I);  | 
|
290  | 
||
291  | 
val Span ((cmd_name, cmd_pos, cmd_tags), srcs, span_newline) = span;  | 
|
292  | 
||
293  | 
val (tag, tags) = tag_stack;  | 
|
| 36950 | 294  | 
val tag' = try hd (fold Keyword.update_tags cmd_tags (the_list tag));  | 
| 22124 | 295  | 
|
296  | 
val active_tag' =  | 
|
297  | 
if is_some tag' then tag'  | 
|
298  | 
else if cmd_name = "end" andalso not (Toplevel.is_toplevel state') then NONE  | 
|
299  | 
else try hd (default_tags cmd_name);  | 
|
300  | 
val edge = (active_tag, active_tag');  | 
|
301  | 
||
302  | 
val newline' =  | 
|
303  | 
if is_none active_tag' then span_newline else newline;  | 
|
304  | 
||
305  | 
val nesting = Toplevel.level state' - Toplevel.level state;  | 
|
306  | 
val tag_stack' =  | 
|
307  | 
if nesting = 0 andalso not (Toplevel.is_proof state) then tag_stack  | 
|
308  | 
else if nesting >= 0 then (tag', replicate nesting tag @ tags)  | 
|
309  | 
else  | 
|
| 33957 | 310  | 
(case drop (~ nesting - 1) tags of  | 
| 22124 | 311  | 
tgs :: tgss => (tgs, tgss)  | 
312  | 
| [] => err_bad_nesting (Position.str_of cmd_pos));  | 
|
313  | 
||
314  | 
val buffer' =  | 
|
315  | 
buffer  | 
|
316  | 
|> end_tag edge  | 
|
317  | 
|> close_delim (fst present_cont) edge  | 
|
318  | 
|> snd present_cont  | 
|
319  | 
|> open_delim (present (#1 srcs)) edge  | 
|
320  | 
|> begin_tag edge  | 
|
321  | 
|> present (#2 srcs);  | 
|
322  | 
val present_cont' =  | 
|
323  | 
if newline then (present (#3 srcs), present (#4 srcs))  | 
|
324  | 
else (I, present (#3 srcs) #> present (#4 srcs));  | 
|
325  | 
in (tag_stack', active_tag', newline', buffer', present_cont') end;  | 
|
326  | 
||
327  | 
fun present_trailer ((_, tags), active_tag, _, buffer, present_cont) =  | 
|
328  | 
if not (null tags) then err_bad_nesting " at end of theory"  | 
|
329  | 
else  | 
|
330  | 
buffer  | 
|
331  | 
|> end_tag (active_tag, NONE)  | 
|
332  | 
|> close_delim (fst present_cont) (active_tag, NONE)  | 
|
333  | 
|> snd present_cont;  | 
|
334  | 
||
335  | 
end;  | 
|
336  | 
||
337  | 
||
| 
28427
 
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
 
wenzelm 
parents: 
28273 
diff
changeset
 | 
338  | 
(* present_thy *)  | 
| 22124 | 339  | 
|
340  | 
datatype markup = Markup | MarkupEnv | Verbatim;  | 
|
341  | 
||
342  | 
local  | 
|
343  | 
||
344  | 
val space_proper =  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
345  | 
Scan.one Token.is_blank -- Scan.many Token.is_comment -- Scan.one Token.is_proper;  | 
| 22124 | 346  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
347  | 
val is_improper = not o (Token.is_proper orf Token.is_begin_ignore orf Token.is_end_ignore);  | 
| 22124 | 348  | 
val improper = Scan.many is_improper;  | 
349  | 
val improper_end = Scan.repeat (Scan.unless space_proper (Scan.one is_improper));  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
350  | 
val blank_end = Scan.repeat (Scan.unless space_proper (Scan.one Token.is_blank));  | 
| 22124 | 351  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
352  | 
val opt_newline = Scan.option (Scan.one Token.is_newline);  | 
| 22124 | 353  | 
|
354  | 
val ignore =  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
355  | 
Scan.depend (fn d => opt_newline |-- Scan.one Token.is_begin_ignore  | 
| 22124 | 356  | 
>> pair (d + 1)) ||  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
357  | 
Scan.depend (fn d => Scan.one Token.is_end_ignore --|  | 
| 
43947
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
43564 
diff
changeset
 | 
358  | 
(if d = 0 then Scan.fail_with (K (fn () => "Bad nesting of meta-comments")) else opt_newline)  | 
| 22124 | 359  | 
>> pair (d - 1));  | 
360  | 
||
| 36950 | 361  | 
val tag = (improper -- Parse.$$$ "%" -- improper) |-- Parse.!!! (Parse.tag_name --| blank_end);  | 
| 22124 | 362  | 
|
363  | 
val locale =  | 
|
| 36950 | 364  | 
  Scan.option ((Parse.$$$ "(" -- improper -- Parse.$$$ "in") |--
 | 
365  | 
Parse.!!! (improper |-- Parse.xname --| (improper -- Parse.$$$ ")")));  | 
|
| 22124 | 366  | 
|
367  | 
in  | 
|
368  | 
||
| 
46811
 
03a2dc9e0624
clarified command span: include trailing whitespace/comments and thus reduce number of ignored spans with associated transactions and states (factor 2);
 
wenzelm 
parents: 
46261 
diff
changeset
 | 
369  | 
fun present_thy lex default_tags is_markup command_results toks =  | 
| 22124 | 370  | 
let  | 
371  | 
(* tokens *)  | 
|
372  | 
||
373  | 
val ignored = Scan.state --| ignore  | 
|
374  | 
      >> (fn d => (NONE, (NoToken, ("", d))));
 | 
|
375  | 
||
376  | 
fun markup mark mk flag = Scan.peek (fn d =>  | 
|
| 36950 | 377  | 
improper |--  | 
| 
46811
 
03a2dc9e0624
clarified command span: include trailing whitespace/comments and thus reduce number of ignored spans with associated transactions and states (factor 2);
 
wenzelm 
parents: 
46261 
diff
changeset
 | 
378  | 
Parse.position (Scan.one (Token.is_command andf is_markup mark o Token.content_of)) --  | 
| 22124 | 379  | 
Scan.repeat tag --  | 
| 36950 | 380  | 
Parse.!!!! ((improper -- locale -- improper) |-- Parse.doc_source --| improper_end)  | 
| 22124 | 381  | 
>> (fn (((tok, pos), tags), txt) =>  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
382  | 
let val name = Token.content_of tok  | 
| 22124 | 383  | 
in (SOME (name, pos, tags), (mk (name, txt), (flag, d))) end));  | 
384  | 
||
385  | 
val command = Scan.peek (fn d =>  | 
|
| 
46811
 
03a2dc9e0624
clarified command span: include trailing whitespace/comments and thus reduce number of ignored spans with associated transactions and states (factor 2);
 
wenzelm 
parents: 
46261 
diff
changeset
 | 
386  | 
Parse.position (Scan.one (Token.is_command)) --  | 
| 22124 | 387  | 
Scan.repeat tag  | 
388  | 
>> (fn ((tok, pos), tags) =>  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
389  | 
let val name = Token.content_of tok  | 
| 22124 | 390  | 
in (SOME (name, pos, tags), (BasicToken tok, (Latex.markup_false, d))) end));  | 
391  | 
||
392  | 
val cmt = Scan.peek (fn d =>  | 
|
| 36950 | 393  | 
Parse.$$$ "--" |-- Parse.!!!! (improper |-- Parse.doc_source)  | 
| 22124 | 394  | 
      >> (fn txt => (NONE, (MarkupToken ("cmt", txt), ("", d)))));
 | 
395  | 
||
396  | 
val other = Scan.peek (fn d =>  | 
|
| 36950 | 397  | 
       Parse.not_eof >> (fn tok => (NONE, (BasicToken tok, ("", d)))));
 | 
| 22124 | 398  | 
|
399  | 
val token =  | 
|
400  | 
ignored ||  | 
|
401  | 
markup Markup MarkupToken Latex.markup_true ||  | 
|
402  | 
markup MarkupEnv MarkupEnvToken Latex.markup_true ||  | 
|
403  | 
markup Verbatim (VerbatimToken o #2) "" ||  | 
|
404  | 
command || cmt || other;  | 
|
405  | 
||
406  | 
||
407  | 
(* spans *)  | 
|
408  | 
||
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
409  | 
val is_eof = fn (_, (BasicToken x, _)) => Token.is_eof x | _ => false;  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
410  | 
    val stopper = Scan.stopper (K (NONE, (BasicToken Token.eof, ("", 0)))) is_eof;
 | 
| 22124 | 411  | 
|
412  | 
val cmd = Scan.one (is_some o fst);  | 
|
| 27732 | 413  | 
val non_cmd = Scan.one (is_none o fst andf not o is_eof) >> #2;  | 
| 22124 | 414  | 
|
415  | 
val comments = Scan.many (comment_token o fst o snd);  | 
|
416  | 
val blank = Scan.one (blank_token o fst o snd);  | 
|
417  | 
val newline = Scan.one (newline_token o fst o snd);  | 
|
418  | 
val before_cmd =  | 
|
419  | 
Scan.option (newline -- comments) --  | 
|
420  | 
Scan.option (newline -- comments) --  | 
|
421  | 
Scan.option (blank -- comments) -- cmd;  | 
|
422  | 
||
423  | 
val span =  | 
|
424  | 
Scan.repeat non_cmd -- cmd --  | 
|
425  | 
Scan.repeat (Scan.unless before_cmd non_cmd) --  | 
|
426  | 
Scan.option (newline >> (single o snd))  | 
|
427  | 
>> (fn (((toks1, (cmd, tok2)), toks3), tok4) =>  | 
|
428  | 
make_span (the cmd) (toks1 @ (tok2 :: (toks3 @ the_default [] tok4))));  | 
|
429  | 
||
430  | 
val spans =  | 
|
| 
46811
 
03a2dc9e0624
clarified command span: include trailing whitespace/comments and thus reduce number of ignored spans with associated transactions and states (factor 2);
 
wenzelm 
parents: 
46261 
diff
changeset
 | 
431  | 
Source.of_list (filter_out Token.is_semicolon toks)  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
432  | 
|> Source.source' 0 Token.stopper (Scan.error (Scan.bulk token)) NONE  | 
| 22124 | 433  | 
|> Source.source stopper (Scan.error (Scan.bulk span)) NONE  | 
434  | 
|> Source.exhaust;  | 
|
| 
28427
 
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
 
wenzelm 
parents: 
28273 
diff
changeset
 | 
435  | 
|
| 
 
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
 
wenzelm 
parents: 
28273 
diff
changeset
 | 
436  | 
|
| 
 
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
 
wenzelm 
parents: 
28273 
diff
changeset
 | 
437  | 
(* present commands *)  | 
| 
 
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
 
wenzelm 
parents: 
28273 
diff
changeset
 | 
438  | 
|
| 
 
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
 
wenzelm 
parents: 
28273 
diff
changeset
 | 
439  | 
fun present_command tr span st st' =  | 
| 
 
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
 
wenzelm 
parents: 
28273 
diff
changeset
 | 
440  | 
Toplevel.setmp_thread_position tr (present_span lex default_tags span st st');  | 
| 
 
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
 
wenzelm 
parents: 
28273 
diff
changeset
 | 
441  | 
|
| 
 
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
 
wenzelm 
parents: 
28273 
diff
changeset
 | 
442  | 
fun present _ [] = I  | 
| 
 
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
 
wenzelm 
parents: 
28273 
diff
changeset
 | 
443  | 
| present st (((tr, st'), span) :: rest) = present_command tr span st st' #> present st' rest;  | 
| 22124 | 444  | 
in  | 
| 
28427
 
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
 
wenzelm 
parents: 
28273 
diff
changeset
 | 
445  | 
if length command_results = length spans then  | 
| 22124 | 446  | 
((NONE, []), NONE, true, Buffer.empty, (I, I))  | 
| 
28427
 
cc9f7d99fb73
turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
 
wenzelm 
parents: 
28273 
diff
changeset
 | 
447  | 
|> present Toplevel.toplevel (command_results ~~ spans)  | 
| 22124 | 448  | 
|> present_trailer  | 
449  | 
else error "Messed-up outer syntax for presentation"  | 
|
450  | 
end;  | 
|
451  | 
||
452  | 
end;  | 
|
453  | 
||
454  | 
||
455  | 
||
456  | 
(** setup default output **)  | 
|
457  | 
||
458  | 
(* options *)  | 
|
459  | 
||
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
460  | 
val _ =  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
461  | 
Context.>> (Context.map_theory  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
462  | 
(add_option (Binding.name "show_types") (Config.put show_types o boolean) #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
463  | 
add_option (Binding.name "show_sorts") (Config.put show_sorts o boolean) #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
464  | 
add_option (Binding.name "show_structs") (Config.put show_structs o boolean) #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
465  | 
add_option (Binding.name "show_question_marks") (Config.put show_question_marks o boolean) #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
466  | 
add_option (Binding.name "show_abbrevs") (Config.put show_abbrevs o boolean) #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
467  | 
add_option (Binding.name "names_long") (Config.put Name_Space.names_long o boolean) #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
468  | 
add_option (Binding.name "names_short") (Config.put Name_Space.names_short o boolean) #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
469  | 
add_option (Binding.name "names_unique") (Config.put Name_Space.names_unique o boolean) #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
470  | 
add_option (Binding.name "eta_contract") (Config.put Syntax_Trans.eta_contract o boolean) #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
471  | 
add_option (Binding.name "display") (Config.put display o boolean) #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
472  | 
add_option (Binding.name "break") (Config.put break o boolean) #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
473  | 
add_option (Binding.name "quotes") (Config.put quotes o boolean) #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
474  | 
add_option (Binding.name "mode") (add_wrapper o Print_Mode.with_modes o single) #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
475  | 
add_option (Binding.name "margin") (add_wrapper o setmp_CRITICAL Pretty.margin_default o integer) #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
476  | 
add_option (Binding.name "indent") (Config.put indent o integer) #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
477  | 
add_option (Binding.name "source") (Config.put source o boolean) #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
478  | 
add_option (Binding.name "goals_limit") (Config.put Goal_Display.goals_limit o integer)));  | 
| 22124 | 479  | 
|
480  | 
||
481  | 
(* basic pretty printing *)  | 
|
482  | 
||
| 
38767
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
483  | 
fun tweak_line ctxt s =  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
484  | 
if Config.get ctxt display then s else Symbol.strip_blanks s;  | 
| 22124 | 485  | 
|
| 
38767
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
486  | 
fun pretty_text ctxt =  | 
| 46195 | 487  | 
Pretty.chunks o map Pretty.str o map (tweak_line ctxt) o split_lines;  | 
| 22124 | 488  | 
|
| 
26710
 
f79aa228c582
pretty_term: no revert_skolems here, but auto_fixes (token translations will do the rest);
 
wenzelm 
parents: 
26455 
diff
changeset
 | 
489  | 
fun pretty_term ctxt t = Syntax.pretty_term (Variable.auto_fixes t ctxt) t;  | 
| 24920 | 490  | 
|
| 
32898
 
e871d897969c
term styles also cover antiquotations term_type and typeof
 
haftmann 
parents: 
32890 
diff
changeset
 | 
491  | 
fun pretty_thm ctxt = pretty_term ctxt o Thm.full_prop_of;  | 
| 
 
e871d897969c
term styles also cover antiquotations term_type and typeof
 
haftmann 
parents: 
32890 
diff
changeset
 | 
492  | 
|
| 
 
e871d897969c
term styles also cover antiquotations term_type and typeof
 
haftmann 
parents: 
32890 
diff
changeset
 | 
493  | 
fun pretty_term_style ctxt (style, t) =  | 
| 
 
e871d897969c
term styles also cover antiquotations term_type and typeof
 
haftmann 
parents: 
32890 
diff
changeset
 | 
494  | 
pretty_term ctxt (style t);  | 
| 22124 | 495  | 
|
| 
32898
 
e871d897969c
term styles also cover antiquotations term_type and typeof
 
haftmann 
parents: 
32890 
diff
changeset
 | 
496  | 
fun pretty_thm_style ctxt (style, th) =  | 
| 
 
e871d897969c
term styles also cover antiquotations term_type and typeof
 
haftmann 
parents: 
32890 
diff
changeset
 | 
497  | 
pretty_term ctxt (style (Thm.full_prop_of th));  | 
| 22124 | 498  | 
|
| 
32898
 
e871d897969c
term styles also cover antiquotations term_type and typeof
 
haftmann 
parents: 
32890 
diff
changeset
 | 
499  | 
fun pretty_term_typ ctxt (style, t) =  | 
| 
 
e871d897969c
term styles also cover antiquotations term_type and typeof
 
haftmann 
parents: 
32890 
diff
changeset
 | 
500  | 
let val t' = style t  | 
| 39288 | 501  | 
in pretty_term ctxt (Type.constraint (Term.fastype_of t') t') end;  | 
| 
32898
 
e871d897969c
term styles also cover antiquotations term_type and typeof
 
haftmann 
parents: 
32890 
diff
changeset
 | 
502  | 
|
| 
 
e871d897969c
term styles also cover antiquotations term_type and typeof
 
haftmann 
parents: 
32890 
diff
changeset
 | 
503  | 
fun pretty_term_typeof ctxt (style, t) =  | 
| 
 
e871d897969c
term styles also cover antiquotations term_type and typeof
 
haftmann 
parents: 
32890 
diff
changeset
 | 
504  | 
Syntax.pretty_typ ctxt (Term.fastype_of (style t));  | 
| 22124 | 505  | 
|
| 
25373
 
ccbf65080fdf
@{const}: improved ProofContext.read_const does the job;
 
wenzelm 
parents: 
25241 
diff
changeset
 | 
506  | 
fun pretty_const ctxt c =  | 
| 
 
ccbf65080fdf
@{const}: improved ProofContext.read_const does the job;
 
wenzelm 
parents: 
25241 
diff
changeset
 | 
507  | 
let  | 
| 42360 | 508  | 
val t = Const (c, Consts.type_scheme (Proof_Context.consts_of ctxt) c)  | 
| 
25373
 
ccbf65080fdf
@{const}: improved ProofContext.read_const does the job;
 
wenzelm 
parents: 
25241 
diff
changeset
 | 
509  | 
handle TYPE (msg, _, _) => error msg;  | 
| 
 
ccbf65080fdf
@{const}: improved ProofContext.read_const does the job;
 
wenzelm 
parents: 
25241 
diff
changeset
 | 
510  | 
val ([t'], _) = Variable.import_terms true [t] ctxt;  | 
| 
 
ccbf65080fdf
@{const}: improved ProofContext.read_const does the job;
 
wenzelm 
parents: 
25241 
diff
changeset
 | 
511  | 
in pretty_term ctxt t' end;  | 
| 22124 | 512  | 
|
| 
25407
 
2859cf34aaf0
abbrev: bypass full term check via ProofContext.standard_infer_types (prevents forced expansion);
 
wenzelm 
parents: 
25373 
diff
changeset
 | 
513  | 
fun pretty_abbrev ctxt s =  | 
| 22124 | 514  | 
let  | 
| 42399 | 515  | 
val t = Syntax.read_term (Proof_Context.set_mode Proof_Context.mode_abbrev ctxt) s;  | 
| 24920 | 516  | 
    fun err () = error ("Abbreviated constant expected: " ^ Syntax.string_of_term ctxt t);
 | 
| 22124 | 517  | 
val (head, args) = Term.strip_comb t;  | 
518  | 
val (c, T) = Term.dest_Const head handle TERM _ => err ();  | 
|
| 42360 | 519  | 
val (U, u) = Consts.the_abbreviation (Proof_Context.consts_of ctxt) c  | 
| 22124 | 520  | 
handle TYPE _ => err ();  | 
521  | 
val t' = Term.betapplys (Envir.expand_atom T (U, u), args);  | 
|
| 
32898
 
e871d897969c
term styles also cover antiquotations term_type and typeof
 
haftmann 
parents: 
32890 
diff
changeset
 | 
522  | 
val eq = Logic.mk_equals (t, t');  | 
| 
 
e871d897969c
term styles also cover antiquotations term_type and typeof
 
haftmann 
parents: 
32890 
diff
changeset
 | 
523  | 
val ctxt' = Variable.auto_fixes eq ctxt;  | 
| 42360 | 524  | 
in Proof_Context.pretty_term_abbrev ctxt' eq end;  | 
| 22124 | 525  | 
|
| 
39305
 
d4fa19eb0822
'class' and 'type' are now antiquoations by default
 
haftmann 
parents: 
39134 
diff
changeset
 | 
526  | 
fun pretty_class ctxt =  | 
| 42360 | 527  | 
Pretty.str o Proof_Context.extern_class ctxt o Proof_Context.read_class ctxt;  | 
| 
39305
 
d4fa19eb0822
'class' and 'type' are now antiquoations by default
 
haftmann 
parents: 
39134 
diff
changeset
 | 
528  | 
|
| 
39309
 
74469faa27ca
type antiquotation: allow arbitrary type abbreviations, but fail with user-space exception on bad input
 
haftmann 
parents: 
39308 
diff
changeset
 | 
529  | 
fun pretty_type ctxt s =  | 
| 42360 | 530  | 
let val Type (name, _) = Proof_Context.read_type_name_proper ctxt false s  | 
531  | 
in Pretty.str (Proof_Context.extern_type ctxt name) end;  | 
|
| 
39305
 
d4fa19eb0822
'class' and 'type' are now antiquoations by default
 
haftmann 
parents: 
39134 
diff
changeset
 | 
532  | 
|
| 33388 | 533  | 
fun pretty_prf full ctxt = Proof_Syntax.pretty_proof_of ctxt full;  | 
| 22124 | 534  | 
|
535  | 
fun pretty_theory ctxt name =  | 
|
| 42360 | 536  | 
(Theory.requires (Proof_Context.theory_of ctxt) name "presentation"; Pretty.str name);  | 
| 22124 | 537  | 
|
538  | 
||
| 
30390
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
539  | 
(* default output *)  | 
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
540  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
541  | 
val str_of_source = space_implode " " o map Token.unparse o #2 o #1 o Args.dest_src;  | 
| 22124 | 542  | 
|
| 
38767
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
543  | 
fun maybe_pretty_source pretty ctxt src xs =  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
544  | 
map (pretty ctxt) xs (*always pretty in order to exhibit errors!*)  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
545  | 
|> (if Config.get ctxt source then K [pretty_text ctxt (str_of_source src)] else I);  | 
| 
30390
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
546  | 
|
| 
38767
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
547  | 
fun output ctxt prts =  | 
| 
30390
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
548  | 
prts  | 
| 
38767
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
549  | 
|> (if Config.get ctxt quotes then map Pretty.quote else I)  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
550  | 
|> (if Config.get ctxt display then  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
551  | 
map (Output.output o Pretty.string_of o Pretty.indent (Config.get ctxt indent))  | 
| 22124 | 552  | 
#> space_implode "\\isasep\\isanewline%\n"  | 
553  | 
    #> enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
 | 
|
554  | 
else  | 
|
| 
38767
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
555  | 
map (Output.output o (if Config.get ctxt break then Pretty.string_of else Pretty.str_of))  | 
| 22124 | 556  | 
#> space_implode "\\isasep\\isanewline%\n"  | 
557  | 
    #> enclose "\\isa{" "}");
 | 
|
558  | 
||
| 
30390
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
559  | 
|
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
560  | 
|
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
561  | 
(** concrete antiquotations **)  | 
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
562  | 
|
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
563  | 
(* basic entities *)  | 
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
564  | 
|
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
565  | 
local  | 
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
566  | 
|
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
567  | 
fun basic_entities name scan pretty = antiquotation name scan  | 
| 
38767
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
568  | 
  (fn {source, context, ...} => output context o maybe_pretty_source pretty context source);
 | 
| 
30390
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
569  | 
|
| 
32890
 
77df12652210
generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
 
haftmann 
parents: 
32738 
diff
changeset
 | 
570  | 
fun basic_entities_style name scan pretty = antiquotation name scan  | 
| 
 
77df12652210
generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
 
haftmann 
parents: 
32738 
diff
changeset
 | 
571  | 
  (fn {source, context, ...} => fn (style, xs) =>
 | 
| 
38767
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
572  | 
output context  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
573  | 
(maybe_pretty_source (fn ctxt => fn x => pretty ctxt (style, x)) context source xs));  | 
| 
32890
 
77df12652210
generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
 
haftmann 
parents: 
32738 
diff
changeset
 | 
574  | 
|
| 
30390
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
575  | 
fun basic_entity name scan = basic_entities name (scan >> single);  | 
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
576  | 
|
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
577  | 
in  | 
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
578  | 
|
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
579  | 
val _ =  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
580  | 
Context.>> (Context.map_theory  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
581  | 
(basic_entities_style (Binding.name "thm") (Term_Style.parse -- Attrib.thms) pretty_thm_style #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
582  | 
basic_entity (Binding.name "prop") (Term_Style.parse -- Args.prop) pretty_term_style #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
583  | 
basic_entity (Binding.name "term") (Term_Style.parse -- Args.term) pretty_term_style #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
584  | 
basic_entity (Binding.name "term_type") (Term_Style.parse -- Args.term) pretty_term_typ #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
585  | 
basic_entity (Binding.name "typeof") (Term_Style.parse -- Args.term) pretty_term_typeof #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
586  | 
basic_entity (Binding.name "const") (Args.const_proper false) pretty_const #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
587  | 
basic_entity (Binding.name "abbrev") (Scan.lift Args.name_source) pretty_abbrev #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
588  | 
basic_entity (Binding.name "typ") Args.typ_abbrev Syntax.pretty_typ #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
589  | 
basic_entity (Binding.name "class") (Scan.lift Args.name) pretty_class #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
590  | 
basic_entity (Binding.name "type") (Scan.lift Args.name) pretty_type #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
591  | 
basic_entity (Binding.name "text") (Scan.lift Args.name) pretty_text #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
592  | 
basic_entities (Binding.name "prf") Attrib.thms (pretty_prf false) #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
593  | 
basic_entities (Binding.name "full_prf") Attrib.thms (pretty_prf true) #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
594  | 
basic_entity (Binding.name "theory") (Scan.lift Args.name) pretty_theory #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
595  | 
basic_entities_style (Binding.name "thm_style")  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
596  | 
(Term_Style.parse_bare -- Attrib.thms) pretty_thm_style #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
597  | 
basic_entity (Binding.name "term_style")  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
598  | 
(Term_Style.parse_bare -- Args.term) pretty_term_style));  | 
| 
32898
 
e871d897969c
term styles also cover antiquotations term_type and typeof
 
haftmann 
parents: 
32890 
diff
changeset
 | 
599  | 
|
| 
30390
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
600  | 
end;  | 
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
601  | 
|
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
602  | 
|
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
603  | 
(* goal state *)  | 
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
604  | 
|
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
605  | 
local  | 
| 22124 | 606  | 
|
| 30367 | 607  | 
fun proof_state state =  | 
608  | 
(case try Toplevel.proof_of state of  | 
|
609  | 
SOME prf => prf  | 
|
| 22124 | 610  | 
| _ => error "No proof state");  | 
611  | 
||
| 
30390
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
612  | 
fun goal_state name main_goal = antiquotation name (Scan.succeed ())  | 
| 
39125
 
f45d332a90e3
pretty_goals: turned some global references and function arguments into configuration options (goals_limit = 10, goals_total = true, show_main_goal = false) depending on the context;
 
wenzelm 
parents: 
38980 
diff
changeset
 | 
613  | 
  (fn {state, context = ctxt, ...} => fn () => output ctxt
 | 
| 
39129
 
976af4e27cde
recovered options for goal antiquotations from f45d332a90e3: actually pass context to Proof.pretty_goals (see also 45facd8f358e);
 
wenzelm 
parents: 
39128 
diff
changeset
 | 
614  | 
[Pretty.chunks  | 
| 
 
976af4e27cde
recovered options for goal antiquotations from f45d332a90e3: actually pass context to Proof.pretty_goals (see also 45facd8f358e);
 
wenzelm 
parents: 
39128 
diff
changeset
 | 
615  | 
(Proof.pretty_goals main_goal (Proof.map_context (K ctxt) (proof_state state)))]);  | 
| 22124 | 616  | 
|
| 
30390
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
617  | 
in  | 
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
618  | 
|
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
619  | 
val _ =  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
620  | 
Context.>> (Context.map_theory  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
621  | 
(goal_state (Binding.name "goals") true #>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
622  | 
goal_state (Binding.name "subgoals") false));  | 
| 
30390
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
623  | 
|
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
624  | 
end;  | 
| 22124 | 625  | 
|
626  | 
||
| 
30390
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
627  | 
(* embedded lemma *)  | 
| 27522 | 628  | 
|
| 36950 | 629  | 
val _ = Keyword.keyword "by";  | 
| 27381 | 630  | 
|
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
631  | 
val _ =  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
632  | 
Context.>> (Context.map_theory  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
633  | 
(antiquotation (Binding.name "lemma")  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
634  | 
(Args.prop -- Scan.lift (Args.$$$ "by" |-- Method.parse -- Scan.option Method.parse))  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
635  | 
    (fn {source, context, ...} => fn (prop, methods) =>
 | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
636  | 
let  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
637  | 
val prop_src =  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
638  | 
(case Args.dest_src source of ((a, arg :: _), pos) => Args.src ((a, [arg]), pos));  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
639  | 
(* FIXME check proof!? *)  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
640  | 
val _ = context  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
641  | 
|> Proof.theorem NONE (K I) [[(prop, [])]]  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
642  | 
|> Proof.global_terminal_proof methods;  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
643  | 
in output context (maybe_pretty_source pretty_term context prop_src [prop]) end)));  | 
| 
30390
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
644  | 
|
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
645  | 
|
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
646  | 
(* ML text *)  | 
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
647  | 
|
| 40801 | 648  | 
val verb_text =  | 
649  | 
split_lines  | 
|
650  | 
#> map (space_implode "\\verb,|," o map (enclose "\\verb|" "|") o space_explode "|")  | 
|
651  | 
#> space_implode "\\isasep\\isanewline%\n";  | 
|
652  | 
||
| 
30390
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
653  | 
local  | 
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
654  | 
|
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
655  | 
fun ml_text name ml = antiquotation name (Scan.lift Args.name_source_position)  | 
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
656  | 
  (fn {context, ...} => fn (txt, pos) =>
 | 
| 
37198
 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 
wenzelm 
parents: 
37146 
diff
changeset
 | 
657  | 
(ML_Context.eval_in (SOME context) false pos (ml pos txt);  | 
| 30573 | 658  | 
Symbol_Pos.content (Symbol_Pos.explode (txt, pos))  | 
| 
38767
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
659  | 
|> (if Config.get context quotes then quote else I)  | 
| 
 
d8da44a8dd25
proper context for various Thy_Output options, via official configuration options in ML and Isar;
 
wenzelm 
parents: 
38766 
diff
changeset
 | 
660  | 
    |> (if Config.get context display then enclose "\\begin{verbatim}\n" "\n\\end{verbatim}"
 | 
| 40801 | 661  | 
else verb_text)));  | 
| 
30390
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
662  | 
|
| 
37198
 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 
wenzelm 
parents: 
37146 
diff
changeset
 | 
663  | 
fun ml_enclose bg en pos txt =  | 
| 
 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 
wenzelm 
parents: 
37146 
diff
changeset
 | 
664  | 
ML_Lex.read Position.none bg @ ML_Lex.read pos txt @ ML_Lex.read Position.none en;  | 
| 
 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 
wenzelm 
parents: 
37146 
diff
changeset
 | 
665  | 
|
| 
30390
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
666  | 
in  | 
| 
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
667  | 
|
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
668  | 
val _ =  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
669  | 
Context.>> (Context.map_theory  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
670  | 
   (ml_text (Binding.name "ML") (ml_enclose "fn _ => (" ");") #>
 | 
| 46261 | 671  | 
ml_text (Binding.name "ML_op") (ml_enclose "fn _ => (op " ");") #>  | 
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
672  | 
    ml_text (Binding.name "ML_type") (ml_enclose "val _ = NONE : (" ") option;") #>
 | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
673  | 
ml_text (Binding.name "ML_struct")  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
674  | 
(ml_enclose "functor XXX() = struct structure XX = " " end;") #>  | 
| 
37198
 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 
wenzelm 
parents: 
37146 
diff
changeset
 | 
675  | 
|
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
676  | 
ml_text (Binding.name "ML_functor") (* FIXME formal treatment of functor name (!?) *)  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
677  | 
(fn pos => fn txt =>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
678  | 
        ML_Lex.read Position.none ("ML_Env.check_functor " ^
 | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
679  | 
ML_Syntax.print_string (Symbol_Pos.content (Symbol_Pos.explode (txt, pos))))) #>  | 
| 
37198
 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 
wenzelm 
parents: 
37146 
diff
changeset
 | 
680  | 
|
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
681  | 
ml_text (Binding.name "ML_text") (K (K []))));  | 
| 22124 | 682  | 
|
683  | 
end;  | 
|
| 
30390
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
684  | 
|
| 40801 | 685  | 
|
686  | 
(* files *)  | 
|
687  | 
||
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
688  | 
val _ =  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
689  | 
Context.>> (Context.map_theory  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
690  | 
(antiquotation (Binding.name "file") (Scan.lift Args.name)  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
691  | 
(fn _ => fn path =>  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
692  | 
if File.exists (Path.explode path) then verb_text path  | 
| 
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42669 
diff
changeset
 | 
693  | 
      else error ("Bad file: " ^ quote path))));
 | 
| 40801 | 694  | 
|
| 
30390
 
ad591ee76c78
simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
 
wenzelm 
parents: 
30381 
diff
changeset
 | 
695  | 
end;  |