src/Pure/ML/ml_antiquotation.ML
author wenzelm
Thu, 22 Nov 2018 20:23:47 +0100
changeset 69329 8bbde4dba926
parent 67146 909dcdec2122
child 69349 7cef9e386ffe
permissions -rw-r--r--
tuned error; tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
56072
31e427387ab5 tuned signature -- clarified module name;
wenzelm
parents: 56071
diff changeset
     1
(*  Title:      Pure/ML/ml_antiquotation.ML
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
     3
56205
ceb8a93460b7 clarified modules;
wenzelm
parents: 56072
diff changeset
     4
ML antiquotations.
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
     5
*)
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
     6
56072
31e427387ab5 tuned signature -- clarified module name;
wenzelm
parents: 56071
diff changeset
     7
signature ML_ANTIQUOTATION =
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
     8
sig
56205
ceb8a93460b7 clarified modules;
wenzelm
parents: 56072
diff changeset
     9
  val declaration: binding -> 'a context_parser ->
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 57832
diff changeset
    10
    (Token.src -> 'a -> Proof.context -> ML_Context.decl * Proof.context) ->
56205
ceb8a93460b7 clarified modules;
wenzelm
parents: 56072
diff changeset
    11
    theory -> theory
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
    12
  val inline: binding -> string context_parser -> theory -> theory
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
    13
  val value: binding -> string context_parser -> theory -> theory
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    14
end;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    15
56072
31e427387ab5 tuned signature -- clarified module name;
wenzelm
parents: 56071
diff changeset
    16
structure ML_Antiquotation: ML_ANTIQUOTATION =
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    17
struct
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    18
56205
ceb8a93460b7 clarified modules;
wenzelm
parents: 56072
diff changeset
    19
(* define antiquotations *)
ceb8a93460b7 clarified modules;
wenzelm
parents: 56072
diff changeset
    20
ceb8a93460b7 clarified modules;
wenzelm
parents: 56072
diff changeset
    21
fun declaration name scan body =
ceb8a93460b7 clarified modules;
wenzelm
parents: 56072
diff changeset
    22
  ML_Context.add_antiquotation name
ceb8a93460b7 clarified modules;
wenzelm
parents: 56072
diff changeset
    23
    (fn src => fn orig_ctxt =>
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 57832
diff changeset
    24
      let val (x, _) = Token.syntax scan src orig_ctxt
56205
ceb8a93460b7 clarified modules;
wenzelm
parents: 56072
diff changeset
    25
      in body src x orig_ctxt end);
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    26
53169
e2d31807cbf6 clarified type ML_Context.antiq: context parser maintains compilation context, declaration is applied to final context;
wenzelm
parents: 53163
diff changeset
    27
fun inline name scan =
56205
ceb8a93460b7 clarified modules;
wenzelm
parents: 56072
diff changeset
    28
  declaration name scan (fn _ => fn s => fn ctxt => (K ("", s), ctxt));
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    29
53169
e2d31807cbf6 clarified type ML_Context.antiq: context parser maintains compilation context, declaration is applied to final context;
wenzelm
parents: 53163
diff changeset
    30
fun value name scan =
59112
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 58011
diff changeset
    31
  declaration name scan (fn _ => ML_Context.value_decl (Binding.name_of name));
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    32
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    33
56205
ceb8a93460b7 clarified modules;
wenzelm
parents: 56072
diff changeset
    34
(* basic antiquotations *)
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    35
53171
a5e54d4d9081 added Theory.setup convenience;
wenzelm
parents: 53170
diff changeset
    36
val _ = Theory.setup
64556
851ae0e7b09c more symbols;
wenzelm
parents: 62900
diff changeset
    37
 (declaration (Binding.make ("here", \<^here>)) (Scan.succeed ())
59112
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 58011
diff changeset
    38
    (fn src => fn () =>
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 58011
diff changeset
    39
      ML_Context.value_decl "position" (ML_Syntax.print_position (#2 (Token.name_of_src src)))) #>
56071
2ffdedb0c044 added ML antiquotation @{here};
wenzelm
parents: 56069
diff changeset
    40
64556
851ae0e7b09c more symbols;
wenzelm
parents: 62900
diff changeset
    41
  inline (Binding.make ("make_string", \<^here>)) (Args.context >> K ML_Pretty.make_string_fn) #>
62662
291cc01f56f5 @{make_string} is available during Pure bootstrap;
wenzelm
parents: 59112
diff changeset
    42
64556
851ae0e7b09c more symbols;
wenzelm
parents: 62900
diff changeset
    43
  value (Binding.make ("binding", \<^here>))
67146
909dcdec2122 more embedded cartouche arguments;
wenzelm
parents: 64556
diff changeset
    44
    (Scan.lift (Parse.position Args.embedded) >> ML_Syntax.make_binding) #>
62850
1f1a2c33ccf4 clarified conditional compilation;
wenzelm
parents: 62662
diff changeset
    45
64556
851ae0e7b09c more symbols;
wenzelm
parents: 62900
diff changeset
    46
  value (Binding.make ("cartouche", \<^here>))
62850
1f1a2c33ccf4 clarified conditional compilation;
wenzelm
parents: 62662
diff changeset
    47
    (Scan.lift Args.cartouche_input >> (fn source =>
1f1a2c33ccf4 clarified conditional compilation;
wenzelm
parents: 62662
diff changeset
    48
      "Input.source true " ^ ML_Syntax.print_string (Input.text_of source) ^ " " ^
1f1a2c33ccf4 clarified conditional compilation;
wenzelm
parents: 62662
diff changeset
    49
        ML_Syntax.atomic (ML_Syntax.print_range (Input.range_of source)))));
46948
aae5566d6756 added ML antiquotation @{keyword};
wenzelm
parents: 43794
diff changeset
    50
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    51
end;