src/Pure/ML/ml_context.ML
author wenzelm
Mon, 20 Mar 2023 10:59:27 +0100
changeset 77692 3e746e684f4b
parent 76884 a004c5322ea4
child 78035 bd5f6cee8001
permissions -rw-r--r--
clarified operations for ML object sizes;
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
56070
1bc0bea908c3 ML_Context.check_antiquotation still required;
wenzelm
parents: 56069
diff changeset
     9
  val check_antiquotation: Proof.context -> xstring * Position.T -> string
59127
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
    10
  val struct_name: Proof.context -> string
59112
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 59067
diff changeset
    11
  val variant: string -> Proof.context -> string * Proof.context
73549
a2c589d5e1e4 clarified signature: more detailed token positions for antiquotations;
wenzelm
parents: 69592
diff changeset
    12
  type decl = Proof.context -> ML_Lex.token list * ML_Lex.token list
74562
8403bd51f8b1 clarified modules;
wenzelm
parents: 74561
diff changeset
    13
  type decls = Proof.context -> (ML_Lex.token list * ML_Lex.token list) list
73549
a2c589d5e1e4 clarified signature: more detailed token positions for antiquotations;
wenzelm
parents: 69592
diff changeset
    14
  type antiquotation = Position.range -> Token.src -> Proof.context -> decl * Proof.context
a2c589d5e1e4 clarified signature: more detailed token positions for antiquotations;
wenzelm
parents: 69592
diff changeset
    15
  val add_antiquotation: binding -> bool -> antiquotation -> theory -> theory
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59127
diff changeset
    16
  val print_antiquotations: bool -> Proof.context -> unit
73551
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
    17
  val expand_antiquotes: ML_Lex.token Antiquote.antiquote list ->
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
    18
    Proof.context -> decl * Proof.context
74562
8403bd51f8b1 clarified modules;
wenzelm
parents: 74561
diff changeset
    19
  val expand_antiquotes_list: ML_Lex.token Antiquote.antiquote list list ->
8403bd51f8b1 clarified modules;
wenzelm
parents: 74561
diff changeset
    20
    Proof.context -> decls * Proof.context
56275
600f432ab556 added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents: 56229
diff changeset
    21
  val eval: ML_Compiler.flags -> Position.T -> ML_Lex.token Antiquote.antiquote list -> unit
600f432ab556 added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents: 56229
diff changeset
    22
  val eval_file: ML_Compiler.flags -> Path.T -> unit
59064
a8bcb5a446c8 more abstract type Input.source;
wenzelm
parents: 59058
diff changeset
    23
  val eval_source: ML_Compiler.flags -> Input.source -> unit
56275
600f432ab556 added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents: 56229
diff changeset
    24
  val eval_in: Proof.context option -> ML_Compiler.flags -> Position.T ->
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
    25
    ML_Lex.token Antiquote.antiquote list -> unit
59064
a8bcb5a446c8 more abstract type Input.source;
wenzelm
parents: 59058
diff changeset
    26
  val eval_source_in: Proof.context option -> ML_Compiler.flags -> Input.source -> unit
62913
13252110a6fe eliminated unused simproc identifier;
wenzelm
parents: 62902
diff changeset
    27
  val exec: (unit -> unit) -> Context.generic -> Context.generic
69216
1a52baa70aed clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
wenzelm
parents: 69187
diff changeset
    28
  val expression: Position.T -> ML_Lex.token Antiquote.antiquote list ->
1a52baa70aed clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
wenzelm
parents: 69187
diff changeset
    29
    Context.generic -> Context.generic
25204
36cf92f63a44 replaced Secure.evaluate by ML_Context.evaluate;
wenzelm
parents: 25142
diff changeset
    30
end
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    31
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    32
structure ML_Context: ML_CONTEXT =
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    33
struct
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    34
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    35
(** ML antiquotations **)
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    36
59127
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
    37
(* names for generated environment *)
59112
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 59067
diff changeset
    38
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 59067
diff changeset
    39
structure Names = Proof_Data
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 59067
diff changeset
    40
