src/Doc/antiquote_setup.ML
author wenzelm
Sun, 30 Nov 2014 13:15:04 +0100
changeset 59066 45ab32a542fe
parent 59064 a8bcb5a446c8
child 59067 dd8ec9138112
permissions -rw-r--r--
tuned signature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
48985
5386df44a037 renamed doc-src to src/Doc;
wenzelm
parents: 48899
diff changeset
     1
(*  Title:      Doc/antiquote_setup.ML
21375
ae8a112b62d7 Auxiliary antiquotations for Isabelle manuals.
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
ae8a112b62d7 Auxiliary antiquotations for Isabelle manuals.
wenzelm
parents:
diff changeset
     3
26742
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
     4
Auxiliary antiquotations for the Isabelle manuals.
21375
ae8a112b62d7 Auxiliary antiquotations for Isabelle manuals.
wenzelm
parents:
diff changeset
     5
*)
ae8a112b62d7 Auxiliary antiquotations for Isabelle manuals.
wenzelm
parents:
diff changeset
     6
56059
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
     7
structure Antiquote_Setup: sig end =
26742
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
     8
struct
21375
ae8a112b62d7 Auxiliary antiquotations for Isabelle manuals.
wenzelm
parents:
diff changeset
     9
26742
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
    10
(* misc utils *)
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
    11
29736
ec3fc818b82e clean_string/clean_name: proper treatment of \<dash>;
wenzelm
parents: 29726
diff changeset
    12
fun translate f = Symbol.explode #> map f #> implode;
ec3fc818b82e clean_string/clean_name: proper treatment of \<dash>;
wenzelm
parents: 29726
diff changeset
    13
ec3fc818b82e clean_string/clean_name: proper treatment of \<dash>;
wenzelm
parents: 29726
diff changeset
    14
val clean_string = translate
26853
52cb0e965041 clean_string: map "_" to "\\_" (best used with underscore.sty);
wenzelm
parents: 26843
diff changeset
    15
  (fn "_" => "\\_"
30120
aaa4667285c8 uniform treatment of ML indexing, using general \indexdef macro for formal Isabelle/Isar entities;
wenzelm
parents: 29736
diff changeset
    16
    | "#" => "\\#"
52408
fa2dc6c6c94f updated operations on proof terms;
wenzelm
parents: 50239
diff changeset
    17
    | "$" => "\\$"
fa2dc6c6c94f updated operations on proof terms;
wenzelm
parents: 50239
diff changeset
    18
    | "%" => "\\%"
26897
044619358d3a clean_string: cover <;
wenzelm
parents: 26894
diff changeset
    19
    | "<" => "$<$"
26768
844068d16ba0 clean_string: handle { };
wenzelm
parents: 26756
diff changeset
    20
    | ">" => "$>$"
844068d16ba0 clean_string: handle { };
wenzelm
parents: 26756
diff changeset
    21
    | "{" => "\\{"
30120
aaa4667285c8 uniform treatment of ML indexing, using general \indexdef macro for formal Isabelle/Isar entities;
wenzelm
parents: 29736
diff changeset
    22
    | "|" => "$\\mid$"
26768
844068d16ba0 clean_string: handle { };
wenzelm
parents: 26756
diff changeset
    23
    | "}" => "\\}"
42666
fee67c099d03 use existing \<hyphen>;
wenzelm
parents: 42361
diff changeset
    24
    | "\<hyphen>" => "-"
26768
844068d16ba0 clean_string: handle { };
wenzelm
parents: 26756
diff changeset
    25
    | c => c);
26751
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
    26
31546
d58d6acab331 eliminated escaped symbols;
wenzelm
parents: 31297
diff changeset
    27
fun clean_name "\<dots>" = "dots"
26897
044619358d3a clean_string: cover <;
wenzelm
parents: 26894
diff changeset
    28
  | clean_name ".." = "ddot"
