--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/lib/Tools/build Tue Jul 17 15:56:19 2012 +0200
@@ -0,0 +1,78 @@
+#!/usr/bin/env bash
+#
+# Author: Makarius
+#
+# DESCRIPTION: build and manage Isabelle sessions
+
+
+## diagnostics
+
+PRG="$(basename "$0")"
+
+function usage()
+{
+ echo
+ echo "Usage: isabelle $PRG [OPTIONS] [SESSIONS ...]"
+ echo
+ echo " Options are:"
+ echo " -a all sessions"
+ echo " -b build target images"
+ echo " -l list sessions only"
+ echo " -o OPTION override session configuration OPTION (via NAME=VAL or NAME)"
+ echo
+ echo " Build and manage Isabelle sessions, depending on implicit"
+ echo " ISABELLE_BUILD_OPTIONS=\"$ISABELLE_BUILD_OPTIONS\""
+ echo
+ echo " ML_PLATFORM=\"$ML_PLATFORM\""
+ echo " ML_HOME=\"$ML_HOME\""
+ echo " ML_SYSTEM=\"$ML_SYSTEM\""
+ echo " ML_OPTIONS=\"$ML_OPTIONS\""
+ echo
+ exit 1
+}
+
+function fail()
+{
+ echo "$1" >&2
+ exit 2
+}
+
+
+## process command line
+
+ALL_SESSIONS=false
+BUILD_IMAGES=false
+LIST_ONLY=false
+
+eval "declare -a BUILD_OPTIONS=($ISABELLE_BUILD_OPTIONS)"
+
+while getopts "ablo:" OPT
+do
+ case "$OPT" in
+ a)
+ ALL_SESSIONS="true"
+ ;;
+ b)
+ BUILD_IMAGES="true"
+ ;;
+ l)
+ LIST_ONLY="true"
+ ;;
+ o)
+ BUILD_OPTIONS["${#BUILD_OPTIONS[@]}"]="$OPTARG"
+ ;;
+ \?)
+ usage
+ ;;
+ esac
+done
+
+shift $(($OPTIND - 1))
+
+
+## main
+
+[ -e "$ISABELLE_HOME/Admin/build" ] && { "$ISABELLE_HOME/Admin/build" jars || exit $?; }
+
+exec "$ISABELLE_TOOL" java isabelle.Build \
+ "$ALL_SESSIONS" "$BUILD_IMAGES" "$LIST_ONLY" "${BUILD_OPTIONS[@]}" $'\n' "$@"
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/System/build.scala Tue Jul 17 15:56:19 2012 +0200
@@ -0,0 +1,60 @@
+/* Title: Pure/System/build.scala
+ Author: Makarius
+
+Build and manage Isabelle sessions.
+*/
+
+package isabelle
+
+
+object Build
+{
+ /* command line entry point */
+
+ private object Bool
+ {
+ def unapply(s: String): Option[Boolean] =
+ s match {
+ case "true" => Some(true)
+ case "false" => Some(false)
+ case _ => None
+ }
+ }
+
+ def main(args: Array[String])
+ {
+ def bad_args()
+ {
+ java.lang.System.err.println("Bad arguments: " + args.toString)
+ sys.exit(2)
+ }
+
+ args.toList match {
+ case Bool(all_sessions) :: Bool(build_images) :: Bool(list_only) :: rest =>
+ rest.indexWhere(_ == "\n") match {
+ case -1 => bad_args()
+ case i =>
+ val (options, rest1) = rest.splitAt(i)
+ val sessions = rest1.tail
+ val rc = build(all_sessions, build_images, list_only, options, sessions)
+ sys.exit(rc)
+ }
+ case _ => bad_args()
+ }
+ }
+
+
+ /* build */
+
+ def build(all_sessions: Boolean, build_images: Boolean, list_only: Boolean,
+ options: List[String], sessions: List[String]): Int =
+ {
+ val rc = 1
+
+ println("options = " + options.toString)
+ println("sessions = " + sessions.toString)
+
+ rc
+ }
+}
+
--- a/src/Pure/build-jars Tue Jul 17 14:33:23 2012 +0200
+++ b/src/Pure/build-jars Tue Jul 17 15:56:19 2012 +0200
@@ -39,6 +39,7 @@
PIDE/text.scala
PIDE/xml.scala
PIDE/yxml.scala
+ System/build.scala
System/event_bus.scala
System/gui_setup.scala
System/invoke_scala.scala