--- a/src/Pure/ML-Systems/polyml.ML Sat Jun 05 13:06:28 2004 +0200
+++ b/src/Pure/ML-Systems/polyml.ML Sat Jun 05 13:06:39 2004 +0200
@@ -3,7 +3,7 @@
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
Copyright 1991 University of Cambridge
-Compatibility file for Poly/ML (version 4.0).
+Compatibility file for Poly/ML (version 4.0, 4.1, 4.1.x).
*)
(** ML system related **)
@@ -26,14 +26,17 @@
val explode = SML90.explode;
val implode = SML90.implode;
+
(* compiler-independent timing functions *)
use "ML-Systems/cpu-timer-basis.ML";
+
(* bounded time execution *)
use "ML-Systems/polyml-time-limit.ML";
+
(* prompts *)
fun ml_prompts p1 p2 = (PolyML.Compiler.prompt1 := p1; PolyML.Compiler.prompt2 := p2);
--- a/src/Pure/sorts.ML Sat Jun 05 13:06:28 2004 +0200
+++ b/src/Pure/sorts.ML Sat Jun 05 13:06:39 2004 +0200
@@ -84,8 +84,8 @@
arities: table of association lists of all type arities; (t, ars)
means that type constructor t has the arities ars; an element (c,
Ss) of ars represents the arity t::(Ss)c. "Coregularity" of the
- arities structure requires that for any two declarations t:(Ss1)c1
- and t:(Ss2)c2 such that c1 <= c2 holds Ss1 <= Ss2.
+ arities structure requires that for any two declarations
+ t::(Ss1)c1 and t::(Ss2)c2 such that c1 <= c2 holds Ss1 <= Ss2.
*)
type classes = stamp Graph.T;