NEWS
author haftmann
Thu, 18 Jun 2020 09:07:30 +0000
changeset 71956 a4bffc0de967
parent 71949 5b8b1183c641
child 71957 3e162c63371a
permissions -rw-r--r--
bit operations as distinctive library theory

Isabelle NEWS -- history of user-relevant changes
=================================================

(Note: Isabelle/jEdit shows a tree-view of the NEWS file in Sidekick.)


New in this Isabelle version
----------------------------

*** Isabelle/jEdit Prover IDE ***

* Update to jedit-5.6pre1, the latest pre-release. This version works
properly on macOS by default, without the special MacOSX plugin.


*** Document preparation ***

* Antiquotation @{bash_function} refers to GNU bash functions that are
checked within the Isabelle settings environment.

* Antiquotations @{scala}, @{scala_object}, @{scala_type},
@{scala_method} refer to checked Isabelle/Scala entities.


*** Pure ***

* Session Pure-Examples contains notable examples for Isabelle/Pure
(former entries of HOL-Isar_Examples).

* Definitions in locales produce rule which can be added as congruence
rule to protect foundational terms during simplification.



*** HOL ***

* Session HOL-Examples contains notable examples for Isabelle/HOL
(former entries of HOL-Isar_Examples, HOL-ex etc.).

* Simproc "datatype_no_proper_subterm" rewrites equalities "lhs = rhs"
on datatypes to "False" if either side is a proper subexpression of the
other (for any datatype with a reasonable size function).

* New constant "power_int" for exponentiation with integer exponent,
written as "x powi n".

* Added the "at most 1" quantifier, Uniq.

* For the natural numbers, Sup {} = 0.

* Library theory "Bit_Operations" with generic bit operations.

* Session HOL-Word: Bit operations are based on library
theory "Bit_Operations".  INCOMPATIBILITY.

* Session HOL-Word: Compound operation "bin_split" simplifies by default
into its components "drop_bit" and "take_bit".  INCOMPATIBILITY.

* Session HOL-Word: Operations "bin_last", "bin_rest", "bin_nth",
"bintrunc" and "max_word" are now mere input abbreviations.
Minor INCOMPATIBILITY.


*** FOL ***

* Added the "at most 1" quantifier, Uniq, as in HOL.

* Simproc defined_all and rewrite rules subst_all and subst_all' perform
more aggressive substitution with variables from assumptions.
INCOMPATIBILITY, use something like
"supply subst_all [simp] subst_all' [simp] [[simproc del: defined_all]]"
to leave fragile proofs unaffected.


*** ML ***

* Antiquotations @{scala_function} and @{scala} refer to registered
Isabelle/Scala functions (of type String => String): invocation works
via the PIDE protocol.


*** System ***

* System option "pide_session" is enabled by default, notably for
standard "isabelle build": this allows to invoke Isabelle/Scala
operations from Isabelle/ML.

* The command-line tool "isabelle console" now supports interrupts
properly (on Linux and macOS).

* The command-line tool "isabelle sessions" explores the structure of
Isabelle sessions and prints result names in topological order (on
stdout).

* The Isabelle/Scala "Progress" interface changed slightly and
"No_Progress" has been discontinued. INCOMPATIBILITY, use "new Progress"
instead.

* General support for Isabelle/Scala system services, configured via the
shell function "isabelle_scala_service" in etc/settings (e.g. of an
Isabelle component); see implementations of class
Isabelle_System.Service in Isabelle/Scala. This supersedes former
"isabelle_scala_tools" and "isabelle_file_format": minor
INCOMPATIBILITY.

* Isabelle/Phabricator setup has been updated to follow ongoing
development: libphutil has been discontinued. Minor INCOMPATIBILITY:
existing server installations should remove libphutil from
/usr/local/bin/isabelle-phabricator-upgrade and each installation root
directory (e.g. /var/www/phabricator-vcs/libphutil).



New in Isabelle2020 (April 2020)
--------------------------------

*** General ***

* Session ROOT files need to specify explicit 'directories' for import
of theory files. Directories cannot be shared by different sessions.
(Recall that import of theories from other sessions works via
session-qualified theory names, together with suitable 'sessions'
declarations in the ROOT.)

* Internal derivations record dependencies on oracles and other theorems
accurately, including the implicit type-class reasoning wrt. proven
class relations and type arities. In particular, the formal tagging with
"Pure.skip_proofs" of results stemming from "instance ... sorry" is now
propagated properly to theorems depending on such type instances.

* Command 'sorry' (oracle "Pure.skip_proofs") is more precise about the
actual proposition that is assumed in the goal and proof context. This
requires at least Proofterm.proofs = 1 to show up in theorem
dependencies.

* Command 'thm_oracles' prints all oracles used in given theorems,
covering the full graph of transitive dependencies.

* Command 'thm_deps' prints immediate theorem dependencies of the given
facts. The former graph visualization has been discontinued, because it
was hardly usable.

* Refined treatment of proof terms, including type-class proofs for
minor object-logics (FOL, FOLP, Sequents).

* The inference kernel is now confined to one main module: structure
Thm, without the former circular dependency on structure Axclass.

* Mixfix annotations may use "' " (single quote followed by space) to
separate delimiters (as documented in the isar-ref manual), without
requiring an auxiliary empty block. A literal single quote needs to be
escaped properly. Minor INCOMPATIBILITY.


*** Isar ***

* The proof method combinator (subproofs m) applies the method
expression m consecutively to each subgoal, constructing individual
subproofs internally. This impacts the internal construction of proof
terms: it makes a cascade of let-expressions within the derivation tree
and may thus improve scalability.

* Attribute "trace_locales" activates tracing of locale instances during
roundup. It replaces the diagnostic command 'print_dependencies', which
has been discontinued.


*** Isabelle/jEdit Prover IDE ***

* Prover IDE startup is now much faster, because theory dependencies are
no longer explored in advance. The overall session structure with its
declarations of 'directories' is sufficient to locate theory files. Thus
the "session focus" of option "isabelle jedit -S" has become obsolete
(likewise for "isabelle vscode_server -S"). Existing option "-R" is both
sufficient and more convenient to start editing a particular session.

* Actions isabelle.tooltip (CS+b) and isabelle.message (CS+m) display
tooltip message popups, corresponding to mouse hovering with/without the
CONTROL/COMMAND key pressed.

* The following actions allow to navigate errors within the current
document snapshot:

  isabelle.first-error (CS+a)
  isabelle.last-error (CS+z)
  isabelle.next-error (CS+n)
  isabelle.prev-error (CS+p)

* Support more brackets: \<llangle> \<rrangle> (intended for implicit argument syntax).

* Action isabelle.jconsole (menu item Plugins / Isabelle / Java/VM
Monitor) applies the jconsole tool on the running Isabelle/jEdit
process. This allows to monitor resource usage etc.

* More adequate default font sizes for Linux on HD / UHD displays:
automatic font scaling is usually absent on Linux, in contrast to
Windows and macOS.

* The default value for the jEdit property "view.antiAlias" (menu item
Utilities / Global Options / Text Area / Anti Aliased smooth text) is
now "subpixel HRGB", instead of former "standard". Especially on Linux
this often leads to faster text rendering, but can also cause problems
with odd color shades. An alternative is to switch back to "standard"
here, and set the following Java system property:

    isabelle jedit -Dsun.java2d.opengl=true

This can be made persistent via JEDIT_JAVA_OPTIONS in
$ISABELLE_HOME_USER/etc/settings. For the "Isabelle2020" desktop
application there is a corresponding options file in the same directory.


*** Isabelle/VSCode Prover IDE ***

* Update of State and Preview panels to use new WebviewPanel API of
VSCode.


*** HOL ***

* Improvements of the 'lift_bnf' command:
  - Add support for quotient types.
  - Generate transfer rules for the lifted map/set/rel/pred constants
    (theorems "<type>.<constant>_transfer_raw").

* Term_XML.Encode/Decode.term uses compact representation of Const
"typargs" from the given declaration environment. This also makes more
sense for translations to lambda-calculi with explicit polymorphism.
INCOMPATIBILITY, use Term_XML.Encode/Decode.term_raw in special
applications.

* ASCII membership syntax concerning big operators for infimum and
supremum has been discontinued. INCOMPATIBILITY.

* Removed multiplicativity assumption from class
"normalization_semidom". Introduced various new intermediate classes
with the multiplicativity assumption; many theorem statements
(especially involving GCD/LCM) had to be adapted. This allows for a more
natural instantiation of the algebraic typeclasses for e.g. Gaussian
integers. INCOMPATIBILITY.

* Clear distinction between types for bits (False / True) and Z2 (0 /
1): theory HOL-Library.Bit has been renamed accordingly.
INCOMPATIBILITY.

* Dynamic facts "algebra_split_simps" and "field_split_simps" correspond
to algebra_simps and field_simps but contain more aggressive rules
potentially splitting goals; algebra_split_simps roughly replaces
sign_simps and field_split_simps can be used instead of divide_simps.
INCOMPATIBILITY.

* Theory HOL.Complete_Lattices:
renamed Inf_Sup -> Inf_eq_Sup and Sup_Inf -> Sup_eq_Inf

* Theory HOL-Library.Monad_Syntax: infix operation "bind" (\<bind>)
associates to the left now as is customary.

* Theory HOL-Library.Ramsey: full finite Ramsey's theorem with
multiple colours and arbitrary exponents.

* Session HOL-Proofs: build faster thanks to better treatment of proof
terms in Isabelle/Pure.

* Session HOL-Word: bitwise NOT-operator has proper prefix syntax. Minor
INCOMPATIBILITY.

* Session HOL-Analysis: proof method "metric" implements a decision
procedure for simple linear statements in metric spaces.

* Session HOL-Complex_Analysis has been split off from HOL-Analysis.


*** ML ***

* Theory construction may be forked internally, the operation
Theory.join_theory recovers a single result theory. See also the example
in theory "HOL-ex.Join_Theory".

* Antiquotation @{oracle_name} inlines a formally checked oracle name.

* Minimal support for a soft-type system within the Isabelle logical
framework (module Soft_Type_System).

* Former Variable.auto_fixes has been replaced by slightly more general
Proof_Context.augment: it is subject to an optional soft-type system of
the underlying object-logic. Minor INCOMPATIBILITY.

* More scalable Export.export using XML.tree to avoid premature string
allocations, with convenient shortcut XML.blob. Minor INCOMPATIBILITY.

* Prover IDE support for the underlying Poly/ML compiler (not the basis
library). Open $ML_SOURCES/ROOT.ML in Isabelle/jEdit to browse the
implementation with full markup.


*** System ***

* Standard rendering for more Isabelle symbols: \<llangle> \<rrangle> \<bbar> \<sqdot>

* The command-line tool "isabelle scala_project" creates a Gradle
project configuration for Isabelle/Scala/jEdit, to support Scala IDEs
such as IntelliJ IDEA.

* The command-line tool "isabelle phabricator_setup" facilitates
self-hosting of the Phabricator software-development platform, with
support for Git, Mercurial, Subversion repositories. This helps to avoid
monoculture and to escape the gravity of centralized version control by
Github and/or Bitbucket. For further documentation, see chapter
"Phabricator server administration" in the "system" manual. A notable
example installation is https://isabelle-dev.sketis.net/.

* The command-line tool "isabelle hg_setup" simplifies the setup of
Mercurial repositories, with hosting via Phabricator or SSH file server
access.

* The command-line tool "isabelle imports" has been discontinued: strict
checking of session directories enforces session-qualified theory names
in applications -- users are responsible to specify session ROOT entries
properly.

* The command-line tool "isabelle dump" and its underlying
Isabelle/Scala module isabelle.Dump has become more scalable, by
splitting sessions and supporting a base logic image. Minor
INCOMPATIBILITY in options and parameters.

* The command-line tool "isabelle build_docker" has been slightly
improved: it is now properly documented in the "system" manual.

* Isabelle/Scala support for the Linux platform (Ubuntu): packages,
users, system services.

* Isabelle/Scala support for proof terms (with full type/term
information) in module isabelle.Term.

* Isabelle/Scala: more scalable output of YXML files, e.g. relevant for
"isabelle dump".

* Theory export via Isabelle/Scala has been reworked. The former "fact"
name space is now split into individual "thm" items: names are
potentially indexed, such as "foo" for singleton facts, or "bar(1)",
"bar(2)", "bar(3)" for multi-facts. Theorem dependencies are now
exported as well: this spans an overall dependency graph of internal
inferences; it might help to reconstruct the formal structure of theory
libraries. See also the module isabelle.Export_Theory in Isabelle/Scala.

* Theory export of structured specifications, based on internal
declarations of Spec_Rules by packages like 'definition', 'inductive',
'primrec', 'function'.

* Old settings variables ISABELLE_PLATFORM and ISABELLE_WINDOWS_PLATFORM
have been discontinued -- deprecated since Isabelle2018.

* More complete x86_64 platform support on macOS, notably Catalina where
old x86 has been discontinued.

* Update to GHC stack 2.1.3 with stackage lts-13.19/ghc-8.6.4.

* Update to OCaml Opam 2.0.6 (using ocaml 4.05.0 as before).



New in Isabelle2019 (June 2019)
-------------------------------

*** General ***

* The font collection "Isabelle DejaVu" is systematically derived from
the existing "DejaVu" fonts, with variants "Sans Mono", "Sans", "Serif"
and styles "Normal", "Bold", "Italic/Oblique", "Bold-Italic/Oblique".
The DejaVu base fonts are retricted to well-defined Unicode ranges and
augmented by special Isabelle symbols, taken from the former
"IsabelleText" font (which is no longer provided separately). The line
metrics and overall rendering quality is closer to original DejaVu.
INCOMPATIBILITY with display configuration expecting the old
"IsabelleText" font: use e.g. "Isabelle DejaVu Sans Mono" instead.

* The Isabelle fonts render "\<inverse>" properly as superscript "-1".

* Old-style inner comments (* ... *) within the term language are no
longer supported (legacy feature in Isabelle2018).

* Old-style {* verbatim *} tokens are explicitly marked as legacy
feature and will be removed soon. Use \<open>cartouche\<close> syntax instead, e.g.
via "isabelle update_cartouches -t" (available since Isabelle2015).

* Infix operators that begin or end with a "*" are now parenthesized
without additional spaces, e.g. "(*)" instead of "( * )". Minor
INCOMPATIBILITY.

* Mixfix annotations may use cartouches instead of old-style double
quotes, e.g. (infixl \<open>+\<close> 60). The command-line tool "isabelle update -u
mixfix_cartouches" allows to update existing theory sources
automatically.

* ML setup commands (e.g. 'setup', 'method_setup', 'parse_translation')
need to provide a closed expression -- without trailing semicolon. Minor
INCOMPATIBILITY.

* Commands 'generate_file', 'export_generated_files', and
'compile_generated_files' support a stateless (PIDE-conformant) model
for generated sources and compiled binaries of other languages. The
compilation process is managed in Isabelle/ML, and results exported to
the session database for further use (e.g. with "isabelle export" or
"isabelle build -e").


*** Isabelle/jEdit Prover IDE ***

* Fonts for the text area, gutter, GUI elements etc. use the "Isabelle
DejaVu" collection by default, which provides uniform rendering quality
with the usual Isabelle symbols. Line spacing no longer needs to be
adjusted: properties for the old IsabelleText font had "Global Options /
Text Area / Extra vertical line spacing (in pixels): -2", it now
defaults to 1, but 0 works as well.

* The jEdit File Browser is more prominent in the default GUI layout of
Isabelle/jEdit: various virtual file-systems provide access to Isabelle
resources, notably via "favorites:" (or "Edit Favorites").

* Further markup and rendering for "plain text" (e.g. informal prose)
and "raw text" (e.g. verbatim sources). This improves the visual
appearance of formal comments inside the term language, or in general
for repeated alternation of formal and informal text.

* Action "isabelle-export-browser" points the File Browser to the theory
exports of the current buffer, based on the "isabelle-export:" virtual
file-system. The directory view needs to be reloaded manually to follow
ongoing document processing.

* Action "isabelle-session-browser" points the File Browser to session
information, based on the "isabelle-session:" virtual file-system. Its
entries are structured according to chapter / session names, the open
operation is redirected to the session ROOT file.

* Support for user-defined file-formats via class isabelle.File_Format
in Isabelle/Scala (e.g. see isabelle.Bibtex.File_Format), configured via
the shell function "isabelle_file_format" in etc/settings (e.g. of an
Isabelle component).

* System option "jedit_text_overview" allows to disable the text
overview column.

* Command-line options "-s" and "-u" of "isabelle jedit" override the
default for system option "system_heaps" that determines the heap
storage directory for "isabelle build". Option "-n" is now clearly
separated from option "-s".

* The Isabelle/jEdit desktop application uses the same options as
"isabelle jedit" for its internal "isabelle build" process: the implicit
option "-o system_heaps" (or "-s") has been discontinued. This reduces
the potential for surprise wrt. command-line tools.

* The official download of the Isabelle/jEdit application already
contains heap images for Isabelle/HOL within its main directory: thus
the first encounter becomes faster and more robust (e.g. when run from a
read-only directory).

* Isabelle DejaVu fonts are available with hinting by default, which is
relevant for low-resolution displays. This may be disabled via system
option "isabelle_fonts_hinted = false" in
$ISABELLE_HOME_USER/etc/preferences -- it occasionally yields better
results.

* OpenJDK 11 has quite different font rendering, with better glyph
shapes and improved sub-pixel anti-aliasing. In some situations results
might be *worse* than Oracle Java 8, though -- a proper HiDPI / UHD
display is recommended.

* OpenJDK 11 supports GTK version 2.2 and 3 (according to system
property jdk.gtk.version). The factory default is version 3, but
ISABELLE_JAVA_SYSTEM_OPTIONS includes "-Djdk.gtk.version=2.2" to make
this more conservative (as in Java 8). Depending on the GTK theme
configuration, "-Djdk.gtk.version=3" might work better or worse.


*** Document preparation ***

* Document markers are formal comments of the form \<^marker>\<open>marker_body\<close> that
are stripped from document output: the effect is to modify the semantic
presentation context or to emit markup to the PIDE document. Some
predefined markers are taken from the Dublin Core Metadata Initiative,
e.g. \<^marker>\<open>contributor arg\<close> or \<^marker>\<open>license arg\<close> and produce PIDE markup that
can be retrieved from the document database.

* Old-style command tags %name are re-interpreted as markers with
proof-scope \<^marker>\<open>tag (proof) name\<close> and produce LaTeX environments as
before. Potential INCOMPATIBILITY: multiple markers are composed in
canonical order, resulting in a reversed list of tags in the
presentation context.

