diff --git a/src/Pure/Thy/document_output.ML b/src/Pure/Thy/document_output.ML --- a/src/Pure/Thy/document_output.ML +++ b/src/Pure/Thy/document_output.ML @@ -1,594 +1,594 @@ (* Title: Pure/Thy/document_output.ML Author: Makarius Theory document output. *) signature DOCUMENT_OUTPUT = sig val document_reports: Input.source -> Position.report list val output_document: Proof.context -> {markdown: bool} -> Input.source -> Latex.text val document_output: {markdown: bool, markup: Latex.text -> Latex.text} -> (xstring * Position.T) option * Input.source -> Toplevel.transition -> Toplevel.transition val check_comments: Proof.context -> Symbol_Pos.T list -> unit val output_token: Proof.context -> Token.T -> Latex.text val output_source: Proof.context -> string -> Latex.text type segment = {span: Command_Span.span, command: Toplevel.transition, prev_state: Toplevel.state, state: Toplevel.state} val present_thy: Options.T -> theory -> segment list -> Latex.text val pretty_term: Proof.context -> term -> Pretty.T val pretty_thm: Proof.context -> thm -> Pretty.T val isabelle: Proof.context -> Latex.text -> Latex.text val isabelle_typewriter: Proof.context -> Latex.text -> Latex.text val typewriter: Proof.context -> string -> Latex.text val verbatim: Proof.context -> string -> Latex.text val source: Proof.context -> {embedded: bool} -> Token.src -> Latex.text val pretty: Proof.context -> Pretty.T -> Latex.text val pretty_source: Proof.context -> {embedded: bool} -> Token.src -> Pretty.T -> Latex.text val pretty_items: Proof.context -> Pretty.T list -> Latex.text val pretty_items_source: Proof.context -> {embedded: bool} -> Token.src -> Pretty.T list -> Latex.text val antiquotation_pretty: binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory val antiquotation_pretty_embedded: binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory val antiquotation_pretty_source: binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory val antiquotation_pretty_source_embedded: binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory val antiquotation_raw: binding -> 'a context_parser -> (Proof.context -> 'a -> Latex.text) -> theory -> theory val antiquotation_raw_embedded: binding -> 'a context_parser -> (Proof.context -> 'a -> Latex.text) -> theory -> theory val antiquotation_verbatim: binding -> 'a context_parser -> (Proof.context -> 'a -> string) -> theory -> theory val antiquotation_verbatim_embedded: binding -> 'a context_parser -> (Proof.context -> 'a -> string) -> theory -> theory end; structure Document_Output: DOCUMENT_OUTPUT = struct (* output document source *) fun document_reports txt = let val pos = Input.pos_of txt in [(pos, Markup.language_document (Input.is_delimited txt)), (pos, Markup.plain_text)] end; fun output_comment ctxt (kind, syms) = (case kind of Comment.Comment => Input.cartouche_content syms |> output_document (ctxt |> Config.put Document_Antiquotation.thy_output_display false) {markdown = false} |> XML.enclose "%\n\\isamarkupcmt{" "%\n}" | Comment.Cancel => Symbol_Pos.cartouche_content syms |> Latex.symbols_output |> XML.enclose "%\n\\isamarkupcancel{" "}" | Comment.Latex => Latex.symbols (Symbol_Pos.cartouche_content syms) | Comment.Marker => []) and output_comment_document ctxt (comment, syms) = (case comment of SOME kind => output_comment ctxt (kind, syms) | NONE => Latex.symbols syms) and output_document_text ctxt syms = Comment.read_body syms |> maps (output_comment_document ctxt) and output_document ctxt {markdown} source = let val pos = Input.pos_of source; val syms = Input.source_explode source; val output_antiquotes = maps (Document_Antiquotation.evaluate (output_document_text ctxt) ctxt); fun output_line line = (if Markdown.line_is_item line then Latex.string "\\item " else []) @ output_antiquotes (Markdown.line_content line); fun output_block (Markdown.Par lines) = separate (XML.Text "\n") (map (Latex.block o output_line) lines) | output_block (Markdown.List {kind, body, ...}) = - Latex.environment_text (Markdown.print_kind kind) (output_blocks body) + Latex.environment (Markdown.print_kind kind) (output_blocks body) and output_blocks blocks = separate (XML.Text "\n\n") (map (Latex.block o output_block) blocks); in if Toplevel.is_skipped_proof (Toplevel.presentation_state ctxt) then [] else if markdown andalso exists (Markdown.is_control o Symbol_Pos.symbol) syms then let val ants = Antiquote.parse_comments pos syms; val reports = Antiquote.antiq_reports ants; val blocks = Markdown.read_antiquotes ants; val _ = Context_Position.reports ctxt (reports @ Markdown.reports blocks); in output_blocks blocks end else let val ants = Antiquote.parse_comments pos (trim (Symbol.is_blank o Symbol_Pos.symbol) syms); val reports = Antiquote.antiq_reports ants; val _ = Context_Position.reports ctxt (reports @ Markdown.text_reports ants); in output_antiquotes ants end end; fun document_output {markdown, markup} (loc, txt) = let fun output st = let val ctxt = Toplevel.presentation_context st; val _ = Context_Position.reports ctxt (document_reports txt); in txt |> output_document ctxt {markdown = markdown} |> markup end; in Toplevel.present (fn st => (case loc of NONE => output st | SOME (_, pos) => error ("Illegal target specification -- not a theory context" ^ Position.here pos))) o Toplevel.present_local_theory loc output end; (* output tokens with formal comments *) local val output_symbols_antiq = (fn Antiquote.Text syms => Latex.symbols_output syms | Antiquote.Control {name = (name, _), body, ...} => Latex.string (Latex.output_symbols [Symbol.encode (Symbol.Control name)]) @ Latex.symbols_output body | Antiquote.Antiq {body, ...} => XML.enclose "%\n\\isaantiq\n" "{}%\n\\endisaantiq\n" (Latex.symbols_output body)); fun output_comment_symbols ctxt {antiq} (comment, syms) = (case (comment, antiq) of (NONE, false) => Latex.symbols_output syms | (NONE, true) => Antiquote.parse_comments (#1 (Symbol_Pos.range syms)) syms |> maps output_symbols_antiq | (SOME comment, _) => output_comment ctxt (comment, syms)); fun output_body ctxt antiq bg en syms = Comment.read_body syms |> maps (output_comment_symbols ctxt {antiq = antiq}) |> XML.enclose bg en; in fun output_token ctxt tok = let fun output antiq bg en = output_body ctxt antiq bg en (Input.source_explode (Token.input_of tok)); in (case Token.kind_of tok of Token.Comment NONE => [] | Token.Comment (SOME Comment.Marker) => [] | Token.Command => output false "\\isacommand{" "}" | Token.Keyword => if Symbol.is_ascii_identifier (Token.content_of tok) then output false "\\isakeyword{" "}" else output false "" "" | Token.String => output false "{\\isachardoublequoteopen}" "{\\isachardoublequoteclose}" | Token.Alt_String => output false "{\\isacharbackquoteopen}" "{\\isacharbackquoteclose}" | Token.Verbatim => output true "{\\isacharverbatimopen}" "{\\isacharverbatimclose}" | Token.Cartouche => output false "{\\isacartoucheopen}" "{\\isacartoucheclose}" | Token.Control control => output_body ctxt false "" "" (Antiquote.control_symbols control) | _ => output false "" "") end handle ERROR msg => error (msg ^ Position.here (Token.pos_of tok)); fun output_source ctxt s = output_body ctxt false "" "" (Symbol_Pos.explode (s, Position.none)); fun check_comments ctxt = Comment.read_body #> List.app (fn (comment, syms) => let val pos = #1 (Symbol_Pos.range syms); val _ = comment |> Option.app (fn kind => Context_Position.reports ctxt (map (pair pos) (Comment.kind_markups kind))); val _ = output_comment_symbols ctxt {antiq = false} (comment, syms); in if comment = SOME Comment.Comment then check_comments ctxt syms else () end); end; (** present theory source **) (* presentation tokens *) datatype token = Ignore | Token of Token.T | Output of Latex.text; fun token_with pred (Token tok) = pred tok | token_with _ _ = false; val white_token = token_with Document_Source.is_white; val white_comment_token = token_with Document_Source.is_white_comment; val blank_token = token_with Token.is_blank; val newline_token = token_with Token.is_newline; fun present_token ctxt tok = (case tok of Ignore => [] | Token tok => output_token ctxt tok | Output output => output); (* command spans *) type command = string * Position.T; (*name, position*) type source = (token * (string * int)) list; (*token, markup flag, meta-comment depth*) datatype span = Span of command * (source * source * source * source) * bool; fun make_span cmd src = let fun chop_newline (tok :: toks) = if newline_token (fst tok) then ([tok], toks, true) else ([], tok :: toks, false) | chop_newline [] = ([], [], false); val (((src_prefix, src_main), src_suffix1), (src_suffix2, src_appendix, newline)) = src |> chop_prefix (white_token o fst) ||>> chop_suffix (white_token o fst) ||>> chop_prefix (white_comment_token o fst) ||> chop_newline; in Span (cmd, (src_prefix, src_main, src_suffix1 @ src_suffix2, src_appendix), newline) end; (* present spans *) local fun err_bad_nesting here = error ("Bad nesting of commands in presentation" ^ here); fun edge which f (x: string option, y) = if x = y then I else (case which (x, y) of NONE => I | SOME txt => fold cons (Latex.string (f txt))); val markup_tag = YXML.output_markup o Markup.latex_tag; val markup_delim = YXML.output_markup o Markup.latex_delim; val bg_delim = #1 o markup_delim; val en_delim = #2 o markup_delim; val begin_tag = edge #2 (#1 o markup_tag); val end_tag = edge #1 (#2 o markup_tag); fun open_delim delim e = edge #2 bg_delim e #> delim #> edge #2 en_delim e; fun close_delim delim e = edge #1 bg_delim e #> delim #> edge #1 en_delim e; fun document_tag cmd_pos state state' tagging_stack = let val ctxt' = Toplevel.presentation_context state'; val nesting = Toplevel.level state' - Toplevel.level state; val (tagging, taggings) = tagging_stack; val (tag', tagging') = Document_Source.update_tagging ctxt' tagging; val tagging_stack' = if nesting = 0 andalso not (Toplevel.is_proof state) then tagging_stack else if nesting >= 0 then (tagging', replicate nesting tagging @ taggings) else (case drop (~ nesting - 1) taggings of tg :: tgs => (tg, tgs) | [] => err_bad_nesting (Position.here cmd_pos)); in (tag', tagging_stack') end; fun read_tag s = (case space_explode "%" s of ["", b] => (SOME b, NONE) | [a, b] => (NONE, SOME (a, b)) | _ => error ("Bad document_tags specification: " ^ quote s)); in fun make_command_tag options keywords = let val document_tags = map read_tag (space_explode "," (Options.string options \<^system_option>\document_tags\)); val document_tags_default = map_filter #1 document_tags; val document_tags_command = map_filter #2 document_tags; in fn cmd_name => fn state => fn state' => fn active_tag => let val keyword_tags = if cmd_name = "end" andalso Toplevel.is_end_theory state' then ["theory"] else Keyword.command_tags keywords cmd_name; val command_tags = the_list (AList.lookup (op =) document_tags_command cmd_name) @ keyword_tags @ document_tags_default; val active_tag' = (case command_tags of default_tag :: _ => SOME default_tag | [] => if Keyword.is_vacuous keywords cmd_name andalso Toplevel.is_proof state then active_tag else NONE); in active_tag' end end; fun present_span command_tag span state state' (tagging_stack, active_tag, newline, latex, present_cont) = let val ctxt' = Toplevel.presentation_context state'; val present = fold (fn (tok, (flag, 0)) => fold cons (present_token ctxt' tok) #> fold cons (Latex.string flag) | _ => I); val Span ((cmd_name, cmd_pos), srcs, span_newline) = span; val (tag', tagging_stack') = document_tag cmd_pos state state' tagging_stack; val active_tag' = if is_some tag' then Option.map #1 tag' else command_tag cmd_name state state' active_tag; val edge = (active_tag, active_tag'); val newline' = if is_none active_tag' then span_newline else newline; val latex' = latex |> end_tag edge |> close_delim (fst present_cont) edge |> snd present_cont |> open_delim (present (#1 srcs)) edge |> begin_tag edge |> present (#2 srcs); val present_cont' = if newline then (present (#3 srcs), present (#4 srcs)) else (I, present (#3 srcs) #> present (#4 srcs)); in (tagging_stack', active_tag', newline', latex', present_cont') end; fun present_trailer ((_, tags), active_tag, _, latex, present_cont) = if not (null tags) then err_bad_nesting " at end of theory" else latex |> end_tag (active_tag, NONE) |> close_delim (fst present_cont) (active_tag, NONE) |> snd present_cont; end; (* present_thy *) type segment = {span: Command_Span.span, command: Toplevel.transition, prev_state: Toplevel.state, state: Toplevel.state}; local val markup_true = "\\isamarkuptrue%\n"; val markup_false = "\\isamarkupfalse%\n"; fun command_output output tok = if Token.is_command tok then SOME (Token.put_output output tok) else NONE; fun segment_content (segment: segment) = let val toks = Command_Span.content (#span segment) in (case Toplevel.output_of (#state segment) of NONE => toks | SOME output => map_filter (command_output output) toks) end; fun output_command keywords = Scan.some (fn tok => if Token.is_command tok then let val name = Token.content_of tok; val is_document = Keyword.is_document keywords name; val is_document_raw = Keyword.is_document_raw keywords name; val flag = if is_document andalso not is_document_raw then markup_true else ""; in if is_document andalso is_some (Token.get_output tok) then SOME ((name, Token.pos_of tok), the (Token.get_output tok), flag) else NONE end else NONE); val opt_newline = Scan.option (Scan.one Token.is_newline); val ignore = Scan.depend (fn d => opt_newline |-- Scan.one Token.is_begin_ignore >> pair (d + 1)) || Scan.depend (fn d => Scan.one Token.is_end_ignore --| (if d = 0 then Scan.fail_with (K (fn () => "Bad nesting of meta-comments")) else opt_newline) >> pair (d - 1)); in fun present_thy options thy (segments: segment list) = let val keywords = Thy_Header.get_keywords thy; (* tokens *) val ignored = Scan.state --| ignore >> (fn d => [(NONE, (Ignore, ("", d)))]); val output = Scan.peek (fn d => Document_Source.improper |-- output_command keywords --| Document_Source.improper_end >> (fn (kind, body, flag) => [(SOME kind, (Output body, (flag, d)))])); val command = Scan.peek (fn d => Scan.optional (Scan.one Token.is_command_modifier ::: Document_Source.improper) [] -- Scan.one Token.is_command --| Document_Source.annotation >> (fn (cmd_mod, cmd) => map (fn tok => (NONE, (Token tok, ("", d)))) cmd_mod @ [(SOME (Token.content_of cmd, Token.pos_of cmd), (Token cmd, (markup_false, d)))])); val cmt = Scan.peek (fn d => Scan.one Document_Source.is_black_comment >> (fn tok => [(NONE, (Token tok, ("", d)))])); val other = Scan.peek (fn d => Parse.not_eof >> (fn tok => [(NONE, (Token tok, ("", d)))])); val tokens = ignored || output || command || cmt || other; (* spans *) val is_eof = fn (_, (Token x, _)) => Token.is_eof x | _ => false; val stopper = Scan.stopper (K (NONE, (Token Token.eof, ("", 0)))) is_eof; val cmd = Scan.one (is_some o fst); val non_cmd = Scan.one (is_none o fst andf not o is_eof) >> #2; val white_comments = Scan.many (white_comment_token o fst o snd); val blank = Scan.one (blank_token o fst o snd); val newline = Scan.one (newline_token o fst o snd); val before_cmd = Scan.option (newline -- white_comments) -- Scan.option (newline -- white_comments) -- Scan.option (blank -- white_comments) -- cmd; val span = Scan.repeat non_cmd -- cmd -- Scan.repeat (Scan.unless before_cmd non_cmd) -- Scan.option (newline >> (single o snd)) >> (fn (((toks1, (cmd, tok2)), toks3), tok4) => make_span (the cmd) (toks1 @ (tok2 :: (toks3 @ the_default [] tok4)))); val spans = segments |> maps segment_content |> drop_suffix Token.is_space |> Source.of_list |> Source.source' 0 Token.stopper (Scan.error (Scan.bulk tokens >> flat)) |> Source.source stopper (Scan.error (Scan.bulk span)) |> Source.exhaust; val command_results = segments |> map_filter (fn {command, state, ...} => if Toplevel.is_ignored command then NONE else SOME (command, state)); (* present commands *) val command_tag = make_command_tag options keywords; fun present_command tr span st st' = Toplevel.setmp_thread_position tr (present_span command_tag span st st'); fun present _ [] = I | present st ((span, (tr, st')) :: rest) = present_command tr span st st' #> present st' rest; in if length command_results = length spans then (([], []), NONE, true, [], (I, I)) |> present (Toplevel.init_toplevel ()) (spans ~~ command_results) |> present_trailer |> rev else error "Messed-up outer syntax for presentation" end; end; (** standard output operations **) (* pretty printing *) fun pretty_term ctxt t = Syntax.pretty_term (Proof_Context.augment t ctxt) t; fun pretty_thm ctxt = pretty_term ctxt o Thm.full_prop_of; (* default output *) fun isabelle ctxt body = if Config.get ctxt Document_Antiquotation.thy_output_display then Latex.environment "isabelle" body else Latex.macro "isa" body; fun isabelle_typewriter ctxt body = if Config.get ctxt Document_Antiquotation.thy_output_display then Latex.environment "isabellett" body else Latex.macro "isatt" body; fun typewriter ctxt s = isabelle_typewriter ctxt (Latex.string (Latex.output_ascii s)); fun verbatim ctxt = if Config.get ctxt Document_Antiquotation.thy_output_display then Document_Antiquotation.indent_lines ctxt #> typewriter ctxt else split_lines #> map (typewriter ctxt #> Latex.block) #> separate (XML.Text "\\isanewline%\n"); fun token_source ctxt {embedded} tok = if Token.is_kind Token.Cartouche tok andalso embedded andalso not (Config.get ctxt Document_Antiquotation.thy_output_source_cartouche) then Token.content_of tok else Token.unparse tok; fun is_source ctxt = Config.get ctxt Document_Antiquotation.thy_output_source orelse Config.get ctxt Document_Antiquotation.thy_output_source_cartouche; fun source ctxt embedded = Token.args_of_src #> map (token_source ctxt embedded #> Document_Antiquotation.prepare_lines ctxt) #> space_implode " " #> output_source ctxt #> isabelle ctxt; fun pretty ctxt = Document_Antiquotation.output ctxt #> Latex.string #> isabelle ctxt; fun pretty_source ctxt embedded src prt = if is_source ctxt then source ctxt embedded src else pretty ctxt prt; fun pretty_items ctxt = map (Document_Antiquotation.output ctxt #> XML.Text) #> separate (XML.Text "\\isasep\\isanewline%\n") #> isabelle ctxt; fun pretty_items_source ctxt embedded src prts = if is_source ctxt then source ctxt embedded src else pretty_items ctxt prts; (* antiquotation variants *) local fun gen_setup name embedded = if embedded then Document_Antiquotation.setup_embedded name else Document_Antiquotation.setup name; fun gen_antiquotation_pretty name embedded scan f = gen_setup name embedded scan (fn {context = ctxt, argument = x, ...} => pretty ctxt (f ctxt x)); fun gen_antiquotation_pretty_source name embedded scan f = gen_setup name embedded scan (fn {context = ctxt, source = src, argument = x} => pretty_source ctxt {embedded = embedded} src (f ctxt x)); fun gen_antiquotation_raw name embedded scan f = gen_setup name embedded scan (fn {context = ctxt, argument = x, ...} => f ctxt x); fun gen_antiquotation_verbatim name embedded scan f = gen_antiquotation_raw name embedded scan (fn ctxt => verbatim ctxt o f ctxt); in fun antiquotation_pretty name = gen_antiquotation_pretty name false; fun antiquotation_pretty_embedded name = gen_antiquotation_pretty name true; fun antiquotation_pretty_source name = gen_antiquotation_pretty_source name false; fun antiquotation_pretty_source_embedded name = gen_antiquotation_pretty_source name true; fun antiquotation_raw name = gen_antiquotation_raw name false; fun antiquotation_raw_embedded name = gen_antiquotation_raw name true; fun antiquotation_verbatim name = gen_antiquotation_verbatim name false; fun antiquotation_verbatim_embedded name = gen_antiquotation_verbatim name true; end; end; diff --git a/src/Pure/Thy/latex.ML b/src/Pure/Thy/latex.ML --- a/src/Pure/Thy/latex.ML +++ b/src/Pure/Thy/latex.ML @@ -1,248 +1,242 @@ (* Title: Pure/Thy/latex.ML Author: Makarius LaTeX output of theory sources. *) signature LATEX = sig type text = XML.body val text: string * Position.T -> text val string: string -> text val block: text -> XML.tree val output: text -> text val macro0: string -> text val macro: string -> text -> text val environment: string -> text -> text val output_name: string -> string val output_ascii: string -> string val output_ascii_breakable: string -> string -> string val output_symbols: Symbol.symbol list -> string val output_syms: string -> string val symbols: Symbol_Pos.T list -> text val symbols_output: Symbol_Pos.T list -> text - val environment_text: string -> text -> text val isabelle_body: string -> text -> text val theory_entry: string -> string type index_item = {text: text, like: string} type index_entry = {items: index_item list, def: bool} val index_entry: index_entry -> text val index_variants: (binding -> bool option -> 'a -> 'a) -> binding -> 'a -> 'a val latexN: string val latex_output: string -> string * int val latex_markup: string * Properties.T -> Markup.output val latex_indent: string -> int -> string end; structure Latex: LATEX = struct (* text with positions *) type text = XML.body; fun text (s, pos) = if s = "" then [] else if pos = Position.none then [XML.Text s] else [XML.Elem (Position.markup pos Markup.document_latex, [XML.Text s])]; fun string s = text (s, Position.none); fun block body = XML.Elem (Markup.document_latex, body); fun output body = [XML.Elem (Markup.latex_output, body)]; fun macro0 name = [XML.Elem (Markup.latex_macro0 name, [])]; fun macro name body = [XML.Elem (Markup.latex_macro name, body)]; fun environment name body = [XML.Elem (Markup.latex_environment name, body)]; (* output name for LaTeX macros *) val output_name = translate_string (fn "_" => "UNDERSCORE" | "'" => "PRIME" | "0" => "ZERO" | "1" => "ONE" | "2" => "TWO" | "3" => "THREE" | "4" => "FOUR" | "5" => "FIVE" | "6" => "SIX" | "7" => "SEVEN" | "8" => "EIGHT" | "9" => "NINE" | s => s); fun enclose_name bg en = enclose bg en o output_name; (* output verbatim ASCII *) val output_ascii = translate_string (fn " " => "\\ " | "\t" => "\\ " | "\n" => "\\isanewline\n" | s => s |> exists_string (fn s' => s = s') "\"#$%&',-<>\\^_`{}~" ? enclose "{\\char`\\" "}" |> suffix "{\\kern0pt}"); fun output_ascii_breakable sep = space_explode sep #> map output_ascii #> space_implode (output_ascii sep ^ "\\discretionary{}{}{}"); (* output symbols *) local val char_table = Symtab.make [("\007", "{\\isacharbell}"), ("!", "{\\isacharbang}"), ("\"", "{\\isachardoublequote}"), ("#", "{\\isacharhash}"), ("$", "{\\isachardollar}"), ("%", "{\\isacharpercent}"), ("&", "{\\isacharampersand}"), ("'", "{\\isacharprime}"), ("(", "{\\isacharparenleft}"), (")", "{\\isacharparenright}"), ("*", "{\\isacharasterisk}"), ("+", "{\\isacharplus}"), (",", "{\\isacharcomma}"), ("-", "{\\isacharminus}"), (".", "{\\isachardot}"), ("/", "{\\isacharslash}"), (":", "{\\isacharcolon}"), (";", "{\\isacharsemicolon}"), ("<", "{\\isacharless}"), ("=", "{\\isacharequal}"), (">", "{\\isachargreater}"), ("?", "{\\isacharquery}"), ("@", "{\\isacharat}"), ("[", "{\\isacharbrackleft}"), ("\\", "{\\isacharbackslash}"), ("]", "{\\isacharbrackright}"), ("^", "{\\isacharcircum}"), ("_", "{\\isacharunderscore}"), ("`", "{\\isacharbackquote}"), ("{", "{\\isacharbraceleft}"), ("|", "{\\isacharbar}"), ("}", "{\\isacharbraceright}"), ("~", "{\\isachartilde}")]; fun output_chr " " = "\\ " | output_chr "\t" = "\\ " | output_chr "\n" = "\\isanewline\n" | output_chr c = (case Symtab.lookup char_table c of SOME s => s ^ "{\\kern0pt}" | NONE => if Symbol.is_ascii_digit c then enclose "{\\isadigit{" "}}" c else c); fun output_sym sym = (case Symbol.decode sym of Symbol.Char s => output_chr s | Symbol.UTF8 s => s | Symbol.Sym s => enclose_name "{\\isasym" "}" s | Symbol.Control s => enclose_name "\\isactrl" " " s | Symbol.Malformed s => error (Symbol.malformed_msg s) | Symbol.EOF => error "Bad EOF symbol"); open Basic_Symbol_Pos; val scan_latex_length = Scan.many1 (fn (s, _) => s <> Symbol.latex andalso Symbol.not_eof s) >> (Symbol.length o map Symbol_Pos.symbol) || $$ Symbol.latex -- Scan.option (Scan.permissive Symbol_Pos.scan_cartouche "") >> K 0; val scan_latex = $$ Symbol.latex |-- Symbol_Pos.scan_cartouche_content "Embedded LaTeX: " >> (implode o map Symbol_Pos.symbol) || Scan.one (Symbol.not_eof o Symbol_Pos.symbol) >> (output_sym o Symbol_Pos.symbol); fun read scan syms = Scan.read Symbol_Pos.stopper (Scan.repeat scan) (map (rpair Position.none) syms); in fun length_symbols syms = fold Integer.add (these (read scan_latex_length syms)) 0; fun output_symbols syms = if member (op =) syms Symbol.latex then (case read scan_latex syms of SOME ss => implode ss | NONE => error ("Malformed embedded LaTeX: " ^ quote (Symbol.beginning 10 syms))) else implode (map output_sym syms); val output_syms = output_symbols o Symbol.explode; end; fun symbols syms = text (Symbol_Pos.content syms, #1 (Symbol_Pos.range syms)); fun symbols_output syms = text (output_symbols (map Symbol_Pos.symbol syms), #1 (Symbol_Pos.range syms)); (* theory presentation *) -fun environment_text name = - XML.enclose - ("%\n\\begin{" ^ output_name name ^ "}%\n") - ("%\n\\end{" ^ output_name name ^ "}"); - fun isabelle_body name = XML.enclose ("%\n\\begin{isabellebody}%\n\\setisabellecontext{" ^ output_syms name ^ "}%\n") "%\n\\end{isabellebody}%\n"; fun theory_entry name = "\\input{" ^ name ^ ".tex}\n\n"; (* index entries *) type index_item = {text: text, like: string}; type index_entry = {items: index_item list, def: bool}; fun index_item (item: index_item) = XML.wrap_elem ((Markup.latex_index_item, #text item), XML.string (#like item)); fun index_entry (entry: index_entry) = [XML.Elem (Markup.latex_index_entry (if #def entry then "isaindexdef" else "isaindexref"), map index_item (#items entry))]; fun index_binding NONE = I | index_binding (SOME def) = Binding.map_name (suffix (if def then "_def" else "_ref")); fun index_variants setup binding = fold (fn index => setup (index_binding index binding) index) [NONE, SOME true, SOME false]; (* print mode *) val latexN = "latex"; fun latex_output str = let val syms = Symbol.explode str in (output_symbols syms, length_symbols syms) end; fun latex_markup (s, _: Properties.T) = if s = Markup.commandN orelse s = Markup.keyword1N orelse s = Markup.keyword3N then ("\\isacommand{", "}") else if s = Markup.keyword2N then ("\\isakeyword{", "}") else Markup.no_output; fun latex_indent "" _ = "" | latex_indent s _ = enclose "\\isaindent{" "}" s; val _ = Output.add_mode latexN latex_output (prefix Symbol.latex o cartouche); val _ = Markup.add_mode latexN latex_markup; val _ = Pretty.add_mode latexN latex_indent; end; diff --git a/src/Pure/Tools/rail.ML b/src/Pure/Tools/rail.ML --- a/src/Pure/Tools/rail.ML +++ b/src/Pure/Tools/rail.ML @@ -1,395 +1,395 @@ (* Title: Pure/Tools/rail.ML Author: Michael Kerscher, TU München Author: Makarius Railroad diagrams in LaTeX. *) signature RAIL = sig datatype rails = Cat of int * rail list and rail = Bar of rails list | Plus of rails * rails | Newline of int | Nonterminal of string | Terminal of bool * string | Antiquote of bool * Antiquote.antiq val read: Proof.context -> Input.source -> (string Antiquote.antiquote * rail) list val output_rules: Proof.context -> (string Antiquote.antiquote * rail) list -> Latex.text end; structure Rail: RAIL = struct (** lexical syntax **) (* singleton keywords *) val keywords = Symtab.make [ ("|", Markup.keyword3), ("*", Markup.keyword3), ("+", Markup.keyword3), ("?", Markup.keyword3), ("(", Markup.empty), (")", Markup.empty), ("\", Markup.keyword2), (";", Markup.keyword2), (":", Markup.keyword2), ("@", Markup.keyword1)]; (* datatype token *) datatype kind = Keyword | Ident | String | Space | Comment of Comment.kind | Antiq of Antiquote.antiq | EOF; datatype token = Token of Position.range * (kind * string); fun pos_of (Token ((pos, _), _)) = pos; fun end_pos_of (Token ((_, pos), _)) = pos; fun range_of (toks as tok :: _) = let val pos' = end_pos_of (List.last toks) in Position.range (pos_of tok, pos') end | range_of [] = Position.no_range; fun kind_of (Token (_, (k, _))) = k; fun content_of (Token (_, (_, x))) = x; fun is_proper (Token (_, (Space, _))) = false | is_proper (Token (_, (Comment _, _))) = false | is_proper _ = true; (* diagnostics *) val print_kind = fn Keyword => "rail keyword" | Ident => "identifier" | String => "single-quoted string" | Space => "white space" | Comment _ => "formal comment" | Antiq _ => "antiquotation" | EOF => "end-of-input"; fun print (Token ((pos, _), (k, x))) = (if k = EOF then print_kind k else print_kind k ^ " " ^ quote x) ^ Position.here pos; fun print_keyword x = print_kind Keyword ^ " " ^ quote x; fun reports_of_token (Token ((pos, _), (Keyword, x))) = map (pair pos) (the_list (Symtab.lookup keywords x) @ Completion.suppress_abbrevs x) | reports_of_token (Token ((pos, _), (String, _))) = [(pos, Markup.inner_string)] | reports_of_token (Token (_, (Antiq antiq, _))) = Antiquote.antiq_reports [Antiquote.Antiq antiq] | reports_of_token _ = []; (* stopper *) fun mk_eof pos = Token ((pos, Position.none), (EOF, "")); val eof = mk_eof Position.none; fun is_eof (Token (_, (EOF, _))) = true | is_eof _ = false; val stopper = Scan.stopper (fn [] => eof | toks => mk_eof (end_pos_of (List.last toks))) is_eof; (* tokenize *) local fun token k ss = [Token (Symbol_Pos.range ss, (k, Symbol_Pos.content ss))]; fun antiq_token antiq = [Token (#range antiq, (Antiq antiq, Symbol_Pos.content (#body antiq)))]; val scan_space = Scan.many1 (Symbol.is_blank o Symbol_Pos.symbol); val scan_keyword = Scan.one (Symtab.defined keywords o Symbol_Pos.symbol); val err_prefix = "Rail lexical error: "; val scan_token = scan_space >> token Space || Comment.scan_inner >> (fn (kind, ss) => token (Comment kind) ss) || Antiquote.scan_antiq >> antiq_token || scan_keyword >> (token Keyword o single) || Lexicon.scan_id >> token Ident || Symbol_Pos.scan_string_q err_prefix >> (fn (pos1, (ss, pos2)) => [Token (Position.range (pos1, pos2), (String, Symbol_Pos.content ss))]); val scan = Scan.repeats scan_token --| Symbol_Pos.!!! (fn () => err_prefix ^ "bad input") (Scan.ahead (Scan.one Symbol_Pos.is_eof)); in val tokenize = #1 o Scan.error (Scan.finite Symbol_Pos.stopper scan); end; (** parsing **) (* parser combinators *) fun !!! scan = let val prefix = "Rail syntax error"; fun get_pos [] = " (end-of-input)" | get_pos (tok :: _) = Position.here (pos_of tok); fun err (toks, NONE) = (fn () => prefix ^ get_pos toks) | err (toks, SOME msg) = (fn () => let val s = msg () in if String.isPrefix prefix s then s else prefix ^ get_pos toks ^ ": " ^ s end); in Scan.!! err scan end; fun $$$ x = Scan.one (fn tok => kind_of tok = Keyword andalso content_of tok = x) || Scan.fail_with (fn [] => (fn () => print_keyword x ^ " expected,\nbut end-of-input was found") | tok :: _ => (fn () => print_keyword x ^ " expected,\nbut " ^ print tok ^ " was found")); fun enum1 sep scan = scan ::: Scan.repeat ($$$ sep |-- !!! scan); fun enum sep scan = enum1 sep scan || Scan.succeed []; val ident = Scan.some (fn tok => if kind_of tok = Ident then SOME (content_of tok) else NONE); val string = Scan.some (fn tok => if kind_of tok = String then SOME (content_of tok) else NONE); val antiq = Scan.some (fn tok => (case kind_of tok of Antiq a => SOME a | _ => NONE)); fun RANGE scan = Scan.trace scan >> apsnd range_of; fun RANGE_APP scan = RANGE scan >> (fn (f, r) => f r); (* parse trees *) datatype trees = CAT of tree list * Position.range and tree = BAR of trees list * Position.range | STAR of (trees * trees) * Position.range | PLUS of (trees * trees) * Position.range | MAYBE of tree * Position.range | NEWLINE of Position.range | NONTERMINAL of string * Position.range | TERMINAL of (bool * string) * Position.range | ANTIQUOTE of (bool * Antiquote.antiq) * Position.range; fun reports_of_tree ctxt = if Context_Position.reports_enabled ctxt then let fun reports r = if r = Position.no_range then [] else [(Position.range_position r, Markup.expression "")]; fun trees (CAT (ts, r)) = reports r @ maps tree ts and tree (BAR (Ts, r)) = reports r @ maps trees Ts | tree (STAR ((T1, T2), r)) = reports r @ trees T1 @ trees T2 | tree (PLUS ((T1, T2), r)) = reports r @ trees T1 @ trees T2 | tree (MAYBE (t, r)) = reports r @ tree t | tree (NEWLINE r) = reports r | tree (NONTERMINAL (_, r)) = reports r | tree (TERMINAL (_, r)) = reports r | tree (ANTIQUOTE (_, r)) = reports r; in distinct (op =) o tree end else K []; local val at_mode = Scan.option ($$$ "@") >> (fn NONE => false | _ => true); fun body x = (RANGE (enum1 "|" body1) >> BAR) x and body0 x = (RANGE (enum "|" body1) >> BAR) x and body1 x = (RANGE_APP (body2 :|-- (fn a => $$$ "*" |-- !!! body4e >> (fn b => fn r => CAT ([STAR ((a, b), r)], r)) || $$$ "+" |-- !!! body4e >> (fn b => fn r => CAT ([PLUS ((a, b), r)], r)) || Scan.succeed (K a)))) x and body2 x = (RANGE (Scan.repeat1 body3) >> CAT) x and body3 x = (RANGE_APP (body4 :|-- (fn a => $$$ "?" >> K (curry MAYBE a) || Scan.succeed (K a)))) x and body4 x = ($$$ "(" |-- !!! (body0 --| $$$ ")") || RANGE_APP ($$$ "\" >> K NEWLINE || ident >> curry NONTERMINAL || at_mode -- string >> curry TERMINAL || at_mode -- antiq >> curry ANTIQUOTE)) x and body4e x = (RANGE (Scan.option body4) >> (fn (a, r) => CAT (the_list a, r))) x; val rule_name = ident >> Antiquote.Text || antiq >> Antiquote.Antiq; val rule = rule_name -- ($$$ ":" |-- !!! body) || body >> pair (Antiquote.Text ""); val rules = enum1 ";" (Scan.option rule) >> map_filter I; in fun parse_rules toks = #1 (Scan.error (Scan.finite stopper (rules --| !!! (Scan.ahead (Scan.one is_eof)))) toks); end; (** rail expressions **) (* datatype *) datatype rails = Cat of int * rail list and rail = Bar of rails list | Plus of rails * rails | Newline of int | Nonterminal of string | Terminal of bool * string | Antiquote of bool * Antiquote.antiq; fun is_newline (Newline _) = true | is_newline _ = false; (* prepare *) local fun cat rails = Cat (0, rails); val empty = cat []; fun is_empty (Cat (_, [])) = true | is_empty _ = false; fun bar [Cat (_, [rail])] = rail | bar cats = Bar cats; fun reverse_cat (Cat (y, rails)) = Cat (y, rev (map reverse rails)) and reverse (Bar cats) = Bar (map reverse_cat cats) | reverse (Plus (cat1, cat2)) = Plus (reverse_cat cat1, reverse_cat cat2) | reverse x = x; fun plus (cat1, cat2) = Plus (cat1, reverse_cat cat2); in fun prepare_trees (CAT (ts, _)) = Cat (0, map prepare_tree ts) and prepare_tree (BAR (Ts, _)) = bar (map prepare_trees Ts) | prepare_tree (STAR (Ts, _)) = let val (cat1, cat2) = apply2 prepare_trees Ts in if is_empty cat2 then plus (empty, cat1) else bar [empty, cat [plus (cat1, cat2)]] end | prepare_tree (PLUS (Ts, _)) = plus (apply2 prepare_trees Ts) | prepare_tree (MAYBE (t, _)) = bar [empty, cat [prepare_tree t]] | prepare_tree (NEWLINE _) = Newline 0 | prepare_tree (NONTERMINAL (a, _)) = Nonterminal a | prepare_tree (TERMINAL (a, _)) = Terminal a | prepare_tree (ANTIQUOTE (a, _)) = Antiquote a; end; (* read *) fun read ctxt source = let val _ = Context_Position.report ctxt (Input.pos_of source) Markup.language_rail; val toks = tokenize (Input.source_explode source); val _ = Context_Position.reports ctxt (maps reports_of_token toks); val rules = parse_rules (filter is_proper toks); val _ = Context_Position.reports ctxt (maps (reports_of_tree ctxt o #2) rules); in map (apsnd prepare_tree) rules end; (* latex output *) local fun vertical_range_cat (Cat (_, rails)) y = let val (rails', (_, y')) = fold_map (fn rail => fn (y0, y') => if is_newline rail then (Newline (y' + 1), (y' + 1, y' + 2)) else let val (rail', y0') = vertical_range rail y0; in (rail', (y0, Int.max (y0', y'))) end) rails (y, y + 1) in (Cat (y, rails'), y') end and vertical_range (Bar cats) y = let val (cats', y') = fold_map vertical_range_cat cats y in (Bar cats', Int.max (y + 1, y')) end | vertical_range (Plus (cat1, cat2)) y = let val ([cat1', cat2'], y') = fold_map vertical_range_cat [cat1, cat2] y; in (Plus (cat1', cat2'), Int.max (y + 1, y')) end | vertical_range (Newline _) y = (Newline (y + 2), y + 3) | vertical_range atom y = (atom, y + 1); in fun output_rules ctxt rules = let val output_antiq = Antiquote.Antiq #> Document_Antiquotation.evaluate Latex.symbols ctxt; fun output_text b s = Latex.string (Output.output s) |> b ? Latex.macro "isakeyword" |> Latex.macro "isa"; fun output_cat c (Cat (_, rails)) = outputs c rails and outputs c [rail] = output c rail | outputs _ rails = maps (output "") rails and output _ (Bar []) = [] | output c (Bar [cat]) = output_cat c cat | output _ (Bar (cat :: cats)) = Latex.string ("\\rail@bar\n") @ output_cat "" cat @ maps (fn Cat (y, rails) => Latex.string ("\\rail@nextbar{" ^ string_of_int y ^ "}\n") @ outputs "" rails) cats @ Latex.string "\\rail@endbar\n" | output c (Plus (cat, Cat (y, rails))) = Latex.string "\\rail@plus\n" @ output_cat c cat @ Latex.string ("\\rail@nextplus{" ^ string_of_int y ^ "}\n") @ outputs "c" rails @ Latex.string "\\rail@endplus\n" | output _ (Newline y) = Latex.string ("\\rail@cr{" ^ string_of_int y ^ "}\n") | output c (Nonterminal s) = Latex.string ("\\rail@" ^ c ^ "nont{") @ output_text false s @ Latex.string "}[]\n" | output c (Terminal (b, s)) = Latex.string ("\\rail@" ^ c ^ "term{") @ output_text b s @ Latex.string "}[]\n" | output c (Antiquote (b, a)) = Latex.string ("\\rail@" ^ c ^ (if b then "term{" else "nont{")) @ Latex.output (output_antiq a) @ Latex.string "}[]\n"; fun output_rule (name, rail) = let val (rail', y') = vertical_range rail 0; val out_name = (case name of Antiquote.Text "" => [] | Antiquote.Text s => output_text false s | Antiquote.Antiq a => output_antiq a); in Latex.string ("\\rail@begin{" ^ string_of_int y' ^ "}{") @ out_name @ Latex.string "}\n" @ output "" rail' @ Latex.string "\\rail@end\n" end; - in Latex.environment_text "railoutput" (maps output_rule rules) end; + in Latex.environment "railoutput" (maps output_rule rules) end; val _ = Theory.setup (Document_Output.antiquotation_raw_embedded \<^binding>\rail\ (Scan.lift Args.text_input) (fn ctxt => output_rules ctxt o read ctxt)); end; end;