src/Doc/System/Scala.thy
author wenzelm
Tue, 26 May 2020 12:09:36 +0200
changeset 71897 2cf0b0293f0d
parent 71892 dff81ce866d4
child 71907 64c9628b39fc
permissions -rw-r--r--
proper check of example;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
61656
cfabbc083977 more uniform jEdit properties;
wenzelm
parents: 61575
diff changeset
     1
(*:maxLineLen=78:*)
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
     2
47825
4f25960417ae some coverage of Isabelle/Scala tools;
wenzelm
parents:
diff changeset
     3
theory Scala
4f25960417ae some coverage of Isabelle/Scala tools;
wenzelm
parents:
diff changeset
     4
imports Base
4f25960417ae some coverage of Isabelle/Scala tools;
wenzelm
parents:
diff changeset
     5
begin
4f25960417ae some coverage of Isabelle/Scala tools;
wenzelm
parents:
diff changeset
     6
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
     7
chapter \<open>Isabelle/Scala development tools\<close>
47825
4f25960417ae some coverage of Isabelle/Scala tools;
wenzelm
parents:
diff changeset
     8
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
     9
text \<open>
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    10
  Isabelle/ML and Isabelle/Scala are the two main language environments for
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    11
  Isabelle tool implementations. There are some basic command-line tools to
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    12
  work with the underlying Java Virtual Machine, the Scala toplevel and
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    13
  compiler. Note that Isabelle/jEdit @{cite "isabelle-jedit"} provides a Scala
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    14
  Console for interactive experimentation within the running application.
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    15
\<close>
47825
4f25960417ae some coverage of Isabelle/Scala tools;
wenzelm
parents:
diff changeset
    16
4f25960417ae some coverage of Isabelle/Scala tools;
wenzelm
parents:
diff changeset
    17
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
    18
section \<open>Java Runtime Environment within Isabelle \label{sec:tool-java}\<close>
47825
4f25960417ae some coverage of Isabelle/Scala tools;
wenzelm
parents:
diff changeset
    19
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    20
text \<open>
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    21
  The @{tool_def java} tool is a direct wrapper for the Java Runtime
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    22
  Environment, within the regular Isabelle settings environment
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    23
  (\secref{sec:settings}). The command line arguments are that of the
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    24
  underlying Java version. It is run in \<^verbatim>\<open>-server\<close> mode if possible, to
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    25
  improve performance (at the cost of extra startup time).
47825
4f25960417ae some coverage of Isabelle/Scala tools;
wenzelm
parents:
diff changeset
    26
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    27
  The \<^verbatim>\<open>java\<close> executable is the one within @{setting ISABELLE_JDK_HOME},
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    28
  according to the standard directory layout for official JDK distributions.
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    29
  The class loader is augmented such that the name space of
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    30
  \<^verbatim>\<open>Isabelle/Pure.jar\<close> is available, which is the main Isabelle/Scala module.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
    31
\<close>
47825
4f25960417ae some coverage of Isabelle/Scala tools;
wenzelm
parents:
diff changeset
    32
4f25960417ae some coverage of Isabelle/Scala tools;
wenzelm
parents:
diff changeset
    33
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
    34
section \<open>Scala toplevel \label{sec:tool-scala}\<close>
47825
4f25960417ae some coverage of Isabelle/Scala tools;
wenzelm
parents:
diff changeset
    35
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    36
text \<open>
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    37
  The @{tool_def scala} tool is a direct wrapper for the Scala toplevel; see
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    38
  also @{tool java} above. The command line arguments are that of the
71897
2cf0b0293f0d proper check of example;
wenzelm
parents: 71892
diff changeset
    39
  underlying Scala version. This allows to interact with Isabelle/Scala in TTY
2cf0b0293f0d proper check of example;
wenzelm
parents: 71892
diff changeset
    40
  mode. An alternative is to use the \<^verbatim>\<open>Console/Scala\<close> plugin of Isabelle/jEdit
2cf0b0293f0d proper check of example;
wenzelm
parents: 71892
diff changeset
    41
  @{cite "isabelle-jedit"}.
