| author | wenzelm | 
| Sat, 08 Dec 2018 21:13:47 +0100 | |
| changeset 69427 | ff2f39a221d4 | 
| parent 69349 | 7cef9e386ffe | 
| child 69592 | a80d8ec6c998 | 
| 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 | |
| 10 | val thy_output_quotes: bool Config.T | |
| 11 | val thy_output_margin: int Config.T | |
| 12 | val thy_output_indent: int Config.T | |
| 13 | val thy_output_source: bool Config.T | |
| 14 | val thy_output_break: bool Config.T | |
| 15 | val thy_output_modes: string Config.T | |
| 67463 | 16 | val trim_blanks: Proof.context -> string -> string | 
| 17 | val trim_lines: Proof.context -> string -> string | |
| 18 | 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: 
67472diff
changeset | 19 | val prepare_lines: Proof.context -> string -> string | 
| 67463 | 20 | val quote: Proof.context -> Pretty.T -> Pretty.T | 
| 21 | val indent: Proof.context -> Pretty.T -> Pretty.T | |
| 22 | val format: Proof.context -> Pretty.T -> string | |
| 23 | val output: Proof.context -> Pretty.T -> Output.output | |
| 67386 | 24 | val print_antiquotations: bool -> Proof.context -> unit | 
| 25 | val check: Proof.context -> xstring * Position.T -> string | |
| 26 | val check_option: Proof.context -> xstring * Position.T -> string | |
| 27 | val setup: binding -> 'a context_parser -> | |
| 67463 | 28 |     ({context: Proof.context, source: Token.src, argument: 'a} -> Latex.text) -> theory -> theory
 | 
| 67386 | 29 | val setup_option: binding -> (string -> Proof.context -> Proof.context) -> theory -> theory | 
| 30 | val boolean: string -> bool | |
| 31 | val integer: string -> int | |
| 67571 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67474diff
changeset | 32 | val evaluate: (Symbol_Pos.T list -> Latex.text list) -> Proof.context -> | 
| 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67474diff
changeset | 33 | Antiquote.text_antiquote -> Latex.text list | 
| 67386 | 34 | end; | 
| 35 | ||
| 36 | structure Document_Antiquotation: DOCUMENT_ANTIQUOTATION = | |
| 37 | struct | |
| 38 | ||
| 39 | (* options *) | |
| 40 | ||
| 41 | val thy_output_display = Attrib.setup_option_bool ("thy_output_display", \<^here>);
 | |
| 42 | val thy_output_break = Attrib.setup_option_bool ("thy_output_break", \<^here>);
 | |
| 43 | val thy_output_quotes = Attrib.setup_option_bool ("thy_output_quotes", \<^here>);
 | |
| 44 | val thy_output_margin = Attrib.setup_option_int ("thy_output_margin", \<^here>);
 | |
| 45 | val thy_output_indent = Attrib.setup_option_int ("thy_output_indent", \<^here>);
 | |
| 46 | val thy_output_source = Attrib.setup_option_bool ("thy_output_source", \<^here>);
 | |
| 47 | val thy_output_modes = Attrib.setup_option_string ("thy_output_modes", \<^here>);
 | |
| 48 | ||
| 49 | ||
| 67463 | 50 | (* blanks *) | 
| 51 | ||
| 52 | fun trim_blanks ctxt = | |
| 53 | not (Config.get ctxt thy_output_display) ? Symbol.trim_blanks; | |
| 54 | ||
| 55 | fun trim_lines ctxt = | |
| 56 | if not (Config.get ctxt thy_output_display) then | |
| 57 | split_lines #> map Symbol.trim_blanks #> space_implode " " | |
| 58 | else I; | |
| 59 | ||
| 60 | fun indent_lines ctxt = | |
| 61 | if Config.get ctxt thy_output_display then | |
| 62 | prefix_lines (Symbol.spaces (Config.get ctxt thy_output_indent)) | |
| 63 | else I; | |
| 64 | ||
| 67474 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67472diff
changeset | 65 | 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: 
67472diff
changeset | 66 | |
| 67463 | 67 | |
| 68 | (* output *) | |
| 69 | ||
| 70 | fun quote ctxt = | |
| 71 | Config.get ctxt thy_output_quotes ? Pretty.quote; | |
| 72 | ||
| 73 | fun indent ctxt = | |
| 74 | Config.get ctxt thy_output_display ? Pretty.indent (Config.get ctxt thy_output_indent); | |
| 75 | ||
| 76 | fun format ctxt = | |
| 77 | if Config.get ctxt thy_output_display orelse Config.get ctxt thy_output_break | |
| 78 | then Pretty.string_of_margin (Config.get ctxt thy_output_margin) | |
| 79 | else Pretty.unformatted_string_of; | |
| 80 | ||
| 81 | fun output ctxt = quote ctxt #> indent ctxt #> format ctxt #> Output.output; | |
| 82 | ||
| 83 | ||
| 67386 | 84 | (* theory data *) | 
| 85 | ||
| 86 | structure Data = Theory_Data | |
| 87 | ( | |
| 88 | type T = | |
| 67463 | 89 | (Token.src -> Proof.context -> Latex.text) Name_Space.table * | 
| 67386 | 90 | (string -> Proof.context -> Proof.context) Name_Space.table; | 
| 91 | val empty : T = | |
| 92 | (Name_Space.empty_table Markup.document_antiquotationN, | |
| 93 | Name_Space.empty_table Markup.document_antiquotation_optionN); | |
| 94 | val extend = I; | |
| 95 | fun merge ((commands1, options1), (commands2, options2)) : T = | |
| 96 | (Name_Space.merge_tables (commands1, commands2), | |
| 97 | Name_Space.merge_tables (options1, options2)); | |
| 98 | ); | |
| 99 | ||
| 100 | val get_antiquotations = Data.get o Proof_Context.theory_of; | |
| 101 | ||
| 102 | fun print_antiquotations verbose ctxt = | |
| 103 | let | |
| 104 | val (commands, options) = get_antiquotations ctxt; | |
| 105 | val command_names = map #1 (Name_Space.markup_table verbose ctxt commands); | |
| 106 | val option_names = map #1 (Name_Space.markup_table verbose ctxt options); | |
| 107 | in | |
| 108 | [Pretty.big_list "document antiquotations:" (map Pretty.mark_str command_names), | |
| 109 | Pretty.big_list "document antiquotation options:" (map Pretty.mark_str option_names)] | |
| 110 | end |> Pretty.writeln_chunks; | |
| 111 | ||
| 112 | fun check ctxt = #1 o Name_Space.check (Context.Proof ctxt) (#1 (get_antiquotations ctxt)); | |
| 113 | fun check_option ctxt = #1 o Name_Space.check (Context.Proof ctxt) (#2 (get_antiquotations ctxt)); | |
| 114 | ||
| 115 | fun setup name scan body thy = | |
| 116 | let | |
| 117 | fun cmd src ctxt = | |
| 118 | let val (x, ctxt') = Token.syntax scan src ctxt | |
| 67463 | 119 |       in body {context = ctxt', source = src, argument = x} end;
 | 
| 67386 | 120 | in thy |> Data.map (apfst (Name_Space.define (Context.Theory thy) true (name, cmd) #> #2)) end; | 
| 121 | ||
| 122 | fun setup_option name opt thy = thy | |
| 123 | |> Data.map (apsnd (Name_Space.define (Context.Theory thy) true (name, opt) #> #2)); | |
| 124 | ||
| 125 | ||
| 126 | (* syntax *) | |
| 127 | ||
| 128 | local | |
| 129 | ||
| 130 | 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: 
68223diff
changeset | 131 | Parse.name_position -- Scan.optional (Parse.$$$ "=" |-- Parse.!!! Parse.name) ""; | 
| 67386 | 132 | |
| 133 | val properties = | |
| 134 | Scan.optional (Parse.$$$ "[" |-- Parse.!!! (Parse.enum "," property --| Parse.$$$ "]")) []; | |
| 135 | ||
| 136 | in | |
| 137 | ||
| 67466 | 138 | val parse_antiq = | 
| 67386 | 139 | Parse.!!! | 
| 140 | (Parse.token Parse.liberal_name -- properties -- Parse.args --| Scan.ahead Parse.eof) | |
| 141 | >> (fn ((name, props), args) => (props, name :: args)); | |
| 142 | ||
| 143 | end; | |
| 144 | ||
| 145 | ||
| 146 | (* evaluate *) | |
| 147 | ||
| 148 | local | |
| 149 | ||
| 150 | fun command src ctxt = | |
| 151 | let val (src', f) = Token.check_src ctxt (#1 o get_antiquotations) src | |
| 152 | in f src' ctxt end; | |
| 153 | ||
| 154 | fun option ((xname, pos), s) ctxt = | |
| 155 | let | |
| 156 | val (_, opt) = | |
| 157 | Name_Space.check (Context.Proof ctxt) (#2 (get_antiquotations ctxt)) (xname, pos); | |
| 158 | in opt s ctxt end; | |
| 159 | ||
| 160 | fun eval ctxt (opts, src) = | |
| 161 | let | |
| 68223 
88dd06301dd3
override default of Isabelle_Process, notably for PIDE export of "document.tex";
 wenzelm parents: 
67571diff
changeset | 162 | val preview_ctxt = fold option opts (Config.put show_markup false ctxt); | 
| 67472 
7ed8d4cdfb13
recovered antiquotation check without latex mode (cf. dfc93f2b01ea);
 wenzelm parents: 
67466diff
changeset | 163 | val _ = command src preview_ctxt; | 
| 
7ed8d4cdfb13
recovered antiquotation check without latex mode (cf. dfc93f2b01ea);
 wenzelm parents: 
67466diff
changeset | 164 | |
| 67386 | 165 | val print_ctxt = Context_Position.set_visible false preview_ctxt; | 
| 166 | val print_modes = space_explode "," (Config.get print_ctxt thy_output_modes) @ [Latex.latexN]; | |
| 67463 | 167 | in [Print_Mode.with_modes print_modes (fn () => command src print_ctxt) ()] end; | 
| 67386 | 168 | |
| 169 | in | |
| 170 | ||
| 67571 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67474diff
changeset | 171 | fun evaluate eval_text ctxt antiq = | 
| 67466 | 172 | (case antiq of | 
| 67571 
f858fe5531ac
more uniform treatment of formal comments within document source;
 wenzelm parents: 
67474diff
changeset | 173 | Antiquote.Text ss => eval_text ss | 
| 67466 | 174 |   | Antiquote.Control {name, body, ...} =>
 | 
| 67386 | 175 | eval ctxt ([], Token.make_src name (if null body then [] else [Token.read_cartouche body])) | 
| 67466 | 176 |   | Antiquote.Antiq {range = (pos, _), body, ...} =>
 | 
| 67386 | 177 | let val keywords = Thy_Header.get_keywords' ctxt; | 
| 67466 | 178 | in eval ctxt (Token.read_antiq keywords parse_antiq (body, pos)) end); | 
| 67386 | 179 | |
| 180 | end; | |
| 181 | ||
| 182 | ||
| 183 | (* option syntax *) | |
| 184 | ||
| 185 | fun boolean "" = true | |
| 186 | | boolean "true" = true | |
| 187 | | boolean "false" = false | |
| 67463 | 188 |   | boolean s = error ("Bad boolean value: " ^ Library.quote s);
 | 
| 67386 | 189 | |
| 190 | fun integer s = | |
| 191 | let | |
| 192 | fun int ss = | |
| 193 | (case Library.read_int ss of (i, []) => i | |
| 67463 | 194 |       | _ => error ("Bad integer value: " ^ Library.quote s));
 | 
| 67386 | 195 | in (case Symbol.explode s of "-" :: ss => ~ (int ss) | ss => int ss) end; | 
| 196 | ||
| 197 | val _ = Theory.setup | |
| 198 | (setup_option \<^binding>\<open>show_types\<close> (Config.put show_types o boolean) #> | |
| 199 | setup_option \<^binding>\<open>show_sorts\<close> (Config.put show_sorts o boolean) #> | |
| 200 | setup_option \<^binding>\<open>show_structs\<close> (Config.put show_structs o boolean) #> | |
| 201 | setup_option \<^binding>\<open>show_question_marks\<close> (Config.put show_question_marks o boolean) #> | |
| 202 | setup_option \<^binding>\<open>show_abbrevs\<close> (Config.put show_abbrevs o boolean) #> | |
| 203 | setup_option \<^binding>\<open>names_long\<close> (Config.put Name_Space.names_long o boolean) #> | |
| 204 | setup_option \<^binding>\<open>names_short\<close> (Config.put Name_Space.names_short o boolean) #> | |
| 205 | setup_option \<^binding>\<open>names_unique\<close> (Config.put Name_Space.names_unique o boolean) #> | |
| 206 | setup_option \<^binding>\<open>eta_contract\<close> (Config.put Syntax_Trans.eta_contract o boolean) #> | |
| 207 | setup_option \<^binding>\<open>display\<close> (Config.put thy_output_display o boolean) #> | |
| 208 | setup_option \<^binding>\<open>break\<close> (Config.put thy_output_break o boolean) #> | |
| 209 | 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: 
67386diff
changeset | 210 | setup_option \<^binding>\<open>mode\<close> (Config.put thy_output_modes) #> | 
| 67386 | 211 | setup_option \<^binding>\<open>margin\<close> (Config.put thy_output_margin o integer) #> | 
| 212 | setup_option \<^binding>\<open>indent\<close> (Config.put thy_output_indent o integer) #> | |
| 213 | setup_option \<^binding>\<open>source\<close> (Config.put thy_output_source o boolean) #> | |
| 214 | setup_option \<^binding>\<open>goals_limit\<close> (Config.put Goal_Display.goals_limit o integer)); | |
| 215 | ||
| 216 | end; |