src/Pure/ML/ml_env.ML
author wenzelm
Fri, 16 Apr 2010 11:39:08 +0200
changeset 36163 823c9400eb62
parent 33519 e31a85f92ce9
child 36165 310a3f2f0e7e
permissions -rw-r--r--
proper checking of ML functors (in Poly/ML 5.2 or later); eliminated pathetic comments;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31325
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/ML/ml_env.ML
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
     3
31470
01e7a8bb9e5b tuned comments;
wenzelm
parents: 31430
diff changeset
     4
Local environment of ML results.
31325
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
     5
*)
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
     6
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
     7
signature ML_ENV =
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
     8
sig
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
     9
  val inherit: Context.generic -> Context.generic -> Context.generic
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    10
  val name_space: ML_Name_Space.T
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    11
  val local_context: use_context
36163
823c9400eb62 proper checking of ML functors (in Poly/ML 5.2 or later);
wenzelm
parents: 33519
diff changeset
    12
  val check_functor: string -> unit
31325
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    13
end
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    14
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    15
structure ML_Env: ML_ENV =
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    16
struct
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    17
31328
0d3aa0aeb96f maintain tokens within common ML environment;
wenzelm
parents: 31325
diff changeset
    18
(* context data *)
0d3aa0aeb96f maintain tokens within common ML environment;
wenzelm
parents: 31325
diff changeset
    19
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 31470
diff changeset
    20
structure Env = Generic_Data
31325
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    21
(
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    22
  type T =
31430
04192aa43112 eliminated costly registration of tokens;
wenzelm
parents: 31331
diff changeset
    23
    ML_Name_Space.valueVal Symtab.table *
04192aa43112 eliminated costly registration of tokens;
wenzelm
parents: 31331
diff changeset
    24
    ML_Name_Space.typeVal Symtab.table *
04192aa43112 eliminated costly registration of tokens;
wenzelm
parents: 31331
diff changeset
    25
    ML_Name_Space.fixityVal Symtab.table *
04192aa43112 eliminated costly registration of tokens;
wenzelm
parents: 31331
diff changeset
    26
    ML_Name_Space.structureVal Symtab.table *
04192aa43112 eliminated costly registration of tokens;
wenzelm
parents: 31331
diff changeset
    27
    ML_Name_Space.signatureVal Symtab.table *
04192aa43112 eliminated costly registration of tokens;
wenzelm
parents: 31331
diff changeset
    28
    ML_Name_Space.functorVal Symtab.table;
04192aa43112 eliminated costly registration of tokens;
wenzelm
parents: 31331
diff changeset
    29
  val empty = (Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty);
31325
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    30
  val extend = I;
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 31470
diff changeset
    31
  fun merge
31430
04192aa43112 eliminated costly registration of tokens;
wenzelm
parents: 31331
diff changeset
    32
   ((val1, type1, fixity1, structure1, signature1, functor1),
04192aa43112 eliminated costly registration of tokens;
wenzelm
parents: 31331
diff changeset
    33
    (val2, type2, fixity2, structure2, signature2, functor2)) : T =
31328
0d3aa0aeb96f maintain tokens within common ML environment;
wenzelm
parents: 31325
diff changeset
    34
    (Symtab.merge (K true) (val1, val2),
0d3aa0aeb96f maintain tokens within common ML environment;
wenzelm
parents: 31325
diff changeset
    35
     Symtab.merge (K true) (type1, type2),
0d3aa0aeb96f maintain tokens within common ML environment;
wenzelm
parents: 31325
diff changeset
    36
     Symtab.merge (K true) (fixity1, fixity2),
0d3aa0aeb96f maintain tokens within common ML environment;
wenzelm
parents: 31325
diff changeset
    37
     Symtab.merge (K true) (structure1, structure2),
0d3aa0aeb96f maintain tokens within common ML environment;
wenzelm
parents: 31325
diff changeset
    38
     Symtab.merge (K true) (signature1, signature2),
31430
04192aa43112 eliminated costly registration of tokens;
wenzelm
parents: 31331
diff changeset
    39
     Symtab.merge (K true) (functor1, functor2));
31325
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    40
);
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    41
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    42
val inherit = Env.put o Env.get;
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    43
31328
0d3aa0aeb96f maintain tokens within common ML environment;
wenzelm
parents: 31325
diff changeset
    44
0d3aa0aeb96f maintain tokens within common ML environment;
wenzelm
parents: 31325
diff changeset
    45
