src/Pure/ML/ml_context.ML
author wenzelm
Thu, 20 Mar 2014 19:24:51 +0100
changeset 56229 f61eaab6bec3
parent 56205 ceb8a93460b7
child 56275 600f432ab556
permissions -rw-r--r--
tuned error, according to "use" in General/secure.ML;
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
    Author:     Makarius
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
     3
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
     4
ML context and antiquotations.
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
     5
*)
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
signature ML_CONTEXT =
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
     8
sig
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
     9
  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
    10
  val the_global_context: unit -> theory
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    11
  val the_local_context: unit -> Proof.context
39164
e7e12555e763 ML_Context.thm and ML_Context.thms no longer pervasive;
wenzelm
parents: 39140
diff changeset
    12
  val thm: xstring -> thm
e7e12555e763 ML_Context.thm and ML_Context.thms no longer pervasive;
wenzelm
parents: 39140
diff changeset
    13
  val thms: xstring -> thm list
26455
1757a6e049f4 reorganized signature of ML_Context;
wenzelm
parents: 26432
diff changeset
    14
  val exec: (unit -> unit) -> Context.generic -> Context.generic
56070
1bc0bea908c3 ML_Context.check_antiquotation still required;
wenzelm
parents: 56069
diff changeset
    15
  val check_antiquotation: Proof.context -> xstring * Position.T -> string
56205
ceb8a93460b7 clarified modules;
wenzelm
parents: 56203
diff changeset
    16
  type decl = Proof.context -> string * string
ceb8a93460b7 clarified modules;
wenzelm
parents: 56203
diff changeset
    17
  val add_antiquotation: binding -> (Args.src -> Proof.context -> decl * Proof.context) ->
ceb8a93460b7 clarified modules;
wenzelm
parents: 56203
diff changeset
    18
    theory -> theory
56069
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
    19
  val print_antiquotations: Proof.context -> unit
41375
7a89b4b94817 configuration option "ML_trace";
wenzelm
parents: 39911
diff changeset
    20
  val trace_raw: Config.raw
7a89b4b94817 configuration option "ML_trace";
wenzelm
parents: 39911
diff changeset
    21
  val trace: bool Config.T
30637
3e3c2cd88cf1 export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents: 30614
diff changeset
    22
  val eval_antiquotes: ML_Lex.token Antiquote.antiquote list * Position.T ->
3e3c2cd88cf1 export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents: 30614
diff changeset
    23
    Context.generic option -> (ML_Lex.token list * ML_Lex.token list) * Context.generic option
37198
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 36959
diff changeset
    24
  val eval: bool -> Position.T -> ML_Lex.token Antiquote.antiquote list -> unit
56203
76c72f4d0667 clarified bootstrap process: switch to ML with context and antiquotations earlier;
wenzelm
parents: 56201
diff changeset
    25
  val eval_file: bool -> Path.T -> unit
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 55743
diff changeset
    26
  val eval_source: bool -> Symbol_Pos.source -> unit
37198
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 36959
diff changeset
    27
  val eval_in: Proof.context option -> bool -> Position.T ->
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 36959
diff changeset
    28
    ML_Lex.token Antiquote.antiquote list -> unit
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 55743
diff changeset
    29
  val eval_source_in: Proof.context option -> bool -> Symbol_Pos.source -> unit
37198
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 36959
diff changeset
    30
  val expression: Position.T -> string -> string -> ML_Lex.token Antiquote.antiquote list ->
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 36959
diff changeset
    31
    Context.generic -> Context.generic
25204
36cf92f63a44 replaced Secure.evaluate by ML_Context.evaluate;
wenzelm
parents: 25142
diff changeset
    32
end
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    33
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    34
structure ML_Context: ML_CONTEXT =
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    35
struct
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    36
26416
a66f86ef7bb9 added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents: 26405
diff changeset
    37
(** implicit ML context **)
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    38
26416
a66f86ef7bb9 added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents: 26405
diff changeset
    39
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
    40
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
    41
val the_local_context = Context.proof_of o the_generic_context;
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    42
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 41485
diff changeset
    43
fun thm name = Proof_Context.get_thm (the_local_context ()) name;
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 41485
diff changeset
    44
fun thms name = Proof_Context.get_thms (the_local_context ()) name;
39164
e7e12555e763 ML_Context.thm and ML_Context.thms no longer pervasive;
wenzelm
parents: 39140
diff changeset
    45
26455
1757a6e049f4 reorganized signature of ML_Context;
wenzelm
parents: 26432
diff changeset
    46
