src/Pure/Thy/thy_output.ML
author wenzelm
Tue, 22 Jan 2019 19:36:17 +0100
changeset 69719 331ef175a112
parent 69592 a80d8ec6c998
child 69876 b49bd228ac8a
permissions -rw-r--r--
Backed out changeset 1bc422c08209 -- obsolete in AFP/5d11846ac6ab;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Thy/thy_output.ML
67386
998e01d6f8fd clarified modules;
wenzelm
parents: 67381
diff changeset
     2
    Author:     Makarius
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
     3
67386
998e01d6f8fd clarified modules;
wenzelm
parents: 67381
diff changeset
     4
Theory document output.
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
     5
*)
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
     6
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
     7
signature THY_OUTPUT =
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
     8
sig
67569
5d71b114e7a3 avoid proliferation of language_document reports;
wenzelm
parents: 67567
diff changeset
     9
  val output_document: Proof.context -> {markdown: bool} -> Input.source -> Latex.text list
67381
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67378
diff changeset
    10
  val check_comments: Proof.context -> Symbol_Pos.T list -> unit
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67378
diff changeset
    11
  val output_token: Proof.context -> Token.T -> Latex.text list
67474
90def2b06305 more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents: 67465
diff changeset
    12
  val output_source: Proof.context -> string -> Latex.text list
68182
fa3cf61121ee tuned signature (see Command.eval_state);
wenzelm
parents: 68178
diff changeset
    13
  type segment = {span: Command_Span.span, command: Toplevel.transition, state: Toplevel.state}
68506
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
    14
  val present_thy: Options.T -> theory -> segment list -> Latex.text list
28644
e2ae4a6cf166 ThyOutput: export some auxiliary operations;
wenzelm
parents: 28427
diff changeset
    15
  val pretty_term: Proof.context -> term -> Pretty.T
e2ae4a6cf166 ThyOutput: export some auxiliary operations;
wenzelm
parents: 28427
diff changeset
    16
  val pretty_thm: Proof.context -> thm -> Pretty.T
67508
189ab2c3026b verbatim output consists of plain lines;
wenzelm
parents: 67506
diff changeset
    17
  val lines: Latex.text list -> Latex.text list
67505
ceb324e34c14 clarified signature: items with \isasep are special;
wenzelm
parents: 67475
diff changeset
    18
  val items: Latex.text list -> Latex.text list
67463
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
    19
  val isabelle: Proof.context -> Latex.text list -> Latex.text
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
    20
  val isabelle_typewriter: Proof.context -> Latex.text list -> Latex.text
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
    21
  val typewriter: Proof.context -> string -> Latex.text
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
    22
  val verbatim: Proof.context -> string -> Latex.text
67474
90def2b06305 more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents: 67465
diff changeset
    23
  val source: Proof.context -> Token.src -> Latex.text
67505
ceb324e34c14 clarified signature: items with \isasep are special;
wenzelm
parents: 67475
diff changeset
    24
  val pretty: Proof.context -> Pretty.T -> Latex.text
ceb324e34c14 clarified signature: items with \isasep are special;
wenzelm
parents: 67475
diff changeset
    25
  val pretty_source: Proof.context -> Token.src -> Pretty.T -> Latex.text
ceb324e34c14 clarified signature: items with \isasep are special;
wenzelm
parents: 67475
diff changeset
    26
  val pretty_items: Proof.context -> Pretty.T list -> Latex.text
ceb324e34c14 clarified signature: items with \isasep are special;
wenzelm
parents: 67475
diff changeset
    27
  val pretty_items_source: Proof.context -> Token.src -> Pretty.T list -> Latex.text
67463
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
    28
  val antiquotation_pretty:
67505
ceb324e34c14 clarified signature: items with \isasep are special;
wenzelm
parents: 67475
diff changeset
    29
    binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory
69592
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
    30
  val antiquotation_pretty_embedded:
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
    31
    binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory
67463
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
    32
  val antiquotation_pretty_source:
67505
ceb324e34c14 clarified signature: items with \isasep are special;
wenzelm
parents: 67475
diff changeset
    33
    binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory
69592
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
    34
  val antiquotation_pretty_source_embedded:
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
    35
    binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory
67463
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
    36
  val antiquotation_raw:
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
    37
    binding -> 'a context_parser -> (Proof.context -> 'a -> Latex.text) -> theory -> theory
69592
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
    38
  val antiquotation_raw_embedded:
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
    39
    binding -> 'a context_parser -> (Proof.context -> 'a -> Latex.text) -> theory -> theory
67463
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
    40
  val antiquotation_verbatim:
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
    41
    binding -> 'a context_parser -> (Proof.context -> 'a -> string) -> theory -> theory
69592
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
    42
  val antiquotation_verbatim_embedded:
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
    43
    binding -> 'a context_parser -> (Proof.context -> 'a -> string) -> theory -> theory
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
    44
end;
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
    45
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 37198
diff changeset
    46
structure Thy_Output: THY_OUTPUT =
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
    47
struct
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
    48
67569
5d71b114e7a3 avoid proliferation of language_document reports;
wenzelm
parents: 67567
diff changeset
    49
(* output document source *)
56548
ae6870efc28d markup for prose words within formal comments;
wenzelm
parents: 56438
diff changeset
    50
67571
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    51
val output_symbols = single o Latex.symbols_output;
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    52
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    53
fun output_comment ctxt (kind, syms) =
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    54
  (case kind of
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    55
    Comment.Comment =>
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    56
      Input.cartouche_content syms
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    57
      |> output_document (ctxt |> Config.put Document_Antiquotation.thy_output_display false)
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    58
          {markdown = false}
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    59
      |> Latex.enclose_body "%\n\\isamarkupcmt{" "%\n}"
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    60
  | Comment.Cancel =>
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    61
      Symbol_Pos.cartouche_content syms
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    62
      |> output_symbols
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    63
      |> Latex.enclose_body "%\n\\isamarkupcancel{" "}"
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    64
  | Comment.Latex =>
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    65
      [Latex.symbols (Symbol_Pos.cartouche_content syms)])
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    66
and output_comment_document ctxt (comment, syms) =
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    67
  (case comment of
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    68
    SOME kind => output_comment ctxt (kind, syms)
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    69
  | NONE => [Latex.symbols syms])
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    70
and output_document_text ctxt syms =
67572
a93cf1d6ba87 clarified signature;
wenzelm
parents: 67571
diff changeset
    71
  Comment.read_body syms |> maps (output_comment_document ctxt)
67571
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    72
and output_document ctxt {markdown} source =
61457
3e21699bb83b clarified Antiquote.antiq_reports;
wenzelm
parents: 61456
diff changeset
    73
  let
3e21699bb83b clarified Antiquote.antiq_reports;
wenzelm
parents: 61456
diff changeset
    74
    val pos = Input.pos_of source;
