tuned signature;
authorwenzelm
Sun, 28 Feb 2016 17:40:01 +0100
changeset 62454 38c89353b349
parent 62453 b93cc7d73431
child 62455 2026ef279d1e
tuned signature;
src/Pure/System/getopts.scala
src/Pure/System/options.scala
src/Pure/Tools/build_doc.scala
src/Pure/Tools/check_sources.scala
src/Pure/Tools/doc.scala
src/Pure/Tools/update_cartouches.scala
src/Pure/Tools/update_header.scala
src/Pure/Tools/update_then.scala
src/Pure/Tools/update_theorems.scala
--- a/src/Pure/System/getopts.scala	Sun Feb 28 17:37:20 2016 +0100
+++ b/src/Pure/System/getopts.scala	Sun Feb 28 17:40:01 2016 +0100
@@ -9,7 +9,7 @@
 
 object Getopts
 {
-  def apply(usage_text: () => String, option_specs: (String, String => Unit)*): Getopts =
+  def apply(usage_text: String, option_specs: (String, String => Unit)*): Getopts =
   {
     val options =
       (Map.empty[Char, (Boolean, String => Unit)] /: option_specs) {
@@ -25,11 +25,11 @@
   }
 }
 
-class Getopts private(usage_text: () => String, options: Map[Char, (Boolean, String => Unit)])
+class Getopts private(usage_text: String, options: Map[Char, (Boolean, String => Unit)])
 {
   def usage(): Nothing =
   {
-    Console.println(usage_text())
+    Console.println(usage_text)
     sys.exit(1)
   }
 
--- a/src/Pure/System/options.scala	Sun Feb 28 17:37:20 2016 +0100
+++ b/src/Pure/System/options.scala	Sun Feb 28 17:40:01 2016 +0100
@@ -150,7 +150,7 @@
       var list_options = false
       var export_file = ""
 
-      val getopts = Getopts(() => """
+      val getopts = Getopts("""
 Usage: isabelle options [OPTIONS] [MORE_OPTIONS ...]
 
   Options are:
--- a/src/Pure/Tools/build_doc.scala	Sun Feb 28 17:37:20 2016 +0100
+++ b/src/Pure/Tools/build_doc.scala	Sun Feb 28 17:40:01 2016 +0100
@@ -77,7 +77,7 @@
       var system_mode = false
 
       val getopts =
-        Getopts(() => """
+        Getopts("""
 Usage: isabelle build_doc [OPTIONS] [DOCS ...]"
 
   Options are:
--- a/src/Pure/Tools/check_sources.scala	Sun Feb 28 17:37:20 2016 +0100
+++ b/src/Pure/Tools/check_sources.scala	Sun Feb 28 17:40:01 2016 +0100
@@ -54,7 +54,7 @@
   def main(args: Array[String])
   {
     Command_Line.tool0 {
-      val getopts = Getopts(() => """
+      val getopts = Getopts("""
 Usage: isabelle check_sources [ROOT_DIRS...]
 
   Check .thy, .ML, ROOT files from manifest of Mercurial ROOT_DIRS.
--- a/src/Pure/Tools/doc.scala	Sun Feb 28 17:37:20 2016 +0100
+++ b/src/Pure/Tools/doc.scala	Sun Feb 28 17:40:01 2016 +0100
@@ -96,7 +96,7 @@
   def main(args: Array[String])
   {
     Command_Line.tool0 {
-      val getopts = Getopts(() => """
+      val getopts = Getopts("""
 Usage: isabelle doc [DOC ...]
 
   View Isabelle documentation.
--- a/src/Pure/Tools/update_cartouches.scala	Sun Feb 28 17:37:20 2016 +0100
+++ b/src/Pure/Tools/update_cartouches.scala	Sun Feb 28 17:40:01 2016 +0100
@@ -90,7 +90,7 @@
       var replace_comment = false
       var replace_text = false
 
-      val getopts = Getopts(() => """
+      val getopts = Getopts("""
 Usage: isabelle update_cartouches [FILES|DIRS...]
 
   Options are:
--- a/src/Pure/Tools/update_header.scala	Sun Feb 28 17:37:20 2016 +0100
+++ b/src/Pure/Tools/update_header.scala	Sun Feb 28 17:40:01 2016 +0100
@@ -33,7 +33,7 @@
     Command_Line.tool0 {
       var section = "section"
 
-      val getopts = Getopts(() => """
+      val getopts = Getopts("""
 Usage: isabelle update_header [FILES|DIRS...]
 
   Options are:
--- a/src/Pure/Tools/update_then.scala	Sun Feb 28 17:37:20 2016 +0100
+++ b/src/Pure/Tools/update_then.scala	Sun Feb 28 17:40:01 2016 +0100
@@ -33,7 +33,7 @@
   def main(args: Array[String])
   {
     Command_Line.tool0 {
-      val getopts = Getopts(() => """
+      val getopts = Getopts("""
 Usage: isabelle update_then [FILES|DIRS...]
 
   Recursively find .thy files and expand old Isar command conflations:
--- a/src/Pure/Tools/update_theorems.scala	Sun Feb 28 17:37:20 2016 +0100
+++ b/src/Pure/Tools/update_theorems.scala	Sun Feb 28 17:40:01 2016 +0100
@@ -34,7 +34,7 @@
   def main(args: Array[String])
   {
     Command_Line.tool0 {
-      val getopts = Getopts(() => """
+      val getopts = Getopts("""
 Usage: isabelle update_theorems [FILES|DIRS...]
 
   Recursively find .thy files and update toplevel theorem keywords: