src/Pure/Tools/build_history.scala
changeset 64049 ac3ed62c53c3
parent 64048 b0c52944978e
child 64050 68fcd61b87ae
equal deleted inserted replaced
64048:b0c52944978e 64049:ac3ed62c53c3
    11 import java.util.Calendar
    11 import java.util.Calendar
    12 
    12 
    13 
    13 
    14 object Build_History
    14 object Build_History
    15 {
    15 {
       
    16   /* other Isabelle environment */
       
    17 
       
    18   private class Other_Isabelle(progress: Progress, isabelle_home: Path, isabelle_identifier: String)
       
    19   {
       
    20     other_isabelle =>
       
    21 
       
    22     def bash(script: String, redirect: Boolean = false, echo: Boolean = false): Process_Result =
       
    23       Isabelle_System.bash(
       
    24         "export ISABELLE_IDENTIFIER=" + File.bash_string(isabelle_identifier) + "\n" + script,
       
    25         env = null, cwd = isabelle_home.file, redirect = redirect,
       
    26         progress_stdout = progress.echo_if(echo, _),
       
    27         progress_stderr = progress.echo_if(echo, _))
       
    28 
       
    29     def apply(cmdline: String, redirect: Boolean = false, echo: Boolean = false): Process_Result =
       
    30       bash("bin/isabelle " + cmdline, redirect, echo)
       
    31 
       
    32     def resolve_components(echo: Boolean): Unit =
       
    33       other_isabelle("components -a", redirect = true, echo = echo).check
       
    34 
       
    35     val isabelle_home_user =
       
    36       Path.explode(other_isabelle("getenv -b ISABELLE_HOME_USER").check.out)
       
    37   }
       
    38 
       
    39 
       
    40   /* settings environment */
       
    41 
       
    42   private def init_settings(
       
    43     isabelle_home_user: Path, components_base: String, nonfree: Boolean): Path =
       
    44   {
       
    45     val etc_settings = isabelle_home_user + Path.explode("etc/settings")
       
    46 
       
    47     if (etc_settings.is_file && !File.read(etc_settings).startsWith("# generated by Isabelle"))
       
    48       error("Cannot proceed with existing user settings file: " + etc_settings)
       
    49 
       
    50     Isabelle_System.mkdirs(etc_settings.dir)
       
    51     File.write(etc_settings,
       
    52       "# generated by Isabelle " + Calendar.getInstance.getTime + "\n" +
       
    53       "#-*- shell-script -*- :mode=shellscript:\n")
       
    54 
       
    55     val component_settings =
       
    56     {
       
    57       val components_base_path =
       
    58         if (components_base == "") isabelle_home_user.dir + Path.explode("contrib")
       
    59         else Path.explode(components_base).expand
       
    60 
       
    61       val catalogs =
       
    62         if (nonfree) List("main", "optional", "nonfree") else List("main", "optional")
       
    63 
       
    64       catalogs.map(catalog =>
       
    65         "init_components " + File.bash_path(components_base_path) +
       
    66           " \"$ISABELLE_HOME/Admin/components/" + catalog + "\"")
       
    67     }
       
    68     File.append(etc_settings, "\n" + Library.terminate_lines(component_settings))
       
    69 
       
    70     etc_settings
       
    71   }
       
    72 
       
    73 
    16   /* build_history */
    74   /* build_history */
    17 
    75 
    18   private val default_rev = "tip"
    76   private val default_rev = "tip"
    19   private val default_threads = 1
    77   private val default_threads = 1
    20   private val default_heap = 1000
    78   private val default_heap = 1000
    34     max_heap: Option[Int] = None,
    92     max_heap: Option[Int] = None,
    35     more_settings: List[String] = Nil,
    93     more_settings: List[String] = Nil,
    36     verbose: Boolean = false,
    94     verbose: Boolean = false,
    37     build_args: List[String] = Nil): Process_Result =
    95     build_args: List[String] = Nil): Process_Result =
    38   {
    96   {
    39     /* output */
       
    40 
       
    41     def output(msg: String) { progress.echo(msg) }
       
    42     def output_if(b: Boolean, msg: String) { if (b) output(msg) }
       
    43 
       
    44 
       
    45     /* sanity checks */
    97     /* sanity checks */
    46 
    98 
    47     if (File.eq(Path.explode("~~").file, hg.root.file))
    99     if (File.eq(Path.explode("~~").file, hg.root.file))
    48       error("Repository coincides with ISABELLE_HOME=" + Path.explode("~~").expand)
   100       error("Repository coincides with ISABELLE_HOME=" + Path.explode("~~").expand)
    49 
   101 
    58       case null | "" =>
   110       case null | "" =>
    59       case _ => error("Cannot run build_history within existing Isabelle settings environment")
   111       case _ => error("Cannot run build_history within existing Isabelle settings environment")
    60     }
   112     }
    61 
   113 
    62 
   114 
    63     /* purge repository */
   115     /* init repository */
    64 
   116 
    65     hg.update(rev = rev, clean = true)
   117     hg.update(rev = rev, clean = true)
    66     output_if(verbose, hg.log(rev, options = "-l1"))
   118     progress.echo_if(verbose, hg.log(rev, options = "-l1"))
    67 
   119 
    68 
   120 
    69     /* invoke isabelle tools */
   121     /* init settings */
    70 
   122 
    71     def bash(script: String, redirect: Boolean = false, echo: Boolean = false): Process_Result =
   123     val other_isabelle = new Other_Isabelle(progress, hg.root, isabelle_identifier)
    72       Isabelle_System.bash(
   124 
    73         "export ISABELLE_IDENTIFIER=" + File.bash_string(isabelle_identifier) + "\n" + script,
   125     val etc_settings = init_settings(other_isabelle.isabelle_home_user, components_base, nonfree)
    74         env = null, cwd = hg.root.file, redirect = redirect,
   126     other_isabelle.resolve_components(verbose)
    75         progress_stdout = output_if(echo, _), progress_stderr = output_if(echo, _))
       
    76 
       
    77     def isabelle(cmdline: String, redirect: Boolean = false, echo: Boolean = false): Process_Result =
       
    78       bash("bin/isabelle " + cmdline, redirect, echo)
       
    79 
       
    80     val isabelle_home_user = Path.explode(isabelle("getenv -b ISABELLE_HOME_USER").check.out)
       
    81     val isabelle_output = Path.explode(isabelle("getenv -b ISABELLE_OUTPUT").check.out)
       
    82 
       
    83 
       
    84     /* reset settings */
       
    85 
       
    86     val etc_settings = isabelle_home_user + Path.explode("etc/settings")
       
    87 
       
    88     if (etc_settings.is_file && !File.read(etc_settings).startsWith("# generated by Isabelle"))
       
    89       error("Cannot proceed with existing user settings file: " + etc_settings)
       
    90 
       
    91     Isabelle_System.mkdirs(etc_settings.dir)
       
    92 
       
    93     File.write(etc_settings,
       
    94       "# generated by Isabelle " + Calendar.getInstance.getTime + "\n" +
       
    95       "#-*- shell-script -*- :mode=shellscript:\n")
       
    96 
       
    97 
       
    98     /* initial settings */
       
    99 
       
   100     val component_settings =
       
   101     {
       
   102       val components_base_path =
       
   103         if (components_base == "") isabelle_home_user.dir + Path.explode("contrib")
       
   104         else Path.explode(components_base).expand
       
   105 
       
   106       val catalogs =
       
   107         if (nonfree) List("main", "optional", "nonfree") else List("main", "optional")
       
   108 
       
   109       catalogs.map(catalog =>
       
   110         "init_components " + File.bash_path(components_base_path) +
       
   111           " \"$ISABELLE_HOME/Admin/components/" + catalog + "\"")
       
   112     }
       
   113 
       
   114     File.append(etc_settings, "\n" + Library.terminate_lines(component_settings))
       
   115 
       
   116 
       
   117     isabelle("components -a", redirect = true, echo = verbose).check
       
   118 
   127 
   119 
   128 
   120     /* augmented settings */
   129     /* augmented settings */
   121 
   130 
   122     val ml_settings =
   131     val ml_settings =
   123     {
   132     {
   124       val windows_32 = "x86-windows"
   133       val windows_32 = "x86-windows"
   125       val windows_64 = "x86_64-windows"
   134       val windows_64 = "x86_64-windows"
   126       val platform_32 = isabelle("getenv -b ISABELLE_PLATFORM32").check.out
   135       val platform_32 = other_isabelle("getenv -b ISABELLE_PLATFORM32").check.out
   127       val platform_64 = isabelle("getenv -b ISABELLE_PLATFORM64").check.out
   136       val platform_64 = other_isabelle("getenv -b ISABELLE_PLATFORM64").check.out
   128       val platform_family = isabelle("getenv -b ISABELLE_PLATFORM_FAMILY").check.out
   137       val platform_family = other_isabelle("getenv -b ISABELLE_PLATFORM_FAMILY").check.out
   129 
   138 
   130       val polyml_home =
   139       val polyml_home =
   131         try { Path.explode(isabelle("getenv -b ML_HOME").check.out).dir }
   140         try { Path.explode(other_isabelle("getenv -b ML_HOME").check.out).dir }
   132         catch { case ERROR(msg) => error("Bad ML_HOME: " + msg) }
   141         catch { case ERROR(msg) => error("Bad ML_HOME: " + msg) }
   133 
   142 
   134       def ml_home(platform: String): Path = polyml_home + Path.explode(platform)
   143       def ml_home(platform: String): Path = polyml_home + Path.explode(platform)
   135 
   144 
   136       def err(platform: String): Nothing =
   145       def err(platform: String): Nothing =
   175 
   184 
   176     if (more_settings.nonEmpty)
   185     if (more_settings.nonEmpty)
   177       File.append(etc_settings, "\n" + Library.terminate_lines(more_settings))
   186       File.append(etc_settings, "\n" + Library.terminate_lines(more_settings))
   178 
   187 
   179 
   188 
   180     isabelle("components -a", redirect = true, echo = verbose).check
   189     other_isabelle.resolve_components(verbose)
   181 
   190 
   182 
   191 
   183     /* build */
   192     /* build */
   184 
   193 
   185     bash("env PATH=\"" + File.bash_path(Path.explode("~~/lib/dummy_stty").expand) + ":$PATH\" " +
   194     other_isabelle.bash(
       
   195       "env PATH=\"" + File.bash_path(Path.explode("~~/lib/dummy_stty").expand) + ":$PATH\" " +
   186       "bin/isabelle jedit -b" + (if (fresh) " -f" else ""), redirect = true, echo = verbose).check
   196       "bin/isabelle jedit -b" + (if (fresh) " -f" else ""), redirect = true, echo = verbose).check
   187 
   197 
       
   198     val isabelle_output = Path.explode(other_isabelle("getenv -b ISABELLE_OUTPUT").check.out)
   188     Isabelle_System.rm_tree(isabelle_output.file)
   199     Isabelle_System.rm_tree(isabelle_output.file)
   189     isabelle("build " + File.bash_args(build_args), redirect = true, echo = true)
   200     other_isabelle("build " + File.bash_args(build_args), redirect = true, echo = true)
   190   }
   201   }
   191 
   202 
   192 
   203 
   193   /* command line entry point */
   204   /* command line entry point */
   194 
   205