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