3e21699bb83b clarified Antiquote.antiq_reports;
wenzelm
parents: 61456
diff changeset
    75
    val syms = Input.source_explode source;
3e21699bb83b clarified Antiquote.antiq_reports;
wenzelm
parents: 61456
diff changeset
    76
67571
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    77
    val output_antiquotes =
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    78
      maps (Document_Antiquotation.evaluate (output_document_text ctxt) ctxt);
61036
f6f2959bed67 clarified language context, e.g. relevant for symbols;
wenzelm
parents: 60100
diff changeset
    79
61461
77c9643a6353 more explicit output of list items;
wenzelm
parents: 61459
diff changeset
    80
    fun output_line line =
67194
1c0a6a957114 positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents: 67191
diff changeset
    81
      (if Markdown.line_is_item line then [Latex.string "\\item "] else []) @
61461
77c9643a6353 more explicit output of list items;
wenzelm
parents: 61459
diff changeset
    82
        output_antiquotes (Markdown.line_content line);
77c9643a6353 more explicit output of list items;
wenzelm
parents: 61459
diff changeset
    83
67194
1c0a6a957114 positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents: 67191
diff changeset
    84
    fun output_block (Markdown.Par lines) =
1c0a6a957114 positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents: 67191
diff changeset
    85
          Latex.block (separate (Latex.string "\n") (map (Latex.block o output_line) lines))
61459
5f2ddeb15c06 clarified nesting of paragraphs: indentation is taken into account more uniformly;
wenzelm
parents: 61458
diff changeset
    86
      | output_block (Markdown.List {kind, body, ...}) =
67194
1c0a6a957114 positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents: 67191
diff changeset
    87
          Latex.environment_block (Markdown.print_kind kind) (output_blocks body)
1c0a6a957114 positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents: 67191
diff changeset
    88
    and output_blocks blocks = separate (Latex.string "\n\n") (map output_block blocks);
67184
ecc786cb3b7b more robust range on preceding comment-line;
wenzelm
parents: 67173
diff changeset
    89
  in
67381
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67378
diff changeset
    90
    if Toplevel.is_skipped_proof (Toplevel.presentation_state ctxt) then []
67184
ecc786cb3b7b more robust range on preceding comment-line;
wenzelm
parents: 67173
diff changeset
    91
    else if markdown andalso exists (Markdown.is_control o Symbol_Pos.symbol) syms
ecc786cb3b7b more robust range on preceding comment-line;
wenzelm
parents: 67173
diff changeset
    92
    then
ecc786cb3b7b more robust range on preceding comment-line;
wenzelm
parents: 67173
diff changeset
    93
      let
67571
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
    94
        val ants = Antiquote.parse_comments pos syms;
67184
ecc786cb3b7b more robust range on preceding comment-line;
wenzelm
parents: 67173
diff changeset
    95
        val reports = Antiquote.antiq_reports ants;
ecc786cb3b7b more robust range on preceding comment-line;
wenzelm
parents: 67173
diff changeset
    96
        val blocks = Markdown.read_antiquotes ants;
67381
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67378
diff changeset
    97
        val _ = Context_Position.reports ctxt (reports @ Markdown.reports blocks);
67186
a58bbe66ac81 avoid excessive whitespace between antiquotations and text;
wenzelm
parents: 67184
diff changeset
    98
      in output_blocks blocks end
67184
ecc786cb3b7b more robust range on preceding comment-line;
wenzelm
parents: 67173
diff changeset
    99
    else
ecc786cb3b7b more robust range on preceding comment-line;
wenzelm
parents: 67173
diff changeset
   100
      let
67571
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
   101
        val ants = Antiquote.parse_comments pos (trim (Symbol.is_blank o Symbol_Pos.symbol) syms);
67184
ecc786cb3b7b more robust range on preceding comment-line;
wenzelm
parents: 67173
diff changeset
   102
        val reports = Antiquote.antiq_reports ants;
67381
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67378
diff changeset
   103
        val _ = Context_Position.reports ctxt (reports @ Markdown.text_reports ants);
67186
a58bbe66ac81 avoid excessive whitespace between antiquotations and text;
wenzelm
parents: 67184
diff changeset
   104
      in output_antiquotes ants end
67184
ecc786cb3b7b more robust range on preceding comment-line;
wenzelm
parents: 67173
diff changeset
   105
  end;
48918
6e5fd4585512 check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
wenzelm
parents: 47005
diff changeset
   106
6e5fd4585512 check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
wenzelm
parents: 47005
diff changeset
   107
67425
7d4a088dbc0e clarified modules: uniform notion of formal comments;
wenzelm
parents: 67421
diff changeset
   108
(* output tokens with formal comments *)
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   109
67353
95f5f4bec7af clarified modules;
wenzelm
parents: 67194
diff changeset
   110
local
95f5f4bec7af clarified modules;
wenzelm
parents: 67194
diff changeset
   111
67358
wenzelm
parents: 67357
diff changeset
   112
val output_symbols_antiq =
wenzelm
parents: 67357
diff changeset
   113
  (fn Antiquote.Text syms => output_symbols syms
67353
95f5f4bec7af clarified modules;
wenzelm
parents: 67194
diff changeset
   114
    | Antiquote.Control {name = (name, _), body, ...} =>
67374
5a049cf98438 clarified output (see also 909dcdec2122, 34d1913f0b20);
wenzelm
parents: 67372
diff changeset
   115
        Latex.string (Latex.output_symbols [Symbol.encode (Symbol.Control name)]) ::
67358
wenzelm
parents: 67357
diff changeset
   116
          output_symbols body
67353
95f5f4bec7af clarified modules;
wenzelm
parents: 67194
diff changeset
   117
    | Antiquote.Antiq {body, ...} =>
67358
wenzelm
parents: 67357
diff changeset
   118
        Latex.enclose_body "%\n\\isaantiq\n" "{}%\n\\endisaantiq\n" (output_symbols body));
67353
95f5f4bec7af clarified modules;
wenzelm
parents: 67194
diff changeset
   119
67571
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
   120
