src/Pure/Thy/bibtex.ML
author wenzelm
Fri, 13 Jan 2023 15:57:11 +0100
changeset 76964 c044306db39f
parent 76963 a8566127d43b
child 76985 aafe49b63205
permissions -rw-r--r--
support embedded syntax, for use with control symbols;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
67274
4588f714a78a clarified directories;
wenzelm
parents: 67147
diff changeset
     1
(*  Title:      Pure/Thy/bibtex.ML
58544
340f130b3d38 bibtex support in ML: document antiquotation @{cite} with markup;
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
340f130b3d38 bibtex support in ML: document antiquotation @{cite} with markup;
wenzelm
parents:
diff changeset
     3
340f130b3d38 bibtex support in ML: document antiquotation @{cite} with markup;
wenzelm
parents:
diff changeset
     4
BibTeX support.
340f130b3d38 bibtex support in ML: document antiquotation @{cite} with markup;
wenzelm
parents:
diff changeset
     5
*)
340f130b3d38 bibtex support in ML: document antiquotation @{cite} with markup;
wenzelm
parents:
diff changeset
     6
340f130b3d38 bibtex support in ML: document antiquotation @{cite} with markup;
wenzelm
parents:
diff changeset
     7
signature BIBTEX =
340f130b3d38 bibtex support in ML: document antiquotation @{cite} with markup;
wenzelm
parents:
diff changeset
     8
sig
67301
e255c76db052 clarified signature;
wenzelm
parents: 67297
diff changeset
     9
  val check_database:
e255c76db052 clarified signature;
wenzelm
parents: 67297
diff changeset
    10
    Position.T -> string -> (string * Position.T) list * (string * Position.T) list
67275
5e427586cb57 check bibtex database on ML side -- for semantic PIDE editing;
wenzelm
parents: 67274
diff changeset
    11
  val check_database_output: Position.T -> string -> unit
58544
340f130b3d38 bibtex support in ML: document antiquotation @{cite} with markup;
wenzelm
parents:
diff changeset
    12
  val cite_macro: string Config.T
76959
3d491a0af6ef clarified signature: more generic operations;
wenzelm
parents: 76958
diff changeset
    13
  val cite_antiquotation: binding -> (Proof.context -> string) -> theory -> theory
58544
340f130b3d38 bibtex support in ML: document antiquotation @{cite} with markup;
wenzelm
parents:
diff changeset
    14
end;
340f130b3d38 bibtex support in ML: document antiquotation @{cite} with markup;
wenzelm
parents:
diff changeset
    15
340f130b3d38 bibtex support in ML: document antiquotation @{cite} with markup;
wenzelm
parents:
diff changeset
    16
structure Bibtex: BIBTEX =
340f130b3d38 bibtex support in ML: document antiquotation @{cite} with markup;
wenzelm
parents:
diff changeset
    17
struct
340f130b3d38 bibtex support in ML: document antiquotation @{cite} with markup;
wenzelm
parents:
diff changeset
    18
67275
5e427586cb57 check bibtex database on ML side -- for semantic PIDE editing;
wenzelm
parents: 67274
diff changeset
    19
(* check database *)
5e427586cb57 check bibtex database on ML side -- for semantic PIDE editing;
wenzelm
parents: 67274
diff changeset
    20
67301
e255c76db052 clarified signature;
wenzelm
parents: 67297
diff changeset
    21
type message = string * Position.T;
67275
5e427586cb57 check bibtex database on ML side -- for semantic PIDE editing;
wenzelm
parents: 67274
diff changeset
    22
5e427586cb57 check bibtex database on ML side -- for semantic PIDE editing;
wenzelm
parents: 67274
diff changeset
    23
fun check_database pos0 database =
72194
eef421b724c0 clarified names;
wenzelm
parents: 71881
diff changeset
    24
  \<^scala>\<open>bibtex_check_database\<close> database
67275
5e427586cb57 check bibtex database on ML side -- for semantic PIDE editing;
wenzelm
parents: 67274
diff changeset
    25
  |> YXML.parse_body
67301
e255c76db052 clarified signature;
wenzelm
parents: 67297
diff changeset
    26
  |> let open XML.Decode in pair (list (pair string properties)) (list (pair string properties)) end
e255c76db052 clarified signature;
wenzelm
parents: 67297
diff changeset
    27
  |> (apply2 o map o apsnd) (fn pos => Position.of_properties (pos @ Position.get_props pos0));
67275
5e427586cb57 check bibtex database on ML side -- for semantic PIDE editing;
wenzelm
parents: 67274
diff changeset
    28
5e427586cb57 check bibtex database on ML side -- for semantic PIDE editing;
wenzelm
parents: 67274
diff changeset
    29
fun check_database_output pos0 database =
67301
e255c76db052 clarified signature;
wenzelm
parents: 67297
diff changeset
    30
  let val (errors, warnings) = check_database pos0 database in
e255c76db052 clarified signature;
wenzelm
parents: 67297
diff changeset
    31
    errors |> List.app (fn (msg, pos) =>
67275
5e427586cb57 check bibtex database on ML side -- for semantic PIDE editing;
wenzelm
parents: 67274
diff changeset
    32
      Output.error_message ("Bibtex error" ^ Position.here pos ^ ":\n  " ^ msg));
67301
e255c76db052 clarified signature;
wenzelm
parents: 67297
diff changeset
    33
    warnings |> List.app (fn (msg, pos) =>
67275
5e427586cb57 check bibtex database on ML side -- for semantic PIDE editing;
wenzelm
parents: 67274
diff changeset
    34
      warning ("Bibtex warning" ^ Position.here pos ^ ":\n  " ^ msg))