2cf0b0293f0d proper check of example;
wenzelm
parents: 71892
diff changeset
    42
\<close>
2cf0b0293f0d proper check of example;
wenzelm
parents: 71892
diff changeset
    43
2cf0b0293f0d proper check of example;
wenzelm
parents: 71892
diff changeset
    44
subsubsection \<open>Example\<close>
47825
4f25960417ae some coverage of Isabelle/Scala tools;
wenzelm
parents:
diff changeset
    45
71897
2cf0b0293f0d proper check of example;
wenzelm
parents: 71892
diff changeset
    46
text \<open>
2cf0b0293f0d proper check of example;
wenzelm
parents: 71892
diff changeset
    47
  Explore the Isabelle system environment in Scala:
2cf0b0293f0d proper check of example;
wenzelm
parents: 71892
diff changeset
    48
  @{scala [display]
2cf0b0293f0d proper check of example;
wenzelm
parents: 71892
diff changeset
    49
\<open>import isabelle._
2cf0b0293f0d proper check of example;
wenzelm
parents: 71892
diff changeset
    50
2cf0b0293f0d proper check of example;
wenzelm
parents: 71892
diff changeset
    51
val isabelle_home = Isabelle_System.getenv("ISABELLE_HOME")
2cf0b0293f0d proper check of example;
wenzelm
parents: 71892
diff changeset
    52
2cf0b0293f0d proper check of example;
wenzelm
parents: 71892
diff changeset
    53
val options = Options.init()
2cf0b0293f0d proper check of example;
wenzelm
parents: 71892
diff changeset
    54
options.bool("browser_info")
2cf0b0293f0d proper check of example;
wenzelm
parents: 71892
diff changeset
    55
options.string("document")\<close>}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
    56
\<close>
47825
4f25960417ae some coverage of Isabelle/Scala tools;
wenzelm
parents:
diff changeset
    57
4f25960417ae some coverage of Isabelle/Scala tools;
wenzelm
parents:
diff changeset
    58
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
    59
section \<open>Scala compiler \label{sec:tool-scalac}\<close>
47825
4f25960417ae some coverage of Isabelle/Scala tools;
wenzelm
parents:
diff changeset
    60
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    61
text \<open>
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    62
  The @{tool_def scalac} tool is a direct wrapper for the Scala compiler; see
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    63
  also @{tool scala} above. The command line arguments are that of the
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    64
  underlying Scala version.
47825
4f25960417ae some coverage of Isabelle/Scala tools;
wenzelm
parents:
diff changeset
    65
4f25960417ae some coverage of Isabelle/Scala tools;
wenzelm
parents:
diff changeset
    66
  This allows to compile further Scala modules, depending on existing
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    67
  Isabelle/Scala functionality. The resulting class or jar files can be added
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    68
  to the Java classpath using the \<^verbatim>\<open>classpath\<close> Bash function that is provided
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    69
  by the Isabelle process environment. Thus add-on components can register
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    70
  themselves in a modular manner, see also \secref{sec:components}.
47825
4f25960417ae some coverage of Isabelle/Scala tools;
wenzelm
parents:
diff changeset
    71
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    72
  Note that jEdit @{cite "isabelle-jedit"} has its own mechanisms for adding
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    73
  plugin components, which needs special attention since it overrides the
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    74
  standard Java class loader.
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    75
\<close>
47825
4f25960417ae some coverage of Isabelle/Scala tools;
wenzelm
parents:
diff changeset
    76
52116
abf9fcfa65cf added isabelle_scala_script wrapper -- NB: portable hash-bang allows exactly one executable, without additional arguments;
wenzelm
parents: 48985
diff changeset
    77
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
    78
section \<open>Scala script wrapper\<close>
52116
abf9fcfa65cf added isabelle_scala_script wrapper -- NB: portable hash-bang allows exactly one executable, without additional arguments;
wenzelm
parents: 48985
diff changeset
    79
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    80
text \<open>
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    81
  The executable @{executable "$ISABELLE_HOME/bin/isabelle_scala_script"}
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    82
  allows to run Isabelle/Scala source files stand-alone programs, by using a