fun output_comment_symbols ctxt {antiq} (comment, syms) =
67425
7d4a088dbc0e clarified modules: uniform notion of formal comments;
wenzelm
parents: 67421
diff changeset
   121
  (case (comment, antiq) of
7d4a088dbc0e clarified modules: uniform notion of formal comments;
wenzelm
parents: 67421
diff changeset
   122
    (NONE, false) => output_symbols syms
7d4a088dbc0e clarified modules: uniform notion of formal comments;
wenzelm
parents: 67421
diff changeset
   123
  | (NONE, true) =>
67571
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
   124
      Antiquote.parse_comments (#1 (Symbol_Pos.range syms)) syms
67413
2555713586c8 added \<^cancel> operator for unused text;
wenzelm
parents: 67393
diff changeset
   125
      |> maps output_symbols_antiq
67571
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
   126
  | (SOME comment, _) => output_comment ctxt (comment, syms));
67372
wenzelm
parents: 67360
diff changeset
   127
67381
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67378
diff changeset
   128
fun output_body ctxt antiq bg en syms =
67572
a93cf1d6ba87 clarified signature;
wenzelm
parents: 67571
diff changeset
   129
  Comment.read_body syms
a93cf1d6ba87 clarified signature;
wenzelm
parents: 67571
diff changeset
   130
  |> maps (output_comment_symbols ctxt {antiq = antiq})
68194
wenzelm
parents: 68182
diff changeset
   131
  |> Latex.enclose_body bg en;
wenzelm
parents: 68182
diff changeset
   132
wenzelm
parents: 68182
diff changeset
   133
in
wenzelm
parents: 68182
diff changeset
   134
wenzelm
parents: 68182
diff changeset
   135
fun output_token ctxt tok =
67353
95f5f4bec7af clarified modules;
wenzelm
parents: 67194
diff changeset
   136
  let
67378
wenzelm
parents: 67377
diff changeset
   137
    fun output antiq bg en =
67381
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67378
diff changeset
   138
      output_body ctxt antiq bg en (Input.source_explode (Token.input_of tok));
67378
wenzelm
parents: 67377
diff changeset
   139
  in
wenzelm
parents: 67377
diff changeset
   140
    (case Token.kind_of tok of
67439
78759a7bd874 more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents: 67429
diff changeset
   141
      Token.Comment NONE => []
67378
wenzelm
parents: 67377
diff changeset
   142
    | Token.Command => output false "\\isacommand{" "}"
wenzelm
parents: 67377
diff changeset
   143
    | Token.Keyword =>
wenzelm
parents: 67377
diff changeset
   144
        if Symbol.is_ascii_identifier (Token.content_of tok)
wenzelm
parents: 67377
diff changeset
   145
        then output false "\\isakeyword{" "}"
wenzelm
parents: 67377
diff changeset
   146
        else output false "" ""
wenzelm
parents: 67377
diff changeset
   147
    | Token.String => output false "{\\isachardoublequoteopen}" "{\\isachardoublequoteclose}"
wenzelm
parents: 67377
diff changeset
   148
    | Token.Alt_String => output false "{\\isacharbackquoteopen}" "{\\isacharbackquoteclose}"
wenzelm
parents: 67377
diff changeset
   149
    | Token.Verbatim => output true "{\\isacharverbatimopen}" "{\\isacharverbatimclose}"
wenzelm
parents: 67377
diff changeset
   150
    | Token.Cartouche => output false "{\\isacartoucheopen}" "{\\isacartoucheclose}"
wenzelm
parents: 67377
diff changeset
   151
    | _ => output false "" "")
wenzelm
parents: 67377
diff changeset
   152
  end handle ERROR msg => error (msg ^ Position.here (Token.pos_of tok));
67353
95f5f4bec7af clarified modules;
wenzelm
parents: 67194
diff changeset
   153
67474
90def2b06305 more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents: 67465
diff changeset
   154
fun output_source ctxt s =
90def2b06305 more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents: 67465
diff changeset
   155
  output_body ctxt false "" "" (Symbol_Pos.explode (s, Position.none));
90def2b06305 more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents: 67465
diff changeset
   156
67381
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67378
diff changeset
   157
fun check_comments ctxt =
67572
a93cf1d6ba87 clarified signature;
wenzelm
parents: 67571
diff changeset
   158
  Comment.read_body #> List.app (fn (comment, syms) =>
67388
5fc0b0c9a735 clarified markup: more like outer syntax side-comment;
wenzelm
parents: 67386
diff changeset
   159
    let
5fc0b0c9a735 clarified markup: more like outer syntax side-comment;
wenzelm
parents: 67386
diff changeset
   160
      val pos = #1 (Symbol_Pos.range syms);
5fc0b0c9a735 clarified markup: more like outer syntax side-comment;
wenzelm
parents: 67386
diff changeset
   161
      val _ =
67429
95877cc6630e allow LaTeX source as formal comment;
wenzelm
parents: 67425
diff changeset
   162
        comment |> Option.app (fn kind =>
67506
30233285270a more markup: disable spell-checker for raw latex;
wenzelm
parents: 67505
diff changeset
   163
          Context_Position.reports ctxt (map (pair pos) (Markup.cartouche :: Comment.markups kind)));
67571
f858fe5531ac more uniform treatment of formal comments within document source;
wenzelm
parents: 67570
diff changeset
   164
      val _ = output_comment_symbols ctxt {antiq = false} (comment, syms);
67425
7d4a088dbc0e clarified modules: uniform notion of formal comments;
wenzelm
parents: 67421
diff changeset
   165
    in if comment = SOME Comment.Comment then check_comments ctxt syms else () end);
67377
143665524d8e check formal comments recursively, within arbitrary cartouches (unknown sublanguages);
wenzelm
parents: 67375
diff changeset
   166
67353
95f5f4bec7af clarified modules;
wenzelm
parents: 67194
diff changeset
   167
end;
95f5f4bec7af clarified modules;
wenzelm
parents: 67194
diff changeset
   168
95f5f4bec7af clarified modules;
wenzelm
parents: 67194
diff changeset
   169
67356
ba226b87c69e output token content with formal comments and antiquotations;
wenzelm
parents: 67354
diff changeset
   170
ba226b87c69e output token content with formal comments and antiquotations;
wenzelm
parents: 67354
diff changeset
   171
(** present theory source **)
ba226b87c69e output token content with formal comments and antiquotations;
wenzelm
parents: 67354
diff changeset
   172
ba226b87c69e output token content with formal comments and antiquotations;
wenzelm
parents: 67354
diff changeset
   173
(*NB: arranging white space around command spans is a black art*)
ba226b87c69e output token content with formal comments and antiquotations;
wenzelm
parents: 67354
diff changeset
   174
67439
78759a7bd874 more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents: 67429
diff changeset
   175
val is_white = Token.is_space orf Token.is_informal_comment;
78759a7bd874 more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents: 67429
diff changeset
   176
val is_black = not o is_white;
78759a7bd874 more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents: 67429
diff changeset
   177
78759a7bd874 more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents: 67429
diff changeset
   178
val is_white_comment = Token.is_informal_comment;
78759a7bd874 more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents: 67429
diff changeset
   179
val is_black_comment = Token.is_formal_comment;
78759a7bd874 more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents: 67429
diff changeset
   180
67356
ba226b87c69e output token content with formal comments and antiquotations;
wenzelm
parents: 67354
diff changeset
   181
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   182
(* presentation tokens *)
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   183
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   184
datatype token =
67360
wenzelm
parents: 67359
diff changeset
   185
    Ignore_Token
