wenzelm@47825: theory Scala wenzelm@47825: imports Base wenzelm@47825: begin wenzelm@47825: wenzelm@58618: chapter \Isabelle/Scala development tools\ wenzelm@47825: wenzelm@58618: text \Isabelle/ML and Isabelle/Scala are the two main language wenzelm@47825: environments for Isabelle tool implementations. There are some basic wenzelm@47825: command-line tools to work with the underlying Java Virtual Machine, wenzelm@47825: the Scala toplevel and compiler. Note that Isabelle/jEdit wenzelm@58553: @{cite "isabelle-jedit"} provides a Scala Console for interactive wenzelm@58618: experimentation within the running application.\ wenzelm@47825: wenzelm@47825: wenzelm@58618: section \Java Runtime Environment within Isabelle \label{sec:tool-java}\ wenzelm@47825: wenzelm@58618: text \The @{tool_def java} tool is a direct wrapper for the Java wenzelm@48602: Runtime Environment, within the regular Isabelle settings wenzelm@47825: environment (\secref{sec:settings}). The command line arguments are wenzelm@47825: that of the underlying Java version. It is run in @{verbatim wenzelm@47825: "-server"} mode if possible, to improve performance (at the cost of wenzelm@47825: extra startup time). wenzelm@47825: wenzelm@47825: The @{verbatim java} executable is the one within @{setting wenzelm@47825: ISABELLE_JDK_HOME}, according to the standard directory layout for wenzelm@47825: official JDK distributions. The class loader is augmented such that wenzelm@47825: the name space of @{verbatim "Isabelle/Pure.jar"} is available, wenzelm@47825: which is the main Isabelle/Scala module. wenzelm@47825: wenzelm@47825: For example, the following command-line invokes the main method of wenzelm@47825: class @{verbatim isabelle.GUI_Setup}, which opens a windows with wenzelm@47825: some diagnostic information about the Isabelle environment: wenzelm@47825: \begin{alltt} wenzelm@47825: isabelle java isabelle.GUI_Setup wenzelm@47825: \end{alltt} wenzelm@58618: \ wenzelm@47825: wenzelm@47825: wenzelm@58618: section \Scala toplevel \label{sec:tool-scala}\ wenzelm@47825: wenzelm@58618: text \The @{tool_def scala} tool is a direct wrapper for the Scala wenzelm@48602: toplevel; see also @{tool java} above. The command line arguments wenzelm@48602: are that of the underlying Scala version. wenzelm@47825: wenzelm@47825: This allows to interact with Isabelle/Scala in TTY mode like this: wenzelm@47825: \begin{alltt} wenzelm@47825: isabelle scala wenzelm@47825: scala> isabelle.Isabelle_System.getenv("ISABELLE_HOME") wenzelm@48815: scala> val options = isabelle.Options.init() wenzelm@48815: scala> options.bool("browser_info") wenzelm@52116: scala> options.string("document") wenzelm@47825: \end{alltt} wenzelm@58618: \ wenzelm@47825: wenzelm@47825: wenzelm@58618: section \Scala compiler \label{sec:tool-scalac}\ wenzelm@47825: wenzelm@58618: text \The @{tool_def scalac} tool is a direct wrapper for the Scala wenzelm@48602: compiler; see also @{tool scala} above. The command line arguments wenzelm@48602: are that of the underlying Scala version. wenzelm@47825: wenzelm@47825: This allows to compile further Scala modules, depending on existing wenzelm@47825: Isabelle/Scala functionality. The resulting class or jar files can wenzelm@57855: be added to the Java classpath using the @{verbatim classpath} Bash wenzelm@53576: function that is provided by the Isabelle process environment. Thus wenzelm@53576: add-on components can register themselves in a modular manner, see wenzelm@53576: also \secref{sec:components}. wenzelm@47825: wenzelm@58553: Note that jEdit @{cite "isabelle-jedit"} has its own mechanisms for wenzelm@47825: adding plugin components, which needs special attention since wenzelm@58618: it overrides the standard Java class loader.\ wenzelm@47825: wenzelm@52116: wenzelm@58618: section \Scala script wrapper\ wenzelm@52116: wenzelm@58618: text \The executable @{executable wenzelm@52116: "$ISABELLE_HOME/bin/isabelle_scala_script"} allows to run wenzelm@52116: Isabelle/Scala source files stand-alone programs, by using a wenzelm@52116: suitable ``hash-bang'' line and executable file permissions. wenzelm@52116: wenzelm@52116: The subsequent example assumes that the main Isabelle binaries have wenzelm@52116: been installed in some directory that is included in @{setting PATH} wenzelm@52116: (see also @{tool "install"}): wenzelm@52116: wenzelm@52116: \begin{alltt} wenzelm@52116: #!/usr/bin/env isabelle_scala_script wenzelm@52116: wenzelm@52116: val options = isabelle.Options.init() wenzelm@52116: Console.println("browser_info = " + options.bool("browser_info")) wenzelm@52116: Console.println("document = " + options.string("document")) wenzelm@52116: \end{alltt} wenzelm@52116: wenzelm@53982: Alternatively the full @{file wenzelm@52116: "$ISABELLE_HOME/bin/isabelle_scala_script"} may be specified in wenzelm@58618: expanded form.\ wenzelm@52116: wenzelm@47825: end