62415
62c03eb38e49 isabelle_scala_script is usually found by PATH;
wenzelm
parents: 61656
diff changeset
    83
  suitable ``hash-bang'' line and executable file permissions. For example:
61407
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 58618
diff changeset
    84
  @{verbatim [display]
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 58618
diff changeset
    85
\<open>#!/usr/bin/env isabelle_scala_script
52116
abf9fcfa65cf added isabelle_scala_script wrapper -- NB: portable hash-bang allows exactly one executable, without additional arguments;
wenzelm
parents: 48985
diff changeset
    86
abf9fcfa65cf added isabelle_scala_script wrapper -- NB: portable hash-bang allows exactly one executable, without additional arguments;
wenzelm
parents: 48985
diff changeset
    87
val options = isabelle.Options.init()
abf9fcfa65cf added isabelle_scala_script wrapper -- NB: portable hash-bang allows exactly one executable, without additional arguments;
wenzelm
parents: 48985
diff changeset
    88
Console.println("browser_info = " + options.bool("browser_info"))
61407
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 58618
diff changeset
    89
Console.println("document = " + options.string("document"))\<close>}
52116
abf9fcfa65cf added isabelle_scala_script wrapper -- NB: portable hash-bang allows exactly one executable, without additional arguments;
wenzelm
parents: 48985
diff changeset
    90
62415
62c03eb38e49 isabelle_scala_script is usually found by PATH;
wenzelm
parents: 61656
diff changeset
    91
  This assumes that the executable may be found via the @{setting PATH} from
62c03eb38e49 isabelle_scala_script is usually found by PATH;
wenzelm
parents: 61656
diff changeset
    92
  the process environment: this is the case when Isabelle settings are active,
62c03eb38e49 isabelle_scala_script is usually found by PATH;
wenzelm
parents: 61656
diff changeset
    93
  e.g.\ in the context of the main Isabelle tool wrapper
63680
6e1e8b5abbfa more symbols;
wenzelm
parents: 62415
diff changeset
    94
  \secref{sec:isabelle-tool}. Alternatively, the full
6e1e8b5abbfa more symbols;
wenzelm
parents: 62415
diff changeset
    95
  \<^file>\<open>$ISABELLE_HOME/bin/isabelle_scala_script\<close> may be specified in expanded
62415
62c03eb38e49 isabelle_scala_script is usually found by PATH;
wenzelm
parents: 61656
diff changeset
    96
  form.
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61503
diff changeset
    97
\<close>
52116
abf9fcfa65cf added isabelle_scala_script wrapper -- NB: portable hash-bang allows exactly one executable, without additional arguments;
wenzelm
parents: 48985
diff changeset
    98
71378
820cf124dced added "isabelle scala_project" to support e.g. IntelliJ IDEA;
wenzelm
parents: 63680
diff changeset
    99
820cf124dced added "isabelle scala_project" to support e.g. IntelliJ IDEA;
wenzelm
parents: 63680
diff changeset
   100
section \<open>Project setup for common Scala IDEs\<close>
820cf124dced added "isabelle scala_project" to support e.g. IntelliJ IDEA;
wenzelm
parents: 63680
diff changeset
   101
820cf124dced added "isabelle scala_project" to support e.g. IntelliJ IDEA;
wenzelm
parents: 63680
diff changeset
   102
text \<open>
71389
wenzelm
parents: 71378
diff changeset
   103
  The @{tool_def scala_project} tool creates a project configuration for
wenzelm
parents: 71378
diff changeset
   104
  Isabelle/Scala/jEdit:
