src/Pure/ML/ml_antiquote.ML
author wenzelm
Thu, 18 Jul 2013 13:12:54 +0200
changeset 52696 38466f4f3483
parent 51979 4f3a5f4c1169
child 53163 7c2b13a53d69
permissions -rw-r--r--
immutable theory values with full stamp record of every update (increase of stamp size for HOL: 20000 -> 100000, JinjaThreads: 65000 -> 300000) -- minimal measurable impact on inference kernel performance;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/ML/ml_antiquote.ML
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
     3
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
     4
Common ML antiquotations.
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
     5
*)
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
     6
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
     7
signature ML_ANTIQUOTE =
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
     8
sig
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
     9
  val variant: string -> Proof.context -> string * Proof.context
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
    10
  val macro: binding -> Proof.context context_parser -> theory -> theory
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
    11
  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
    12
  val declaration: string -> 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
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    16
structure ML_Antiquote: ML_ANTIQUOTE =
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    17
struct
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    18
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    19
(** generic tools **)
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    20
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    21
(* ML names *)
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    22
48776
37cd53e69840 faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents: 48646
diff changeset
    23
val init_context = ML_Syntax.reserved |> Name.declare "ML_context";
37cd53e69840 faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents: 48646
diff changeset
    24
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 36950
diff changeset
    25
structure Names = Proof_Data
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    26
(
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    27
  type T = Name.context;
48776
37cd53e69840 faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents: 48646
diff changeset
    28
  fun init _ = init_context;
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    29
);
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    30
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    31
fun variant a ctxt =
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    32
  let
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 36950
diff changeset
    33
    val names = Names.get ctxt;
