author | wenzelm |
Sat, 20 Nov 2021 20:42:41 +0100 | |
changeset 74826 | 0e4d8aa61ad7 |
parent 74790 | 3ce6fb9db485 |
child 74887 | 56247fdb8bbb |
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) = |
|
73761 | 17 |
Document_Output.pretty_term ctxt (style t); |
67386 | 18 |
|
67505
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67497
diff
changeset
|
19 |
fun pretty_thms_style ctxt (style: style, ths) = |
73761 | 20 |
map (fn th => Document_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 |
73761 | 24 |
in Document_Output.pretty_term ctxt (Type.constraint (Term.fastype_of t') t') end; |
67386 | 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; |
73761 | 34 |
in Document_Output.pretty_term ctxt t' end; |
67386 | 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:
68823
diff
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:
68823
diff
changeset
|
51 |
in Pretty.str (Locale.extern thy (Locale.check thy (name, pos))) end; |
67386 | 52 |
|
74122 | 53 |
fun pretty_bundle ctxt (name, pos) = |
54 |
Pretty.str (Bundle.extern ctxt (Bundle.check ctxt (name, pos))); |
|
55 |
||
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:
68823
diff
changeset
|
56 |
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:
68823
diff
changeset
|
57 |
Pretty.str (Proof_Context.extern_class ctxt (Proof_Context.read_class ctxt s)); |
67386 | 58 |
|
59 |
fun pretty_type ctxt s = |
|
60 |
let val Type (name, _) = Proof_Context.read_type_name {proper = true, strict = false} ctxt s |
|
61 |
in Pretty.str (Proof_Context.extern_type ctxt name) end; |
|
62 |
||
70839
2136e4670ad2
clarified standard_proof_of: prefer expand_proof over somewhat adhoc strip_thm_proof;
wenzelm
parents:
70326
diff
changeset
|
63 |
fun pretty_prf full ctxt = Proof_Syntax.pretty_standard_proof_of ctxt full; |
67386 | 64 |
|
68482 | 65 |
fun pretty_theory ctxt (name, pos) = |
68484 | 66 |
(Theory.check {long = true} ctxt (name, pos); Pretty.str name); |
67386 | 67 |
|
73761 | 68 |
val basic_entity = Document_Output.antiquotation_pretty_source_embedded; |
67386 | 69 |
|
67505
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67497
diff
changeset
|
70 |
fun basic_entities name scan pretty = |
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67497
diff
changeset
|
71 |
Document_Antiquotation.setup name scan |
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67497
diff
changeset
|
72 |
(fn {context = ctxt, source = src, argument = xs} => |
73761 | 73 |
Document_Output.pretty_items_source ctxt {embedded = false} src (map (pretty ctxt) xs)); |
67386 | 74 |
|
75 |
val _ = Theory.setup |
|
67505
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67497
diff
changeset
|
76 |
(basic_entity \<^binding>\<open>prop\<close> (Term_Style.parse -- Args.prop) pretty_term_style #> |
67386 | 77 |
basic_entity \<^binding>\<open>term\<close> (Term_Style.parse -- Args.term) pretty_term_style #> |
78 |
basic_entity \<^binding>\<open>term_type\<close> (Term_Style.parse -- Args.term) pretty_term_typ #> |
|
79 |
basic_entity \<^binding>\<open>typeof\<close> (Term_Style.parse -- Args.term) pretty_term_typeof #> |
|
80 |
basic_entity \<^binding>\<open>const\<close> (Args.const {proper = true, strict = false}) pretty_const #> |
|
74563 | 81 |
basic_entity \<^binding>\<open>abbrev\<close> (Scan.lift Parse.embedded_inner_syntax) pretty_abbrev #> |
67386 | 82 |
basic_entity \<^binding>\<open>typ\<close> Args.typ_abbrev Syntax.pretty_typ #> |
74563 | 83 |
basic_entity \<^binding>\<open>locale\<close> (Scan.lift Parse.embedded_position) pretty_locale #> |
84 |
basic_entity \<^binding>\<open>bundle\<close> (Scan.lift Parse.embedded_position) pretty_bundle #> |
|
85 |
basic_entity \<^binding>\<open>class\<close> (Scan.lift Parse.embedded_inner_syntax) pretty_class #> |
|
86 |
basic_entity \<^binding>\<open>type\<close> (Scan.lift Parse.embedded_inner_syntax) pretty_type #> |
|
87 |
basic_entity \<^binding>\<open>theory\<close> (Scan.lift Parse.embedded_position) pretty_theory #> |
|
67386 | 88 |
basic_entities \<^binding>\<open>prf\<close> Attrib.thms (pretty_prf false) #> |
89 |
basic_entities \<^binding>\<open>full_prf\<close> Attrib.thms (pretty_prf true) #> |
|
67505
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67497
diff
changeset
|
90 |
Document_Antiquotation.setup \<^binding>\<open>thm\<close> (Term_Style.parse -- Attrib.thms) |
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67497
diff
changeset
|
91 |
(fn {context = ctxt, source = src, argument = arg} => |
73761 | 92 |
Document_Output.pretty_items_source ctxt {embedded = false} src (pretty_thms_style ctxt arg))); |
67386 | 93 |
|
73755 | 94 |
in end; |
67386 | 95 |
|
96 |
||
62153
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
wenzelm
parents:
62058
diff
changeset
|
97 |
(* Markdown errors *) |
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
wenzelm
parents:
62058
diff
changeset
|
98 |
|
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
wenzelm
parents:
62058
diff
changeset
|
99 |
local |
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
wenzelm
parents:
62058
diff
changeset
|
100 |
|
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
wenzelm
parents:
62058
diff
changeset
|
101 |
fun markdown_error binding = |
67386 | 102 |
Document_Antiquotation.setup binding (Scan.succeed ()) |
67463 | 103 |
(fn {source = src, ...} => |
62153
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
wenzelm
parents:
62058
diff
changeset
|
104 |
error ("Bad Markdown structure: illegal " ^ quote (Binding.name_of binding) ^ |
67463 | 105 |
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:
62058
diff
changeset
|
106 |
|
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
wenzelm
parents:
62058
diff
changeset
|
107 |
val _ = |
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
wenzelm
parents:
62058
diff
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:
62058
diff
changeset
|
112 |
|
73755 | 113 |
in end; |
62153
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
wenzelm
parents:
62058
diff
changeset
|
114 |
|
df566b87e269
more explicit errors for control symbols that are left-over after Markdown parsing;
wenzelm
parents:
62058
diff
changeset
|
115 |
|
61619 | 116 |
(* control spacing *) |
117 |
||
118 |
val _ = |
|
119 |
Theory.setup |
|
73761 | 120 |
(Document_Output.antiquotation_raw \<^binding>\<open>noindent\<close> (Scan.succeed ()) |
67463 | 121 |
(fn _ => fn () => Latex.string "\\noindent") #> |
73761 | 122 |
Document_Output.antiquotation_raw \<^binding>\<open>smallskip\<close> (Scan.succeed ()) |
67463 | 123 |
(fn _ => fn () => Latex.string "\\smallskip") #> |
73761 | 124 |
Document_Output.antiquotation_raw \<^binding>\<open>medskip\<close> (Scan.succeed ()) |
67463 | 125 |
(fn _ => fn () => Latex.string "\\medskip") #> |
73761 | 126 |
Document_Output.antiquotation_raw \<^binding>\<open>bigskip\<close> (Scan.succeed ()) |
67463 | 127 |
(fn _ => fn () => Latex.string "\\bigskip")); |
61619 | 128 |
|
129 |
||
73753 | 130 |
(* nested document text *) |
61619 | 131 |
|
132 |
local |
|
133 |
||
74790 | 134 |
fun nested_antiquotation name macro = |
73761 | 135 |
Document_Output.antiquotation_raw_embedded name (Scan.lift Args.cartouche_input) |
73752 | 136 |
(fn ctxt => fn txt => |
73761 | 137 |
(Context_Position.reports ctxt (Document_Output.document_reports txt); |
74790 | 138 |
Latex.macro macro (Document_Output.output_document ctxt {markdown = false} txt))); |
61619 | 139 |
|
140 |
val _ = |
|
141 |
Theory.setup |
|
74790 | 142 |
(nested_antiquotation \<^binding>\<open>footnote\<close> "footnote" #> |
143 |
nested_antiquotation \<^binding>\<open>emph\<close> "emph" #> |
|
144 |
nested_antiquotation \<^binding>\<open>bold\<close> "textbf"); |
|
61619 | 145 |
|
73755 | 146 |
in end; |
61619 | 147 |
|
148 |
||
73754 | 149 |
(* index entries *) |
150 |
||
151 |
local |
|
152 |
||
153 |
val index_like = Parse.$$$ "(" |-- Parse.!!! (Parse.$$$ "is" |-- Args.name --| Parse.$$$ ")"); |
|
74563 | 154 |
val index_args = Parse.enum1 "!" (Parse.embedded_input -- Scan.option index_like); |
73754 | 155 |
|
73780
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
wenzelm
parents:
73769
diff
changeset
|
156 |
fun output_text ctxt = Document_Output.output_document ctxt {markdown = false}; |
73754 | 157 |
|
158 |
fun index binding def = |
|
73761 | 159 |
Document_Output.antiquotation_raw binding (Scan.lift index_args) |
73754 | 160 |
(fn ctxt => fn args => |
73756 | 161 |
let |
73761 | 162 |
val _ = Context_Position.reports ctxt (maps (Document_Output.document_reports o #1) args); |
73756 | 163 |
fun make_item (txt, opt_like) = |
164 |
let |
|
165 |
val text = output_text ctxt txt; |
|
166 |
val like = |
|
167 |
(case opt_like of |
|
168 |
SOME s => s |
|
73767 | 169 |
| NONE => Document_Antiquotation.approx_content ctxt (Input.string_of txt)); |
73756 | 170 |
val _ = |
171 |
if is_none opt_like andalso Context_Position.is_visible ctxt then |
|
172 |
writeln ("(" ^ Markup.markup Markup.keyword2 "is" ^ " " ^ quote like ^ ")" ^ |
|
173 |
Position.here (Input.pos_of txt)) |
|
174 |
else (); |
|
175 |
in {text = text, like = like} end; |
|
176 |
in Latex.index_entry {items = map make_item args, def = def} end); |
|
73754 | 177 |
|
178 |
val _ = Theory.setup (index \<^binding>\<open>index_ref\<close> false #> index \<^binding>\<open>index_def\<close> true); |
|
179 |
||
73755 | 180 |
in end; |
73754 | 181 |
|
182 |
||
61619 | 183 |
(* quasi-formal text (unchecked) *) |
184 |
||
185 |
local |
|
186 |
||
67474
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67471
diff
changeset
|
187 |
fun report_text ctxt text = |
69965
da5e7278286b
more markup for various text kinds, notably for nested formal comments;
wenzelm
parents:
69592
diff
changeset
|
188 |
let val pos = Input.pos_of text in |
da5e7278286b
more markup for various text kinds, notably for nested formal comments;
wenzelm
parents:
69592
diff
changeset
|
189 |
Context_Position.reports ctxt |
da5e7278286b
more markup for various text kinds, notably for nested formal comments;
wenzelm
parents:
69592
diff
changeset
|
190 |
[(pos, Markup.language_text (Input.is_delimited text)), |
da5e7278286b
more markup for various text kinds, notably for nested formal comments;
wenzelm
parents:
69592
diff
changeset
|
191 |
(pos, Markup.raw_text)] |
da5e7278286b
more markup for various text kinds, notably for nested formal comments;
wenzelm
parents:
69592
diff
changeset
|
192 |
end; |
67474
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67471
diff
changeset
|
193 |
|
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67471
diff
changeset
|
194 |
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:
68823
diff
changeset
|
195 |
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:
67471
diff
changeset
|
196 |
|
61619 | 197 |
fun text_antiquotation name = |
73761 | 198 |
Document_Output.antiquotation_raw_embedded name (Scan.lift Args.text_input) |
67463 | 199 |
(fn ctxt => fn text => |
200 |
let |
|
67474
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67471
diff
changeset
|
201 |
val _ = report_text ctxt text; |
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67471
diff
changeset
|
202 |
in |
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67471
diff
changeset
|
203 |
prepare_text ctxt text |
73761 | 204 |
|> Document_Output.output_source ctxt |
205 |
|> Document_Output.isabelle ctxt |
|
67474
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67471
diff
changeset
|
206 |
end); |
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67471
diff
changeset
|
207 |
|
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67471
diff
changeset
|
208 |
val theory_text_antiquotation = |
73761 | 209 |
Document_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:
67471
diff
changeset
|
210 |
(fn ctxt => fn text => |
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67471
diff
changeset
|
211 |
let |
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67471
diff
changeset
|
212 |
val keywords = Thy_Header.get_keywords' ctxt; |
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67471
diff
changeset
|
213 |
|
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67471
diff
changeset
|
214 |
val _ = report_text ctxt text; |
67463 | 215 |
val _ = |
67474
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67471
diff
changeset
|
216 |
Input.source_explode text |
67497 | 217 |
|> Token.tokenize keywords {strict = true} |
67474
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67471
diff
changeset
|
218 |
|> maps (Token.reports keywords) |
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67471
diff
changeset
|
219 |
|> Context_Position.reports_text ctxt; |
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67471
diff
changeset
|
220 |
in |
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67471
diff
changeset
|
221 |
prepare_text ctxt text |
67495 | 222 |
|> Token.explode0 keywords |
73761 | 223 |
|> maps (Document_Output.output_token ctxt) |
224 |
|> Document_Output.isabelle ctxt |
|
67474
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67471
diff
changeset
|
225 |
end); |
61619 | 226 |
|
227 |
val _ = |
|
228 |
Theory.setup |
|
67147 | 229 |
(text_antiquotation \<^binding>\<open>text\<close> #> |
67474
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67471
diff
changeset
|
230 |
text_antiquotation \<^binding>\<open>cartouche\<close> #> |
90def2b06305
more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents:
67471
diff
changeset
|
231 |
theory_text_antiquotation); |
61619 | 232 |
|
73755 | 233 |
in end; |
61619 | 234 |
|
235 |
||
236 |
(* goal state *) |
|
237 |
||
238 |
local |
|
239 |
||
67463 | 240 |
fun goal_state name main = |
73761 | 241 |
Document_Output.antiquotation_pretty name (Scan.succeed ()) |
67463 | 242 |
(fn ctxt => fn () => |
67505
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67497
diff
changeset
|
243 |
Goal_Display.pretty_goal |
61619 | 244 |
(Config.put Goal_Display.show_main_goal main ctxt) |
67505
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67497
diff
changeset
|
245 |
(#goal (Proof.goal (Toplevel.proof_of (Toplevel.presentation_state ctxt))))); |
61619 | 246 |
|
247 |
val _ = Theory.setup |
|
67147 | 248 |
(goal_state \<^binding>\<open>goals\<close> true #> |
249 |
goal_state \<^binding>\<open>subgoals\<close> false); |
|
61619 | 250 |
|
73755 | 251 |
in end; |
61619 | 252 |
|
253 |
||
254 |
(* embedded lemma *) |
|
255 |
||
256 |
val _ = Theory.setup |
|
67386 | 257 |
(Document_Antiquotation.setup \<^binding>\<open>lemma\<close> |
74601
b7804cff55d9
clarified keywords: major take precedence for commands, but not used for antiquotations;
wenzelm
parents:
74563
diff
changeset
|
258 |
(Scan.lift (Scan.ahead Parse.not_eof) -- Args.prop -- Scan.lift Method.parse_by) |
b7804cff55d9
clarified keywords: major take precedence for commands, but not used for antiquotations;
wenzelm
parents:
74563
diff
changeset
|
259 |
(fn {context = ctxt, source = src, argument = ((prop_tok, prop), (methods, reports))} => |
61619 | 260 |
let |
261 |
val _ = Context_Position.reports ctxt reports; |
|
262 |
||
263 |
(* FIXME check proof!? *) |
|
264 |
val _ = ctxt |
|
265 |
|> Proof.theorem NONE (K I) [[(prop, [])]] |
|
74601
b7804cff55d9
clarified keywords: major take precedence for commands, but not used for antiquotations;
wenzelm
parents:
74563
diff
changeset
|
266 |
|> Proof.global_terminal_proof methods; |
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:
69965
diff
changeset
|
267 |
in |
73761 | 268 |
Document_Output.pretty_source ctxt {embedded = false} |
269 |
[hd src, prop_tok] (Document_Output.pretty_term ctxt prop) |
|
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:
69965
diff
changeset
|
270 |
end)); |
61619 | 271 |
|
272 |
||
273 |
(* verbatim text *) |
|
274 |
||
67463 | 275 |
val _ = Theory.setup |
73761 | 276 |
(Document_Output.antiquotation_verbatim_embedded \<^binding>\<open>verbatim\<close> (Scan.lift Args.text_input) |
67463 | 277 |
(fn ctxt => fn text => |
278 |
let |
|
69965
da5e7278286b
more markup for various text kinds, notably for nested formal comments;
wenzelm
parents:
69592
diff
changeset
|
279 |
val pos = Input.pos_of text; |
67463 | 280 |
val _ = |
69965
da5e7278286b
more markup for various text kinds, notably for nested formal comments;
wenzelm
parents:
69592
diff
changeset
|
281 |
Context_Position.reports ctxt |
da5e7278286b
more markup for various text kinds, notably for nested formal comments;
wenzelm
parents:
69592
diff
changeset
|
282 |
[(pos, Markup.language_verbatim (Input.is_delimited text)), |
da5e7278286b
more markup for various text kinds, notably for nested formal comments;
wenzelm
parents:
69592
diff
changeset
|
283 |
(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:
68823
diff
changeset
|
284 |
in #1 (Input.source_content text) end)); |
61619 | 285 |
|
286 |
||
71902
1529336eaedc
check bash functions against Isabelle settings environment;
wenzelm
parents:
71519
diff
changeset
|
287 |
(* bash functions *) |
1529336eaedc
check bash functions against Isabelle settings environment;
wenzelm
parents:
71519
diff
changeset
|
288 |
|
1529336eaedc
check bash functions against Isabelle settings environment;
wenzelm
parents:
71519
diff
changeset
|
289 |
val _ = Theory.setup |
73761 | 290 |
(Document_Output.antiquotation_verbatim_embedded \<^binding>\<open>bash_function\<close> |
74563 | 291 |
(Scan.lift Parse.embedded_position) Isabelle_System.check_bash_function); |
71902
1529336eaedc
check bash functions against Isabelle settings environment;
wenzelm
parents:
71519
diff
changeset
|
292 |
|
1529336eaedc
check bash functions against Isabelle settings environment;
wenzelm
parents:
71519
diff
changeset
|
293 |
|
71146 | 294 |
(* system options *) |
295 |
||
296 |
val _ = Theory.setup |
|
73761 | 297 |
(Document_Output.antiquotation_verbatim_embedded \<^binding>\<open>system_option\<close> |
74563 | 298 |
(Scan.lift Parse.embedded_position) |
71146 | 299 |
(fn ctxt => fn (name, pos) => |
300 |
let val _ = Completion.check_option (Options.default ()) ctxt (name, pos); |
|
301 |
in name end)); |
|
302 |
||
303 |
||
61619 | 304 |
(* ML text *) |
305 |
||
306 |
local |
|
307 |
||
73765
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
308 |
fun test_val (ml1, []) = ML_Lex.read "fn _ => (" @ ml1 @ ML_Lex.read ");" |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
309 |
| test_val (ml1, ml2) = ML_Lex.read "fn _ => (" @ ml1 @ ML_Lex.read " : " @ ml2 @ ML_Lex.read ");"; |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
310 |
|
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
311 |
fun test_op (ml1, ml2) = test_val (ML_Lex.read "op " @ ml1, ml2); |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
312 |
|
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
313 |
fun test_type (ml1, []) = ML_Lex.read "val _ = NONE : (" @ ml1 @ ML_Lex.read ") option;" |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
314 |
| test_type (ml1, ml2) = |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
315 |
ML_Lex.read "val _ = [NONE : (" @ ml1 @ ML_Lex.read ") option, NONE : (" @ |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
316 |
ml2 @ ML_Lex.read ") option];"; |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
317 |
|
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
318 |
fun test_exn (ml1, []) = ML_Lex.read "fn _ => (" @ ml1 @ ML_Lex.read " : exn);" |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
319 |
| test_exn (ml1, ml2) = |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
320 |
ML_Lex.read "fn _ => (" @ ml1 @ ML_Lex.read " : " @ ml2 @ ML_Lex.read " -> exn);"; |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
321 |
|
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
322 |
fun test_struct (ml, _) = |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
323 |
ML_Lex.read "functor XXX() = struct structure XX = " @ ml @ ML_Lex.read " end;"; |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
324 |
|
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
325 |
fun test_functor (Antiquote.Text tok :: _, _) = |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
326 |
ML_Lex.read "ML_Env.check_functor " @ |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
327 |
ML_Lex.read (ML_Syntax.print_string (ML_Lex.content_of tok)) |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
328 |
| test_functor _ = raise Fail "Bad ML functor specification"; |
61619 | 329 |
|
73769 | 330 |
val parse_ml0 = Args.text_input >> (fn source => ("", (source, Input.empty))); |
331 |
||
332 |
val parse_ml = |
|
333 |
Args.text_input -- Scan.optional (Args.colon |-- Args.text_input) Input.empty >> pair ""; |
|
334 |
||
335 |
val parse_exn = |
|
336 |
Args.text_input -- Scan.optional (Args.$$$ "of" |-- Args.text_input) Input.empty >> pair ""; |
|
337 |
||
338 |
val parse_type = |
|
339 |
(Parse.type_args >> (fn [] => "" | [a] => a ^ " " | bs => enclose "(" ") " (commas bs))) -- |
|
340 |
(Args.text_input -- Scan.optional (Args.$$$ "=" |-- Args.text_input) Input.empty); |
|
73765
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
341 |
|
73768 | 342 |
fun eval ctxt pos ml = |
343 |
ML_Context.eval_in (SOME ctxt) ML_Compiler.flags pos ml |
|
344 |
handle ERROR msg => error (msg ^ Position.here pos); |
|
345 |
||
346 |
fun make_text sep sources = |
|
347 |
let |
|
348 |
val (txt1, txt2) = apply2 (#1 o Input.source_content) sources; |
|
349 |
val is_ident = |
|
350 |
(case try List.last (Symbol.explode txt1) of |
|
351 |
NONE => false |
|
352 |
| SOME s => Symbol.is_ascii_letdig s); |
|
353 |
val txt = |
|
354 |
if txt2 = "" then txt1 |
|
355 |
else if sep = ":" andalso is_ident then txt1 ^ ": " ^ txt2 |
|
356 |
else txt1 ^ " " ^ sep ^ " " ^ txt2 |
|
357 |
in (txt, txt1) end; |
|
358 |
||
73765
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
359 |
fun antiquotation_ml parse test kind show_kind binding index = |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
360 |
Document_Output.antiquotation_raw binding (Scan.lift parse) |
73769 | 361 |
(fn ctxt => fn (txt0, sources) => |
73765
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
362 |
let |
73769 | 363 |
val (ml1, ml2) = apply2 ML_Lex.read_source sources; |
364 |
val ml0 = ML_Lex.read_source (Input.string txt0); |
|
365 |
val _ = test (ml0 @ ml1, ml2) |> eval ctxt (Input.pos_of (#1 sources)); |
|
61619 | 366 |
|
73765
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
367 |
val sep = if kind = "type" then "=" else if kind = "exception" then "of" else ":"; |
73768 | 368 |
val (txt, idx) = make_text sep sources; |
73765
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
369 |
|
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
370 |
val main_text = |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
371 |
Document_Output.verbatim ctxt |
73769 | 372 |
((if kind = "" orelse not show_kind then "" else kind ^ " ") ^ txt0 ^ txt); |
73765
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
373 |
val index_text = |
73780
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
wenzelm
parents:
73769
diff
changeset
|
374 |
(case index of |
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
wenzelm
parents:
73769
diff
changeset
|
375 |
NONE => [] |
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
wenzelm
parents:
73769
diff
changeset
|
376 |
| SOME def => |
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
wenzelm
parents:
73769
diff
changeset
|
377 |
let |
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
wenzelm
parents:
73769
diff
changeset
|
378 |
val ctxt' = Config.put Document_Antiquotation.thy_output_display false ctxt; |
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
wenzelm
parents:
73769
diff
changeset
|
379 |
val kind' = if kind = "" then " (ML)" else " (ML " ^ kind ^ ")"; |
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
wenzelm
parents:
73769
diff
changeset
|
380 |
val txt' = Document_Output.verbatim ctxt' idx @ Latex.string kind'; |
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
wenzelm
parents:
73769
diff
changeset
|
381 |
val like = Document_Antiquotation.approx_content ctxt' idx; |
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
wenzelm
parents:
73769
diff
changeset
|
382 |
in Latex.index_entry {items = [{text = txt', like = like}], def = def} end); |
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
wenzelm
parents:
73769
diff
changeset
|
383 |
in index_text @ main_text end); |
61619 | 384 |
|
73765
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
385 |
fun antiquotation_ml0 test kind = |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
386 |
antiquotation_ml parse_ml0 test kind false; |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
387 |
|
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
388 |
fun antiquotation_ml1 parse test kind binding = |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
389 |
antiquotation_ml parse test kind true binding (SOME true); |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
390 |
|
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
391 |
in |
61619 | 392 |
|
73765
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
393 |
val _ = |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
394 |
Theory.setup |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
395 |
(Latex.index_variants (antiquotation_ml0 test_val "") \<^binding>\<open>ML\<close> #> |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
396 |
Latex.index_variants (antiquotation_ml0 test_op "infix") \<^binding>\<open>ML_infix\<close> #> |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
397 |
Latex.index_variants (antiquotation_ml0 test_type "type") \<^binding>\<open>ML_type\<close> #> |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
398 |
Latex.index_variants (antiquotation_ml0 test_struct "structure") \<^binding>\<open>ML_structure\<close> #> |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
399 |
Latex.index_variants (antiquotation_ml0 test_functor "functor") \<^binding>\<open>ML_functor\<close> #> |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
400 |
antiquotation_ml0 (K []) "text" \<^binding>\<open>ML_text\<close> NONE #> |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
401 |
antiquotation_ml1 parse_ml test_val "" \<^binding>\<open>define_ML\<close> #> |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
402 |
antiquotation_ml1 parse_ml test_op "infix" \<^binding>\<open>define_ML_infix\<close> #> |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
403 |
antiquotation_ml1 parse_type test_type "type" \<^binding>\<open>define_ML_type\<close> #> |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
404 |
antiquotation_ml1 parse_exn test_exn "exception" \<^binding>\<open>define_ML_exception\<close> #> |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
405 |
antiquotation_ml1 parse_ml0 test_struct "structure" \<^binding>\<open>define_ML_structure\<close> #> |
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
406 |
antiquotation_ml1 parse_ml0 test_functor "functor" \<^binding>\<open>define_ML_functor\<close>); |
61619 | 407 |
|
73765
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
wenzelm
parents:
73762
diff
changeset
|
408 |
end; |
61619 | 409 |
|
410 |
||
411 |
(* URLs *) |
|
412 |
||
71519
8a96a11e0cf5
escape some special chars, notably for URL#NAME form;
wenzelm
parents:
71146
diff
changeset
|
413 |
val escape_url = |
8a96a11e0cf5
escape some special chars, notably for URL#NAME form;
wenzelm
parents:
71146
diff
changeset
|
414 |
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:
71146
diff
changeset
|
415 |
|
61619 | 416 |
val _ = Theory.setup |
74563 | 417 |
(Document_Output.antiquotation_raw_embedded \<^binding>\<open>url\<close> (Scan.lift Parse.embedded_input) |
72841 | 418 |
(fn ctxt => fn source => |
419 |
let |
|
420 |
val url = Input.string_of source; |
|
421 |
val pos = Input.pos_of source; |
|
422 |
val delimited = Input.is_delimited source; |
|
423 |
val _ = |
|
424 |
Context_Position.reports ctxt |
|
72843 | 425 |
[(pos, Markup.language_url delimited), (pos, Markup.url url)]; |
74790 | 426 |
in Latex.macro "url" (Latex.string (escape_url url)) end)); |
61619 | 427 |
|
61623 | 428 |
|
429 |
(* formal entities *) |
|
430 |
||
67463 | 431 |
local |
432 |
||
74790 | 433 |
fun entity_antiquotation name check macro = |
73761 | 434 |
Document_Output.antiquotation_raw name (Scan.lift Args.name_position) |
67463 | 435 |
(fn ctxt => fn (name, pos) => |
436 |
let val _ = check ctxt (name, pos) |
|
74790 | 437 |
in Latex.macro macro (Latex.string (Output.output name)) end); |
67463 | 438 |
|
61623 | 439 |
val _ = |
440 |
Theory.setup |
|
74790 | 441 |
(entity_antiquotation \<^binding>\<open>command\<close> Outer_Syntax.check_command "isacommand" #> |
442 |
entity_antiquotation \<^binding>\<open>method\<close> Method.check_name "isa" #> |
|
443 |
entity_antiquotation \<^binding>\<open>attribute\<close> Attrib.check_name "isa"); |
|
61623 | 444 |
|
73755 | 445 |
in end; |
67463 | 446 |
|
447 |
end; |