doc-src/antiquote_setup.ML
author wenzelm
Sun, 30 May 2010 21:34:19 +0200
changeset 37198 3af985b10550
parent 36973 b0033a307d1f
child 37216 3165bc303f66
permissions -rw-r--r--
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information; ML_Context.eval/expression expect explicit ML_Lex source, which allows surrounding further text without loosing position information; fall back on ML_Context.eval_text if there is no position or no surrounding text; proper Args.name_source_position for method "tactic" and "raw_tactic"; tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
     1
(*  Title:      doc-src/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
26742
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
     7
structure AntiquoteSetup: sig end =
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
    | "#" => "\\#"
26897
044619358d3a clean_string: cover <;
wenzelm
parents: 26894
diff changeset
    17
    | "<" => "$<$"
26768
844068d16ba0 clean_string: handle { };
wenzelm
parents: 26756
diff changeset
    18
    | ">" => "$>$"
844068d16ba0 clean_string: handle { };
wenzelm
parents: 26756
diff changeset
    19
    | "{" => "\\{"
30120
aaa4667285c8 uniform treatment of ML indexing, using general \indexdef macro for formal Isabelle/Isar entities;
wenzelm
parents: 29736
diff changeset
    20
    | "|" => "$\\mid$"
26768
844068d16ba0 clean_string: handle { };
wenzelm
parents: 26756
diff changeset
    21
    | "}" => "\\}"
31546
d58d6acab331 eliminated escaped symbols;
wenzelm
parents: 31297
diff changeset
    22
    | "\<dash>" => "-"
26768
844068d16ba0 clean_string: handle { };
wenzelm
parents: 26756
diff changeset
    23
    | c => c);
26751
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
    24
31546
d58d6acab331 eliminated escaped symbols;
wenzelm
parents: 31297
diff changeset
    25
fun clean_name "\<dots>" = "dots"
26897
044619358d3a clean_string: cover <;
wenzelm
parents: 26894
diff changeset
    26
  | clean_name ".." = "ddot"
044619358d3a clean_string: cover <;
wenzelm
parents: 26894
diff changeset
    27
  | clean_name "." = "dot"
26910
aa6357b39212 tuned clean_name (underscore);
wenzelm
parents: 26903
diff changeset
    28
  | clean_name "_" = "underscore"
26897
044619358d3a clean_string: cover <;
wenzelm
parents: 26894
diff changeset
    29
  | clean_name "{" = "braceleft"
044619358d3a clean_string: cover <;
wenzelm
parents: 26894
diff changeset
    30
  | clean_name "}" = "braceright"
31546
d58d6acab331 eliminated escaped symbols;
wenzelm
parents: 31297
diff changeset
    31
  | clean_name s = s |> translate (fn "_" => "-" | "\<dash>" => "-" | c => c);
26897
044619358d3a clean_string: cover <;
wenzelm
parents: 26894
diff changeset
    32
26742
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
    33
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
    34
(* verbatim text *)
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
    35
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
    36
val verbatim = space_implode "\\verb,|," o map (enclose "\\verb|" "|") o space_explode "|";
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
    37
30382
910290f04692 adapted ThyOutput.antiquotation;
wenzelm
parents: 30369
diff changeset
    38
