src/Pure/ML/ml_test.ML
author wenzelm
Sat, 23 May 2009 17:39:19 +0200
changeset 31235 67c796138bf0
parent 30744 50ccaef52871
child 31239 dcbf876f5592
permissions -rw-r--r--
adapted to Poly/ML SVN 719; removed obsolete add_prefix;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/ML/ml_test.ML
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
     3
31235
67c796138bf0 adapted to Poly/ML SVN 719;
wenzelm
parents: 30744
diff changeset
     4
Test of advanced ML compiler invocation in Poly/ML 5.3 (SVN 719).
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
     5
*)
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
     6
30646
d26ad4576d5c proper signature;
wenzelm
parents: 30644
diff changeset
     7
signature ML_TEST =
d26ad4576d5c proper signature;
wenzelm
parents: 30644
diff changeset
     8
sig
30744
50ccaef52871 export position_of;
wenzelm
parents: 30709
diff changeset
     9
  val position_of: Proof.context -> PolyML.location -> Position.T
30680
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
    10
  val eval: bool -> Position.T -> ML_Lex.token list -> unit
30646
d26ad4576d5c proper signature;
wenzelm
parents: 30644
diff changeset
    11
end
d26ad4576d5c proper signature;
wenzelm
parents: 30644
diff changeset
    12
d26ad4576d5c proper signature;
wenzelm
parents: 30644
diff changeset
    13
structure ML_Test: ML_TEST =
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
    14
struct
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
    15
30704
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    16
(* extra ML environment *)
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
    17
30704
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    18
structure Extra_Env = GenericDataFun
30646
d26ad4576d5c proper signature;
wenzelm
parents: 30644
diff changeset
    19
(
30704
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    20
  type T = Position.T Inttab.table;  (*position of registered tokens*)
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    21
  val empty = Inttab.empty;
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    22
  val extend = I;
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    23
  fun merge _ = Inttab.merge (K true);
30646
d26ad4576d5c proper signature;
wenzelm
parents: 30644
diff changeset
    24
);
d26ad4576d5c proper signature;
wenzelm
parents: 30644
diff changeset
    25