59065
8ce02aafc363 tuned signature -- prefer Input.source;
wenzelm
parents: 59064
diff changeset
   186
  | Basic_Token of Token.T
8ce02aafc363 tuned signature -- prefer Input.source;
wenzelm
parents: 59064
diff changeset
   187
  | Markup_Token of string * Input.source
8ce02aafc363 tuned signature -- prefer Input.source;
wenzelm
parents: 59064
diff changeset
   188
  | Markup_Env_Token of string * Input.source
61456
b521b8b400f7 trim_blanks after read, before eval;
wenzelm
parents: 61455
diff changeset
   189
  | Raw_Token of Input.source;
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   190
59065
8ce02aafc363 tuned signature -- prefer Input.source;
wenzelm
parents: 59064
diff changeset
   191
fun basic_token pred (Basic_Token tok) = pred tok
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   192
  | basic_token _ _ = false;
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   193
67439
78759a7bd874 more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents: 67429
diff changeset
   194
val white_token = basic_token is_white;
78759a7bd874 more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents: 67429
diff changeset
   195
val white_comment_token = basic_token is_white_comment;
36959
f5417836dbea renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents: 36950
diff changeset
   196
val blank_token = basic_token Token.is_blank;
f5417836dbea renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents: 36950
diff changeset
   197
val newline_token = basic_token Token.is_newline;
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   198
67381
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67378
diff changeset
   199
fun present_token ctxt tok =
61457
3e21699bb83b clarified Antiquote.antiq_reports;
wenzelm
parents: 61456
diff changeset
   200
  (case tok of
67360
wenzelm
parents: 67359
diff changeset
   201
    Ignore_Token => []
67381
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67378
diff changeset
   202
  | Basic_Token tok => output_token ctxt tok
67357
d7c6054b2ab1 clarified output: avoid extra space;
wenzelm
parents: 67356
diff changeset
   203
  | Markup_Token (cmd, source) =>
d7c6054b2ab1 clarified output: avoid extra space;
wenzelm
parents: 67356
diff changeset
   204
      Latex.enclose_body ("%\n\\isamarkup" ^ cmd ^ "{") "%\n}\n"
67569
5d71b114e7a3 avoid proliferation of language_document reports;
wenzelm
parents: 67567
diff changeset
   205
        (output_document ctxt {markdown = false} source)
61457
3e21699bb83b clarified Antiquote.antiq_reports;
wenzelm
parents: 61456
diff changeset
   206
  | Markup_Env_Token (cmd, source) =>
67569
5d71b114e7a3 avoid proliferation of language_document reports;
wenzelm
parents: 67567
diff changeset
   207
      [Latex.environment_block ("isamarkup" ^ cmd) (output_document ctxt {markdown = true} source)]
67186
a58bbe66ac81 avoid excessive whitespace between antiquotations and text;
wenzelm
parents: 67184
diff changeset
   208
  | Raw_Token source =>
67569
5d71b114e7a3 avoid proliferation of language_document reports;
wenzelm
parents: 67567
diff changeset
   209
      Latex.string "%\n" :: output_document ctxt {markdown = true} source @ [Latex.string "\n"]);
61455
0e4c257358cf clarified modules;
wenzelm
parents: 61435
diff changeset
   210
0e4c257358cf clarified modules;
wenzelm
parents: 61435
diff changeset
   211
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   212
(* command spans *)
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   213
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   214
type command = string * Position.T * string list;   (*name, position, tags*)
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   215
type source = (token * (string * int)) list;        (*token, markup flag, meta-comment depth*)
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   216
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   217
datatype span = Span of command * (source * source * source * source) * bool;
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   218
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   219
fun make_span cmd src =
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   220
  let
67522
9e712280cc37 clarified take/drop/chop prefix/suffix;
wenzelm
parents: 67508
diff changeset
   221
    fun chop_newline (tok :: toks) =
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   222
          if newline_token (fst tok) then ([tok], toks, true)
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   223
          else ([], tok :: toks, false)
67522
9e712280cc37 clarified take/drop/chop prefix/suffix;
wenzelm
parents: 67508
diff changeset
   224
      | chop_newline [] = ([], [], false);
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   225
    val (((src_prefix, src_main), src_suffix1), (src_suffix2, src_appendix, newline)) =
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   226
      src
67522
9e712280cc37 clarified take/drop/chop prefix/suffix;
wenzelm
parents: 67508
diff changeset
   227
      |> chop_prefix (white_token o fst)
9e712280cc37 clarified take/drop/chop prefix/suffix;
wenzelm
parents: 67508
diff changeset
   228
      ||>> chop_suffix (white_token o fst)
9e712280cc37 clarified take/drop/chop prefix/suffix;
wenzelm
parents: 67508
diff changeset
   229
      ||>> chop_prefix (white_comment_token o fst)
9e712280cc37 clarified take/drop/chop prefix/suffix;
wenzelm
parents: 67508
diff changeset
   230
      ||> chop_newline;
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   231
  in Span (cmd, (src_prefix, src_main, src_suffix1 @ src_suffix2, src_appendix), newline) end;
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   232
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   233
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   234
(* present spans *)
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   235
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   236
local
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   237
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   238
fun err_bad_nesting pos =
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   239
  error ("Bad nesting of commands in presentation" ^ pos);
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   240
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   241
fun edge which f (x: string option, y) =
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   242
  if x = y then I
67194
1c0a6a957114 positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents: 67191
diff changeset
   243
  else (case which (x, y) of NONE => I | SOME txt => cons (Latex.string (f txt)));
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   244
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   245
val begin_tag = edge #2 Latex.begin_tag;
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   246
val end_tag = edge #1 Latex.end_tag;
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   247
fun open_delim delim e = edge #2 Latex.begin_delim e #> delim #> edge #2 Latex.end_delim e;
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   248
fun close_delim delim e = edge #1 Latex.begin_delim e #> delim #> edge #1 Latex.end_delim e;
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   249
68507
60668de02229 more flexible document_tags;
wenzelm
parents: 68506
diff changeset
   250
fun read_tag s =
68512
16ae55c77bcb clarified syntax;
wenzelm
parents: 68510
diff changeset
   251
  (case space_explode "%" s of
16ae55c77bcb clarified syntax;
wenzelm
parents: 68510
diff changeset
   252
    ["", b] => (SOME b, NONE)
68510
795f39bfe4e1 simplified: allow only command names, with dummy for default;
wenzelm
parents: 68509
diff changeset
   253
  | [a, b] => (NONE, SOME (a, b))
68507
60668de02229 more flexible document_tags;
wenzelm
parents: 68506
diff changeset
   254
  | _ => error ("Bad document_tags specification: " ^ quote s));
