| author | wenzelm | 
| Thu, 06 May 2021 23:20:02 +0200 | |
| changeset 73638 | a6a9162f3ec1 | 
| parent 72843 | dd56ba1974e6 | 
| child 73752 | eeb076fc569f | 
| permissions | -rw-r--r-- | 
| 62058 | 1 | (* Title: Pure/Thy/document_antiquotations.ML | 
| 61619 | 2 | Author: Makarius | 
| 3 | ||
| 4 | Miscellaneous document antiquotations. | |
| 5 | *) | |
| 6 | ||
| 7 | structure Document_Antiquotations: sig end = | |
| 8 | struct | |
| 9 | ||
| 67386 | 10 | (* basic entities *) | 
| 11 | ||
| 12 | local | |
| 13 | ||
| 67463 | 14 | type style = term -> term; | 
| 15 | ||
| 16 | fun pretty_term_style ctxt (style: style, t) = | |
| 67386 | 17 | Thy_Output.pretty_term ctxt (style t); | 
| 18 | ||
| 67505 
ceb324e34c14
clarified signature: items with \isasep are special;
 wenzelm parents: 
67497diff
changeset | 19 | fun pretty_thms_style ctxt (style: style, ths) = | 
| 
ceb324e34c14
clarified signature: items with \isasep are special;
 wenzelm parents: 
