src/Pure/ML/ml_antiquote.ML
author wenzelm
Thu, 05 Mar 2009 12:08:00 +0100
changeset 30280 eb98b49ef835
parent 30231 b3f3ad327d4d
child 30364 577edc39b501
permissions -rw-r--r--
renamed NameSpace.base to NameSpace.base_name; renamed NameSpace.map_base to NameSpace.map_base_name; eliminated alias Sign.base_name = NameSpace.base_name;
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
27379
c706b7201826 added macro interface;
wenzelm
parents: 27370
diff changeset
     9
  val macro: string ->
c706b7201826 added macro interface;
wenzelm
parents: 27370
diff changeset
    10
    (Context.generic * Args.T list -> Proof.context * (Context.generic * Args.T list)) -> unit
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    11
  val variant: string -> Proof.context -> string * Proof.context
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    12
  val inline: string ->
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    13
    (Context.generic * Args.T list -> string * (Context.generic * Args.T list)) -> unit
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    14
  val declaration: string -> string ->
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    15
    (Context.generic * Args.T list -> string * (Context.generic * Args.T list)) -> unit
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    16
  val value: string ->
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    17
    (Context.generic * Args.T list -> string * (Context.generic * Args.T list)) -> unit
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    18
end;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    19
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    20
structure ML_Antiquote: ML_ANTIQUOTE =
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    21
struct
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    22
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    23
(** generic tools **)
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    24
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    25
(* ML names *)
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    26
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    27
structure NamesData = ProofDataFun
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    28
(
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    29
  type T = Name.context;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    30
  fun init _ = ML_Syntax.reserved;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    31
);
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    32
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    33
fun variant a ctxt =
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    34
  let
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    35
    val names = NamesData.get ctxt;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    36
    val ([b], names') = Name.variants [a] names;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    37
    val ctxt' = NamesData.put names' ctxt;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    38
  in (b, ctxt') end;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    39
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    40
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    41
(* specific antiquotations *)
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    42
27379
c706b7201826 added macro interface;
wenzelm
parents: 27370
diff changeset
    43
fun macro name scan = ML_Context.add_antiq name
27868
a28b3cd0077b ML_Context.add_antiq: pass position;
wenzelm
parents: 27809
diff changeset
    44
  (fn _ => scan :|-- (fn ctxt => Scan.depend (fn _ => Scan.succeed
27379
c706b7201826 added macro interface;
wenzelm
parents: 27370
diff changeset
    45
    (Context.Proof ctxt, fn {background, ...} => (K ("", ""), background)))));
c706b7201826 added macro interface;
wenzelm
parents: 27370
diff changeset
    46
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    47
fun inline name scan = ML_Context.add_antiq name
27868
a28b3cd0077b ML_Context.add_antiq: pass position;
wenzelm
parents: 27809
diff changeset
    48
  (fn _ => scan >> (fn s => fn {struct_name, background} => (K ("", s), background)));
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    49
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    50
fun declaration kind name scan = ML_Context.add_antiq name
27868
a28b3cd0077b ML_Context.add_antiq: pass position;
wenzelm
parents: 27809
diff changeset
    51
  (fn _ => scan >> (fn s => fn {struct_name, background} =>
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    52
    let
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    53
      val (a, background') = variant name background;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    54
      val env = kind ^ " " ^ a ^ " = " ^ s ^ ";\n";
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    55
      val body = struct_name ^ "." ^ a;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    56
    in (K (env, body), background') end));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    57
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    58
val value = declaration "val";
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    59
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    60
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    61
30231
b3f3ad327d4d added @{binding} ML antiquotations;
wenzelm
parents: 29606
diff changeset
    62
(** misc antiquotations **)
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    63
27809
a1e409db516b unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27390
diff changeset
    64
structure P = OuterParse;
a1e409db516b unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27390
diff changeset
    65
30231
b3f3ad327d4d added @{binding} ML antiquotations;
wenzelm
parents: 29606
diff changeset
    66
val _ = inline "binding" (Scan.lift (P.position Args.name) >> (fn b =>
b3f3ad327d4d added @{binding} ML antiquotations;
wenzelm
parents: 29606
diff changeset
    67
  ML_Syntax.atomic ("Binding.make " ^
b3f3ad327d4d added @{binding} ML antiquotations;
wenzelm
parents: 29606
diff changeset
    68
    ML_Syntax.print_pair ML_Syntax.print_string ML_Syntax.print_position b)));
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    69
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    70
val _ = value "theory"
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    71
  (Scan.lift Args.name >> (fn name => "ThyInfo.get_theory " ^ ML_Syntax.print_string name)
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    72
  || Scan.succeed "ML_Context.the_global_context ()");
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    73
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    74
val _ = value "theory_ref"
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    75
  (Scan.lift Args.name >> (fn name =>
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    76
    "Theory.check_thy (ThyInfo.theory " ^ ML_Syntax.print_string name ^ ")")
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    77
  || Scan.succeed "Theory.check_thy (ML_Context.the_global_context ())");
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    78
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    79
val _ = value "context" (Scan.succeed "ML_Context.the_local_context ()");
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    80
27882
eaa9fef9f4c1 Args.name_source(_position) for proper position information;
wenzelm
parents: 27868
diff changeset
    81
val _ = inline "sort" (Args.context -- Scan.lift Args.name_source >> (fn (ctxt, s) =>
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    82
  ML_Syntax.atomic (ML_Syntax.print_sort (Syntax.read_sort ctxt s))));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    83
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    84
val _ = inline "typ" (Args.typ >> (ML_Syntax.atomic o ML_Syntax.print_typ));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    85
val _ = inline "term" (Args.term >> (ML_Syntax.atomic o ML_Syntax.print_term));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    86
val _ = inline "prop" (Args.prop >> (ML_Syntax.atomic o ML_Syntax.print_term));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    87
27379
c706b7201826 added macro interface;
wenzelm
parents: 27370
diff changeset
    88
val _ = macro "let" (Args.context --
27882
eaa9fef9f4c1 Args.name_source(_position) for proper position information;
wenzelm
parents: 27868
diff changeset
    89
  Scan.lift (P.and_list1 (P.and_list1 Args.name_source -- (Args.$$$ "=" |-- Args.name_source)))
27379
c706b7201826 added macro interface;
wenzelm
parents: 27370
diff changeset
    90
    >> (fn (ctxt, args) => #2 (ProofContext.match_bind true args ctxt)));
c706b7201826 added macro interface;
wenzelm
parents: 27370
diff changeset
    91
c706b7201826 added macro interface;
wenzelm
parents: 27370
diff changeset
    92
val _ = macro "note" (Args.context :|-- (fn ctxt =>
27809
a1e409db516b unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27390
diff changeset
    93
  P.and_list1' (Scan.lift (Args.opt_thm_name I "=") -- Attrib.thms >> (fn ((a, srcs), ths) =>
27379
c706b7201826 added macro interface;
wenzelm
parents: 27370
diff changeset
    94
    ((a, map (Attrib.attribute (ProofContext.theory_of ctxt)) srcs), [(ths, [])])))
c706b7201826 added macro interface;
wenzelm
parents: 27370
diff changeset
    95
  >> (fn args => #2 (ProofContext.note_thmss_i "" args ctxt))));
c706b7201826 added macro interface;
wenzelm
parents: 27370
diff changeset
    96
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    97
val _ = value "ctyp" (Args.typ >> (fn T =>
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    98
  "Thm.ctyp_of (ML_Context.the_global_context ()) " ^ ML_Syntax.atomic (ML_Syntax.print_typ T)));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    99
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   100
val _ = value "cterm" (Args.term >> (fn t =>
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   101
  "Thm.cterm_of (ML_Context.the_global_context ()) " ^ ML_Syntax.atomic (ML_Syntax.print_term t)));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   102
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   103
val _ = value "cprop" (Args.prop >> (fn t =>
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   104
  "Thm.cterm_of (ML_Context.the_global_context ()) " ^ ML_Syntax.atomic (ML_Syntax.print_term t)));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   105
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   106
val _ = value "cpat"
27882
eaa9fef9f4c1 Args.name_source(_position) for proper position information;
wenzelm
parents: 27868
diff changeset
   107
  (Args.context -- Scan.lift Args.name_source >> uncurry ProofContext.read_term_pattern >> (fn t =>
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   108
    "Thm.cterm_of (ML_Context.the_global_context ()) " ^ ML_Syntax.atomic (ML_Syntax.print_term t)));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   109
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   110
27882
eaa9fef9f4c1 Args.name_source(_position) for proper position information;
wenzelm
parents: 27868
diff changeset
   111
fun type_ syn = (Args.context -- Scan.lift Args.name_source >> (fn (ctxt, c) =>
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   112
    #1 (Term.dest_Type (ProofContext.read_tyname ctxt c))
30280
eb98b49ef835 renamed NameSpace.base to NameSpace.base_name;
wenzelm
parents: 30231
diff changeset
   113
    |> syn ? NameSpace.base_name
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   114
    |> ML_Syntax.print_string));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   115
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   116
val _ = inline "type_name" (type_ false);
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   117
val _ = inline "type_syntax" (type_ true);
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   118
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   119
27882
eaa9fef9f4c1 Args.name_source(_position) for proper position information;
wenzelm
parents: 27868
diff changeset
   120
fun const syn = Args.context -- Scan.lift Args.name_source >> (fn (ctxt, c) =>
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   121
  #1 (Term.dest_Const (ProofContext.read_const_proper ctxt c))
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   122
  |> syn ? ProofContext.const_syntax_name ctxt
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   123
  |> ML_Syntax.print_string);
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   124
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   125
val _ = inline "const_name" (const false);
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   126
val _ = inline "const_syntax" (const true);
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   127
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   128
val _ = inline "const"
27882
eaa9fef9f4c1 Args.name_source(_position) for proper position information;
wenzelm
parents: 27868
diff changeset
   129
  (Args.context -- Scan.lift Args.name_source -- Scan.optional
27809
a1e409db516b unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27390
diff changeset
   130
      (Scan.lift (Args.$$$ "(") |-- OuterParse.enum1' "," Args.typ --| Scan.lift (Args.$$$ ")")) []
27340
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   131
    >> (fn ((ctxt, c), Ts) =>
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   132
      let val (c, _) = Term.dest_Const (ProofContext.read_const_proper ctxt c)
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   133
      in ML_Syntax.atomic (ML_Syntax.print_term (ProofContext.mk_const ctxt (c, Ts))) end));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   134
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   135
end;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   136