60668de02229 more flexible document_tags;
wenzelm
parents: 68506
diff changeset
   255
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   256
in
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   257
68506
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   258
fun make_command_tag options keywords =
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   259
  let
68507
60668de02229 more flexible document_tags;
wenzelm
parents: 68506
diff changeset
   260
    val document_tags =
60668de02229 more flexible document_tags;
wenzelm
parents: 68506
diff changeset
   261
      map read_tag (space_explode "," (Options.string options \<^system_option>\<open>document_tags\<close>));
60668de02229 more flexible document_tags;
wenzelm
parents: 68506
diff changeset
   262
    val document_tags_default = map_filter #1 document_tags;
68510
795f39bfe4e1 simplified: allow only command names, with dummy for default;
wenzelm
parents: 68509
diff changeset
   263
    val document_tags_command = map_filter #2 document_tags;
68506
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   264
  in
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   265
    fn {cmd_name, cmd_tags, tag, active_tag} => fn state => fn state' =>
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   266
      let
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   267
        val tag' = try hd (fold (update (op =)) cmd_tags (the_list tag));
68508
wenzelm
parents: 68507
diff changeset
   268
68509
0f91ff642658 clarified: more uniform keyword_tags;
wenzelm
parents: 68508
diff changeset
   269
        val keyword_tags =
0f91ff642658 clarified: more uniform keyword_tags;
wenzelm
parents: 68508
diff changeset
   270
          if cmd_name = "end" andalso Toplevel.is_end_theory state' then ["theory"]
0f91ff642658 clarified: more uniform keyword_tags;
wenzelm
parents: 68508
diff changeset
   271
          else Keyword.command_tags keywords cmd_name;
68508
wenzelm
parents: 68507
diff changeset
   272
        val command_tags =
wenzelm
parents: 68507
diff changeset
   273
          the_list (AList.lookup (op =) document_tags_command cmd_name) @
wenzelm
parents: 68507
diff changeset
   274
          keyword_tags @ document_tags_default;
wenzelm
parents: 68507
diff changeset
   275
68506
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   276
        val active_tag' =
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   277
          if is_some tag' then tag'
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   278
          else
68508
wenzelm
parents: 68507
diff changeset
   279
            (case command_tags of
68506
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   280
              default_tag :: _ => SOME default_tag
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   281
            | [] =>
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   282
                if Keyword.is_vacuous keywords cmd_name andalso Toplevel.is_proof state
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   283
                then active_tag
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   284
                else NONE);
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   285
      in {tag' = tag', active_tag' = active_tag'} end
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   286
  end;
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   287
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   288
fun present_span thy command_tag span state state'
67194
1c0a6a957114 positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents: 67191
diff changeset
   289
    (tag_stack, active_tag, newline, latex, present_cont) =
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   290
  let
67381
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67378
diff changeset
   291
    val ctxt' =
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67378
diff changeset
   292
      Toplevel.presentation_context state'