67497diff
changeset | 20 | map (fn th => Thy_Output.pretty_term ctxt (style (Thm.full_prop_of th))) ths; | 
| 67386 | 21 | |
| 67463 | 22 | fun pretty_term_typ ctxt (style: style, t) = | 
| 67386 | 23 | let val t' = style t | 
| 24 | in Thy_Output.pretty_term ctxt (Type.constraint (Term.fastype_of t') t') end; | |
| 25 | ||
| 67463 | 26 | fun pretty_term_typeof ctxt (style: style, t) = | 
| 67386 | 27 | Syntax.pretty_typ ctxt (Term.fastype_of (style t)); | 
| 28 | ||
| 29 | fun pretty_const ctxt c = | |
| 30 | let | |
| 31 | val t = Const (c, Consts.type_scheme (Proof_Context.consts_of ctxt) c) | |
| 32 | handle TYPE (msg, _, _) => error msg; | |
| 70326 | 33 | val (t', _) = yield_singleton (Variable.import_terms true) t ctxt; | 
| 67386 | 34 | in Thy_Output.pretty_term ctxt t' end; | 
| 35 | ||
| 36 | fun pretty_abbrev ctxt s = | |
| 37 | let | |
| 38 | val t = Syntax.read_term (Proof_Context.set_mode Proof_Context.mode_abbrev ctxt) s; | |
| 39 |     fun err () = error ("Abbreviated constant expected: " ^ Syntax.string_of_term ctxt t);
 | |
| 40 | val (head, args) = Term.strip_comb t; | |
| 41 | val (c, T) = Term.dest_Const head handle TERM _ => err (); | |
| 42 | val (U, u) = Consts.the_abbreviation (Proof_Context.consts_of ctxt) c | |
| 43 | handle TYPE _ => err (); | |
| 44 | val t' = Term.betapplys (Envir.expand_atom T (U, u), args); | |
| 45 | val eq = Logic.mk_equals (t, t'); | |
| 70308 | 46 | val ctxt' = Proof_Context.augment eq ctxt; | 
| 67386 | 47 | in Proof_Context.pretty_term_abbrev ctxt' eq end; | 
| 48 | ||
| 49 | fun pretty_locale ctxt (name, pos) = | |
| 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: 
68823diff
changeset | 50 | let val thy = Proof_Context.theory_of ctxt | 
| 
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: 
68823diff
changeset | 51 | in Pretty.str (Locale.extern thy (Locale.check thy (name, pos))) end; | 
| 67386 | 52 | |
| 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: 
68823diff
changeset | 53 | fun pretty_class ctxt s = | 
| 
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: 
68823diff
changeset | 54 | Pretty.str (Proof_Context.extern_class ctxt (Proof_Context.read_class ctxt s)); | 
| 67386 | 55 | |
| 56 | fun pretty_type ctxt s = | |
| 57 |   let val Type (name, _) = Proof_Context.read_type_name {proper = true, strict = false} ctxt s
 | |
| 58 | in Pretty.str (Proof_Context.extern_type ctxt name) end; | |
| 59 | ||
| 70839 
2136e4670ad2
clarified standard_proof_of: prefer expand_proof over somewhat adhoc strip_thm_proof;
 wenzelm parents: 
70326diff
changeset | 60 | fun pretty_prf full ctxt = Proof_Syntax.pretty_standard_proof_of ctxt full; | 
| 67386 | 61 | |
| 68482 | 62 | fun pretty_theory ctxt (name, pos) = | 
| 68484 | 63 |   (Theory.check {long = true} ctxt (name, pos); Pretty.str name);
 | 
| 67386 | 64 | |
| 69592 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
69349diff
changeset | 65 | val basic_entity = Thy_Output.antiquotation_pretty_source_embedded; | 
| 67386 | 66 | |
| 67505 
ceb324e34c14
clarified signature: items with \isasep are special;
 wenzelm parents: 
67497diff
changeset | 67 | fun basic_entities name scan pretty = | 
| 
ceb324e34c14
clarified signature: items with \isasep are special;
 wenzelm parents: 
67497diff
changeset | 68 | Document_Antiquotation.setup name scan | 
| 
ceb324e34c14
clarified signature: items with \isasep are special;
 wenzelm parents: 
67497diff
changeset | 69 |     (fn {context = ctxt, source = src, argument = xs} =>
 | 
| 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: 
69965diff
changeset | 70 |       Thy_Output.pretty_items_source ctxt {embedded = false} src (map (pretty ctxt) xs));
 | 
| 67386 | 71 | |
| 72 | in | |
| 73 | ||
| 74 | val _ = Theory.setup | |
| 67505 
ceb324e34c14
clarified signature: items with \isasep are special;
 wenzelm parents: 
67497diff
changeset | 75 | (basic_entity \<^binding>\<open>prop\<close> (Term_Style.parse -- Args.prop) pretty_term_style #> | 
| 67386 | 76 | basic_entity \<^binding>\<open>term\<close> (Term_Style.parse -- Args.term) pretty_term_style #> | 
| 77 | basic_entity \<^binding>\<open>term_type\<close> (Term_Style.parse -- Args.term) pretty_term_typ #> | |
| 78 | basic_entity \<^binding>\<open>typeof\<close> (Term_Style.parse -- Args.term) pretty_term_typeof #> | |
| 79 |   basic_entity \<^binding>\<open>const\<close> (Args.const {proper = true, strict = false}) pretty_const #>
 | |
| 80 | basic_entity \<^binding>\<open>abbrev\<close> (Scan.lift Args.embedded_inner_syntax) pretty_abbrev #> | |
| 81 | basic_entity \<^binding>\<open>typ\<close> Args.typ_abbrev Syntax.pretty_typ #> | |
| 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: 
68823diff
changeset | 82 | basic_entity \<^binding>\<open>locale\<close> (Scan.lift Args.embedded_position) pretty_locale #> | 
| 67386 | 83 | basic_entity \<^binding>\<open>class\<close> (Scan.lift Args.embedded_inner_syntax) pretty_class #> | 
| 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: 
68823diff
changeset | 84 | basic_entity \<^binding>\<open>type\<close> (Scan.lift Args.embedded_inner_syntax) pretty_type #> | 
| 
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: 
68823diff
changeset | 85 | basic_entity \<^binding>\<open>theory\<close> (Scan.lift Args.embedded_position) pretty_theory #> | 
| 67386 | 86 | basic_entities \<^binding>\<open>prf\<close> Attrib.thms (pretty_prf false) #> | 
| 87 | basic_entities \<^binding>\<open>full_prf\<close> Attrib.thms (pretty_prf true) #> | |
| 67505 
ceb324e34c14
clarified signature: items with \isasep are special;
 wenzelm parents: 
67497diff
changeset | 88 | Document_Antiquotation.setup \<^binding>\<open>thm\<close> (Term_Style.parse -- Attrib.thms) | 
| 
ceb324e34c14
clarified signature: items with \isasep are special;
 wenzelm parents: 
67497diff
changeset | 89 |     (fn {context = ctxt, source = src, argument = arg} =>
 | 
| 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: 
69965diff
changeset | 90 |       Thy_Output.pretty_items_source ctxt {embedded = false} src (pretty_thms_style ctxt arg)));
 | 
| 67386 | 91 | |
| 92 | end; | |
| 93 | ||
| 94 | ||
| 62153 
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
 wenzelm parents: 
62058diff
changeset | 95 | (* Markdown errors *) | 
| 
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
 wenzelm parents: 
62058diff
changeset | 96 | |
| 
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
 wenzelm parents: 
62058diff
changeset | 97 | local | 
| 
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
 wenzelm parents: 
62058diff
changeset | 98 | |
| 
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
 wenzelm parents: 
62058diff
changeset | 99 | fun markdown_error binding = | 
| 67386 | 100 | Document_Antiquotation.setup binding (Scan.succeed ()) | 
| 67463 | 101 |     (fn {source = src, ...} =>
 | 
| 62153 
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
 wenzelm parents: 
62058diff
changeset | 102 |       error ("Bad Markdown structure: illegal " ^ quote (Binding.name_of binding) ^
 | 
| 67463 | 103 | Position.here (Position.no_range_position (#1 (Token.range_of src))))) | 
| 62153 
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
 wenzelm parents: 
62058diff
changeset | 104 | |
| 
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
 wenzelm parents: 
62058diff
changeset | 105 | in | 
| 
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
 wenzelm parents: 
62058diff
changeset | 106 | |
| 
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
 wenzelm parents: 
62058diff
changeset | 107 | val _ = | 
| 
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
 wenzelm parents: 
62058diff
changeset | 108 | Theory.setup | 
| 67147 | 109 | (markdown_error \<^binding>\<open>item\<close> #> | 
| 110 | markdown_error \<^binding>\<open>enum\<close> #> | |
| 111 | markdown_error \<^binding>\<open>descr\<close>); | |
| 62153 
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
 wenzelm parents: 
62058diff
changeset | 112 | |
| 
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
 wenzelm parents: 
62058diff
changeset | 113 | end; | 
| 
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
 wenzelm parents: 
62058diff
changeset | 114 | |
| 
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
 wenzelm parents: 
62058diff
changeset | 115 | |
| 61619 | 116 | (* control spacing *) | 
| 117 | ||
| 118 | val _ = | |
| 119 | Theory.setup | |
| 67463 | 120 | (Thy_Output.antiquotation_raw \<^binding>\<open>noindent\<close> (Scan.succeed ()) | 
| 121 | (fn _ => fn () => Latex.string "\\noindent") #> | |
| 122 | Thy_Output.antiquotation_raw \<^binding>\<open>smallskip\<close> (Scan.succeed ()) | |
| 123 | (fn _ => fn () => Latex.string "\\smallskip") #> | |
| 124 | Thy_Output.antiquotation_raw \<^binding>\<open>medskip\<close> (Scan.succeed ()) | |
| 125 | (fn _ => fn () => Latex.string "\\medskip") #> | |
| 126 | Thy_Output.antiquotation_raw \<^binding>\<open>bigskip\<close> (Scan.succeed ()) | |
| 127 | (fn _ => fn () => Latex.string "\\bigskip")); | |
| 61619 | 128 | |
| 129 | ||
| 130 | (* control style *) | |
| 131 | ||
| 132 | local | |
| 133 | ||
| 134 | fun control_antiquotation name s1 s2 = | |
| 69592 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
69349diff
changeset | 135 | Thy_Output.antiquotation_raw_embedded name (Scan.lift Args.cartouche_input) | 
| 67569 
5d71b114e7a3
avoid proliferation of language_document reports;
 wenzelm parents: 
67505diff
changeset | 136 |     (fn ctxt => Latex.enclose_block s1 s2 o Thy_Output.output_document ctxt {markdown = false});
 | 
| 61619 | 137 | |
| 138 | in | |
| 139 | ||
| 140 | val _ = | |
| 141 | Theory.setup | |
| 67147 | 142 |    (control_antiquotation \<^binding>\<open>footnote\<close> "\\footnote{" "}" #>
 | 
| 143 |     control_antiquotation \<^binding>\<open>emph\<close> "\\emph{" "}" #>
 | |
| 144 |     control_antiquotation \<^binding>\<open>bold\<close> "\\textbf{" "}");
 | |
| 61619 | 145 | |
| 146 | end; | |
| 147 | ||
| 148 | ||
| 149 | (* quasi-formal text (unchecked) *) | |
| 150 | ||
| 151 | local | |
| 152 | ||
| 67474 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 153 | fun report_text ctxt text = | 
| 69965 
da5e7278286b
more markup for various text kinds, notably for nested formal comments;
 wenzelm parents: 
69592diff
changeset | 154 | let val pos = Input.pos_of text in | 
| 
da5e7278286b
more markup for various text kinds, notably for nested formal comments;
 wenzelm parents: 
69592diff
changeset | 155 | Context_Position.reports ctxt | 
| 
da5e7278286b
more markup for various text kinds, notably for nested formal comments;
 wenzelm parents: 
69592diff
changeset | 156 | [(pos, Markup.language_text (Input.is_delimited text)), | 
| 
da5e7278286b
more markup for various text kinds, notably for nested formal comments;
 wenzelm parents: 
69592diff
changeset | 157 | (pos, Markup.raw_text)] | 
| 
da5e7278286b
more markup for various text kinds, notably for nested formal comments;
 wenzelm parents: 
69592diff
changeset | 158 | end; | 
| 67474 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 159 | |
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 160 | fun prepare_text ctxt = | 
| 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: 
68823diff
changeset | 161 | Input.source_content #> #1 #> Document_Antiquotation.prepare_lines ctxt; | 
| 67474 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 162 | |
| 61619 | 163 | fun text_antiquotation name = | 
| 69592 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
69349diff
changeset | 164 | Thy_Output.antiquotation_raw_embedded name (Scan.lift Args.text_input) | 
| 67463 | 165 | (fn ctxt => fn text => | 
| 166 | let | |
| 67474 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 167 | val _ = report_text ctxt text; | 
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 168 | in | 
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 169 | prepare_text ctxt text | 
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 170 | |> Thy_Output.output_source ctxt | 
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 171 | |> Thy_Output.isabelle ctxt | 
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 172 | end); | 
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 173 | |
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 174 | val theory_text_antiquotation = | 
| 69592 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
69349diff
changeset | 175 | Thy_Output.antiquotation_raw_embedded \<^binding>\<open>theory_text\<close> (Scan.lift Args.text_input) | 
| 67474 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 176 | (fn ctxt => fn text => | 
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 177 | let | 
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 178 | val keywords = Thy_Header.get_keywords' ctxt; | 
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 179 | |
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 180 | val _ = report_text ctxt text; | 
| 67463 | 181 | val _ = | 
| 67474 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 182 | Input.source_explode text | 
| 67497 | 183 |           |> Token.tokenize keywords {strict = true}
 | 
| 67474 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 184 | |> maps (Token.reports keywords) | 
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 185 | |> Context_Position.reports_text ctxt; | 
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 186 | in | 
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 187 | prepare_text ctxt text | 
| 67495 | 188 | |> Token.explode0 keywords | 
| 67474 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 189 | |> maps (Thy_Output.output_token ctxt) | 
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 190 | |> Thy_Output.isabelle ctxt | 
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 191 | end); | 
| 61619 | 192 | |
| 193 | in | |
| 194 | ||
| 195 | val _ = | |
| 196 | Theory.setup | |
| 67147 | 197 | (text_antiquotation \<^binding>\<open>text\<close> #> | 
| 67474 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 198 | text_antiquotation \<^binding>\<open>cartouche\<close> #> | 
| 
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
 wenzelm parents: 
67471diff
changeset | 199 | theory_text_antiquotation); | 
| 61619 | 200 | |
| 201 | end; | |
| 202 | ||
| 203 | ||
| 204 | ||
| 205 | ||
| 206 | (* goal state *) | |
| 207 | ||
| 208 | local | |
| 209 | ||
| 67463 | 210 | fun goal_state name main = | 
| 211 | Thy_Output.antiquotation_pretty name (Scan.succeed ()) | |
| 212 | (fn ctxt => fn () => | |
| 67505 
ceb324e34c14
clarified signature: items with \isasep are special;
 wenzelm parents: 
67497diff
changeset | 213 | Goal_Display.pretty_goal | 
| 61619 | 214 | (Config.put Goal_Display.show_main_goal main ctxt) | 
| 67505 
ceb324e34c14
clarified signature: items with \isasep are special;
 wenzelm parents: 
67497diff
changeset | 215 | (#goal (Proof.goal (Toplevel.proof_of (Toplevel.presentation_state ctxt))))); | 
| 61619 | 216 | |
| 217 | in | |
| 218 | ||
| 219 | val _ = Theory.setup | |
| 67147 | 220 | (goal_state \<^binding>\<open>goals\<close> true #> | 
| 221 | goal_state \<^binding>\<open>subgoals\<close> false); | |
| 61619 | 222 | |
| 223 | end; | |
| 224 | ||
| 225 | ||
| 226 | (* embedded lemma *) | |
| 227 | ||
| 228 | val _ = Theory.setup | |
| 67386 | 229 | (Document_Antiquotation.setup \<^binding>\<open>lemma\<close> | 
| 61619 | 230 | (Scan.lift (Scan.ahead Parse.not_eof) -- Args.prop -- | 
| 231 | Scan.lift (Parse.position (Parse.reserved "by") -- Method.parse -- Scan.option Method.parse)) | |
| 67463 | 232 |     (fn {context = ctxt, source = src, argument = ((prop_tok, prop), (((_, by_pos), m1), m2))} =>
 | 
| 61619 | 233 | let | 
| 66067 | 234 | val reports = | 
| 235 | (by_pos, Markup.keyword1 |> Markup.keyword_properties) :: | |
| 236 | maps Method.reports_of (m1 :: the_list m2); | |
| 61619 | 237 | val _ = Context_Position.reports ctxt reports; | 
| 238 | ||
| 239 | (* FIXME check proof!? *) | |
| 240 | val _ = ctxt | |
| 241 | |> Proof.theorem NONE (K I) [[(prop, [])]] | |
| 242 | |> Proof.global_terminal_proof (m1, m2); | |
| 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: 
69965diff
changeset | 243 | in | 
| 
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: 
69965diff
changeset | 244 |         Thy_Output.pretty_source ctxt {embedded = false}
 | 
| 
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: 
69965diff
changeset | 245 | [hd src, prop_tok] (Thy_Output.pretty_term ctxt prop) | 
| 
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: 
69965diff
changeset | 246 | end)); | 
| 61619 | 247 | |
| 248 | ||
| 249 | (* verbatim text *) | |
| 250 | ||
| 67463 | 251 | val _ = Theory.setup | 
| 69592 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
69349diff
changeset | 252 | (Thy_Output.antiquotation_verbatim_embedded \<^binding>\<open>verbatim\<close> (Scan.lift Args.text_input) | 
| 67463 | 253 | (fn ctxt => fn text => | 
| 254 | let | |
| 69965 
da5e7278286b
more markup for various text kinds, notably for nested formal comments;
 wenzelm parents: 
69592diff
changeset | 255 | val pos = Input.pos_of text; | 
| 67463 | 256 | val _ = | 
| 69965 
da5e7278286b
more markup for various text kinds, notably for nested formal comments;
 wenzelm parents: 
69592diff
changeset | 257 | Context_Position.reports ctxt | 
| 
da5e7278286b
more markup for various text kinds, notably for nested formal comments;
 wenzelm parents: 
69592diff
changeset | 258 | [(pos, Markup.language_verbatim (Input.is_delimited text)), | 
| 
da5e7278286b
more markup for various text kinds, notably for nested formal comments;
 wenzelm parents: 
69592diff
changeset | 259 | (pos, Markup.raw_text)]; | 
| 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: 
68823diff
changeset | 260 | in #1 (Input.source_content text) end)); | 
| 61619 | 261 | |
| 262 | ||
| 71902 
1529336eaedc
check bash functions against Isabelle settings environment;
 wenzelm parents: 
71519diff
changeset | 263 | (* bash functions *) | 
| 
1529336eaedc
check bash functions against Isabelle settings environment;
 wenzelm parents: 
71519diff
changeset | 264 | |
| 
1529336eaedc
check bash functions against Isabelle settings environment;
 wenzelm parents: 
71519diff
changeset | 265 | val _ = Theory.setup | 
| 
1529336eaedc
check bash functions against Isabelle settings environment;
 wenzelm parents: 
71519diff
changeset | 266 | (Thy_Output.antiquotation_verbatim_embedded \<^binding>\<open>bash_function\<close> | 
| 
1529336eaedc
check bash functions against Isabelle settings environment;
 wenzelm parents: 
71519diff
changeset | 267 | (Scan.lift Args.embedded_position) Isabelle_System.check_bash_function); | 
| 
1529336eaedc
check bash functions against Isabelle settings environment;
 wenzelm parents: 
71519diff
changeset | 268 | |
| 
1529336eaedc
check bash functions against Isabelle settings environment;
 wenzelm parents: 
71519diff
changeset | 269 | |
| 71146 | 270 | (* system options *) | 
| 271 | ||
| 272 | val _ = Theory.setup | |
| 273 | (Thy_Output.antiquotation_verbatim_embedded \<^binding>\<open>system_option\<close> | |
| 274 | (Scan.lift Args.embedded_position) | |
| 275 | (fn ctxt => fn (name, pos) => | |
| 276 | let val _ = Completion.check_option (Options.default ()) ctxt (name, pos); | |
| 277 | in name end)); | |
| 278 | ||
| 279 | ||
| 61619 | 280 | (* ML text *) | 
| 281 | ||
| 282 | local | |
| 283 | ||
| 67463 | 284 | fun ml_text name ml = | 
| 69592 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
69349diff
changeset | 285 | Thy_Output.antiquotation_verbatim_embedded name (Scan.lift Args.text_input) | 
| 67463 | 286 | (fn ctxt => fn text => | 
| 287 | let val _ = ML_Context.eval_in (SOME ctxt) ML_Compiler.flags (Input.pos_of text) (ml text) | |
| 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: 
68823diff
changeset | 288 | in #1 (Input.source_content text) end); | 
| 61619 | 289 | |
| 290 | fun ml_enclose bg en source = | |
| 68823 | 291 | ML_Lex.read bg @ ML_Lex.read_source source @ ML_Lex.read en; | 
| 61619 | 292 | |
| 293 | in | |
| 294 | ||
| 295 | val _ = Theory.setup | |
| 67147 | 296 |  (ml_text \<^binding>\<open>ML\<close> (ml_enclose "fn _ => (" ");") #>
 | 
| 297 | ml_text \<^binding>\<open>ML_op\<close> (ml_enclose "fn _ => (op " ");") #> | |
| 298 |   ml_text \<^binding>\<open>ML_type\<close> (ml_enclose "val _ = NONE : (" ") option;") #>
 | |
| 299 | ml_text \<^binding>\<open>ML_structure\<close> | |
| 61619 | 300 | (ml_enclose "functor XXX() = struct structure XX = " " end;") #> | 
| 301 | ||
| 67147 | 302 | ml_text \<^binding>\<open>ML_functor\<close> (* FIXME formal treatment of functor name (!?) *) | 
| 61619 | 303 | (fn source => | 
| 304 |       ML_Lex.read ("ML_Env.check_functor " ^
 | |
| 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: 
68823diff
changeset | 305 | ML_Syntax.print_string (#1 (Input.source_content source)))) #> | 
| 61619 | 306 | |
| 67147 | 307 | ml_text \<^binding>\<open>ML_text\<close> (K [])); | 
| 61619 | 308 | |
| 309 | end; | |
| 310 | ||
| 311 | ||
| 312 | (* URLs *) | |
| 313 | ||
| 71519 
8a96a11e0cf5
escape some special chars, notably for URL#NAME form;
 wenzelm parents: 
71146diff
changeset | 314 | val escape_url = | 
| 
8a96a11e0cf5
escape some special chars, notably for URL#NAME form;
 wenzelm parents: 
71146diff
changeset | 315 | translate_string (fn c => if c = "%" orelse c = "#" orelse c = "^" then "\\" ^ c else c); | 
| 
8a96a11e0cf5
escape some special chars, notably for URL#NAME form;
 wenzelm parents: 
71146diff
changeset | 316 | |
| 61619 | 317 | val _ = Theory.setup | 
| 72841 | 318 | (Thy_Output.antiquotation_raw_embedded \<^binding>\<open>url\<close> (Scan.lift Args.embedded_input) | 
| 319 | (fn ctxt => fn source => | |
| 320 | let | |
| 321 | val url = Input.string_of source; | |
| 322 | val pos = Input.pos_of source; | |
| 323 | val delimited = Input.is_delimited source; | |
| 324 | val _ = | |
| 325 | Context_Position.reports ctxt | |
| 72843 | 326 | [(pos, Markup.language_url delimited), (pos, Markup.url url)]; | 
| 71519 
8a96a11e0cf5
escape some special chars, notably for URL#NAME form;
 wenzelm parents: 
71146diff
changeset | 327 |       in Latex.enclose_block "\\url{" "}" [Latex.string (escape_url url)] end));
 | 
| 61619 | 328 | |
| 61623 | 329 | |
| 330 | (* formal entities *) | |
| 331 | ||
| 67463 | 332 | local | 
| 333 | ||
| 334 | fun entity_antiquotation name check bg en = | |
| 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: 
68823diff
changeset | 335 | Thy_Output.antiquotation_raw name (Scan.lift Args.name_position) | 
| 67463 | 336 | (fn ctxt => fn (name, pos) => | 
| 337 | let val _ = check ctxt (name, pos) | |
| 338 | in Latex.enclose_block bg en [Latex.string (Output.output name)] end); | |
| 339 | ||
| 340 | in | |
| 61623 | 341 | |
| 342 | val _ = | |
| 343 | Theory.setup | |
| 67463 | 344 |    (entity_antiquotation \<^binding>\<open>command\<close> Outer_Syntax.check_command "\\isacommand{" "}" #>
 | 
| 345 |     entity_antiquotation \<^binding>\<open>method\<close> Method.check_name "\\isa{" "}" #>
 | |
| 346 |     entity_antiquotation \<^binding>\<open>attribute\<close> Attrib.check_name "\\isa{" "}");
 | |
| 61623 | 347 | |
| 61619 | 348 | end; | 
| 67463 | 349 | |
| 350 | end; |