merged
authorwenzelm
Sat, 21 Mar 2009 20:39:38 +0100
changeset 30634 bc3c1b7df4ec
parent 30632 5bfea958f893 (current diff)
parent 30633 cc18ae3c1c7f (diff)
child 30635 a7d175c48228
child 30647 ef8f46c3158a
merged
--- a/src/Pure/IsaMakefile	Sat Mar 21 20:22:13 2009 +0100
+++ b/src/Pure/IsaMakefile	Sat Mar 21 20:39:38 2009 +0100
@@ -69,7 +69,8 @@
   Isar/spec_parse.ML Isar/specification.ML Isar/theory_target.ML	\
   Isar/toplevel.ML Isar/value_parse.ML ML/ml_antiquote.ML		\
   ML/ml_context.ML ML/ml_lex.ML ML/ml_parse.ML ML/ml_syntax.ML		\
-  ML/ml_thms.ML ML-Systems/install_pp_polyml.ML Proof/extraction.ML	\
+  ML/ml_thms.ML ML-Systems/install_pp_polyml.ML				\
+  ML-Systems/install_pp_polyml-experimental.ML Proof/extraction.ML	\
   Proof/proof_rewrite_rules.ML Proof/proof_syntax.ML			\
   Proof/proofchecker.ML Proof/reconstruct.ML ProofGeneral/ROOT.ML	\
   ProofGeneral/pgip.ML ProofGeneral/pgip_input.ML			\
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/ML-Systems/install_pp_polyml-experimental.ML	Sat Mar 21 20:39:38 2009 +0100
@@ -0,0 +1,27 @@
+(*  Title:      Pure/ML-Systems/install_pp_polyml-experimental.ML
+
+Extra toplevel pretty-printing for Poly/ML; experimental version for
+Poly/ML 5.3.
+*)
+
+local
+
+fun pretty_future depth (pretty: 'a * int -> pretty) (x: 'a future) =
+  (case Future.peek x of
+    NONE => PrettyString "<future>"
+  | SOME (Exn.Exn _) => PrettyString "<failed>"
+  | SOME (Exn.Result y) => pretty (y, depth));
+
+fun pretty_lazy depth (pretty: 'a * int -> pretty) (x: 'a lazy) =
+  (case Lazy.peek x of
+    NONE => PrettyString "<lazy>"
+  | SOME (Exn.Exn _) => PrettyString "<failed>"
+  | SOME (Exn.Result y) => pretty (y, depth));
+
+in
+
+val _ = addPrettyPrinter pretty_future;
+val _ = addPrettyPrinter pretty_lazy;
+
+end;
+