30704
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    26
fun inherit_env context =
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    27
  ML_Context.inherit_env context #>
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    28
  Extra_Env.put (Extra_Env.get context);
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    29
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    30
fun register_tokens toks context =
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    31
  let
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    32
    val regs = map (fn tok => (serial (), tok)) toks;
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    33
    val context' = context
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    34
      |> Extra_Env.map (fold (fn (i, tok) => Inttab.update (i, ML_Lex.pos_of tok)) regs);
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    35
  in (regs, context') end;
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    36
30744
50ccaef52871 export position_of;
wenzelm
parents: 30709
diff changeset
    37
fun position_of ctxt
30704
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    38
    ({file, startLine = line, startPosition = i, endPosition = j, ...}: location) =
30744
50ccaef52871 export position_of;
wenzelm
parents: 30709
diff changeset
    39
  (case pairself (Inttab.lookup (Extra_Env.get (Context.Proof ctxt))) (i, j) of
30704
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    40
    (SOME pos1, SOME pos2) => Position.encode_range (pos1, pos2)
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    41
  | (SOME pos, NONE) => pos
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    42
  | _ => Position.line_file line file);
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    43
30646
d26ad4576d5c proper signature;
wenzelm
parents: 30644
diff changeset
    44
30704
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    45
(* parse trees *)
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    46
30709
d9ca766bf24c report ML typing;
wenzelm
parents: 30705
diff changeset
    47
fun report_parse_tree context depth =
30704
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    48
  let
30744
50ccaef52871 export position_of;
wenzelm
parents: 30709
diff changeset
    49
    val pos_of = position_of (Context.proof_of context);
30709
d9ca766bf24c report ML typing;
wenzelm
parents: 30705
diff changeset
    50
    fun report loc (PTtype types) =
d9ca766bf24c report ML typing;
wenzelm
parents: 30705
diff changeset
    51
          PolyML.NameSpace.displayTypeExpression (types, depth)
d9ca766bf24c report ML typing;
wenzelm
parents: 30705
diff changeset
    52
          |> pretty_ml |> Pretty.from_ML |> Pretty.string_of
d9ca766bf24c report ML typing;
wenzelm
parents: 30705
diff changeset
    53
          |> Position.report_text Markup.ML_typing (pos_of loc)
30704
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    54
      | report loc (PTdeclaredAt decl) =
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    55
          Markup.markup (Markup.properties (Position.properties_of (pos_of decl)) Markup.ML_def) ""
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    56
          |> Position.report_text Markup.ML_ref (pos_of loc)
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    57
      | report _ (PTnextSibling tree) = report_tree (tree ())
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    58
      | report _ (PTfirstChild tree) = report_tree (tree ())
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    59
      | report _ _ = ()
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    60
    and report_tree (loc, props) = List.app (report loc) props;
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    61
  in report_tree end;
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    62
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    63
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    64
(* eval ML source tokens *)
30680
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
    65
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
    66
fun use_text ({name_space = space, print, error, ...}: use_context) verbose pos toks =
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
    67
  let
30680
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
    68
    (* input *)
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
    69
30704
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    70
    val input = Context.>>> (register_tokens toks);
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    71
    val input_buffer = ref (map (apsnd (String.explode o ML_Lex.text_of)) input);
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    72
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    73
    val current_line = ref (the_default 1 (Position.line_of pos));
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
    74
30705
e8ab35c6ade6 get_index: produce index of next pending token, not the last one;
wenzelm
parents: 30704
diff changeset
    75
    fun get_index () =
e8ab35c6ade6 get_index: produce index of next pending token, not the last one;
wenzelm
parents: 30704
diff changeset
    76
      the_default 0 (get_first (fn (_, []) => NONE | (i, _) => SOME i) (! input_buffer));
e8ab35c6ade6 get_index: produce index of next pending token, not the last one;
wenzelm
parents: 30704
diff changeset
    77
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
    78
    fun get () =
30704
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    79
      (case ! input_buffer of
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
    80
        [] => NONE
30704
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    81
      | (_, []) :: rest => (input_buffer := rest; get ())
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
    82
      | (i, c :: cs) :: rest =>
30704
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    83
          (input_buffer := (i, cs) :: rest;
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
    84
           if c = #"\n" then current_line := ! current_line + 1 else ();
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
    85
           SOME c));
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
    86
30680
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
    87
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
    88
    (* output *)
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
    89
30704
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    90
    val output_buffer = ref Buffer.empty;
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    91
    fun output () = Buffer.content (! output_buffer);
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
    92
    fun put s = change output_buffer (Buffer.add s);
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
    93
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
    94
    fun put_message {message, hard, location, context = _} =
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
    95
     (put (if hard then "Error: " else "Warning: ");
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
    96
      put (Pretty.string_of (Pretty.from_ML (pretty_ml message)));
30744
50ccaef52871 export position_of;
wenzelm
parents: 30709
diff changeset
    97
      put (Position.str_of
50ccaef52871 export position_of;
wenzelm
parents: 30709
diff changeset
    98
        (position_of (Context.proof_of (the (Context.thread_data ()))) location) ^ "\n"));
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
    99
30680
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   100
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   101
    (* results *)
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   102
30709
d9ca766bf24c report ML typing;
wenzelm
parents: 30705
diff changeset
   103
    val depth = get_print_depth ();
d9ca766bf24c report ML typing;
wenzelm
parents: 30705
diff changeset
   104
30680
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   105
    fun apply_result {fixes, types, signatures, structures, functors, values} =
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   106
      let
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   107
        fun display disp x =
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   108
          if depth > 0 then
31235
67c796138bf0 adapted to Poly/ML SVN 719;
wenzelm
parents: 30744
diff changeset
   109
            (disp x |> pretty_ml |> Pretty.from_ML |> Pretty.string_of |> put; put "\n")
30680
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   110
          else ();
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   111
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   112
        fun apply_fix (a, b) =
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   113
          (display PolyML.NameSpace.displayFix (a, b); #enterFix space (a, b));
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   114
        fun apply_type (a, b) =
31235
67c796138bf0 adapted to Poly/ML SVN 719;
wenzelm
parents: 30744
diff changeset
   115
          (display PolyML.NameSpace.displayType (b, depth, space); #enterType space (a, b));
30680
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   116
        fun apply_sig (a, b) =
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   117
          (display PolyML.NameSpace.displaySig (b, depth, space); #enterSig space (a, b));
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   118
        fun apply_struct (a, b) =
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   119
          (display PolyML.NameSpace.displayStruct (b, depth, space); #enterStruct space (a, b));
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   120
        fun apply_funct (a, b) =
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   121
          (display PolyML.NameSpace.displayFunct (b, depth, space); #enterFunct space (a, b));
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   122
        fun apply_val (a, b) =
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   123
          (display PolyML.NameSpace.displayVal (b, depth, space); #enterVal space (a, b));
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   124
      in
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   125
        List.app apply_fix fixes;
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   126
        List.app apply_type types;
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   127
        List.app apply_sig signatures;
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   128
        List.app apply_struct structures;
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   129
        List.app apply_funct functors;
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   130
        List.app apply_val values
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   131
      end;
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   132
30704
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
   133
    fun result_fun (phase1, phase2) () =
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
   134
     (case phase1 of NONE => ()
30709
d9ca766bf24c report ML typing;
wenzelm
parents: 30705
diff changeset
   135
      | SOME parse_tree => report_parse_tree (the (Context.thread_data ())) depth parse_tree;
30704
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
   136
      case phase2 of NONE => error "Static Errors"
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
   137
      | SOME code => apply_result (Toplevel.program code));
30680
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   138
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   139
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   140
    (* compiler invocation *)
30646
d26ad4576d5c proper signature;
wenzelm
parents: 30644
diff changeset
   141
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   142
    val parameters =
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   143
     [PolyML.Compiler.CPOutStream put,
30680
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   144
      PolyML.Compiler.CPNameSpace space,
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   145
      PolyML.Compiler.CPErrorMessageProc put_message,
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   146
      PolyML.Compiler.CPLineNo (fn () => ! current_line),
30704
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
   147
      PolyML.Compiler.CPFileName (the_default "ML" (Position.file_of pos)),
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   148
      PolyML.Compiler.CPLineOffset get_index,
30680
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   149
      PolyML.Compiler.CPCompilerResultFun result_fun];
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   150
    val _ =
30704
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
   151
      (while not (List.null (! input_buffer)) do
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   152
        PolyML.compiler (get, parameters) ())
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   153
      handle exn =>
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   154
       (put ("Exception- " ^ General.exnMessage exn ^ " raised");
30677
df6ca2f50199 eliminated Output.ml_output;
wenzelm
parents: 30668
diff changeset
   155
        error (output ()); raise exn);
30680
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   156
  in if verbose then print (output ()) else () end;
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   157
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   158
val eval = use_text ML_Context.local_context;
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   159
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   160
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   161
(* ML test command *)
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   162
30680
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   163
fun ML_test (txt, pos) =
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   164
  let
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   165
    val _ = Position.report Markup.ML_source pos;
30644
2948f4494e86 ML_Lex.read_antiq;
wenzelm
parents: 30640
diff changeset
   166
    val ants = ML_Lex.read_antiq (Symbol_Pos.explode (txt, pos), pos);
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   167
    val ((env, body), env_ctxt) = ML_Context.eval_antiquotes (ants, pos) (Context.thread_data ());
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   168
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   169
    val _ = Context.setmp_thread_data env_ctxt
30680
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   170
        (fn () => (eval false Position.none env; Context.thread_data ())) ()
30704
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
   171
      |> (fn NONE => () | SOME context' => Context.>> (inherit_env context'));
30680
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   172
    val _ = eval true pos body;
0863bb353065 more systematic type use_context;
wenzelm
parents: 30677
diff changeset
   173
    val _ = eval false Position.none (ML_Lex.tokenize "structure Isabelle = struct end");
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   174
  in () end;
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   175
30644
2948f4494e86 ML_Lex.read_antiq;
wenzelm
parents: 30640
diff changeset
   176
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   177
local structure P = OuterParse and K = OuterKeyword in
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   178
30704
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
   179
fun propagate_env (context as Context.Proof lthy) =
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
   180
      Context.Proof (LocalTheory.map_contexts (inherit_env context) lthy)
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
   181
  | propagate_env context = context;
30644
2948f4494e86 ML_Lex.read_antiq;
wenzelm
parents: 30640
diff changeset
   182
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   183
val _ =
30644
2948f4494e86 ML_Lex.read_antiq;
wenzelm
parents: 30640
diff changeset
   184
  OuterSyntax.command "ML_test" "advanced ML compiler test" (K.tag_ml K.thy_decl)
2948f4494e86 ML_Lex.read_antiq;
wenzelm
parents: 30640
diff changeset
   185
    (P.ML_source >> (fn src =>
30704
d6d4828e74a2 register token positions persistently with context;
wenzelm
parents: 30681
diff changeset
   186
      Toplevel.generic_theory (ML_Context.exec (fn () => ML_test src) #> propagate_env)));
30640
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   187
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   188
end;
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   189
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   190
end;
3f3d1e218b64 Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff changeset
   191