src/Pure/Isar/isar_output.ML
author wenzelm
Thu, 03 Aug 2000 18:43:35 +0200
changeset 9512 c30f6d0f81d2
parent 9504 8168600e88a5
child 9732 c32c7ef228c6
permissions -rw-r--r--
added unknown_theory/proof/context;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9140
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/isar_output.ML
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
     3
    Author:     Markus Wenzel, TU Muenchen
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
     4
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
     5
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
     6
Isar theory output.
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
     7
*)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
     8
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
     9
signature ISAR_OUTPUT =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    10
sig
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    11
  val add_commands: (string * (Args.src -> Proof.context -> string)) list -> unit
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    12
  val add_options: (string * (string -> (unit -> string) -> unit -> string)) list -> unit
9220
d0506ced27c8 added options "eta_contract", "long_names";
wenzelm
parents: 9213
diff changeset
    13
  val print_antiquotations: unit -> unit
9140
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    14
  val boolean: string -> bool
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    15
  val integer: string -> int
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    16
  val args: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) ->
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    17
    (Proof.context -> 'a -> string) -> Args.src -> Proof.context -> string
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    18
  datatype markup = Markup | MarkupEnv | Verbatim
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    19
  val interest_level: int ref
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    20
  val parse_thy: (markup -> OuterLex.token list -> OuterLex.token * OuterLex.token list) ->
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    21
    Scan.lexicon -> Toplevel.transition list -> (OuterLex.token, 'a) Source.source ->
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    22
      (Toplevel.transition * (Toplevel.state -> Buffer.T -> Buffer.T)) list * Buffer.T
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    23
  val display: bool ref
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    24
  val quotes: bool ref
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    25
end;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    26
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    27
structure IsarOutput: ISAR_OUTPUT =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    28
struct
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    29
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    30
structure T = OuterLex;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    31
structure P = OuterParse;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    32
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    33
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    34
(** maintain global commands **)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    35
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    36
local
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    37
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    38
val global_commands =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    39
  ref (Symtab.empty: (Args.src -> Proof.context -> string) Symtab.table);
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    40
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    41
val global_options =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    42
  ref (Symtab.empty: (string -> (unit -> string) -> unit -> string) Symtab.table);
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    43
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    44
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    45
fun add_item kind (tab, (name, x)) =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    46
 (if is_none (Symtab.lookup (tab, name)) then ()
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    47
  else warning ("Redefined antiquotation output " ^ kind ^ ": " ^ quote name);
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    48
  Symtab.update ((name, x), tab));
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    49
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    50
fun add_items kind xs tab = foldl (add_item kind) (tab, xs);
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    51
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    52
in
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    53
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    54
val add_commands = Library.change global_commands o (add_items "command");
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    55
val add_options = Library.change global_options o (add_items "option");
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    56
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    57
fun command src =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    58
  let val ((name, _), pos) = Args.dest_src src in
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    59
    (case Symtab.lookup (! global_commands, name) of
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    60
      None => error ("Unknown antiquotation command: " ^ quote name ^ Position.str_of pos)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    61
    | Some f => transform_failure (curry Comment.OUTPUT_FAIL (name, pos)) (f src))
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    62
  end;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    63
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    64
fun option (name, s) f () =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    65
  (case Symtab.lookup (! global_options, name) of
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    66
    None => error ("Unknown antiquotation option: " ^ quote name)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    67
  | Some opt => opt s f ());
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    68
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    69
fun options [] f = f
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    70
  | options (opt :: opts) f = option opt (options opts f);
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    71
9213
2651a4db8883 help_antiquotations;
wenzelm
parents: 9140
diff changeset
    72
9220
d0506ced27c8 added options "eta_contract", "long_names";
wenzelm
parents: 9213
diff changeset
    73
fun print_antiquotations () =
d0506ced27c8 added options "eta_contract", "long_names";
wenzelm
parents: 9213
diff changeset
    74
 [Pretty.big_list "antiquotation commands:" (map Pretty.str (Symtab.keys (! global_commands))),
d0506ced27c8 added options "eta_contract", "long_names";
wenzelm
parents: 9213
diff changeset
    75
  Pretty.big_list "antiquotation options:" (map Pretty.str (Symtab.keys (! global_options)))]
d0506ced27c8 added options "eta_contract", "long_names";
wenzelm
parents: 9213
diff changeset
    76
 |> Pretty.chunks |> Pretty.writeln;
9213
2651a4db8883 help_antiquotations;
wenzelm
parents: 9140
diff changeset
    77
9140
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    78
end;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    79
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    80
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    81
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    82
(** syntax of antiquotations **)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    83
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    84
(* option values *)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    85
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    86
fun boolean "" = true
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    87
  | boolean "true" = true
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    88
  | boolean "false" = false
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    89
  | boolean s = error ("Bad boolean value: " ^ quote s);
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    90
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    91
fun integer s =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    92
  let
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    93
    fun int ss =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    94
      (case Term.read_int ss of (i, []) => i | _ => error ("Bad integer value: " ^ quote s));
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    95
  in (case Symbol.explode s of "-" :: ss => ~ (int ss) | ss => int ss) end;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    96
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    97
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    98
(* args syntax *)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
    99
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   100
fun syntax (scan: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list))) =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   101
  Args.syntax "antiquotation" scan;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   102
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   103
fun args scan f src ctxt : string =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   104
  let val (ctxt', x) = syntax scan src ctxt
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   105
  in f ctxt' x end;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   106
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   107
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   108
(* outer syntax *)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   109
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   110
local
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   111
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   112
val property = P.xname -- Scan.optional (P.$$$ "=" |-- P.!!! P.xname) "";
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   113
val properties = Scan.optional (P.$$$ "[" |-- P.!!! (P.enum "," property --| P.$$$ "]")) [];
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   114
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   115
val antiq = P.position P.xname -- properties -- P.arguments --| Scan.ahead P.eof
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   116
  >> (fn (((x, pos), y), z) => (y, Args.src ((x, z), pos)));
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   117
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   118
fun antiq_args_aux keyword_lexicon (str, pos) =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   119
  Source.of_string str
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   120
  |> Symbol.source false
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   121
  |> T.source false (K (keyword_lexicon, Scan.empty_lexicon)) pos
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   122
  |> T.source_proper
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   123
  |> Source.source T.stopper (Scan.error (Scan.bulk (P.!!! antiq))) None
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   124
  |> Source.exhaust;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   125
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   126
in
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   127
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   128
fun antiq_args lex (s, pos) =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   129
  (case antiq_args_aux lex (s, pos) of [x] => x | _ => raise ERROR) handle ERROR =>
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   130
    error ("Malformed antiquotation: " ^ quote ("@{" ^ s ^ "}") ^ Position.str_of pos);
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   131
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   132
end;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   133
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   134
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   135
(* eval_antiquote *)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   136
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   137
fun eval_antiquote lex state (str, pos) =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   138
  let
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   139
    fun expand (Antiquote.Text s) = s
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   140
      | expand (Antiquote.Antiq x) =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   141
          let val (opts, src) = antiq_args lex x in
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   142
            Library.setmp print_mode Latex.modes   (*note: lazy lookup of context below!*)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   143
              (options opts (fn () => command src (Toplevel.context_of state))) () end;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   144
    val ants = Antiquote.antiquotes_of (str, pos);
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   145
  in
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   146
    if Toplevel.is_toplevel state andalso exists Antiquote.is_antiq ants then
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   147
      error ("Cannot expand antiquotations at top-level" ^ Position.str_of pos)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   148
    else implode (map expand ants)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   149
  end;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   150
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   151
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   152
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   153
(** present theory source **)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   154
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   155
(* present_tokens *)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   156
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   157
val interest_level = ref 0;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   158
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   159
fun present_tokens lex toks state =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   160
  toks
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   161
  |> mapfilter (fn (tk, i) => if i > ! interest_level then None else Some tk)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   162
  |> map (apsnd (eval_antiquote lex state))
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   163
  |> Latex.output_tokens
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   164
  |> Buffer.add;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   165
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   166
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   167
(* parse_thy *)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   168
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   169
datatype markup = Markup | MarkupEnv | Verbatim;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   170
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   171
local
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   172
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   173
val opt_newline = Scan.option (Scan.one T.is_newline);
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   174
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   175
fun check_level i =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   176
  if i > 0 then Scan.succeed ()
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   177
  else Scan.fail_with (K "Bad nesting of meta-comments");
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   178
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   179
val ignore =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   180
  Scan.depend (fn i => opt_newline |-- P.position (Scan.one T.is_begin_ignore) >> pair (i + 1)) ||
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   181
  Scan.depend (fn i => P.position (Scan.one T.is_end_ignore) --|
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   182
    (opt_newline -- check_level i) >> pair (i - 1));
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   183
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   184
val ignore_cmd = Scan.state -- ignore
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   185
  >> (fn (i, (x, pos)) => (false, ((Latex.Basic x, ("", pos)), i)));
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   186
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   187
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   188
val is_improper = not o (T.is_proper orf T.is_begin_ignore orf T.is_end_ignore);
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   189
val improper = Scan.any is_improper;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   190
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   191
val improper_keep_indent = Scan.repeat
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   192
  (Scan.unless (Scan.one T.is_indent -- Scan.one T.is_proper) (Scan.one is_improper));
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   193
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   194
val improper_end =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   195
  (improper -- P.semicolon) |-- improper_keep_indent ||
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   196
  improper_keep_indent;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   197
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   198
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   199
val stopper =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   200
  ((false, ((Latex.Basic (#1 T.stopper), ("", Position.none)), 0)),
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   201
    fn (_, ((Latex.Basic x, _), _)) => (#2 T.stopper x) | _ => false);
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   202
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   203
in
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   204
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   205
fun parse_thy markup lex trs src =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   206
  let
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   207
    val text = P.position P.text;
9220
d0506ced27c8 added options "eta_contract", "long_names";
wenzelm
parents: 9213
diff changeset
   208
    val token = Scan.depend (fn i =>
9140
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   209
     (improper |-- markup Markup -- P.!!!! (improper |-- text --| improper_end)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   210
        >> (fn (x, y) => (true, ((Latex.Markup (T.val_of x), y), i))) ||
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   211
      improper |-- markup MarkupEnv -- P.!!!! (improper |-- text --| improper_end)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   212
        >> (fn (x, y) => (true, ((Latex.MarkupEnv (T.val_of x), y), i))) ||
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   213
      (P.$$$ "--" |-- P.!!!! (improper |-- text))
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   214
        >> (fn y => (false, ((Latex.Markup "cmt", y), i))) ||
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   215
      (improper -- markup Verbatim) |-- P.!!!! (improper |-- text --| improper_end)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   216
        >> (fn y => (true, ((Latex.Verbatim, y), i))) ||
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   217
      P.position (Scan.one T.not_eof)
9220
d0506ced27c8 added options "eta_contract", "long_names";
wenzelm
parents: 9213
diff changeset
   218
        >> (fn (x, pos) => (T.is_kind T.Command x, ((Latex.Basic x, ("", pos)), i))))
d0506ced27c8 added options "eta_contract", "long_names";
wenzelm
parents: 9213
diff changeset
   219
      >> pair i);
9140
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   220
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   221
    val body = Scan.any (not o fst andf not o #2 stopper);
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   222
    val section = body -- Scan.one fst -- body >> (fn ((x, y), z) => map snd (x @ (y :: z)));
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   223
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   224
    val cmds =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   225
      src
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   226
      |> Source.filter (not o T.is_semicolon)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   227
      |> Source.source' 0 T.stopper (Scan.error (Scan.bulk (ignore_cmd || token))) None
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   228
      |> Source.source stopper (Scan.error (Scan.bulk section)) None
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   229
      |> Source.exhaust;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   230
  in
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   231
    if length cmds = length trs then
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   232
      (trs ~~ map (present_tokens lex) cmds, Buffer.empty)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   233
    else error "Messed up outer syntax for presentation"
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   234
  end;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   235
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   236
end;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   237
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   238
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   239
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   240
(** setup default output **)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   241
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   242
(* options *)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   243
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   244
val display = ref false;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   245
val quotes = ref false;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   246
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   247
val _ = add_options
9220
d0506ced27c8 added options "eta_contract", "long_names";
wenzelm
parents: 9213
diff changeset
   248
 [("show_types", Library.setmp Syntax.show_types o boolean),
d0506ced27c8 added options "eta_contract", "long_names";
wenzelm
parents: 9213
diff changeset
   249
  ("show_sorts", Library.setmp Syntax.show_sorts o boolean),
d0506ced27c8 added options "eta_contract", "long_names";
wenzelm
parents: 9213
diff changeset
   250
  ("eta_contract", Library.setmp Syntax.eta_contract o boolean),
d0506ced27c8 added options "eta_contract", "long_names";
wenzelm
parents: 9213
diff changeset
   251
  ("long_names", Library.setmp NameSpace.long_names o boolean),
9140
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   252
  ("display", Library.setmp display o boolean),
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   253
  ("quotes", Library.setmp quotes o boolean),
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   254
  ("mode", fn s => fn f => fn () => Library.setmp print_mode (s :: ! print_mode) f ()),
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   255
  ("margin", Pretty.setmp_margin o integer)];
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   256
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   257
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   258
(* commands *)
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   259
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   260
local
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   261
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   262
fun cond_quote prt =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   263
  if ! quotes then Pretty.quote prt else prt;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   264
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   265
fun cond_display prt =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   266
  if ! display then
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   267
    Pretty.string_of prt
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   268
    |> enclose "\n\\begin{isabelle}%\n" "\n\\end{isabelle}%\n"
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   269
  else
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   270
    Pretty.str_of prt
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   271
    |> enclose "\\isa{" "}";
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   272
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   273
val string_of = cond_display o cond_quote;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   274
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   275
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   276
fun pretty_typ ctxt T =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   277
  Sign.pretty_typ (ProofContext.sign_of ctxt) T;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   278
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   279
fun pretty_term ctxt t =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   280
  Sign.pretty_term (ProofContext.sign_of ctxt) (ProofContext.extern_skolem ctxt t);
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   281
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   282
fun pretty_thm ctxt thms =
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   283
  Pretty.chunks (map (pretty_term ctxt o #prop o Thm.rep_thm) thms);
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   284
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   285
in
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   286
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   287
val _ = add_commands
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   288
 [("thm", args Attrib.local_thms (string_of oo pretty_thm)),
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   289
  ("prop", args Args.local_prop (string_of oo pretty_term)),
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   290
  ("term", args Args.local_term (string_of oo pretty_term)),
9504
8168600e88a5 typ_no_norm;
wenzelm
parents: 9220
diff changeset
   291
  ("typ", args Args.local_typ_no_norm (string_of oo pretty_typ))];
9140
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   292
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   293
end;
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   294
69e8938c2409 Isar theory output.
wenzelm
parents:
diff changeset
   295
end;