43326
47cf4bc789aa simplified Name.variant -- discontinued builtin fold_map;
wenzelm
parents: 42468
diff changeset
    34
    val (b, names') = Name.variant a names;
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 36950
diff changeset
    35
    val ctxt' = Names.put names' ctxt;
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    36
  in (b, ctxt') end;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    37
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    38
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    39
(* specific antiquotations *)
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    40
27379
c706b7201826 added macro interface;
wenzelm
parents: 27370
diff changeset
    41
fun macro name scan = ML_Context.add_antiq name
27868
a28b3cd0077b ML_Context.add_antiq: pass position;
wenzelm
parents: 27809
diff changeset
    42
  (fn _ => scan :|-- (fn ctxt => Scan.depend (fn _ => Scan.succeed
35019
1ec0a3ff229e simplified interface for ML antiquotations, struct_name is always "Isabelle";
wenzelm
parents: 33519
diff changeset
    43
    (Context.Proof ctxt, fn background => (K ("", ""), background)))));
27379
c706b7201826 added macro interface;
wenzelm
parents: 27370
diff changeset
    44
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    45
fun inline name scan = ML_Context.add_antiq name
35019
1ec0a3ff229e simplified interface for ML antiquotations, struct_name is always "Isabelle";
wenzelm
parents: 33519
diff changeset
    46
  (fn _ => scan >> (fn s => fn background => (K ("", s), background)));
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    47
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    48
fun declaration kind name scan = ML_Context.add_antiq name
35019
1ec0a3ff229e simplified interface for ML antiquotations, struct_name is always "Isabelle";
wenzelm
parents: 33519
diff changeset
    49
  (fn _ => scan >> (fn s => fn background =>
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    50
    let
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
    51
      val (a, background') =
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
    52
        variant (translate_string (fn "." => "_" | c => c) (Binding.name_of name)) background;
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    53
      val env = kind ^ " " ^ a ^ " = " ^ s ^ ";\n";
41486
82c1e348bc18 reverted 08240feb69c7 -- breaks positions of reports;
wenzelm
parents: 41376
diff changeset
    54
      val body = "Isabelle." ^ a;
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    55
    in (K (env, body), background') end));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    56
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    57
val value = declaration "val";
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    58
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    59
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    60
30231
b3f3ad327d4d added @{binding} ML antiquotations;
wenzelm
parents: 29606
diff changeset
    61
(** misc antiquotations **)
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    62
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
    63
val _ = Context.>> (Context.map_theory
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
    64
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
    65
 (inline (Binding.name "assert")
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
    66
    (Scan.succeed "(fn b => if b then () else raise General.Fail \"Assertion failed\")") #>
40110
93e7935d4cb5 added ML antiquotation @{assert};
wenzelm
parents: 37868
diff changeset
    67
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
    68
  inline (Binding.name "make_string") (Scan.succeed ml_make_string) #>
36162
0bd034a80a9a added ML antiquotation @{make_string}, which produces proper pretty printed version in Poly/ML 5.3.0 or later;
wenzelm
parents: 35670
diff changeset
    69
51979
4f3a5f4c1169 more antiquotations;
wenzelm
parents: 51946
diff changeset
    70
  value (Binding.name "option") (Scan.lift (Parse.position Args.name) >> (fn (name, pos) =>
4f3a5f4c1169 more antiquotations;
wenzelm
parents: 51946
diff changeset
    71
    (Options.default_typ name handle ERROR msg => error (msg ^ Position.here pos);
4f3a5f4c1169 more antiquotations;
wenzelm
parents: 51946
diff changeset
    72
     ML_Syntax.print_string name))) #>
51942
527e39becafc more systematic access to options default;
wenzelm
parents: 51686
diff changeset
    73
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
    74
  value (Binding.name "binding")
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46948
diff changeset
    75
    (Scan.lift (Parse.position Args.name) >> ML_Syntax.make_binding) #>
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    76
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
    77
  value (Binding.name "theory")
48927
ef462b5558eb theory def/ref position reports, which enable hyperlinks etc.;
wenzelm
parents: 48864
diff changeset
    78
    (Args.theory -- Scan.lift (Parse.position Args.name) >> (fn (thy, (name, pos)) =>
ef462b5558eb theory def/ref position reports, which enable hyperlinks etc.;
wenzelm
parents: 48864
diff changeset
    79
      (Position.report pos (Theory.get_markup (Context.get_theory thy name));
ef462b5558eb theory def/ref position reports, which enable hyperlinks etc.;
wenzelm
parents: 48864
diff changeset
    80
        "Context.get_theory (Proof_Context.theory_of ML_context) " ^ ML_Syntax.print_string name))
48781
21af4b8103fa more static antiquotations;
wenzelm
parents: 48776
diff changeset
    81
    || Scan.succeed "Proof_Context.theory_of ML_context") #>
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    82
51686
532e0ac5a66d added ML antiquotation @{theory_context};
wenzelm
parents: 48992
diff changeset
    83
  value (Binding.name "theory_context")
532e0ac5a66d added ML antiquotation @{theory_context};
wenzelm
parents: 48992
diff changeset
    84
    (Args.theory -- Scan.lift (Parse.position Args.name) >> (fn (thy, (name, pos)) =>
532e0ac5a66d added ML antiquotation @{theory_context};
wenzelm
parents: 48992
diff changeset
    85
      (Position.report pos (Theory.get_markup (Context.get_theory thy name));
532e0ac5a66d added ML antiquotation @{theory_context};
wenzelm
parents: 48992
diff changeset
    86
        "Proof_Context.get_global (Proof_Context.theory_of ML_context) " ^
532e0ac5a66d added ML antiquotation @{theory_context};
wenzelm
parents: 48992
diff changeset
    87
        ML_Syntax.print_string name))) #>
532e0ac5a66d added ML antiquotation @{theory_context};
wenzelm
parents: 48992
diff changeset
    88
48781
21af4b8103fa more static antiquotations;
wenzelm
parents: 48776
diff changeset
    89
  inline (Binding.name "context") (Scan.succeed "Isabelle.ML_context") #>
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    90
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
    91
  inline (Binding.name "typ") (Args.typ >> (ML_Syntax.atomic o ML_Syntax.print_typ)) #>
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
    92
  inline (Binding.name "term") (Args.term >> (ML_Syntax.atomic o ML_Syntax.print_term)) #>
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
    93
  inline (Binding.name "prop") (Args.prop >> (ML_Syntax.atomic o ML_Syntax.print_term)) #>
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    94
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
    95
  macro (Binding.name "let")
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
    96
    (Args.context --
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
    97
      Scan.lift
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
    98
        (Parse.and_list1 (Parse.and_list1 Args.name_source -- (Args.$$$ "=" |-- Args.name_source)))
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
    99
        >> (fn (ctxt, args) => #2 (Proof_Context.match_bind true args ctxt))) #>
27379
c706b7201826 added macro interface;
wenzelm
parents: 27370
diff changeset
   100
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   101
  macro (Binding.name "note")
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   102
    (Args.context :|-- (fn ctxt =>
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   103
      Parse.and_list1' (Scan.lift (Args.opt_thm_name I "=") -- Attrib.thms
47815
43f677b3ae91 clarified signature;
wenzelm
parents: 46987
diff changeset
   104
        >> (fn ((a, srcs), ths) => ((a, map (Attrib.attribute_cmd ctxt) srcs), [(ths, [])])))
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   105
      >> (fn args => #2 (Proof_Context.note_thmss "" args ctxt)))) #>
27379
c706b7201826 added macro interface;
wenzelm
parents: 27370
diff changeset
   106
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   107
  value (Binding.name "ctyp") (Args.typ >> (fn T =>
48781
21af4b8103fa more static antiquotations;
wenzelm
parents: 48776
diff changeset
   108
    "Thm.ctyp_of (Proof_Context.theory_of ML_context) " ^
21af4b8103fa more static antiquotations;
wenzelm
parents: 48776
diff changeset
   109
      ML_Syntax.atomic (ML_Syntax.print_typ T))) #>
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   110
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   111
  value (Binding.name "cterm") (Args.term >> (fn t =>
48781
21af4b8103fa more static antiquotations;
wenzelm
parents: 48776
diff changeset
   112
    "Thm.cterm_of (Proof_Context.theory_of ML_context) " ^
21af4b8103fa more static antiquotations;
wenzelm
parents: 48776
diff changeset
   113
     ML_Syntax.atomic (ML_Syntax.print_term t))) #>
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   114
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   115
  value (Binding.name "cprop") (Args.prop >> (fn t =>
48781
21af4b8103fa more static antiquotations;
wenzelm
parents: 48776
diff changeset
   116
    "Thm.cterm_of (Proof_Context.theory_of ML_context) " ^
21af4b8103fa more static antiquotations;
wenzelm
parents: 48776
diff changeset
   117
     ML_Syntax.atomic (ML_Syntax.print_term t))) #>
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   118
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   119
  value (Binding.name "cpat")
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   120
    (Args.context --
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   121
      Scan.lift Args.name_source >> uncurry Proof_Context.read_term_pattern >> (fn t =>
48781
21af4b8103fa more static antiquotations;
wenzelm
parents: 48776
diff changeset
   122
        "Thm.cterm_of (Proof_Context.theory_of ML_context) " ^
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   123
          ML_Syntax.atomic (ML_Syntax.print_term t)))));
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   124
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   125
35396
041bb8d18916 ML antiquotations for type classes;
wenzelm
parents: 35361
diff changeset
   126
(* type classes *)
041bb8d18916 ML antiquotations for type classes;
wenzelm
parents: 35361
diff changeset
   127
35670
3007b46c1660 ProofContext.read_class/read_type_name_proper;
wenzelm
parents: 35429
diff changeset
   128
fun class syn = Args.context -- Scan.lift Args.name_source >> (fn (ctxt, s) =>
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42298
diff changeset
   129
  Proof_Context.read_class ctxt s
42290
b1f544c84040 discontinued special treatment of structure Lexicon;
wenzelm
parents: 41486
diff changeset
   130
  |> syn ? Lexicon.mark_class
35396
041bb8d18916 ML antiquotations for type classes;
wenzelm
parents: 35361
diff changeset
   131
  |> ML_Syntax.print_string);
041bb8d18916 ML antiquotations for type classes;
wenzelm
parents: 35361
diff changeset
   132
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   133
val _ = Context.>> (Context.map_theory
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   134
 (inline (Binding.name "class") (class false) #>
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   135
  inline (Binding.name "class_syntax") (class true) #>
35396
041bb8d18916 ML antiquotations for type classes;
wenzelm
parents: 35361
diff changeset
   136
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   137
  inline (Binding.name "sort")
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   138
    (Args.context -- Scan.lift Args.name_source >> (fn (ctxt, s) =>
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   139
      ML_Syntax.atomic (ML_Syntax.print_sort (Syntax.read_sort ctxt s))))));
35396
041bb8d18916 ML antiquotations for type classes;
wenzelm
parents: 35361
diff changeset
   140
041bb8d18916 ML antiquotations for type classes;
wenzelm
parents: 35361
diff changeset
   141
35361
4c7c849b70aa more orthogonal antiquotations for type constructors;
wenzelm
parents: 35262
diff changeset
   142
(* type constructors *)
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   143
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36162
diff changeset
   144
fun type_name kind check = Args.context -- Scan.lift (Parse.position Args.name_source)
35401
bfcbab8592ba clarified @{const_name} (only logical consts) vs. @{const_abbrev};
wenzelm
parents: 35396
diff changeset
   145
  >> (fn (ctxt, (s, pos)) =>
bfcbab8592ba clarified @{const_name} (only logical consts) vs. @{const_abbrev};
wenzelm
parents: 35396
diff changeset
   146
    let
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42298
diff changeset
   147
      val Type (c, _) = Proof_Context.read_type_name_proper ctxt false s;
42468
aea61c5f88c3 clarified Type.the_decl;
wenzelm
parents: 42360
diff changeset
   148
      val decl = Type.the_decl (Proof_Context.tsig_of ctxt) (c, pos);
35401
bfcbab8592ba clarified @{const_name} (only logical consts) vs. @{const_abbrev};
wenzelm
parents: 35396
diff changeset
   149
      val res =
bfcbab8592ba clarified @{const_name} (only logical consts) vs. @{const_abbrev};
wenzelm
parents: 35396
diff changeset
   150
        (case try check (c, decl) of
bfcbab8592ba clarified @{const_name} (only logical consts) vs. @{const_abbrev};
wenzelm
parents: 35396
diff changeset
   151
          SOME res => res
48992
0518bf89c777 renamed Position.str_of to Position.here;
wenzelm
parents: 48927
diff changeset
   152
        | NONE => error ("Not a " ^ kind ^ ": " ^ quote c ^ Position.here pos));
35401
bfcbab8592ba clarified @{const_name} (only logical consts) vs. @{const_abbrev};
wenzelm
parents: 35396
diff changeset
   153
    in ML_Syntax.print_string res end);
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   154
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   155
val _ = Context.>> (Context.map_theory
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   156
 (inline (Binding.name "type_name")
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   157
    (type_name "logical type" (fn (c, Type.LogicalType _) => c)) #>
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   158
  inline (Binding.name "type_abbrev")
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   159
    (type_name "type abbreviation" (fn (c, Type.Abbreviation _) => c)) #>
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   160
  inline (Binding.name "nonterminal")
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   161
    (type_name "nonterminal" (fn (c, Type.Nonterminal) => c)) #>
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   162
  inline (Binding.name "type_syntax")
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   163
    (type_name "type" (fn (c, _) => Lexicon.mark_type c))));
35361
4c7c849b70aa more orthogonal antiquotations for type constructors;
wenzelm
parents: 35262
diff changeset
   164
4c7c849b70aa more orthogonal antiquotations for type constructors;
wenzelm
parents: 35262
diff changeset
   165
4c7c849b70aa more orthogonal antiquotations for type constructors;
wenzelm
parents: 35262
diff changeset
   166
(* constants *)
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   167
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36162
diff changeset
   168
fun const_name check = Args.context -- Scan.lift (Parse.position Args.name_source)
35401
bfcbab8592ba clarified @{const_name} (only logical consts) vs. @{const_abbrev};
wenzelm
parents: 35396
diff changeset
   169
  >> (fn (ctxt, (s, pos)) =>
bfcbab8592ba clarified @{const_name} (only logical consts) vs. @{const_abbrev};
wenzelm
parents: 35396
diff changeset
   170
    let
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42298
diff changeset
   171
      val Const (c, _) = Proof_Context.read_const_proper ctxt false s;
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42298
diff changeset
   172
      val res = check (Proof_Context.consts_of ctxt, c)
48992
0518bf89c777 renamed Position.str_of to Position.here;
wenzelm
parents: 48927
diff changeset
   173
        handle TYPE (msg, _, _) => error (msg ^ Position.here pos);
35401
bfcbab8592ba clarified @{const_name} (only logical consts) vs. @{const_abbrev};
wenzelm
parents: 35396
diff changeset
   174
    in ML_Syntax.print_string res end);
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   175
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   176
val _ = Context.>> (Context.map_theory
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   177
 (inline (Binding.name "const_name")
43794
49cbbe2768a8 sub-structural sharing after Syntax.check phase, with global interning of logical entities (the latter is relevant when bypassing default parsing via YXML);
wenzelm
parents: 43560
diff changeset
   178
    (const_name (fn (consts, c) => (Consts.the_const consts c; c))) #>
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   179
  inline (Binding.name "const_abbrev")
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   180
    (const_name (fn (consts, c) => (Consts.the_abbreviation consts c; c))) #>
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   181
  inline (Binding.name "const_syntax")
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   182
    (const_name (fn (_, c) => Lexicon.mark_const c)) #>
35401
bfcbab8592ba clarified @{const_name} (only logical consts) vs. @{const_abbrev};
wenzelm
parents: 35396
diff changeset
   183
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   184
  inline (Binding.name "syntax_const")
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   185
    (Args.context -- Scan.lift (Parse.position Args.name) >> (fn (ctxt, (c, pos)) =>
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   186
      if is_some (Syntax.lookup_const (Proof_Context.syn_of ctxt) c)
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   187
      then ML_Syntax.print_string c
48992
0518bf89c777 renamed Position.str_of to Position.here;
wenzelm
parents: 48927
diff changeset
   188
      else error ("Unknown syntax const: " ^ quote c ^ Position.here pos))) #>
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   189
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   190
  inline (Binding.name "const")
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   191
    (Args.context -- Scan.lift Args.name_source -- Scan.optional
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   192
        (Scan.lift (Args.$$$ "(") |-- Parse.enum1' "," Args.typ --| Scan.lift (Args.$$$ ")")) []
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   193
      >> (fn ((ctxt, raw_c), Ts) =>
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   194
        let
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   195
          val Const (c, _) = Proof_Context.read_const_proper ctxt true raw_c;
46987
15ce93dfe6da tuned message;
wenzelm
parents: 46961
diff changeset
   196
          val consts = Proof_Context.consts_of ctxt;
15ce93dfe6da tuned message;
wenzelm
parents: 46961
diff changeset
   197
          val n = length (Consts.typargs consts (c, Consts.type_scheme consts c));
15ce93dfe6da tuned message;
wenzelm
parents: 46961
diff changeset
   198
          val _ = length Ts <> n andalso
15ce93dfe6da tuned message;
wenzelm
parents: 46961
diff changeset
   199
            error ("Constant requires " ^ string_of_int n ^ " type argument(s): " ^
15ce93dfe6da tuned message;
wenzelm
parents: 46961
diff changeset
   200
              quote c ^ enclose "(" ")" (commas (replicate n "_")));
15ce93dfe6da tuned message;
wenzelm
parents: 46961
diff changeset
   201
          val const = Const (c, Consts.instance consts (c, Ts));
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43326
diff changeset
   202
        in ML_Syntax.atomic (ML_Syntax.print_term const) end))));
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   203
46948
aae5566d6756 added ML antiquotation @{keyword};
wenzelm
parents: 43794
diff changeset
   204
aae5566d6756 added ML antiquotation @{keyword};
wenzelm
parents: 43794
diff changeset
   205
(* outer syntax *)
aae5566d6756 added ML antiquotation @{keyword};
wenzelm
parents: 43794
diff changeset
   206
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46948
diff changeset
   207
fun with_keyword f =
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46948
diff changeset
   208
  Args.theory -- Scan.lift (Parse.position Parse.string) >> (fn (thy, (name, pos)) =>
48646
91281e9472d8 more official command specifications, including source position;
wenzelm
parents: 47815
diff changeset
   209
    (f ((name, Thy_Header.the_keyword thy name), pos)
48992
0518bf89c777 renamed Position.str_of to Position.here;
wenzelm
parents: 48927
diff changeset
   210
      handle ERROR msg => error (msg ^ Position.here pos)));
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46948
diff changeset
   211
46948
aae5566d6756 added ML antiquotation @{keyword};
wenzelm
parents: 43794
diff changeset
   212
val _ = Context.>> (Context.map_theory
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46948
diff changeset
   213
 (value (Binding.name "keyword")
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46948
diff changeset
   214
    (with_keyword
48646
91281e9472d8 more official command specifications, including source position;
wenzelm
parents: 47815
diff changeset
   215
      (fn ((name, NONE), _) => "Parse.$$$ " ^ ML_Syntax.print_string name
91281e9472d8 more official command specifications, including source position;
wenzelm
parents: 47815
diff changeset
   216
        | ((name, SOME _), pos) =>
48992
0518bf89c777 renamed Position.str_of to Position.here;
wenzelm
parents: 48927
diff changeset
   217
            error ("Expected minor keyword " ^ quote name ^ Position.here pos))) #>
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46948
diff changeset
   218
  value (Binding.name "command_spec")
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46948
diff changeset
   219
    (with_keyword
48646
91281e9472d8 more official command specifications, including source position;
wenzelm
parents: 47815
diff changeset
   220
      (fn ((name, SOME kind), pos) =>
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46948
diff changeset
   221
          "Keyword.command_spec " ^ ML_Syntax.atomic
48646
91281e9472d8 more official command specifications, including source position;
wenzelm
parents: 47815
diff changeset
   222
            ((ML_Syntax.print_pair
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46948
diff changeset
   223
              (ML_Syntax.print_pair ML_Syntax.print_string
48864
3ee314ae1e0a added keyword kind "thy_load" (with optional list of file extensions);
wenzelm
parents: 48781
diff changeset
   224
                (ML_Syntax.print_pair
3ee314ae1e0a added keyword kind "thy_load" (with optional list of file extensions);
wenzelm
parents: 48781
diff changeset
   225
                  (ML_Syntax.print_pair ML_Syntax.print_string
3ee314ae1e0a added keyword kind "thy_load" (with optional list of file extensions);
wenzelm
parents: 48781
diff changeset
   226
                    (ML_Syntax.print_list ML_Syntax.print_string))
48646
91281e9472d8 more official command specifications, including source position;
wenzelm
parents: 47815
diff changeset
   227
                  (ML_Syntax.print_list ML_Syntax.print_string)))
91281e9472d8 more official command specifications, including source position;
wenzelm
parents: 47815
diff changeset
   228
              ML_Syntax.print_position) ((name, kind), pos))
91281e9472d8 more official command specifications, including source position;
wenzelm
parents: 47815
diff changeset
   229
        | ((name, NONE), pos) =>
48992
0518bf89c777 renamed Position.str_of to Position.here;
wenzelm
parents: 48927
diff changeset
   230
            error ("Expected command keyword " ^ quote name ^ Position.here pos)))));
46948
aae5566d6756 added ML antiquotation @{keyword};
wenzelm
parents: 43794
diff changeset
   231
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   232
end;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   233