71378
820cf124dced added "isabelle scala_project" to support e.g. IntelliJ IDEA;
wenzelm
parents: 63680
diff changeset
   105
  @{verbatim [display]
820cf124dced added "isabelle scala_project" to support e.g. IntelliJ IDEA;
wenzelm
parents: 63680
diff changeset
   106
\<open>Usage: isabelle scala_project [OPTIONS] PROJECT_DIR
820cf124dced added "isabelle scala_project" to support e.g. IntelliJ IDEA;
wenzelm
parents: 63680
diff changeset
   107
820cf124dced added "isabelle scala_project" to support e.g. IntelliJ IDEA;
wenzelm
parents: 63680
diff changeset
   108
  Options are:
71503
df7494f14388 clarified command line;
wenzelm
parents: 71389
diff changeset
   109
    -L           make symlinks to original scala files
71378
820cf124dced added "isabelle scala_project" to support e.g. IntelliJ IDEA;
wenzelm
parents: 63680
diff changeset
   110
820cf124dced added "isabelle scala_project" to support e.g. IntelliJ IDEA;
wenzelm
parents: 63680
diff changeset
   111
  Setup Gradle project for Isabelle/Scala/jEdit --- to support Scala IDEs
820cf124dced added "isabelle scala_project" to support e.g. IntelliJ IDEA;
wenzelm
parents: 63680
diff changeset
   112
  such as IntelliJ IDEA.\<close>}
820cf124dced added "isabelle scala_project" to support e.g. IntelliJ IDEA;
wenzelm
parents: 63680
diff changeset
   113
71389
wenzelm
parents: 71378
diff changeset
   114
  The generated configuration is for Gradle\<^footnote>\<open>\<^url>\<open>https://gradle.org\<close>\<close>, but the
wenzelm
parents: 71378
diff changeset
   115
  main purpose is to import it into common Scala IDEs, such as IntelliJ
wenzelm
parents: 71378
diff changeset
   116
  IDEA\<^footnote>\<open>\<^url>\<open>https://www.jetbrains.com/idea\<close>\<close>. This allows to explore the
wenzelm
parents: 71378
diff changeset
   117
  sources with static analysis and other hints in real-time.
71378
820cf124dced added "isabelle scala_project" to support e.g. IntelliJ IDEA;
wenzelm
parents: 63680
diff changeset
   118
820cf124dced added "isabelle scala_project" to support e.g. IntelliJ IDEA;
wenzelm
parents: 63680
diff changeset
   119
  The specified project directory must not exist yet. The generated files
71389
wenzelm
parents: 71378
diff changeset
   120
  refer to physical file-system locations, using the path notation of the
wenzelm
parents: 71378
diff changeset
   121
  underlying OS platform. Thus the project needs to be recreated whenever the
wenzelm
parents: 71378
diff changeset
   122
  Isabelle installation is changed or moved.
71378
820cf124dced added "isabelle scala_project" to support e.g. IntelliJ IDEA;
wenzelm
parents: 63680
diff changeset
   123
820cf124dced added "isabelle scala_project" to support e.g. IntelliJ IDEA;
wenzelm
parents: 63680
diff changeset
   124
  \<^medskip> By default, Scala sources are \<^emph>\<open>copied\<close> from the Isabelle distribution and
820cf124dced added "isabelle scala_project" to support e.g. IntelliJ IDEA;
wenzelm
parents: 63680
diff changeset
   125
  editing them within the IDE has no permanent effect.
820cf124dced added "isabelle scala_project" to support e.g. IntelliJ IDEA;
wenzelm
parents: 63680
diff changeset
   126
71523
abe3c309998b proper option;
wenzelm
parents: 71503
diff changeset
   127
  Option \<^verbatim>\<open>-L\<close> produces \<^emph>\<open>symlinks\<close> to the original files: this allows to
71389
wenzelm
parents: 71378
diff changeset
   128
  develop Isabelle/Scala/jEdit within an external Scala IDE. Note that
wenzelm
parents: 71378
diff changeset
   129
  building the result always requires \<^verbatim>\<open>isabelle jedit -b\<close> on the
wenzelm
parents: 71378
diff changeset
   130
  command-line.
71378
820cf124dced added "isabelle scala_project" to support e.g. IntelliJ IDEA;
wenzelm
parents: 63680
diff changeset
   131
\<close>
820cf124dced added "isabelle scala_project" to support e.g. IntelliJ IDEA;
wenzelm
parents: 63680
diff changeset
   132
47825
4f25960417ae some coverage of Isabelle/Scala tools;
wenzelm
parents:
diff changeset
   133
end