fun exec (e: unit -> unit) context =
1757a6e049f4 reorganized signature of ML_Context;
wenzelm
parents: 26432
diff changeset
    47
  (case Context.setmp_thread_data (SOME context) (fn () => (e (); Context.thread_data ())) () of
1757a6e049f4 reorganized signature of ML_Context;
wenzelm
parents: 26432
diff changeset
    48
    SOME context' => context'
1757a6e049f4 reorganized signature of ML_Context;
wenzelm
parents: 26432
diff changeset
    49
  | NONE => error "Missing context after execution");
26416
a66f86ef7bb9 added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents: 26405
diff changeset
    50
a66f86ef7bb9 added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents: 26405
diff changeset
    51
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    52
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    53
(** ML antiquotations **)
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    54
56069
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
    55
(* theory data *)
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
    56
56069
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
    57
type decl = Proof.context -> string * string;  (*final context -> ML env, ML body*)
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
    58
structure Antiquotations = Theory_Data
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 42360
diff changeset
    59
(
56069
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
    60
  type T = (Args.src -> Proof.context -> decl * Proof.context) Name_Space.table;
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 48992
diff changeset
    61
  val empty : T = Name_Space.empty_table Markup.ML_antiquotationN;
43560
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 42360
diff changeset
    62
  val extend = I;
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 42360
diff changeset
    63
  fun merge data : T = Name_Space.merge_tables data;
d1650e3720fd ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 42360
diff changeset
    64
);
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
    65
56069
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
    66
val get_antiquotations = Antiquotations.get o Proof_Context.theory_of;
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
    67
56070
1bc0bea908c3 ML_Context.check_antiquotation still required;
wenzelm
parents: 56069
diff changeset
    68
fun check_antiquotation ctxt =
1bc0bea908c3 ML_Context.check_antiquotation still required;
wenzelm
parents: 56069
diff changeset
    69
  #1 o Name_Space.check (Context.Proof ctxt) (get_antiquotations ctxt);
1bc0bea908c3 ML_Context.check_antiquotation still required;
wenzelm
parents: 56069
diff changeset
    70
56069
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
    71