(
59127
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
    41
  type T = string * Name.context;
62900
c641bf9402fd simplified default print_depth: context is usually available, in contrast to 0d295e339f52;
wenzelm
parents: 62889
diff changeset
    42
  val init_names = ML_Syntax.reserved |> Name.declare "ML_context";
59127
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
    43
  fun init _ = ("Isabelle0", init_names);
59112
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 59067
diff changeset
    44
);
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 59067
diff changeset
    45
59127
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
    46
fun struct_name ctxt = #1 (Names.get ctxt);
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
    47
val struct_begin = (Names.map o apfst) (fn _ => "Isabelle" ^ serial_string ());
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
    48
59112
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 59067
diff changeset
    49
fun variant a ctxt =
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 59067
diff changeset
    50
  let
59127
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
    51
    val names = #2 (Names.get ctxt);
59112
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 59067
diff changeset
    52
    val (b, names') = Name.variant (Name.desymbolize (SOME false) a) names;
59127
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
    53
    val ctxt' = (Names.map o apsnd) (K names') ctxt;
59112
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 59067
diff changeset
    54
  in (b, ctxt') end;
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 59067
diff changeset
    55
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 59067
diff changeset
    56
73549
a2c589d5e1e4 clarified signature: more detailed token positions for antiquotations;
wenzelm
parents: 69592
diff changeset
    57
(* theory data *)
59112
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 59067
diff changeset
    58
73549
a2c589d5e1e4 clarified signature: more detailed token positions for antiquotations;
wenzelm
parents: 69592
diff changeset
    59
type decl =
a2c589d5e1e4 clarified signature: more detailed token positions for antiquotations;
wenzelm
parents: 69592
diff changeset
    60
  Proof.context -> ML_Lex.token list * ML_Lex.token list;  (*final context -> ML env, ML body*)
59112
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 59067
diff changeset
    61
74562
8403bd51f8b1 clarified modules;
wenzelm
parents: 74561
diff changeset
    62
type decls = Proof.context -> (ML_Lex.token list * ML_Lex.token list) list;
8403bd51f8b1 clarified modules;
wenzelm
parents: 74561
diff changeset
    63
73549
a2c589d5e1e4 clarified signature: more detailed token positions for antiquotations;
wenzelm
parents: 69592
diff changeset
    64
type antiquotation =
a2c589d5e1e4 clarified signature: more detailed token positions for antiquotations;
wenzelm
parents: 69592
diff changeset
    65
  Position.range -> Token.src -> Proof.context -> decl * Proof.context;
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
    66
56069
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
    67
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
    68
(
73549
a2c589d5e1e4 clarified signature: more detailed token positions for antiquotations;
wenzelm
parents: 69592
diff changeset
    69
  type T = (bool * antiquotation) Name_Space.table;
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 48992
diff changeset
    70
  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
    71
  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
    72
);
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
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
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
    75
56070
1bc0bea908c3 ML_Context.check_antiquotation still required;
wenzelm
parents: 56069
diff changeset
    76
fun check_antiquotation ctxt =
1bc0bea908c3 ML_Context.check_antiquotation still required;
wenzelm
parents: 56069
diff changeset
    77
  #1 o Name_Space.check (Context.Proof ctxt) (get_antiquotations ctxt);
1bc0bea908c3 ML_Context.check_antiquotation still required;
wenzelm
parents: 56069
diff changeset
    78
73549
a2c589d5e1e4 clarified signature: more detailed token positions for antiquotations;
wenzelm
parents: 69592
diff changeset
    79
fun add_antiquotation name embedded antiquotation thy =
a2c589d5e1e4 clarified signature: more detailed token positions for antiquotations;
wenzelm
parents: 69592
diff changeset
    80
  thy |> Antiquotations.map
a2c589d5e1e4 clarified signature: more detailed token positions for antiquotations;
wenzelm
parents: 69592
diff changeset
    81
    (Name_Space.define (Context.Theory thy) true (name, (embedded, antiquotation)) #> #2);
55743
225a060e7445 proper context for global data;
wenzelm
parents: 55526
diff changeset
    82
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59127
diff changeset
    83
fun print_antiquotations verbose ctxt =
56069
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
    84
  Pretty.big_list "ML antiquotations:"
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59127
diff changeset
    85
    (map (Pretty.mark_str o #1) (Name_Space.markup_table verbose ctxt (get_antiquotations ctxt)))
56069
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 56032
diff changeset
    86
  |> Pretty.writeln;
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
    87
73549
a2c589d5e1e4 clarified signature: more detailed token positions for antiquotations;
wenzelm
parents: 69592
diff changeset
    88
fun apply_antiquotation range src ctxt =
69592
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
    89
  let
73549
a2c589d5e1e4 clarified signature: more detailed token positions for antiquotations;
wenzelm
parents: 69592
diff changeset
    90
    val (src', (embedded, antiquotation)) = Token.check_src ctxt get_antiquotations src;
69592
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
    91
    val _ =
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
    92
      if Options.default_bool "update_control_cartouches" then
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
    93
        Context_Position.reports_text ctxt
73549
a2c589d5e1e4 clarified signature: more detailed token positions for antiquotations;
wenzelm
parents: 69592
diff changeset
    94
          (Antiquote.update_reports embedded (#1 range) (map Token.content_of src'))
69592
a80d8ec6c998 support for isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
    95
      else ();
73549
a2c589d5e1e4 clarified signature: more detailed token positions for antiquotations;
wenzelm
parents: 69592
diff changeset
    96
  in antiquotation range src' ctxt end;
43563
aeabb735883a proper checking of @{ML_antiquotation};
wenzelm
parents: 43560
diff changeset
    97
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
    98
73551
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
    99
(* parsing *)
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   100
27343
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   101
local
4b28b80dd1f8 add_antiq: more general notion of ML antiquotation;
wenzelm
parents: 26881
diff changeset
   102
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   103
val antiq =
63671
eb4f59275c05 liberal name as in document antiquotations;
wenzelm
parents: 62969
diff changeset
   104
  Parse.!!! ((Parse.token Parse.liberal_name ::: Parse.args) --| Scan.ahead Parse.eof);
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   105
73551
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   106
fun expand_antiquote ant ctxt =
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   107
  (case ant of
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   108
    Antiquote.Text tok => (K ([], [tok]), ctxt)
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   109
  | Antiquote.Control {name, range, body} =>
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   110
      ctxt |> apply_antiquotation range
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   111
        (Token.make_src name (if null body then [] else [Token.read_cartouche body]))
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   112
  | Antiquote.Antiq {range, body, ...} =>
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   113
      ctxt |> apply_antiquotation range
74564
0a66a61e740c clarified modules;
wenzelm
parents: 74562
diff changeset
   114
        (Parse.read_antiq (Thy_Header.get_keywords' ctxt) antiq (body, #1 range)));
73551
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   115
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   116
in
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   117
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   118
fun expand_antiquotes ants ctxt =
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   119
  let
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   120
    val (decls, ctxt') = fold_map expand_antiquote ants ctxt;
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   121
    fun decl ctxt'' = decls |> map (fn d => d ctxt'') |> split_list |> apply2 flat;
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   122
  in (decl, ctxt') end;
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   123
74562
8403bd51f8b1 clarified modules;
wenzelm
parents: 74561
diff changeset
   124
fun expand_antiquotes_list antss ctxt =
8403bd51f8b1 clarified modules;
wenzelm
parents: 74561
diff changeset
   125
  let
8403bd51f8b1 clarified modules;
wenzelm
parents: 74561
diff changeset
   126
    val (decls, ctxt') = fold_map expand_antiquotes antss ctxt;
8403bd51f8b1 clarified modules;
wenzelm
parents: 74561
diff changeset
   127
    fun decl' ctxt'' = map (fn decl => decl ctxt'') decls;
8403bd51f8b1 clarified modules;
wenzelm
parents: 74561
diff changeset
   128
  in (decl', ctxt') end
8403bd51f8b1 clarified modules;
wenzelm
parents: 74561
diff changeset
   129
73551
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   130
end;
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   131
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   132
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   133
(* evaluation *)
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   134
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   135
local
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   136
59127
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
   137
fun make_env name visible =
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
   138
  (ML_Lex.tokenize
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
   139
    ("structure " ^ name ^ " =\nstruct\n\
52663
6e71d43775e5 retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents: 50201
diff changeset
   140
     \val ML_context = Context_Position.set_visible " ^ Bool.toString visible ^
62900
c641bf9402fd simplified default print_depth: context is usually available, in contrast to 0d295e339f52;
wenzelm
parents: 62889
diff changeset
   141
     " (Context.the_local_context ());\n"),
59127
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
   142
   ML_Lex.tokenize "end;");
48776
37cd53e69840 faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents: 47005
diff changeset
   143
59127
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
   144
fun reset_env name = ML_Lex.tokenize ("structure " ^ name ^ " = struct end");
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   145
73551
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   146
fun eval_antiquotes ants opt_context =
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   147
  if forall Antiquote.is_text ants orelse is_none opt_context then
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   148
    (([], map (Antiquote.the_text "No context -- cannot expand ML antiquotations") ants),
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   149
      Option.map Context.proof_of opt_context)
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   150
  else
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   151
    let
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   152
      val context = the opt_context;
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   153
      val visible =
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   154
        (case context of
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   155
          Context.Proof ctxt => Context_Position.is_visible ctxt
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   156
        | _ => true);
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   157
      val ctxt = struct_begin (Context.proof_of context);
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   158
      val (begin_env, end_env) = make_env (struct_name ctxt) visible;
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   159
      val (decl, ctxt') = expand_antiquotes ants ctxt;
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   160
      val (ml_env, ml_body) = decl ctxt';
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   161
    in ((begin_env @ ml_env @ end_env, ml_body), SOME ctxt') end;
59127
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
   162
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
   163
in
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   164
56275
600f432ab556 added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents: 56229
diff changeset
   165
fun eval flags pos ants =
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   166
  let
56304
40274e4f5ebf redirect ML_Compiler reports more directly: only the (big) parse tree report is deferred via Execution.print (NB: this does not work for asynchronous "diag" commands);
wenzelm
parents: 56303
diff changeset
   167
    val non_verbose = ML_Compiler.verbose false flags;
56275
600f432ab556 added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents: 56229
diff changeset
   168
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   169
    (*prepare source text*)
73551
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 73549
diff changeset
   170
    val ((env, body), env_ctxt) = eval_antiquotes ants (Context.get_generic_context ());
31334
999fa9e1dbdd structure ML_Compiler;
wenzelm
parents: 31325
diff changeset
   171
    val _ =
59127
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
   172
      (case env_ctxt of
41485
6c0de045d127 ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents: 41375
diff changeset
   173
        SOME ctxt =>
56303
4cc3f4db3447 clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
wenzelm
parents: 56294
diff changeset
   174
          if Config.get ctxt ML_Options.source_trace andalso Context_Position.is_visible ctxt
56294
85911b8a6868 prefer Context_Position where a context is available;
wenzelm
parents: 56279
diff changeset
   175
          then tracing (cat_lines [ML_Lex.flatten env, ML_Lex.flatten body])
41375
7a89b4b94817 configuration option "ML_trace";
wenzelm
parents: 39911
diff changeset
   176
          else ()
7a89b4b94817 configuration option "ML_trace";
wenzelm
parents: 39911
diff changeset
   177
      | NONE => ());
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   178
59127
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
   179
    (*prepare environment*)
41485
6c0de045d127 ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents: 41375
diff changeset
   180
    val _ =
62889
99c7f31615c2 clarified modules;
wenzelm
parents: 62878
diff changeset
   181
      Context.setmp_generic_context
59127
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
   182
        (Option.map (Context.Proof o Context_Position.set_visible false) env_ctxt)
62889
99c7f31615c2 clarified modules;
wenzelm
parents: 62878
diff changeset
   183
        (fn () =>
99c7f31615c2 clarified modules;
wenzelm
parents: 62878
diff changeset
   184
          (ML_Compiler.eval non_verbose Position.none env; Context.get_generic_context ())) ()
68865
dd44e31ca2c6 support multiple inheritance of ML environments, with canonical merge order as in Context.begin_theory;
wenzelm
parents: 68823
diff changeset
   185
      |> (fn NONE => () | SOME context' => Context.>> (ML_Env.inherit [context']));
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   186
59127
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
   187
    (*eval body*)
56275
600f432ab556 added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents: 56229
diff changeset
   188
    val _ = ML_Compiler.eval flags pos body;
59127
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
   189
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
   190
    (*clear environment*)
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
   191
    val _ =
62889
99c7f31615c2 clarified modules;
wenzelm
parents: 62878
diff changeset
   192
      (case (env_ctxt, is_some (Context.get_generic_context ())) of
59127
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
   193
        (SOME ctxt, true) =>
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
   194
          let
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
   195
            val name = struct_name ctxt;
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
   196
            val _ = ML_Compiler.eval non_verbose Position.none (reset_env name);
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
   197
            val _ = Context.>> (ML_Env.forget_structure name);
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
   198
          in () end
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
   199
      | _ => ());
28270
7ada24ebea94 explicit handling of ML environment within generic context;
wenzelm
parents: 27895
diff changeset
   200
  in () end;
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   201
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   202
end;
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   203
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   204
30672
beaadd5af500 more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents: 30643
diff changeset
   205
(* derived versions *)
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   206
56275
600f432ab556 added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents: 56229
diff changeset
   207
fun eval_file flags path =
76884
a004c5322ea4 clarified modules;
wenzelm
parents: 76880
diff changeset
   208
  let val pos = Position.file (File.symbolic_path path)
68823
5e7b1ae10eb8 clarified signature;
wenzelm
parents: 68821
diff changeset
   209
  in eval flags pos (ML_Lex.read_text (File.read path, pos)) end;
56203
76c72f4d0667 clarified bootstrap process: switch to ML with context and antiquotations earlier;
wenzelm
parents: 56201
diff changeset
   210
56275
600f432ab556 added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents: 56229
diff changeset
   211
fun eval_source flags source =
68820
2e4df245754e clarified environment: allow "read>write" specification;
wenzelm
parents: 68816
diff changeset
   212
  let
2e4df245754e clarified environment: allow "read>write" specification;
wenzelm
parents: 68816
diff changeset
   213
    val opt_context = Context.get_generic_context ();
68821
877534be1930 explicit setup of operations: avoid hardwired stuff;
wenzelm
parents: 68820
diff changeset
   214
    val {read_source, ...} = ML_Env.operations opt_context (#environment flags);
877534be1930 explicit setup of operations: avoid hardwired stuff;
wenzelm
parents: 68820
diff changeset
   215
  in eval flags (Input.pos_of source) (read_source source) end;
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
   216
56275
600f432ab556 added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents: 56229
diff changeset
   217
fun eval_in ctxt flags pos ants =
62889
99c7f31615c2 clarified modules;
wenzelm
parents: 62878
diff changeset
   218
  Context.setmp_generic_context (Option.map Context.Proof ctxt)
56275
600f432ab556 added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents: 56229
diff changeset
   219
    (fn () => eval flags pos ants) ();
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   220
56275
600f432ab556 added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents: 56229
diff changeset
   221
fun eval_source_in ctxt flags source =
62889
99c7f31615c2 clarified modules;
wenzelm
parents: 62878
diff changeset
   222
  Context.setmp_generic_context (Option.map Context.Proof ctxt)
56275
600f432ab556 added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents: 56229
diff changeset
   223
    (fn () => eval_source flags 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
   224
62913
13252110a6fe eliminated unused simproc identifier;
wenzelm
parents: 62902
diff changeset
   225
fun exec (e: unit -> unit) context =
13252110a6fe eliminated unused simproc identifier;
wenzelm
parents: 62902
diff changeset
   226
  (case Context.setmp_generic_context (SOME context)
13252110a6fe eliminated unused simproc identifier;
wenzelm
parents: 62902
diff changeset
   227
      (fn () => (e (); Context.get_generic_context ())) () of
13252110a6fe eliminated unused simproc identifier;
wenzelm
parents: 62902
diff changeset
   228
    SOME context' => context'
13252110a6fe eliminated unused simproc identifier;
wenzelm
parents: 62902
diff changeset
   229
  | NONE => error "Missing context after execution");
13252110a6fe eliminated unused simproc identifier;
wenzelm
parents: 62902
diff changeset
   230
69216
1a52baa70aed clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
wenzelm
parents: 69187
diff changeset
   231
fun expression pos ants = exec (fn () => eval ML_Compiler.flags pos ants);
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
   232
24574
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   233
end;
e840872e9c7c moved ML_XXX.ML files to Pure/ML;
wenzelm
parents:
diff changeset
   234
62931
09b516854210 flags as in 'ML' command;
wenzelm
parents: 62913
diff changeset
   235
val ML = ML_Context.eval_source (ML_Compiler.verbose true ML_Compiler.flags);