68482
cb84beb84ca9 clarified signature;
wenzelm
parents: 68194
diff changeset
   293
        handle Toplevel.UNDEF => Proof_Context.get_global thy Context.PureN;
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   294
    val present = fold (fn (tok, (flag, 0)) =>
67381
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67378
diff changeset
   295
        fold cons (present_token ctxt' tok)
67194
1c0a6a957114 positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents: 67191
diff changeset
   296
        #> cons (Latex.string flag)
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   297
      | _ => I);
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   298
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   299
    val Span ((cmd_name, cmd_pos, cmd_tags), srcs, span_newline) = span;
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   300
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   301
    val (tag, tags) = tag_stack;
68506
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   302
    val {tag', active_tag'} =
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   303
      command_tag {cmd_name = cmd_name, cmd_tags = cmd_tags, tag = tag, active_tag = active_tag}
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   304
        state state';
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   305
    val edge = (active_tag, active_tag');
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   306
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58992
diff changeset
   307
    val nesting = Toplevel.level state' - Toplevel.level state;
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58992
diff changeset
   308
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   309
    val newline' =
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   310
      if is_none active_tag' then span_newline else newline;
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   311
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   312
    val tag_stack' =
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   313
      if nesting = 0 andalso not (Toplevel.is_proof state) then tag_stack
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   314
      else if nesting >= 0 then (tag', replicate nesting tag @ tags)
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   315
      else
33957
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
   316
        (case drop (~ nesting - 1) tags of
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58992
diff changeset
   317
          tg :: tgs => (tg, tgs)
48992
0518bf89c777 renamed Position.str_of to Position.here;
wenzelm
parents: 48927
diff changeset
   318
        | [] => err_bad_nesting (Position.here cmd_pos));
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   319
67194
1c0a6a957114 positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents: 67191
diff changeset
   320
    val latex' =
1c0a6a957114 positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents: 67191
diff changeset
   321
      latex
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   322
      |> end_tag edge
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   323
      |> close_delim (fst present_cont) edge
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   324
      |> snd present_cont
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   325
      |> open_delim (present (#1 srcs)) edge
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   326
      |> begin_tag edge
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   327
      |> present (#2 srcs);
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   328
    val present_cont' =
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   329
      if newline then (present (#3 srcs), present (#4 srcs))
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   330
      else (I, present (#3 srcs) #> present (#4 srcs));
67194
1c0a6a957114 positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents: 67191
diff changeset
   331
  in (tag_stack', active_tag', newline', latex', present_cont') end;
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   332
67194
1c0a6a957114 positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents: 67191
diff changeset
   333
fun present_trailer ((_, tags), active_tag, _, latex, present_cont) =
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   334
  if not (null tags) then err_bad_nesting " at end of theory"
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   335
  else
67194
1c0a6a957114 positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents: 67191
diff changeset
   336
    latex
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   337
    |> end_tag (active_tag, NONE)
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   338
    |> close_delim (fst present_cont) (active_tag, NONE)
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   339
    |> snd present_cont;
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   340
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   341
end;
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   342
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   343
28427
cc9f7d99fb73 turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
wenzelm
parents: 28273
diff changeset
   344
(* present_thy *)
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   345
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   346
local
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   347
61455
0e4c257358cf clarified modules;
wenzelm
parents: 61435
diff changeset
   348
val markup_true = "\\isamarkuptrue%\n";
0e4c257358cf clarified modules;
wenzelm
parents: 61435
diff changeset
   349
val markup_false = "\\isamarkupfalse%\n";
0e4c257358cf clarified modules;
wenzelm
parents: 61435
diff changeset
   350
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   351
val space_proper =
67439
78759a7bd874 more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents: 67429
diff changeset
   352
  Scan.one Token.is_blank -- Scan.many is_white_comment -- Scan.one is_black;
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   353
67439
78759a7bd874 more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents: 67429
diff changeset
   354
val is_improper = not o (is_black orf Token.is_begin_ignore orf Token.is_end_ignore);
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   355
val improper = Scan.many is_improper;
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   356
val improper_end = Scan.repeat (Scan.unless space_proper (Scan.one is_improper));
36959
f5417836dbea renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents: 36950
diff changeset
   357
val blank_end = Scan.repeat (Scan.unless space_proper (Scan.one Token.is_blank));
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   358
36959
f5417836dbea renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents: 36950
diff changeset
   359
val opt_newline = Scan.option (Scan.one Token.is_newline);
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   360
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   361
val ignore =
36959
f5417836dbea renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents: 36950
diff changeset
   362
  Scan.depend (fn d => opt_newline |-- Scan.one Token.is_begin_ignore
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   363
    >> pair (d + 1)) ||
36959
f5417836dbea renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents: 36950
diff changeset
   364
  Scan.depend (fn d => Scan.one Token.is_end_ignore --|
43947
9b00f09f7721 defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
wenzelm
parents: 43564
diff changeset
   365
    (if d = 0 then Scan.fail_with (K (fn () => "Bad nesting of meta-comments")) else opt_newline)
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   366
    >> pair (d - 1));
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   367
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36745
diff changeset
   368
val tag = (improper -- Parse.$$$ "%" -- improper) |-- Parse.!!! (Parse.tag_name --| blank_end);
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   369
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   370
val locale =
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36745
diff changeset
   371
  Scan.option ((Parse.$$$ "(" -- improper -- Parse.$$$ "in") |--
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 62749
diff changeset
   372
    Parse.!!! (improper |-- Parse.name --| (improper -- Parse.$$$ ")")));
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   373
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   374
in
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   375
68182
fa3cf61121ee tuned signature (see Command.eval_state);
wenzelm
parents: 68178
diff changeset
   376
type segment = {span: Command_Span.span, command: Toplevel.transition, state: Toplevel.state};
68178
e3dd94d04eee more explicit type Thy_Output.segment;
wenzelm
parents: 67572
diff changeset
   377
68506
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   378
fun present_thy options thy (segments: segment list) =
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   379
  let
58928
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58923
diff changeset
   380
    val keywords = Thy_Header.get_keywords thy;
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58923
diff changeset
   381
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58923
diff changeset
   382
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   383
    (* tokens *)
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   384
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   385
    val ignored = Scan.state --| ignore
67360
wenzelm
parents: 67359
diff changeset
   386
      >> (fn d => (NONE, (Ignore_Token, ("", d))));
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   387
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58992
diff changeset
   388
    fun markup pred mk flag = Scan.peek (fn d =>
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36745
diff changeset
   389
      improper |--
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58992
diff changeset
   390
        Parse.position (Scan.one (fn tok =>
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58992
diff changeset
   391
          Token.is_command tok andalso pred keywords (Token.content_of tok))) --
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   392
      Scan.repeat tag --
51627
589daaf48dba tuned signature -- agree with markup terminology;
wenzelm
parents: 51626
diff changeset
   393
      Parse.!!!! ((improper -- locale -- improper) |-- Parse.document_source --| improper_end)
59064
a8bcb5a446c8 more abstract type Input.source;
wenzelm
parents: 58999
diff changeset
   394
      >> (fn (((tok, pos'), tags), source) =>
59065
8ce02aafc363 tuned signature -- prefer Input.source;
wenzelm
parents: 59064
diff changeset
   395
        let val name = Token.content_of tok
8ce02aafc363 tuned signature -- prefer Input.source;
wenzelm
parents: 59064
diff changeset
   396
        in (SOME (name, pos', tags), (mk (name, source), (flag, d))) end));
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   397
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   398
    val command = Scan.peek (fn d =>
61814
1ca1142e1711 clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents: 61748
diff changeset
   399
      Scan.optional (Scan.one Token.is_command_modifier ::: improper) [] --
59924
801b979ec0c2 more general notion of command span: command keyword not necessarily at start;
wenzelm
parents: 59917
diff changeset
   400
      Scan.one Token.is_command -- Scan.repeat tag
59939
7d46aa03696e support for 'restricted' modifier: only qualified accesses outside the local scope;
wenzelm
parents: 59924
diff changeset
   401
      >> (fn ((cmd_mod, cmd), tags) =>
7d46aa03696e support for 'restricted' modifier: only qualified accesses outside the local scope;
wenzelm
parents: 59924
diff changeset
   402
        map (fn tok => (NONE, (Basic_Token tok, ("", d)))) cmd_mod @
59924
801b979ec0c2 more general notion of command span: command keyword not necessarily at start;
wenzelm
parents: 59917
diff changeset
   403
          [(SOME (Token.content_of cmd, Token.pos_of cmd, tags),
61455
0e4c257358cf clarified modules;
wenzelm
parents: 61435
diff changeset
   404
            (Basic_Token cmd, (markup_false, d)))]));
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   405
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   406
    val cmt = Scan.peek (fn d =>
67446
1f4d167b6ac9 discontinued old form of marginal comments;
wenzelm
parents: 67439
diff changeset
   407
      Scan.one is_black_comment >> (fn tok => (NONE, (Basic_Token tok, ("", d)))));
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   408
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   409
    val other = Scan.peek (fn d =>
59065
8ce02aafc363 tuned signature -- prefer Input.source;
wenzelm
parents: 59064
diff changeset
   410
       Parse.not_eof >> (fn tok => (NONE, (Basic_Token tok, ("", d)))));
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   411
59924
801b979ec0c2 more general notion of command span: command keyword not necessarily at start;
wenzelm
parents: 59917
diff changeset
   412
    val tokens =
801b979ec0c2 more general notion of command span: command keyword not necessarily at start;
wenzelm
parents: 59917
diff changeset
   413
      (ignored ||
61455
0e4c257358cf clarified modules;
wenzelm
parents: 61435
diff changeset
   414
        markup Keyword.is_document_heading Markup_Token markup_true ||
0e4c257358cf clarified modules;
wenzelm
parents: 61435
diff changeset
   415
        markup Keyword.is_document_body Markup_Env_Token markup_true ||
61456
b521b8b400f7 trim_blanks after read, before eval;
wenzelm
parents: 61455
diff changeset
   416
        markup Keyword.is_document_raw (Raw_Token o #2) "") >> single ||
59924
801b979ec0c2 more general notion of command span: command keyword not necessarily at start;
wenzelm
parents: 59917
diff changeset
   417
      command ||
801b979ec0c2 more general notion of command span: command keyword not necessarily at start;
wenzelm
parents: 59917
diff changeset
   418
      (cmt || other) >> single;
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   419
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   420
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   421
    (* spans *)
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   422
59065
8ce02aafc363 tuned signature -- prefer Input.source;
wenzelm
parents: 59064
diff changeset
   423
    val is_eof = fn (_, (Basic_Token x, _)) => Token.is_eof x | _ => false;
8ce02aafc363 tuned signature -- prefer Input.source;
wenzelm
parents: 59064
diff changeset
   424
    val stopper = Scan.stopper (K (NONE, (Basic_Token Token.eof, ("", 0)))) is_eof;
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   425
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   426
    val cmd = Scan.one (is_some o fst);
27732
8dbf5761a24a abstract type Scan.stopper;
wenzelm
parents: 27566
diff changeset
   427
    val non_cmd = Scan.one (is_none o fst andf not o is_eof) >> #2;
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   428
67439
78759a7bd874 more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents: 67429
diff changeset
   429
    val white_comments = Scan.many (white_comment_token o fst o snd);
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   430
    val blank = Scan.one (blank_token o fst o snd);
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   431
    val newline = Scan.one (newline_token o fst o snd);
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   432
    val before_cmd =
67439
78759a7bd874 more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents: 67429
diff changeset
   433
      Scan.option (newline -- white_comments) --
78759a7bd874 more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents: 67429
diff changeset
   434
      Scan.option (newline -- white_comments) --
78759a7bd874 more uniform support for formal comments in outer syntax, notably \<^cancel> and \<^latex>;
wenzelm
parents: 67429
diff changeset
   435
      Scan.option (blank -- white_comments) -- cmd;
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   436
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   437
    val span =
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   438
      Scan.repeat non_cmd -- cmd --
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   439
        Scan.repeat (Scan.unless before_cmd non_cmd) --
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   440
        Scan.option (newline >> (single o snd))
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   441
      >> (fn (((toks1, (cmd, tok2)), toks3), tok4) =>
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   442
          make_span (the cmd) (toks1 @ (tok2 :: (toks3 @ the_default [] tok4))));
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   443
68178
e3dd94d04eee more explicit type Thy_Output.segment;
wenzelm
parents: 67572
diff changeset
   444
    val spans = segments
e3dd94d04eee more explicit type Thy_Output.segment;
wenzelm
parents: 67572
diff changeset
   445
      |> maps (Command_Span.content o #span)
67522
9e712280cc37 clarified take/drop/chop prefix/suffix;
wenzelm
parents: 67508
diff changeset
   446
      |> drop_suffix Token.is_space
57080
0e5fa27d3293 strip trailing white space, to avoid notorious problems of jEdit with last line;
wenzelm
parents: 56548
diff changeset
   447
      |> Source.of_list
59924
801b979ec0c2 more general notion of command span: command keyword not necessarily at start;
wenzelm
parents: 59917
diff changeset
   448
      |> Source.source' 0 Token.stopper (Scan.error (Scan.bulk tokens >> flat))
58864
505a8150368a recover via scanner;
wenzelm
parents: 58861
diff changeset
   449
      |> Source.source stopper (Scan.error (Scan.bulk span))
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   450
      |> Source.exhaust;
28427
cc9f7d99fb73 turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
wenzelm
parents: 28273
diff changeset
   451
68178
e3dd94d04eee more explicit type Thy_Output.segment;
wenzelm
parents: 67572
diff changeset
   452
    val command_results =
68182
fa3cf61121ee tuned signature (see Command.eval_state);
wenzelm
parents: 68178
diff changeset
   453
      segments |> map_filter (fn {command, state, ...} =>
fa3cf61121ee tuned signature (see Command.eval_state);
wenzelm
parents: 68178
diff changeset
   454
        if Toplevel.is_ignored command then NONE else SOME (command, state));
68178
e3dd94d04eee more explicit type Thy_Output.segment;
wenzelm
parents: 67572
diff changeset
   455
28427
cc9f7d99fb73 turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
wenzelm
parents: 28273
diff changeset
   456
cc9f7d99fb73 turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
wenzelm
parents: 28273
diff changeset
   457
    (* present commands *)
cc9f7d99fb73 turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
wenzelm
parents: 28273
diff changeset
   458
68506
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   459
    val command_tag = make_command_tag options keywords;
67138
82283d52b4d6 system option for default command tags;
wenzelm
parents: 66021
diff changeset
   460
28427
cc9f7d99fb73 turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
wenzelm
parents: 28273
diff changeset
   461
    fun present_command tr span st st' =
68506
cef6c6b009fb prefer explicit options;
wenzelm
parents: 68505
diff changeset
   462
      Toplevel.setmp_thread_position tr (present_span thy command_tag span st st');
28427
cc9f7d99fb73 turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
wenzelm
parents: 28273
diff changeset
   463
cc9f7d99fb73 turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
wenzelm
parents: 28273
diff changeset
   464
    fun present _ [] = I
68178
e3dd94d04eee more explicit type Thy_Output.segment;
wenzelm
parents: 67572
diff changeset
   465
      | present st ((span, (tr, st')) :: rest) = present_command tr span st st' #> present st' rest;
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   466
  in
28427
cc9f7d99fb73 turned process_thy into present_thy, which merely does presentation (wrt. persistent intermediate states);
wenzelm
parents: 28273
diff changeset
   467
    if length command_results = length spans then
67194
1c0a6a957114 positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents: 67191
diff changeset
   468
      ((NONE, []), NONE, true, [], (I, I))
68178
e3dd94d04eee more explicit type Thy_Output.segment;
wenzelm
parents: 67572
diff changeset
   469
      |> present Toplevel.toplevel (spans ~~ command_results)
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   470
      |> present_trailer
67194
1c0a6a957114 positions as postlude: avoid intrusion of odd %-forms into main tex source;
wenzelm
parents: 67191
diff changeset
   471
      |> rev
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   472
    else error "Messed-up outer syntax for presentation"
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   473
  end;
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   474
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   475
end;
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   476
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   477
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   478
67386
998e01d6f8fd clarified modules;
wenzelm
parents: 67381
diff changeset
   479
(** standard output operations **)
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   480
67463
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   481
(* pretty printing *)
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   482
67463
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   483
fun pretty_term ctxt t =
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   484
  Syntax.pretty_term (Variable.auto_fixes t ctxt) t;
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 24680
diff changeset
   485
32898
e871d897969c term styles also cover antiquotations term_type and typeof
haftmann
parents: 32890
diff changeset
   486
fun pretty_thm ctxt = pretty_term ctxt o Thm.full_prop_of;
e871d897969c term styles also cover antiquotations term_type and typeof
haftmann
parents: 32890
diff changeset
   487
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   488
30390
ad591ee76c78 simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
wenzelm
parents: 30381
diff changeset
   489
(* default output *)
ad591ee76c78 simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
wenzelm
parents: 30381
diff changeset
   490
67508
189ab2c3026b verbatim output consists of plain lines;
wenzelm
parents: 67506
diff changeset
   491
val lines = separate (Latex.string "\\isanewline%\n");
67505
ceb324e34c14 clarified signature: items with \isasep are special;
wenzelm
parents: 67475
diff changeset
   492
val items = separate (Latex.string "\\isasep\\isanewline%\n");
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   493
67463
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   494
fun isabelle ctxt body =
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   495
  if Config.get ctxt Document_Antiquotation.thy_output_display
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   496
  then Latex.environment_block "isabelle" body
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   497
  else Latex.block (Latex.enclose_body "\\isa{" "}" body);
59175
bf465f335e85 system option "pretty_margin" is superseded by "thy_output_margin";
wenzelm
parents: 59067
diff changeset
   498
67463
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   499
fun isabelle_typewriter ctxt body =
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   500
  if Config.get ctxt Document_Antiquotation.thy_output_display
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   501
  then Latex.environment_block "isabellett" body
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   502
  else Latex.block (Latex.enclose_body "\\isatt{" "}" body);
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   503
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   504
fun typewriter ctxt s =
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   505
  isabelle_typewriter ctxt [Latex.string (Latex.output_ascii s)];
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   506
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   507
fun verbatim ctxt =
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   508
  if Config.get ctxt Document_Antiquotation.thy_output_display
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   509
  then Document_Antiquotation.indent_lines ctxt #> typewriter ctxt
67508
189ab2c3026b verbatim output consists of plain lines;
wenzelm
parents: 67506
diff changeset
   510
  else split_lines #> map (typewriter ctxt) #> lines #> Latex.block;
22124
27b674312b2f renamed Isar/isar_output.ML to Thy/thy_output.ML;
wenzelm
parents:
diff changeset
   511
67474
90def2b06305 more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents: 67465
diff changeset
   512
fun source ctxt =
90def2b06305 more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents: 67465
diff changeset
   513
  Token.args_of_src
90def2b06305 more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents: 67465
diff changeset
   514
  #> map (Token.unparse #> Document_Antiquotation.prepare_lines ctxt)
90def2b06305 more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents: 67465
diff changeset
   515
  #> space_implode " "
90def2b06305 more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents: 67465
diff changeset
   516
  #> output_source ctxt
90def2b06305 more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents: 67465
diff changeset
   517
  #> isabelle ctxt;
90def2b06305 more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents: 67465
diff changeset
   518
67463
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   519
fun pretty ctxt =
67505
ceb324e34c14 clarified signature: items with \isasep are special;
wenzelm
parents: 67475
diff changeset
   520
  Document_Antiquotation.output ctxt #> Latex.string #> single #> isabelle ctxt;
ceb324e34c14 clarified signature: items with \isasep are special;
wenzelm
parents: 67475
diff changeset
   521
ceb324e34c14 clarified signature: items with \isasep are special;
wenzelm
parents: 67475
diff changeset
   522
fun pretty_source ctxt src prt =
ceb324e34c14 clarified signature: items with \isasep are special;
wenzelm
parents: 67475
diff changeset
   523
  if Config.get ctxt Document_Antiquotation.thy_output_source
ceb324e34c14 clarified signature: items with \isasep are special;
wenzelm
parents: 67475
diff changeset
   524
  then source ctxt src else pretty ctxt prt;
67463
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   525
67505
ceb324e34c14 clarified signature: items with \isasep are special;
wenzelm
parents: 67475
diff changeset
   526
fun pretty_items ctxt =
ceb324e34c14 clarified signature: items with \isasep are special;
wenzelm
parents: 67475
diff changeset
   527
  map (Document_Antiquotation.output ctxt #> Latex.string) #> items #> isabelle ctxt;
ceb324e34c14 clarified signature: items with \isasep are special;
wenzelm
parents: 67475
diff changeset
   528
ceb324e34c14 clarified signature: items with \isasep are special;
wenzelm
parents: 67475
diff changeset
   529
fun pretty_items_source ctxt src prts =
67474
90def2b06305 more uniform output of source / text / theory_text, with handling of formal comments etc.;
wenzelm
parents: 67465
diff changeset
   530
  if Config.get ctxt Document_Antiquotation.thy_output_source
67505
ceb324e34c14 clarified signature: items with \isasep are special;
wenzelm
parents: 67475
diff changeset
   531
  then source ctxt src else pretty_items ctxt prts;
61473
34d1913f0b20 clarified control antiquotations: decode control symbol to get name;
wenzelm
parents: 61471
diff changeset
   532
67465
wenzelm
parents: 67463
diff changeset
   533
wenzelm
parents: 67463
diff changeset
   534
(* antiquotation variants *)
wenzelm
parents: 67463
diff changeset
   535
69592
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   536
local
67463
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   537
69592
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   538
fun gen_setup name embedded =
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   539
  if embedded
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   540
  then Document_Antiquotation.setup_embedded name
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   541
  else Document_Antiquotation.setup name;
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   542
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   543
fun gen_antiquotation_pretty name embedded scan f =
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   544
  gen_setup name embedded scan (fn {context = ctxt, argument = x, ...} => pretty ctxt (f ctxt x));
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   545
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   546
fun gen_antiquotation_pretty_source name embedded scan f =
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   547
  gen_setup name embedded scan
67463
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   548
    (fn {context = ctxt, source = src, argument = x} => pretty_source ctxt src (f ctxt x));
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   549
69592
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   550
fun gen_antiquotation_raw name embedded scan f =
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   551
  gen_setup name embedded scan (fn {context = ctxt, argument = x, ...} => f ctxt x);
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   552
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   553
fun gen_antiquotation_verbatim name embedded scan f =
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   554
  gen_antiquotation_raw name embedded scan (fn ctxt => verbatim ctxt o f ctxt);
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   555
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   556
in
67463
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67461
diff changeset
   557
69592
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   558
fun antiquotation_pretty name = gen_antiquotation_pretty name false;
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   559
fun antiquotation_pretty_embedded name = gen_antiquotation_pretty name true;
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   560
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   561
fun antiquotation_pretty_source name = gen_antiquotation_pretty_source name false;
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   562
fun antiquotation_pretty_source_embedded name = gen_antiquotation_pretty_source name true;
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   563
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   564
fun antiquotation_raw name = gen_antiquotation_raw name false;
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   565
fun antiquotation_raw_embedded name = gen_antiquotation_raw name true;
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   566
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   567
fun antiquotation_verbatim name = gen_antiquotation_verbatim name false;
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   568
fun antiquotation_verbatim_embedded name = gen_antiquotation_verbatim name true;
61491
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61473
diff changeset
   569
30390
ad591ee76c78 simplified interface to define document antiquotations, cf. antiquotatation, maybe_pretty_source, output;
wenzelm
parents: 30381
diff changeset
   570
end;
69592
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   571
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 68512
diff changeset
   572
end;