* Marker \<^marker>\<open>tag name\<close> does not apply to the proof of a top-level goal
statement by default (e.g. 'theorem', 'lemma'). This is a subtle change
of semantics wrt. old-style %name.

* In Isabelle/jEdit, the string "\tag" may be completed to a "\<^marker>\<open>tag \<close>"
template.

* Document antiquotation option "cartouche" indicates if the output
should be delimited as cartouche; this takes precedence over the
analogous option "quotes".

* Many document antiquotations are internally categorized as "embedded"
and expect one cartouche argument, which is typically used with the
\<^control>\<open>cartouche\<close> notation (e.g. \<^term>\<open>\<lambda>x y. x\<close>). The cartouche
delimiters are stripped in output of the source (antiquotation option
"source"), but it is possible to enforce delimiters via option
"source_cartouche", e.g. @{term [source_cartouche] \<open>\<lambda>x y. x\<close>}.


*** Isar ***

* Implicit cases goal1, goal2, goal3, etc. have been discontinued
(legacy feature since Isabelle2016).

* More robust treatment of structural errors: begin/end blocks take
precedence over goal/proof. This is particularly relevant for the
headless PIDE session and server.

* Command keywords of kind thy_decl / thy_goal may be more specifically
fit into the traditional document model of "definition-statement-proof"
via thy_defn / thy_stmt / thy_goal_defn / thy_goal_stmt.


*** HOL ***

* Command 'export_code' produces output as logical files within the
theory context, as well as formal session exports that can be
materialized via command-line tools "isabelle export" or "isabelle build
-e" (with 'export_files' in the session ROOT). Isabelle/jEdit also
provides a virtual file-system "isabelle-export:" that can be explored
in the regular file-browser. A 'file_prefix' argument allows to specify
an explicit name prefix for the target file (SML, OCaml, Scala) or
directory (Haskell); the default is "export" with a consecutive number
within each theory.

* Command 'export_code': the 'file' argument is now legacy and will be
removed soon: writing to the physical file-system is not well-defined in
a reactive/parallel application like Isabelle. The empty 'file' argument
has been discontinued already: it is superseded by the file-browser in
Isabelle/jEdit on "isabelle-export:". Minor INCOMPATIBILITY.

* Command 'code_reflect' no longer supports the 'file' argument: it has
been superseded by 'file_prefix' for stateless file management as in
'export_code'. Minor INCOMPATIBILITY.

* Code generation for OCaml: proper strings are used for literals.
Minor INCOMPATIBILITY.

* Code generation for OCaml: Zarith supersedes Nums as library for
proper integer arithmetic. The library is located via standard
invocations of "ocamlfind" (via ISABELLE_OCAMLFIND settings variable).
The environment provided by "isabelle ocaml_setup" already contains this
tool and the required packages. Minor INCOMPATIBILITY.

* Code generation for Haskell: code includes for Haskell must contain
proper module frame, nothing is added magically any longer.
INCOMPATIBILITY.

* Code generation: slightly more conventional syntax for 'code_stmts'
antiquotation. Minor INCOMPATIBILITY.

* Theory List: the precedence of the list_update operator has changed:
"f a [n := x]" now needs to be written "(f a)[n := x]".

* The functions \<Union>, \<Inter>, \<Squnion>, \<Sqinter> (not the corresponding binding operators)
now have the same precedence as any other prefix function symbol. Minor
INCOMPATIBILITY.

* Simplified syntax setup for big operators under image. In rare
situations, type conversions are not inserted implicitly any longer
and need to be given explicitly. Auxiliary abbreviations INFIMUM,
SUPREMUM, UNION, INTER should now rarely occur in output and are just
retained as migration auxiliary. Abbreviations MINIMUM and MAXIMUM
are gone INCOMPATIBILITY.

* The simplifier uses image_cong_simp as a congruence rule. The historic
and not really well-formed congruence rules INF_cong*, SUP_cong*, are
not used by default any longer. INCOMPATIBILITY; consider using declare
image_cong_simp [cong del] in extreme situations.

* INF_image and SUP_image are no default simp rules any longer.
INCOMPATIBILITY, prefer image_comp as simp rule if needed.

* Strong congruence rules (with =simp=> in the premises) for constant f
are now uniformly called f_cong_simp, in accordance with congruence
rules produced for mappers by the datatype package. INCOMPATIBILITY.

* Retired lemma card_Union_image; use the simpler card_UN_disjoint
instead. INCOMPATIBILITY.

* Facts sum_mset.commute and prod_mset.commute have been renamed to
sum_mset.swap and prod_mset.swap, similarly to sum.swap and prod.swap.
INCOMPATIBILITY.

* ML structure Inductive: slightly more conventional naming schema.
Minor INCOMPATIBILITY.

* ML: Various _global variants of specification tools have been removed.
Minor INCOMPATIBILITY, prefer combinators
Named_Target.theory_map[_result] to lift specifications to the global
theory level.

* Theory HOL-Library.Simps_Case_Conv: 'case_of_simps' now supports
overlapping and non-exhaustive patterns and handles arbitrarily nested
patterns. It uses on the same algorithm as HOL-Library.Code_Lazy, which
assumes sequential left-to-right pattern matching. The generated
equation no longer tuples the arguments on the right-hand side.
INCOMPATIBILITY.

* Theory HOL-Library.Multiset: the \<Union># operator now has the same
precedence as any other prefix function symbol.

* Theory HOL-Library.Cardinal_Notations has been discontinued in favor
of the bundle cardinal_syntax (available in theory Main). Minor
INCOMPATIBILITY.

* Session HOL-Library and HOL-Number_Theory: Exponentiation by squaring,
used for computing powers in class "monoid_mult" and modular
exponentiation.

* Session HOL-Computational_Algebra: Formal Laurent series and overhaul
of Formal power series.

* Session HOL-Number_Theory: More material on residue rings in
Carmichael's function, primitive roots, more properties for "ord".

* Session HOL-Analysis: Better organization and much more material
at the level of abstract topological spaces.

* Session HOL-Algebra: Free abelian groups, etc., ported from HOL Light;
 algebraic closure of a field by de Vilhena and Baillon.

* Session HOL-Homology has been added. It is a port of HOL Light's
homology library, with new proofs of "invariance of domain" and related
results.

* Session HOL-SPARK: .prv files are no longer written to the
file-system, but exported to the session database. Results may be
retrieved via "isabelle build -e HOL-SPARK-Examples" on the
command-line.

* Sledgehammer:
  - The URL for SystemOnTPTP, which is used by remote provers, has been
    updated.
  - The machine-learning-based filter MaSh has been optimized to take
    less time (in most cases).

* SMT: reconstruction is now possible using the SMT solver veriT.

* Session HOL-Word:
  * New theory More_Word as comprehensive entrance point.
  * Merged type class bitss into type class bits.
  INCOMPATIBILITY.


*** ML ***

* Command 'generate_file' allows to produce sources for other languages,
with antiquotations in the Isabelle context (only the control-cartouche
form). The default "cartouche" antiquotation evaluates an ML expression
of type string and inlines the result as a string literal of the target
language. For example, this works for Haskell as follows:

  generate_file "Pure.hs" = \<open>
  module Isabelle.Pure where
    allConst, impConst, eqConst :: String
    allConst = \<open>\<^const_name>\<open>Pure.all\<close>\<close>
    impConst = \<open>\<^const_name>\<open>Pure.imp\<close>\<close>
    eqConst = \<open>\<^const_name>\<open>Pure.eq\<close>\<close>
  \<close>

See also commands 'export_generated_files' and 'compile_generated_files'
to use the results.

* ML evaluation (notably via command 'ML' or 'ML_file') is subject to
option ML_environment to select a named environment, such as "Isabelle"
for Isabelle/ML, or "SML" for official Standard ML.

* ML antiquotation @{master_dir} refers to the master directory of the
underlying theory, i.e. the directory of the theory file.

* ML antiquotation @{verbatim} inlines its argument as string literal,
preserving newlines literally. The short form \<^verbatim>\<open>abc\<close> is particularly
useful.

* Local_Theory.reset is no longer available in user space. Regular
definitional packages should use balanced blocks of
Local_Theory.open_target versus Local_Theory.close_target instead, or
the Local_Theory.subtarget(_result) combinator. Rare INCOMPATIBILITY.

* Original PolyML.pointerEq is retained as a convenience for tools that
don't use Isabelle/ML (where this is called "pointer_eq").


*** System ***

* Update to OpenJDK 11: the current long-term support version of Java.

* Update to Poly/ML 5.8 allows to use the native x86_64 platform without
the full overhead of 64-bit values everywhere. This special x86_64_32
mode provides up to 16GB ML heap, while program code and stacks are
allocated elsewhere. Thus approx. 5 times more memory is available for
applications compared to old x86 mode (which is no longer used by
Isabelle). The switch to the x86_64 CPU architecture also avoids
compatibility problems with Linux and macOS, where 32-bit applications
are gradually phased out.

* System option "checkpoint" has been discontinued: obsolete thanks to
improved memory management in Poly/ML.

* System option "system_heaps" determines where to store the session
image of "isabelle build" (and other tools using that internally).
Former option "-s" is superseded by option "-o system_heaps".
INCOMPATIBILITY in command-line syntax.

* Session directory $ISABELLE_HOME/src/Tools/Haskell provides some
source modules for Isabelle tools implemented in Haskell, notably for
Isabelle/PIDE.

* The command-line tool "isabelle build -e" retrieves theory exports
from the session build database, using 'export_files' in session ROOT
entries.

* The command-line tool "isabelle update" uses Isabelle/PIDE in
batch-mode to update theory sources based on semantic markup produced in
Isabelle/ML. Actual updates depend on system options that may be enabled
via "-u OPT" (for "update_OPT"), see also $ISABELLE_HOME/etc/options
section "Theory update". Theory sessions are specified as in "isabelle
dump".

* The command-line tool "isabelle update -u control_cartouches" changes
antiquotations into control-symbol format (where possible): @{NAME}
becomes \<^NAME> and @{NAME ARG} becomes \<^NAME>\<open>ARG\<close>.

* Support for Isabelle command-line tools defined in Isabelle/Scala.
Instances of class Isabelle_Scala_Tools may be configured via the shell
function "isabelle_scala_tools" in etc/settings (e.g. of an Isabelle
component).

* Isabelle Server command "use_theories" supports "nodes_status_delay"
for continuous output of node status information. The time interval is
specified in seconds; a negative value means it is disabled (default).

* Isabelle Server command "use_theories" terminates more robustly in the
presence of structurally broken sources: full consolidation of theories
is no longer required.

* OCaml tools and libraries are now accesed via ISABELLE_OCAMLFIND,
which needs to point to a suitable version of "ocamlfind" (e.g. via
OPAM, see below). INCOMPATIBILITY: settings variables ISABELLE_OCAML and
ISABELLE_OCAMLC are no longer supported.

* Support for managed installations of Glasgow Haskell Compiler and
OCaml via the following command-line tools:

  isabelle ghc_setup
  isabelle ghc_stack

  isabelle ocaml_setup
  isabelle ocaml_opam

The global installation state is determined by the following settings
(and corresponding directory contents):

  ISABELLE_STACK_ROOT
  ISABELLE_STACK_RESOLVER
  ISABELLE_GHC_VERSION

  ISABELLE_OPAM_ROOT
  ISABELLE_OCAML_VERSION

After setup, the following Isabelle settings are automatically
redirected (overriding existing user settings):

  ISABELLE_GHC

  ISABELLE_OCAMLFIND

The old meaning of these settings as locally installed executables may
be recovered by purging the directories ISABELLE_STACK_ROOT /
ISABELLE_OPAM_ROOT, or by resetting these variables in
$ISABELLE_HOME_USER/etc/settings.



New in Isabelle2018 (August 2018)
---------------------------------

*** General ***

* Session-qualified theory names are mandatory: it is no longer possible
to refer to unqualified theories from the parent session.
INCOMPATIBILITY for old developments that have not been updated to
Isabelle2017 yet (using the "isabelle imports" tool).

* Only the most fundamental theory names are global, usually the entry
points to major logic sessions: Pure, Main, Complex_Main, HOLCF, IFOL,
FOL, ZF, ZFC etc. INCOMPATIBILITY, need to use qualified names for
formerly global "HOL-Probability.Probability" and "HOL-SPARK.SPARK".

* Global facts need to be closed: no free variables and no hypotheses.
Rare INCOMPATIBILITY.

* Facts stemming from locale interpretation are subject to lazy
evaluation for improved performance. Rare INCOMPATIBILITY: errors
stemming from interpretation morphisms might be deferred and thus
difficult to locate; enable system option "strict_facts" temporarily to
avoid this.

* Marginal comments need to be written exclusively in the new-style form
"\<comment> \<open>text\<close>", old ASCII variants like "-- {* ... *}" are no longer
supported. INCOMPATIBILITY, use the command-line tool "isabelle
update_comments" to update existing theory files.

* Old-style inner comments (* ... *) within the term language are legacy
and will be discontinued soon: use formal comments "\<comment> \<open>...\<close>" or "\<^cancel>\<open>...\<close>"
instead.

* The "op <infix-op>" syntax for infix operators has been replaced by
"(<infix-op>)". If <infix-op> begins or ends with a "*", there needs to
be a space between the "*" and the corresponding parenthesis.
INCOMPATIBILITY, use the command-line tool "isabelle update_op" to
convert theory and ML files to the new syntax. Because it is based on
regular expression matching, the result may need a bit of manual
postprocessing. Invoking "isabelle update_op" converts all files in the
current directory (recursively). In case you want to exclude conversion
of ML files (because the tool frequently also converts ML's "op"
syntax), use option "-m".

* Theory header 'abbrevs' specifications need to be separated by 'and'.
INCOMPATIBILITY.

* Command 'external_file' declares the formal dependency on the given
file name, such that the Isabelle build process knows about it, but
without specific Prover IDE management.

* Session ROOT entries no longer allow specification of 'files'. Rare
INCOMPATIBILITY, use command 'external_file' within a proper theory
context.

* Session root directories may be specified multiple times: each
accessible ROOT file is processed only once. This facilitates
specification of $ISABELLE_HOME_USER/ROOTS or command-line options like
-d or -D for "isabelle build" and "isabelle jedit". Example:

  isabelle build -D '~~/src/ZF'

* The command 'display_drafts' has been discontinued. INCOMPATIBILITY,
use action "isabelle.draft" (or "print") in Isabelle/jEdit instead.

* In HTML output, the Isabelle symbol "\<hyphen>" is rendered as explicit
Unicode hyphen U+2010, to avoid unclear meaning of the old "soft hyphen"
U+00AD. Rare INCOMPATIBILITY, e.g. copy-paste of historic Isabelle HTML
output.


*** Isabelle/jEdit Prover IDE ***

* The command-line tool "isabelle jedit" provides more flexible options
for session management:

  - option -R builds an auxiliary logic image with all theories from
    other sessions that are not already present in its parent

  - option -S is like -R, with a focus on the selected session and its
    descendants (this reduces startup time for big projects like AFP)

  - option -A specifies an alternative ancestor session for options -R
    and -S

  - option -i includes additional sessions into the name-space of
    theories

  Examples:
    isabelle jedit -R HOL-Number_Theory
    isabelle jedit -R HOL-Number_Theory -A HOL
    isabelle jedit -d '$AFP' -S Formal_SSA -A HOL
    isabelle jedit -d '$AFP' -S Formal_SSA -A HOL-Analysis
    isabelle jedit -d '$AFP' -S Formal_SSA -A HOL-Analysis -i CryptHOL

* PIDE markup for session ROOT files: allows to complete session names,
follow links to theories and document files etc.

* Completion supports theory header imports, using theory base name.
E.g. "Prob" may be completed to "HOL-Probability.Probability".

* Named control symbols (without special Unicode rendering) are shown as
bold-italic keyword. This is particularly useful for the short form of
antiquotations with control symbol: \<^name>\<open>argument\<close>. The action
"isabelle.antiquoted_cartouche" turns an antiquotation with 0 or 1
arguments into this format.

* Completion provides templates for named symbols with arguments,
e.g. "\<comment> \<open>ARGUMENT\<close>" or "\<^emph>\<open>ARGUMENT\<close>".

* Slightly more parallel checking, notably for high priority print
functions (e.g. State output).

* The view title is set dynamically, according to the Isabelle
distribution and the logic session name. The user can override this via
set-view-title (stored persistently in $JEDIT_SETTINGS/perspective.xml).

* System options "spell_checker_include" and "spell_checker_exclude"
supersede former "spell_checker_elements" to determine regions of text
that are subject to spell-checking. Minor INCOMPATIBILITY.

* Action "isabelle.preview" is able to present more file formats,
notably bibtex database files and ML files.

* Action "isabelle.draft" is similar to "isabelle.preview", but shows a
plain-text document draft. Both are available via the menu "Plugins /
Isabelle".

* When loading text files, the Isabelle symbols encoding UTF-8-Isabelle
is only used if there is no conflict with existing Unicode sequences in
the file. Otherwise, the fallback encoding is plain UTF-8 and Isabelle
symbols remain in literal \<symbol> form. This avoids accidental loss of
Unicode content when saving the file.

* Bibtex database files (.bib) are semantically checked.

* Update to jedit-5.5.0, the latest release.


*** Isabelle/VSCode Prover IDE ***

* HTML preview of theories and other file-formats similar to
Isabelle/jEdit.

* Command-line tool "isabelle vscode_server" accepts the same options
-A, -R, -S, -i for session selection as "isabelle jedit". This is
relevant for isabelle.args configuration settings in VSCode. The former
option -A (explore all known session files) has been discontinued: it is
enabled by default, unless option -S is used to focus on a particular
spot in the session structure. INCOMPATIBILITY.


*** Document preparation ***

* Formal comments work uniformly in outer syntax, inner syntax (term
language), Isabelle/ML and some other embedded languages of Isabelle.
See also "Document comments" in the isar-ref manual. The following forms
are supported:

  - marginal text comment: \<comment> \<open>\<dots>\<close>
  - canceled source: \<^cancel>\<open>\<dots>\<close>
  - raw LaTeX: \<^latex>\<open>\<dots>\<close>

* Outside of the inner theory body, the default presentation context is
theory Pure. Thus elementary antiquotations may be used in markup
commands (e.g. 'chapter', 'section', 'text') and formal comments.

* System option "document_tags" specifies alternative command tags. This
is occasionally useful to control the global visibility of commands via
session options (e.g. in ROOT).