(* results *)
0d3aa0aeb96f maintain tokens within common ML environment;
wenzelm
parents: 31325
diff changeset
    46
31325
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    47
val name_space: ML_Name_Space.T =
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    48
  let
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    49
    fun lookup sel1 sel2 name =
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    50
      Context.thread_data ()
31430
04192aa43112 eliminated costly registration of tokens;
wenzelm
parents: 31331
diff changeset
    51
      |> (fn NONE => NONE | SOME context => Symtab.lookup (sel1 (Env.get context)) name)
31325
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    52
      |> (fn NONE => sel2 ML_Name_Space.global name | some => some);
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    53
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    54
    fun all sel1 sel2 () =
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    55
      Context.thread_data ()
31430
04192aa43112 eliminated costly registration of tokens;
wenzelm
parents: 31331
diff changeset
    56
      |> (fn NONE => [] | SOME context => Symtab.dest (sel1 (Env.get context)))
31325
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    57
      |> append (sel2 ML_Name_Space.global ())
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    58
      |> sort_distinct (string_ord o pairself #1);
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    59
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    60
    fun enter ap1 sel2 entry =
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    61
      if is_some (Context.thread_data ()) then
31430
04192aa43112 eliminated costly registration of tokens;
wenzelm
parents: 31331
diff changeset
    62
        Context.>> (Env.map (ap1 (Symtab.update entry)))
31325
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    63
      else sel2 ML_Name_Space.global entry;
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    64
  in
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    65
   {lookupVal    = lookup #1 #lookupVal,
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    66
    lookupType   = lookup #2 #lookupType,
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    67
    lookupFix    = lookup #3 #lookupFix,
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    68
    lookupStruct = lookup #4 #lookupStruct,
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    69
    lookupSig    = lookup #5 #lookupSig,
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    70
    lookupFunct  = lookup #6 #lookupFunct,
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    71
    enterVal     = enter (fn h => fn (a, b, c, d, e, f) => (h a, b, c, d, e, f)) #enterVal,
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    72
    enterType    = enter (fn h => fn (a, b, c, d, e, f) => (a, h b, c, d, e, f)) #enterType,
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    73
    enterFix     = enter (fn h => fn (a, b, c, d, e, f) => (a, b, h c, d, e, f)) #enterFix,
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    74
    enterStruct  = enter (fn h => fn (a, b, c, d, e, f) => (a, b, c, h d, e, f)) #enterStruct,
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    75
    enterSig     = enter (fn h => fn (a, b, c, d, e, f) => (a, b, c, d, h e, f)) #enterSig,
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    76
    enterFunct   = enter (fn h => fn (a, b, c, d, e, f) => (a, b, c, d, e, h f)) #enterFunct,
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    77
    allVal       = all #1 #allVal,
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    78
    allType      = all #2 #allType,
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    79
    allFix       = all #3 #allFix,
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    80
    allStruct    = all #4 #allStruct,
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    81
    allSig       = all #5 #allSig,
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    82
    allFunct     = all #6 #allFunct}
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    83
  end;
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    84
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    85
val local_context: use_context =
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    86
 {tune_source = ML_Parse.fix_ints,
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    87
  name_space = name_space,
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    88
  str_of_pos = Position.str_of oo Position.line_file,
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    89
  print = writeln,
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    90
  error = error};
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    91
36163
823c9400eb62 proper checking of ML functors (in Poly/ML 5.2 or later);
wenzelm
parents: 33519
diff changeset
    92
val is_functor = is_some o #lookupFunct name_space;
823c9400eb62 proper checking of ML functors (in Poly/ML 5.2 or later);
wenzelm
parents: 33519
diff changeset
    93
823c9400eb62 proper checking of ML functors (in Poly/ML 5.2 or later);
wenzelm
parents: 33519
diff changeset
    94
fun check_functor name =
823c9400eb62 proper checking of ML functors (in Poly/ML 5.2 or later);
wenzelm
parents: 33519
diff changeset
    95
  if is_functor "Table" (*lookup actually works*) andalso is_functor name then ()
823c9400eb62 proper checking of ML functors (in Poly/ML 5.2 or later);
wenzelm
parents: 33519
diff changeset
    96
  else error ("Unknown ML functor: " ^ quote name);
823c9400eb62 proper checking of ML functors (in Poly/ML 5.2 or later);
wenzelm
parents: 33519
diff changeset
    97
31325
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    98
end;
700951b53d21 moved local ML environment to separate module ML_Env;
wenzelm
parents:
diff changeset
    99