044619358d3a clean_string: cover <;
wenzelm
parents: 26894
diff changeset
    29
  | clean_name "." = "dot"
26910
aa6357b39212 tuned clean_name (underscore);
wenzelm
parents: 26903
diff changeset
    30
  | clean_name "_" = "underscore"
26897
044619358d3a clean_string: cover <;
wenzelm
parents: 26894
diff changeset
    31
  | clean_name "{" = "braceleft"
044619358d3a clean_string: cover <;
wenzelm
parents: 26894
diff changeset
    32
  | clean_name "}" = "braceright"
42666
fee67c099d03 use existing \<hyphen>;
wenzelm
parents: 42361
diff changeset
    33
  | clean_name s = s |> translate (fn "_" => "-" | "\<hyphen>" => "-" | c => c);
26897
044619358d3a clean_string: cover <;
wenzelm
parents: 26894
diff changeset
    34
26742
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
    35
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
    36
(* ML text *)
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
    37
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
    38
local
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
    39
55831
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    40
val ml_toks = ML_Lex.read Position.none;
21375
ae8a112b62d7 Auxiliary antiquotations for Isabelle manuals.
wenzelm
parents:
diff changeset
    41
55831
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    42
fun ml_val (toks1, []) = ml_toks "fn _ => (" @ toks1 @ ml_toks ");"
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    43
  | ml_val (toks1, toks2) = ml_toks "fn _ => (" @ toks1 @ ml_toks " : " @ toks2 @ ml_toks ");";
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    44
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    45
fun ml_op (toks1, []) = ml_toks "fn _ => (op " @ toks1 @ ml_toks ");"
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    46
  | ml_op (toks1, toks2) = ml_toks "fn _ => (op " @ toks1 @ ml_toks " : " @ toks2 @ ml_toks ");";
46261
b03897da3c90 document antiquotations for ML infix operators;
wenzelm
parents: 45675
diff changeset
    47
55831
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    48
fun ml_type (toks1, []) = ml_toks "val _ = NONE : (" @ toks1 @ ml_toks ") option;"
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    49
  | ml_type (toks1, toks2) =
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    50
      ml_toks "val _ = [NONE : (" @ toks1 @ ml_toks ") option, NONE : (" @
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    51
        toks2 @ ml_toks ") option];";
22289
41ce4f5c97c9 added antiquotation for exceptions
haftmann
parents: 22094
diff changeset
    52
55837
154855d9a564 clarified names of antiquotations and markup;
wenzelm
parents: 55831
diff changeset
    53
fun ml_exception (toks1, []) = ml_toks "fn _ => (" @ toks1 @ ml_toks " : exn);"
154855d9a564 clarified names of antiquotations and markup;
wenzelm
parents: 55831
diff changeset
    54
  | ml_exception (toks1, toks2) =
55831
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    55
      ml_toks "fn _ => (" @ toks1 @ ml_toks " : " @ toks2 @ ml_toks " -> exn);";
22289
41ce4f5c97c9 added antiquotation for exceptions
haftmann
parents: 22094
diff changeset
    56
55831
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    57
fun ml_structure (toks, _) =
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    58
  ml_toks "functor XXX() = struct structure XX = " @ toks @ ml_toks " end;";
21375
ae8a112b62d7 Auxiliary antiquotations for Isabelle manuals.
wenzelm
parents:
diff changeset
    59
55831
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    60
fun ml_functor (Antiquote.Text tok :: _, _) =
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    61
      ml_toks "ML_Env.check_functor " @ ml_toks (ML_Syntax.print_string (ML_Lex.content_of tok))
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    62
  | ml_functor _ = raise Fail "Bad ML functor specification";
21375
ae8a112b62d7 Auxiliary antiquotations for Isabelle manuals.
wenzelm
parents:
diff changeset
    63
39869
c269f6bd0a1f more robust index_ML antiquotations: guess name from text (affects infixes and type constructors in particular);
wenzelm
parents: 39858
diff changeset
    64
