src/Pure/ML/ml_antiquote.ML
author wenzelm
Wed, 25 Jun 2008 17:38:39 +0200
changeset 27359 54b5367a827a
parent 27340 3de9f20f4e28
child 27370 8e8f96dfaf63
permissions -rw-r--r--
re-use official outer keywords;
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
    ID:         $Id$
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
     3
    Author:     Makarius
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
     4
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
     5
Common ML antiquotations.
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
     6
*)
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
     7
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
     8
signature ML_ANTIQUOTE =
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
     9
sig
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    10
  val variant: string -> Proof.context -> string * Proof.context
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    11
  val inline: string ->
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    12
    (Context.generic * Args.T list -> string * (Context.generic * Args.T list)) -> unit
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    13
  val declaration: string -> string ->
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    14
    (Context.generic * Args.T list -> string * (Context.generic * Args.T list)) -> unit
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    15
  val value: string ->
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    16
    (Context.generic * Args.T list -> string * (Context.generic * Args.T list)) -> unit
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    17
  val the_thms: Proof.context -> int -> thm list
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    18
  val the_thm: Proof.context -> int -> thm
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    19
end;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    20
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    21
structure ML_Antiquote: ML_ANTIQUOTE =
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    22
struct
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    23
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    24
(** generic tools **)
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    25
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    26
(* ML names *)
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    27
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    28
structure NamesData = ProofDataFun
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    29
(
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    30
  type T = Name.context;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    31
  fun init _ = ML_Syntax.reserved;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    32
);
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    33
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    34
fun variant a ctxt =
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    35
  let
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    36
    val names = NamesData.get ctxt;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    37
    val ([b], names') = Name.variants [a] names;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    38
    val ctxt' = NamesData.put names' ctxt;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    39
  in (b, ctxt') end;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    40
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    41
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    42
(* specific antiquotations *)
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    43
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    44
fun inline name scan = ML_Context.add_antiq name
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    45
  (scan >> (fn s => fn {struct_name, background} => ((K ("", s)), background)));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    46
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    47
fun declaration kind name scan = ML_Context.add_antiq name
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    48
  (scan >> (fn s => fn {struct_name, background} =>
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    49
    let
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    50
      val (a, background') = variant name background;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    51
      val env = kind ^ " " ^ a ^ " = " ^ s ^ ";\n";
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    52
      val body = struct_name ^ "." ^ a;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    53
    in (K (env, body), background') end));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    54
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    55
val value = declaration "val";
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    56
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    57
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    58
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    59
(** concrete antiquotations **)
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    60
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    61
fun context x = (Scan.state >> Context.proof_of) x;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    62
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    63
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    64
(* misc *)
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    65
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    66
val _ = value "theory"
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    67
  (Scan.lift Args.name >> (fn name => "ThyInfo.get_theory " ^ ML_Syntax.print_string name)
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    68
  || Scan.succeed "ML_Context.the_global_context ()");
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    69
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    70
val _ = value "theory_ref"
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    71
  (Scan.lift Args.name >> (fn name =>
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    72
    "Theory.check_thy (ThyInfo.theory " ^ ML_Syntax.print_string name ^ ")")
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    73
  || Scan.succeed "Theory.check_thy (ML_Context.the_global_context ())");
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    74
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    75
val _ = value "context" (Scan.succeed "ML_Context.the_local_context ()");
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    76
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    77
val _ = inline "sort" (context -- Scan.lift Args.name >> (fn (ctxt, s) =>
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    78
  ML_Syntax.atomic (ML_Syntax.print_sort (Syntax.read_sort ctxt s))));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    79
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    80
val _ = inline "typ" (Args.typ >> (ML_Syntax.atomic o ML_Syntax.print_typ));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    81
val _ = inline "term" (Args.term >> (ML_Syntax.atomic o ML_Syntax.print_term));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    82
val _ = inline "prop" (Args.prop >> (ML_Syntax.atomic o ML_Syntax.print_term));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    83
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    84
val _ = value "ctyp" (Args.typ >> (fn T =>
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    85
  "Thm.ctyp_of (ML_Context.the_global_context ()) " ^ ML_Syntax.atomic (ML_Syntax.print_typ T)));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    86
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    87
val _ = value "cterm" (Args.term >> (fn t =>
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    88
  "Thm.cterm_of (ML_Context.the_global_context ()) " ^ ML_Syntax.atomic (ML_Syntax.print_term t)));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    89
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    90
val _ = value "cprop" (Args.prop >> (fn t =>
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    91
  "Thm.cterm_of (ML_Context.the_global_context ()) " ^ ML_Syntax.atomic (ML_Syntax.print_term t)));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    92
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    93
val _ = value "cpat"
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    94
  (context -- Scan.lift Args.name >> uncurry ProofContext.read_term_pattern >> (fn t =>
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    95
    "Thm.cterm_of (ML_Context.the_global_context ()) " ^ ML_Syntax.atomic (ML_Syntax.print_term t)));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    96
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    97
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    98
fun type_ syn = (context -- Scan.lift Args.name >> (fn (ctxt, c) =>
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
    99
    #1 (Term.dest_Type (ProofContext.read_tyname ctxt c))
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   100
    |> syn ? Sign.base_name
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   101
    |> ML_Syntax.print_string));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   102
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   103
val _ = inline "type_name" (type_ false);
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   104
val _ = inline "type_syntax" (type_ true);
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   105
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   106
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   107
fun const syn = context -- Scan.lift Args.name >> (fn (ctxt, c) =>
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   108
  #1 (Term.dest_Const (ProofContext.read_const_proper ctxt c))
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   109
  |> syn ? ProofContext.const_syntax_name ctxt
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   110
  |> ML_Syntax.print_string);
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   111
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   112
val _ = inline "const_name" (const false);
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   113
val _ = inline "const_syntax" (const true);
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   114
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   115
val _ = inline "const"
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   116
  (context -- Scan.lift Args.name --
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   117
    Scan.optional (Scan.lift (Args.$$$ "(") |-- Args.enum1 "," Args.typ --| Scan.lift (Args.$$$ ")")) []
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   118
    >> (fn ((ctxt, c), Ts) =>
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   119
      let val (c, _) = Term.dest_Const (ProofContext.read_const_proper ctxt c)
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   120
      in ML_Syntax.atomic (ML_Syntax.print_term (ProofContext.mk_const ctxt (c, Ts))) end));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   121
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   122
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   123
(* abstract fact values *)
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   124
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   125
structure AuxFacts = ProofDataFun
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   126
(
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   127
  type T = thm list Inttab.table;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   128
  fun init _ = Inttab.empty;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   129
);
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   130
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   131
fun the_thms ctxt name = the (Inttab.lookup (AuxFacts.get ctxt) name);
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   132
fun the_thm ctxt name = the_single (the_thms ctxt name);
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   133
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   134
fun put_thms ths ctxt =
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   135
  let val i = serial ()
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   136
  in (i, AuxFacts.map (Inttab.update (i, ths)) ctxt) end;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   137
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   138
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   139
fun thm_antiq kind scan = ML_Context.add_antiq kind
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   140
  (scan >> (fn ths => fn {struct_name, background} =>
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   141
    let
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   142
      val ((a, i), background') = background |> variant kind ||>> put_thms ths;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   143
      val env = "val " ^ a ^ " = ML_Antiquote.the_" ^ kind ^
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   144
        " (ML_Context.the_local_context ()) " ^ string_of_int i ^ ";\n";
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   145
      val body = struct_name ^ "." ^ a;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   146
    in (K (env, body), background') end));
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   147
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   148
val _ = thm_antiq "thm" (Attrib.thm >> single);
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   149
val _ = thm_antiq "thms" Attrib.thms;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   150
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   151
end;
3de9f20f4e28 Common ML antiquotations.
wenzelm
parents:
diff changeset
   152