src/Pure/General/secure.ML
author wenzelm
Wed, 25 Jul 2007 17:05:45 +0200
changeset 23978 6e8d5a05ffe8
parent 23922 707639e9497d
child 24058 81aafd465662
permissions -rw-r--r--
added use_noncritical;

(*  Title:      Pure/General/secure.ML
    ID:         $Id$
    Author:     Makarius

Secure critical operations.
*)

signature SECURE =
sig
  val set_secure: unit -> unit
  val is_secure: unit -> bool
  val use_text: string -> (string -> unit) * (string -> 'a) -> bool -> string -> unit
  val use_file: (string -> unit) * (string -> 'a) -> bool -> string -> unit
  val use_noncritical: string -> unit
  val use: string -> unit
  val commit: unit -> unit
  val execute: string -> string
  val system: string -> int
end;

structure Secure: SECURE =
struct

(** secure flag **)

val secure = ref false;

fun set_secure () = secure := true;
fun is_secure () = ! secure;

fun deny_secure msg = if is_secure () then error msg else ();


(** critical operations **)

(* ML evaluation *)

fun secure_mltext () = deny_secure "Cannot evaluate ML source in secure mode";

val orig_use_text = use_text;
val orig_use_file = use_file;

fun use_text name pr verbose txt = CRITICAL (fn () =>
  (secure_mltext (); orig_use_text name pr verbose txt));

fun use_file pr verbose name = CRITICAL (fn () =>
  (secure_mltext (); orig_use_file pr verbose name));

fun use name = use_file Output.ml_output true name;

fun use_noncritical name =
  (secure_mltext (); orig_use_file Output.ml_output true name);

(*commit is dynamically bound!*)
fun commit () = CRITICAL (fn () => orig_use_text "" Output.ml_output false "commit();");


(* shell commands *)

fun secure_shell () = deny_secure "Cannot execute shell commands in secure mode";

val orig_execute = execute;
val orig_system = system;

fun execute s = CRITICAL (fn () => (secure_shell (); orig_execute s));
fun system s = CRITICAL (fn () => (secure_shell (); orig_system s));

end;

(*override previous toplevel bindings!*)
val use_text = Secure.use_text;
val use_file = Secure.use_file;
val use = Secure.use;
val execute = Secure.execute;
val system = Secure.system;