val is_name = ML_Lex.kind_of #> (fn kind => kind = ML_Lex.Ident orelse kind = ML_Lex.LongIdent);
c269f6bd0a1f more robust index_ML antiquotations: guess name from text (affects infixes and type constructors in particular);
wenzelm
parents: 39858
diff changeset
    65
c269f6bd0a1f more robust index_ML antiquotations: guess name from text (affects infixes and type constructors in particular);
wenzelm
parents: 39858
diff changeset
    66
fun ml_name txt =
c269f6bd0a1f more robust index_ML antiquotations: guess name from text (affects infixes and type constructors in particular);
wenzelm
parents: 39858
diff changeset
    67
  (case filter is_name (ML_Lex.tokenize txt) of
c269f6bd0a1f more robust index_ML antiquotations: guess name from text (affects infixes and type constructors in particular);
wenzelm
parents: 39858
diff changeset
    68
    toks as [_] => ML_Lex.flatten toks
c269f6bd0a1f more robust index_ML antiquotations: guess name from text (affects infixes and type constructors in particular);
wenzelm
parents: 39858
diff changeset
    69
  | _ => error ("Single ML name expected in input: " ^ quote txt));
c269f6bd0a1f more robust index_ML antiquotations: guess name from text (affects infixes and type constructors in particular);
wenzelm
parents: 39858
diff changeset
    70
55831
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    71
fun prep_ml source =
59066
45ab32a542fe tuned signature;
wenzelm
parents: 59064
diff changeset
    72
  (Input.source_content source, ML_Lex.read_source false source);
55831
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    73
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 37198
diff changeset
    74
fun index_ml name kind ml = Thy_Output.antiquotation name
58069
0255436b3d85 more liberal embedded "text", which includes cartouches;
wenzelm
parents: 57918
diff changeset
    75
  (Scan.lift (Args.text_source_position -- Scan.option (Args.colon |-- Args.text_source_position)))
