src/Pure/ML/ml_name_space.ML
author wenzelm
Sun, 03 Apr 2016 23:28:48 +0200
changeset 62839 ea9f12e422c7
parent 62819 d3ff367a16a0
child 62851 07eea2843b82
permissions -rw-r--r--
clarified SML name space: no access to structure PolyML;

(*  Title:      Pure/ML/ml_name_space.ML
    Author:     Makarius

Name space for Poly/ML.
*)

structure ML_Name_Space =
struct
  val critical_structures = ["CInterface", "Foreign", "RunCall", "RuntimeCalls", "Signal"];

  open PolyML.NameSpace;

  type T = PolyML.NameSpace.nameSpace;

  val global = PolyML.globalNameSpace;
  fun global_values values : T =
   {lookupVal = #lookupVal global,
    lookupType = #lookupType global,
    lookupStruct = #lookupStruct global,
    lookupFix = #lookupFix global,
    lookupSig = #lookupSig global,
    lookupFunct = #lookupFunct global,
    enterVal =
      fn (x, value) =>
        (case List.find (fn (y, _) => x = y) values of
          SOME (_, x') => #enterVal global (x', value)
        | NONE => ()),
    enterType = fn _ => (),
    enterFix = fn _ => (),
    enterStruct = fn _ => (),
    enterSig = fn _ => (),
    enterFunct = fn _ => (),
    allVal = #allVal global,
    allType = #allType global,
    allFix = #allFix global,
    allStruct = #allStruct global,
    allSig = #allSig global,
    allFunct = #allFunct global};

  type valueVal = Values.value;
  fun displayVal (x, depth, space) = Values.printWithType (x, depth, SOME space);
  fun displayTypeExpression (x, depth, space) = Values.printType (x, depth, SOME space);
  val forget_val = PolyML.Compiler.forgetValue;

  type typeVal = TypeConstrs.typeConstr;
  fun displayType (x, depth, space) = TypeConstrs.print (x, depth, SOME space);
  val forget_type = PolyML.Compiler.forgetType;

  type fixityVal = Infixes.fixity;
  fun displayFix (_: string, x) = Infixes.print x;

  type structureVal = Structures.structureVal;
  fun displayStruct (x, depth, space) = Structures.print (x, depth, SOME space);
  val forget_structure = PolyML.Compiler.forgetStructure;

  type signatureVal = Signatures.signatureVal;
  fun displaySig (x, depth, space) = Signatures.print (x, depth, SOME space);

  type functorVal = Functors.functorVal;
  fun displayFunct (x, depth, space) = Functors.print (x, depth, SOME space);


  (* Standard ML name space *)

  val sml_val = List.filter (fn (a, _) => a <> "use" andalso a <> "exit") (#allVal global ());
  val sml_type = #allType global ();
  val sml_fixity = #allFix global ();
  val sml_structure =
    List.filter (fn (a, _) => a <> "PolyML" andalso a <> "ML_System" andalso
      not (List.exists (fn b => a = b) critical_structures)) (#allStruct global ());
  val sml_signature =
    List.filter (fn (a, _) => a <> "ML_SYSTEM" andalso a <> "ML_NAME_SPACE") (#allSig global ());
  val sml_functor = #allFunct global ();
end;