src/Pure/ML/ml_test.ML
changeset 31333 fcd010617e6c
parent 31332 9639a6d4d714
child 31334 999fa9e1dbdd
equal deleted inserted replaced
31332:9639a6d4d714 31333:fcd010617e6c
     1 (*  Title:      Pure/ML/ml_test.ML
       
     2     Author:     Makarius
       
     3 
       
     4 Test of advanced ML compiler invocation in Poly/ML 5.3 (SVN 762).
       
     5 *)
       
     6 
       
     7 signature ML_TEST =
       
     8 sig
       
     9   val position_of: Proof.context -> PolyML.location -> Position.T
       
    10   val eval: bool -> Position.T -> ML_Lex.token list -> unit
       
    11 end
       
    12 
       
    13 structure ML_Test: ML_TEST =
       
    14 struct
       
    15 
       
    16 (* extra ML environment *)
       
    17 
       
    18 fun position_of ctxt
       
    19     ({file, startLine = line, startPosition = i, endPosition = j, ...}: PolyML.location) =
       
    20   (case pairself (ML_Env.token_position ctxt) (i, j) of
       
    21     (SOME pos1, SOME pos2) => Position.encode_range (pos1, pos2)
       
    22   | (SOME pos, NONE) => pos
       
    23   | _ => Position.line_file line file);
       
    24 
       
    25 
       
    26 (* parse trees *)
       
    27 
       
    28 fun report_parse_tree context depth space =
       
    29   let
       
    30     val pos_of = position_of (Context.proof_of context);
       
    31     fun report loc (PolyML.PTtype types) =
       
    32           PolyML.NameSpace.displayTypeExpression (types, depth, space)
       
    33           |> pretty_ml |> Pretty.from_ML |> Pretty.string_of
       
    34           |> Position.report_text Markup.ML_typing (pos_of loc)
       
    35       | report loc (PolyML.PTdeclaredAt decl) =
       
    36           Markup.markup (Markup.properties (Position.properties_of (pos_of decl)) Markup.ML_def) ""
       
    37           |> Position.report_text Markup.ML_ref (pos_of loc)
       
    38       | report _ (PolyML.PTnextSibling tree) = report_tree (tree ())
       
    39       | report _ (PolyML.PTfirstChild tree) = report_tree (tree ())
       
    40       | report _ _ = ()
       
    41     and report_tree (loc, props) = List.app (report loc) props;
       
    42   in report_tree end;
       
    43 
       
    44 
       
    45 (* eval ML source tokens *)
       
    46 
       
    47 fun use_text ({name_space = space, print, error, ...}: use_context) verbose pos toks =
       
    48   let
       
    49     (* input *)
       
    50 
       
    51     val input = Context.>>> (ML_Env.register_tokens toks);
       
    52     val input_buffer = ref (map (apsnd (String.explode o ML_Lex.text_of)) input);
       
    53 
       
    54     val current_line = ref (the_default 1 (Position.line_of pos));
       
    55 
       
    56     fun get_index () =
       
    57       the_default 0 (get_first (fn (_, []) => NONE | (i, _) => SOME i) (! input_buffer));
       
    58 
       
    59     fun get () =
       
    60       (case ! input_buffer of
       
    61         [] => NONE
       
    62       | (_, []) :: rest => (input_buffer := rest; get ())
       
    63       | (i, c :: cs) :: rest =>
       
    64           (input_buffer := (i, cs) :: rest;
       
    65            if c = #"\n" then current_line := ! current_line + 1 else ();
       
    66            SOME c));
       
    67 
       
    68 
       
    69     (* output *)
       
    70 
       
    71     val output_buffer = ref Buffer.empty;
       
    72     fun output () = Buffer.content (! output_buffer);
       
    73     fun put s = change output_buffer (Buffer.add s);
       
    74 
       
    75     fun put_message {message, hard, location, context = _} =
       
    76      (put (if hard then "Error: " else "Warning: ");
       
    77       put (Pretty.string_of (Pretty.from_ML (pretty_ml message)));
       
    78       put (Position.str_of
       
    79         (position_of (Context.proof_of (the (Context.thread_data ()))) location) ^ "\n"));
       
    80 
       
    81 
       
    82     (* results *)
       
    83 
       
    84     val depth = get_print_depth ();
       
    85 
       
    86     fun apply_result {fixes, types, signatures, structures, functors, values} =
       
    87       let
       
    88         fun display disp x =
       
    89           if depth > 0 then
       
    90             (disp x |> pretty_ml |> Pretty.from_ML |> Pretty.string_of |> put; put "\n")
       
    91           else ();
       
    92 
       
    93         fun apply_fix (a, b) =
       
    94           (display PolyML.NameSpace.displayFix (a, b); #enterFix space (a, b));
       
    95         fun apply_type (a, b) =
       
    96           (display PolyML.NameSpace.displayType (b, depth, space); #enterType space (a, b));
       
    97         fun apply_sig (a, b) =
       
    98           (display PolyML.NameSpace.displaySig (b, depth, space); #enterSig space (a, b));
       
    99         fun apply_struct (a, b) =
       
   100           (display PolyML.NameSpace.displayStruct (b, depth, space); #enterStruct space (a, b));
       
   101         fun apply_funct (a, b) =
       
   102           (display PolyML.NameSpace.displayFunct (b, depth, space); #enterFunct space (a, b));
       
   103         fun apply_val (a, b) =
       
   104           (display PolyML.NameSpace.displayVal (b, depth, space); #enterVal space (a, b));
       
   105       in
       
   106         List.app apply_fix fixes;
       
   107         List.app apply_type types;
       
   108         List.app apply_sig signatures;
       
   109         List.app apply_struct structures;
       
   110         List.app apply_funct functors;
       
   111         List.app apply_val values
       
   112       end;
       
   113 
       
   114     fun result_fun (phase1, phase2) () =
       
   115      (case phase1 of NONE => ()
       
   116       | SOME parse_tree => report_parse_tree (the (Context.thread_data ())) depth space parse_tree;
       
   117       case phase2 of NONE => error "Static Errors"
       
   118       | SOME code => apply_result (Toplevel.program code));
       
   119 
       
   120 
       
   121     (* compiler invocation *)
       
   122 
       
   123     val parameters =
       
   124      [PolyML.Compiler.CPOutStream put,
       
   125       PolyML.Compiler.CPNameSpace space,
       
   126       PolyML.Compiler.CPErrorMessageProc put_message,
       
   127       PolyML.Compiler.CPLineNo (fn () => ! current_line),
       
   128       PolyML.Compiler.CPFileName (the_default "ML" (Position.file_of pos)),
       
   129       PolyML.Compiler.CPLineOffset get_index,
       
   130       PolyML.Compiler.CPCompilerResultFun result_fun];
       
   131     val _ =
       
   132       (while not (List.null (! input_buffer)) do
       
   133         PolyML.compiler (get, parameters) ())
       
   134       handle exn =>
       
   135        (put ("Exception- " ^ General.exnMessage exn ^ " raised");
       
   136         error (output ()); raise exn);
       
   137   in if verbose then print (output ()) else () end;
       
   138 
       
   139 val eval = use_text ML_Env.local_context;
       
   140 
       
   141 
       
   142 (* ML test command *)
       
   143 
       
   144 fun ML_test (txt, pos) =
       
   145   let
       
   146     val _ = Position.report Markup.ML_source pos;
       
   147     val ants = ML_Lex.read_antiq (Symbol_Pos.explode (txt, pos), pos);
       
   148     val ((env, body), env_ctxt) = ML_Context.eval_antiquotes (ants, pos) (Context.thread_data ());
       
   149 
       
   150     val _ = Context.setmp_thread_data env_ctxt
       
   151         (fn () => (eval false Position.none env; Context.thread_data ())) ()
       
   152       |> (fn NONE => () | SOME context' => Context.>> (ML_Env.inherit context'));
       
   153     val _ = eval true pos body;
       
   154     val _ = eval false Position.none (ML_Lex.tokenize "structure Isabelle = struct end");
       
   155   in () end;
       
   156 
       
   157 
       
   158 local structure P = OuterParse and K = OuterKeyword in
       
   159 
       
   160 fun propagate_env (context as Context.Proof lthy) =
       
   161       Context.Proof (LocalTheory.map_contexts (ML_Env.inherit context) lthy)
       
   162   | propagate_env context = context;
       
   163 
       
   164 val _ =
       
   165   OuterSyntax.command "ML_test" "advanced ML compiler test" (K.tag_ml K.thy_decl)
       
   166     (P.ML_source >> (fn src =>
       
   167       Toplevel.generic_theory (ML_Context.exec (fn () => ML_test src) #> propagate_env)));
       
   168 
       
   169 end;
       
   170 
       
   171 end;
       
   172