* Document markup commands ('section', 'text' etc.) are implicitly
tagged as "document" and visible by default. This avoids the application
of option "document_tags" to these commands.

* Isabelle names are mangled into LaTeX macro names to allow the full
identifier syntax with underscore, prime, digits. This is relevant for
antiquotations in control symbol notation, e.g. \<^const_name> becomes
\isactrlconstUNDERSCOREname.

* Document preparation with skip_proofs option now preserves the content
more accurately: only terminal proof steps ('by' etc.) are skipped.

* Document antiquotation @{theory name} requires the long
session-qualified theory name: this is what users reading the text
normally need to import.

* Document antiquotation @{session name} checks and prints the given
session name verbatim.

* Document antiquotation @{cite} now checks the given Bibtex entries
against the Bibtex database files -- only in batch-mode session builds.

* Command-line tool "isabelle document" has been re-implemented in
Isabelle/Scala, with simplified arguments and explicit errors from the
latex and bibtex process. Minor INCOMPATIBILITY.

* Session ROOT entry: empty 'document_files' means there is no document
for this session. There is no need to specify options [document = false]
anymore.


*** Isar ***

* Command 'interpret' no longer exposes resulting theorems as literal
facts, notably for the \<open>prop\<close> notation or the "fact" proof method. This
improves modularity of proofs and scalability of locale interpretation.
Rare INCOMPATIBILITY, need to refer to explicitly named facts instead
(e.g. use 'find_theorems' or 'try' to figure this out).

* The old 'def' command has been discontinued (legacy since
Isbelle2016-1). INCOMPATIBILITY, use 'define' instead -- usually with
object-logic equality or equivalence.


*** Pure ***

* The inner syntax category "sort" now includes notation "_" for the
dummy sort: it is effectively ignored in type-inference.

* Rewrites clauses (keyword 'rewrites') were moved into the locale
expression syntax, where they are part of locale instances. In
interpretation commands rewrites clauses now need to occur before 'for'
and 'defines'. Rare INCOMPATIBILITY; definitions immediately subject to
rewriting may need to be pulled up into the surrounding theory.

* For 'rewrites' clauses, if activating a locale instance fails, fall
back to reading the clause first. This helps avoid qualification of
locale instances where the qualifier's sole purpose is avoiding
duplicate constant declarations.

* Proof method "simp" now supports a new modifier "flip:" followed by a
list of theorems. Each of these theorems is removed from the simpset
(without warning if it is not there) and the symmetric version of the
theorem (i.e. lhs and rhs exchanged) is added to the simpset. For "auto"
and friends the modifier is "simp flip:".


*** HOL ***

* Sledgehammer: bundled version of "vampire" (for non-commercial users)
helps to avoid fragility of "remote_vampire" service.

* Clarified relationship of characters, strings and code generation:

  - Type "char" is now a proper datatype of 8-bit values.

  - Conversions "nat_of_char" and "char_of_nat" are gone; use more
    general conversions "of_char" and "char_of" with suitable type
    constraints instead.

  - The zero character is just written "CHR 0x00", not "0" any longer.

  - Type "String.literal" (for code generation) is now isomorphic to
    lists of 7-bit (ASCII) values; concrete values can be written as
    "STR ''...''" for sequences of printable characters and "STR 0x..."
    for one single ASCII code point given as hexadecimal numeral.

  - Type "String.literal" supports concatenation "... + ..." for all
    standard target languages.

  - Theory HOL-Library.Code_Char is gone; study the explanations
    concerning "String.literal" in the tutorial on code generation to
    get an idea how target-language string literals can be converted to
    HOL string values and vice versa.

  - Session Imperative-HOL: operation "raise" directly takes a value of
    type "String.literal" as argument, not type "string".

INCOMPATIBILITY.

* Code generation: Code generation takes an explicit option
"case_insensitive" to accomodate case-insensitive file systems.

* Abstract bit operations as part of Main: push_bit, take_bit, drop_bit.

* New, more general, axiomatization of complete_distrib_lattice. The
former axioms:

  "sup x (Inf X) = Inf (sup x ` X)" and "inf x (Sup X) = Sup (inf x ` X)"

are replaced by:

  "Inf (Sup ` A) <= Sup (Inf ` {f ` A | f . (! Y \<in> A . f Y \<in> Y)})"

The instantiations of sets and functions as complete_distrib_lattice are
moved to Hilbert_Choice.thy because their proofs need the Hilbert choice
operator. The dual of this property is also proved in theory
HOL.Hilbert_Choice.

* New syntax for the minimum/maximum of a function over a finite set:
MIN x\<in>A. B and even MIN x. B (only useful for finite types), also MAX.

* Clarifed theorem names:

  Min.antimono ~> Min.subset_imp
  Max.antimono ~> Max.subset_imp

Minor INCOMPATIBILITY.

* SMT module:

  - The 'smt_oracle' option is now necessary when using the 'smt' method
    with a solver other than Z3. INCOMPATIBILITY.

  - The encoding to first-order logic is now more complete in the
    presence of higher-order quantifiers. An 'smt_explicit_application'
    option has been added to control this. INCOMPATIBILITY.

* Facts sum.commute(_restrict) and prod.commute(_restrict) renamed to
sum.swap(_restrict) and prod.swap(_restrict), to avoid name clashes on
interpretation of abstract locales. INCOMPATIBILITY.

* Predicate coprime is now a real definition, not a mere abbreviation.
INCOMPATIBILITY.

* Predicate pairwise_coprime abolished, use "pairwise coprime" instead.
INCOMPATIBILITY.

* The relator rel_filter on filters has been strengthened to its
canonical categorical definition with better properties.
INCOMPATIBILITY.

* Generalized linear algebra involving linear, span, dependent, dim
from type class real_vector to locales module and vector_space.
Renamed:

  span_inc ~> span_superset
  span_superset ~> span_base
  span_eq ~> span_eq_iff

INCOMPATIBILITY.

* Class linordered_semiring_1 covers zero_less_one also, ruling out
pathologic instances. Minor INCOMPATIBILITY.

* Theory HOL.List: functions "sorted_wrt" and "sorted" now compare every
element in a list to all following elements, not just the next one.

* Theory HOL.List syntax:

  - filter-syntax "[x <- xs. P]" is no longer output syntax, but only
    input syntax

  - list comprehension syntax now supports tuple patterns in "pat <- xs"

* Theory Map: "empty" must now be qualified as "Map.empty".

* Removed nat-int transfer machinery. Rare INCOMPATIBILITY.

* Fact mod_mult_self4 (on nat) renamed to Suc_mod_mult_self3, to avoid
clash with fact mod_mult_self4 (on more generic semirings).
INCOMPATIBILITY.

* Eliminated some theorem aliasses:
  even_times_iff ~> even_mult_iff
  mod_2_not_eq_zero_eq_one_nat ~> not_mod_2_eq_0_eq_1
  even_of_nat ~> even_int_iff

INCOMPATIBILITY.

* Eliminated some theorem duplicate variations:

  - dvd_eq_mod_eq_0_numeral can be replaced by dvd_eq_mod_eq_0
  - mod_Suc_eq_Suc_mod can be replaced by mod_Suc
  - mod_Suc_eq_Suc_mod [symmetrict] can be replaced by mod_simps
  - mod_eq_0_iff can be replaced by mod_eq_0_iff_dvd and dvd_def
  - the witness of mod_eqD can be given directly as "_ div _"

INCOMPATIBILITY.

* Classical setup: Assumption "m mod d = 0" (for m d :: nat) is no
longer aggresively destroyed to "\<exists>q. m = d * q". INCOMPATIBILITY, adding
"elim!: dvd" to classical proof methods in most situations restores
broken proofs.

* Theory HOL-Library.Conditional_Parametricity provides command
'parametric_constant' for proving parametricity of non-recursive
definitions. For constants that are not fully parametric the command
will infer conditions on relations (e.g., bi_unique, bi_total, or type
class conditions such as "respects 0") sufficient for parametricity. See
theory HOL-ex.Conditional_Parametricity_Examples for some examples.

* Theory HOL-Library.Code_Lazy provides a new preprocessor for the code
generator to generate code for algebraic types with lazy evaluation
semantics even in call-by-value target languages. See the theories
HOL-ex.Code_Lazy_Demo and HOL-Codegenerator_Test.Code_Lazy_Test for some
examples.

* Theory HOL-Library.Landau_Symbols has been moved here from AFP.

* Theory HOL-Library.Old_Datatype no longer provides the legacy command
'old_datatype'. INCOMPATIBILITY.

* Theory HOL-Computational_Algebra.Polynomial_Factorial does not provide
instances of rat, real, complex as factorial rings etc. Import
HOL-Computational_Algebra.Field_as_Ring explicitly in case of need.
INCOMPATIBILITY.

* Session HOL-Algebra: renamed (^) to [^] to avoid conflict with new
infix/prefix notation.

* Session HOL-Algebra: revamped with much new material. The set of
isomorphisms between two groups is now denoted iso rather than iso_set.
INCOMPATIBILITY.

* Session HOL-Analysis: the Arg function now respects the same interval
as Ln, namely (-pi,pi]; the old Arg function has been renamed Arg2pi.
INCOMPATIBILITY.

* Session HOL-Analysis: the functions zorder, zer_poly, porder and
pol_poly have been redefined. All related lemmas have been reworked.
INCOMPATIBILITY.

* Session HOL-Analysis: infinite products, Moebius functions, the
Riemann mapping theorem, the Vitali covering theorem,
change-of-variables results for integration and measures.

* Session HOL-Real_Asymp: proof method "real_asymp" proves asymptotics
or real-valued functions (limits, "Big-O", etc.) automatically.
See also ~~/src/HOL/Real_Asymp/Manual for some documentation.

* Session HOL-Types_To_Sets: more tool support (unoverload_type combines
internalize_sorts and unoverload) and larger experimental application
(type based linear algebra transferred to linear algebra on subspaces).


*** ML ***

* Operation Export.export emits theory exports (arbitrary blobs), which
are stored persistently in the session build database.

* Command 'ML_export' exports ML toplevel bindings to the global
bootstrap environment of the ML process. This allows ML evaluation
without a formal theory context, e.g. in command-line tools like
"isabelle process".


*** System ***

* Mac OS X 10.10 Yosemite is now the baseline version; Mavericks is no
longer supported.

* Linux and Windows/Cygwin is for x86_64 only, old 32bit platform
support has been discontinued.

* Java runtime is for x86_64 only. Corresponding Isabelle settings have
been renamed to ISABELLE_TOOL_JAVA_OPTIONS and JEDIT_JAVA_OPTIONS,
instead of former 32/64 variants. INCOMPATIBILITY.

* Old settings ISABELLE_PLATFORM and ISABELLE_WINDOWS_PLATFORM should be
phased out due to unclear preference of 32bit vs. 64bit architecture.
Explicit GNU bash expressions are now preferred, for example (with
quotes):

  #Posix executables (Unix or Cygwin), with preference for 64bit
  "${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM32}"

  #native Windows or Unix executables, with preference for 64bit
  "${ISABELLE_WINDOWS_PLATFORM64:-${ISABELLE_WINDOWS_PLATFORM32:-${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM32}}}"

  #native Windows (32bit) or Unix executables (preference for 64bit)
  "${ISABELLE_WINDOWS_PLATFORM32:-${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM32}}"

* Command-line tool "isabelle build" supports new options:
  - option -B NAME: include session NAME and all descendants
  - option -S: only observe changes of sources, not heap images
  - option -f: forces a fresh build

* Command-line tool "isabelle build" options -c -x -B refer to
descendants wrt. the session parent or import graph. Subtle
INCOMPATIBILITY: options -c -x used to refer to the session parent graph
only.

* Command-line tool "isabelle build" takes "condition" options with the
corresponding environment values into account, when determining the
up-to-date status of a session.

* The command-line tool "dump" dumps information from the cumulative
PIDE session database: many sessions may be loaded into a given logic
image, results from all loaded theories are written to the output
directory.

* Command-line tool "isabelle imports -I" also reports actual session
imports. This helps to minimize the session dependency graph.

* The command-line tool "export" and 'export_files' in session ROOT
entries retrieve theory exports from the session build database.

* The command-line tools "isabelle server" and "isabelle client" provide
access to the Isabelle Server: it supports responsive session management
and concurrent use of theories, based on Isabelle/PIDE infrastructure.
See also the "system" manual.

* The command-line tool "isabelle update_comments" normalizes formal
comments in outer syntax as follows: \<comment> \<open>text\<close> (whith a single space to
approximate the appearance in document output). This is more specific
than former "isabelle update_cartouches -c": the latter tool option has
been discontinued.

* The command-line tool "isabelle mkroot" now always produces a document
outline: its options have been adapted accordingly. INCOMPATIBILITY.

* The command-line tool "isabelle mkroot -I" initializes a Mercurial
repository for the generated session files.

