src/Pure/ML/ml_context.ML
author wenzelm
Sat, 13 Dec 2008 15:00:39 +0100
changeset 29091 b81fe045e799
parent 28412 0608c04858c7
child 29579 cb520b766e00
permissions -rw-r--r--
Context.display_names;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24581
6491d89ba76c tuned comments;
wenzelm
parents: 24574
diff changeset
     1
(*  Title:      Pure/ML/ml_context.ML
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
     3
    Author:     Makarius
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
     4
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
     5
ML context and antiquotations.
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
     6
*)
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
     7
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
     8
signature BASIC_ML_CONTEXT =
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
     9
sig
26416
a66f86ef7bb9 added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents: 26405
diff changeset
    10
  val bind_thm: string * thm -> unit
a66f86ef7bb9 added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents: 26405
diff changeset
    11
  val bind_thms: string * thm list -> unit
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    12
  val thm: xstring -> thm
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    13
  val thms: xstring -> thm list
25204
36cf92f63a44 replaced Secure.evaluate by ML_Context.evaluate;
wenzelm
parents: 25142
diff changeset
    14
end
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    15
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    16
signature ML_CONTEXT =
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    17
sig
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    18
  include BASIC_ML_CONTEXT
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    19
  val the_generic_context: unit -> Context.generic
26432
095e448b95a0 renamed ML_Context.the_context to ML_Context.the_global_context;
wenzelm
parents: 26416
diff changeset
    20
  val the_global_context: unit -> theory
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    21
  val the_local_context: unit -> Proof.context
26455
1757a6e049f4 reorganized signature of ML_Context;
wenzelm
parents: 26432
diff changeset
    22
  val exec: (unit -> unit) -> Context.generic -> Context.generic
28279
7d56de7e2305 added inherit_env;
wenzelm
parents: 28270
diff changeset
    23
  val inherit_env: Proof.context -> Proof.context -> Proof.context
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    24
  val name_space: ML_NameSpace.nameSpace
26416
a66f86ef7bb9 added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents: 26405
diff changeset
    25
  val stored_thms: thm list ref
a66f86ef7bb9 added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents: 26405
diff changeset
    26
  val ml_store_thm: string * thm -> unit
a66f86ef7bb9 added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents: 26405
diff changeset
    27
  val ml_store_thms: string * thm list -> unit
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
    28
  type antiq =
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
    29
    {struct_name: string, background: Proof.context} ->
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
    30
      (Proof.context -> string * string) * Proof.context
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
    31
  val add_antiq: string ->
27868
a28b3cd0077b ML_Context.add_antiq: pass position;
wenzelm
parents: 27809
diff changeset
    32
    (Position.T -> Context.generic * Args.T list -> antiq * (Context.generic * Args.T list)) -> unit
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    33
  val trace: bool ref
