src/Pure/Tools/build.scala
changeset 68213 bb93511c7e8f
parent 68212 5a59fded83c7
child 68214 b0e2a19df95b
equal deleted inserted replaced
68212:5a59fded83c7 68213:bb93511c7e8f
   481     }
   481     }
   482 
   482 
   483     // scheduler loop
   483     // scheduler loop
   484     case class Result(
   484     case class Result(
   485       current: Boolean,
   485       current: Boolean,
   486       heap_stamp: Option[String],
   486       heap_digest: Option[String],
   487       process: Option[Process_Result],
   487       process: Option[Process_Result],
   488       info: Sessions.Info)
   488       info: Sessions.Info)
   489     {
   489     {
   490       def ok: Boolean =
   490       def ok: Boolean =
   491         process match {
   491         process match {
   532                   (if (tail == 0) log_lines else log_lines.drop(log_lines.length - tail max 0)))
   532                   (if (tail == 0) log_lines else log_lines.drop(log_lines.length - tail max 0)))
   533             }
   533             }
   534 
   534 
   535 
   535 
   536             // write log file
   536             // write log file
   537             val heap_stamp =
   537             val heap_digest =
   538               if (process_result.ok) {
   538               if (process_result.ok) {
   539                 val heap_stamp =
   539                 val heap_digest =
   540                   for (path <- job.output_path if path.is_file)
   540                   for (path <- job.output_path if path.is_file)
   541                     yield Sessions.write_heap_digest(path)
   541                     yield Sessions.write_heap_digest(path)
   542 
   542 
   543                 File.write_gzip(store.output_log_gz(name), terminate_lines(log_lines))
   543                 File.write_gzip(store.output_log_gz(name), terminate_lines(log_lines))
   544 
   544 
   545                 heap_stamp
   545                 heap_digest
   546               }
   546               }
   547               else {
   547               else {
   548                 File.write(store.output_log(name), terminate_lines(log_lines))
   548                 File.write(store.output_log(name), terminate_lines(log_lines))
   549 
   549 
   550                 None
   550                 None
   563               using(SQLite.open_database(store.output_database(name)))(db =>
   563               using(SQLite.open_database(store.output_database(name)))(db =>
   564                 store.write_session_info(db, name,
   564                 store.write_session_info(db, name,
   565                   build_log =
   565                   build_log =
   566                     if (process_result.timeout) build_log.error("Timeout") else build_log,
   566                     if (process_result.timeout) build_log.error("Timeout") else build_log,
   567                   build =
   567                   build =
   568                     Session_Info(sources_stamp(deps, name), input_heaps, heap_stamp,
   568                     Session_Info(sources_stamp(deps, name), input_heaps, heap_digest,
   569                       process_result.rc)))
   569                       process_result.rc)))
   570             }
   570             }
   571 
   571 
   572             // messages
   572             // messages
   573             process_result.err_lines.foreach(progress.echo(_))
   573             process_result.err_lines.foreach(progress.echo(_))
   578               progress.echo(name + " FAILED")
   578               progress.echo(name + " FAILED")
   579               if (!process_result.interrupted) progress.echo(process_result_tail.out)
   579               if (!process_result.interrupted) progress.echo(process_result_tail.out)
   580             }
   580             }
   581 
   581 
   582             loop(pending - name, running - name,
   582             loop(pending - name, running - name,
   583               results + (name -> Result(false, heap_stamp, Some(process_result_tail), job.info)))
   583               results + (name -> Result(false, heap_digest, Some(process_result_tail), job.info)))
   584             //}}}
   584             //}}}
   585           case None if running.size < (max_jobs max 1) =>
   585           case None if running.size < (max_jobs max 1) =>
   586             //{{{ check/start next job
   586             //{{{ check/start next job
   587             pending.dequeue(running.isDefinedAt(_)) match {
   587             pending.dequeue(running.isDefinedAt(_)) match {
   588               case Some((name, info)) =>
   588               case Some((name, info)) =>
   589                 val ancestor_results =
   589                 val ancestor_results =
   590                   deps.sessions_structure.build_requirements(List(name)).filterNot(_ == name).
   590                   deps.sessions_structure.build_requirements(List(name)).filterNot(_ == name).
   591                     map(results(_))
   591                     map(results(_))
   592                 val ancestor_heaps = ancestor_results.flatMap(_.heap_stamp)
   592                 val ancestor_heaps = ancestor_results.flatMap(_.heap_digest)
   593 
   593 
   594                 val do_output = build_heap || Sessions.is_pure(name) || queue.is_inner(name)
   594                 val do_output = build_heap || Sessions.is_pure(name) || queue.is_inner(name)
   595 
   595 
   596                 val (current, heap_stamp) =
   596                 val (current, heap_digest) =
   597                 {
   597                 {
   598                   store.try_open_database(name) match {
   598                   store.try_open_database(name) match {
   599                     case Some(db) =>
   599                     case Some(db) =>
   600                       try {
   600                       try {
   601                         store.read_build(db, name) match {
   601                         store.read_build(db, name) match {
   602                           case Some(build) =>
   602                           case Some(build) =>
   603                             val heap_stamp = store.find_heap_digest(name)
   603                             val heap_digest = store.find_heap_digest(name)
   604                             val current =
   604                             val current =
   605                               !fresh_build &&
   605                               !fresh_build &&
   606                               build.ok &&
   606                               build.ok &&
   607                               build.sources == sources_stamp(deps, name) &&
   607                               build.sources == sources_stamp(deps, name) &&
   608                               build.input_heaps == ancestor_heaps &&
   608                               build.input_heaps == ancestor_heaps &&
   609                               build.output_heap == heap_stamp &&
   609                               build.output_heap == heap_digest &&
   610                               !(do_output && heap_stamp.isEmpty)
   610                               !(do_output && heap_digest.isEmpty)
   611                             (current, heap_stamp)
   611                             (current, heap_digest)
   612                           case None => (false, None)
   612                           case None => (false, None)
   613                         }
   613                         }
   614                       } finally { db.close }
   614                       } finally { db.close }
   615                     case None => (false, None)
   615                     case None => (false, None)
   616                   }
   616                   }
   617                 }
   617                 }
   618                 val all_current = current && ancestor_results.forall(_.current)
   618                 val all_current = current && ancestor_results.forall(_.current)
   619 
   619 
   620                 if (all_current)
   620                 if (all_current)
   621                   loop(pending - name, running,
   621                   loop(pending - name, running,
   622                     results + (name -> Result(true, heap_stamp, Some(Process_Result(0)), info)))
   622                     results + (name -> Result(true, heap_digest, Some(Process_Result(0)), info)))
   623                 else if (no_build) {
   623                 else if (no_build) {
   624                   if (verbose) progress.echo("Skipping " + name + " ...")
   624                   if (verbose) progress.echo("Skipping " + name + " ...")
   625                   loop(pending - name, running,
   625                   loop(pending - name, running,
   626                     results + (name -> Result(false, heap_stamp, Some(Process_Result(1)), info)))
   626                     results + (name -> Result(false, heap_digest, Some(Process_Result(1)), info)))
   627                 }
   627                 }
   628                 else if (ancestor_results.forall(_.ok) && !progress.stopped) {
   628                 else if (ancestor_results.forall(_.ok) && !progress.stopped) {
   629                   progress.echo((if (do_output) "Building " else "Running ") + name + " ...")
   629                   progress.echo((if (do_output) "Building " else "Running ") + name + " ...")
   630 
   630 
   631                   cleanup(name)
   631                   cleanup(name)
   639                   loop(pending, running + (name -> (ancestor_heaps, job)), results)
   639                   loop(pending, running + (name -> (ancestor_heaps, job)), results)
   640                 }
   640                 }
   641                 else {
   641                 else {
   642                   progress.echo(name + " CANCELLED")
   642                   progress.echo(name + " CANCELLED")
   643                   loop(pending - name, running,
   643                   loop(pending - name, running,
   644                     results + (name -> Result(false, heap_stamp, None, info)))
   644                     results + (name -> Result(false, heap_digest, None, info)))
   645                 }
   645                 }
   646               case None => sleep(); loop(pending, running, results)
   646               case None => sleep(); loop(pending, running, results)
   647             }
   647             }
   648             ///}}}
   648             ///}}}
   649           case None => sleep(); loop(pending, running, results)
   649           case None => sleep(); loop(pending, running, results)