| author | wenzelm | 
| Sun, 16 Jul 2023 19:13:08 +0200 | |
| changeset 78369 | ba71ea02d965 | 
| parent 76069 | 79094d7b6f22 | 
| child 80841 | 1beb2dc3bf14 | 
| permissions | -rw-r--r-- | 
| 67386 | 1  | 
(* Title: Pure/Thy/document_antiquotation.ML  | 
2  | 
Author: Makarius  | 
|
3  | 
||
4  | 
Document antiquotations.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
signature DOCUMENT_ANTIQUOTATION =  | 
|
8  | 
sig  | 
|
9  | 
val thy_output_display: bool Config.T  | 
|
| 70121 | 10  | 
val thy_output_cartouche: bool Config.T  | 
| 67386 | 11  | 
val thy_output_quotes: bool Config.T  | 
12  | 
val thy_output_margin: int Config.T  | 
|
13  | 
val thy_output_indent: int Config.T  | 
|
14  | 
val thy_output_source: bool Config.T  | 
|
| 
70122
 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 
wenzelm 
parents: 
70121 
diff
changeset
 | 
15  | 
val thy_output_source_cartouche: bool Config.T  | 
| 67386 | 16  | 
val thy_output_break: bool Config.T  | 
17  | 
val thy_output_modes: string Config.T  | 
|
| 67463 | 18  | 
val trim_blanks: Proof.context -> string -> string  | 
19  | 
val trim_lines: Proof.context -> string -> string  | 
|
20  | 
val indent_lines: Proof.context -> string -> string  | 
|
| 
67474
 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 
wenzelm 
parents: 
67472 
diff
changeset
 | 
21  | 
val prepare_lines: Proof.context -> string -> string  | 
| 70121 | 22  | 
val delimit: Proof.context -> Pretty.T -> Pretty.T  | 
| 67463 | 23  | 
val indent: Proof.context -> Pretty.T -> Pretty.T  | 
24  | 
val format: Proof.context -> Pretty.T -> string  | 
|
25  | 
val output: Proof.context -> Pretty.T -> Output.output  | 
|
| 67386 | 26  | 
val print_antiquotations: bool -> Proof.context -> unit  | 
27  | 
val check: Proof.context -> xstring * Position.T -> string  | 
|
28  | 
val check_option: Proof.context -> xstring * Position.T -> string  | 
|
29  | 
val setup: binding -> 'a context_parser ->  | 
|
| 67463 | 30  | 
    ({context: Proof.context, source: Token.src, argument: 'a} -> Latex.text) -> theory -> theory
 | 
| 
69592
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
31  | 
val setup_embedded: binding -> 'a context_parser ->  | 
| 
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
32  | 
    ({context: Proof.context, source: Token.src, argument: 'a} -> Latex.text) -> theory -> theory
 | 
| 67386 | 33  | 
val setup_option: binding -> (string -> Proof.context -> Proof.context) -> theory -> theory  | 
34  | 
val boolean: string -> bool  | 
|
35  | 
val integer: string -> int  | 
|
| 
73780
 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 
wenzelm 
parents: 
73767 
diff
changeset
 | 
36  | 
val evaluate: (Symbol_Pos.T list -> Latex.text) -> Proof.context ->  | 
| 
 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 
wenzelm 
parents: 
73767 
diff
changeset
 | 
37  | 
Antiquote.text_antiquote -> Latex.text  | 
| 73767 | 38  | 
val approx_content: Proof.context -> string -> string  | 
| 67386 | 39  | 
end;  | 
40  | 
||
41  | 
structure Document_Antiquotation: DOCUMENT_ANTIQUOTATION =  | 
|
42  | 
struct  | 
|
43  | 
||
44  | 
(* options *)  | 
|
45  | 
||
| 76069 | 46  | 
val thy_output_display = Attrib.setup_option_bool (\<^system_option>\<open>thy_output_display\<close>, \<^here>);  | 
47  | 
val thy_output_break = Attrib.setup_option_bool (\<^system_option>\<open>thy_output_break\<close>, \<^here>);  | 
|
48  | 
val thy_output_cartouche = Attrib.setup_option_bool (\<^system_option>\<open>thy_output_cartouche\<close>, \<^here>);  | 
|
49  | 
val thy_output_quotes = Attrib.setup_option_bool (\<^system_option>\<open>thy_output_quotes\<close>, \<^here>);  | 
|
50  | 
val thy_output_margin = Attrib.setup_option_int (\<^system_option>\<open>thy_output_margin\<close>, \<^here>);  | 
|
51  | 
val thy_output_indent = Attrib.setup_option_int (\<^system_option>\<open>thy_output_indent\<close>, \<^here>);  | 
|
52  | 
val thy_output_source = Attrib.setup_option_bool (\<^system_option>\<open>thy_output_source\<close>, \<^here>);  | 
|
53  | 
val thy_output_source_cartouche = Attrib.setup_option_bool (\<^system_option>\<open>thy_output_source_cartouche\<close>, \<^here>);  | 
|
54  | 
val thy_output_modes = Attrib.setup_option_string (\<^system_option>\<open>thy_output_modes\<close>, \<^here>);  | 
|
| 67386 | 55  | 
|
56  | 
||
| 67463 | 57  | 
(* blanks *)  | 
58  | 
||
59  | 
fun trim_blanks ctxt =  | 
|
60  | 
not (Config.get ctxt thy_output_display) ? Symbol.trim_blanks;  | 
|
61  | 
||
62  | 
fun trim_lines ctxt =  | 
|
63  | 
if not (Config.get ctxt thy_output_display) then  | 
|
64  | 
split_lines #> map Symbol.trim_blanks #> space_implode " "  | 
|
65  | 
else I;  | 
|
66  | 
||
67  | 
fun indent_lines ctxt =  | 
|
68  | 
if Config.get ctxt thy_output_display then  | 
|
69  | 
prefix_lines (Symbol.spaces (Config.get ctxt thy_output_indent))  | 
|
70  | 
else I;  | 
|
71  | 
||
| 
67474
 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 
wenzelm 
parents: 
67472 
diff
changeset
 | 
72  | 
fun prepare_lines ctxt = trim_lines ctxt #> indent_lines ctxt;  | 
| 
 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 
wenzelm 
parents: 
67472 
diff
changeset
 | 
73  | 
|
| 67463 | 74  | 
|
75  | 
(* output *)  | 
|
76  | 
||
| 70121 | 77  | 
fun delimit ctxt =  | 
78  | 
if Config.get ctxt thy_output_cartouche then Pretty.cartouche  | 
|
79  | 
else if Config.get ctxt thy_output_quotes then Pretty.quote  | 
|
80  | 
else I;  | 
|
| 67463 | 81  | 
|
82  | 
fun indent ctxt =  | 
|
83  | 
Config.get ctxt thy_output_display ? Pretty.indent (Config.get ctxt thy_output_indent);  | 
|
84  | 
||
85  | 
fun format ctxt =  | 
|
86  | 
if Config.get ctxt thy_output_display orelse Config.get ctxt thy_output_break  | 
|
87  | 
then Pretty.string_of_margin (Config.get ctxt thy_output_margin)  | 
|
88  | 
else Pretty.unformatted_string_of;  | 
|
89  | 
||
| 70121 | 90  | 
fun output ctxt = delimit ctxt #> indent ctxt #> format ctxt #> Output.output;  | 
| 67463 | 91  | 
|
92  | 
||
| 67386 | 93  | 
(* theory data *)  | 
94  | 
||
95  | 
structure Data = Theory_Data  | 
|
96  | 
(  | 
|
97  | 
type T =  | 
|
| 
69592
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
98  | 
(bool * (Token.src -> Proof.context -> Latex.text)) Name_Space.table *  | 
| 67386 | 99  | 
(string -> Proof.context -> Proof.context) Name_Space.table;  | 
100  | 
val empty : T =  | 
|
101  | 
(Name_Space.empty_table Markup.document_antiquotationN,  | 
|
102  | 
Name_Space.empty_table Markup.document_antiquotation_optionN);  | 
|
103  | 
fun merge ((commands1, options1), (commands2, options2)) : T =  | 
|
104  | 
(Name_Space.merge_tables (commands1, commands2),  | 
|
105  | 
Name_Space.merge_tables (options1, options2));  | 
|
106  | 
);  | 
|
107  | 
||
108  | 
val get_antiquotations = Data.get o Proof_Context.theory_of;  | 
|
109  | 
||
110  | 
fun print_antiquotations verbose ctxt =  | 
|
111  | 
let  | 
|
112  | 
val (commands, options) = get_antiquotations ctxt;  | 
|
113  | 
val command_names = map #1 (Name_Space.markup_table verbose ctxt commands);  | 
|
114  | 
val option_names = map #1 (Name_Space.markup_table verbose ctxt options);  | 
|
115  | 
in  | 
|
116  | 
[Pretty.big_list "document antiquotations:" (map Pretty.mark_str command_names),  | 
|
117  | 
Pretty.big_list "document antiquotation options:" (map Pretty.mark_str option_names)]  | 
|
118  | 
end |> Pretty.writeln_chunks;  | 
|
119  | 
||
120  | 
fun check ctxt = #1 o Name_Space.check (Context.Proof ctxt) (#1 (get_antiquotations ctxt));  | 
|
121  | 
fun check_option ctxt = #1 o Name_Space.check (Context.Proof ctxt) (#2 (get_antiquotations ctxt));  | 
|
122  | 
||
| 
69592
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
123  | 
fun gen_setup name embedded scan body thy =  | 
| 67386 | 124  | 
let  | 
125  | 
fun cmd src ctxt =  | 
|
126  | 
let val (x, ctxt') = Token.syntax scan src ctxt  | 
|
| 67463 | 127  | 
      in body {context = ctxt', source = src, argument = x} end;
 | 
| 
69592
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
128  | 
val entry = (name, (embedded, cmd));  | 
| 
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
129  | 
in thy |> Data.map (apfst (Name_Space.define (Context.Theory thy) true entry #> #2)) end;  | 
| 
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
130  | 
|
| 
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
131  | 
fun setup name = gen_setup name false;  | 
| 
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
132  | 
fun setup_embedded name = gen_setup name true;  | 
| 67386 | 133  | 
|
134  | 
fun setup_option name opt thy = thy  | 
|
135  | 
|> Data.map (apsnd (Name_Space.define (Context.Theory thy) true (name, opt) #> #2));  | 
|
136  | 
||
137  | 
||
138  | 
(* syntax *)  | 
|
139  | 
||
140  | 
local  | 
|
141  | 
||
142  | 
val property =  | 
|
| 
69349
 
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
 
wenzelm 
parents: 
68223 
diff
changeset
 | 
143  | 
Parse.name_position -- Scan.optional (Parse.$$$ "=" |-- Parse.!!! Parse.name) "";  | 
| 67386 | 144  | 
|
145  | 
val properties =  | 
|
146  | 
Scan.optional (Parse.$$$ "[" |-- Parse.!!! (Parse.enum "," property --| Parse.$$$ "]")) [];  | 
|
147  | 
||
148  | 
in  | 
|
149  | 
||
| 67466 | 150  | 
val parse_antiq =  | 
| 67386 | 151  | 
Parse.!!!  | 
152  | 
(Parse.token Parse.liberal_name -- properties -- Parse.args --| Scan.ahead Parse.eof)  | 
|
| 
69592
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
153  | 
>> (fn ((name, opts), args) => (opts, name :: args));  | 
| 67386 | 154  | 
|
155  | 
end;  | 
|
156  | 
||
157  | 
||
158  | 
(* evaluate *)  | 
|
159  | 
||
160  | 
local  | 
|
161  | 
||
| 
69592
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
162  | 
fun command pos (opts, src) ctxt =  | 
| 
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
163  | 
let  | 
| 
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
164  | 
val (src', (embedded, scan)) = Token.check_src ctxt (#1 o get_antiquotations) src;  | 
| 
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
165  | 
val _ =  | 
| 
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
166  | 
if null opts andalso Options.default_bool "update_control_cartouches" then  | 
| 
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
167  | 
Context_Position.reports_text ctxt  | 
| 
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
168  | 
(Antiquote.update_reports embedded pos (map Token.content_of src'))  | 
| 
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
169  | 
else ();  | 
| 
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
170  | 
in scan src' ctxt end;  | 
| 67386 | 171  | 
|
172  | 
fun option ((xname, pos), s) ctxt =  | 
|
173  | 
let  | 
|
174  | 
val (_, opt) =  | 
|
175  | 
Name_Space.check (Context.Proof ctxt) (#2 (get_antiquotations ctxt)) (xname, pos);  | 
|
176  | 
in opt s ctxt end;  | 
|
177  | 
||
| 
69592
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
178  | 
fun eval ctxt pos (opts, src) =  | 
| 67386 | 179  | 
let  | 
| 
68223
 
88dd06301dd3
override default of Isabelle_Process, notably for PIDE export of "document.tex";
 
wenzelm 
parents: 
67571 
diff
changeset
 | 
180  | 
val preview_ctxt = fold option opts (Config.put show_markup false ctxt);  | 
| 
69592
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
181  | 
val _ = command pos (opts, src) preview_ctxt;  | 
| 
67472
 
7ed8d4cdfb13
recovered antiquotation check without latex mode (cf. dfc93f2b01ea);
 
wenzelm 
parents: 
67466 
diff
changeset
 | 
182  | 
|
| 67386 | 183  | 
val print_ctxt = Context_Position.set_visible false preview_ctxt;  | 
| 
72075
 
9c0b835d4cc2
proper pretty printing for latex output, notably for pide_session=true (default);
 
wenzelm 
parents: 
70122 
diff
changeset
 | 
184  | 
val print_modes =  | 
| 
 
9c0b835d4cc2
proper pretty printing for latex output, notably for pide_session=true (default);
 
wenzelm 
parents: 
70122 
diff
changeset
 | 
185  | 
space_explode "," (Config.get print_ctxt thy_output_modes) @ [Latex.latexN, Pretty.regularN];  | 
| 
73780
 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 
wenzelm 
parents: 
73767 
diff
changeset
 | 
186  | 
in Print_Mode.with_modes print_modes (fn () => command pos (opts, src) print_ctxt) () end;  | 
| 67386 | 187  | 
|
188  | 
in  | 
|
189  | 
||
| 73756 | 190  | 
fun read_antiq ctxt ({range = (pos, _), body, ...}: Antiquote.antiq) =
 | 
| 74564 | 191  | 
Parse.read_antiq (Thy_Header.get_keywords' ctxt) parse_antiq (body, pos);  | 
| 73756 | 192  | 
|
| 
67571
 
f858fe5531ac
more uniform treatment of formal comments within document source;
 
wenzelm 
parents: 
67474 
diff
changeset
 | 
193  | 
fun evaluate eval_text ctxt antiq =  | 
| 67466 | 194  | 
(case antiq of  | 
| 
67571
 
f858fe5531ac
more uniform treatment of formal comments within document source;
 
wenzelm 
parents: 
67474 
diff
changeset
 | 
195  | 
Antiquote.Text ss => eval_text ss  | 
| 67466 | 196  | 
  | Antiquote.Control {name, body, ...} =>
 | 
| 
69592
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
197  | 
eval ctxt Position.none  | 
| 
 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 
wenzelm 
parents: 
69349 
diff
changeset
 | 
198  | 
([], Token.make_src name (if null body then [] else [Token.read_cartouche body]))  | 
| 73756 | 199  | 
| Antiquote.Antiq antiq => eval ctxt (#1 (#range antiq)) (read_antiq ctxt antiq));  | 
| 67386 | 200  | 
|
201  | 
end;  | 
|
202  | 
||
203  | 
||
| 73762 | 204  | 
(* approximative content, e.g. for index entries *)  | 
205  | 
||
206  | 
local  | 
|
207  | 
||
208  | 
val strip_symbols = [Symbol.open_, Symbol.close, "'", "\"", "`"];  | 
|
209  | 
||
210  | 
fun strip_symbol sym =  | 
|
211  | 
if member (op =) strip_symbols sym then ""  | 
|
212  | 
else  | 
|
213  | 
(case Symbol.decode sym of  | 
|
214  | 
Symbol.Char s => s  | 
|
215  | 
| Symbol.UTF8 s => s  | 
|
216  | 
| Symbol.Sym s => s  | 
|
217  | 
| Symbol.Control s => s  | 
|
218  | 
| _ => "");  | 
|
219  | 
||
220  | 
fun strip_antiq _ (Antiquote.Text body) = map Symbol_Pos.symbol body  | 
|
221  | 
  | strip_antiq _ (Antiquote.Control {body, ...}) = map Symbol_Pos.symbol body
 | 
|
222  | 
| strip_antiq ctxt (Antiquote.Antiq antiq) =  | 
|
223  | 
read_antiq ctxt antiq |> #2 |> tl  | 
|
224  | 
|> maps (Symbol.explode o Token.content_of);  | 
|
225  | 
||
226  | 
in  | 
|
227  | 
||
| 73767 | 228  | 
fun approx_content ctxt =  | 
229  | 
Symbol_Pos.explode0  | 
|
230  | 
#> trim (Symbol.is_blank o Symbol_Pos.symbol)  | 
|
231  | 
#> Antiquote.parse_comments Position.none  | 
|
232  | 
#> maps (strip_antiq ctxt)  | 
|
233  | 
#> map strip_symbol  | 
|
234  | 
#> implode;  | 
|
| 73762 | 235  | 
|
236  | 
end;  | 
|
237  | 
||
238  | 
||
| 67386 | 239  | 
(* option syntax *)  | 
240  | 
||
241  | 
fun boolean "" = true  | 
|
242  | 
| boolean "true" = true  | 
|
243  | 
| boolean "false" = false  | 
|
| 67463 | 244  | 
  | boolean s = error ("Bad boolean value: " ^ Library.quote s);
 | 
| 67386 | 245  | 
|
246  | 
fun integer s =  | 
|
247  | 
let  | 
|
248  | 
fun int ss =  | 
|
249  | 
(case Library.read_int ss of (i, []) => i  | 
|
| 67463 | 250  | 
      | _ => error ("Bad integer value: " ^ Library.quote s));
 | 
| 67386 | 251  | 
in (case Symbol.explode s of "-" :: ss => ~ (int ss) | ss => int ss) end;  | 
252  | 
||
253  | 
val _ = Theory.setup  | 
|
254  | 
(setup_option \<^binding>\<open>show_types\<close> (Config.put show_types o boolean) #>  | 
|
255  | 
setup_option \<^binding>\<open>show_sorts\<close> (Config.put show_sorts o boolean) #>  | 
|
256  | 
setup_option \<^binding>\<open>show_structs\<close> (Config.put show_structs o boolean) #>  | 
|
257  | 
setup_option \<^binding>\<open>show_question_marks\<close> (Config.put show_question_marks o boolean) #>  | 
|
258  | 
setup_option \<^binding>\<open>show_abbrevs\<close> (Config.put show_abbrevs o boolean) #>  | 
|
259  | 
setup_option \<^binding>\<open>names_long\<close> (Config.put Name_Space.names_long o boolean) #>  | 
|
260  | 
setup_option \<^binding>\<open>names_short\<close> (Config.put Name_Space.names_short o boolean) #>  | 
|
261  | 
setup_option \<^binding>\<open>names_unique\<close> (Config.put Name_Space.names_unique o boolean) #>  | 
|
262  | 
setup_option \<^binding>\<open>eta_contract\<close> (Config.put Syntax_Trans.eta_contract o boolean) #>  | 
|
263  | 
setup_option \<^binding>\<open>display\<close> (Config.put thy_output_display o boolean) #>  | 
|
264  | 
setup_option \<^binding>\<open>break\<close> (Config.put thy_output_break o boolean) #>  | 
|
| 70121 | 265  | 
setup_option \<^binding>\<open>cartouche\<close> (Config.put thy_output_cartouche o boolean) #>  | 
| 67386 | 266  | 
setup_option \<^binding>\<open>quotes\<close> (Config.put thy_output_quotes o boolean) #>  | 
| 
67460
 
dfc93f2b01ea
discontinued unused wrapper: print_mode is provided directly;
 
wenzelm 
parents: 
67386 
diff
changeset
 | 
267  | 
setup_option \<^binding>\<open>mode\<close> (Config.put thy_output_modes) #>  | 
| 67386 | 268  | 
setup_option \<^binding>\<open>margin\<close> (Config.put thy_output_margin o integer) #>  | 
269  | 
setup_option \<^binding>\<open>indent\<close> (Config.put thy_output_indent o integer) #>  | 
|
270  | 
setup_option \<^binding>\<open>source\<close> (Config.put thy_output_source o boolean) #>  | 
|
| 
70122
 
a0b21b4b7a4a
strip cartouches from arguments of "embedded" document antiquotations, corresponding to automated update via "isabelle update -u control_cartouches" -- e.g. relevant for documents with thy_output_source (e.g. doc "isar-ref", "jedit", "system");
 
wenzelm 
parents: 
70121 
diff
changeset
 | 
271  | 
setup_option \<^binding>\<open>source_cartouche\<close> (Config.put thy_output_source_cartouche o boolean) #>  | 
| 67386 | 272  | 
setup_option \<^binding>\<open>goals_limit\<close> (Config.put Goal_Display.goals_limit o integer));  | 
273  | 
||
274  | 
end;  |