55831
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    76
  (fn {context = ctxt, ...} => fn (source1, opt_source2) =>
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    77
    let
55831
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    78
      val (txt1, toks1) = prep_ml source1;
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    79
      val (txt2, toks2) =
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    80
        (case opt_source2 of
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    81
          SOME source => prep_ml source
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    82
        | NONE => ("", []));
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    83
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    84
      val txt =
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    85
        if txt2 = "" then txt1
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    86
        else if kind = "type" then txt1 ^ " = " ^ txt2
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    87
        else if kind = "exception" then txt1 ^ " of " ^ txt2
50239
fb579401dc26 tuned signature;
wenzelm
parents: 50201
diff changeset
    88
        else if Symbol_Pos.is_identifier (Long_Name.base_name (ml_name txt1))
42290
b1f544c84040 discontinued special treatment of structure Lexicon;
wenzelm
parents: 40801
diff changeset
    89
        then txt1 ^ ": " ^ txt2
39858
5be7a57c3b4e more robust treatment of symbolic indentifiers (which may contain colons);
wenzelm
parents: 39829
diff changeset
    90
        else txt1 ^ " : " ^ txt2;
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    91
      val txt' = if kind = "" then txt else kind ^ " " ^ txt;
55831
3a9386b32211 more markup for ML source;
wenzelm
parents: 55828
diff changeset
    92
59064
a8bcb5a446c8 more abstract type Input.source;
wenzelm
parents: 58978
diff changeset
    93
      val pos = Input.pos_of source1;
56275
600f432ab556 added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents: 56208
diff changeset
    94
      val _ =
57477
c3b5cb53ea79 redundant error position, to ensure the message is attached somewhere, despite the distortion of positions due to glued tokens;
wenzelm
parents: 57334
diff changeset
    95
        ML_Context.eval_in (SOME ctxt) ML_Compiler.flags pos (ml (toks1, toks2))
c3b5cb53ea79 redundant error position, to ensure the message is attached somewhere, despite the distortion of positions due to glued tokens;
wenzelm
parents: 57334
diff changeset
    96
          handle ERROR msg => error (msg ^ Position.here pos);
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    97
      val kind' = if kind = "" then "ML" else "ML " ^ kind;
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    98
    in
39869
c269f6bd0a1f more robust index_ML antiquotations: guess name from text (affects infixes and type constructors in particular);
wenzelm
parents: 39858
diff changeset
    99
      "\\indexdef{}{" ^ kind' ^ "}{" ^ clean_string (ml_name txt1) ^ "}" ^
58716
23a380cc45f4 official support for "tt" style variants, avoid fragile \verb in LaTeX;
wenzelm
parents: 58069
diff changeset
   100
      (Thy_Output.verbatim_text ctxt txt')
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   101
    end);
26742
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
   102
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
   103
in
21375
ae8a112b62d7 Auxiliary antiquotations for Isabelle manuals.
wenzelm
parents:
diff changeset
   104
56059
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   105
val _ =
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   106
  Theory.setup
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   107
   (index_ml @{binding index_ML} "" ml_val #>
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   108
    index_ml @{binding index_ML_op} "infix" ml_op #>
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   109
    index_ml @{binding index_ML_type} "type" ml_type #>
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   110
    index_ml @{binding index_ML_exception} "exception" ml_exception #>
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   111
    index_ml @{binding index_ML_structure} "structure" ml_structure #>
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   112
    index_ml @{binding index_ML_functor} "functor" ml_functor);
21375
ae8a112b62d7 Auxiliary antiquotations for Isabelle manuals.
wenzelm
parents:
diff changeset
   113
26742
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
   114
end;
21375
ae8a112b62d7 Auxiliary antiquotations for Isabelle manuals.
wenzelm
parents:
diff changeset
   115
23846
bfedd1a024fc Added named_thms antiquotation.
berghofe
parents: 23651
diff changeset
   116
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   117
(* named theorems *)
23846
bfedd1a024fc Added named_thms antiquotation.
berghofe
parents: 23651
diff changeset
   118
56059
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   119
val _ =
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   120
  Theory.setup (Thy_Output.antiquotation @{binding named_thms}
43564
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43563
diff changeset
   121
    (Scan.repeat (Attrib.thm -- Scan.lift (Args.parens Args.name)))
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43563
diff changeset
   122
    (fn {context = ctxt, ...} =>
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43563
diff changeset
   123
      map (apfst (Thy_Output.pretty_thm ctxt))
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43563
diff changeset
   124
      #> (if Config.get ctxt Thy_Output.quotes then map (apfst Pretty.quote) else I)
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43563
diff changeset
   125
      #> (if Config.get ctxt Thy_Output.display
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43563
diff changeset
   126
          then
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43563
diff changeset
   127
            map (fn (p, name) =>
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43563
diff changeset
   128
              Output.output (Pretty.string_of (Pretty.indent (Config.get ctxt Thy_Output.indent) p)) ^
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43563
diff changeset
   129
              "\\rulename{" ^ Output.output (Pretty.str_of (Thy_Output.pretty_text ctxt name)) ^ "}")
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43563
diff changeset
   130
            #> space_implode "\\par\\smallskip%\n"
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43563
diff changeset
   131
            #> enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43563
diff changeset
   132
          else
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43563
diff changeset
   133
            map (fn (p, name) =>
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43563
diff changeset
   134
              Output.output (Pretty.str_of p) ^
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43563
diff changeset
   135
              "\\rulename{" ^ Output.output (Pretty.str_of (Thy_Output.pretty_text ctxt name)) ^ "}")
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43563
diff changeset
   136
            #> space_implode "\\par\\smallskip%\n"
56059
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   137
            #> enclose "\\isa{" "}")));
26742
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
   138
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
   139
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   140
(* theory file *)
26742
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
   141
56059
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   142
val _ =
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   143
  Theory.setup (Thy_Output.antiquotation @{binding thy_file} (Scan.lift Args.name)
43564
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43563
diff changeset
   144
    (fn {context = ctxt, ...} =>
56208
06cc31dff138 clarifed module name;
wenzelm
parents: 56185
diff changeset
   145
      fn name => (Resources.check_thy Path.current name; Thy_Output.output ctxt [Pretty.str name])));
56059
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   146
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   147
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   148
(* Isabelle/jEdit elements *)
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   149
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   150
local
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   151
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   152
fun parse_named a (XML.Elem ((b, props), _)) =
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   153
      (case Properties.get props "NAME" of
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   154
        SOME name => if a = b then [name] else []
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   155
      | NONE => [])
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   156
  | parse_named _ _ = [];
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   157
57334
54c6b73774d2 formal check of jEdit actions;
wenzelm
parents: 56467
diff changeset
   158
val isabelle_jedit_actions =
56135
efa24d31e595 added ML antiquotation @{path};
wenzelm
parents: 56070
diff changeset
   159
  (case XML.parse (File.read @{path "~~/src/Tools/jEdit/src/actions.xml"}) of
56059
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   160
    XML.Elem (("ACTIONS", _), body) => maps (parse_named "ACTION") body
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   161
  | _ => []);
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   162
57334
54c6b73774d2 formal check of jEdit actions;
wenzelm
parents: 56467
diff changeset
   163
val isabelle_jedit_dockables =
56135
efa24d31e595 added ML antiquotation @{path};
wenzelm
parents: 56070
diff changeset
   164
  (case XML.parse (File.read @{path "~~/src/Tools/jEdit/src/dockables.xml"}) of
56059
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   165
    XML.Elem (("DOCKABLES", _), body) => maps (parse_named "DOCKABLE") body
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   166
  | _ => []);
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   167
57334
54c6b73774d2 formal check of jEdit actions;
wenzelm
parents: 56467
diff changeset
   168
val jedit_actions =
54c6b73774d2 formal check of jEdit actions;
wenzelm
parents: 56467
diff changeset
   169
  Lazy.lazy (fn () =>
54c6b73774d2 formal check of jEdit actions;
wenzelm
parents: 56467
diff changeset
   170
    (case Isabelle_System.bash_output
54c6b73774d2 formal check of jEdit actions;
wenzelm
parents: 56467
diff changeset
   171
      "unzip -p \"$JEDIT_HOME/dist/jedit.jar\" org/gjt/sp/jedit/actions.xml" of
54c6b73774d2 formal check of jEdit actions;
wenzelm
parents: 56467
diff changeset
   172
      (txt, 0) =>
54c6b73774d2 formal check of jEdit actions;
wenzelm
parents: 56467
diff changeset
   173
        (case XML.parse txt of
54c6b73774d2 formal check of jEdit actions;
wenzelm
parents: 56467
diff changeset
   174
          XML.Elem (("ACTIONS", _), body) => maps (parse_named "ACTION") body
54c6b73774d2 formal check of jEdit actions;
wenzelm
parents: 56467
diff changeset
   175
        | _ => [])
54c6b73774d2 formal check of jEdit actions;
wenzelm
parents: 56467
diff changeset
   176
    | (_, rc) => error ("Cannot unzip jedit.jar\nreturn code = " ^ string_of_int rc)));
56059
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   177
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   178
in
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   179
57334
54c6b73774d2 formal check of jEdit actions;
wenzelm
parents: 56467
diff changeset
   180
fun is_action a =
54c6b73774d2 formal check of jEdit actions;
wenzelm
parents: 56467
diff changeset
   181
  member (op =) isabelle_jedit_actions a orelse
54c6b73774d2 formal check of jEdit actions;
wenzelm
parents: 56467
diff changeset
   182
  member (op =) isabelle_jedit_dockables a orelse
54c6b73774d2 formal check of jEdit actions;
wenzelm
parents: 56467
diff changeset
   183
  member (op =) (Lazy.force jedit_actions) a;
56059
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   184
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   185
end;
26742
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
   186
26751
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   187
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   188
(* Isabelle/Isar entities (with index) *)
26751
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   189
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   190
local
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   191
26894
1120f6cc10b0 proper checking of various Isar elements;
wenzelm
parents: 26853
diff changeset
   192
fun no_check _ _ = true;
1120f6cc10b0 proper checking of various Isar elements;
wenzelm
parents: 26853
diff changeset
   193
58928
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58923
diff changeset
   194
fun is_keyword ctxt (name, _) =
58931
3097ec653547 clarified keyword categories;
wenzelm
parents: 58928
diff changeset
   195
  Keyword.is_keyword (Thy_Header.get_keywords' ctxt) name;
58919
82a71046dce8 prefer explicit Keyword.keywords;
wenzelm
parents: 58903
diff changeset
   196
56294
85911b8a6868 prefer Context_Position where a context is available;
wenzelm
parents: 56278
diff changeset
   197
fun check_command ctxt (name, pos) =
58928
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58923
diff changeset
   198
  let
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58923
diff changeset
   199
    val thy = Proof_Context.theory_of ctxt;
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58923
diff changeset
   200
    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
   201
  in
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58923
diff changeset
   202
    Keyword.is_command keywords name andalso
58923
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58919
diff changeset
   203
      let
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58919
diff changeset
   204
        val markup =
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58919
diff changeset
   205
          Outer_Syntax.scan keywords Position.none name
58928
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58923
diff changeset
   206
          |> maps (Outer_Syntax.command_reports thy)
58923
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58919
diff changeset
   207
          |> map (snd o fst);
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58919
diff changeset
   208
        val _ = Context_Position.reports ctxt (map (pair pos) markup);
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58919
diff changeset
   209
      in true end
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58919
diff changeset
   210
  end;
53061
417cb0f713e0 more markup;
wenzelm
parents: 53045
diff changeset
   211
56467
8d7d6f17c6a7 more uniform ML/document antiquotations;
wenzelm
parents: 56304
diff changeset
   212
fun check_system_option ctxt (name, pos) =
8d7d6f17c6a7 more uniform ML/document antiquotations;
wenzelm
parents: 56304
diff changeset
   213
  (Context_Position.report ctxt pos (Options.default_markup (name, pos)); true)
8d7d6f17c6a7 more uniform ML/document antiquotations;
wenzelm
parents: 56304
diff changeset
   214
    handle ERROR _ => false;
8d7d6f17c6a7 more uniform ML/document antiquotations;
wenzelm
parents: 56304
diff changeset
   215
56294
85911b8a6868 prefer Context_Position where a context is available;
wenzelm
parents: 56278
diff changeset
   216
fun check_tool ctxt (name, pos) =
53044
be27b6be8027 more markup via Name_Space.check;
wenzelm
parents: 52408
diff changeset
   217
  let
be27b6be8027 more markup via Name_Space.check;
wenzelm
parents: 52408
diff changeset
   218
    fun tool dir =
be27b6be8027 more markup via Name_Space.check;
wenzelm
parents: 52408
diff changeset
   219
      let val path = Path.append dir (Path.basic name)
be27b6be8027 more markup via Name_Space.check;
wenzelm
parents: 52408
diff changeset
   220
      in if File.exists path then SOME path else NONE end;
be27b6be8027 more markup via Name_Space.check;
wenzelm
parents: 52408
diff changeset
   221
  in
53045
4c297ee47c28 check_tool wrt. official ISABELLE_TOOLS;
wenzelm
parents: 53044
diff changeset
   222
    (case get_first tool (Path.split (getenv "ISABELLE_TOOLS")) of
53044
be27b6be8027 more markup via Name_Space.check;
wenzelm
parents: 52408
diff changeset
   223
      NONE => false
56294
85911b8a6868 prefer Context_Position where a context is available;
wenzelm
parents: 56278
diff changeset
   224
    | SOME path => (Context_Position.report ctxt pos (Markup.path (Path.implode path)); true))
53044
be27b6be8027 more markup via Name_Space.check;
wenzelm
parents: 52408
diff changeset
   225
  end;
28237
f1fc11c73569 check setting and tool;
wenzelm
parents: 28218
diff changeset
   226
26751
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   227
val arg = enclose "{" "}" o clean_string;
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   228
56185
851c7b05eb92 more antiquotations;
wenzelm
parents: 56135
diff changeset
   229
fun entity check markup binding index =
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 37198
diff changeset
   230
  Thy_Output.antiquotation
56185
851c7b05eb92 more antiquotations;
wenzelm
parents: 56135
diff changeset
   231
    (binding |> Binding.map_name (fn name => name ^
43564
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43563
diff changeset
   232
      (case index of NONE => "" | SOME true => "_def" | SOME false => "_ref")))
53044
be27b6be8027 more markup via Name_Space.check;
wenzelm
parents: 52408
diff changeset
   233
    (Scan.lift (Scan.optional (Args.parens Args.name) "" -- Parse.position Args.name))
be27b6be8027 more markup via Name_Space.check;
wenzelm
parents: 52408
diff changeset
   234
    (fn {context = ctxt, ...} => fn (logic, (name, pos)) =>
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   235
      let
56185
851c7b05eb92 more antiquotations;
wenzelm
parents: 56135
diff changeset
   236
        val kind = translate (fn "_" => " " | c => c) (Binding.name_of binding);
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   237
        val hyper_name =
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   238
          "{" ^ Long_Name.append kind (Long_Name.append logic (clean_name name)) ^ "}";
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   239
        val hyper =
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   240
          enclose ("\\hyperlink" ^ hyper_name ^ "{") "}" #>
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   241
          index = SOME true ? enclose ("\\hypertarget" ^ hyper_name ^ "{") "}";
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   242
        val idx =
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   243
          (case index of
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   244
            NONE => ""
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   245
          | SOME is_def =>
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   246
              "\\index" ^ (if is_def then "def" else "ref") ^ arg logic ^ arg kind ^ arg name);
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   247
      in
53044
be27b6be8027 more markup via Name_Space.check;
wenzelm
parents: 52408
diff changeset
   248
        if check ctxt (name, pos) then
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   249
          idx ^
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   250
          (Output.output name
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   251
            |> (if markup = "" then I else enclose ("\\" ^ markup ^ "{") "}")
38767
d8da44a8dd25 proper context for various Thy_Output options, via official configuration options in ML and Isar;
wenzelm
parents: 37982
diff changeset
   252
            |> (if Config.get ctxt Thy_Output.quotes then quote else I)
d8da44a8dd25 proper context for various Thy_Output options, via official configuration options in ML and Isar;
wenzelm
parents: 37982
diff changeset
   253
            |> (if Config.get ctxt Thy_Output.display
d8da44a8dd25 proper context for various Thy_Output options, via official configuration options in ML and Isar;
wenzelm
parents: 37982
diff changeset
   254
                then enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   255
                else hyper o enclose "\\mbox{\\isa{" "}}"))
53044
be27b6be8027 more markup via Name_Space.check;
wenzelm
parents: 52408
diff changeset
   256
        else error ("Bad " ^ kind ^ " " ^ quote name ^ Position.here pos)
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   257
      end);
26897
044619358d3a clean_string: cover <;
wenzelm
parents: 26894
diff changeset
   258
26894
1120f6cc10b0 proper checking of various Isar elements;
wenzelm
parents: 26853
diff changeset
   259
fun entity_antiqs check markup kind =
43564
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43563
diff changeset
   260
  entity check markup kind NONE #>
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43563
diff changeset
   261
  entity check markup kind (SOME true) #>
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43563
diff changeset
   262
  entity check markup kind (SOME false);
26751
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   263
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   264
in
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   265
56059
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   266
val _ =
2390391584c2 some document antiquotations for Isabelle/jEdit elements;
wenzelm
parents: 55997
diff changeset
   267
  Theory.setup
56185
851c7b05eb92 more antiquotations;
wenzelm
parents: 56135
diff changeset
   268
   (entity_antiqs no_check "" @{binding syntax} #>
56294
85911b8a6868 prefer Context_Position where a context is available;
wenzelm
parents: 56278
diff changeset
   269
    entity_antiqs check_command "isacommand" @{binding command} #>
58919
82a71046dce8 prefer explicit Keyword.keywords;
wenzelm
parents: 58903
diff changeset
   270
    entity_antiqs is_keyword "isakeyword" @{binding keyword} #>
82a71046dce8 prefer explicit Keyword.keywords;
wenzelm
parents: 58903
diff changeset
   271
    entity_antiqs is_keyword "isakeyword" @{binding element} #>
56185
851c7b05eb92 more antiquotations;
wenzelm
parents: 56135
diff changeset
   272
    entity_antiqs (can o Method.check_name) "" @{binding method} #>
851c7b05eb92 more antiquotations;
wenzelm
parents: 56135
diff changeset
   273
    entity_antiqs (can o Attrib.check_name) "" @{binding attribute} #>
851c7b05eb92 more antiquotations;
wenzelm
parents: 56135
diff changeset
   274
    entity_antiqs no_check "" @{binding fact} #>
851c7b05eb92 more antiquotations;
wenzelm
parents: 56135
diff changeset
   275
    entity_antiqs no_check "" @{binding variable} #>
851c7b05eb92 more antiquotations;
wenzelm
parents: 56135
diff changeset
   276
    entity_antiqs no_check "" @{binding case} #>
851c7b05eb92 more antiquotations;
wenzelm
parents: 56135
diff changeset
   277
    entity_antiqs (can o Thy_Output.check_command) "" @{binding antiquotation} #>
851c7b05eb92 more antiquotations;
wenzelm
parents: 56135
diff changeset
   278
    entity_antiqs (can o Thy_Output.check_option) "" @{binding antiquotation_option} #>
58716
23a380cc45f4 official support for "tt" style variants, avoid fragile \verb in LaTeX;
wenzelm
parents: 58069
diff changeset
   279
    entity_antiqs no_check "isasystem" @{binding setting} #>
23a380cc45f4 official support for "tt" style variants, avoid fragile \verb in LaTeX;
wenzelm
parents: 58069
diff changeset
   280
    entity_antiqs check_system_option "isasystem" @{binding system_option} #>
56185
851c7b05eb92 more antiquotations;
wenzelm
parents: 56135
diff changeset
   281
    entity_antiqs no_check "" @{binding inference} #>
58716
23a380cc45f4 official support for "tt" style variants, avoid fragile \verb in LaTeX;
wenzelm
parents: 58069
diff changeset
   282
    entity_antiqs no_check "isasystem" @{binding executable} #>
56294
85911b8a6868 prefer Context_Position where a context is available;
wenzelm
parents: 56278
diff changeset
   283
    entity_antiqs check_tool "isatool" @{binding tool} #>
56185
851c7b05eb92 more antiquotations;
wenzelm
parents: 56135
diff changeset
   284
    entity_antiqs (can o ML_Context.check_antiquotation) "" @{binding ML_antiquotation} #>
58716
23a380cc45f4 official support for "tt" style variants, avoid fragile \verb in LaTeX;
wenzelm
parents: 58069
diff changeset
   285
    entity_antiqs (K (is_action o #1)) "isasystem" @{binding action});
26751
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   286
26742
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
   287
end;
26751
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   288
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   289
end;