make SML/NJ happy;
authorwenzelm
Sun, 25 Oct 2009 13:04:06 +0100
changeset 33159 369da293bbd4
parent 33156 57222d336c86
child 33160 ccbd4ad5a965
make SML/NJ happy;
src/Pure/Isar/attrib.ML
src/Pure/Isar/locale.ML
src/Pure/Isar/method.ML
src/Pure/simplifier.ML
src/Pure/theory.ML
src/Pure/thm.ML
--- a/src/Pure/Isar/attrib.ML	Sun Oct 25 11:58:11 2009 +0100
+++ b/src/Pure/Isar/attrib.ML	Sun Oct 25 13:04:06 2009 +0100
@@ -68,7 +68,7 @@
 structure Attributes = TheoryDataFun
 (
   type T = ((src -> attribute) * string) Name_Space.table;
-  val empty = Name_Space.empty_table "attribute";
+  val empty : T = Name_Space.empty_table "attribute";
   val copy = I;
   val extend = I;
   fun merge _ tables : T = Name_Space.merge_tables tables;
--- a/src/Pure/Isar/locale.ML	Sun Oct 25 11:58:11 2009 +0100
+++ b/src/Pure/Isar/locale.ML	Sun Oct 25 13:04:06 2009 +0100
@@ -125,7 +125,7 @@
 structure Locales = TheoryDataFun
 (
   type T = locale Name_Space.table;
-  val empty = Name_Space.empty_table "locale";
+  val empty : T = Name_Space.empty_table "locale";
   val copy = I;
   val extend = I;
   fun merge _ = Name_Space.join_tables (K merge_locale);
--- a/src/Pure/Isar/method.ML	Sun Oct 25 11:58:11 2009 +0100
+++ b/src/Pure/Isar/method.ML	Sun Oct 25 13:04:06 2009 +0100
@@ -323,7 +323,7 @@
 structure Methods = TheoryDataFun
 (
   type T = ((src -> Proof.context -> method) * string) Name_Space.table;
-  val empty = Name_Space.empty_table "method";
+  val empty : T = Name_Space.empty_table "method";
   val copy = I;
   val extend = I;
   fun merge _ tables : T = Name_Space.merge_tables tables;
--- a/src/Pure/simplifier.ML	Sun Oct 25 11:58:11 2009 +0100
+++ b/src/Pure/simplifier.ML	Sun Oct 25 13:04:06 2009 +0100
@@ -148,7 +148,7 @@
 structure Simprocs = GenericDataFun
 (
   type T = simproc Name_Space.table;
-  val empty = Name_Space.empty_table "simproc";
+  val empty : T = Name_Space.empty_table "simproc";
   val extend = I;
   fun merge _ simprocs = Name_Space.merge_tables simprocs;
 );
--- a/src/Pure/theory.ML	Sun Oct 25 11:58:11 2009 +0100
+++ b/src/Pure/theory.ML	Sun Oct 25 13:04:06 2009 +0100
@@ -89,7 +89,7 @@
 structure ThyData = TheoryDataFun
 (
   type T = thy;
-  val empty_axioms = Name_Space.empty_table "axiom";
+  val empty_axioms = Name_Space.empty_table "axiom" : term Name_Space.table;
   val empty = make_thy (empty_axioms, Defs.empty, ([], []));
   val copy = I;
 
--- a/src/Pure/thm.ML	Sun Oct 25 11:58:11 2009 +0100
+++ b/src/Pure/thm.ML	Sun Oct 25 13:04:06 2009 +0100
@@ -1727,7 +1727,7 @@
 structure Oracles = TheoryDataFun
 (
   type T = unit Name_Space.table;
-  val empty = Name_Space.empty_table "oracle";
+  val empty : T = Name_Space.empty_table "oracle";
   val copy = I;
   val extend = I;
   fun merge _ oracles : T = Name_Space.merge_tables oracles;