5e427586cb57 check bibtex database on ML side -- for semantic PIDE editing;
wenzelm
parents: 67274
diff changeset
    35
  end;
5e427586cb57 check bibtex database on ML side -- for semantic PIDE editing;
wenzelm
parents: 67274
diff changeset
    36
5e427586cb57 check bibtex database on ML side -- for semantic PIDE editing;
wenzelm
parents: 67274
diff changeset
    37
5e427586cb57 check bibtex database on ML side -- for semantic PIDE editing;
wenzelm
parents: 67274
diff changeset
    38
(* document antiquotations *)
5e427586cb57 check bibtex database on ML side -- for semantic PIDE editing;
wenzelm
parents: 67274
diff changeset
    39
76961
d756f4f78dc7 clarified default: final value is provided in Isabelle/Scala Latex.Cite.unapply;
wenzelm
parents: 76960
diff changeset
    40
val cite_macro = Attrib.setup_config_string \<^binding>\<open>cite_macro\<close> (K "");
76963
wenzelm
parents: 76961
diff changeset
    41
fun get_cite_macro ctxt = Config.get ctxt cite_macro;
wenzelm
parents: 76961
diff changeset
    42
wenzelm
parents: 76961
diff changeset
    43
val _ =
wenzelm
parents: 76961
diff changeset
    44
  Theory.setup (Document_Antiquotation.setup_option \<^binding>\<open>cite_macro\<close> (Config.put cite_macro));
58544
340f130b3d38 bibtex support in ML: document antiquotation @{cite} with markup;
wenzelm
parents:
diff changeset
    45
76964
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    46
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    47
local
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    48
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    49
val parse_citations = Parse.and_list1 Args.name_position;
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    50
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    51
fun cite_command ctxt get_kind (opt_loc, citations) =
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    52
  let
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    53
    val loc = the_default Input.empty opt_loc;
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    54
    val _ =
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    55
      Context_Position.reports ctxt
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    56
        (Document_Output.document_reports loc @
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    57
          map (fn (name, pos) => (pos, Markup.citation name)) citations);
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    58
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    59
    val thy_name = Context.theory_long_name (Proof_Context.theory_of ctxt);
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    60
    val bibtex_entries = Resources.theory_bibtex_entries thy_name;
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    61
    val _ =
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    62
      if null bibtex_entries andalso thy_name <> Context.PureN then ()
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    63
      else
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    64
        citations |> List.app (fn (name, pos) =>
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    65
          if member (op =) bibtex_entries name orelse name = "*" then ()
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    66
          else error ("Unknown Bibtex entry " ^ quote name ^ Position.here pos));
76959
3d491a0af6ef clarified signature: more generic operations;
wenzelm
parents: 76958
diff changeset
    67
76964
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    68
    val kind = get_kind ctxt;
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    69
    val location = Document_Output.output_document ctxt {markdown = false} loc;
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    70
  in Latex.cite {kind = kind, citations = citations, location = location} end;
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    71
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    72
fun cite_command_embedded ctxt get_kind input =
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    73
  let
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    74
    val keywords = Keyword.no_major_keywords (Thy_Header.get_keywords' ctxt);
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    75
    val parser = Scan.option (Parse.embedded_input --| Parse.$$$ "in") -- parse_citations;
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    76
    val args = Parse.read_embedded ctxt keywords parser input;
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    77
  in cite_command ctxt get_kind args end;
76959
3d491a0af6ef clarified signature: more generic operations;
wenzelm
parents: 76958
diff changeset
    78
76964
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    79
fun cite_command_parser get_kind =
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    80
  Scan.option Args.cartouche_input -- parse_citations
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    81
    >> (fn args => fn ctxt => cite_command ctxt get_kind args) ||
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    82
  Parse.embedded_input
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    83
    >> (fn arg => fn ctxt => cite_command_embedded ctxt get_kind arg);
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    84
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    85
in
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    86
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    87
fun cite_antiquotation binding get_kind =
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    88
  Document_Output.antiquotation_raw binding (Scan.lift (cite_command_parser get_kind))
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    89
    (fn ctxt => fn cmd => cmd ctxt);
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    90
c044306db39f support embedded syntax, for use with control symbols;
wenzelm
parents: 76963
diff changeset
    91
end;
76959
3d491a0af6ef clarified signature: more generic operations;
wenzelm
parents: 76958
diff changeset
    92
58544
340f130b3d38 bibtex support in ML: document antiquotation @{cite} with markup;
wenzelm
parents:
diff changeset
    93
val _ =
340f130b3d38 bibtex support in ML: document antiquotation @{cite} with markup;
wenzelm
parents:
diff changeset
    94
  Theory.setup
76963
wenzelm
parents: 76961
diff changeset
    95
   (cite_antiquotation \<^binding>\<open>cite\<close> get_cite_macro #>
76960
6c623c517a6e more "cite" antiquotations;
wenzelm
parents: 76959
diff changeset
    96
    cite_antiquotation \<^binding>\<open>nocite\<close> (K "nocite") #>
6c623c517a6e more "cite" antiquotations;
wenzelm
parents: 76959
diff changeset
    97
    cite_antiquotation \<^binding>\<open>citet\<close> (K "citet") #>
6c623c517a6e more "cite" antiquotations;
wenzelm
parents: 76959
diff changeset
    98
    cite_antiquotation \<^binding>\<open>citep\<close> (K "citep"));
58544
340f130b3d38 bibtex support in ML: document antiquotation @{cite} with markup;
wenzelm
parents:
diff changeset
    99
340f130b3d38 bibtex support in ML: document antiquotation @{cite} with markup;
wenzelm
parents:
diff changeset
   100
end;