* Settings ISABELLE_HEAPS + ISABELLE_BROWSER_INFO (or
ISABELLE_HEAPS_SYSTEM + ISABELLE_BROWSER_INFO_SYSTEM in "system build
mode") determine the directory locations of the main build artefacts --
instead of hard-wired directories in ISABELLE_HOME_USER (or
ISABELLE_HOME).

* Settings ISABELLE_PATH and ISABELLE_OUTPUT have been discontinued:
heap images and session databases are always stored in
$ISABELLE_HEAPS/$ML_IDENTIFIER (command-line default) or
$ISABELLE_HEAPS_SYSTEM/$ML_IDENTIFIER (main Isabelle application or
"isabelle jedit -s" or "isabelle build -s").

* ISABELLE_LATEX and ISABELLE_PDFLATEX now include platform-specific
options for improved error reporting. Potential INCOMPATIBILITY with
unusual LaTeX installations, may have to adapt these settings.

* Update to Poly/ML 5.7.1 with slightly improved performance and PIDE
markup for identifier bindings. It now uses The GNU Multiple Precision
Arithmetic Library (libgmp) on all platforms, notably Mac OS X with
32/64 bit.



New in Isabelle2017 (October 2017)
----------------------------------

*** General ***

* Experimental support for Visual Studio Code (VSCode) as alternative
Isabelle/PIDE front-end, see also
https://marketplace.visualstudio.com/items?itemName=makarius.Isabelle2017

VSCode is a new type of application that continues the concepts of
"programmer's editor" and "integrated development environment" towards
fully semantic editing and debugging -- in a relatively light-weight
manner. Thus it fits nicely on top of the Isabelle/PIDE infrastructure.
Technically, VSCode is based on the Electron application framework
(Node.js + Chromium browser + V8), which is implemented in JavaScript
and TypeScript, while Isabelle/VSCode mainly consists of Isabelle/Scala
modules around a Language Server implementation.

* Theory names are qualified by the session name that they belong to.
This affects imports, but not the theory name space prefix (which is
just the theory base name as before).

In order to import theories from other sessions, the ROOT file format
provides a new 'sessions' keyword. In contrast, a theory that is
imported in the old-fashioned manner via an explicit file-system path
belongs to the current session, and might cause theory name conflicts
later on. Theories that are imported from other sessions are excluded
from the current session document. The command-line tool "isabelle
imports" helps to update theory imports.

* The main theory entry points for some non-HOL sessions have changed,
to avoid confusion with the global name "Main" of the session HOL. This
leads to the follow renamings:

  CTT/Main.thy    ~>  CTT/CTT.thy
  ZF/Main.thy     ~>  ZF/ZF.thy
  ZF/Main_ZF.thy  ~>  ZF/ZF.thy
  ZF/Main_ZFC.thy ~>  ZF/ZFC.thy
  ZF/ZF.thy       ~>  ZF/ZF_Base.thy

INCOMPATIBILITY.

* Commands 'alias' and 'type_alias' introduce aliases for constants and
type constructors, respectively. This allows adhoc changes to name-space
accesses within global or local theory contexts, e.g. within a 'bundle'.

* Document antiquotations @{prf} and @{full_prf} output proof terms
(again) in the same way as commands 'prf' and 'full_prf'.

* Computations generated by the code generator can be embedded directly
into ML, alongside with @{code} antiquotations, using the following
antiquotations:

  @{computation ... terms: ... datatypes: ...} :
    ((term -> term) -> 'ml option -> 'a) -> Proof.context -> term -> 'a
  @{computation_conv ... terms: ... datatypes: ...} :
    (Proof.context -> 'ml -> conv) -> Proof.context -> conv
  @{computation_check terms: ... datatypes: ...} : Proof.context -> conv

See src/HOL/ex/Computations.thy,
src/HOL/Decision_Procs/Commutative_Ring.thy and
src/HOL/Decision_Procs/Reflective_Field.thy for examples and the
tutorial on code generation.


*** Prover IDE -- Isabelle/Scala/jEdit ***

* Session-qualified theory imports allow the Prover IDE to process
arbitrary theory hierarchies independently of the underlying logic
session image (e.g. option "isabelle jedit -l"), but the directory
structure needs to be known in advance (e.g. option "isabelle jedit -d"
or a line in the file $ISABELLE_HOME_USER/ROOTS).

* The PIDE document model maintains file content independently of the
status of jEdit editor buffers. Reloading jEdit buffers no longer causes
changes of formal document content. Theory dependencies are always
resolved internally, without the need for corresponding editor buffers.
The system option "jedit_auto_load" has been discontinued: it is
effectively always enabled.

* The Theories dockable provides a "Purge" button, in order to restrict
the document model to theories that are required for open editor
buffers.

* The Theories dockable indicates the overall status of checking of each
entry. When all forked tasks of a theory are finished, the border is
painted with thick lines; remaining errors in this situation are
represented by a different border color.

* Automatic indentation is more careful to avoid redundant spaces in
intermediate situations. Keywords are indented after input (via typed
characters or completion); see also option "jedit_indent_input".

* Action "isabelle.preview" opens an HTML preview of the current theory
document in the default web browser.

* Command-line invocation "isabelle jedit -R -l LOGIC" opens the ROOT
entry of the specified logic session in the editor, while its parent is
used for formal checking.

* The main Isabelle/jEdit plugin may be restarted manually (using the
jEdit Plugin Manager), as long as the "Isabelle Base" plugin remains
enabled at all times.

* Update to current jedit-5.4.0.


*** Pure ***

* Deleting the last code equations for a particular function using
[code del] results in function with no equations (runtime abort) rather
than an unimplemented function (generation time abort). Use explicit
[[code drop:]] to enforce the latter. Minor INCOMPATIBILITY.

* Proper concept of code declarations in code.ML:
  - Regular code declarations act only on the global theory level, being
    ignored with warnings if syntactically malformed.
  - Explicitly global code declarations yield errors if syntactically
    malformed.
  - Default code declarations are silently ignored if syntactically
    malformed.
Minor INCOMPATIBILITY.

* Clarified and standardized internal data bookkeeping of code
declarations: history of serials allows to track potentially
non-monotonous declarations appropriately. Minor INCOMPATIBILITY.


*** HOL ***

* The Nunchaku model finder is now part of "Main".

* SMT module:
  - A new option, 'smt_nat_as_int', has been added to translate 'nat' to
    'int' and benefit from the SMT solver's theory reasoning. It is
    disabled by default.
  - The legacy module "src/HOL/Library/Old_SMT.thy" has been removed.
  - Several small issues have been rectified in the 'smt' command.

* (Co)datatype package: The 'size_gen_o_map' lemma is no longer
generated for datatypes with type class annotations. As a result, the
tactic that derives it no longer fails on nested datatypes. Slight
INCOMPATIBILITY.

* Command and antiquotation "value" with modified default strategy:
terms without free variables are always evaluated using plain evaluation
only, with no fallback on normalization by evaluation. Minor
INCOMPATIBILITY.

* Theories "GCD" and "Binomial" are already included in "Main" (instead
of "Complex_Main").

* Constant "surj" is a full input/output abbreviation (again).
Minor INCOMPATIBILITY.

* Dropped aliasses RangeP, DomainP for Rangep, Domainp respectively.
INCOMPATIBILITY.

* Renamed ii to imaginary_unit in order to free up ii as a variable
name. The syntax \<i> remains available. INCOMPATIBILITY.

* Dropped abbreviations transP, antisymP, single_valuedP; use constants
transp, antisymp, single_valuedp instead. INCOMPATIBILITY.

* Constant "subseq" in Topological_Spaces has been removed -- it is
subsumed by "strict_mono". Some basic lemmas specific to "subseq" have
been renamed accordingly, e.g. "subseq_o" -> "strict_mono_o" etc.

* Theory List: "sublist" renamed to "nths" in analogy with "nth", and
"sublisteq" renamed to "subseq". Minor INCOMPATIBILITY.

* Theory List: new generic function "sorted_wrt".

* Named theorems mod_simps covers various congruence rules concerning
mod, replacing former zmod_simps. INCOMPATIBILITY.

* Swapped orientation of congruence rules mod_add_left_eq,
mod_add_right_eq, mod_add_eq, mod_mult_left_eq, mod_mult_right_eq,
mod_mult_eq, mod_minus_eq, mod_diff_left_eq, mod_diff_right_eq,
mod_diff_eq. INCOMPATIBILITY.

* Generalized some facts:
    measure_induct_rule
    measure_induct
    zminus_zmod ~> mod_minus_eq
    zdiff_zmod_left ~> mod_diff_left_eq
    zdiff_zmod_right ~> mod_diff_right_eq
    zmod_eq_dvd_iff ~> mod_eq_dvd_iff
INCOMPATIBILITY.

* Algebraic type class hierarchy of euclidean (semi)rings in HOL:
euclidean_(semi)ring, euclidean_(semi)ring_cancel,
unique_euclidean_(semi)ring; instantiation requires provision of a
euclidean size.

* Theory "HOL-Number_Theory.Euclidean_Algorithm" has been reworked:
  - Euclidean induction is available as rule eucl_induct.
  - Constants Euclidean_Algorithm.gcd, Euclidean_Algorithm.lcm,
    Euclidean_Algorithm.Gcd and Euclidean_Algorithm.Lcm allow
    easy instantiation of euclidean (semi)rings as GCD (semi)rings.
  - Coefficients obtained by extended euclidean algorithm are
    available as "bezout_coefficients".
INCOMPATIBILITY.

* Theory "Number_Theory.Totient" introduces basic notions about Euler's
totient function previously hidden as solitary example in theory
Residues. Definition changed so that "totient 1 = 1" in agreement with
the literature. Minor INCOMPATIBILITY.

* New styles in theory "HOL-Library.LaTeXsugar":
  - "dummy_pats" for printing equations with "_" on the lhs;
  - "eta_expand" for printing eta-expanded terms.

* Theory "HOL-Library.Permutations": theorem bij_swap_ompose_bij has
been renamed to bij_swap_compose_bij. INCOMPATIBILITY.

* New theory "HOL-Library.Going_To_Filter" providing the "f going_to F"
filter for describing points x such that f(x) is in the filter F.

* Theory "HOL-Library.Formal_Power_Series": constants X/E/L/F have been
renamed to fps_X/fps_exp/fps_ln/fps_hypergeo to avoid polluting the name
space. INCOMPATIBILITY.

* Theory "HOL-Library.FinFun" has been moved to AFP (again).
INCOMPATIBILITY.

* Theory "HOL-Library.FuncSet": some old and rarely used ASCII
replacement syntax has been removed. INCOMPATIBILITY, standard syntax
with symbols should be used instead. The subsequent commands help to
reproduce the old forms, e.g. to simplify porting old theories:

syntax (ASCII)
  "_PiE" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<Rightarrow> 'b) set"  ("(3PIE _:_./ _)" 10)
  "_Pi"  :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<Rightarrow> 'b) set"  ("(3PI _:_./ _)" 10)
  "_lam" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow> 'b)"  ("(3%_:_./ _)" [0,0,3] 3)

* Theory "HOL-Library.Multiset": the simprocs on subsets operators of
multisets have been renamed:

  msetless_cancel_numerals ~> msetsubset_cancel
  msetle_cancel_numerals ~> msetsubset_eq_cancel

INCOMPATIBILITY.

* Theory "HOL-Library.Pattern_Aliases" provides input and output syntax
for pattern aliases as known from Haskell, Scala and ML.

* Theory "HOL-Library.Uprod" formalizes the type of unordered pairs.

* Session HOL-Analysis: more material involving arcs, paths, covering
spaces, innessential maps, retracts, infinite products, simplicial
complexes. Baire Category theorem. Major results include the Jordan
Curve Theorem and the Great Picard Theorem.

* Session HOL-Algebra has been extended by additional lattice theory:
the Knaster-Tarski fixed point theorem and Galois Connections.

* Sessions HOL-Computational_Algebra and HOL-Number_Theory: new notions
of squarefreeness, n-th powers, and prime powers.

* Session "HOL-Computional_Algebra" covers many previously scattered
theories, notably Euclidean_Algorithm, Factorial_Ring,
Formal_Power_Series, Fraction_Field, Fundamental_Theorem_Algebra,
Normalized_Fraction, Polynomial_FPS, Polynomial, Primes. Minor
INCOMPATIBILITY.


*** System ***

* Isabelle/Scala: the SQL module supports access to relational
databases, either as plain file (SQLite) or full-scale server
(PostgreSQL via local port or remote ssh connection).

* Results of "isabelle build" are recorded as SQLite database (i.e.
"Application File Format" in the sense of
https://www.sqlite.org/appfileformat.html). This allows systematic
access via operations from module Sessions.Store in Isabelle/Scala.

* System option "parallel_proofs" is 1 by default (instead of more
aggressive 2). This requires less heap space and avoids burning parallel
CPU cycles, while full subproof parallelization is enabled for repeated
builds (according to parallel_subproofs_threshold).

* System option "record_proofs" allows to change the global
Proofterm.proofs variable for a session. Regular values are are 0, 1, 2;
a negative value means the current state in the ML heap image remains
unchanged.

* Isabelle settings variable ISABELLE_SCALA_BUILD_OPTIONS has been
renamed to ISABELLE_SCALAC_OPTIONS. Rare INCOMPATIBILITY.

* Isabelle settings variables ISABELLE_WINDOWS_PLATFORM,
ISABELLE_WINDOWS_PLATFORM32, ISABELLE_WINDOWS_PLATFORM64 indicate the
native Windows platform (independently of the Cygwin installation). This
is analogous to ISABELLE_PLATFORM, ISABELLE_PLATFORM32,
ISABELLE_PLATFORM64.

* Command-line tool "isabelle build_docker" builds a Docker image from
the Isabelle application bundle for Linux. See also
https://hub.docker.com/r/makarius/isabelle

* Command-line tool "isabelle vscode_server" provides a Language Server
Protocol implementation, e.g. for the Visual Studio Code editor. It
serves as example for alternative PIDE front-ends.

* Command-line tool "isabelle imports" helps to maintain theory imports
wrt. session structure. Examples for the main Isabelle distribution:

  isabelle imports -I -a
  isabelle imports -U -a
  isabelle imports -U -i -a
  isabelle imports -M -a -d '~~/src/Benchmarks'



New in Isabelle2016-1 (December 2016)
-------------------------------------

*** General ***

* Splitter in proof methods "simp", "auto" and friends:
  - The syntax "split add" has been discontinued, use plain "split",
    INCOMPATIBILITY.
  - For situations with many conditional or case expressions, there is
    an alternative splitting strategy that can be much faster. It is
    selected by writing "split!" instead of "split". It applies safe
    introduction and elimination rules after each split rule. As a
    result the subgoal may be split into several subgoals.

* Command 'bundle' provides a local theory target to define a bundle
from the body of specification commands (such as 'declare',
'declaration', 'notation', 'lemmas', 'lemma'). For example:

bundle foo
begin
  declare a [simp]
  declare b [intro]
end

* Command 'unbundle' is like 'include', but works within a local theory
context. Unlike "context includes ... begin", the effect of 'unbundle'
on the target context persists, until different declarations are given.

* Simplified outer syntax: uniform category "name" includes long
identifiers. Former "xname" / "nameref" / "name reference" has been
discontinued.

* Embedded content (e.g. the inner syntax of types, terms, props) may be
delimited uniformly via cartouches. This works better than old-fashioned
quotes when sub-languages are nested.

* Mixfix annotations support general block properties, with syntax
"(\<open>x=a y=b z \<dots>\<close>". Notable property names are "indent", "consistent",
"unbreakable", "markup". The existing notation "(DIGITS" is equivalent
to "(\<open>indent=DIGITS\<close>". The former notation "(00" for unbreakable blocks
is superseded by "(\<open>unbreabable\<close>" --- rare INCOMPATIBILITY.

* Proof method "blast" is more robust wrt. corner cases of Pure
statements without object-logic judgment.

* Commands 'prf' and 'full_prf' are somewhat more informative (again):
proof terms are reconstructed and cleaned from administrative thm nodes.

* Code generator: config option "code_timing" triggers measurements of
different phases of code generation. See src/HOL/ex/Code_Timing.thy for
examples.

* Code generator: implicits in Scala (stemming from type class
instances) are generated into companion object of corresponding type
class, to resolve some situations where ambiguities may occur.

* Solve direct: option "solve_direct_strict_warnings" gives explicit
warnings for lemma statements with trivial proofs.


*** Prover IDE -- Isabelle/Scala/jEdit ***

* More aggressive flushing of machine-generated input, according to
system option editor_generated_input_delay (in addition to existing
editor_input_delay for regular user edits). This may affect overall PIDE
reactivity and CPU usage.

* Syntactic indentation according to Isabelle outer syntax. Action
"indent-lines" (shortcut C+i) indents the current line according to
command keywords and some command substructure. Action
"isabelle.newline" (shortcut ENTER) indents the old and the new line
according to command keywords only; see also option
"jedit_indent_newline".

* Semantic indentation for unstructured proof scripts ('apply' etc.) via
number of subgoals. This requires information of ongoing document
processing and may thus lag behind, when the user is editing too
quickly; see also option "jedit_script_indent" and
"jedit_script_indent_limit".

* Refined folding mode "isabelle" based on Isar syntax: 'next' and 'qed'
are treated as delimiters for fold structure; 'begin' and 'end'
structure of theory specifications is treated as well.

* Command 'proof' provides information about proof outline with cases,
e.g. for proof methods "cases", "induct", "goal_cases".

* Completion templates for commands involving "begin ... end" blocks,
e.g. 'context', 'notepad'.

* Sidekick parser "isabelle-context" shows nesting of context blocks
according to 'begin' and 'end' structure.

* Highlighting of entity def/ref positions wrt. cursor.

* Action "isabelle.select-entity" (shortcut CS+ENTER) selects all
occurrences of the formal entity at the caret position. This facilitates
systematic renaming.

* PIDE document markup works across multiple Isar commands, e.g. the
results established at the end of a proof are properly identified in the
theorem statement.

* Cartouche abbreviations work both for " and ` to accomodate typical
situations where old ASCII notation may be updated.

* Dockable window "Symbols" also provides access to 'abbrevs' from the
outer syntax of the current theory buffer. This provides clickable
syntax templates, including entries with empty abbrevs name (which are
inaccessible via keyboard completion).

* IDE support for the Isabelle/Pure bootstrap process, with the
following independent stages:

  src/Pure/ROOT0.ML
  src/Pure/ROOT.ML
  src/Pure/Pure.thy
  src/Pure/ML_Bootstrap.thy

The ML ROOT files act like quasi-theories in the context of theory
ML_Bootstrap: this allows continuous checking of all loaded ML files.
The theory files are presented with a modified header to import Pure
from the running Isabelle instance. Results from changed versions of
each stage are *not* propagated to the next stage, and isolated from the
actual Isabelle/Pure that runs the IDE itself. The sequential
dependencies of the above files are only observed for batch build.

* Isabelle/ML and Standard ML files are presented in Sidekick with the
tree structure of section headings: this special comment format is
described in "implementation" chapter 0, e.g. (*** section ***).

* Additional abbreviations for syntactic completion may be specified
within the theory header as 'abbrevs'. The theory syntax for 'keywords'
has been simplified accordingly: optional abbrevs need to go into the
new 'abbrevs' section.

* Global abbreviations via $ISABELLE_HOME/etc/abbrevs and
$ISABELLE_HOME_USER/etc/abbrevs are no longer supported. Minor
INCOMPATIBILITY, use 'abbrevs' within theory header instead.

* Action "isabelle.keymap-merge" asks the user to resolve pending
Isabelle keymap changes that are in conflict with the current jEdit
keymap; non-conflicting changes are always applied implicitly. This
action is automatically invoked on Isabelle/jEdit startup and thus
increases chances that users see new keyboard shortcuts when re-using
old keymaps.

* ML and document antiquotations for file-systems paths are more uniform
and diverse:

  @{path NAME}   -- no file-system check
  @{file NAME}   -- check for plain file
  @{dir NAME}    -- check for directory

Minor INCOMPATIBILITY, former uses of @{file} and @{file_unchecked} may
have to be changed.


*** Document preparation ***

* New symbol \<circle>, e.g. for temporal operator.

* New document and ML antiquotation @{locale} for locales, similar to
existing antiquotation @{class}.

* Mixfix annotations support delimiters like \<^control>\<open>cartouche\<close> --
this allows special forms of document output.

* Raw LaTeX output now works via \<^latex>\<open>...\<close> instead of raw control
symbol \<^raw:...>. INCOMPATIBILITY, notably for LaTeXsugar.thy and its
derivatives.

* \<^raw:...> symbols are no longer supported.

* Old 'header' command is no longer supported (legacy since
Isabelle2015).


*** Isar ***

* Many specification elements support structured statements with 'if' /
'for' eigen-context, e.g. 'axiomatization', 'abbreviation',
'definition', 'inductive', 'function'.

* Toplevel theorem statements support eigen-context notation with 'if' /
'for' (in postfix), which corresponds to 'assumes' / 'fixes' in the
traditional long statement form (in prefix). Local premises are called
"that" or "assms", respectively. Empty premises are *not* bound in the
context: INCOMPATIBILITY.

* Command 'define' introduces a local (non-polymorphic) definition, with
optional abstraction over local parameters. The syntax resembles
'definition' and 'obtain'. It fits better into the Isar language than
old 'def', which is now a legacy feature.

* Command 'obtain' supports structured statements with 'if' / 'for'
context.

* Command '\<proof>' is an alias for 'sorry', with different
typesetting. E.g. to produce proof holes in examples and documentation.

* The defining position of a literal fact \<open>prop\<close> is maintained more
carefully, and made accessible as hyperlink in the Prover IDE.

* Commands 'finally' and 'ultimately' used to expose the result as
literal fact: this accidental behaviour has been discontinued. Rare
INCOMPATIBILITY, use more explicit means to refer to facts in Isar.

* Command 'axiomatization' has become more restrictive to correspond
better to internal axioms as singleton facts with mandatory name. Minor
INCOMPATIBILITY.

* Proof methods may refer to the main facts via the dynamic fact
"method_facts". This is particularly useful for Eisbach method
definitions.

* Proof method "use" allows to modify the main facts of a given method
expression, e.g.

  (use facts in simp)
  (use facts in \<open>simp add: ...\<close>)

* The old proof method "default" has been removed (legacy since
Isabelle2016). INCOMPATIBILITY, use "standard" instead.


*** Pure ***

* Pure provides basic versions of proof methods "simp" and "simp_all"
that only know about meta-equality (==). Potential INCOMPATIBILITY in
theory imports that merge Pure with e.g. Main of Isabelle/HOL: the order
is relevant to avoid confusion of Pure.simp vs. HOL.simp.

* The command 'unfolding' and proof method "unfold" include a second
stage where given equations are passed through the attribute "abs_def"
before rewriting. This ensures that definitions are fully expanded,
regardless of the actual parameters that are provided. Rare
INCOMPATIBILITY in some corner cases: use proof method (simp only:)
instead, or declare [[unfold_abs_def = false]] in the proof context.

* Type-inference improves sorts of newly introduced type variables for
the object-logic, using its base sort (i.e. HOL.type for Isabelle/HOL).
Thus terms like "f x" or "\<And>x. P x" without any further syntactic context
produce x::'a::type in HOL instead of x::'a::{} in Pure. Rare
INCOMPATIBILITY, need to provide explicit type constraints for Pure
types where this is really intended.


*** HOL ***

* New proof method "argo" using the built-in Argo solver based on SMT
technology. The method can be used to prove goals of quantifier-free
propositional logic, goals based on a combination of quantifier-free
propositional logic with equality, and goals based on a combination of
quantifier-free propositional logic with linear real arithmetic
including min/max/abs. See HOL/ex/Argo_Examples.thy for examples.

* The new "nunchaku" command integrates the Nunchaku model finder. The
tool is experimental. See ~~/src/HOL/Nunchaku/Nunchaku.thy for details.

* Metis: The problem encoding has changed very slightly. This might
break existing proofs. INCOMPATIBILITY.

* Sledgehammer:
  - The MaSh relevance filter is now faster than before.
  - Produce syntactically correct Vampire 4.0 problem files.

* (Co)datatype package:
  - New commands for defining corecursive functions and reasoning about
    them in "~~/src/HOL/Library/BNF_Corec.thy": 'corec', 'corecursive',
    'friend_of_corec', and 'corecursion_upto'; and 'corec_unique' proof
    method. See 'isabelle doc corec'.
  - The predicator :: ('a \<Rightarrow> bool) \<Rightarrow> 'a F \<Rightarrow> bool is now a first-class
    citizen in bounded natural functors.
  - 'primrec' now allows nested calls through the predicator in addition
    to the map function.
  - 'bnf' automatically discharges reflexive proof obligations.
  - 'bnf' outputs a slightly modified proof obligation expressing rel in
       terms of map and set
       (not giving a specification for rel makes this one reflexive).
  - 'bnf' outputs a new proof obligation expressing pred in terms of set
       (not giving a specification for pred makes this one reflexive).
    INCOMPATIBILITY: manual 'bnf' declarations may need adjustment.
  - Renamed lemmas:
      rel_prod_apply ~> rel_prod_inject
      pred_prod_apply ~> pred_prod_inject
    INCOMPATIBILITY.
  - The "size" plugin has been made compatible again with locales.
  - The theorems about "rel" and "set" may have a slightly different (but
    equivalent) form.
    INCOMPATIBILITY.

* The 'coinductive' command produces a proper coinduction rule for
mutual coinductive predicates. This new rule replaces the old rule,
which exposed details of the internal fixpoint construction and was
hard to use. INCOMPATIBILITY.

* New abbreviations for negated existence (but not bounded existence):

  \<nexists>x. P x \<equiv> \<not> (\<exists>x. P x)
  \<nexists>!x. P x \<equiv> \<not> (\<exists>!x. P x)

* The print mode "HOL" for ASCII syntax of binders "!", "?", "?!", "@"
has been removed for output. It is retained for input only, until it is
eliminated altogether.

* The unique existence quantifier no longer provides 'binder' syntax,
but uses syntax translations (as for bounded unique existence). Thus
iterated quantification \<exists>!x y. P x y with its slightly confusing
sequential meaning \<exists>!x. \<exists>!y. P x y is no longer possible. Instead,
pattern abstraction admits simultaneous unique existence \<exists>!(x, y). P x y
(analogous to existing notation \<exists>!(x, y)\<in>A. P x y). Potential
INCOMPATIBILITY in rare situations.

* Conventional syntax "%(). t" for unit abstractions. Slight syntactic
INCOMPATIBILITY.

* Renamed constants and corresponding theorems:

    setsum ~> sum
    setprod ~> prod
    listsum ~> sum_list
    listprod ~> prod_list

INCOMPATIBILITY.

* Sligthly more standardized theorem names:
    sgn_times ~> sgn_mult
    sgn_mult' ~> Real_Vector_Spaces.sgn_mult
    divide_zero_left ~> div_0
    zero_mod_left ~> mod_0
    divide_zero ~> div_by_0
    divide_1 ~> div_by_1
    nonzero_mult_divide_cancel_left ~> nonzero_mult_div_cancel_left
    div_mult_self1_is_id ~> nonzero_mult_div_cancel_left
    nonzero_mult_divide_cancel_right ~> nonzero_mult_div_cancel_right
    div_mult_self2_is_id ~> nonzero_mult_div_cancel_right
    is_unit_divide_mult_cancel_left ~> is_unit_div_mult_cancel_left
    is_unit_divide_mult_cancel_right ~> is_unit_div_mult_cancel_right
    mod_div_equality ~> div_mult_mod_eq
    mod_div_equality2 ~> mult_div_mod_eq
    mod_div_equality3 ~> mod_div_mult_eq
    mod_div_equality4 ~> mod_mult_div_eq
    minus_div_eq_mod ~> minus_div_mult_eq_mod
    minus_div_eq_mod2 ~> minus_mult_div_eq_mod
    minus_mod_eq_div ~> minus_mod_eq_div_mult
    minus_mod_eq_div2 ~> minus_mod_eq_mult_div
    div_mod_equality' ~> minus_mod_eq_div_mult [symmetric]
    mod_div_equality' ~> minus_div_mult_eq_mod [symmetric]
    zmod_zdiv_equality ~> mult_div_mod_eq [symmetric]
    zmod_zdiv_equality' ~> minus_div_mult_eq_mod [symmetric]
    Divides.mult_div_cancel ~> minus_mod_eq_mult_div [symmetric]
    mult_div_cancel ~> minus_mod_eq_mult_div [symmetric]
    zmult_div_cancel ~> minus_mod_eq_mult_div [symmetric]
    div_1 ~> div_by_Suc_0
    mod_1 ~> mod_by_Suc_0
INCOMPATIBILITY.

* New type class "idom_abs_sgn" specifies algebraic properties
of sign and absolute value functions.  Type class "sgn_if" has
disappeared.  Slight INCOMPATIBILITY.

* Dedicated syntax LENGTH('a) for length of types.

* Characters (type char) are modelled as finite algebraic type
corresponding to {0..255}.

  - Logical representation:
    * 0 is instantiated to the ASCII zero character.
    * All other characters are represented as "Char n"
      with n being a raw numeral expression less than 256.
    * Expressions of the form "Char n" with n greater than 255
      are non-canonical.
  - Printing and parsing:
    * Printable characters are printed and parsed as "CHR ''\<dots>''"
      (as before).
    * The ASCII zero character is printed and parsed as "0".
    * All other canonical characters are printed as "CHR 0xXX"
      with XX being the hexadecimal character code.  "CHR n"
      is parsable for every numeral expression n.
    * Non-canonical characters have no special syntax and are
      printed as their logical representation.
  - Explicit conversions from and to the natural numbers are
    provided as char_of_nat, nat_of_char (as before).
  - The auxiliary nibble type has been discontinued.

INCOMPATIBILITY.

* Type class "div" with operation "mod" renamed to type class "modulo"
with operation "modulo", analogously to type class "divide". This
eliminates the need to qualify any of those names in the presence of
infix "mod" syntax. INCOMPATIBILITY.

* Statements and proofs of Knaster-Tarski fixpoint combinators lfp/gfp
have been clarified. The fixpoint properties are lfp_fixpoint, its
symmetric lfp_unfold (as before), and the duals for gfp. Auxiliary items
for the proof (lfp_lemma2 etc.) are no longer exported, but can be
easily recovered by composition with eq_refl. Minor INCOMPATIBILITY.

* Constant "surj" is a mere input abbreviation, to avoid hiding an
equation in term output. Minor INCOMPATIBILITY.

* Command 'code_reflect' accepts empty constructor lists for datatypes,
which renders those abstract effectively.

* Command 'export_code' checks given constants for abstraction
violations: a small guarantee that given constants specify a safe
interface for the generated code.

* Code generation for Scala: ambiguous implicts in class diagrams are
spelt out explicitly.

* Static evaluators (Code_Evaluation.static_* in Isabelle/ML) rely on
explicitly provided auxiliary definitions for required type class
dictionaries rather than half-working magic. INCOMPATIBILITY, see the
tutorial on code generation for details.

* Theory Set_Interval: substantial new theorems on indexed sums and
products.

* Locale bijection establishes convenient default simp rules such as
"inv f (f a) = a" for total bijections.

* Abstract locales semigroup, abel_semigroup, semilattice,
semilattice_neutr, ordering, ordering_top, semilattice_order,
semilattice_neutr_order, comm_monoid_set, semilattice_set,
semilattice_neutr_set, semilattice_order_set,
semilattice_order_neutr_set monoid_list, comm_monoid_list,
comm_monoid_list_set, comm_monoid_mset, comm_monoid_fun use boldified
syntax uniformly that does not clash with corresponding global syntax.
INCOMPATIBILITY.

* Former locale lifting_syntax is now a bundle, which is easier to
include in a local context or theorem statement, e.g. "context includes
lifting_syntax begin ... end". Minor INCOMPATIBILITY.

* Some old / obsolete theorems have been renamed / removed, potential
INCOMPATIBILITY.

  nat_less_cases  --  removed, use linorder_cases instead
  inv_image_comp  --  removed, use image_inv_f_f instead
  image_surj_f_inv_f  ~>  image_f_inv_f

* Some theorems about groups and orders have been generalised from
  groups to semi-groups that are also monoids:
    le_add_same_cancel1
    le_add_same_cancel2
    less_add_same_cancel1
    less_add_same_cancel2
    add_le_same_cancel1
    add_le_same_cancel2
    add_less_same_cancel1
    add_less_same_cancel2

* Some simplifications theorems about rings have been removed, since
  superseeded by a more general version:
    less_add_cancel_left_greater_zero ~> less_add_same_cancel1
    less_add_cancel_right_greater_zero ~> less_add_same_cancel2
    less_eq_add_cancel_left_greater_eq_zero ~> le_add_same_cancel1
    less_eq_add_cancel_right_greater_eq_zero ~> le_add_same_cancel2
    less_eq_add_cancel_left_less_eq_zero ~> add_le_same_cancel1
    less_eq_add_cancel_right_less_eq_zero ~> add_le_same_cancel2
    less_add_cancel_left_less_zero ~> add_less_same_cancel1
    less_add_cancel_right_less_zero ~> add_less_same_cancel2
INCOMPATIBILITY.

* Renamed split_if -> if_split and split_if_asm -> if_split_asm to
resemble the f.split naming convention, INCOMPATIBILITY.

* Added class topological_monoid.

* The following theorems have been renamed:

  setsum_left_distrib ~> sum_distrib_right
  setsum_right_distrib ~> sum_distrib_left

INCOMPATIBILITY.

* Compound constants INFIMUM and SUPREMUM are mere abbreviations now.
INCOMPATIBILITY.

* "Gcd (f ` A)" and "Lcm (f ` A)" are printed with optional
comprehension-like syntax analogously to "Inf (f ` A)" and "Sup (f `
A)".

* Class semiring_Lcd merged into semiring_Gcd. INCOMPATIBILITY.

* The type class ordered_comm_monoid_add is now called
ordered_cancel_comm_monoid_add. A new type class ordered_comm_monoid_add
is introduced as the combination of ordered_ab_semigroup_add +
comm_monoid_add. INCOMPATIBILITY.

* Introduced the type classes canonically_ordered_comm_monoid_add and
dioid.

* Introduced the type class ordered_ab_semigroup_monoid_add_imp_le. When
instantiating linordered_semiring_strict and ordered_ab_group_add, an
explicit instantiation of ordered_ab_semigroup_monoid_add_imp_le might
be required. INCOMPATIBILITY.

* Dropped various legacy fact bindings, whose replacements are often
of a more general type also:
  lcm_left_commute_nat ~> lcm.left_commute
  lcm_left_commute_int ~> lcm.left_commute
  gcd_left_commute_nat ~> gcd.left_commute
  gcd_left_commute_int ~> gcd.left_commute
  gcd_greatest_iff_nat ~> gcd_greatest_iff
  gcd_greatest_iff_int ~> gcd_greatest_iff
  coprime_dvd_mult_nat ~> coprime_dvd_mult
  coprime_dvd_mult_int ~> coprime_dvd_mult
  zpower_numeral_even ~> power_numeral_even
  gcd_mult_cancel_nat ~> gcd_mult_cancel
  gcd_mult_cancel_int ~> gcd_mult_cancel
  div_gcd_coprime_nat ~> div_gcd_coprime
  div_gcd_coprime_int ~> div_gcd_coprime
  zpower_numeral_odd ~> power_numeral_odd
  zero_less_int_conv ~> of_nat_0_less_iff
  gcd_greatest_nat ~> gcd_greatest
  gcd_greatest_int ~> gcd_greatest
  coprime_mult_nat ~> coprime_mult
  coprime_mult_int ~> coprime_mult
  lcm_commute_nat ~> lcm.commute
  lcm_commute_int ~> lcm.commute
  int_less_0_conv ~> of_nat_less_0_iff
  gcd_commute_nat ~> gcd.commute
  gcd_commute_int ~> gcd.commute
  Gcd_insert_nat ~> Gcd_insert
  Gcd_insert_int ~> Gcd_insert
  of_int_int_eq ~> of_int_of_nat_eq
  lcm_least_nat ~> lcm_least
  lcm_least_int ~> lcm_least
  lcm_assoc_nat ~> lcm.assoc
  lcm_assoc_int ~> lcm.assoc
  int_le_0_conv ~> of_nat_le_0_iff
  int_eq_0_conv ~> of_nat_eq_0_iff
  Gcd_empty_nat ~> Gcd_empty
  Gcd_empty_int ~> Gcd_empty
  gcd_assoc_nat ~> gcd.assoc
  gcd_assoc_int ~> gcd.assoc
  zero_zle_int ~> of_nat_0_le_iff
  lcm_dvd2_nat ~> dvd_lcm2
  lcm_dvd2_int ~> dvd_lcm2
  lcm_dvd1_nat ~> dvd_lcm1
  lcm_dvd1_int ~> dvd_lcm1
  gcd_zero_nat ~> gcd_eq_0_iff
  gcd_zero_int ~> gcd_eq_0_iff
  gcd_dvd2_nat ~> gcd_dvd2
  gcd_dvd2_int ~> gcd_dvd2
  gcd_dvd1_nat ~> gcd_dvd1
  gcd_dvd1_int ~> gcd_dvd1
  int_numeral ~> of_nat_numeral
  lcm_ac_nat ~> ac_simps
  lcm_ac_int ~> ac_simps
  gcd_ac_nat ~> ac_simps
  gcd_ac_int ~> ac_simps
  abs_int_eq ~> abs_of_nat
  zless_int ~> of_nat_less_iff
  zdiff_int ~> of_nat_diff
  zadd_int ~> of_nat_add
  int_mult ~> of_nat_mult
  int_Suc ~> of_nat_Suc
  inj_int ~> inj_of_nat
  int_1 ~> of_nat_1
  int_0 ~> of_nat_0
  Lcm_empty_nat ~> Lcm_empty
  Lcm_empty_int ~> Lcm_empty
  Lcm_insert_nat ~> Lcm_insert
  Lcm_insert_int ~> Lcm_insert
  comp_fun_idem_gcd_nat ~> comp_fun_idem_gcd
  comp_fun_idem_gcd_int ~> comp_fun_idem_gcd
  comp_fun_idem_lcm_nat ~> comp_fun_idem_lcm
  comp_fun_idem_lcm_int ~> comp_fun_idem_lcm
  Lcm_eq_0 ~> Lcm_eq_0_I
  Lcm0_iff ~> Lcm_0_iff
  Lcm_dvd_int ~> Lcm_least
  divides_mult_nat ~> divides_mult
  divides_mult_int ~> divides_mult
  lcm_0_nat ~> lcm_0_right
  lcm_0_int ~> lcm_0_right
  lcm_0_left_nat ~> lcm_0_left
  lcm_0_left_int ~> lcm_0_left
  dvd_gcd_D1_nat ~> dvd_gcdD1
  dvd_gcd_D1_int ~> dvd_gcdD1
  dvd_gcd_D2_nat ~> dvd_gcdD2
  dvd_gcd_D2_int ~> dvd_gcdD2
  coprime_dvd_mult_iff_nat ~> coprime_dvd_mult_iff
  coprime_dvd_mult_iff_int ~> coprime_dvd_mult_iff
  realpow_minus_mult ~> power_minus_mult
  realpow_Suc_le_self ~> power_Suc_le_self
  dvd_Gcd, dvd_Gcd_nat, dvd_Gcd_int removed in favour of Gcd_greatest
INCOMPATIBILITY.

* Renamed HOL/Quotient_Examples/FSet.thy to
HOL/Quotient_Examples/Quotient_FSet.thy INCOMPATIBILITY.

* Session HOL-Library: theory FinFun bundles "finfun_syntax" and
"no_finfun_syntax" allow to control optional syntax in local contexts;
this supersedes former theory FinFun_Syntax. INCOMPATIBILITY, e.g. use
"unbundle finfun_syntax" to imitate import of
"~~/src/HOL/Library/FinFun_Syntax".

* Session HOL-Library: theory Multiset_Permutations (executably) defines
the set of permutations of a given set or multiset, i.e. the set of all
lists that contain every element of the carrier (multi-)set exactly
once.

* Session HOL-Library: multiset membership is now expressed using
set_mset rather than count.

  - Expressions "count M a > 0" and similar simplify to membership
    by default.

  - Converting between "count M a = 0" and non-membership happens using
    equations count_eq_zero_iff and not_in_iff.

  - Rules count_inI and in_countE obtain facts of the form
    "count M a = n" from membership.

  - Rules count_in_diffI and in_diff_countE obtain facts of the form
    "count M a = n + count N a" from membership on difference sets.

INCOMPATIBILITY.

* Session HOL-Library: theory LaTeXsugar uses new-style "dummy_pats" for
displaying equations in functional programming style --- variables
present on the left-hand but not on the righ-hand side are replaced by
underscores.

* Session HOL-Library: theory Combinator_PER provides combinator to
build partial equivalence relations from a predicate and an equivalence
relation.

* Session HOL-Library: theory Perm provides basic facts about almost
everywhere fix bijections.

* Session HOL-Library: theory Normalized_Fraction allows viewing an
element of a field of fractions as a normalized fraction (i.e. a pair of
numerator and denominator such that the two are coprime and the
denominator is normalized wrt. unit factors).

* Session HOL-NSA has been renamed to HOL-Nonstandard_Analysis.

* Session HOL-Multivariate_Analysis has been renamed to HOL-Analysis.

* Session HOL-Analysis: measure theory has been moved here from
HOL-Probability. When importing HOL-Analysis some theorems need
additional name spaces prefixes due to name clashes. INCOMPATIBILITY.

* Session HOL-Analysis: more complex analysis including Cauchy's
inequality, Liouville theorem, open mapping theorem, maximum modulus
principle, Residue theorem, Schwarz Lemma.

* Session HOL-Analysis: Theory of polyhedra: faces, extreme points,
polytopes, and the Krein–Milman Minkowski theorem.

* Session HOL-Analysis: Numerous results ported from the HOL Light
libraries: homeomorphisms, continuous function extensions, invariance of
domain.

* Session HOL-Probability: the type of emeasure and nn_integral was
changed from ereal to ennreal, INCOMPATIBILITY.

  emeasure :: 'a measure \<Rightarrow> 'a set \<Rightarrow> ennreal
  nn_integral :: 'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> ennreal

* Session HOL-Probability: Code generation and QuickCheck for
Probability Mass Functions.

* Session HOL-Probability: theory Random_Permutations contains some
theory about choosing a permutation of a set uniformly at random and
folding over a list in random order.

* Session HOL-Probability: theory SPMF formalises discrete
subprobability distributions.

* Session HOL-Library: the names of multiset theorems have been
normalised to distinguish which ordering the theorems are about

    mset_less_eqI ~> mset_subset_eqI
    mset_less_insertD ~> mset_subset_insertD
    mset_less_eq_count ~> mset_subset_eq_count
    mset_less_diff_self ~> mset_subset_diff_self
    mset_le_exists_conv ~> mset_subset_eq_exists_conv
    mset_le_mono_add_right_cancel ~> mset_subset_eq_mono_add_right_cancel
    mset_le_mono_add_left_cancel ~> mset_subset_eq_mono_add_left_cancel
    mset_le_mono_add ~> mset_subset_eq_mono_add
    mset_le_add_left ~> mset_subset_eq_add_left
    mset_le_add_right ~> mset_subset_eq_add_right
    mset_le_single ~> mset_subset_eq_single
    mset_le_multiset_union_diff_commute ~> mset_subset_eq_multiset_union_diff_commute
    diff_le_self ~> diff_subset_eq_self
    mset_leD ~> mset_subset_eqD
    mset_lessD ~> mset_subsetD
    mset_le_insertD ~> mset_subset_eq_insertD
    mset_less_of_empty ~> mset_subset_of_empty
    mset_less_size ~> mset_subset_size
    wf_less_mset_rel ~> wf_subset_mset_rel
    count_le_replicate_mset_le ~> count_le_replicate_mset_subset_eq
    mset_remdups_le ~> mset_remdups_subset_eq
    ms_lesseq_impl ~> subset_eq_mset_impl

Some functions have been renamed:
    ms_lesseq_impl -> subset_eq_mset_impl

* HOL-Library: multisets are now ordered with the multiset ordering
    #\<subseteq># ~> \<le>
    #\<subset># ~> <
    le_multiset ~> less_eq_multiset
    less_multiset ~> le_multiset
INCOMPATIBILITY.

* Session HOL-Library: the prefix multiset_order has been discontinued:
the theorems can be directly accessed. As a consequence, the lemmas
"order_multiset" and "linorder_multiset" have been discontinued, and the
interpretations "multiset_linorder" and "multiset_wellorder" have been
replaced by instantiations. INCOMPATIBILITY.

* Session HOL-Library: some theorems about the multiset ordering have
been renamed:

    le_multiset_def ~> less_eq_multiset_def
    less_multiset_def ~> le_multiset_def
    less_eq_imp_le_multiset ~> subset_eq_imp_le_multiset
    mult_less_not_refl ~> mset_le_not_refl
    mult_less_trans ~> mset_le_trans
    mult_less_not_sym ~> mset_le_not_sym
    mult_less_asym ~> mset_le_asym
    mult_less_irrefl ~> mset_le_irrefl
    union_less_mono2{,1,2} ~> union_le_mono2{,1,2}

    le_multiset\<^sub>H\<^sub>O ~> less_eq_multiset\<^sub>H\<^sub>O
    le_multiset_total ~> less_eq_multiset_total
    less_multiset_right_total ~> subset_eq_imp_le_multiset
    le_multiset_empty_left ~> less_eq_multiset_empty_left
    le_multiset_empty_right ~> less_eq_multiset_empty_right
    less_multiset_empty_right ~> le_multiset_empty_left
    less_multiset_empty_left ~> le_multiset_empty_right
    union_less_diff_plus ~> union_le_diff_plus
    ex_gt_count_imp_less_multiset ~> ex_gt_count_imp_le_multiset
    less_multiset_plus_left_nonempty ~> le_multiset_plus_left_nonempty
    le_multiset_plus_right_nonempty ~> le_multiset_plus_right_nonempty
INCOMPATIBILITY.

* Session HOL-Library: the lemma mset_map has now the attribute [simp].
INCOMPATIBILITY.

* Session HOL-Library: some theorems about multisets have been removed.
INCOMPATIBILITY, use the following replacements:

    le_multiset_plus_plus_left_iff ~> add_less_cancel_right
    less_multiset_plus_plus_left_iff ~> add_less_cancel_right
    le_multiset_plus_plus_right_iff ~> add_less_cancel_left
    less_multiset_plus_plus_right_iff ~> add_less_cancel_left
    add_eq_self_empty_iff ~> add_cancel_left_right
    mset_subset_add_bothsides ~> subset_mset.add_less_cancel_right
    mset_less_add_bothsides ~> subset_mset.add_less_cancel_right
    mset_le_add_bothsides ~> subset_mset.add_less_cancel_right
    empty_inter ~> subset_mset.inf_bot_left
    inter_empty ~> subset_mset.inf_bot_right
    empty_sup ~> subset_mset.sup_bot_left
    sup_empty ~> subset_mset.sup_bot_right
    bdd_below_multiset ~> subset_mset.bdd_above_bot
    subset_eq_empty ~> subset_mset.le_zero_eq
    le_empty ~> subset_mset.le_zero_eq
    mset_subset_empty_nonempty ~> subset_mset.zero_less_iff_neq_zero
    mset_less_empty_nonempty ~> subset_mset.zero_less_iff_neq_zero

* Session HOL-Library: some typeclass constraints about multisets have
been reduced from ordered or linordered to preorder. Multisets have the
additional typeclasses order_bot, no_top,
ordered_ab_semigroup_add_imp_le, ordered_cancel_comm_monoid_add,
linordered_cancel_ab_semigroup_add, and
ordered_ab_semigroup_monoid_add_imp_le. INCOMPATIBILITY.

* Session HOL-Library: there are some new simplification rules about
multisets, the multiset ordering, and the subset ordering on multisets.
INCOMPATIBILITY.

* Session HOL-Library: the subset ordering on multisets has now the
interpretations ordered_ab_semigroup_monoid_add_imp_le and
bounded_lattice_bot. INCOMPATIBILITY.

* Session HOL-Library, theory Multiset: single has been removed in favor
of add_mset that roughly corresponds to Set.insert. Some theorems have
removed or changed:

  single_not_empty ~> add_mset_not_empty or empty_not_add_mset
  fold_mset_insert ~> fold_mset_add_mset
  image_mset_insert ~> image_mset_add_mset
  union_single_eq_diff
  multi_self_add_other_not_self
  diff_single_eq_union
INCOMPATIBILITY.

* Session HOL-Library, theory Multiset: some theorems have been changed
to use add_mset instead of single:

  mset_add
  multi_self_add_other_not_self
  diff_single_eq_union
  union_single_eq_diff
  union_single_eq_member
  add_eq_conv_diff
  insert_noteq_member
  add_eq_conv_ex
  multi_member_split
  multiset_add_sub_el_shuffle
  mset_subset_eq_insertD
  mset_subset_insertD
  insert_subset_eq_iff
  insert_union_subset_iff
  multi_psub_of_add_self
  inter_add_left1
  inter_add_left2
  inter_add_right1
  inter_add_right2
  sup_union_left1
  sup_union_left2
  sup_union_right1
  sup_union_right2
  size_eq_Suc_imp_eq_union
  multi_nonempty_split
  mset_insort
  mset_update
  mult1I
  less_add
  mset_zip_take_Cons_drop_twice
  rel_mset_Zero
  msed_map_invL
  msed_map_invR
  msed_rel_invL
  msed_rel_invR
  le_multiset_right_total
  multiset_induct
  multiset_induct2_size
  multiset_induct2
INCOMPATIBILITY.

* Session HOL-Library, theory Multiset: the definitions of some
constants have changed to use add_mset instead of adding a single
element:

  image_mset
  mset
  replicate_mset
  mult1
  pred_mset
  rel_mset'
  mset_insort

INCOMPATIBILITY.

* Session HOL-Library, theory Multiset: due to the above changes, the
attributes of some multiset theorems have been changed:

  insert_DiffM  [] ~> [simp]
  insert_DiffM2 [simp] ~> []
  diff_add_mset_swap [simp]
  fold_mset_add_mset [simp]
  diff_diff_add [simp] (for multisets only)
  diff_cancel [simp] ~> []
  count_single [simp] ~> []
  set_mset_single [simp] ~> []
  size_multiset_single [simp] ~> []
  size_single [simp] ~> []
  image_mset_single [simp] ~> []
  mset_subset_eq_mono_add_right_cancel [simp] ~> []
  mset_subset_eq_mono_add_left_cancel [simp] ~> []
  fold_mset_single [simp] ~> []
  subset_eq_empty [simp] ~> []
  empty_sup [simp] ~> []
  sup_empty [simp] ~> []
  inter_empty [simp] ~> []
  empty_inter [simp] ~> []
INCOMPATIBILITY.

* Session HOL-Library, theory Multiset: the order of the variables in
the second cases of multiset_induct, multiset_induct2_size,
multiset_induct2 has been changed (e.g. Add A a ~> Add a A).
INCOMPATIBILITY.

* Session HOL-Library, theory Multiset: there is now a simplification
procedure on multisets. It mimics the behavior of the procedure on
natural numbers. INCOMPATIBILITY.

* Session HOL-Library, theory Multiset: renamed sums and products of
multisets:

  msetsum ~> sum_mset
  msetprod ~> prod_mset

* Session HOL-Library, theory Multiset: the notation for intersection
and union of multisets have been changed:

  #\<inter> ~> \<inter>#
  #\<union> ~> \<union>#

INCOMPATIBILITY.

* Session HOL-Library, theory Multiset: the lemma
one_step_implies_mult_aux on multisets has been removed, use
one_step_implies_mult instead. INCOMPATIBILITY.

* Session HOL-Library: theory Complete_Partial_Order2 provides reasoning
support for monotonicity and continuity in chain-complete partial orders
and about admissibility conditions for fixpoint inductions.

* Session HOL-Library: theory Library/Polynomial contains also
derivation of polynomials (formerly in Library/Poly_Deriv) but not
gcd/lcm on polynomials over fields. This has been moved to a separate
theory Library/Polynomial_GCD_euclidean.thy, to pave way for a possible
future different type class instantiation for polynomials over factorial
rings. INCOMPATIBILITY.

* Session HOL-Library: theory Sublist provides function "prefixes" with
the following renaming

  prefixeq -> prefix
  prefix -> strict_prefix
  suffixeq -> suffix
  suffix -> strict_suffix

Added theory of longest common prefixes.

* Session HOL-Number_Theory: algebraic foundation for primes:
Generalisation of predicate "prime" and introduction of predicates
"prime_elem", "irreducible", a "prime_factorization" function, and the
"factorial_ring" typeclass with instance proofs for nat, int, poly. Some
theorems now have different names, most notably "prime_def" is now
"prime_nat_iff". INCOMPATIBILITY.

* Session Old_Number_Theory has been removed, after porting remaining
theories.

* Session HOL-Types_To_Sets provides an experimental extension of
Higher-Order Logic to allow translation of types to sets.


*** ML ***

* Integer.gcd and Integer.lcm use efficient operations from the Poly/ML
library (notably for big integers). Subtle change of semantics:
Integer.gcd and Integer.lcm both normalize the sign, results are never
negative. This coincides with the definitions in HOL/GCD.thy.
INCOMPATIBILITY.

* Structure Rat for rational numbers is now an integral part of
Isabelle/ML, with special notation @int/nat or @int for numerals (an
abbreviation for antiquotation @{Pure.rat argument}) and ML pretty
printing. Standard operations on type Rat.rat are provided via ad-hoc
overloading of + - * / < <= > >= ~ abs. INCOMPATIBILITY, need to
use + instead of +/ etc. Moreover, exception Rat.DIVZERO has been
superseded by General.Div.

* ML antiquotation @{path} is superseded by @{file}, which ensures that
the argument is a plain file. Minor INCOMPATIBILITY.

* Antiquotation @{make_string} is available during Pure bootstrap --
with approximative output quality.

* Low-level ML system structures (like PolyML and RunCall) are no longer
exposed to Isabelle/ML user-space. Potential INCOMPATIBILITY.

* The ML function "ML" provides easy access to run-time compilation.
This is particularly useful for conditional compilation, without
requiring separate files.

* Option ML_exception_debugger controls detailed exception trace via the
Poly/ML debugger. Relevant ML modules need to be compiled beforehand
with ML_file_debug, or with ML_file and option ML_debugger enabled. Note
debugger information requires consirable time and space: main
Isabelle/HOL with full debugger support may need ML_system_64.

* Local_Theory.restore has been renamed to Local_Theory.reset to
emphasize its disruptive impact on the cumulative context, notably the
scope of 'private' or 'qualified' names. Note that Local_Theory.reset is
only appropriate when targets are managed, e.g. starting from a global
theory and returning to it. Regular definitional packages should use
balanced blocks of Local_Theory.open_target versus
Local_Theory.close_target instead. Rare INCOMPATIBILITY.

* Structure TimeLimit (originally from the SML/NJ library) has been
replaced by structure Timeout, with slightly different signature.
INCOMPATIBILITY.

* Discontinued cd and pwd operations, which are not well-defined in a
multi-threaded environment. Note that files are usually located
relatively to the master directory of a theory (see also
File.full_path). Potential INCOMPATIBILITY.

* Binding.empty_atts supersedes Thm.empty_binding and
Attrib.empty_binding. Minor INCOMPATIBILITY.


*** System ***

* SML/NJ and old versions of Poly/ML are no longer supported.

* Poly/ML heaps now follow the hierarchy of sessions, and thus require
much less disk space.

* The Isabelle ML process is now managed directly by Isabelle/Scala, and
shell scripts merely provide optional command-line access. In
particular:

  . Scala module ML_Process to connect to the raw ML process,
    with interaction via stdin/stdout/stderr or in batch mode;
  . command-line tool "isabelle console" as interactive wrapper;
  . command-line tool "isabelle process" as batch mode wrapper.

* The executable "isabelle_process" has been discontinued. Tools and
prover front-ends should use ML_Process or Isabelle_Process in
Isabelle/Scala. INCOMPATIBILITY.

* New command-line tool "isabelle process" supports ML evaluation of
literal expressions (option -e) or files (option -f) in the context of a
given heap image. Errors lead to premature exit of the ML process with
return code 1.

* The command-line tool "isabelle build" supports option -N for cyclic
shuffling of NUMA CPU nodes. This may help performance tuning on Linux
servers with separate CPU/memory modules.

* System option "threads" (for the size of the Isabelle/ML thread farm)
is also passed to the underlying ML runtime system as --gcthreads,
unless there is already a default provided via ML_OPTIONS settings.

* System option "checkpoint" helps to fine-tune the global heap space
management of isabelle build. This is relevant for big sessions that may
exhaust the small 32-bit address space of the ML process (which is used
by default).

* System option "profiling" specifies the mode for global ML profiling
in "isabelle build". Possible values are "time", "allocations". The
command-line tool "isabelle profiling_report" helps to digest the
resulting log files.

* System option "ML_process_policy" specifies an optional command prefix
for the underlying ML process, e.g. to control CPU affinity on
multiprocessor systems. The "isabelle jedit" tool allows to override the
implicit default via option -p.

* Command-line tool "isabelle console" provides option -r to help to
bootstrapping Isabelle/Pure interactively.

* Command-line tool "isabelle yxml" has been discontinued.
INCOMPATIBILITY, use operations from the modules "XML" and "YXML" in
Isabelle/ML or Isabelle/Scala.

* Many Isabelle tools that require a Java runtime system refer to the
settings ISABELLE_TOOL_JAVA_OPTIONS32 / ISABELLE_TOOL_JAVA_OPTIONS64,
depending on the underlying platform. The settings for "isabelle build"
ISABELLE_BUILD_JAVA_OPTIONS32 / ISABELLE_BUILD_JAVA_OPTIONS64 have been
discontinued. Potential INCOMPATIBILITY.

* The Isabelle system environment always ensures that the main
executables are found within the shell search $PATH: "isabelle" and
"isabelle_scala_script".

* Isabelle tools may consist of .scala files: the Scala compiler is
invoked on the spot. The source needs to define some object that extends
Isabelle_Tool.Body.

* File.bash_string, File.bash_path etc. represent Isabelle/ML and
Isabelle/Scala strings authentically within GNU bash. This is useful to
produce robust shell scripts under program control, without worrying
about spaces or special characters. Note that user output works via
Path.print (ML) or Path.toString (Scala). INCOMPATIBILITY, the old (and
less versatile) operations File.shell_quote, File.shell_path etc. have
been discontinued.

* The isabelle_java executable allows to run a Java process within the
name space of Java and Scala components that are bundled with Isabelle,
but without the Isabelle settings environment.

* Isabelle/Scala: the SSH module supports ssh and sftp connections, for
remote command-execution and file-system access. This resembles
operations from module File and Isabelle_System to some extent. Note
that Path specifications need to be resolved remotely via
ssh.remote_path instead of File.standard_path: the implicit process
environment is different, Isabelle settings are not available remotely.

* Isabelle/Scala: the Mercurial module supports repositories via the
regular hg command-line interface. The repositroy clone and working
directory may reside on a local or remote file-system (via ssh
connection).



New in Isabelle2016 (February 2016)
-----------------------------------

*** General ***

* Eisbach is now based on Pure instead of HOL. Objects-logics may import
either the theory ~~/src/HOL/Eisbach/Eisbach (for HOL etc.) or
~~/src/HOL/Eisbach/Eisbach_Old_Appl_Syntax (for FOL, ZF etc.). Note that
the HOL-Eisbach session located in ~~/src/HOL/Eisbach/ contains further
examples that do require HOL.

* Better resource usage on all platforms (Linux, Windows, Mac OS X) for
both Isabelle/ML and Isabelle/Scala.  Slightly reduced heap space usage.

* Former "xsymbols" syntax with Isabelle symbols is used by default,
without any special print mode. Important ASCII replacement syntax
remains available under print mode "ASCII", but less important syntax
has been removed (see below).

* Support for more arrow symbols, with rendering in LaTeX and Isabelle
fonts: \<Lleftarrow> \<Rrightarrow> \<longlongleftarrow> \<longlongrightarrow> \<longlonglongleftarrow> \<longlonglongrightarrow>.

* Special notation \<struct> for the first implicit 'structure' in the
context has been discontinued. Rare INCOMPATIBILITY, use explicit
structure name instead, notably in indexed notation with block-subscript
(e.g. \<odot>\<^bsub>A\<^esub>).

* The glyph for \<diamond> in the IsabelleText font now corresponds better to its
counterpart \<box> as quantifier-like symbol. A small diamond is available as
\<diamondop>; the old symbol \<struct> loses this rendering and any special
meaning.

* Syntax for formal comments "-- text" now also supports the symbolic
form "\<comment> text". Command-line tool "isabelle update_cartouches -c" helps
to update old sources.

* Toplevel theorem statements have been simplified as follows:

  theorems             ~>  lemmas
  schematic_lemma      ~>  schematic_goal
  schematic_theorem    ~>  schematic_goal
  schematic_corollary  ~>  schematic_goal

Command-line tool "isabelle update_theorems" updates theory sources
accordingly.

* Toplevel theorem statement 'proposition' is another alias for
'theorem'.

* The old 'defs' command has been removed (legacy since Isabelle2014).
INCOMPATIBILITY, use regular 'definition' instead. Overloaded and/or
deferred definitions require a surrounding 'overloading' block.


*** Prover IDE -- Isabelle/Scala/jEdit ***

* IDE support for the source-level debugger of Poly/ML, to work with
Isabelle/ML and official Standard ML. Option "ML_debugger" and commands
'ML_file_debug', 'ML_file_no_debug', 'SML_file_debug',
'SML_file_no_debug' control compilation of sources with or without
debugging information. The Debugger panel allows to set breakpoints (via
context menu), step through stopped threads, evaluate local ML
expressions etc. At least one Debugger view needs to be active to have
any effect on the running ML program.

* The State panel manages explicit proof state output, with dynamic
auto-update according to cursor movement. Alternatively, the jEdit
action "isabelle.update-state" (shortcut S+ENTER) triggers manual
update.

* The Output panel no longer shows proof state output by default, to
avoid GUI overcrowding. INCOMPATIBILITY, use the State panel instead or
enable option "editor_output_state".

* The text overview column (status of errors, warnings etc.) is updated
asynchronously, leading to much better editor reactivity. Moreover, the
full document node content is taken into account. The width of the
column is scaled according to the main text area font, for improved
visibility.

* The main text area no longer changes its color hue in outdated
situations. The text overview column takes over the role to indicate
unfinished edits in the PIDE pipeline. This avoids flashing text display
due to ad-hoc updates by auxiliary GUI components, such as the State
panel.

* Slightly improved scheduling for urgent print tasks (e.g. command
state output, interactive queries) wrt. long-running background tasks.

* Completion of symbols via prefix of \<name> or \<^name> or \name is
always possible, independently of the language context. It is never
implicit: a popup will show up unconditionally.

* Additional abbreviations for syntactic completion may be specified in
$ISABELLE_HOME/etc/abbrevs and $ISABELLE_HOME_USER/etc/abbrevs, with
support for simple templates using ASCII 007 (bell) as placeholder.

* Symbols \<oplus>, \<Oplus>, \<otimes>, \<Otimes>, \<odot>, \<Odot>, \<ominus>, \<oslash> no longer provide abbreviations for
completion like "+o", "*o", ".o" etc. -- due to conflicts with other
ASCII syntax. INCOMPATIBILITY, use plain backslash-completion or define
suitable abbreviations in $ISABELLE_HOME_USER/etc/abbrevs.

* Action "isabelle-emph" (with keyboard shortcut C+e LEFT) controls
emphasized text style; the effect is visible in document output, not in
the editor.

* Action "isabelle-reset" now uses keyboard shortcut C+e BACK_SPACE,
instead of former C+e LEFT.

* The command-line tool "isabelle jedit" and the isabelle.Main
application wrapper treat the default $USER_HOME/Scratch.thy more
uniformly, and allow the dummy file argument ":" to open an empty buffer
instead.

* New command-line tool "isabelle jedit_client" allows to connect to an
already running Isabelle/jEdit process. This achieves the effect of
single-instance applications seen on common GUI desktops.

* The default look-and-feel for Linux is the traditional "Metal", which
works better with GUI scaling for very high-resolution displays (e.g.
4K). Moreover, it is generally more robust than "Nimbus".

* Update to jedit-5.3.0, with improved GUI scaling and support of
high-resolution displays (e.g. 4K).

* The main Isabelle executable is managed as single-instance Desktop
application uniformly on all platforms: Linux, Windows, Mac OS X.


*** Document preparation ***

* Commands 'paragraph' and 'subparagraph' provide additional section
headings. Thus there are 6 levels of standard headings, as in HTML.

* Command 'text_raw' has been clarified: input text is processed as in
'text' (with antiquotations and control symbols). The key difference is
the lack of the surrounding isabelle markup environment in output.

* Text is structured in paragraphs and nested lists, using notation that
is similar to Markdown. The control symbols for list items are as
follows:

  \<^item>  itemize
  \<^enum>  enumerate
  \<^descr>  description

* There is a new short form for antiquotations with a single argument
that is a cartouche: \<^name>\<open>...\<close> is equivalent to @{name \<open>...\<close>} and
\<open>...\<close> without control symbol is equivalent to @{cartouche \<open>...\<close>}.
\<^name> without following cartouche is equivalent to @{name}. The
standard Isabelle fonts provide glyphs to render important control
symbols, e.g. "\<^verbatim>", "\<^emph>", "\<^bold>".

* Antiquotations @{noindent}, @{smallskip}, @{medskip}, @{bigskip} with
corresponding control symbols \<^noindent>, \<^smallskip>, \<^medskip>, \<^bigskip> specify spacing formally, using
standard LaTeX macros of the same names.

* Antiquotation @{cartouche} in Isabelle/Pure is the same as @{text}.
Consequently, \<open>...\<close> without any decoration prints literal quasi-formal
text. Command-line tool "isabelle update_cartouches -t" helps to update
old sources, by approximative patching of the content of string and
cartouche tokens seen in theory sources.

* The @{text} antiquotation now ignores the antiquotation option
"source". The given text content is output unconditionally, without any
surrounding quotes etc. Subtle INCOMPATIBILITY, put quotes into the
argument where they are really intended, e.g. @{text \<open>"foo"\<close>}. Initial
or terminal spaces are ignored.

* Antiquotations @{emph} and @{bold} output LaTeX source recursively,
adding appropriate text style markup. These may be used in the short
form \<^emph>\<open>...\<close> and \<^bold>\<open>...\<close>.

* Document antiquotation @{footnote} outputs LaTeX source recursively,
marked as \footnote{}. This may be used in the short form \<^footnote>\<open>...\<close>.

* Antiquotation @{verbatim [display]} supports option "indent".

* Antiquotation @{theory_text} prints uninterpreted theory source text
(Isar outer syntax with command keywords etc.). This may be used in the
short form \<^theory_text>\<open>...\<close>. @{theory_text [display]} supports option "indent".

* Antiquotation @{doc ENTRY} provides a reference to the given
documentation, with a hyperlink in the Prover IDE.

* Antiquotations @{command}, @{method}, @{attribute} print checked
entities of the Isar language.

* HTML presentation uses the standard IsabelleText font and Unicode
rendering of Isabelle symbols like Isabelle/Scala/jEdit.  The former
print mode "HTML" loses its special meaning.


*** Isar ***

* Local goals ('have', 'show', 'hence', 'thus') allow structured rule
statements like fixes/assumes/shows in theorem specifications, but the
notation is postfix with keywords 'if' (or 'when') and 'for'. For
example:

  have result: "C x y"
    if "A x" and "B y"
    for x :: 'a and y :: 'a
    <proof>

The local assumptions are bound to the name "that". The result is
exported from context of the statement as usual. The above roughly
corresponds to a raw proof block like this:

  {
    fix x :: 'a and y :: 'a
    assume that: "A x" "B y"
    have "C x y" <proof>
  }
  note result = this

The keyword 'when' may be used instead of 'if', to indicate 'presume'
instead of 'assume' above.

* Assumptions ('assume', 'presume') allow structured rule statements
using 'if' and 'for', similar to 'have' etc. above. For example:

  assume result: "C x y"
    if "A x" and "B y"
    for x :: 'a and y :: 'a

This assumes "\<And>x y::'a. A x \<Longrightarrow> B y \<Longrightarrow> C x y" and produces a general
result as usual: "A ?x \<Longrightarrow> B ?y \<Longrightarrow> C ?x ?y".

Vacuous quantification in assumptions is omitted, i.e. a for-context
only effects propositions according to actual use of variables. For
example:

  assume "A x" and "B y" for x and y

is equivalent to:

  assume "\<And>x. A x" and "\<And>y. B y"

* The meaning of 'show' with Pure rule statements has changed: premises
are treated in the sense of 'assume', instead of 'presume'. This means,
a goal like "\<And>x. A x \<Longrightarrow> B x \<Longrightarrow> C x" can be solved completely as
follows:

  show "\<And>x. A x \<Longrightarrow> B x \<Longrightarrow> C x"

or:

  show "C x" if "A x" "B x" for x

Rare INCOMPATIBILITY, the old behaviour may be recovered as follows:

  show "C x" when "A x" "B x" for x

* New command 'consider' states rules for generalized elimination and
case splitting. This is like a toplevel statement "theorem obtains" used
within a proof body; or like a multi-branch 'obtain' without activation
of the local context elements yet.

* Proof method "cases" allows to specify the rule as first entry of
chained facts.  This is particularly useful with 'consider':

  consider (a) A | (b) B | (c) C <proof>
  then have something
  proof cases
    case a
    then show ?thesis <proof>
  next
    case b
    then show ?thesis <proof>
  next
    case c
    then show ?thesis <proof>
  qed

* Command 'case' allows fact name and attribute specification like this:

  case a: (c xs)
  case a [attributes]: (c xs)

Facts that are introduced by invoking the case context are uniformly
qualified by "a"; the same name is used for the cumulative fact. The old
form "case (c xs) [attributes]" is no longer supported. Rare
INCOMPATIBILITY, need to adapt uses of case facts in exotic situations,
and always put attributes in front.

* The standard proof method of commands 'proof' and '..' is now called
"standard" to make semantically clear what it is; the old name "default"
is still available as legacy for some time. Documentation now explains
'..' more accurately as "by standard" instead of "by rule".

* Nesting of Isar goal structure has been clarified: the context after
the initial backwards refinement is retained for the whole proof, within
all its context sections (as indicated via 'next'). This is e.g.
relevant for 'using', 'including', 'supply':

  have "A \<and> A" if a: A for A
    supply [simp] = a
  proof
    show A by simp
  next
    show A by simp
  qed

* Command 'obtain' binds term abbreviations (via 'is' patterns) in the
proof body as well, abstracted over relevant parameters.

* Improved type-inference for theorem statement 'obtains': separate
parameter scope for of each clause.

* Term abbreviations via 'is' patterns also work for schematic
statements: result is abstracted over unknowns.

* Command 'subgoal' allows to impose some structure on backward
refinements, to avoid proof scripts degenerating into long of 'apply'
sequences. Further explanations and examples are given in the isar-ref
manual.

* Command 'supply' supports fact definitions during goal refinement
('apply' scripts).

* Proof method "goal_cases" turns the current subgoals into cases within
the context; the conclusion is bound to variable ?case in each case. For
example:

lemma "\<And>x. A x \<Longrightarrow> B x \<Longrightarrow> C x"
  and "\<And>y z. U y \<Longrightarrow> V z \<Longrightarrow> W y z"
proof goal_cases
  case (1 x)
  then show ?case using \<open>A x\<close> \<open>B x\<close> sorry
next
  case (2 y z)
  then show ?case using \<open>U y\<close> \<open>V z\<close> sorry
qed

lemma "\<And>x. A x \<Longrightarrow> B x \<Longrightarrow> C x"
  and "\<And>y z. U y \<Longrightarrow> V z \<Longrightarrow> W y z"
proof goal_cases
  case prems: 1
  then show ?case using prems sorry
next
  case prems: 2
  then show ?case using prems sorry
qed

* The undocumented feature of implicit cases goal1, goal2, goal3, etc.
is marked as legacy, and will be removed eventually. The proof method
"goals" achieves a similar effect within regular Isar; often it can be
done more adequately by other means (e.g. 'consider').

* The vacuous fact "TERM x" may be established "by fact" or as `TERM x`
as well, not just "by this" or "." as before.

* Method "sleep" succeeds after a real-time delay (in seconds). This is
occasionally useful for demonstration and testing purposes.


*** Pure ***

* Qualifiers in locale expressions default to mandatory ('!') regardless
of the command. Previously, for 'locale' and 'sublocale' the default was
optional ('?'). The old synatx '!' has been discontinued.
INCOMPATIBILITY, remove '!' and add '?' as required.

* Keyword 'rewrites' identifies rewrite morphisms in interpretation
commands. Previously, the keyword was 'where'. INCOMPATIBILITY.

* More gentle suppression of syntax along locale morphisms while
printing terms. Previously 'abbreviation' and 'notation' declarations
would be suppressed for morphisms except term identity. Now
'abbreviation' is also kept for morphims that only change the involved
parameters, and only 'notation' is suppressed. This can be of great help
when working with complex locale hierarchies, because proof states are
displayed much more succinctly. It also means that only notation needs
to be redeclared if desired, as illustrated by this example:

  locale struct = fixes composition :: "'a => 'a => 'a" (infixl "\<cdot>" 65)
  begin
    definition derived (infixl "\<odot>" 65) where ...
  end

  locale morphism =
    left: struct composition + right: struct composition'
    for composition (infix "\<cdot>" 65) and composition' (infix "\<cdot>''" 65)
  begin
    notation right.derived ("\<odot>''")
  end

* Command 'global_interpretation' issues interpretations into global
theories, with optional rewrite definitions following keyword 'defines'.

* Command 'sublocale' accepts optional rewrite definitions after keyword
'defines'.

* Command 'permanent_interpretation' has been discontinued. Use
'global_interpretation' or 'sublocale' instead. INCOMPATIBILITY.

* Command 'print_definitions' prints dependencies of definitional
specifications. This functionality used to be part of 'print_theory'.

* Configuration option rule_insts_schematic has been discontinued
(intermediate legacy feature in Isabelle2015). INCOMPATIBILITY.

* Abbreviations in type classes now carry proper sort constraint. Rare
INCOMPATIBILITY in situations where the previous misbehaviour has been
exploited.

* Refinement of user-space type system in type classes: pseudo-local
operations behave more similar to abbreviations. Potential
INCOMPATIBILITY in exotic situations.


*** HOL ***

* The 'typedef' command has been upgraded from a partially checked
"axiomatization", to a full definitional specification that takes the
global collection of overloaded constant / type definitions into
account. Type definitions with open dependencies on overloaded
definitions need to be specified as "typedef (overloaded)". This
provides extra robustness in theory construction. Rare INCOMPATIBILITY.

* Qualification of various formal entities in the libraries is done more
uniformly via "context begin qualified definition ... end" instead of
old-style "hide_const (open) ...". Consequently, both the defined
constant and its defining fact become qualified, e.g. Option.is_none and
Option.is_none_def. Occasional INCOMPATIBILITY in applications.

* Some old and rarely used ASCII replacement syntax has been removed.
INCOMPATIBILITY, standard syntax with symbols should be used instead.
The subsequent commands help to reproduce the old forms, e.g. to
simplify porting old theories:

  notation iff  (infixr "<->" 25)

  notation Times  (infixr "<*>" 80)

  type_notation Map.map  (infixr "~=>" 0)
  notation Map.map_comp  (infixl "o'_m" 55)

  type_notation FinFun.finfun ("(_ =>f /_)" [22, 21] 21)

  notation FuncSet.funcset  (infixr "->" 60)
  notation FuncSet.extensional_funcset  (infixr "->\<^sub>E" 60)

  notation Omega_Words_Fun.conc (infixr "conc" 65)

  notation Preorder.equiv ("op ~~")
    and Preorder.equiv ("(_/ ~~ _)" [51, 51] 50)

  notation (in topological_space) tendsto (infixr "--->" 55)
  notation (in topological_space) LIMSEQ ("((_)/ ----> (_))" [60, 60] 60)
  notation LIM ("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60)

  notation NSA.approx (infixl "@=" 50)
  notation NSLIMSEQ ("((_)/ ----NS> (_))" [60, 60] 60)
  notation NSLIM ("((_)/ -- (_)/ --NS> (_))" [60, 0, 60] 60)

* The alternative notation "\<Colon>" for type and sort constraints has been
removed: in LaTeX document output it looks the same as "::".
INCOMPATIBILITY, use plain "::" instead.

* Commands 'inductive' and 'inductive_set' work better when names for
intro rules are omitted: the "cases" and "induct" rules no longer
declare empty case_names, but no case_names at all. This allows to use
numbered cases in proofs, without requiring method "goal_cases".

* Inductive definitions ('inductive', 'coinductive', etc.) expose
low-level facts of the internal construction only if the option
"inductive_internals" is enabled. This refers to the internal predicate
definition and its monotonicity result. Rare INCOMPATIBILITY.

* Recursive function definitions ('fun', 'function', 'partial_function')
expose low-level facts of the internal construction only if the option
"function_internals" is enabled. Its internal inductive definition is
also subject to "inductive_internals". Rare INCOMPATIBILITY.

* BNF datatypes ('datatype', 'codatatype', etc.) expose low-level facts
of the internal construction only if the option "bnf_internals" is
enabled. This supersedes the former option "bnf_note_all". Rare
INCOMPATIBILITY.

* Combinator to represent case distinction on products is named
"case_prod", uniformly, discontinuing any input aliasses. Very popular
theorem aliasses have been retained.

Consolidated facts:
  PairE ~> prod.exhaust
  Pair_eq ~> prod.inject
  pair_collapse ~> prod.collapse
  Pair_fst_snd_eq ~> prod_eq_iff
  split_twice ~> prod.case_distrib
  split_weak_cong ~> prod.case_cong_weak
  split_split ~> prod.split
  split_split_asm ~> prod.split_asm
  splitI ~> case_prodI
  splitD ~> case_prodD
  splitI2 ~> case_prodI2
  splitI2' ~> case_prodI2'
  splitE ~> case_prodE
  splitE' ~> case_prodE'
  split_pair ~> case_prod_Pair
  split_eta ~> case_prod_eta
  split_comp ~> case_prod_comp
  mem_splitI ~> mem_case_prodI
  mem_splitI2 ~> mem_case_prodI2
  mem_splitE ~> mem_case_prodE
  The_split ~> The_case_prod
  cond_split_eta ~> cond_case_prod_eta
  Collect_split_in_rel_leE ~> Collect_case_prod_in_rel_leE
  Collect_split_in_rel_leI ~> Collect_case_prod_in_rel_leI
  in_rel_Collect_split_eq ~> in_rel_Collect_case_prod_eq
  Collect_split_Grp_eqD ~> Collect_case_prod_Grp_eqD
  Collect_split_Grp_inD ~> Collect_case_prod_Grp_in
  Domain_Collect_split ~> Domain_Collect_case_prod
  Image_Collect_split ~> Image_Collect_case_prod
  Range_Collect_split ~> Range_Collect_case_prod
  Eps_split ~> Eps_case_prod
  Eps_split_eq ~> Eps_case_prod_eq
  split_rsp ~> case_prod_rsp
  curry_split ~> curry_case_prod
  split_curry ~> case_prod_curry

Changes in structure HOLogic:
  split_const ~> case_prod_const
  mk_split ~> mk_case_prod
  mk_psplits ~> mk_ptupleabs
  strip_psplits ~> strip_ptupleabs

INCOMPATIBILITY.

* The coercions to type 'real' have been reorganised. The function
'real' is no longer overloaded, but has type 'nat => real' and
abbreviates of_nat for that type. Also 'real_of_int :: int => real'
abbreviates of_int for that type. Other overloaded instances of 'real'
have been replaced by 'real_of_ereal' and 'real_of_float'.

Consolidated facts (among others):
  real_of_nat_le_iff -> of_nat_le_iff
  real_of_nat_numeral of_nat_numeral
  real_of_int_zero of_int_0
  real_of_nat_zero of_nat_0
  real_of_one of_int_1
  real_of_int_add of_int_add
  real_of_nat_add of_nat_add
  real_of_int_diff of_int_diff
  real_of_nat_diff of_nat_diff
  floor_subtract floor_diff_of_int
  real_of_int_inject of_int_eq_iff
  real_of_int_gt_zero_cancel_iff of_int_0_less_iff
  real_of_int_ge_zero_cancel_iff of_int_0_le_iff
  real_of_nat_ge_zero of_nat_0_le_iff
  real_of_int_ceiling_ge le_of_int_ceiling
  ceiling_less_eq ceiling_less_iff
  ceiling_le_eq ceiling_le_iff
  less_floor_eq less_floor_iff
  floor_less_eq floor_less_iff
  floor_divide_eq_div floor_divide_of_int_eq
  real_of_int_zero_cancel of_nat_eq_0_iff
  ceiling_real_of_int ceiling_of_int

INCOMPATIBILITY.

* Theory Map: lemma map_of_is_SomeD was a clone of map_of_SomeD and has
been removed. INCOMPATIBILITY.

* Quickcheck setup for finite sets.

* Discontinued simp_legacy_precond. Potential INCOMPATIBILITY.

* Sledgehammer:
  - The MaSh relevance filter has been sped up.
  - Proof reconstruction has been improved, to minimize the incidence of
    cases where Sledgehammer gives a proof that does not work.
  - Auto Sledgehammer now minimizes and preplays the results.
  - Handle Vampire 4.0 proof output without raising exception.
  - Eliminated "MASH" environment variable. Use the "MaSh" option in
    Isabelle/jEdit instead. INCOMPATIBILITY.
  - Eliminated obsolete "blocking" option and related subcommands.

* Nitpick:
  - Fixed soundness bug in translation of "finite" predicate.
  - Fixed soundness bug in "destroy_constrs" optimization.
  - Fixed soundness bug in translation of "rat" type.
  - Removed "check_potential" and "check_genuine" options.
  - Eliminated obsolete "blocking" option.

* (Co)datatype package:
  - New commands "lift_bnf" and "copy_bnf" for lifting (copying) a BNF
    structure on the raw type to an abstract type defined using typedef.
  - Always generate "case_transfer" theorem.
  - For mutual types, generate slightly stronger "rel_induct",
    "rel_coinduct", and "coinduct" theorems. INCOMPATIBILITY.
  - Allow discriminators and selectors with the same name as the type
    being defined.
  - Avoid various internal name clashes (e.g., 'datatype f = f').

* Transfer: new methods for interactive debugging of 'transfer' and
'transfer_prover': 'transfer_start', 'transfer_step', 'transfer_end',
'transfer_prover_start' and 'transfer_prover_end'.

* New diagnostic command print_record for displaying record definitions.

* Division on integers is bootstrapped directly from division on
naturals and uses generic numeral algorithm for computations. Slight
INCOMPATIBILITY, simproc numeral_divmod replaces and generalizes former
simprocs binary_int_div and binary_int_mod

* Tightened specification of class semiring_no_zero_divisors. Minor
INCOMPATIBILITY.

* Class algebraic_semidom introduces common algebraic notions of
integral (semi)domains, particularly units. Although logically subsumed
by fields, is is not a super class of these in order not to burden
fields with notions that are trivial there.

* Class normalization_semidom specifies canonical representants for
equivalence classes of associated elements in an integral (semi)domain.
This formalizes associated elements as well.

* Abstract specification of gcd/lcm operations in classes semiring_gcd,
semiring_Gcd, semiring_Lcd. Minor INCOMPATIBILITY: facts gcd_nat.commute
and gcd_int.commute are subsumed by gcd.commute, as well as
gcd_nat.assoc and gcd_int.assoc by gcd.assoc.

* Former constants Fields.divide (_ / _) and Divides.div (_ div _) are
logically unified to Rings.divide in syntactic type class Rings.divide,
with infix syntax (_ div _). Infix syntax (_ / _) for field division is
added later as abbreviation in class Fields.inverse. INCOMPATIBILITY,
instantiations must refer to Rings.divide rather than the former
separate constants, hence infix syntax (_ / _) is usually not available
during instantiation.

* New cancellation simprocs for boolean algebras to cancel complementary
terms for sup and inf. For example, "sup x (sup y (- x))" simplifies to
"top". INCOMPATIBILITY.

* Class uniform_space introduces uniform spaces btw topological spaces
and metric spaces. Minor INCOMPATIBILITY: open_<type>_def needs to be
introduced in the form of an uniformity. Some constants are more general
now, it may be necessary to add type class constraints.

  open_real_def \<leadsto> open_dist
  open_complex_def \<leadsto> open_dist

* Library/Monad_Syntax: notation uses symbols \<bind> and \<then>. INCOMPATIBILITY.

* Library/Multiset:
  - Renamed multiset inclusion operators:
      < ~> <#
      > ~> >#
      <= ~> <=#
      >= ~> >=#
      \<le> ~> \<le>#
      \<ge> ~> \<ge>#
    INCOMPATIBILITY.
  - Added multiset inclusion operator syntax:
      \<subset>#
      \<subseteq>#
      \<supset>#
      \<supseteq>#
  - "'a multiset" is no longer an instance of the "order",
    "ordered_ab_semigroup_add_imp_le", "ordered_cancel_comm_monoid_diff",
    "semilattice_inf", and "semilattice_sup" type classes. The theorems
    previously provided by these type classes (directly or indirectly)
    are now available through the "subset_mset" interpretation
    (e.g. add_mono ~> subset_mset.add_mono).
    INCOMPATIBILITY.
  - Renamed conversions:
      multiset_of ~> mset
      multiset_of_set ~> mset_set
      set_of ~> set_mset
    INCOMPATIBILITY
  - Renamed lemmas:
      mset_le_def ~> subseteq_mset_def
      mset_less_def ~> subset_mset_def
      less_eq_multiset.rep_eq ~> subseteq_mset_def
    INCOMPATIBILITY
  - Removed lemmas generated by lift_definition:
    less_eq_multiset.abs_eq, less_eq_multiset.rsp,
    less_eq_multiset.transfer, less_eq_multiset_def
    INCOMPATIBILITY

* Library/Omega_Words_Fun: Infinite words modeled as functions nat \<Rightarrow> 'a.

* Library/Bourbaki_Witt_Fixpoint: Added formalisation of the
Bourbaki-Witt fixpoint theorem for increasing functions in
chain-complete partial orders.

* Library/Old_Recdef: discontinued obsolete 'defer_recdef' command.
Minor INCOMPATIBILITY, use 'function' instead.

* Library/Periodic_Fun: a locale that provides convenient lemmas for
periodic functions.

* Library/Formal_Power_Series: proper definition of division (with
remainder) for formal power series; instances for Euclidean Ring and
GCD.

* HOL-Imperative_HOL: obsolete theory Legacy_Mrec has been removed.

* HOL-Statespace: command 'statespace' uses mandatory qualifier for
import of parent, as for general 'locale' expressions. INCOMPATIBILITY,
remove '!' and add '?' as required.

* HOL-Decision_Procs: The "approximation" method works with "powr"
(exponentiation on real numbers) again.

* HOL-Multivariate_Analysis: theory Cauchy_Integral_Thm with Contour
integrals (= complex path integrals), Cauchy's integral theorem, winding
numbers and Cauchy's integral formula, Liouville theorem, Fundamental
Theorem of Algebra. Ported from HOL Light.

* HOL-Multivariate_Analysis: topological concepts such as connected
components, homotopic paths and the inside or outside of a set.

* HOL-Multivariate_Analysis: radius of convergence of power series and
various summability tests; Harmonic numbers and the Euler–Mascheroni
constant; the Generalised Binomial Theorem; the complex and real
Gamma/log-Gamma/Digamma/ Polygamma functions and their most important
properties.

* HOL-Probability: The central limit theorem based on Levy's uniqueness
and continuity theorems, weak convergence, and characterisitc functions.

* HOL-Data_Structures: new and growing session of standard data
structures.


*** ML ***

* The following combinators for low-level profiling of the ML runtime
system are available:

  profile_time          (*CPU time*)
  profile_time_thread   (*CPU time on this thread*)
  profile_allocations   (*overall heap allocations*)

* Antiquotation @{undefined} or \<^undefined> inlines (raise Match).

* Antiquotation @{method NAME} inlines the (checked) name of the given
Isar proof method.

* Pretty printing of Poly/ML compiler output in Isabelle has been
improved: proper treatment of break offsets and blocks with consistent
breaks.

* The auxiliary module Pure/display.ML has been eliminated. Its
elementary thm print operations are now in Pure/more_thm.ML and thus
called Thm.pretty_thm, Thm.string_of_thm etc. INCOMPATIBILITY.

* Simproc programming interfaces have been simplified:
Simplifier.make_simproc and Simplifier.define_simproc supersede various
forms of Simplifier.mk_simproc, Simplifier.simproc_global etc. Note that
term patterns for the left-hand sides are specified with implicitly
fixed variables, like top-level theorem statements. INCOMPATIBILITY.

* Instantiation rules have been re-organized as follows:

  Thm.instantiate  (*low-level instantiation with named arguments*)
  Thm.instantiate' (*version with positional arguments*)

  Drule.infer_instantiate  (*instantiation with type inference*)
  Drule.infer_instantiate'  (*version with positional arguments*)

The LHS only requires variable specifications, instead of full terms.
Old cterm_instantiate is superseded by infer_instantiate.
INCOMPATIBILITY, need to re-adjust some ML names and types accordingly.

* Old tactic shorthands atac, rtac, etac, dtac, ftac have been
discontinued. INCOMPATIBILITY, use regular assume_tac, resolve_tac etc.
instead (with proper context).

* Thm.instantiate (and derivatives) no longer require the LHS of the
instantiation to be certified: plain variables are given directly.

* Subgoal.SUBPROOF and Subgoal.FOCUS combinators use anonymous
quasi-bound variables (like the Simplifier), instead of accidentally
named local fixes. This has the potential to improve stability of proof
tools, but can also cause INCOMPATIBILITY for tools that don't observe
the proof context discipline.

* Isar proof methods are based on a slightly more general type
context_tactic, which allows to change the proof context dynamically
(e.g. to update cases) and indicate explicit Seq.Error results. Former
METHOD_CASES is superseded by CONTEXT_METHOD; further combinators are
provided in src/Pure/Isar/method.ML for convenience. INCOMPATIBILITY.


*** System ***

* Command-line tool "isabelle console" enables print mode "ASCII".

* Command-line tool "isabelle update_then" expands old Isar command
conflations:

    hence  ~>  then have
    thus   ~>  then show

This syntax is more orthogonal and improves readability and
maintainability of proofs.

* Global session timeout is multiplied by timeout_scale factor. This
allows to adjust large-scale tests (e.g. AFP) to overall hardware
performance.

* Property values in etc/symbols may contain spaces, if written with the
replacement character "␣" (Unicode point 0x2324). For example:

    \<star>  code: 0x0022c6  group: operator  font: Deja␣Vu␣Sans␣Mono

* Java runtime environment for x86_64-windows allows to use larger heap
space.

* Java runtime options are determined separately for 32bit vs. 64bit
platforms as follows.

  - Isabelle desktop application: platform-specific files that are
    associated with the main app bundle

  - isabelle jedit: settings
    JEDIT_JAVA_SYSTEM_OPTIONS
    JEDIT_JAVA_OPTIONS32 vs. JEDIT_JAVA_OPTIONS64

  - isabelle build: settings
    ISABELLE_BUILD_JAVA_OPTIONS32 vs. ISABELLE_BUILD_JAVA_OPTIONS64

* Bash shell function "jvmpath" has been renamed to "platform_path": it
is relevant both for Poly/ML and JVM processes.

* Poly/ML default platform architecture may be changed from 32bit to
64bit via system option ML_system_64. A system restart (and rebuild) is
required after change.

* Poly/ML 5.6 runs natively on x86-windows and x86_64-windows, which
both allow larger heap space than former x86-cygwin.

* Heap images are 10-15% smaller due to less wasteful persistent theory
content (using ML type theory_id instead of theory);



New in Isabelle2015 (May 2015)
------------------------------

*** General ***

* Local theory specification commands may have a 'private' or
'qualified' modifier to restrict name space accesses to the local scope,
as provided by some "context begin ... end" block. For example:

  context
  begin

  private definition ...
  private lemma ...

  qualified definition ...
  qualified lemma ...

  lemma ...
  theorem ...

  end

* Command 'experiment' opens an anonymous locale context with private
naming policy.

* Command 'notepad' requires proper nesting of begin/end and its proof
structure in the body: 'oops' is no longer supported here. Minor
INCOMPATIBILITY, use 'sorry' instead.

* Command 'named_theorems' declares a dynamic fact within the context,
together with an attribute to maintain the content incrementally. This
supersedes functor Named_Thms in Isabelle/ML, but with a subtle change
of semantics due to external visual order vs. internal reverse order.

* 'find_theorems': search patterns which are abstractions are
schematically expanded before search. Search results match the naive
expectation more closely, particularly wrt. abbreviations.
INCOMPATIBILITY.

* Commands 'method_setup' and 'attribute_setup' now work within a local
theory context.

* Outer syntax commands are managed authentically within the theory
context, without implicit global state. Potential for accidental
INCOMPATIBILITY, make sure that required theories are really imported.

* Historical command-line terminator ";" is no longer accepted (and
already used differently in Isar). Minor INCOMPATIBILITY, use "isabelle
update_semicolons" to remove obsolete semicolons from old theory
sources.

* Structural composition of proof methods (meth1; meth2) in Isar
corresponds to (tac1 THEN_ALL_NEW tac2) in ML.

* The Eisbach proof method language allows to define new proof methods
by combining existing ones with their usual syntax. The "match" proof
method provides basic fact/term matching in addition to
premise/conclusion matching through Subgoal.focus, and binds fact names
from matches as well as term patterns within matches. The Isabelle
documentation provides an entry "eisbach" for the Eisbach User Manual.
Sources and various examples are in ~~/src/HOL/Eisbach/.


*** Prover IDE -- Isabelle/Scala/jEdit ***

* Improved folding mode "isabelle" based on Isar syntax. Alternatively,
the "sidekick" mode may be used for document structure.

* Extended bracket matching based on Isar language structure. System
option jedit_structure_limit determines maximum number of lines to scan
in the buffer.

* Support for BibTeX files: context menu, context-sensitive token
marker, SideKick parser.

* Document antiquotation @{cite} provides formal markup, which is
interpreted semi-formally based on .bib files that happen to be open in
the editor (hyperlinks, completion etc.).

* Less waste of vertical space via negative line spacing (see Global
Options / Text Area).

* Improved graphview panel with optional output of PNG or PDF, for
display of 'thy_deps', 'class_deps' etc.

* The commands 'thy_deps' and 'class_deps' allow optional bounds to
restrict the visualized hierarchy.

* Improved scheduling for asynchronous print commands (e.g. provers
managed by the Sledgehammer panel) wrt. ongoing document processing.


*** Document preparation ***

* Document markup commands 'chapter', 'section', 'subsection',
'subsubsection', 'text', 'txt', 'text_raw' work uniformly in any
context, even before the initial 'theory' command. Obsolete proof
commands 'sect', 'subsect', 'subsubsect', 'txt_raw' have been
discontinued, use 'section', 'subsection', 'subsubsection', 'text_raw'
instead. The old 'header' command is still retained for some time, but
should be replaced by 'chapter', 'section' etc. (using "isabelle
update_header"). Minor INCOMPATIBILITY.

* Official support for "tt" style variants, via \isatt{...} or
\begin{isabellett}...\end{isabellett}. The somewhat fragile \verb or
verbatim environment of LaTeX is no longer used. This allows @{ML} etc.
as argument to other macros (such as footnotes).

* Document antiquotation @{verbatim} prints ASCII text literally in "tt"
style.

* Discontinued obsolete option "document_graph": session_graph.pdf is
produced unconditionally for HTML browser_info and PDF-LaTeX document.

* Diagnostic commands and document markup commands within a proof do not
affect the command tag for output. Thus commands like 'thm' are subject
to proof document structure, and no longer "stick out" accidentally.
Commands 'text' and 'txt' merely differ in the LaTeX style, not their
tags. Potential INCOMPATIBILITY in exotic situations.

* System option "pretty_margin" is superseded by "thy_output_margin",
which is also accessible via document antiquotation option "margin".
Only the margin for document output may be changed, but not the global
pretty printing: that is 76 for plain console output, and adapted
dynamically in GUI front-ends. Implementations of document
antiquotations need to observe the margin explicitly according to
Thy_Output.string_of_margin. Minor INCOMPATIBILITY.

* Specification of 'document_files' in the session ROOT file is
mandatory for document preparation. The legacy mode with implicit
copying of the document/ directory is no longer supported. Minor
INCOMPATIBILITY.


*** Pure ***

* Proof methods with explicit instantiation ("rule_tac", "subgoal_tac"
etc.) allow an optional context of local variables ('for' declaration):
these variables become schematic in the instantiated theorem; this
behaviour is analogous to 'for' in attributes "where" and "of".
Configuration option rule_insts_schematic (default false) controls use
of schematic variables outside the context. Minor INCOMPATIBILITY,
declare rule_insts_schematic = true temporarily and update to use local
variable declarations or dummy patterns instead.

* Explicit instantiation via attributes "where", "of", and proof methods
"rule_tac" with derivatives like "subgoal_tac" etc. admit dummy patterns
("_") that stand for anonymous local variables.

* Generated schematic variables in standard format of exported facts are
incremented to avoid material in the proof context. Rare