src/HOL/Tools/semiring_normalizer.ML
changeset 58826 2ed2eaabe3df
parent 54249 ce00f2fef556
child 59058 a78612c67ec0
--- a/src/HOL/Tools/semiring_normalizer.ML	Wed Oct 29 17:01:44 2014 +0100
+++ b/src/HOL/Tools/semiring_normalizer.ML	Wed Oct 29 19:01:49 2014 +0100
@@ -41,8 +41,6 @@
        main: Proof.context -> conv,
        pow: Proof.context -> conv,
        sub: Proof.context -> conv}
-
-  val setup: theory -> theory
 end
 
 structure Semiring_Normalizer: SEMIRING_NORMALIZER = 
@@ -912,20 +910,21 @@
 
 in
 
-val setup =
-  Attrib.setup @{binding normalizer}
-    (Scan.lift (Args.$$$ delN >> K del) ||
-      ((keyword2 semiringN opsN |-- terms) --
-       (keyword2 semiringN rulesN |-- thms)) --
-      (optional (keyword2 ringN opsN |-- terms) --
-       optional (keyword2 ringN rulesN |-- thms)) --
-      (optional (keyword2 fieldN opsN |-- terms) --
-       optional (keyword2 fieldN rulesN |-- thms)) --
-      optional (keyword2 idomN rulesN |-- thms) --
-      optional (keyword2 idealN rulesN |-- thms)
-      >> (fn ((((sr, r), f), id), idl) => 
-             add {semiring = sr, ring = r, field = f, idom = id, ideal = idl}))
-    "semiring normalizer data";
+val _ =
+  Theory.setup
+    (Attrib.setup @{binding normalizer}
+      (Scan.lift (Args.$$$ delN >> K del) ||
+        ((keyword2 semiringN opsN |-- terms) --
+         (keyword2 semiringN rulesN |-- thms)) --
+        (optional (keyword2 ringN opsN |-- terms) --
+         optional (keyword2 ringN rulesN |-- thms)) --
+        (optional (keyword2 fieldN opsN |-- terms) --
+         optional (keyword2 fieldN rulesN |-- thms)) --
+        optional (keyword2 idomN rulesN |-- thms) --
+        optional (keyword2 idealN rulesN |-- thms)
+        >> (fn ((((sr, r), f), id), idl) => 
+               add {semiring = sr, ring = r, field = f, idom = id, ideal = idl}))
+      "semiring normalizer data");
 
 end;