val _ = ThyOutput.antiquotation "verbatim" (Scan.lift Args.name)
910290f04692 adapted ThyOutput.antiquotation;
wenzelm
parents: 30369
diff changeset
    39
  (K (split_lines #> map verbatim #> space_implode "\\isasep\\isanewline%\n"));
26742
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
    40
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
    41
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
    42
(* ML text *)
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
    43
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
    44
local
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
    45
21375
ae8a112b62d7 Auxiliary antiquotations for Isabelle manuals.
wenzelm
parents:
diff changeset
    46
fun ml_val (txt1, "") = "fn _ => (" ^ txt1 ^ ");"
ae8a112b62d7 Auxiliary antiquotations for Isabelle manuals.
wenzelm
parents:
diff changeset
    47
  | ml_val (txt1, txt2) = "fn _ => (" ^ txt1 ^ ": " ^ txt2 ^ ");";
ae8a112b62d7 Auxiliary antiquotations for Isabelle manuals.
wenzelm
parents:
diff changeset
    48
ae8a112b62d7 Auxiliary antiquotations for Isabelle manuals.
wenzelm
parents:
diff changeset
    49
fun ml_type (txt1, "") = "val _ = NONE : (" ^ txt1 ^ ") option;"
ae8a112b62d7 Auxiliary antiquotations for Isabelle manuals.
wenzelm
parents:
diff changeset
    50
  | ml_type (txt1, txt2) = "val _ = [NONE : (" ^ txt1 ^ ") option, NONE : (" ^ txt2 ^ ") option];";
22289
41ce4f5c97c9 added antiquotation for exceptions
haftmann
parents: 22094
diff changeset
    51
23651
6e0b8b6012c9 renamed ML_exc to ML_exn;
wenzelm
parents: 22289
diff changeset
    52
fun ml_exn (txt1, "") = "fn _ => (" ^ txt1 ^ ": exn);"
6e0b8b6012c9 renamed ML_exc to ML_exn;
wenzelm
parents: 22289
diff changeset
    53
  | ml_exn (txt1, txt2) = "fn _ => (" ^ txt1 ^ ": " ^ txt2 ^ " -> exn);";
22289
41ce4f5c97c9 added antiquotation for exceptions
haftmann
parents: 22094
diff changeset
    54
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    55
fun ml_structure (txt, _) = "functor XXX() = struct structure XX = " ^ txt ^ " end;";
21375
ae8a112b62d7 Auxiliary antiquotations for Isabelle manuals.
wenzelm
parents:
diff changeset
    56
36163
823c9400eb62 proper checking of ML functors (in Poly/ML 5.2 or later);
wenzelm
parents: 31546
diff changeset
    57
fun ml_functor (txt, _) = "ML_Env.check_functor " ^ ML_Syntax.print_string txt;
21375
ae8a112b62d7 Auxiliary antiquotations for Isabelle manuals.
wenzelm
parents:
diff changeset
    58
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    59
fun index_ml name kind ml = ThyOutput.antiquotation name
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    60
  (Scan.lift (Args.name -- Scan.optional (Args.colon |-- Args.name) ""))
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    61
  (fn {context = ctxt, ...} => fn (txt1, txt2) =>
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    62
    let
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    63
      val txt =
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    64
        if txt2 = "" then txt1
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    65
        else if kind = "type" then txt1 ^ " = " ^ txt2
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    66
        else if kind = "exception" then txt1 ^ " of " ^ txt2
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    67
        else txt1 ^ ": " ^ txt2;
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    68
      val txt' = if kind = "" then txt else kind ^ " " ^ txt;
37198
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 36973
diff changeset
    69
      val _ = ML_Context.eval_text_in (SOME ctxt) false Position.none (ml (txt1, txt2));  (* ML_Lex.read (!?) *)
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    70
      val kind' = if kind = "" then "ML" else "ML " ^ kind;
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    71
    in
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    72
      "\\indexdef{}{" ^ kind' ^ "}{" ^ clean_string txt1 ^ "}" ^
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    73
      (txt'
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    74
      |> (if ! ThyOutput.quotes then quote else I)
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    75
      |> (if ! ThyOutput.display then enclose "\\begin{verbatim}\n" "\n\\end{verbatim}"
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    76
          else split_lines #> map verbatim #> space_implode "\\isasep\\isanewline%\n"))
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    77
    end);
26742
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
    78
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
    79
in
21375
ae8a112b62d7 Auxiliary antiquotations for Isabelle manuals.
wenzelm
parents:
diff changeset
    80
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    81
val _ = index_ml "index_ML" "" ml_val;
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    82
val _ = index_ml "index_ML_type" "type" ml_type;
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    83
val _ = index_ml "index_ML_exn" "exception" ml_exn;
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    84
val _ = index_ml "index_ML_structure" "structure" ml_structure;
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    85
val _ = index_ml "index_ML_functor" "functor" ml_functor;
21375
ae8a112b62d7 Auxiliary antiquotations for Isabelle manuals.
wenzelm
parents:
diff changeset
    86
26742
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
    87
end;
21375
ae8a112b62d7 Auxiliary antiquotations for Isabelle manuals.
wenzelm
parents:
diff changeset
    88
23846
bfedd1a024fc Added named_thms antiquotation.
berghofe
parents: 23651
diff changeset
    89
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    90
(* named theorems *)
23846
bfedd1a024fc Added named_thms antiquotation.
berghofe
parents: 23651
diff changeset
    91
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    92
val _ = ThyOutput.antiquotation "named_thms"
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    93
  (Scan.repeat (Attrib.thm -- Scan.lift (Args.parens Args.name)))
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    94
  (fn {context = ctxt, ...} =>
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    95
    map (apfst (ThyOutput.pretty_thm ctxt))
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    96
    #> (if ! ThyOutput.quotes then map (apfst Pretty.quote) else I)
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    97
    #> (if ! ThyOutput.display
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    98
        then
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
    99
          map (fn (p, name) =>
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   100
            Output.output (Pretty.string_of (Pretty.indent (! ThyOutput.indent) p)) ^
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   101
            "\\rulename{" ^ Output.output (Pretty.str_of (ThyOutput.pretty_text name)) ^ "}")
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   102
          #> space_implode "\\par\\smallskip%\n"
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   103
          #> enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   104
        else
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   105
          map (fn (p, name) =>
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   106
            Output.output (Pretty.str_of p) ^
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   107
            "\\rulename{" ^ Output.output (Pretty.str_of (ThyOutput.pretty_text name)) ^ "}")
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   108
          #> space_implode "\\par\\smallskip%\n"
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   109
          #> enclose "\\isa{" "}"));
26742
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
   110
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
   111
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   112
(* theory file *)
26742
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
   113
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   114
val _ = ThyOutput.antiquotation "thy_file" (Scan.lift Args.name)
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   115
  (fn _ => fn name => (ThyLoad.check_thy Path.current name; ThyOutput.output [Pretty.str name]));
26742
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
   116
26751
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   117
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   118
(* Isabelle/Isar entities (with index) *)
26751
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   119
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   120
local
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   121
26894
1120f6cc10b0 proper checking of various Isar elements;
wenzelm
parents: 26853
diff changeset
   122
fun no_check _ _ = true;
1120f6cc10b0 proper checking of various Isar elements;
wenzelm
parents: 26853
diff changeset
   123
1120f6cc10b0 proper checking of various Isar elements;
wenzelm
parents: 26853
diff changeset
   124
fun thy_check intern defined ctxt =
1120f6cc10b0 proper checking of various Isar elements;
wenzelm
parents: 26853
diff changeset
   125
  let val thy = ProofContext.theory_of ctxt
1120f6cc10b0 proper checking of various Isar elements;
wenzelm
parents: 26853
diff changeset
   126
  in defined thy o intern thy end;
1120f6cc10b0 proper checking of various Isar elements;
wenzelm
parents: 26853
diff changeset
   127
28237
f1fc11c73569 check setting and tool;
wenzelm
parents: 28218
diff changeset
   128
fun check_tool name =
f1fc11c73569 check setting and tool;
wenzelm
parents: 28218
diff changeset
   129
  File.exists (Path.append (Path.explode "~~/lib/Tools") (Path.basic name));
f1fc11c73569 check setting and tool;
wenzelm
parents: 28218
diff changeset
   130
26751
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   131
val arg = enclose "{" "}" o clean_string;
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   132
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   133
fun entity check markup kind index =
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   134
  ThyOutput.antiquotation
30396
841ce0fcbe14 fornal markup for antiquotation options;
wenzelm
parents: 30394
diff changeset
   135
    (translate (fn " " => "_" | c => c) kind ^
841ce0fcbe14 fornal markup for antiquotation options;
wenzelm
parents: 30394
diff changeset
   136
      (case index of NONE => "" | SOME true => "_def" | SOME false => "_ref"))
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   137
    (Scan.lift (Scan.optional (Args.parens Args.name) "" -- Args.name))
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   138
    (fn {context = ctxt, ...} => fn (logic, name) =>
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   139
      let
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   140
        val hyper_name =
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   141
          "{" ^ 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
   142
        val hyper =
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   143
          enclose ("\\hyperlink" ^ hyper_name ^ "{") "}" #>
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   144
          index = SOME true ? enclose ("\\hypertarget" ^ hyper_name ^ "{") "}";
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   145
        val idx =
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   146
          (case index of
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   147
            NONE => ""
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   148
          | SOME is_def =>
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   149
              "\\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
   150
      in
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   151
        if check ctxt name then
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   152
          idx ^
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   153
          (Output.output name
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   154
            |> (if markup = "" then I else enclose ("\\" ^ markup ^ "{") "}")
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   155
            |> (if ! ThyOutput.quotes then quote else I)
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   156
            |> (if ! ThyOutput.display then enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   157
                else hyper o enclose "\\mbox{\\isa{" "}}"))
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   158
        else error ("Bad " ^ kind ^ " " ^ quote name)
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   159
      end);
26897
044619358d3a clean_string: cover <;
wenzelm
parents: 26894
diff changeset
   160
26894
1120f6cc10b0 proper checking of various Isar elements;
wenzelm
parents: 26853
diff changeset
   161
fun entity_antiqs check markup kind =
31297
wenzelm
parents: 30396
diff changeset
   162
 ((entity check markup kind NONE);
wenzelm
parents: 30396
diff changeset
   163
  (entity check markup kind (SOME true));
wenzelm
parents: 30396
diff changeset
   164
  (entity check markup kind (SOME false)));
26751
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   165
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   166
in
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   167
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   168
val _ = entity_antiqs no_check "" "syntax";
36973
b0033a307d1f prefer structure Keyword and Parse;
wenzelm
parents: 36163
diff changeset
   169
val _ = entity_antiqs (K (is_some o Keyword.command_keyword)) "isacommand" "command";
b0033a307d1f prefer structure Keyword and Parse;
wenzelm
parents: 36163
diff changeset
   170
val _ = entity_antiqs (K Keyword.is_keyword) "isakeyword" "keyword";
b0033a307d1f prefer structure Keyword and Parse;
wenzelm
parents: 36163
diff changeset
   171
val _ = entity_antiqs (K Keyword.is_keyword) "isakeyword" "element";
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   172
val _ = entity_antiqs (thy_check Method.intern Method.defined) "" "method";
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   173
val _ = entity_antiqs (thy_check Attrib.intern Attrib.defined) "" "attribute";
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   174
val _ = entity_antiqs no_check "" "fact";
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   175
val _ = entity_antiqs no_check "" "variable";
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   176
val _ = entity_antiqs no_check "" "case";
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   177
val _ = entity_antiqs (K ThyOutput.defined_command) "" "antiquotation";
30396
841ce0fcbe14 fornal markup for antiquotation options;
wenzelm
parents: 30394
diff changeset
   178
val _ = entity_antiqs (K ThyOutput.defined_option) "" "antiquotation option";
30394
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   179
val _ = entity_antiqs (fn _ => fn name => is_some (OS.Process.getEnv name)) "isatt" "setting";
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   180
val _ = entity_antiqs no_check "" "inference";
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   181
val _ = entity_antiqs no_check "isatt" "executable";
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   182
val _ = entity_antiqs (K check_tool) "isatt" "tool";
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   183
val _ = entity_antiqs (K (File.exists o Path.explode)) "isatt" "file";
c11a1e65a2ed moved @{ML_functor} and @{ML_text} to Pure;
wenzelm
parents: 30382
diff changeset
   184
val _ = entity_antiqs (K ThyInfo.known_thy) "" "theory";
26751
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   185
26742
5a86bc79431c misc cleanup;
wenzelm
parents: 26710
diff changeset
   186
end;
26751
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   187
2b97ea3130c2 added setup for Isar entities;
wenzelm
parents: 26742
diff changeset
   188
end;