fun add_antiquotation name f thy = thy
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
    72
  |> Antiquotations.map (Name_Space.define (Context.Theory thy) true (name, f) #> snd);
55743
225a060e7445 proper context for global data;
wenzelm
parents: 55526
diff changeset
    73
56069
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
    74
fun print_antiquotations ctxt =
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
    75
  Pretty.big_list "ML antiquotations:"
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
    76
    (map (Pretty.mark_str o #1) (Name_Space.markup_table ctxt (get_antiquotations ctxt)))
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
    77
  |> Pretty.writeln;
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
    78
56069
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
    79
fun apply_antiquotation src ctxt =
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
    80
  let val (src', f) = Args.check_src ctxt (get_antiquotations ctxt) src
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
    81
  in f src' ctxt end;
43563
aeabb735883a proper checking of @{ML_antiquotation};
wenzelm
parents: 43560
diff changeset
    82
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    83
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
    84
(* parsing and evaluation *)
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    85
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
    86
local
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
    87
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    88
val antiq =
56201
dd2df97b379b tuned signature;
wenzelm
parents: 56199
diff changeset
    89
  Parse.!!! (Parse.position Parse.xname -- Parse.args --| Scan.ahead Parse.eof)
56029
8bedca4bd5a3 clarified Args.src: more abstract type, position refers to name only;
wenzelm
parents: 55828
diff changeset
    90
  >> uncurry Args.src;
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    91
48776
37cd53e69840 faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents: 47005
diff changeset
    92
val begin_env0 = ML_Lex.tokenize "structure Isabelle =\nstruct\n";
52663
6e71d43775e5 retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents: 50201
diff changeset
    93
6e71d43775e5 retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents: 50201
diff changeset
    94
fun begin_env visible =
48776
37cd53e69840 faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents: 47005
diff changeset
    95
  ML_Lex.tokenize
52663
6e71d43775e5 retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents: 50201
diff changeset
    96
    ("structure Isabelle =\nstruct\n\
6e71d43775e5 retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents: 50201
diff changeset
    97
     \val ML_context = Context_Position.set_visible " ^ Bool.toString visible ^
6e71d43775e5 retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents: 50201
diff changeset
    98
     " (ML_Context.the_local_context ());\n");
48776
37cd53e69840 faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents: 47005
diff changeset
    99
30637
3e3c2cd88cf1 export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents: 30614
diff changeset
   100
val end_env = ML_Lex.tokenize "end;";
31334
999fa9e1dbdd structure ML_Compiler;
wenzelm
parents: 31325
diff changeset
   101
val reset_env = ML_Lex.tokenize "structure Isabelle = struct end";
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   102
30637
3e3c2cd88cf1 export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents: 30614
diff changeset
   103
in
3e3c2cd88cf1 export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents: 30614
diff changeset
   104
3e3c2cd88cf1 export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents: 30614
diff changeset
   105
fun eval_antiquotes (ants, pos) opt_context =
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   106
  let
52663
6e71d43775e5 retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents: 50201
diff changeset
   107
    val visible =
6e71d43775e5 retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents: 50201
diff changeset
   108
      (case opt_context of
6e71d43775e5 retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents: 50201
diff changeset
   109
        SOME (Context.Proof ctxt) => Context_Position.is_visible ctxt
6e71d43775e5 retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents: 50201
diff changeset
   110
      | _ => true);
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   111
    val opt_ctxt = Option.map (Context.Proof o Context.proof_of) opt_context;
52663
6e71d43775e5 retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents: 50201
diff changeset
   112
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   113
    val ((ml_env, ml_body), opt_ctxt') =
30637
3e3c2cd88cf1 export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents: 30614
diff changeset
   114
      if forall Antiquote.is_text ants
48776
37cd53e69840 faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents: 47005
diff changeset
   115
      then ((begin_env0, map (fn Antiquote.Text tok => tok) ants), opt_ctxt)
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   116
      else
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   117
        let
53167
4e7ddd76e632 discontinued unused antiquotation blocks;
wenzelm
parents: 53044
diff changeset
   118
          val lex = #1 (Keyword.get_lexicons ());
4e7ddd76e632 discontinued unused antiquotation blocks;
wenzelm
parents: 53044
diff changeset
   119
          fun no_decl _ = ([], []);
4e7ddd76e632 discontinued unused antiquotation blocks;
wenzelm
parents: 53044
diff changeset
   120
4e7ddd76e632 discontinued unused antiquotation blocks;
wenzelm
parents: 53044
diff changeset
   121
          fun expand (Antiquote.Text tok) ctxt = (K ([], [tok]), ctxt)
55526
39708e59f4b0 more markup;
wenzelm
parents: 53171
diff changeset
   122
            | expand (Antiquote.Antiq (ss, {range, ...})) ctxt =
53167
4e7ddd76e632 discontinued unused antiquotation blocks;
wenzelm
parents: 53044
diff changeset
   123
                let
56069
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
   124
                  val (decl, ctxt') =
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
   125
                    apply_antiquotation (Token.read_antiq lex antiq (ss, #1 range)) ctxt;
53167
4e7ddd76e632 discontinued unused antiquotation blocks;
wenzelm
parents: 53044
diff changeset
   126
                  val decl' = decl #> pairself (ML_Lex.tokenize #> map (ML_Lex.set_range range));
4e7ddd76e632 discontinued unused antiquotation blocks;
wenzelm
parents: 53044
diff changeset
   127
                in (decl', ctxt') end;
4e7ddd76e632 discontinued unused antiquotation blocks;
wenzelm
parents: 53044
diff changeset
   128
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   129
          val ctxt =
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   130
            (case opt_ctxt of
48992
0518bf89c777 renamed Position.str_of to Position.here;
wenzelm
parents: 48781
diff changeset
   131
              NONE => error ("No context -- cannot expand ML antiquotations" ^ Position.here pos)
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   132
            | SOME ctxt => Context.proof_of ctxt);
27378
wenzelm
parents: 27359
diff changeset
   133
53167
4e7ddd76e632 discontinued unused antiquotation blocks;
wenzelm
parents: 53044
diff changeset
   134
          val (decls, ctxt') = fold_map expand ants ctxt;
52663
6e71d43775e5 retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents: 50201
diff changeset
   135
          val (ml_env, ml_body) =
6e71d43775e5 retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents: 50201
diff changeset
   136
            decls |> map (fn decl => decl ctxt') |> split_list |> pairself flat;
6e71d43775e5 retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents: 50201
diff changeset
   137
        in ((begin_env visible @ ml_env, ml_body), SOME (Context.Proof ctxt')) end;
48776
37cd53e69840 faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents: 47005
diff changeset
   138
  in ((ml_env @ end_env, ml_body), opt_ctxt') end;
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   139
41375
7a89b4b94817 configuration option "ML_trace";
wenzelm
parents: 39911
diff changeset
   140
val trace_raw = Config.declare "ML_trace" (fn _ => Config.Bool false);
7a89b4b94817 configuration option "ML_trace";
wenzelm
parents: 39911
diff changeset
   141
val trace = Config.bool trace_raw;
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   142
37198
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 36959
diff changeset
   143
fun eval verbose pos ants =
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   144
  let
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   145
    (*prepare source text*)
41485
6c0de045d127 ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents: 41375
diff changeset
   146
    val ((env, body), env_ctxt) = eval_antiquotes (ants, pos) (Context.thread_data ());
31334
999fa9e1dbdd structure ML_Compiler;
wenzelm
parents: 31325
diff changeset
   147
    val _ =
41485
6c0de045d127 ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents: 41375
diff changeset
   148
      (case Option.map Context.proof_of env_ctxt of
6c0de045d127 ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents: 41375
diff changeset
   149
        SOME ctxt =>
6c0de045d127 ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents: 41375
diff changeset
   150
          if Config.get ctxt trace then
6c0de045d127 ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents: 41375
diff changeset
   151
            Context_Position.if_visible ctxt
6c0de045d127 ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents: 41375
diff changeset
   152
              tracing (cat_lines [ML_Lex.flatten env, ML_Lex.flatten body])
41375
7a89b4b94817 configuration option "ML_trace";
wenzelm
parents: 39911
diff changeset
   153
          else ()
7a89b4b94817 configuration option "ML_trace";
wenzelm
parents: 39911
diff changeset
   154
      | NONE => ());
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   155
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   156
    (*prepare static ML environment*)
41485
6c0de045d127 ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents: 41375
diff changeset
   157
    val _ =
6c0de045d127 ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents: 41375
diff changeset
   158
      Context.setmp_thread_data
6c0de045d127 ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents: 41375
diff changeset
   159
        (Option.map (Context.mapping I (Context_Position.set_visible false)) env_ctxt)
31334
999fa9e1dbdd structure ML_Compiler;
wenzelm
parents: 31325
diff changeset
   160
        (fn () => (ML_Compiler.eval false Position.none env; Context.thread_data ())) ()
31325
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents: 30683
diff changeset
   161
      |> (fn NONE => () | SOME context' => Context.>> (ML_Env.inherit context'));
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   162
31334
999fa9e1dbdd structure ML_Compiler;
wenzelm
parents: 31325
diff changeset
   163
    val _ = ML_Compiler.eval verbose pos body;
999fa9e1dbdd structure ML_Compiler;
wenzelm
parents: 31325
diff changeset
   164
    val _ = ML_Compiler.eval false Position.none reset_env;
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   165
  in () end;
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   166
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   167
end;
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   168
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   169
30672
beaadd5af500 more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents: 30643
diff changeset
   170
(* derived versions *)
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   171
56203
76c72f4d0667 clarified bootstrap process: switch to ML with context and antiquotations earlier;
wenzelm
parents: 56201
diff changeset
   172
fun eval_file verbose path =
76c72f4d0667 clarified bootstrap process: switch to ML with context and antiquotations earlier;
wenzelm
parents: 56201
diff changeset
   173
  let val pos = Path.position path
76c72f4d0667 clarified bootstrap process: switch to ML with context and antiquotations earlier;
wenzelm
parents: 56201
diff changeset
   174
  in eval verbose pos (ML_Lex.read pos (File.read path)) end;
76c72f4d0667 clarified bootstrap process: switch to ML with context and antiquotations earlier;
wenzelm
parents: 56201
diff changeset
   175
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 55743
diff changeset
   176
fun eval_source verbose source =
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 55743
diff changeset
   177
  eval verbose (#pos source) (ML_Lex.read_source source);
37198
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 36959
diff changeset
   178
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 36959
diff changeset
   179
fun eval_in ctxt verbose pos ants =
56069
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
   180
  Context.setmp_thread_data (Option.map Context.Proof ctxt)
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
   181
    (fn () => eval verbose pos ants) ();
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   182
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 55743
diff changeset
   183
fun eval_source_in ctxt verbose source =
56069
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
   184
  Context.setmp_thread_data (Option.map Context.Proof ctxt)
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
   185
    (fn () => eval_source verbose source) ();
37198
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 36959
diff changeset
   186
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 36959
diff changeset
   187
fun expression pos bind body ants =
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 36959
diff changeset
   188
  exec (fn () => eval false pos
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 36959
diff changeset
   189
    (ML_Lex.read Position.none ("Context.set_thread_data (SOME (let " ^ bind ^ " = ") @ ants @
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 36959
diff changeset
   190
      ML_Lex.read Position.none (" in " ^ body ^ " end (ML_Context.the_generic_context ())));")));
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 36959
diff changeset
   191
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   192
end;
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   193
56229
f61eaab6bec3 tuned error, according to "use" in General/secure.ML;
wenzelm
parents: 56205
diff changeset
   194
fun use s =
f61eaab6bec3 tuned error, according to "use" in General/secure.ML;
wenzelm
parents: 56205
diff changeset
   195
  ML_Context.eval_file true (Path.explode s)
f61eaab6bec3 tuned error, according to "use" in General/secure.ML;
wenzelm
parents: 56205
diff changeset
   196
    handle ERROR msg => (writeln msg; error "ML error");
f61eaab6bec3 tuned error, according to "use" in General/secure.ML;
wenzelm
parents: 56205
diff changeset
   197