27723
ce8f79b91ed1 Exported eval_wrapper.
berghofe
parents: 27378
diff changeset
    34
  val eval_wrapper: (string -> unit) * (string -> 'a) -> bool -> Position.T -> string -> unit
26455
1757a6e049f4 reorganized signature of ML_Context;
wenzelm
parents: 26432
diff changeset
    35
  val eval: bool -> Position.T -> string -> unit
1757a6e049f4 reorganized signature of ML_Context;
wenzelm
parents: 26432
diff changeset
    36
  val eval_file: Path.T -> unit
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    37
  val eval_in: Proof.context option -> bool -> Position.T -> string -> unit
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    38
  val evaluate: Proof.context -> (string -> unit) * (string -> 'b) -> bool ->
25204
36cf92f63a44 replaced Secure.evaluate by ML_Context.evaluate;
wenzelm
parents: 25142
diff changeset
    39
    string * (unit -> 'a) option ref -> string -> 'a
26455
1757a6e049f4 reorganized signature of ML_Context;
wenzelm
parents: 26432
diff changeset
    40
  val expression: Position.T -> string -> string -> string -> Context.generic -> Context.generic
25204
36cf92f63a44 replaced Secure.evaluate by ML_Context.evaluate;
wenzelm
parents: 25142
diff changeset
    41
end
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    42
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    43
structure ML_Context: ML_CONTEXT =
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    44
struct
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    45
26416
a66f86ef7bb9 added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents: 26405
diff changeset
    46
(** implicit ML context **)
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    47
26416
a66f86ef7bb9 added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents: 26405
diff changeset
    48
val the_generic_context = Context.the_thread_data;
26432
095e448b95a0 renamed ML_Context.the_context to ML_Context.the_global_context;
wenzelm
parents: 26416
diff changeset
    49
val the_global_context = Context.theory_of o the_generic_context;
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    50
val the_local_context = Context.proof_of o the_generic_context;
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    51
26455
1757a6e049f4 reorganized signature of ML_Context;
wenzelm
parents: 26432
diff changeset
    52
fun exec (e: unit -> unit) context =
1757a6e049f4 reorganized signature of ML_Context;
wenzelm
parents: 26432
diff changeset
    53
  (case Context.setmp_thread_data (SOME context) (fn () => (e (); Context.thread_data ())) () of
1757a6e049f4 reorganized signature of ML_Context;
wenzelm
parents: 26432
diff changeset
    54
    SOME context' => context'
1757a6e049f4 reorganized signature of ML_Context;
wenzelm
parents: 26432
diff changeset
    55
  | NONE => error "Missing context after execution");
26416
a66f86ef7bb9 added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents: 26405
diff changeset
    56
a66f86ef7bb9 added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents: 26405
diff changeset
    57
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    58
(* ML name space *)
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    59
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    60
structure ML_Env = GenericDataFun
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    61
(
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    62
  type T =
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    63
    ML_NameSpace.valueVal Symtab.table *
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    64
    ML_NameSpace.typeVal Symtab.table *
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    65
    ML_NameSpace.fixityVal Symtab.table *
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    66
    ML_NameSpace.structureVal Symtab.table *
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    67
    ML_NameSpace.signatureVal Symtab.table *
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    68
    ML_NameSpace.functorVal Symtab.table;
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    69
  val empty = (Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty);
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    70
  val extend = I;
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    71
  fun merge _
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    72
   ((val1, type1, fixity1, structure1, signature1, functor1),
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    73
    (val2, type2, fixity2, structure2, signature2, functor2)) : T =
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    74
   (Symtab.merge (K true) (val1, val2),
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    75
    Symtab.merge (K true) (type1, type2),
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    76
    Symtab.merge (K true) (fixity1, fixity2),
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    77
    Symtab.merge (K true) (structure1, structure2),
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    78
    Symtab.merge (K true) (signature1, signature2),
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    79
    Symtab.merge (K true) (functor1, functor2));
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    80
);
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    81
28279
7d56de7e2305 added inherit_env;
wenzelm
parents: 28270
diff changeset
    82
val inherit_env = Context.proof_map o ML_Env.put o ML_Env.get o Context.Proof;
7d56de7e2305 added inherit_env;
wenzelm
parents: 28270
diff changeset
    83
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    84
val name_space: ML_NameSpace.nameSpace =
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    85
  let
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    86
    fun lookup sel1 sel2 name =
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    87
      Context.thread_data ()
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    88
      |> (fn NONE => NONE | SOME context => Symtab.lookup (sel1 (ML_Env.get context)) name)
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    89
      |> (fn NONE => sel2 ML_NameSpace.global name | some => some);
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    90
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    91
    fun all sel1 sel2 () =
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    92
      Context.thread_data ()
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    93
      |> (fn NONE => [] | SOME context => Symtab.dest (sel1 (ML_Env.get context)))
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    94
      |> append (sel2 ML_NameSpace.global ())
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    95
      |> sort_distinct (string_ord o pairself #1);
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    96
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    97
    fun enter ap1 sel2 entry =
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    98
      if is_some (Context.thread_data ()) then
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
    99
        Context.>> (ML_Env.map (ap1 (Symtab.update entry)))
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   100
      else sel2 ML_NameSpace.global entry;
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   101
  in
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   102
   {lookupVal    = lookup #1 #lookupVal,
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   103
    lookupType   = lookup #2 #lookupType,
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   104
    lookupFix    = lookup #3 #lookupFix,
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   105
    lookupStruct = lookup #4 #lookupStruct,
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   106
    lookupSig    = lookup #5 #lookupSig,
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   107
    lookupFunct  = lookup #6 #lookupFunct,
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   108
    enterVal     = enter (fn h => fn (a, b, c, d, e, f) => (h a, b, c, d, e, f)) #enterVal,
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   109
    enterType    = enter (fn h => fn (a, b, c, d, e, f) => (a, h b, c, d, e, f)) #enterType,
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   110
    enterFix     = enter (fn h => fn (a, b, c, d, e, f) => (a, b, h c, d, e, f)) #enterFix,
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   111
    enterStruct  = enter (fn h => fn (a, b, c, d, e, f) => (a, b, c, h d, e, f)) #enterStruct,
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   112
    enterSig     = enter (fn h => fn (a, b, c, d, e, f) => (a, b, c, d, h e, f)) #enterSig,
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   113
    enterFunct   = enter (fn h => fn (a, b, c, d, e, f) => (a, b, c, d, e, h f)) #enterFunct,
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   114
    allVal       = all #1 #allVal,
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   115
    allType      = all #2 #allType,
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   116
    allFix       = all #3 #allFix,
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   117
    allStruct    = all #4 #allStruct,
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   118
    allSig       = all #5 #allSig,
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   119
    allFunct     = all #6 #allFunct}
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   120
  end;
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   121
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   122
26416
a66f86ef7bb9 added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents: 26405
diff changeset
   123
(* theorem bindings *)
a66f86ef7bb9 added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents: 26405
diff changeset
   124
a66f86ef7bb9 added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents: 26405
diff changeset
   125
val stored_thms: thm list ref = ref [];
a66f86ef7bb9 added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents: 26405
diff changeset
   126
26492
6e87cc839632 removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents: 26473
diff changeset
   127
fun ml_store sel (name, ths) =
6e87cc839632 removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents: 26473
diff changeset
   128
  let
6e87cc839632 removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents: 26473
diff changeset
   129
    val ths' = Context.>>> (Context.map_theory_result (PureThy.store_thms (name, ths)));
6e87cc839632 removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents: 26473
diff changeset
   130
    val _ =
6e87cc839632 removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents: 26473
diff changeset
   131
      if name = "" then ()
6e87cc839632 removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents: 26473
diff changeset
   132
      else if not (ML_Syntax.is_identifier name) then
6e87cc839632 removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents: 26473
diff changeset
   133
        error ("Cannot bind theorem(s) " ^ quote name ^ " as ML value")
6e87cc839632 removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents: 26473
diff changeset
   134
      else setmp stored_thms ths' (fn () =>
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   135
        use_text name_space (0, "") Output.ml_output true
26492
6e87cc839632 removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents: 26473
diff changeset
   136
          ("val " ^ name ^ " = " ^ sel ^ "(! ML_Context.stored_thms);")) ();
6e87cc839632 removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents: 26473
diff changeset
   137
  in () end;
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   138
26492
6e87cc839632 removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents: 26473
diff changeset
   139
val ml_store_thms = ml_store "";
6e87cc839632 removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents: 26473
diff changeset
   140
fun ml_store_thm (name, th) = ml_store "hd" (name, [th]);
26416
a66f86ef7bb9 added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents: 26405
diff changeset
   141
a66f86ef7bb9 added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents: 26405
diff changeset
   142
fun bind_thm (name, thm) = ml_store_thm (name, Drule.standard thm);
a66f86ef7bb9 added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents: 26405
diff changeset
   143
fun bind_thms (name, thms) = ml_store_thms (name, map Drule.standard thms);
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   144
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   145
fun thm name = ProofContext.get_thm (the_local_context ()) name;
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   146
fun thms name = ProofContext.get_thms (the_local_context ()) name;
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   147
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   148
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   149
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   150
(** ML antiquotations **)
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   151
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   152
(* antiquotation commands *)
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   153
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   154
type antiq =
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   155
  {struct_name: string, background: Proof.context} ->
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   156
    (Proof.context -> string * string) * Proof.context;
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   157
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   158
local
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   159
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   160
val global_parsers = ref (Symtab.empty:
27868
a28b3cd0077b ML_Context.add_antiq: pass position;
wenzelm
parents: 27809
diff changeset
   161
  (Position.T -> Context.generic * Args.T list -> antiq * (Context.generic * Args.T list))
a28b3cd0077b ML_Context.add_antiq: pass position;
wenzelm
parents: 27809
diff changeset
   162
    Symtab.table);
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   163
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   164
in
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   165
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   166
fun add_antiq name scan = CRITICAL (fn () =>
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   167
  change global_parsers (fn tab =>
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   168
   (if not (Symtab.defined tab name) then ()
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   169
    else warning ("Redefined ML antiquotation: " ^ quote name);
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   170
    Symtab.update (name, scan) tab)));
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   171
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   172
fun antiquotation src ctxt =
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   173
  let val ((name, _), pos) = Args.dest_src src in
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   174
    (case Symtab.lookup (! global_parsers) name of
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   175
      NONE => error ("Unknown ML antiquotation command: " ^ quote name ^ Position.str_of pos)
28412
0608c04858c7 back to plain Position.report for regular references;
wenzelm
parents: 28407
diff changeset
   176
    | SOME scan => (Position.report (Markup.ML_antiq name) pos;
27895
e4f8763b971b report antiquotation names;
wenzelm
parents: 27878
diff changeset
   177
        Args.context_syntax "ML antiquotation" (scan pos) src ctxt))
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   178
  end;
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   179
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   180
end;
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   181
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   182
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   183
(* parsing and evaluation *)
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   184
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   185
local
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   186
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   187
structure P = OuterParse;
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   188
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   189
val antiq =
27809
a1e409db516b unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27781
diff changeset
   190
  P.!!! (P.position P.xname -- Args.parse --| Scan.ahead P.eof)
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   191
  >> (fn ((x, pos), y) => Args.src ((x, y), pos));
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   192
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   193
val struct_name = "Isabelle";
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   194
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   195
fun eval_antiquotes (txt, pos) opt_context =
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   196
  let
27874
f0364f1c0ecf antiquotes: proper SymbolPos decoding, adapted Antiquote.read/Antiq;
wenzelm
parents: 27868
diff changeset
   197
    val syms = SymbolPos.explode (txt, pos);
f0364f1c0ecf antiquotes: proper SymbolPos decoding, adapted Antiquote.read/Antiq;
wenzelm
parents: 27868
diff changeset
   198
    val ants = Antiquote.read (syms, pos);
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   199
    val opt_ctxt = Option.map (Context.Proof o Context.proof_of) opt_context;
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   200
    val ((ml_env, ml_body), opt_ctxt') =
27874
f0364f1c0ecf antiquotes: proper SymbolPos decoding, adapted Antiquote.read/Antiq;
wenzelm
parents: 27868
diff changeset
   201
      if not (exists Antiquote.is_antiq ants)
f0364f1c0ecf antiquotes: proper SymbolPos decoding, adapted Antiquote.read/Antiq;
wenzelm
parents: 27868
diff changeset
   202
      then (("", Symbol.escape (SymbolPos.content syms)), opt_ctxt)
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   203
      else
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   204
        let
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   205
          val ctxt =
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   206
            (case opt_ctxt of
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   207
              NONE => error ("Unknown context -- cannot expand ML antiquotations" ^
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   208
                Position.str_of pos)
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   209
            | SOME ctxt => Context.proof_of ctxt);
27378
wenzelm
parents: 27359
diff changeset
   210
27359
54b5367a827a re-use official outer keywords;
wenzelm
parents: 27343
diff changeset
   211
          val lex = #1 (OuterKeyword.get_lexicons ());
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   212
          fun no_decl _ = ("", "");
27378
wenzelm
parents: 27359
diff changeset
   213
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   214
          fun expand (Antiquote.Text s) state = (K ("", Symbol.escape s), state)
27781
5a82ee34e9fc simplified Antiquote signature;
wenzelm
parents: 27771
diff changeset
   215
            | expand (Antiquote.Antiq x) (scope, background) =
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   216
                let
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   217
                  val context = Stack.top scope;
27781
5a82ee34e9fc simplified Antiquote signature;
wenzelm
parents: 27771
diff changeset
   218
                  val (f, context') = antiquotation (Antiquote.read_antiq lex antiq x) context;
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   219
                  val (decl, background') = f {background = background, struct_name = struct_name};
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   220
                in (decl, (Stack.map_top (K context') scope, background')) end
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   221
            | expand (Antiquote.Open _) (scope, background) =
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   222
                (no_decl, (Stack.push scope, background))
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   223
            | expand (Antiquote.Close _) (scope, background) =
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   224
                (no_decl, (Stack.pop scope, background));
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   225
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   226
          val (decls, (_, ctxt')) = fold_map expand ants (Stack.init ctxt, ctxt);
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   227
          val ml = decls |> map (fn decl => decl ctxt') |> split_list |> pairself implode;
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   228
        in (ml, SOME (Context.Proof ctxt')) end;
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   229
  in (("structure " ^ struct_name ^ " =\nstruct\n" ^ ml_env ^ "end;", ml_body), opt_ctxt') end;
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   230
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   231
in
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   232
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   233
val trace = ref false;
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   234
26385
ae7564661e76 ML runtime compilation: pass position, tuned signature;
wenzelm
parents: 26374
diff changeset
   235
fun eval_wrapper pr verbose pos txt =
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   236
  let
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   237
    fun eval_raw p = use_text name_space
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   238
      (the_default 1 (Position.line_of p), the_default "ML" (Position.file_of p)) pr;
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   239
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   240
    (*prepare source text*)
27878
1ba19c9edd18 report ML_source;
wenzelm
parents: 27874
diff changeset
   241
    val _ = Position.report Markup.ML_source pos;
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   242
    val ((env, body), env_ctxt) = eval_antiquotes (txt, pos) (Context.thread_data ());
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   243
    val _ = if ! trace then tracing (cat_lines [env, body]) else ();
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   244
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   245
    (*prepare static ML environment*)
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   246
    val _ = Context.setmp_thread_data env_ctxt
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   247
        (fn () => (eval_raw Position.none false env; Context.thread_data ())) ()
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   248
      |> (fn NONE => () | SOME context' => Context.>> (ML_Env.put (ML_Env.get context')));
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   249
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   250
    (*eval ML*)
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   251
    val _ = eval_raw pos verbose body;
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   252
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   253
    (*reset static ML environment*)
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   254
    val _ = eval_raw Position.none false "structure Isabelle = struct end";
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   255
  in () end;
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   256
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   257
end;
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   258
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   259
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   260
(* ML evaluation *)
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   261
26455
1757a6e049f4 reorganized signature of ML_Context;
wenzelm
parents: 26432
diff changeset
   262
val eval = eval_wrapper Output.ml_output;
26881
bb68f50644a9 renamed Position.path to Path.position;
wenzelm
parents: 26658
diff changeset
   263
fun eval_file path = eval true (Path.position path) (File.read path);
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   264
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   265
fun eval_in ctxt verbose pos txt =
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   266
  Context.setmp_thread_data (Option.map Context.Proof ctxt) (fn () => eval verbose pos txt) ();
25204
36cf92f63a44 replaced Secure.evaluate by ML_Context.evaluate;
wenzelm
parents: 25142
diff changeset
   267
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   268
fun evaluate ctxt pr verbose (ref_name, r) txt = NAMED_CRITICAL "ML" (fn () =>
25204
36cf92f63a44 replaced Secure.evaluate by ML_Context.evaluate;
wenzelm
parents: 25142
diff changeset
   269
  let
36cf92f63a44 replaced Secure.evaluate by ML_Context.evaluate;
wenzelm
parents: 25142
diff changeset
   270
    val _ = r := NONE;
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   271
    val _ = Context.setmp_thread_data (SOME (Context.Proof ctxt)) (fn () =>
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   272
      eval_wrapper pr verbose Position.none
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   273
        ("val _ = (" ^ ref_name ^ " := SOME (fn () => " ^ txt ^ "))")) ();
25204
36cf92f63a44 replaced Secure.evaluate by ML_Context.evaluate;
wenzelm
parents: 25142
diff changeset
   274
  in (case ! r of NONE => error ("Bad evaluation for " ^ ref_name) | SOME e => e) end) ();
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   275
26455
1757a6e049f4 reorganized signature of ML_Context;
wenzelm
parents: 26432
diff changeset
   276
fun expression pos bind body txt =
1757a6e049f4 reorganized signature of ML_Context;
wenzelm
parents: 26432
diff changeset
   277
  exec (fn () => eval false pos
1757a6e049f4 reorganized signature of ML_Context;
wenzelm
parents: 26432
diff changeset
   278
    ("Context.set_thread_data (SOME (let " ^ bind ^ " = " ^ txt ^ " in " ^ body ^
1757a6e049f4 reorganized signature of ML_Context;
wenzelm
parents: 26432
diff changeset
   279
      " end (ML_Context.the_generic_context ())));"));
1757a6e049f4 reorganized signature of ML_Context;
wenzelm
parents: 26432
diff changeset
   280
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   281
end;
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   282
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   283
structure Basic_ML_Context: BASIC_ML_CONTEXT = ML_Context;
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   284
open Basic_ML_Context;